| author | wenzelm | 
| Mon, 29 Jul 2019 11:09:37 +0200 | |
| changeset 70436 | 251f1fb44ccd | 
| parent 70365 | 4df0628e8545 | 
| child 70532 | fcf3b891ccb1 | 
| permissions | -rw-r--r-- | 
| 52265 | 1 | (* Title: HOL/Limits.thy | 
| 51526 | 2 | Author: Brian Huffman | 
| 3 | Author: Jacques D. Fleuriot, University of Cambridge | |
| 4 | Author: Lawrence C Paulson | |
| 5 | Author: Jeremy Avigad | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 6 | *) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 7 | |
| 60758 | 8 | section \<open>Limits on Real Vector Spaces\<close> | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 9 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 10 | theory Limits | 
| 63546 | 11 | imports Real_Vector_Spaces | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 12 | begin | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 13 | |
| 60758 | 14 | subsection \<open>Filter going to infinity norm\<close> | 
| 51526 | 15 | |
| 63546 | 16 | definition at_infinity :: "'a::real_normed_vector filter" | 
| 17 |   where "at_infinity = (INF r. principal {x. r \<le> norm x})"
 | |
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 18 | |
| 57276 | 19 | lemma eventually_at_infinity: "eventually P at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. b \<le> norm x \<longrightarrow> P x)" | 
| 20 | unfolding at_infinity_def | |
| 21 | by (subst eventually_INF_base) | |
| 22 | (auto simp: subset_eq eventually_principal intro!: exI[of _ "max a b" for a b]) | |
| 31392 | 23 | |
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62369diff
changeset | 24 | corollary eventually_at_infinity_pos: | 
| 63546 | 25 | "eventually p at_infinity \<longleftrightarrow> (\<exists>b. 0 < b \<and> (\<forall>x. norm x \<ge> b \<longrightarrow> p x))" | 
| 68614 | 26 | unfolding eventually_at_infinity | 
| 27 | by (meson le_less_trans norm_ge_zero not_le zero_less_one) | |
| 63546 | 28 | |
| 29 | lemma at_infinity_eq_at_top_bot: "(at_infinity :: real filter) = sup at_top at_bot" | |
| 68614 | 30 | proof - | 
| 31 | have 1: "\<lbrakk>\<forall>n\<ge>u. A n; \<forall>n\<le>v. A n\<rbrakk> | |
| 32 | \<Longrightarrow> \<exists>b. \<forall>x. b \<le> \<bar>x\<bar> \<longrightarrow> A x" for A and u v::real | |
| 33 | by (rule_tac x="max (- v) u" in exI) (auto simp: abs_real_def) | |
| 34 | have 2: "\<forall>x. u \<le> \<bar>x\<bar> \<longrightarrow> A x \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. A n" for A and u::real | |
| 35 | by (meson abs_less_iff le_cases less_le_not_le) | |
| 36 | have 3: "\<forall>x. u \<le> \<bar>x\<bar> \<longrightarrow> A x \<Longrightarrow> \<exists>N. \<forall>n\<le>N. A n" for A and u::real | |
| 37 | by (metis (full_types) abs_ge_self abs_minus_cancel le_minus_iff order_trans) | |
| 38 | show ?thesis | |
| 68615 | 39 | by (auto simp: filter_eq_iff eventually_sup eventually_at_infinity | 
| 68614 | 40 | eventually_at_top_linorder eventually_at_bot_linorder intro: 1 2 3) | 
| 41 | qed | |
| 50325 | 42 | |
| 57276 | 43 | lemma at_top_le_at_infinity: "at_top \<le> (at_infinity :: real filter)" | 
| 50325 | 44 | unfolding at_infinity_eq_at_top_bot by simp | 
| 45 | ||
| 57276 | 46 | lemma at_bot_le_at_infinity: "at_bot \<le> (at_infinity :: real filter)" | 
| 50325 | 47 | unfolding at_infinity_eq_at_top_bot by simp | 
| 48 | ||
| 63546 | 49 | lemma filterlim_at_top_imp_at_infinity: "filterlim f at_top F \<Longrightarrow> filterlim f at_infinity F" | 
| 50 | for f :: "_ \<Rightarrow> real" | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56541diff
changeset | 51 | by (rule filterlim_mono[OF _ at_top_le_at_infinity order_refl]) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56541diff
changeset | 52 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 53 | lemma filterlim_real_at_infinity_sequentially: "filterlim real at_infinity sequentially" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 54 | by (simp add: filterlim_at_top_imp_at_infinity filterlim_real_sequentially) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 55 | |
| 63546 | 56 | lemma lim_infinity_imp_sequentially: "(f \<longlongrightarrow> l) at_infinity \<Longrightarrow> ((\<lambda>n. f(n)) \<longlongrightarrow> l) sequentially" | 
| 57 | by (simp add: filterlim_at_top_imp_at_infinity filterlim_compose filterlim_real_sequentially) | |
| 59613 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 58 | |
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 59 | |
| 60758 | 60 | subsubsection \<open>Boundedness\<close> | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 61 | |
| 63546 | 62 | definition Bfun :: "('a \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a filter \<Rightarrow> bool"
 | 
| 63 | where Bfun_metric_def: "Bfun f F = (\<exists>y. \<exists>K>0. eventually (\<lambda>x. dist (f x) y \<le> K) F)" | |
| 64 | ||
| 65 | abbreviation Bseq :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" | |
| 66 | where "Bseq X \<equiv> Bfun X sequentially" | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 67 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 68 | lemma Bseq_conv_Bfun: "Bseq X \<longleftrightarrow> Bfun X sequentially" .. | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 69 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 70 | lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 71 | unfolding Bfun_metric_def by (subst eventually_sequentially_seg) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 72 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 73 | lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 74 | unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg) | 
| 31355 | 75 | |
| 63546 | 76 | lemma Bfun_def: "Bfun f F \<longleftrightarrow> (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) F)" | 
| 51474 
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
 hoelzl parents: 
51472diff
changeset | 77 | unfolding Bfun_metric_def norm_conv_dist | 
| 
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
 hoelzl parents: 
51472diff
changeset | 78 | proof safe | 
| 63546 | 79 | fix y K | 
| 80 | assume K: "0 < K" and *: "eventually (\<lambda>x. dist (f x) y \<le> K) F" | |
| 51474 
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
 hoelzl parents: 
51472diff
changeset | 81 | moreover have "eventually (\<lambda>x. dist (f x) 0 \<le> dist (f x) y + dist 0 y) F" | 
| 
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
 hoelzl parents: 
51472diff
changeset | 82 | by (intro always_eventually) (metis dist_commute dist_triangle) | 
| 
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
 hoelzl parents: 
51472diff
changeset | 83 | with * have "eventually (\<lambda>x. dist (f x) 0 \<le> K + dist 0 y) F" | 
| 
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
 hoelzl parents: 
51472diff
changeset | 84 | by eventually_elim auto | 
| 60758 | 85 | with \<open>0 < K\<close> show "\<exists>K>0. eventually (\<lambda>x. dist (f x) 0 \<le> K) F" | 
| 51474 
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
 hoelzl parents: 
51472diff
changeset | 86 | by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto | 
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62369diff
changeset | 87 | qed (force simp del: norm_conv_dist [symmetric]) | 
| 31355 | 88 | |
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 89 | lemma BfunI: | 
| 63546 | 90 | assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) F" | 
| 91 | shows "Bfun f F" | |
| 92 | unfolding Bfun_def | |
| 31355 | 93 | proof (intro exI conjI allI) | 
| 94 | show "0 < max K 1" by simp | |
| 44195 | 95 | show "eventually (\<lambda>x. norm (f x) \<le> max K 1) F" | 
| 63546 | 96 | using K by (rule eventually_mono) simp | 
| 31355 | 97 | qed | 
| 98 | ||
| 99 | lemma BfunE: | |
| 44195 | 100 | assumes "Bfun f F" | 
| 101 | obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) F" | |
| 63546 | 102 | using assms unfolding Bfun_def by blast | 
| 31355 | 103 | |
| 68614 | 104 | lemma Cauchy_Bseq: | 
| 105 | assumes "Cauchy X" shows "Bseq X" | |
| 106 | proof - | |
| 107 | have "\<exists>y K. 0 < K \<and> (\<exists>N. \<forall>n\<ge>N. dist (X n) y \<le> K)" | |
| 108 | if "\<And>m n. \<lbrakk>m \<ge> M; n \<ge> M\<rbrakk> \<Longrightarrow> dist (X m) (X n) < 1" for M | |
| 109 | by (meson order.order_iff_strict that zero_less_one) | |
| 110 | with assms show ?thesis | |
| 111 | by (force simp: Cauchy_def Bfun_metric_def eventually_sequentially) | |
| 112 | qed | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 113 | |
| 60758 | 114 | subsubsection \<open>Bounded Sequences\<close> | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 115 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 116 | lemma BseqI': "(\<And>n. norm (X n) \<le> K) \<Longrightarrow> Bseq X" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 117 | by (intro BfunI) (auto simp: eventually_sequentially) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 118 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 119 | lemma BseqI2': "\<forall>n\<ge>N. norm (X n) \<le> K \<Longrightarrow> Bseq X" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 120 | by (intro BfunI) (auto simp: eventually_sequentially) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 121 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 122 | lemma Bseq_def: "Bseq X \<longleftrightarrow> (\<exists>K>0. \<forall>n. norm (X n) \<le> K)" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 123 | unfolding Bfun_def eventually_sequentially | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 124 | proof safe | 
| 63546 | 125 | fix N K | 
| 126 | assume "0 < K" "\<forall>n\<ge>N. norm (X n) \<le> K" | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 127 | then show "\<exists>K>0. \<forall>n. norm (X n) \<le> K" | 
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
54263diff
changeset | 128 |     by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] max.strict_coboundedI2)
 | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 129 | (auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 130 | qed auto | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 131 | |
| 63546 | 132 | lemma BseqE: "Bseq X \<Longrightarrow> (\<And>K. 0 < K \<Longrightarrow> \<forall>n. norm (X n) \<le> K \<Longrightarrow> Q) \<Longrightarrow> Q" | 
| 133 | unfolding Bseq_def by auto | |
| 134 | ||
| 135 | lemma BseqD: "Bseq X \<Longrightarrow> \<exists>K. 0 < K \<and> (\<forall>n. norm (X n) \<le> K)" | |
| 136 | by (simp add: Bseq_def) | |
| 137 | ||
| 138 | lemma BseqI: "0 < K \<Longrightarrow> \<forall>n. norm (X n) \<le> K \<Longrightarrow> Bseq X" | |
| 68615 | 139 | by (auto simp: Bseq_def) | 
| 63546 | 140 | |
| 141 | lemma Bseq_bdd_above: "Bseq X \<Longrightarrow> bdd_above (range X)" | |
| 142 | for X :: "nat \<Rightarrow> real" | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 143 | proof (elim BseqE, intro bdd_aboveI2) | 
| 63546 | 144 | fix K n | 
| 145 | assume "0 < K" "\<forall>n. norm (X n) \<le> K" | |
| 146 | then show "X n \<le> K" | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 147 | by (auto elim!: allE[of _ n]) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 148 | qed | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 149 | |
| 63546 | 150 | lemma Bseq_bdd_above': "Bseq X \<Longrightarrow> bdd_above (range (\<lambda>n. norm (X n)))" | 
| 151 | for X :: "nat \<Rightarrow> 'a :: real_normed_vector" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 152 | proof (elim BseqE, intro bdd_aboveI2) | 
| 63546 | 153 | fix K n | 
| 154 | assume "0 < K" "\<forall>n. norm (X n) \<le> K" | |
| 155 | then show "norm (X n) \<le> K" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 156 | by (auto elim!: allE[of _ n]) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 157 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 158 | |
| 63546 | 159 | lemma Bseq_bdd_below: "Bseq X \<Longrightarrow> bdd_below (range X)" | 
| 160 | for X :: "nat \<Rightarrow> real" | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 161 | proof (elim BseqE, intro bdd_belowI2) | 
| 63546 | 162 | fix K n | 
| 163 | assume "0 < K" "\<forall>n. norm (X n) \<le> K" | |
| 164 | then show "- K \<le> X n" | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 165 | by (auto elim!: allE[of _ n]) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 166 | qed | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 167 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 168 | lemma Bseq_eventually_mono: | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 169 | assumes "eventually (\<lambda>n. norm (f n) \<le> norm (g n)) sequentially" "Bseq g" | 
| 63546 | 170 | shows "Bseq f" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 171 | proof - | 
| 67958 
732c0b059463
tuned proofs and generalized some lemmas about limits
 huffman parents: 
67950diff
changeset | 172 | from assms(2) obtain K where "0 < K" and "eventually (\<lambda>n. norm (g n) \<le> K) sequentially" | 
| 
732c0b059463
tuned proofs and generalized some lemmas about limits
 huffman parents: 
67950diff
changeset | 173 | unfolding Bfun_def by fast | 
| 
732c0b059463
tuned proofs and generalized some lemmas about limits
 huffman parents: 
67950diff
changeset | 174 | with assms(1) have "eventually (\<lambda>n. norm (f n) \<le> K) sequentially" | 
| 
732c0b059463
tuned proofs and generalized some lemmas about limits
 huffman parents: 
67950diff
changeset | 175 | by (fast elim: eventually_elim2 order_trans) | 
| 69272 | 176 | with \<open>0 < K\<close> show "Bseq f" | 
| 67958 
732c0b059463
tuned proofs and generalized some lemmas about limits
 huffman parents: 
67950diff
changeset | 177 | unfolding Bfun_def by fast | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 178 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 179 | |
| 63546 | 180 | lemma lemma_NBseq_def: "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) \<longleftrightarrow> (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))" | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 181 | proof safe | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 182 | fix K :: real | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 183 | from reals_Archimedean2 obtain n :: nat where "K < real n" .. | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 184 | then have "K \<le> real (Suc n)" by auto | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 185 | moreover assume "\<forall>m. norm (X m) \<le> K" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 186 | ultimately have "\<forall>m. norm (X m) \<le> real (Suc n)" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 187 | by (blast intro: order_trans) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 188 | then show "\<exists>N. \<forall>n. norm (X n) \<le> real (Suc N)" .. | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 189 | next | 
| 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 190 | show "\<And>N. \<forall>n. norm (X n) \<le> real (Suc N) \<Longrightarrow> \<exists>K>0. \<forall>n. norm (X n) \<le> K" | 
| 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 191 | using of_nat_0_less_iff by blast | 
| 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 192 | qed | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 193 | |
| 63546 | 194 | text \<open>Alternative definition for \<open>Bseq\<close>.\<close> | 
| 195 | lemma Bseq_iff: "Bseq X \<longleftrightarrow> (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))" | |
| 196 | by (simp add: Bseq_def) (simp add: lemma_NBseq_def) | |
| 197 | ||
| 198 | lemma lemma_NBseq_def2: "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))" | |
| 68614 | 199 | proof - | 
| 200 | have *: "\<And>N. \<forall>n. norm (X n) \<le> 1 + real N \<Longrightarrow> | |
| 201 | \<exists>N. \<forall>n. norm (X n) < 1 + real N" | |
| 202 | by (metis add.commute le_less_trans less_add_one of_nat_Suc) | |
| 203 | then show ?thesis | |
| 204 | unfolding lemma_NBseq_def | |
| 205 | by (metis less_le_not_le not_less_iff_gr_or_eq of_nat_Suc) | |
| 206 | qed | |
| 63546 | 207 | |
| 208 | text \<open>Yet another definition for Bseq.\<close> | |
| 209 | lemma Bseq_iff1a: "Bseq X \<longleftrightarrow> (\<exists>N. \<forall>n. norm (X n) < real (Suc N))" | |
| 210 | by (simp add: Bseq_def lemma_NBseq_def2) | |
| 211 | ||
| 212 | subsubsection \<open>A Few More Equivalence Theorems for Boundedness\<close> | |
| 213 | ||
| 214 | text \<open>Alternative formulation for boundedness.\<close> | |
| 215 | lemma Bseq_iff2: "Bseq X \<longleftrightarrow> (\<exists>k > 0. \<exists>x. \<forall>n. norm (X n + - x) \<le> k)" | |
| 68614 | 216 | by (metis BseqE BseqI' add.commute add_cancel_right_left add_uminus_conv_diff norm_add_leD | 
| 217 | norm_minus_cancel norm_minus_commute) | |
| 63546 | 218 | |
| 219 | text \<open>Alternative formulation for boundedness.\<close> | |
| 220 | lemma Bseq_iff3: "Bseq X \<longleftrightarrow> (\<exists>k>0. \<exists>N. \<forall>n. norm (X n + - X N) \<le> k)" | |
| 221 | (is "?P \<longleftrightarrow> ?Q") | |
| 53602 | 222 | proof | 
| 223 | assume ?P | |
| 63546 | 224 | then obtain K where *: "0 < K" and **: "\<And>n. norm (X n) \<le> K" | 
| 68615 | 225 | by (auto simp: Bseq_def) | 
| 53602 | 226 | from * have "0 < K + norm (X 0)" by (rule order_less_le_trans) simp | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53602diff
changeset | 227 | from ** have "\<forall>n. norm (X n - X 0) \<le> K + norm (X 0)" | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53602diff
changeset | 228 | by (auto intro: order_trans norm_triangle_ineq4) | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53602diff
changeset | 229 | then have "\<forall>n. norm (X n + - X 0) \<le> K + norm (X 0)" | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53602diff
changeset | 230 | by simp | 
| 60758 | 231 | with \<open>0 < K + norm (X 0)\<close> show ?Q by blast | 
| 53602 | 232 | next | 
| 63546 | 233 | assume ?Q | 
| 68615 | 234 | then show ?P by (auto simp: Bseq_iff2) | 
| 53602 | 235 | qed | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 236 | |
| 63546 | 237 | |
| 238 | subsubsection \<open>Upper Bounds and Lubs of Bounded Sequences\<close> | |
| 239 | ||
| 240 | lemma Bseq_minus_iff: "Bseq (\<lambda>n. - (X n) :: 'a::real_normed_vector) \<longleftrightarrow> Bseq X" | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 241 | by (simp add: Bseq_def) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 242 | |
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 243 | lemma Bseq_add: | 
| 63546 | 244 | fixes f :: "nat \<Rightarrow> 'a::real_normed_vector" | 
| 245 | assumes "Bseq f" | |
| 246 | shows "Bseq (\<lambda>x. f x + c)" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 247 | proof - | 
| 63546 | 248 | from assms obtain K where K: "\<And>x. norm (f x) \<le> K" | 
| 249 | unfolding Bseq_def by blast | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 250 |   {
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 251 | fix x :: nat | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 252 | have "norm (f x + c) \<le> norm (f x) + norm c" by (rule norm_triangle_ineq) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 253 | also have "norm (f x) \<le> K" by (rule K) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 254 | finally have "norm (f x + c) \<le> K + norm c" by simp | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 255 | } | 
| 63546 | 256 | then show ?thesis by (rule BseqI') | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 257 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 258 | |
| 63546 | 259 | lemma Bseq_add_iff: "Bseq (\<lambda>x. f x + c) \<longleftrightarrow> Bseq f" | 
| 260 | for f :: "nat \<Rightarrow> 'a::real_normed_vector" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 261 | using Bseq_add[of f c] Bseq_add[of "\<lambda>x. f x + c" "-c"] by auto | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 262 | |
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 263 | lemma Bseq_mult: | 
| 63546 | 264 | fixes f g :: "nat \<Rightarrow> 'a::real_normed_field" | 
| 265 | assumes "Bseq f" and "Bseq g" | |
| 266 | shows "Bseq (\<lambda>x. f x * g x)" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 267 | proof - | 
| 63546 | 268 | from assms obtain K1 K2 where K: "norm (f x) \<le> K1" "K1 > 0" "norm (g x) \<le> K2" "K2 > 0" | 
| 269 | for x | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 270 | unfolding Bseq_def by blast | 
| 63546 | 271 | then have "norm (f x * g x) \<le> K1 * K2" for x | 
| 272 | by (auto simp: norm_mult intro!: mult_mono) | |
| 273 | then show ?thesis by (rule BseqI') | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 274 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 275 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 276 | lemma Bfun_const [simp]: "Bfun (\<lambda>_. c) F" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 277 | unfolding Bfun_metric_def by (auto intro!: exI[of _ c] exI[of _ "1::real"]) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 278 | |
| 63546 | 279 | lemma Bseq_cmult_iff: | 
| 280 | fixes c :: "'a::real_normed_field" | |
| 281 | assumes "c \<noteq> 0" | |
| 282 | shows "Bseq (\<lambda>x. c * f x) \<longleftrightarrow> Bseq f" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 283 | proof | 
| 63546 | 284 | assume "Bseq (\<lambda>x. c * f x)" | 
| 285 | with Bfun_const have "Bseq (\<lambda>x. inverse c * (c * f x))" | |
| 286 | by (rule Bseq_mult) | |
| 287 | with \<open>c \<noteq> 0\<close> show "Bseq f" | |
| 288 | by (simp add: divide_simps) | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 289 | qed (intro Bseq_mult Bfun_const) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 290 | |
| 63546 | 291 | lemma Bseq_subseq: "Bseq f \<Longrightarrow> Bseq (\<lambda>x. f (g x))" | 
| 292 | for f :: "nat \<Rightarrow> 'a::real_normed_vector" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 293 | unfolding Bseq_def by auto | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 294 | |
| 63546 | 295 | lemma Bseq_Suc_iff: "Bseq (\<lambda>n. f (Suc n)) \<longleftrightarrow> Bseq f" | 
| 296 | for f :: "nat \<Rightarrow> 'a::real_normed_vector" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 297 | using Bseq_offset[of f 1] by (auto intro: Bseq_subseq) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 298 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 299 | lemma increasing_Bseq_subseq_iff: | 
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
65680diff
changeset | 300 | assumes "\<And>x y. x \<le> y \<Longrightarrow> norm (f x :: 'a::real_normed_vector) \<le> norm (f y)" "strict_mono g" | 
| 63546 | 301 | shows "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 302 | proof | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 303 | assume "Bseq (\<lambda>x. f (g x))" | 
| 63546 | 304 | then obtain K where K: "\<And>x. norm (f (g x)) \<le> K" | 
| 305 | unfolding Bseq_def by auto | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 306 |   {
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 307 | fix x :: nat | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 308 | from filterlim_subseq[OF assms(2)] obtain y where "g y \<ge> x" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 309 | by (auto simp: filterlim_at_top eventually_at_top_linorder) | 
| 63546 | 310 | then have "norm (f x) \<le> norm (f (g y))" | 
| 311 | using assms(1) by blast | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 312 | also have "norm (f (g y)) \<le> K" by (rule K) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 313 | finally have "norm (f x) \<le> K" . | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 314 | } | 
| 63546 | 315 | then show "Bseq f" by (rule BseqI') | 
| 316 | qed (use Bseq_subseq[of f g] in simp_all) | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 317 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 318 | lemma nonneg_incseq_Bseq_subseq_iff: | 
| 63546 | 319 | fixes f :: "nat \<Rightarrow> real" | 
| 320 | and g :: "nat \<Rightarrow> nat" | |
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
65680diff
changeset | 321 | assumes "\<And>x. f x \<ge> 0" "incseq f" "strict_mono g" | 
| 63546 | 322 | shows "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 323 | using assms by (intro increasing_Bseq_subseq_iff) (auto simp: incseq_def) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 324 | |
| 63546 | 325 | lemma Bseq_eq_bounded: "range f \<subseteq> {a..b} \<Longrightarrow> Bseq f"
 | 
| 326 | for a b :: real | |
| 68614 | 327 | proof (rule BseqI'[where K="max (norm a) (norm b)"]) | 
| 328 |   fix n assume "range f \<subseteq> {a..b}"
 | |
| 329 |   then have "f n \<in> {a..b}"
 | |
| 330 | by blast | |
| 331 | then show "norm (f n) \<le> max (norm a) (norm b)" | |
| 332 | by auto | |
| 333 | qed | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 334 | |
| 63546 | 335 | lemma incseq_bounded: "incseq X \<Longrightarrow> \<forall>i. X i \<le> B \<Longrightarrow> Bseq X" | 
| 336 | for B :: real | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 337 | by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 338 | |
| 63546 | 339 | lemma decseq_bounded: "decseq X \<Longrightarrow> \<forall>i. B \<le> X i \<Longrightarrow> Bseq X" | 
| 340 | for B :: real | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 341 | by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 342 | |
| 63546 | 343 | |
| 60758 | 344 | subsection \<open>Convergence to Zero\<close> | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 345 | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 346 | definition Zfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
 | 
| 44195 | 347 | where "Zfun f F = (\<forall>r>0. eventually (\<lambda>x. norm (f x) < r) F)" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 348 | |
| 63546 | 349 | lemma ZfunI: "(\<And>r. 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F) \<Longrightarrow> Zfun f F" | 
| 350 | by (simp add: Zfun_def) | |
| 351 | ||
| 352 | lemma ZfunD: "Zfun f F \<Longrightarrow> 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F" | |
| 353 | by (simp add: Zfun_def) | |
| 354 | ||
| 355 | lemma Zfun_ssubst: "eventually (\<lambda>x. f x = g x) F \<Longrightarrow> Zfun g F \<Longrightarrow> Zfun f F" | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 356 | unfolding Zfun_def by (auto elim!: eventually_rev_mp) | 
| 31355 | 357 | |
| 44195 | 358 | lemma Zfun_zero: "Zfun (\<lambda>x. 0) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 359 | unfolding Zfun_def by simp | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 360 | |
| 44195 | 361 | lemma Zfun_norm_iff: "Zfun (\<lambda>x. norm (f x)) F = Zfun (\<lambda>x. f x) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 362 | unfolding Zfun_def by simp | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 363 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 364 | lemma Zfun_imp_Zfun: | 
| 44195 | 365 | assumes f: "Zfun f F" | 
| 63546 | 366 | and g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F" | 
| 44195 | 367 | shows "Zfun (\<lambda>x. g x) F" | 
| 63546 | 368 | proof (cases "0 < K") | 
| 369 | case K: True | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 370 | show ?thesis | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 371 | proof (rule ZfunI) | 
| 63546 | 372 | fix r :: real | 
| 373 | assume "0 < r" | |
| 374 | then have "0 < r / K" using K by simp | |
| 44195 | 375 | then have "eventually (\<lambda>x. norm (f x) < r / K) F" | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 376 | using ZfunD [OF f] by blast | 
| 44195 | 377 | with g show "eventually (\<lambda>x. norm (g x) < r) F" | 
| 46887 | 378 | proof eventually_elim | 
| 379 | case (elim x) | |
| 63546 | 380 | then have "norm (f x) * K < r" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 381 | by (simp add: pos_less_divide_eq K) | 
| 63546 | 382 | then show ?case | 
| 46887 | 383 | by (simp add: order_le_less_trans [OF elim(1)]) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 384 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 385 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 386 | next | 
| 63546 | 387 | case False | 
| 388 | then have K: "K \<le> 0" by (simp only: not_less) | |
| 31355 | 389 | show ?thesis | 
| 390 | proof (rule ZfunI) | |
| 391 | fix r :: real | |
| 392 | assume "0 < r" | |
| 44195 | 393 | from g show "eventually (\<lambda>x. norm (g x) < r) F" | 
| 46887 | 394 | proof eventually_elim | 
| 395 | case (elim x) | |
| 396 | also have "norm (f x) * K \<le> norm (f x) * 0" | |
| 31355 | 397 | using K norm_ge_zero by (rule mult_left_mono) | 
| 46887 | 398 | finally show ?case | 
| 60758 | 399 | using \<open>0 < r\<close> by simp | 
| 31355 | 400 | qed | 
| 401 | qed | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 402 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 403 | |
| 63546 | 404 | lemma Zfun_le: "Zfun g F \<Longrightarrow> \<forall>x. norm (f x) \<le> norm (g x) \<Longrightarrow> Zfun f F" | 
| 405 | by (erule Zfun_imp_Zfun [where K = 1]) simp | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 406 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 407 | lemma Zfun_add: | 
| 63546 | 408 | assumes f: "Zfun f F" | 
| 409 | and g: "Zfun g F" | |
| 44195 | 410 | shows "Zfun (\<lambda>x. f x + g x) F" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 411 | proof (rule ZfunI) | 
| 63546 | 412 | fix r :: real | 
| 413 | assume "0 < r" | |
| 414 | then have r: "0 < r / 2" by simp | |
| 44195 | 415 | have "eventually (\<lambda>x. norm (f x) < r/2) F" | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 416 | using f r by (rule ZfunD) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 417 | moreover | 
| 44195 | 418 | have "eventually (\<lambda>x. norm (g x) < r/2) F" | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 419 | using g r by (rule ZfunD) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 420 | ultimately | 
| 44195 | 421 | show "eventually (\<lambda>x. norm (f x + g x) < r) F" | 
| 46887 | 422 | proof eventually_elim | 
| 423 | case (elim x) | |
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 424 | have "norm (f x + g x) \<le> norm (f x) + norm (g x)" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 425 | by (rule norm_triangle_ineq) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 426 | also have "\<dots> < r/2 + r/2" | 
| 46887 | 427 | using elim by (rule add_strict_mono) | 
| 428 | finally show ?case | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 429 | by simp | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 430 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 431 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 432 | |
| 44195 | 433 | lemma Zfun_minus: "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. - f x) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 434 | unfolding Zfun_def by simp | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 435 | |
| 63546 | 436 | lemma Zfun_diff: "Zfun f F \<Longrightarrow> Zfun g F \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53602diff
changeset | 437 | using Zfun_add [of f F "\<lambda>x. - g x"] by (simp add: Zfun_minus) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 438 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 439 | lemma (in bounded_linear) Zfun: | 
| 44195 | 440 | assumes g: "Zfun g F" | 
| 441 | shows "Zfun (\<lambda>x. f (g x)) F" | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 442 | proof - | 
| 63546 | 443 | obtain K where "norm (f x) \<le> norm x * K" for x | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 444 | using bounded by blast | 
| 44195 | 445 | then have "eventually (\<lambda>x. norm (f (g x)) \<le> norm (g x) * K) F" | 
| 31355 | 446 | by simp | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 447 | with g show ?thesis | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 448 | by (rule Zfun_imp_Zfun) | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 449 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 450 | |
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 451 | lemma (in bounded_bilinear) Zfun: | 
| 44195 | 452 | assumes f: "Zfun f F" | 
| 63546 | 453 | and g: "Zfun g F" | 
| 44195 | 454 | shows "Zfun (\<lambda>x. f x ** g x) F" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 455 | proof (rule ZfunI) | 
| 63546 | 456 | fix r :: real | 
| 457 | assume r: "0 < r" | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 458 | obtain K where K: "0 < K" | 
| 63546 | 459 | and norm_le: "norm (x ** y) \<le> norm x * norm y * K" for x y | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 460 | using pos_bounded by blast | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 461 | from K have K': "0 < inverse K" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 462 | by (rule positive_imp_inverse_positive) | 
| 44195 | 463 | have "eventually (\<lambda>x. norm (f x) < r) F" | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 464 | using f r by (rule ZfunD) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 465 | moreover | 
| 44195 | 466 | have "eventually (\<lambda>x. norm (g x) < inverse K) F" | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 467 | using g K' by (rule ZfunD) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 468 | ultimately | 
| 44195 | 469 | show "eventually (\<lambda>x. norm (f x ** g x) < r) F" | 
| 46887 | 470 | proof eventually_elim | 
| 471 | case (elim x) | |
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 472 | have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K" | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 473 | by (rule norm_le) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 474 | also have "norm (f x) * norm (g x) * K < r * inverse K * K" | 
| 46887 | 475 | by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero elim K) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 476 | also from K have "r * inverse K * K = r" | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 477 | by simp | 
| 46887 | 478 | finally show ?case . | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 479 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 480 | qed | 
| 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 481 | |
| 63546 | 482 | lemma (in bounded_bilinear) Zfun_left: "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. f x ** a) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 483 | by (rule bounded_linear_left [THEN bounded_linear.Zfun]) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 484 | |
| 63546 | 485 | lemma (in bounded_bilinear) Zfun_right: "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. a ** f x) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 486 | by (rule bounded_linear_right [THEN bounded_linear.Zfun]) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 487 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 488 | lemmas Zfun_mult = bounded_bilinear.Zfun [OF bounded_bilinear_mult] | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 489 | lemmas Zfun_mult_right = bounded_bilinear.Zfun_right [OF bounded_bilinear_mult] | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 490 | lemmas Zfun_mult_left = bounded_bilinear.Zfun_left [OF bounded_bilinear_mult] | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 491 | |
| 61973 | 492 | lemma tendsto_Zfun_iff: "(f \<longlongrightarrow> a) F = Zfun (\<lambda>x. f x - a) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 493 | by (simp only: tendsto_iff Zfun_def dist_norm) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 494 | |
| 63546 | 495 | lemma tendsto_0_le: | 
| 496 | "(f \<longlongrightarrow> 0) F \<Longrightarrow> eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F \<Longrightarrow> (g \<longlongrightarrow> 0) F" | |
| 56366 | 497 | by (simp add: Zfun_imp_Zfun tendsto_Zfun_iff) | 
| 498 | ||
| 63546 | 499 | |
| 60758 | 500 | subsubsection \<open>Distance and norms\<close> | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 501 | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 502 | lemma tendsto_dist [tendsto_intros]: | 
| 63546 | 503 | fixes l m :: "'a::metric_space" | 
| 504 | assumes f: "(f \<longlongrightarrow> l) F" | |
| 505 | and g: "(g \<longlongrightarrow> m) F" | |
| 61973 | 506 | shows "((\<lambda>x. dist (f x) (g x)) \<longlongrightarrow> dist l m) F" | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 507 | proof (rule tendstoI) | 
| 63546 | 508 | fix e :: real | 
| 509 | assume "0 < e" | |
| 510 | then have e2: "0 < e/2" by simp | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 511 | from tendstoD [OF f e2] tendstoD [OF g e2] | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 512 | show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) F" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 513 | proof (eventually_elim) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 514 | case (elim x) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 515 | then show "dist (dist (f x) (g x)) (dist l m) < e" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 516 | unfolding dist_real_def | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 517 | using dist_triangle2 [of "f x" "g x" "l"] | 
| 63546 | 518 | and dist_triangle2 [of "g x" "l" "m"] | 
| 519 | and dist_triangle3 [of "l" "m" "f x"] | |
| 520 | and dist_triangle [of "f x" "m" "g x"] | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 521 | by arith | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 522 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 523 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 524 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 525 | lemma continuous_dist[continuous_intros]: | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 526 | fixes f g :: "_ \<Rightarrow> 'a :: metric_space" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 527 | shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. dist (f x) (g x))" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 528 | unfolding continuous_def by (rule tendsto_dist) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 529 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56366diff
changeset | 530 | lemma continuous_on_dist[continuous_intros]: | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 531 | fixes f g :: "_ \<Rightarrow> 'a :: metric_space" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 532 | shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. dist (f x) (g x))" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 533 | unfolding continuous_on_def by (auto intro: tendsto_dist) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 534 | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 535 | lemma continuous_at_dist: "isCont (dist a) b" | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 536 | using continuous_on_dist [OF continuous_on_const continuous_on_id] continuous_on_eq_continuous_within by blast | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 537 | |
| 63546 | 538 | lemma tendsto_norm [tendsto_intros]: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. norm (f x)) \<longlongrightarrow> norm a) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 539 | unfolding norm_conv_dist by (intro tendsto_intros) | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 540 | |
| 63546 | 541 | lemma continuous_norm [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. norm (f x))" | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 542 | unfolding continuous_def by (rule tendsto_norm) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 543 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56366diff
changeset | 544 | lemma continuous_on_norm [continuous_intros]: | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 545 | "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. norm (f x))" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 546 | unfolding continuous_on_def by (auto intro: tendsto_norm) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 547 | |
| 63546 | 548 | lemma tendsto_norm_zero: "(f \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) F" | 
| 549 | by (drule tendsto_norm) simp | |
| 550 | ||
| 551 | lemma tendsto_norm_zero_cancel: "((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> 0) F" | |
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 552 | unfolding tendsto_iff dist_norm by simp | 
| 36662 
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
 huffman parents: 
36656diff
changeset | 553 | |
| 63546 | 554 | lemma tendsto_norm_zero_iff: "((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) F \<longleftrightarrow> (f \<longlongrightarrow> 0) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 555 | unfolding tendsto_iff dist_norm by simp | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 556 | |
| 63546 | 557 | lemma tendsto_rabs [tendsto_intros]: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> \<bar>l\<bar>) F" | 
| 558 | for l :: real | |
| 559 | by (fold real_norm_def) (rule tendsto_norm) | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 560 | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 561 | lemma continuous_rabs [continuous_intros]: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 562 | "continuous F f \<Longrightarrow> continuous F (\<lambda>x. \<bar>f x :: real\<bar>)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 563 | unfolding real_norm_def[symmetric] by (rule continuous_norm) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 564 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56366diff
changeset | 565 | lemma continuous_on_rabs [continuous_intros]: | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 566 | "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. \<bar>f x :: real\<bar>)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 567 | unfolding real_norm_def[symmetric] by (rule continuous_on_norm) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 568 | |
| 63546 | 569 | lemma tendsto_rabs_zero: "(f \<longlongrightarrow> (0::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> 0) F" | 
| 570 | by (fold real_norm_def) (rule tendsto_norm_zero) | |
| 571 | ||
| 572 | lemma tendsto_rabs_zero_cancel: "((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> (0::real)) F \<Longrightarrow> (f \<longlongrightarrow> 0) F" | |
| 573 | by (fold real_norm_def) (rule tendsto_norm_zero_cancel) | |
| 574 | ||
| 575 | lemma tendsto_rabs_zero_iff: "((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> (0::real)) F \<longleftrightarrow> (f \<longlongrightarrow> 0) F" | |
| 576 | by (fold real_norm_def) (rule tendsto_norm_zero_iff) | |
| 577 | ||
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 578 | |
| 62368 | 579 | subsection \<open>Topological Monoid\<close> | 
| 580 | ||
| 581 | class topological_monoid_add = topological_space + monoid_add + | |
| 582 | assumes tendsto_add_Pair: "LIM x (nhds a \<times>\<^sub>F nhds b). fst x + snd x :> nhds (a + b)" | |
| 583 | ||
| 584 | class topological_comm_monoid_add = topological_monoid_add + comm_monoid_add | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 585 | |
| 31565 | 586 | lemma tendsto_add [tendsto_intros]: | 
| 62368 | 587 | fixes a b :: "'a::topological_monoid_add" | 
| 588 | shows "(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> ((\<lambda>x. f x + g x) \<longlongrightarrow> a + b) F" | |
| 589 | using filterlim_compose[OF tendsto_add_Pair, of "\<lambda>x. (f x, g x)" a b F] | |
| 590 | by (simp add: nhds_prod[symmetric] tendsto_Pair) | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 591 | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 592 | lemma continuous_add [continuous_intros]: | 
| 62368 | 593 | fixes f g :: "_ \<Rightarrow> 'b::topological_monoid_add" | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 594 | shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x + g x)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 595 | unfolding continuous_def by (rule tendsto_add) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 596 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56366diff
changeset | 597 | lemma continuous_on_add [continuous_intros]: | 
| 62368 | 598 | fixes f g :: "_ \<Rightarrow> 'b::topological_monoid_add" | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 599 | shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f x + g x)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 600 | unfolding continuous_on_def by (auto intro: tendsto_add) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 601 | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 602 | lemma tendsto_add_zero: | 
| 62368 | 603 | fixes f g :: "_ \<Rightarrow> 'b::topological_monoid_add" | 
| 63546 | 604 | shows "(f \<longlongrightarrow> 0) F \<Longrightarrow> (g \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f x + g x) \<longlongrightarrow> 0) F" | 
| 605 | by (drule (1) tendsto_add) simp | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 606 | |
| 64267 | 607 | lemma tendsto_sum [tendsto_intros]: | 
| 62368 | 608 | fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::topological_comm_monoid_add" | 
| 63915 | 609 | shows "(\<And>i. i \<in> I \<Longrightarrow> (f i \<longlongrightarrow> a i) F) \<Longrightarrow> ((\<lambda>x. \<Sum>i\<in>I. f i x) \<longlongrightarrow> (\<Sum>i\<in>I. a i)) F" | 
| 610 | by (induct I rule: infinite_finite_induct) (simp_all add: tendsto_add) | |
| 62368 | 611 | |
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 612 | lemma tendsto_null_sum: | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 613 | fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::topological_comm_monoid_add" | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 614 | assumes "\<And>i. i \<in> I \<Longrightarrow> ((\<lambda>x. f x i) \<longlongrightarrow> 0) F" | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 615 | shows "((\<lambda>i. sum (f i) I) \<longlongrightarrow> 0) F" | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 616 | using tendsto_sum [of I "\<lambda>x y. f y x" "\<lambda>x. 0"] assms by simp | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 617 | |
| 64267 | 618 | lemma continuous_sum [continuous_intros]: | 
| 62368 | 619 | fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::topological_comm_monoid_add" | 
| 63301 | 620 | shows "(\<And>i. i \<in> I \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Sum>i\<in>I. f i x)" | 
| 64267 | 621 | unfolding continuous_def by (rule tendsto_sum) | 
| 622 | ||
| 623 | lemma continuous_on_sum [continuous_intros]: | |
| 62368 | 624 | fixes f :: "'a \<Rightarrow> 'b::topological_space \<Rightarrow> 'c::topological_comm_monoid_add" | 
| 63301 | 625 | shows "(\<And>i. i \<in> I \<Longrightarrow> continuous_on S (f i)) \<Longrightarrow> continuous_on S (\<lambda>x. \<Sum>i\<in>I. f i x)" | 
| 64267 | 626 | unfolding continuous_on_def by (auto intro: tendsto_sum) | 
| 62368 | 627 | |
| 62369 | 628 | instance nat :: topological_comm_monoid_add | 
| 63546 | 629 | by standard | 
| 630 | (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) | |
| 62369 | 631 | |
| 632 | instance int :: topological_comm_monoid_add | |
| 63546 | 633 | by standard | 
| 634 | (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) | |
| 635 | ||
| 62369 | 636 | |
| 63081 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 637 | subsubsection \<open>Topological group\<close> | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 638 | |
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 639 | class topological_group_add = topological_monoid_add + group_add + | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 640 | assumes tendsto_uminus_nhds: "(uminus \<longlongrightarrow> - a) (nhds a)" | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 641 | begin | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 642 | |
| 63546 | 643 | lemma tendsto_minus [tendsto_intros]: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. - f x) \<longlongrightarrow> - a) F" | 
| 63081 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 644 | by (rule filterlim_compose[OF tendsto_uminus_nhds]) | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 645 | |
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 646 | end | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 647 | |
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 648 | class topological_ab_group_add = topological_group_add + ab_group_add | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 649 | |
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 650 | instance topological_ab_group_add < topological_comm_monoid_add .. | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 651 | |
| 63546 | 652 | lemma continuous_minus [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. - f x)" | 
| 653 | for f :: "'a::t2_space \<Rightarrow> 'b::topological_group_add" | |
| 63081 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 654 | unfolding continuous_def by (rule tendsto_minus) | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 655 | |
| 63546 | 656 | lemma continuous_on_minus [continuous_intros]: "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. - f x)" | 
| 657 | for f :: "_ \<Rightarrow> 'b::topological_group_add" | |
| 63081 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 658 | unfolding continuous_on_def by (auto intro: tendsto_minus) | 
| 62368 | 659 | |
| 63546 | 660 | lemma tendsto_minus_cancel: "((\<lambda>x. - f x) \<longlongrightarrow> - a) F \<Longrightarrow> (f \<longlongrightarrow> a) F" | 
| 661 | for a :: "'a::topological_group_add" | |
| 662 | by (drule tendsto_minus) simp | |
| 63081 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 663 | |
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 664 | lemma tendsto_minus_cancel_left: | 
| 63546 | 665 | "(f \<longlongrightarrow> - (y::_::topological_group_add)) F \<longleftrightarrow> ((\<lambda>x. - f x) \<longlongrightarrow> y) F" | 
| 63081 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 666 | using tendsto_minus_cancel[of f "- y" F] tendsto_minus[of f "- y" F] | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 667 | by auto | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 668 | |
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 669 | lemma tendsto_diff [tendsto_intros]: | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 670 | fixes a b :: "'a::topological_group_add" | 
| 63546 | 671 | shows "(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> ((\<lambda>x. f x - g x) \<longlongrightarrow> a - b) F" | 
| 63081 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 672 | using tendsto_add [of f a F "\<lambda>x. - g x" "- b"] by (simp add: tendsto_minus) | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 673 | |
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 674 | lemma continuous_diff [continuous_intros]: | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 675 | fixes f g :: "'a::t2_space \<Rightarrow> 'b::topological_group_add" | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 676 | shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x - g x)" | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 677 | unfolding continuous_def by (rule tendsto_diff) | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 678 | |
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 679 | lemma continuous_on_diff [continuous_intros]: | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 680 | fixes f g :: "_ \<Rightarrow> 'b::topological_group_add" | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 681 | shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f x - g x)" | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 682 | unfolding continuous_on_def by (auto intro: tendsto_diff) | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 683 | |
| 67399 | 684 | lemma continuous_on_op_minus: "continuous_on (s::'a::topological_group_add set) ((-) x)" | 
| 63081 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 685 | by (rule continuous_intros | simp)+ | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 686 | |
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 687 | instance real_normed_vector < topological_ab_group_add | 
| 62368 | 688 | proof | 
| 63546 | 689 | fix a b :: 'a | 
| 690 | show "((\<lambda>x. fst x + snd x) \<longlongrightarrow> a + b) (nhds a \<times>\<^sub>F nhds b)" | |
| 62368 | 691 | unfolding tendsto_Zfun_iff add_diff_add | 
| 692 | using tendsto_fst[OF filterlim_ident, of "(a,b)"] tendsto_snd[OF filterlim_ident, of "(a,b)"] | |
| 693 | by (intro Zfun_add) | |
| 68615 | 694 | (auto simp: tendsto_Zfun_iff[symmetric] nhds_prod[symmetric] intro!: tendsto_fst) | 
| 63081 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 695 | show "(uminus \<longlongrightarrow> - a) (nhds a)" | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 696 | unfolding tendsto_Zfun_iff minus_diff_minus | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 697 | using filterlim_ident[of "nhds a"] | 
| 
5a5beb3dbe7e
introduced class topological_group between topological_monoid and real_normed_vector
 immler parents: 
63040diff
changeset | 698 | by (intro Zfun_minus) (simp add: tendsto_Zfun_iff) | 
| 62368 | 699 | qed | 
| 700 | ||
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 701 | lemmas real_tendsto_sandwich = tendsto_sandwich[where 'a=real] | 
| 50999 | 702 | |
| 63546 | 703 | |
| 60758 | 704 | subsubsection \<open>Linear operators and multiplication\<close> | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 705 | |
| 63546 | 706 | lemma linear_times: "linear (\<lambda>x. c * x)" | 
| 707 | for c :: "'a::real_algebra" | |
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61799diff
changeset | 708 | by (auto simp: linearI distrib_left) | 
| 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61799diff
changeset | 709 | |
| 63546 | 710 | lemma (in bounded_linear) tendsto: "(g \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f (g x)) \<longlongrightarrow> f a) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 711 | by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) | 
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 712 | |
| 63546 | 713 | lemma (in bounded_linear) continuous: "continuous F g \<Longrightarrow> continuous F (\<lambda>x. f (g x))" | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 714 | using tendsto[of g _ F] by (auto simp: continuous_def) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 715 | |
| 63546 | 716 | lemma (in bounded_linear) continuous_on: "continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f (g x))" | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 717 | using tendsto[of g] by (auto simp: continuous_on_def) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 718 | |
| 63546 | 719 | lemma (in bounded_linear) tendsto_zero: "(g \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f (g x)) \<longlongrightarrow> 0) F" | 
| 720 | by (drule tendsto) (simp only: zero) | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 721 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 722 | lemma (in bounded_bilinear) tendsto: | 
| 63546 | 723 | "(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> ((\<lambda>x. f x ** g x) \<longlongrightarrow> a ** b) F" | 
| 724 | by (simp only: tendsto_Zfun_iff prod_diff_prod Zfun_add Zfun Zfun_left Zfun_right) | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 725 | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 726 | lemma (in bounded_bilinear) continuous: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 727 | "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x ** g x)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 728 | using tendsto[of f _ F g] by (auto simp: continuous_def) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 729 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 730 | lemma (in bounded_bilinear) continuous_on: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 731 | "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f x ** g x)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 732 | using tendsto[of f _ _ g] by (auto simp: continuous_on_def) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 733 | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 734 | lemma (in bounded_bilinear) tendsto_zero: | 
| 61973 | 735 | assumes f: "(f \<longlongrightarrow> 0) F" | 
| 63546 | 736 | and g: "(g \<longlongrightarrow> 0) F" | 
| 61973 | 737 | shows "((\<lambda>x. f x ** g x) \<longlongrightarrow> 0) F" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 738 | using tendsto [OF f g] by (simp add: zero_left) | 
| 31355 | 739 | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 740 | lemma (in bounded_bilinear) tendsto_left_zero: | 
| 61973 | 741 | "(f \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f x ** c) \<longlongrightarrow> 0) F" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 742 | by (rule bounded_linear.tendsto_zero [OF bounded_linear_left]) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 743 | |
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 744 | lemma (in bounded_bilinear) tendsto_right_zero: | 
| 61973 | 745 | "(f \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. c ** f x) \<longlongrightarrow> 0) F" | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 746 | by (rule bounded_linear.tendsto_zero [OF bounded_linear_right]) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 747 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 748 | lemmas tendsto_of_real [tendsto_intros] = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 749 | bounded_linear.tendsto [OF bounded_linear_of_real] | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 750 | |
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 751 | lemmas tendsto_scaleR [tendsto_intros] = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 752 | bounded_bilinear.tendsto [OF bounded_bilinear_scaleR] | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 753 | |
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 754 | |
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 755 | text\<open>Analogous type class for multiplication\<close> | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 756 | class topological_semigroup_mult = topological_space + semigroup_mult + | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 757 | assumes tendsto_mult_Pair: "LIM x (nhds a \<times>\<^sub>F nhds b). fst x * snd x :> nhds (a * b)" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 758 | |
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 759 | instance real_normed_algebra < topological_semigroup_mult | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 760 | proof | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 761 | fix a b :: 'a | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 762 | show "((\<lambda>x. fst x * snd x) \<longlongrightarrow> a * b) (nhds a \<times>\<^sub>F nhds b)" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 763 | unfolding nhds_prod[symmetric] | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 764 | using tendsto_fst[OF filterlim_ident, of "(a,b)"] tendsto_snd[OF filterlim_ident, of "(a,b)"] | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 765 | by (simp add: bounded_bilinear.tendsto [OF bounded_bilinear_mult]) | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 766 | qed | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 767 | |
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 768 | lemma tendsto_mult [tendsto_intros]: | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 769 | fixes a b :: "'a::topological_semigroup_mult" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 770 | shows "(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> ((\<lambda>x. f x * g x) \<longlongrightarrow> a * b) F" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 771 | using filterlim_compose[OF tendsto_mult_Pair, of "\<lambda>x. (f x, g x)" a b F] | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 772 | by (simp add: nhds_prod[symmetric] tendsto_Pair) | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 773 | |
| 63546 | 774 | lemma tendsto_mult_left: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. c * (f x)) \<longlongrightarrow> c * l) F" | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 775 | for c :: "'a::topological_semigroup_mult" | 
| 63546 | 776 | by (rule tendsto_mult [OF tendsto_const]) | 
| 777 | ||
| 778 | lemma tendsto_mult_right: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. (f x) * c) \<longlongrightarrow> l * c) F" | |
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 779 | for c :: "'a::topological_semigroup_mult" | 
| 63546 | 780 | by (rule tendsto_mult [OF _ tendsto_const]) | 
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61799diff
changeset | 781 | |
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 782 | lemma lim_const_over_n [tendsto_intros]: | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 783 | fixes a :: "'a::real_normed_field" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 784 | shows "(\<lambda>n. a / of_nat n) \<longlonglongrightarrow> 0" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 785 | using tendsto_mult [OF tendsto_const [of a] lim_1_over_n] by simp | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 786 | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 787 | lemmas continuous_of_real [continuous_intros] = | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 788 | bounded_linear.continuous [OF bounded_linear_of_real] | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 789 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 790 | lemmas continuous_scaleR [continuous_intros] = | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 791 | bounded_bilinear.continuous [OF bounded_bilinear_scaleR] | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 792 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 793 | lemmas continuous_mult [continuous_intros] = | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 794 | bounded_bilinear.continuous [OF bounded_bilinear_mult] | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 795 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56366diff
changeset | 796 | lemmas continuous_on_of_real [continuous_intros] = | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 797 | bounded_linear.continuous_on [OF bounded_linear_of_real] | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 798 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56366diff
changeset | 799 | lemmas continuous_on_scaleR [continuous_intros] = | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 800 | bounded_bilinear.continuous_on [OF bounded_bilinear_scaleR] | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 801 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56366diff
changeset | 802 | lemmas continuous_on_mult [continuous_intros] = | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 803 | bounded_bilinear.continuous_on [OF bounded_bilinear_mult] | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 804 | |
| 44568 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 805 | lemmas tendsto_mult_zero = | 
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 806 | bounded_bilinear.tendsto_zero [OF bounded_bilinear_mult] | 
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 807 | |
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 808 | lemmas tendsto_mult_left_zero = | 
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 809 | bounded_bilinear.tendsto_left_zero [OF bounded_bilinear_mult] | 
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 810 | |
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 811 | lemmas tendsto_mult_right_zero = | 
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 812 | bounded_bilinear.tendsto_right_zero [OF bounded_bilinear_mult] | 
| 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44342diff
changeset | 813 | |
| 68296 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 814 | |
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 815 | lemma continuous_mult_left: | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 816 | fixes c::"'a::real_normed_algebra" | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 817 | shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. c * f x)" | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 818 | by (rule continuous_mult [OF continuous_const]) | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 819 | |
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 820 | lemma continuous_mult_right: | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 821 | fixes c::"'a::real_normed_algebra" | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 822 | shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x * c)" | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 823 | by (rule continuous_mult [OF _ continuous_const]) | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 824 | |
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 825 | lemma continuous_on_mult_left: | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 826 | fixes c::"'a::real_normed_algebra" | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 827 | shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. c * f x)" | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 828 | by (rule continuous_on_mult [OF continuous_on_const]) | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 829 | |
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 830 | lemma continuous_on_mult_right: | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 831 | fixes c::"'a::real_normed_algebra" | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 832 | shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x * c)" | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 833 | by (rule continuous_on_mult [OF _ continuous_on_const]) | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 834 | |
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 835 | lemma continuous_on_mult_const [simp]: | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 836 | fixes c::"'a::real_normed_algebra" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68860diff
changeset | 837 | shows "continuous_on s ((*) c)" | 
| 68296 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 838 | by (intro continuous_on_mult_left continuous_on_id) | 
| 
69d680e94961
tidying and reorganisation around Cauchy Integral Theorem
 paulson <lp15@cam.ac.uk> parents: 
68064diff
changeset | 839 | |
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 840 | lemma tendsto_divide_zero: | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 841 | fixes c :: "'a::real_normed_field" | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 842 | shows "(f \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f x / c) \<longlongrightarrow> 0) F" | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 843 | by (cases "c=0") (simp_all add: divide_inverse tendsto_mult_left_zero) | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 844 | |
| 63546 | 845 | lemma tendsto_power [tendsto_intros]: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f x ^ n) \<longlongrightarrow> a ^ n) F" | 
| 846 |   for f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra}"
 | |
| 58729 
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
 hoelzl parents: 
57512diff
changeset | 847 | by (induct n) (simp_all add: tendsto_mult) | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 848 | |
| 65680 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 849 | lemma tendsto_null_power: "\<lbrakk>(f \<longlongrightarrow> 0) F; 0 < n\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x ^ n) \<longlongrightarrow> 0) F" | 
| 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 850 |     for f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra_1}"
 | 
| 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 851 | using tendsto_power [of f 0 F n] by (simp add: power_0_left) | 
| 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 852 | |
| 63546 | 853 | lemma continuous_power [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. (f x)^n)" | 
| 854 |   for f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
 | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 855 | unfolding continuous_def by (rule tendsto_power) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 856 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56366diff
changeset | 857 | lemma continuous_on_power [continuous_intros]: | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 858 |   fixes f :: "_ \<Rightarrow> 'b::{power,real_normed_algebra}"
 | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 859 | shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. (f x)^n)" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 860 | unfolding continuous_on_def by (auto intro: tendsto_power) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 861 | |
| 64272 | 862 | lemma tendsto_prod [tendsto_intros]: | 
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 863 |   fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
 | 
| 63915 | 864 | shows "(\<And>i. i \<in> S \<Longrightarrow> (f i \<longlongrightarrow> L i) F) \<Longrightarrow> ((\<lambda>x. \<Prod>i\<in>S. f i x) \<longlongrightarrow> (\<Prod>i\<in>S. L i)) F" | 
| 865 | by (induct S rule: infinite_finite_induct) (simp_all add: tendsto_mult) | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 866 | |
| 64272 | 867 | lemma continuous_prod [continuous_intros]: | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 868 |   fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
 | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 869 | shows "(\<And>i. i \<in> S \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Prod>i\<in>S. f i x)" | 
| 64272 | 870 | unfolding continuous_def by (rule tendsto_prod) | 
| 871 | ||
| 872 | lemma continuous_on_prod [continuous_intros]: | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 873 |   fixes f :: "'a \<Rightarrow> _ \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
 | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 874 | shows "(\<And>i. i \<in> S \<Longrightarrow> continuous_on s (f i)) \<Longrightarrow> continuous_on s (\<lambda>x. \<Prod>i\<in>S. f i x)" | 
| 64272 | 875 | unfolding continuous_on_def by (auto intro: tendsto_prod) | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 876 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 877 | lemma tendsto_of_real_iff: | 
| 63546 | 878 | "((\<lambda>x. of_real (f x) :: 'a::real_normed_div_algebra) \<longlongrightarrow> of_real c) F \<longleftrightarrow> (f \<longlongrightarrow> c) F" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 879 | unfolding tendsto_iff by simp | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 880 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 881 | lemma tendsto_add_const_iff: | 
| 63546 | 882 | "((\<lambda>x. c + f x :: 'a::real_normed_vector) \<longlongrightarrow> c + d) F \<longleftrightarrow> (f \<longlongrightarrow> d) F" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 883 | using tendsto_add[OF tendsto_const[of c], of f d] | 
| 63546 | 884 | and tendsto_add[OF tendsto_const[of "-c"], of "\<lambda>x. c + f x" "c + d"] by auto | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 885 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 886 | |
| 68860 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 887 | class topological_monoid_mult = topological_semigroup_mult + monoid_mult | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 888 | class topological_comm_monoid_mult = topological_monoid_mult + comm_monoid_mult | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 889 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 890 | lemma tendsto_power_strong [tendsto_intros]: | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 891 | fixes f :: "_ \<Rightarrow> 'b :: topological_monoid_mult" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 892 | assumes "(f \<longlongrightarrow> a) F" "(g \<longlongrightarrow> b) F" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 893 | shows "((\<lambda>x. f x ^ g x) \<longlongrightarrow> a ^ b) F" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 894 | proof - | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 895 | have "((\<lambda>x. f x ^ b) \<longlongrightarrow> a ^ b) F" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 896 | by (induction b) (auto intro: tendsto_intros assms) | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 897 | also from assms(2) have "eventually (\<lambda>x. g x = b) F" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 898 | by (simp add: nhds_discrete filterlim_principal) | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 899 | hence "eventually (\<lambda>x. f x ^ b = f x ^ g x) F" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 900 | by eventually_elim simp | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 901 | hence "((\<lambda>x. f x ^ b) \<longlongrightarrow> a ^ b) F \<longleftrightarrow> ((\<lambda>x. f x ^ g x) \<longlongrightarrow> a ^ b) F" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 902 | by (intro filterlim_cong refl) | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 903 | finally show ?thesis . | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 904 | qed | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 905 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 906 | lemma continuous_mult' [continuous_intros]: | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 907 | fixes f g :: "_ \<Rightarrow> 'b::topological_semigroup_mult" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 908 | shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x * g x)" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 909 | unfolding continuous_def by (rule tendsto_mult) | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 910 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 911 | lemma continuous_power' [continuous_intros]: | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 912 | fixes f :: "_ \<Rightarrow> 'b::topological_monoid_mult" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 913 | shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x ^ g x)" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 914 | unfolding continuous_def by (rule tendsto_power_strong) auto | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 915 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 916 | lemma continuous_on_mult' [continuous_intros]: | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 917 | fixes f g :: "_ \<Rightarrow> 'b::topological_semigroup_mult" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 918 | shows "continuous_on A f \<Longrightarrow> continuous_on A g \<Longrightarrow> continuous_on A (\<lambda>x. f x * g x)" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 919 | unfolding continuous_on_def by (auto intro: tendsto_mult) | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 920 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 921 | lemma continuous_on_power' [continuous_intros]: | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 922 | fixes f :: "_ \<Rightarrow> 'b::topological_monoid_mult" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 923 | shows "continuous_on A f \<Longrightarrow> continuous_on A g \<Longrightarrow> continuous_on A (\<lambda>x. f x ^ g x)" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 924 | unfolding continuous_on_def by (auto intro: tendsto_power_strong) | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 925 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 926 | lemma tendsto_mult_one: | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 927 | fixes f g :: "_ \<Rightarrow> 'b::topological_monoid_mult" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 928 | shows "(f \<longlongrightarrow> 1) F \<Longrightarrow> (g \<longlongrightarrow> 1) F \<Longrightarrow> ((\<lambda>x. f x * g x) \<longlongrightarrow> 1) F" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 929 | by (drule (1) tendsto_mult) simp | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 930 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 931 | lemma tendsto_prod' [tendsto_intros]: | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 932 | fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::topological_comm_monoid_mult" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 933 | shows "(\<And>i. i \<in> I \<Longrightarrow> (f i \<longlongrightarrow> a i) F) \<Longrightarrow> ((\<lambda>x. \<Prod>i\<in>I. f i x) \<longlongrightarrow> (\<Prod>i\<in>I. a i)) F" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 934 | by (induct I rule: infinite_finite_induct) (simp_all add: tendsto_mult) | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 935 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 936 | lemma tendsto_one_prod': | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 937 | fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::topological_comm_monoid_mult" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 938 | assumes "\<And>i. i \<in> I \<Longrightarrow> ((\<lambda>x. f x i) \<longlongrightarrow> 1) F" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 939 | shows "((\<lambda>i. prod (f i) I) \<longlongrightarrow> 1) F" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 940 | using tendsto_prod' [of I "\<lambda>x y. f y x" "\<lambda>x. 1"] assms by simp | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 941 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 942 | lemma continuous_prod' [continuous_intros]: | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 943 | fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::topological_comm_monoid_mult" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 944 | shows "(\<And>i. i \<in> I \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Prod>i\<in>I. f i x)" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 945 | unfolding continuous_def by (rule tendsto_prod') | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 946 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 947 | lemma continuous_on_prod' [continuous_intros]: | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 948 | fixes f :: "'a \<Rightarrow> 'b::topological_space \<Rightarrow> 'c::topological_comm_monoid_mult" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 949 | shows "(\<And>i. i \<in> I \<Longrightarrow> continuous_on S (f i)) \<Longrightarrow> continuous_on S (\<lambda>x. \<Prod>i\<in>I. f i x)" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 950 | unfolding continuous_on_def by (auto intro: tendsto_prod') | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 951 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 952 | instance nat :: topological_comm_monoid_mult | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 953 | by standard | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 954 | (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 955 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 956 | instance int :: topological_comm_monoid_mult | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 957 | by standard | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 958 | (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 959 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 960 | class comm_real_normed_algebra_1 = real_normed_algebra_1 + comm_monoid_mult | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 961 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 962 | context real_normed_field | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 963 | begin | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 964 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 965 | subclass comm_real_normed_algebra_1 | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 966 | proof | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 967 | from norm_mult[of "1 :: 'a" 1] show "norm 1 = 1" by simp | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 968 | qed (simp_all add: norm_mult) | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 969 | |
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 970 | end | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
68721diff
changeset | 971 | |
| 60758 | 972 | subsubsection \<open>Inverse and division\<close> | 
| 31355 | 973 | |
| 974 | lemma (in bounded_bilinear) Zfun_prod_Bfun: | |
| 44195 | 975 | assumes f: "Zfun f F" | 
| 63546 | 976 | and g: "Bfun g F" | 
| 44195 | 977 | shows "Zfun (\<lambda>x. f x ** g x) F" | 
| 31355 | 978 | proof - | 
| 979 | obtain K where K: "0 \<le> K" | |
| 980 | and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K" | |
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 981 | using nonneg_bounded by blast | 
| 31355 | 982 | obtain B where B: "0 < B" | 
| 44195 | 983 | and norm_g: "eventually (\<lambda>x. norm (g x) \<le> B) F" | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 984 | using g by (rule BfunE) | 
| 44195 | 985 | have "eventually (\<lambda>x. norm (f x ** g x) \<le> norm (f x) * (B * K)) F" | 
| 46887 | 986 | using norm_g proof eventually_elim | 
| 987 | case (elim x) | |
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 988 | have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K" | 
| 31355 | 989 | by (rule norm_le) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 990 | also have "\<dots> \<le> norm (f x) * B * K" | 
| 63546 | 991 | by (intro mult_mono' order_refl norm_g norm_ge_zero mult_nonneg_nonneg K elim) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 992 | also have "\<dots> = norm (f x) * (B * K)" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57447diff
changeset | 993 | by (rule mult.assoc) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 994 | finally show "norm (f x ** g x) \<le> norm (f x) * (B * K)" . | 
| 31355 | 995 | qed | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 996 | with f show ?thesis | 
| 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 997 | by (rule Zfun_imp_Zfun) | 
| 31355 | 998 | qed | 
| 999 | ||
| 1000 | lemma (in bounded_bilinear) Bfun_prod_Zfun: | |
| 44195 | 1001 | assumes f: "Bfun f F" | 
| 63546 | 1002 | and g: "Zfun g F" | 
| 44195 | 1003 | shows "Zfun (\<lambda>x. f x ** g x) F" | 
| 44081 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 huffman parents: 
44079diff
changeset | 1004 | using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun) | 
| 31355 | 1005 | |
| 1006 | lemma Bfun_inverse: | |
| 1007 | fixes a :: "'a::real_normed_div_algebra" | |
| 61973 | 1008 | assumes f: "(f \<longlongrightarrow> a) F" | 
| 31355 | 1009 | assumes a: "a \<noteq> 0" | 
| 44195 | 1010 | shows "Bfun (\<lambda>x. inverse (f x)) F" | 
| 31355 | 1011 | proof - | 
| 1012 | from a have "0 < norm a" by simp | |
| 63546 | 1013 | then have "\<exists>r>0. r < norm a" by (rule dense) | 
| 1014 | then obtain r where r1: "0 < r" and r2: "r < norm a" | |
| 1015 | by blast | |
| 44195 | 1016 | have "eventually (\<lambda>x. dist (f x) a < r) F" | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 1017 | using tendstoD [OF f r1] by blast | 
| 63546 | 1018 | then have "eventually (\<lambda>x. norm (inverse (f x)) \<le> inverse (norm a - r)) F" | 
| 46887 | 1019 | proof eventually_elim | 
| 1020 | case (elim x) | |
| 63546 | 1021 | then have 1: "norm (f x - a) < r" | 
| 31355 | 1022 | by (simp add: dist_norm) | 
| 63546 | 1023 | then have 2: "f x \<noteq> 0" using r2 by auto | 
| 1024 | then have "norm (inverse (f x)) = inverse (norm (f x))" | |
| 31355 | 1025 | by (rule nonzero_norm_inverse) | 
| 1026 | also have "\<dots> \<le> inverse (norm a - r)" | |
| 1027 | proof (rule le_imp_inverse_le) | |
| 63546 | 1028 | show "0 < norm a - r" | 
| 1029 | using r2 by simp | |
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 1030 | have "norm a - norm (f x) \<le> norm (a - f x)" | 
| 31355 | 1031 | by (rule norm_triangle_ineq2) | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 1032 | also have "\<dots> = norm (f x - a)" | 
| 31355 | 1033 | by (rule norm_minus_commute) | 
| 1034 | also have "\<dots> < r" using 1 . | |
| 63546 | 1035 | finally show "norm a - r \<le> norm (f x)" | 
| 1036 | by simp | |
| 31355 | 1037 | qed | 
| 31487 
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
 huffman parents: 
31447diff
changeset | 1038 | finally show "norm (inverse (f x)) \<le> inverse (norm a - r)" . | 
| 31355 | 1039 | qed | 
| 63546 | 1040 | then show ?thesis by (rule BfunI) | 
| 31355 | 1041 | qed | 
| 1042 | ||
| 31565 | 1043 | lemma tendsto_inverse [tendsto_intros]: | 
| 31355 | 1044 | fixes a :: "'a::real_normed_div_algebra" | 
| 61973 | 1045 | assumes f: "(f \<longlongrightarrow> a) F" | 
| 63546 | 1046 | and a: "a \<noteq> 0" | 
| 61973 | 1047 | shows "((\<lambda>x. inverse (f x)) \<longlongrightarrow> inverse a) F" | 
| 31355 | 1048 | proof - | 
| 1049 | from a have "0 < norm a" by simp | |
| 44195 | 1050 | with f have "eventually (\<lambda>x. dist (f x) a < norm a) F" | 
| 31355 | 1051 | by (rule tendstoD) | 
| 44195 | 1052 | then have "eventually (\<lambda>x. f x \<noteq> 0) F" | 
| 61810 | 1053 | unfolding dist_norm by (auto elim!: eventually_mono) | 
| 44627 | 1054 | with a have "eventually (\<lambda>x. inverse (f x) - inverse a = | 
| 1055 | - (inverse (f x) * (f x - a) * inverse a)) F" | |
| 61810 | 1056 | by (auto elim!: eventually_mono simp: inverse_diff_inverse) | 
| 44627 | 1057 | moreover have "Zfun (\<lambda>x. - (inverse (f x) * (f x - a) * inverse a)) F" | 
| 1058 | by (intro Zfun_minus Zfun_mult_left | |
| 1059 | bounded_bilinear.Bfun_prod_Zfun [OF bounded_bilinear_mult] | |
| 1060 | Bfun_inverse [OF f a] f [unfolded tendsto_Zfun_iff]) | |
| 1061 | ultimately show ?thesis | |
| 1062 | unfolding tendsto_Zfun_iff by (rule Zfun_ssubst) | |
| 31355 | 1063 | qed | 
| 1064 | ||
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1065 | lemma continuous_inverse: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1066 | fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra" | 
| 63546 | 1067 | assumes "continuous F f" | 
| 1068 | and "f (Lim F (\<lambda>x. x)) \<noteq> 0" | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1069 | shows "continuous F (\<lambda>x. inverse (f x))" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1070 | using assms unfolding continuous_def by (rule tendsto_inverse) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1071 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1072 | lemma continuous_at_within_inverse[continuous_intros]: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1073 | fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra" | 
| 63546 | 1074 | assumes "continuous (at a within s) f" | 
| 1075 | and "f a \<noteq> 0" | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1076 | shows "continuous (at a within s) (\<lambda>x. inverse (f x))" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1077 | using assms unfolding continuous_within by (rule tendsto_inverse) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1078 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56366diff
changeset | 1079 | lemma continuous_on_inverse[continuous_intros]: | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1080 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra" | 
| 63546 | 1081 | assumes "continuous_on s f" | 
| 1082 | and "\<forall>x\<in>s. f x \<noteq> 0" | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1083 | shows "continuous_on s (\<lambda>x. inverse (f x))" | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 1084 | using assms unfolding continuous_on_def by (blast intro: tendsto_inverse) | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1085 | |
| 31565 | 1086 | lemma tendsto_divide [tendsto_intros]: | 
| 31355 | 1087 | fixes a b :: "'a::real_normed_field" | 
| 63546 | 1088 | shows "(f \<longlongrightarrow> a) F \<Longrightarrow> (g \<longlongrightarrow> b) F \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> ((\<lambda>x. f x / g x) \<longlongrightarrow> a / b) F" | 
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44253diff
changeset | 1089 | by (simp add: tendsto_mult tendsto_inverse divide_inverse) | 
| 31355 | 1090 | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1091 | lemma continuous_divide: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1092 | fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_field" | 
| 63546 | 1093 | assumes "continuous F f" | 
| 1094 | and "continuous F g" | |
| 1095 | and "g (Lim F (\<lambda>x. x)) \<noteq> 0" | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1096 | shows "continuous F (\<lambda>x. (f x) / (g x))" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1097 | using assms unfolding continuous_def by (rule tendsto_divide) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1098 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1099 | lemma continuous_at_within_divide[continuous_intros]: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1100 | fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_field" | 
| 63546 | 1101 | assumes "continuous (at a within s) f" "continuous (at a within s) g" | 
| 1102 | and "g a \<noteq> 0" | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1103 | shows "continuous (at a within s) (\<lambda>x. (f x) / (g x))" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1104 | using assms unfolding continuous_within by (rule tendsto_divide) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1105 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1106 | lemma isCont_divide[continuous_intros, simp]: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1107 | fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_field" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1108 | assumes "isCont f a" "isCont g a" "g a \<noteq> 0" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1109 | shows "isCont (\<lambda>x. (f x) / g x) a" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1110 | using assms unfolding continuous_at by (rule tendsto_divide) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1111 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56366diff
changeset | 1112 | lemma continuous_on_divide[continuous_intros]: | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1113 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_field" | 
| 63546 | 1114 | assumes "continuous_on s f" "continuous_on s g" | 
| 1115 | and "\<forall>x\<in>s. g x \<noteq> 0" | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1116 | shows "continuous_on s (\<lambda>x. (f x) / (g x))" | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 1117 | using assms unfolding continuous_on_def by (blast intro: tendsto_divide) | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1118 | |
| 63546 | 1119 | lemma tendsto_sgn [tendsto_intros]: "(f \<longlongrightarrow> l) F \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> ((\<lambda>x. sgn (f x)) \<longlongrightarrow> sgn l) F" | 
| 1120 | for l :: "'a::real_normed_vector" | |
| 44194 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 1121 | unfolding sgn_div_norm by (simp add: tendsto_intros) | 
| 
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
 huffman parents: 
44081diff
changeset | 1122 | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1123 | lemma continuous_sgn: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1124 | fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" | 
| 63546 | 1125 | assumes "continuous F f" | 
| 1126 | and "f (Lim F (\<lambda>x. x)) \<noteq> 0" | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1127 | shows "continuous F (\<lambda>x. sgn (f x))" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1128 | using assms unfolding continuous_def by (rule tendsto_sgn) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1129 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1130 | lemma continuous_at_within_sgn[continuous_intros]: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1131 | fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" | 
| 63546 | 1132 | assumes "continuous (at a within s) f" | 
| 1133 | and "f a \<noteq> 0" | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1134 | shows "continuous (at a within s) (\<lambda>x. sgn (f x))" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1135 | using assms unfolding continuous_within by (rule tendsto_sgn) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1136 | |
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1137 | lemma isCont_sgn[continuous_intros]: | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1138 | fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" | 
| 63546 | 1139 | assumes "isCont f a" | 
| 1140 | and "f a \<noteq> 0" | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1141 | shows "isCont (\<lambda>x. sgn (f x)) a" | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1142 | using assms unfolding continuous_at by (rule tendsto_sgn) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1143 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56366diff
changeset | 1144 | lemma continuous_on_sgn[continuous_intros]: | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1145 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" | 
| 63546 | 1146 | assumes "continuous_on s f" | 
| 1147 | and "\<forall>x\<in>s. f x \<noteq> 0" | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1148 | shows "continuous_on s (\<lambda>x. sgn (f x))" | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 1149 | using assms unfolding continuous_on_def by (blast intro: tendsto_sgn) | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51474diff
changeset | 1150 | |
| 50325 | 1151 | lemma filterlim_at_infinity: | 
| 61076 | 1152 | fixes f :: "_ \<Rightarrow> 'a::real_normed_vector" | 
| 50325 | 1153 | assumes "0 \<le> c" | 
| 1154 | shows "(LIM x F. f x :> at_infinity) \<longleftrightarrow> (\<forall>r>c. eventually (\<lambda>x. r \<le> norm (f x)) F)" | |
| 1155 | unfolding filterlim_iff eventually_at_infinity | |
| 1156 | proof safe | |
| 63546 | 1157 | fix P :: "'a \<Rightarrow> bool" | 
| 1158 | fix b | |
| 50325 | 1159 | assume *: "\<forall>r>c. eventually (\<lambda>x. r \<le> norm (f x)) F" | 
| 63546 | 1160 | assume P: "\<forall>x. b \<le> norm x \<longrightarrow> P x" | 
| 50325 | 1161 | have "max b (c + 1) > c" by auto | 
| 1162 | with * have "eventually (\<lambda>x. max b (c + 1) \<le> norm (f x)) F" | |
| 1163 | by auto | |
| 1164 | then show "eventually (\<lambda>x. P (f x)) F" | |
| 1165 | proof eventually_elim | |
| 63546 | 1166 | case (elim x) | 
| 50325 | 1167 | with P show "P (f x)" by auto | 
| 1168 | qed | |
| 1169 | qed force | |
| 1170 | ||
| 67371 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1171 | lemma filterlim_at_infinity_imp_norm_at_top: | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1172 | fixes F | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1173 | assumes "filterlim f at_infinity F" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1174 | shows "filterlim (\<lambda>x. norm (f x)) at_top F" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1175 | proof - | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1176 |   {
 | 
| 68611 | 1177 | fix r :: real | 
| 1178 | have "\<forall>\<^sub>F x in F. r \<le> norm (f x)" using filterlim_at_infinity[of 0 f F] assms | |
| 1179 | by (cases "r > 0") | |
| 67371 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1180 | (auto simp: not_less intro: always_eventually order.trans[OF _ norm_ge_zero]) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1181 | } | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1182 | thus ?thesis by (auto simp: filterlim_at_top) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1183 | qed | 
| 68611 | 1184 | |
| 67371 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1185 | lemma filterlim_norm_at_top_imp_at_infinity: | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1186 | fixes F | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1187 | assumes "filterlim (\<lambda>x. norm (f x)) at_top F" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1188 | shows "filterlim f at_infinity F" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1189 | using filterlim_at_infinity[of 0 f F] assms by (auto simp: filterlim_at_top) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1190 | |
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1191 | lemma filterlim_norm_at_top: "filterlim norm at_top at_infinity" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1192 | by (rule filterlim_at_infinity_imp_norm_at_top) (rule filterlim_ident) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1193 | |
| 67950 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1194 | lemma filterlim_at_infinity_conv_norm_at_top: | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1195 | "filterlim f at_infinity G \<longleftrightarrow> filterlim (\<lambda>x. norm (f x)) at_top G" | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1196 | by (auto simp: filterlim_at_infinity[OF order.refl] filterlim_at_top_gt[of _ _ 0]) | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1197 | |
| 67371 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1198 | lemma eventually_not_equal_at_infinity: | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1199 |   "eventually (\<lambda>x. x \<noteq> (a :: 'a :: {real_normed_vector})) at_infinity"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1200 | proof - | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1201 | from filterlim_norm_at_top[where 'a = 'a] | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1202 | have "\<forall>\<^sub>F x in at_infinity. norm a < norm (x::'a)" by (auto simp: filterlim_at_top_dense) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1203 | thus ?thesis by eventually_elim auto | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1204 | qed | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1205 | |
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1206 | lemma filterlim_int_of_nat_at_topD: | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1207 | fixes F | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1208 | assumes "filterlim (\<lambda>x. f (int x)) F at_top" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1209 | shows "filterlim f F at_top" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1210 | proof - | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1211 | have "filterlim (\<lambda>x. f (int (nat x))) F at_top" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1212 | by (rule filterlim_compose[OF assms filterlim_nat_sequentially]) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1213 | also have "?this \<longleftrightarrow> filterlim f F at_top" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1214 | by (intro filterlim_cong refl eventually_mono [OF eventually_ge_at_top[of "0::int"]]) auto | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1215 | finally show ?thesis . | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1216 | qed | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1217 | |
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1218 | lemma filterlim_int_sequentially [tendsto_intros]: | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1219 | "filterlim int at_top sequentially" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1220 | unfolding filterlim_at_top | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1221 | proof | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1222 | fix C :: int | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1223 | show "eventually (\<lambda>n. int n \<ge> C) at_top" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1224 | using eventually_ge_at_top[of "nat \<lceil>C\<rceil>"] by eventually_elim linarith | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1225 | qed | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1226 | |
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1227 | lemma filterlim_real_of_int_at_top [tendsto_intros]: | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1228 | "filterlim real_of_int at_top at_top" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1229 | unfolding filterlim_at_top | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1230 | proof | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1231 | fix C :: real | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1232 | show "eventually (\<lambda>n. real_of_int n \<ge> C) at_top" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1233 | using eventually_ge_at_top[of "\<lceil>C\<rceil>"] by eventually_elim linarith | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1234 | qed | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1235 | |
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1236 | lemma filterlim_abs_real: "filterlim (abs::real \<Rightarrow> real) at_top at_top" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1237 | proof (subst filterlim_cong[OF refl refl]) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1238 | from eventually_ge_at_top[of "0::real"] show "eventually (\<lambda>x::real. \<bar>x\<bar> = x) at_top" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1239 | by eventually_elim simp | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1240 | qed (simp_all add: filterlim_ident) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1241 | |
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1242 | lemma filterlim_of_real_at_infinity [tendsto_intros]: | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1243 | "filterlim (of_real :: real \<Rightarrow> 'a :: real_normed_algebra_1) at_infinity at_top" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1244 | by (intro filterlim_norm_at_top_imp_at_infinity) (auto simp: filterlim_abs_real) | 
| 68611 | 1245 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1246 | lemma not_tendsto_and_filterlim_at_infinity: | 
| 63546 | 1247 | fixes c :: "'a::real_normed_vector" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1248 | assumes "F \<noteq> bot" | 
| 63546 | 1249 | and "(f \<longlongrightarrow> c) F" | 
| 1250 | and "filterlim f at_infinity F" | |
| 1251 | shows False | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1252 | proof - | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 1253 | from tendstoD[OF assms(2), of "1/2"] | 
| 63546 | 1254 | have "eventually (\<lambda>x. dist (f x) c < 1/2) F" | 
| 1255 | by simp | |
| 1256 | moreover | |
| 1257 | from filterlim_at_infinity[of "norm c" f F] assms(3) | |
| 1258 | have "eventually (\<lambda>x. norm (f x) \<ge> norm c + 1) F" by simp | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1259 | ultimately have "eventually (\<lambda>x. False) F" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1260 | proof eventually_elim | 
| 63546 | 1261 | fix x | 
| 1262 | assume A: "dist (f x) c < 1/2" | |
| 1263 | assume "norm (f x) \<ge> norm c + 1" | |
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62369diff
changeset | 1264 | also have "norm (f x) = dist (f x) 0" by simp | 
| 63546 | 1265 | also have "\<dots> \<le> dist (f x) c + dist c 0" by (rule dist_triangle) | 
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62369diff
changeset | 1266 | finally show False using A by simp | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1267 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1268 | with assms show False by simp | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1269 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1270 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1271 | lemma filterlim_at_infinity_imp_not_convergent: | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1272 | assumes "filterlim f at_infinity sequentially" | 
| 63546 | 1273 | shows "\<not> convergent f" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1274 | by (rule notI, rule not_tendsto_and_filterlim_at_infinity[OF _ _ assms]) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1275 | (simp_all add: convergent_LIMSEQ_iff) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1276 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1277 | lemma filterlim_at_infinity_imp_eventually_ne: | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1278 | assumes "filterlim f at_infinity F" | 
| 63546 | 1279 | shows "eventually (\<lambda>z. f z \<noteq> c) F" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1280 | proof - | 
| 63546 | 1281 | have "norm c + 1 > 0" | 
| 1282 | by (intro add_nonneg_pos) simp_all | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1283 | with filterlim_at_infinity[OF order.refl, of f F] assms | 
| 63546 | 1284 | have "eventually (\<lambda>z. norm (f z) \<ge> norm c + 1) F" | 
| 1285 | by blast | |
| 1286 | then show ?thesis | |
| 1287 | by eventually_elim auto | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1288 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1289 | |
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 1290 | lemma tendsto_of_nat [tendsto_intros]: | 
| 63546 | 1291 | "filterlim (of_nat :: nat \<Rightarrow> 'a::real_normed_algebra_1) at_infinity sequentially" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1292 | proof (subst filterlim_at_infinity[OF order.refl], intro allI impI) | 
| 63040 | 1293 | fix r :: real | 
| 1294 | assume r: "r > 0" | |
| 1295 | define n where "n = nat \<lceil>r\<rceil>" | |
| 63546 | 1296 | from r have n: "\<forall>m\<ge>n. of_nat m \<ge> r" | 
| 1297 | unfolding n_def by linarith | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1298 | from eventually_ge_at_top[of n] show "eventually (\<lambda>m. norm (of_nat m :: 'a) \<ge> r) sequentially" | 
| 63546 | 1299 | by eventually_elim (use n in simp_all) | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1300 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1301 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1302 | |
| 69593 | 1303 | subsection \<open>Relate \<^const>\<open>at\<close>, \<^const>\<open>at_left\<close> and \<^const>\<open>at_right\<close>\<close> | 
| 50347 | 1304 | |
| 60758 | 1305 | text \<open> | 
| 69593 | 1306 | This lemmas are useful for conversion between \<^term>\<open>at x\<close> to \<^term>\<open>at_left x\<close> and | 
| 1307 | \<^term>\<open>at_right x\<close> and also \<^term>\<open>at_right 0\<close>. | |
| 60758 | 1308 | \<close> | 
| 50347 | 1309 | |
| 51471 | 1310 | lemmas filterlim_split_at_real = filterlim_split_at[where 'a=real] | 
| 50323 | 1311 | |
| 63546 | 1312 | lemma filtermap_nhds_shift: "filtermap (\<lambda>x. x - d) (nhds a) = nhds (a - d)" | 
| 1313 | for a d :: "'a::real_normed_vector" | |
| 60721 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1314 | by (rule filtermap_fun_inverse[where g="\<lambda>x. x + d"]) | 
| 63546 | 1315 | (auto intro!: tendsto_eq_intros filterlim_ident) | 
| 1316 | ||
| 1317 | lemma filtermap_nhds_minus: "filtermap (\<lambda>x. - x) (nhds a) = nhds (- a)" | |
| 1318 | for a :: "'a::real_normed_vector" | |
| 60721 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1319 | by (rule filtermap_fun_inverse[where g=uminus]) | 
| 63546 | 1320 | (auto intro!: tendsto_eq_intros filterlim_ident) | 
| 1321 | ||
| 1322 | lemma filtermap_at_shift: "filtermap (\<lambda>x. x - d) (at a) = at (a - d)" | |
| 1323 | for a d :: "'a::real_normed_vector" | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51531diff
changeset | 1324 | by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) | 
| 50347 | 1325 | |
| 63546 | 1326 | lemma filtermap_at_right_shift: "filtermap (\<lambda>x. x - d) (at_right a) = at_right (a - d)" | 
| 1327 | for a d :: "real" | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51531diff
changeset | 1328 | by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) | 
| 50323 | 1329 | |
| 63546 | 1330 | lemma at_right_to_0: "at_right a = filtermap (\<lambda>x. x + a) (at_right 0)" | 
| 1331 | for a :: real | |
| 50347 | 1332 | using filtermap_at_right_shift[of "-a" 0] by simp | 
| 1333 | ||
| 1334 | lemma filterlim_at_right_to_0: | |
| 63546 | 1335 | "filterlim f F (at_right a) \<longleftrightarrow> filterlim (\<lambda>x. f (x + a)) F (at_right 0)" | 
| 1336 | for a :: real | |
| 50347 | 1337 | unfolding filterlim_def filtermap_filtermap at_right_to_0[of a] .. | 
| 1338 | ||
| 1339 | lemma eventually_at_right_to_0: | |
| 63546 | 1340 | "eventually P (at_right a) \<longleftrightarrow> eventually (\<lambda>x. P (x + a)) (at_right 0)" | 
| 1341 | for a :: real | |
| 50347 | 1342 | unfolding at_right_to_0[of a] by (simp add: eventually_filtermap) | 
| 1343 | ||
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1344 | lemma at_to_0: "at a = filtermap (\<lambda>x. x + a) (at 0)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1345 | for a :: "'a::real_normed_vector" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1346 | using filtermap_at_shift[of "-a" 0] by simp | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1347 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1348 | lemma filterlim_at_to_0: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1349 | "filterlim f F (at a) \<longleftrightarrow> filterlim (\<lambda>x. f (x + a)) F (at 0)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1350 | for a :: "'a::real_normed_vector" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1351 | unfolding filterlim_def filtermap_filtermap at_to_0[of a] .. | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1352 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1353 | lemma eventually_at_to_0: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1354 | "eventually P (at a) \<longleftrightarrow> eventually (\<lambda>x. P (x + a)) (at 0)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1355 | for a :: "'a::real_normed_vector" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1356 | unfolding at_to_0[of a] by (simp add: eventually_filtermap) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1357 | |
| 63546 | 1358 | lemma filtermap_at_minus: "filtermap (\<lambda>x. - x) (at a) = at (- a)" | 
| 1359 | for a :: "'a::real_normed_vector" | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51531diff
changeset | 1360 | by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) | 
| 50347 | 1361 | |
| 63546 | 1362 | lemma at_left_minus: "at_left a = filtermap (\<lambda>x. - x) (at_right (- a))" | 
| 1363 | for a :: real | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51531diff
changeset | 1364 | by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) | 
| 50323 | 1365 | |
| 63546 | 1366 | lemma at_right_minus: "at_right a = filtermap (\<lambda>x. - x) (at_left (- a))" | 
| 1367 | for a :: real | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51531diff
changeset | 1368 | by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) | 
| 50347 | 1369 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1370 | |
| 50347 | 1371 | lemma filterlim_at_left_to_right: | 
| 63546 | 1372 | "filterlim f F (at_left a) \<longleftrightarrow> filterlim (\<lambda>x. f (- x)) F (at_right (-a))" | 
| 1373 | for a :: real | |
| 50347 | 1374 | unfolding filterlim_def filtermap_filtermap at_left_minus[of a] .. | 
| 1375 | ||
| 1376 | lemma eventually_at_left_to_right: | |
| 63546 | 1377 | "eventually P (at_left a) \<longleftrightarrow> eventually (\<lambda>x. P (- x)) (at_right (-a))" | 
| 1378 | for a :: real | |
| 50347 | 1379 | unfolding at_left_minus[of a] by (simp add: eventually_filtermap) | 
| 1380 | ||
| 60721 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1381 | lemma filterlim_uminus_at_top_at_bot: "LIM x at_bot. - x :: real :> at_top" | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1382 | unfolding filterlim_at_top eventually_at_bot_dense | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1383 | by (metis leI minus_less_iff order_less_asym) | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1384 | |
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1385 | lemma filterlim_uminus_at_bot_at_top: "LIM x at_top. - x :: real :> at_bot" | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1386 | unfolding filterlim_at_bot eventually_at_top_dense | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1387 | by (metis leI less_minus_iff order_less_asym) | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1388 | |
| 68611 | 1389 | lemma at_bot_mirror : | 
| 1390 |   shows "(at_bot::('a::{ordered_ab_group_add,linorder} filter)) = filtermap uminus at_top"
 | |
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68296diff
changeset | 1391 | apply (rule filtermap_fun_inverse[of uminus, symmetric]) | 
| 68615 | 1392 | subgoal unfolding filterlim_at_top filterlim_at_bot eventually_at_bot_linorder using le_minus_iff by auto | 
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68296diff
changeset | 1393 | subgoal unfolding filterlim_at_bot eventually_at_top_linorder using minus_le_iff by auto | 
| 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68296diff
changeset | 1394 | by auto | 
| 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68296diff
changeset | 1395 | |
| 68611 | 1396 | lemma at_top_mirror : | 
| 1397 |   shows "(at_top::('a::{ordered_ab_group_add,linorder} filter)) = filtermap uminus at_bot"
 | |
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68296diff
changeset | 1398 | apply (subst at_bot_mirror) | 
| 68615 | 1399 | by (auto simp: filtermap_filtermap) | 
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1400 | |
| 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1401 | lemma filterlim_at_top_mirror: "(LIM x at_top. f x :> F) \<longleftrightarrow> (LIM x at_bot. f (-x::real) :> F)" | 
| 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1402 | unfolding filterlim_def at_top_mirror filtermap_filtermap .. | 
| 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1403 | |
| 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1404 | lemma filterlim_at_bot_mirror: "(LIM x at_bot. f x :> F) \<longleftrightarrow> (LIM x at_top. f (-x::real) :> F)" | 
| 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1405 | unfolding filterlim_def at_bot_mirror filtermap_filtermap .. | 
| 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1406 | |
| 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1407 | lemma filterlim_uminus_at_top: "(LIM x F. f x :> at_top) \<longleftrightarrow> (LIM x F. - (f x) :: real :> at_bot)" | 
| 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1408 | using filterlim_compose[OF filterlim_uminus_at_bot_at_top, of f F] | 
| 63546 | 1409 | and filterlim_compose[OF filterlim_uminus_at_top_at_bot, of "\<lambda>x. - f x" F] | 
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1410 | by auto | 
| 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1411 | |
| 68721 | 1412 | lemma tendsto_at_botI_sequentially: | 
| 1413 | fixes f :: "real \<Rightarrow> 'b::first_countable_topology" | |
| 1414 | assumes *: "\<And>X. filterlim X at_bot sequentially \<Longrightarrow> (\<lambda>n. f (X n)) \<longlonglongrightarrow> y" | |
| 1415 | shows "(f \<longlongrightarrow> y) at_bot" | |
| 1416 | unfolding filterlim_at_bot_mirror | |
| 1417 | proof (rule tendsto_at_topI_sequentially) | |
| 1418 | fix X :: "nat \<Rightarrow> real" assume "filterlim X at_top sequentially" | |
| 1419 | thus "(\<lambda>n. f (-X n)) \<longlonglongrightarrow> y" by (intro *) (auto simp: filterlim_uminus_at_top) | |
| 1420 | qed | |
| 1421 | ||
| 67950 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1422 | lemma filterlim_at_infinity_imp_filterlim_at_top: | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1423 | assumes "filterlim (f :: 'a \<Rightarrow> real) at_infinity F" | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1424 | assumes "eventually (\<lambda>x. f x > 0) F" | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1425 | shows "filterlim f at_top F" | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1426 | proof - | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1427 | from assms(2) have *: "eventually (\<lambda>x. norm (f x) = f x) F" by eventually_elim simp | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1428 | from assms(1) show ?thesis unfolding filterlim_at_infinity_conv_norm_at_top | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1429 | by (subst (asm) filterlim_cong[OF refl refl *]) | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1430 | qed | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1431 | |
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1432 | lemma filterlim_at_infinity_imp_filterlim_at_bot: | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1433 | assumes "filterlim (f :: 'a \<Rightarrow> real) at_infinity F" | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1434 | assumes "eventually (\<lambda>x. f x < 0) F" | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1435 | shows "filterlim f at_bot F" | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1436 | proof - | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1437 | from assms(2) have *: "eventually (\<lambda>x. norm (f x) = -f x) F" by eventually_elim simp | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1438 | from assms(1) have "filterlim (\<lambda>x. - f x) at_top F" | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1439 | unfolding filterlim_at_infinity_conv_norm_at_top | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1440 | by (subst (asm) filterlim_cong[OF refl refl *]) | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1441 | thus ?thesis by (simp add: filterlim_uminus_at_top) | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1442 | qed | 
| 
99eaa5cedbb7
Added some simple facts about limits
 Manuel Eberl <eberlm@in.tum.de> parents: 
67707diff
changeset | 1443 | |
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1444 | lemma filterlim_uminus_at_bot: "(LIM x F. f x :> at_bot) \<longleftrightarrow> (LIM x F. - (f x) :: real :> at_top)" | 
| 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1445 | unfolding filterlim_uminus_at_top by simp | 
| 50323 | 1446 | |
| 50347 | 1447 | lemma filterlim_inverse_at_top_right: "LIM x at_right (0::real). inverse x :> at_top" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51531diff
changeset | 1448 | unfolding filterlim_at_top_gt[where c=0] eventually_at_filter | 
| 50347 | 1449 | proof safe | 
| 63546 | 1450 | fix Z :: real | 
| 1451 | assume [arith]: "0 < Z" | |
| 50347 | 1452 | then have "eventually (\<lambda>x. x < inverse Z) (nhds 0)" | 
| 68615 | 1453 | by (auto simp: eventually_nhds_metric dist_real_def intro!: exI[of _ "\<bar>inverse Z\<bar>"]) | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51531diff
changeset | 1454 |   then show "eventually (\<lambda>x. x \<noteq> 0 \<longrightarrow> x \<in> {0<..} \<longrightarrow> Z \<le> inverse x) (nhds 0)"
 | 
| 61810 | 1455 | by (auto elim!: eventually_mono simp: inverse_eq_divide field_simps) | 
| 50347 | 1456 | qed | 
| 1457 | ||
| 50325 | 1458 | lemma tendsto_inverse_0: | 
| 61076 | 1459 | fixes x :: "_ \<Rightarrow> 'a::real_normed_div_algebra" | 
| 61973 | 1460 | shows "(inverse \<longlongrightarrow> (0::'a)) at_infinity" | 
| 50325 | 1461 | unfolding tendsto_Zfun_iff diff_0_right Zfun_def eventually_at_infinity | 
| 1462 | proof safe | |
| 63546 | 1463 | fix r :: real | 
| 1464 | assume "0 < r" | |
| 50325 | 1465 | show "\<exists>b. \<forall>x. b \<le> norm x \<longrightarrow> norm (inverse x :: 'a) < r" | 
| 1466 | proof (intro exI[of _ "inverse (r / 2)"] allI impI) | |
| 1467 | fix x :: 'a | |
| 60758 | 1468 | from \<open>0 < r\<close> have "0 < inverse (r / 2)" by simp | 
| 50325 | 1469 | also assume *: "inverse (r / 2) \<le> norm x" | 
| 1470 | finally show "norm (inverse x) < r" | |
| 63546 | 1471 | using * \<open>0 < r\<close> | 
| 1472 | by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps) | |
| 50325 | 1473 | qed | 
| 1474 | qed | |
| 1475 | ||
| 61552 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 eberlm parents: 
61531diff
changeset | 1476 | lemma tendsto_add_filterlim_at_infinity: | 
| 63546 | 1477 | fixes c :: "'b::real_normed_vector" | 
| 1478 | and F :: "'a filter" | |
| 1479 | assumes "(f \<longlongrightarrow> c) F" | |
| 1480 | and "filterlim g at_infinity F" | |
| 1481 | shows "filterlim (\<lambda>x. f x + g x) at_infinity F" | |
| 61552 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 eberlm parents: 
61531diff
changeset | 1482 | proof (subst filterlim_at_infinity[OF order_refl], safe) | 
| 63546 | 1483 | fix r :: real | 
| 1484 | assume r: "r > 0" | |
| 1485 | from assms(1) have "((\<lambda>x. norm (f x)) \<longlongrightarrow> norm c) F" | |
| 1486 | by (rule tendsto_norm) | |
| 1487 | then have "eventually (\<lambda>x. norm (f x) < norm c + 1) F" | |
| 1488 | by (rule order_tendstoD) simp_all | |
| 1489 | moreover from r have "r + norm c + 1 > 0" | |
| 1490 | by (intro add_pos_nonneg) simp_all | |
| 61552 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 eberlm parents: 
61531diff
changeset | 1491 | with assms(2) have "eventually (\<lambda>x. norm (g x) \<ge> r + norm c + 1) F" | 
| 63546 | 1492 | unfolding filterlim_at_infinity[OF order_refl] | 
| 1493 | by (elim allE[of _ "r + norm c + 1"]) simp_all | |
| 61552 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 eberlm parents: 
61531diff
changeset | 1494 | ultimately show "eventually (\<lambda>x. norm (f x + g x) \<ge> r) F" | 
| 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 eberlm parents: 
61531diff
changeset | 1495 | proof eventually_elim | 
| 63546 | 1496 | fix x :: 'a | 
| 1497 | assume A: "norm (f x) < norm c + 1" and B: "r + norm c + 1 \<le> norm (g x)" | |
| 1498 | from A B have "r \<le> norm (g x) - norm (f x)" | |
| 1499 | by simp | |
| 1500 | also have "norm (g x) - norm (f x) \<le> norm (g x + f x)" | |
| 1501 | by (rule norm_diff_ineq) | |
| 1502 | finally show "r \<le> norm (f x + g x)" | |
| 1503 | by (simp add: add_ac) | |
| 61552 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 eberlm parents: 
61531diff
changeset | 1504 | qed | 
| 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 eberlm parents: 
61531diff
changeset | 1505 | qed | 
| 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 eberlm parents: 
61531diff
changeset | 1506 | |
| 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 eberlm parents: 
61531diff
changeset | 1507 | lemma tendsto_add_filterlim_at_infinity': | 
| 63546 | 1508 | fixes c :: "'b::real_normed_vector" | 
| 1509 | and F :: "'a filter" | |
| 61552 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 eberlm parents: 
61531diff
changeset | 1510 | assumes "filterlim f at_infinity F" | 
| 63546 | 1511 | and "(g \<longlongrightarrow> c) F" | 
| 1512 | shows "filterlim (\<lambda>x. f x + g x) at_infinity F" | |
| 61552 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 eberlm parents: 
61531diff
changeset | 1513 | by (subst add.commute) (rule tendsto_add_filterlim_at_infinity assms)+ | 
| 
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
 eberlm parents: 
61531diff
changeset | 1514 | |
| 60721 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1515 | lemma filterlim_inverse_at_right_top: "LIM x at_top. inverse x :> at_right (0::real)" | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1516 | unfolding filterlim_at | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1517 | by (auto simp: eventually_at_top_dense) | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1518 | (metis tendsto_inverse_0 filterlim_mono at_top_le_at_infinity order_refl) | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1519 | |
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1520 | lemma filterlim_inverse_at_top: | 
| 61973 | 1521 | "(f \<longlongrightarrow> (0 :: real)) F \<Longrightarrow> eventually (\<lambda>x. 0 < f x) F \<Longrightarrow> LIM x F. inverse (f x) :> at_top" | 
| 60721 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1522 | by (intro filterlim_compose[OF filterlim_inverse_at_top_right]) | 
| 61810 | 1523 | (simp add: filterlim_def eventually_filtermap eventually_mono at_within_def le_principal) | 
| 60721 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1524 | |
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1525 | lemma filterlim_inverse_at_bot_neg: | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1526 | "LIM x (at_left (0::real)). inverse x :> at_bot" | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1527 | by (simp add: filterlim_inverse_at_top_right filterlim_uminus_at_bot filterlim_at_left_to_right) | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1528 | |
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1529 | lemma filterlim_inverse_at_bot: | 
| 61973 | 1530 | "(f \<longlongrightarrow> (0 :: real)) F \<Longrightarrow> eventually (\<lambda>x. f x < 0) F \<Longrightarrow> LIM x F. inverse (f x) :> at_bot" | 
| 60721 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1531 | unfolding filterlim_uminus_at_bot inverse_minus_eq[symmetric] | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1532 | by (rule filterlim_inverse_at_top) (simp_all add: tendsto_minus_cancel_left[symmetric]) | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1533 | |
| 50347 | 1534 | lemma at_right_to_top: "(at_right (0::real)) = filtermap inverse at_top" | 
| 60721 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1535 | by (intro filtermap_fun_inverse[symmetric, where g=inverse]) | 
| 
c1b7793c23a3
generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
 hoelzl parents: 
60182diff
changeset | 1536 | (auto intro: filterlim_inverse_at_top_right filterlim_inverse_at_right_top) | 
| 50347 | 1537 | |
| 1538 | lemma eventually_at_right_to_top: | |
| 1539 | "eventually P (at_right (0::real)) \<longleftrightarrow> eventually (\<lambda>x. P (inverse x)) at_top" | |
| 1540 | unfolding at_right_to_top eventually_filtermap .. | |
| 1541 | ||
| 1542 | lemma filterlim_at_right_to_top: | |
| 1543 | "filterlim f F (at_right (0::real)) \<longleftrightarrow> (LIM x at_top. f (inverse x) :> F)" | |
| 1544 | unfolding filterlim_def at_right_to_top filtermap_filtermap .. | |
| 1545 | ||
| 1546 | lemma at_top_to_right: "at_top = filtermap inverse (at_right (0::real))" | |
| 1547 | unfolding at_right_to_top filtermap_filtermap inverse_inverse_eq filtermap_ident .. | |
| 1548 | ||
| 1549 | lemma eventually_at_top_to_right: | |
| 1550 | "eventually P at_top \<longleftrightarrow> eventually (\<lambda>x. P (inverse x)) (at_right (0::real))" | |
| 1551 | unfolding at_top_to_right eventually_filtermap .. | |
| 1552 | ||
| 1553 | lemma filterlim_at_top_to_right: | |
| 1554 | "filterlim f F at_top \<longleftrightarrow> (LIM x (at_right (0::real)). f (inverse x) :> F)" | |
| 1555 | unfolding filterlim_def at_top_to_right filtermap_filtermap .. | |
| 1556 | ||
| 50325 | 1557 | lemma filterlim_inverse_at_infinity: | 
| 61076 | 1558 |   fixes x :: "_ \<Rightarrow> 'a::{real_normed_div_algebra, division_ring}"
 | 
| 50325 | 1559 | shows "filterlim inverse at_infinity (at (0::'a))" | 
| 1560 | unfolding filterlim_at_infinity[OF order_refl] | |
| 1561 | proof safe | |
| 63546 | 1562 | fix r :: real | 
| 1563 | assume "0 < r" | |
| 50325 | 1564 | then show "eventually (\<lambda>x::'a. r \<le> norm (inverse x)) (at 0)" | 
| 1565 | unfolding eventually_at norm_inverse | |
| 1566 | by (intro exI[of _ "inverse r"]) | |
| 1567 | (auto simp: norm_conv_dist[symmetric] field_simps inverse_eq_divide) | |
| 1568 | qed | |
| 1569 | ||
| 1570 | lemma filterlim_inverse_at_iff: | |
| 61076 | 1571 |   fixes g :: "'a \<Rightarrow> 'b::{real_normed_div_algebra, division_ring}"
 | 
| 50325 | 1572 | shows "(LIM x F. inverse (g x) :> at 0) \<longleftrightarrow> (LIM x F. g x :> at_infinity)" | 
| 1573 | unfolding filterlim_def filtermap_filtermap[symmetric] | |
| 1574 | proof | |
| 1575 | assume "filtermap g F \<le> at_infinity" | |
| 1576 | then have "filtermap inverse (filtermap g F) \<le> filtermap inverse at_infinity" | |
| 1577 | by (rule filtermap_mono) | |
| 1578 | also have "\<dots> \<le> at 0" | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51531diff
changeset | 1579 | using tendsto_inverse_0[where 'a='b] | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51531diff
changeset | 1580 | by (auto intro!: exI[of _ 1] | 
| 63546 | 1581 | simp: le_principal eventually_filtermap filterlim_def at_within_def eventually_at_infinity) | 
| 50325 | 1582 | finally show "filtermap inverse (filtermap g F) \<le> at 0" . | 
| 1583 | next | |
| 1584 | assume "filtermap inverse (filtermap g F) \<le> at 0" | |
| 1585 | then have "filtermap inverse (filtermap inverse (filtermap g F)) \<le> filtermap inverse (at 0)" | |
| 1586 | by (rule filtermap_mono) | |
| 1587 | with filterlim_inverse_at_infinity show "filtermap g F \<le> at_infinity" | |
| 1588 | by (auto intro: order_trans simp: filterlim_def filtermap_filtermap) | |
| 1589 | qed | |
| 1590 | ||
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1591 | lemma tendsto_mult_filterlim_at_infinity: | 
| 63546 | 1592 | fixes c :: "'a::real_normed_field" | 
| 64394 | 1593 | assumes "(f \<longlongrightarrow> c) F" "c \<noteq> 0" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1594 | assumes "filterlim g at_infinity F" | 
| 63546 | 1595 | shows "filterlim (\<lambda>x. f x * g x) at_infinity F" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1596 | proof - | 
| 61973 | 1597 | have "((\<lambda>x. inverse (f x) * inverse (g x)) \<longlongrightarrow> inverse c * 0) F" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1598 | by (intro tendsto_mult tendsto_inverse assms filterlim_compose[OF tendsto_inverse_0]) | 
| 63546 | 1599 | then have "filterlim (\<lambda>x. inverse (f x) * inverse (g x)) (at (inverse c * 0)) F" | 
| 1600 | unfolding filterlim_at | |
| 1601 | using assms | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1602 | by (auto intro: filterlim_at_infinity_imp_eventually_ne tendsto_imp_eventually_ne eventually_conj) | 
| 63546 | 1603 | then show ?thesis | 
| 1604 | by (subst filterlim_inverse_at_iff[symmetric]) simp_all | |
| 68611 | 1605 | qed | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1606 | |
| 61973 | 1607 | lemma tendsto_inverse_0_at_top: "LIM x F. f x :> at_top \<Longrightarrow> ((\<lambda>x. inverse (f x) :: real) \<longlongrightarrow> 0) F" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51531diff
changeset | 1608 | by (metis filterlim_at filterlim_mono[OF _ at_top_le_at_infinity order_refl] filterlim_inverse_at_iff) | 
| 50419 | 1609 | |
| 63556 | 1610 | lemma real_tendsto_divide_at_top: | 
| 1611 | fixes c::"real" | |
| 1612 | assumes "(f \<longlongrightarrow> c) F" | |
| 1613 | assumes "filterlim g at_top F" | |
| 1614 | shows "((\<lambda>x. f x / g x) \<longlongrightarrow> 0) F" | |
| 1615 | by (auto simp: divide_inverse_commute | |
| 1616 | intro!: tendsto_mult[THEN tendsto_eq_rhs] tendsto_inverse_0_at_top assms) | |
| 1617 | ||
| 63546 | 1618 | lemma mult_nat_left_at_top: "c > 0 \<Longrightarrow> filterlim (\<lambda>x. c * x) at_top sequentially" | 
| 1619 | for c :: nat | |
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
65680diff
changeset | 1620 | by (rule filterlim_subseq) (auto simp: strict_mono_def) | 
| 59613 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1621 | |
| 63546 | 1622 | lemma mult_nat_right_at_top: "c > 0 \<Longrightarrow> filterlim (\<lambda>x. x * c) at_top sequentially" | 
| 1623 | for c :: nat | |
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
65680diff
changeset | 1624 | by (rule filterlim_subseq) (auto simp: strict_mono_def) | 
| 63546 | 1625 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1626 | lemma filterlim_times_pos: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1627 | "LIM x F1. c * f x :> at_right l" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1628 | if "filterlim f (at_right p) F1" "0 < c" "l = c * p" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1629 |   for c::"'a::{linordered_field, linorder_topology}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1630 | unfolding filterlim_iff | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1631 | proof safe | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1632 | fix P | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1633 | assume "\<forall>\<^sub>F x in at_right l. P x" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1634 | then obtain d where "c * p < d" "\<And>y. y > c * p \<Longrightarrow> y < d \<Longrightarrow> P y" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1635 | unfolding \<open>l = _ \<close> eventually_at_right_field | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1636 | by auto | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1637 | then have "\<forall>\<^sub>F a in at_right p. P (c * a)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1638 | by (auto simp: eventually_at_right_field \<open>0 < c\<close> field_simps intro!: exI[where x="d/c"]) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1639 | from that(1)[unfolded filterlim_iff, rule_format, OF this] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1640 | show "\<forall>\<^sub>F x in F1. P (c * f x)" . | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1641 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1642 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1643 | lemma filtermap_nhds_times: "c \<noteq> 0 \<Longrightarrow> filtermap (times c) (nhds a) = nhds (c * a)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1644 | for a c :: "'a::real_normed_field" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1645 | by (rule filtermap_fun_inverse[where g="\<lambda>x. inverse c * x"]) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1646 | (auto intro!: tendsto_eq_intros filterlim_ident) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1647 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1648 | lemma filtermap_times_pos_at_right: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1649 |   fixes c::"'a::{linordered_field, linorder_topology}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1650 | assumes "c > 0" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1651 | shows "filtermap (times c) (at_right p) = at_right (c * p)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1652 | using assms | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1653 | by (intro filtermap_fun_inverse[where g="\<lambda>x. inverse c * x"]) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1654 | (auto intro!: filterlim_ident filterlim_times_pos) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1655 | |
| 63546 | 1656 | lemma at_to_infinity: "(at (0::'a::{real_normed_field,field})) = filtermap inverse at_infinity"
 | 
| 59613 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1657 | proof (rule antisym) | 
| 61973 | 1658 | have "(inverse \<longlongrightarrow> (0::'a)) at_infinity" | 
| 59613 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1659 | by (fact tendsto_inverse_0) | 
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1660 | then show "filtermap inverse at_infinity \<le> at (0::'a)" | 
| 68615 | 1661 | using filterlim_def filterlim_ident filterlim_inverse_at_iff by fastforce | 
| 59613 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1662 | next | 
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1663 | have "filtermap inverse (filtermap inverse (at (0::'a))) \<le> filtermap inverse at_infinity" | 
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1664 | using filterlim_inverse_at_infinity unfolding filterlim_def | 
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1665 | by (rule filtermap_mono) | 
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1666 | then show "at (0::'a) \<le> filtermap inverse at_infinity" | 
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1667 | by (simp add: filtermap_ident filtermap_filtermap) | 
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1668 | qed | 
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1669 | |
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1670 | lemma lim_at_infinity_0: | 
| 63546 | 1671 |   fixes l :: "'a::{real_normed_field,field}"
 | 
| 1672 | shows "(f \<longlongrightarrow> l) at_infinity \<longleftrightarrow> ((f \<circ> inverse) \<longlongrightarrow> l) (at (0::'a))" | |
| 1673 | by (simp add: tendsto_compose_filtermap at_to_infinity filtermap_filtermap) | |
| 59613 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1674 | |
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1675 | lemma lim_zero_infinity: | 
| 63546 | 1676 |   fixes l :: "'a::{real_normed_field,field}"
 | 
| 61973 | 1677 | shows "((\<lambda>x. f(1 / x)) \<longlongrightarrow> l) (at (0::'a)) \<Longrightarrow> (f \<longlongrightarrow> l) at_infinity" | 
| 63546 | 1678 | by (simp add: inverse_eq_divide lim_at_infinity_0 comp_def) | 
| 59613 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1679 | |
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 1680 | |
| 60758 | 1681 | text \<open> | 
| 63546 | 1682 | We only show rules for multiplication and addition when the functions are either against a real | 
| 1683 |   value or against infinity. Further rules are easy to derive by using @{thm
 | |
| 1684 | filterlim_uminus_at_top}. | |
| 60758 | 1685 | \<close> | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1686 | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1687 | lemma filterlim_tendsto_pos_mult_at_top: | 
| 63546 | 1688 | assumes f: "(f \<longlongrightarrow> c) F" | 
| 1689 | and c: "0 < c" | |
| 1690 | and g: "LIM x F. g x :> at_top" | |
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1691 | shows "LIM x F. (f x * g x :: real) :> at_top" | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1692 | unfolding filterlim_at_top_gt[where c=0] | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1693 | proof safe | 
| 63546 | 1694 | fix Z :: real | 
| 1695 | assume "0 < Z" | |
| 60758 | 1696 | from f \<open>0 < c\<close> have "eventually (\<lambda>x. c / 2 < f x) F" | 
| 61810 | 1697 | by (auto dest!: tendstoD[where e="c / 2"] elim!: eventually_mono | 
| 63546 | 1698 | simp: dist_real_def abs_real_def split: if_split_asm) | 
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1699 | moreover from g have "eventually (\<lambda>x. (Z / c * 2) \<le> g x) F" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1700 | unfolding filterlim_at_top by auto | 
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1701 | ultimately show "eventually (\<lambda>x. Z \<le> f x * g x) F" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1702 | proof eventually_elim | 
| 63546 | 1703 | case (elim x) | 
| 60758 | 1704 | with \<open>0 < Z\<close> \<open>0 < c\<close> have "c / 2 * (Z / c * 2) \<le> f x * g x" | 
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1705 | by (intro mult_mono) (auto simp: zero_le_divide_iff) | 
| 60758 | 1706 | with \<open>0 < c\<close> show "Z \<le> f x * g x" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1707 | by simp | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1708 | qed | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1709 | qed | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1710 | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1711 | lemma filterlim_at_top_mult_at_top: | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1712 | assumes f: "LIM x F. f x :> at_top" | 
| 63546 | 1713 | and g: "LIM x F. g x :> at_top" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1714 | shows "LIM x F. (f x * g x :: real) :> at_top" | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1715 | unfolding filterlim_at_top_gt[where c=0] | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1716 | proof safe | 
| 63546 | 1717 | fix Z :: real | 
| 1718 | assume "0 < Z" | |
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1719 | from f have "eventually (\<lambda>x. 1 \<le> f x) F" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1720 | unfolding filterlim_at_top by auto | 
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1721 | moreover from g have "eventually (\<lambda>x. Z \<le> g x) F" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1722 | unfolding filterlim_at_top by auto | 
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1723 | ultimately show "eventually (\<lambda>x. Z \<le> f x * g x) F" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1724 | proof eventually_elim | 
| 63546 | 1725 | case (elim x) | 
| 60758 | 1726 | with \<open>0 < Z\<close> have "1 * Z \<le> f x * g x" | 
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1727 | by (intro mult_mono) (auto simp: zero_le_divide_iff) | 
| 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1728 | then show "Z \<le> f x * g x" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1729 | by simp | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1730 | qed | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1731 | qed | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1732 | |
| 63556 | 1733 | lemma filterlim_at_top_mult_tendsto_pos: | 
| 1734 | assumes f: "(f \<longlongrightarrow> c) F" | |
| 1735 | and c: "0 < c" | |
| 1736 | and g: "LIM x F. g x :> at_top" | |
| 1737 | shows "LIM x F. (g x * f x:: real) :> at_top" | |
| 1738 | by (auto simp: mult.commute intro!: filterlim_tendsto_pos_mult_at_top f c g) | |
| 1739 | ||
| 50419 | 1740 | lemma filterlim_tendsto_pos_mult_at_bot: | 
| 63546 | 1741 | fixes c :: real | 
| 1742 | assumes "(f \<longlongrightarrow> c) F" "0 < c" "filterlim g at_bot F" | |
| 50419 | 1743 | shows "LIM x F. f x * g x :> at_bot" | 
| 1744 | using filterlim_tendsto_pos_mult_at_top[OF assms(1,2), of "\<lambda>x. - g x"] assms(3) | |
| 1745 | unfolding filterlim_uminus_at_bot by simp | |
| 1746 | ||
| 60182 
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
 hoelzl parents: 
60141diff
changeset | 1747 | lemma filterlim_tendsto_neg_mult_at_bot: | 
| 63546 | 1748 | fixes c :: real | 
| 1749 | assumes c: "(f \<longlongrightarrow> c) F" "c < 0" and g: "filterlim g at_top F" | |
| 60182 
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
 hoelzl parents: 
60141diff
changeset | 1750 | shows "LIM x F. f x * g x :> at_bot" | 
| 
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
 hoelzl parents: 
60141diff
changeset | 1751 | using c filterlim_tendsto_pos_mult_at_top[of "\<lambda>x. - f x" "- c" F, OF _ _ g] | 
| 
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
 hoelzl parents: 
60141diff
changeset | 1752 | unfolding filterlim_uminus_at_bot tendsto_minus_cancel_left by simp | 
| 
e1ea5a6379c9
generalized tends over powr; added DERIV rule for powr
 hoelzl parents: 
60141diff
changeset | 1753 | |
| 56330 | 1754 | lemma filterlim_pow_at_top: | 
| 63721 | 1755 | fixes f :: "'a \<Rightarrow> real" | 
| 63546 | 1756 | assumes "0 < n" | 
| 1757 | and f: "LIM x F. f x :> at_top" | |
| 56330 | 1758 | shows "LIM x F. (f x)^n :: real :> at_top" | 
| 63546 | 1759 | using \<open>0 < n\<close> | 
| 1760 | proof (induct n) | |
| 1761 | case 0 | |
| 1762 | then show ?case by simp | |
| 1763 | next | |
| 56330 | 1764 | case (Suc n) with f show ?case | 
| 1765 | by (cases "n = 0") (auto intro!: filterlim_at_top_mult_at_top) | |
| 63546 | 1766 | qed | 
| 56330 | 1767 | |
| 1768 | lemma filterlim_pow_at_bot_even: | |
| 1769 | fixes f :: "real \<Rightarrow> real" | |
| 1770 | shows "0 < n \<Longrightarrow> LIM x F. f x :> at_bot \<Longrightarrow> even n \<Longrightarrow> LIM x F. (f x)^n :> at_top" | |
| 1771 | using filterlim_pow_at_top[of n "\<lambda>x. - f x" F] by (simp add: filterlim_uminus_at_top) | |
| 1772 | ||
| 1773 | lemma filterlim_pow_at_bot_odd: | |
| 1774 | fixes f :: "real \<Rightarrow> real" | |
| 1775 | shows "0 < n \<Longrightarrow> LIM x F. f x :> at_bot \<Longrightarrow> odd n \<Longrightarrow> LIM x F. (f x)^n :> at_bot" | |
| 1776 | using filterlim_pow_at_top[of n "\<lambda>x. - f x" F] by (simp add: filterlim_uminus_at_bot) | |
| 1777 | ||
| 67371 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1778 | lemma filterlim_power_at_infinity [tendsto_intros]: | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1779 | fixes F and f :: "'a \<Rightarrow> 'b :: real_normed_div_algebra" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1780 | assumes "filterlim f at_infinity F" "n > 0" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1781 | shows "filterlim (\<lambda>x. f x ^ n) at_infinity F" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1782 | by (rule filterlim_norm_at_top_imp_at_infinity) | 
| 68611 | 1783 | (auto simp: norm_power intro!: filterlim_pow_at_top assms | 
| 67371 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1784 | intro: filterlim_at_infinity_imp_norm_at_top) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 1785 | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1786 | lemma filterlim_tendsto_add_at_top: | 
| 61973 | 1787 | assumes f: "(f \<longlongrightarrow> c) F" | 
| 63546 | 1788 | and g: "LIM x F. g x :> at_top" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1789 | shows "LIM x F. (f x + g x :: real) :> at_top" | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1790 | unfolding filterlim_at_top_gt[where c=0] | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1791 | proof safe | 
| 63546 | 1792 | fix Z :: real | 
| 1793 | assume "0 < Z" | |
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1794 | from f have "eventually (\<lambda>x. c - 1 < f x) F" | 
| 61810 | 1795 | by (auto dest!: tendstoD[where e=1] elim!: eventually_mono simp: dist_real_def) | 
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1796 | moreover from g have "eventually (\<lambda>x. Z - (c - 1) \<le> g x) F" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1797 | unfolding filterlim_at_top by auto | 
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1798 | ultimately show "eventually (\<lambda>x. Z \<le> f x + g x) F" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1799 | by eventually_elim simp | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1800 | qed | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1801 | |
| 50347 | 1802 | lemma LIM_at_top_divide: | 
| 1803 | fixes f g :: "'a \<Rightarrow> real" | |
| 61973 | 1804 | assumes f: "(f \<longlongrightarrow> a) F" "0 < a" | 
| 63546 | 1805 | and g: "(g \<longlongrightarrow> 0) F" "eventually (\<lambda>x. 0 < g x) F" | 
| 50347 | 1806 | shows "LIM x F. f x / g x :> at_top" | 
| 1807 | unfolding divide_inverse | |
| 1808 | by (rule filterlim_tendsto_pos_mult_at_top[OF f]) (rule filterlim_inverse_at_top[OF g]) | |
| 1809 | ||
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1810 | lemma filterlim_at_top_add_at_top: | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1811 | assumes f: "LIM x F. f x :> at_top" | 
| 63546 | 1812 | and g: "LIM x F. g x :> at_top" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1813 | shows "LIM x F. (f x + g x :: real) :> at_top" | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1814 | unfolding filterlim_at_top_gt[where c=0] | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1815 | proof safe | 
| 63546 | 1816 | fix Z :: real | 
| 1817 | assume "0 < Z" | |
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1818 | from f have "eventually (\<lambda>x. 0 \<le> f x) F" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1819 | unfolding filterlim_at_top by auto | 
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1820 | moreover from g have "eventually (\<lambda>x. Z \<le> g x) F" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1821 | unfolding filterlim_at_top by auto | 
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1822 | ultimately show "eventually (\<lambda>x. Z \<le> f x + g x) F" | 
| 50324 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1823 | by eventually_elim simp | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1824 | qed | 
| 
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
 hoelzl parents: 
50323diff
changeset | 1825 | |
| 50331 | 1826 | lemma tendsto_divide_0: | 
| 61076 | 1827 |   fixes f :: "_ \<Rightarrow> 'a::{real_normed_div_algebra, division_ring}"
 | 
| 61973 | 1828 | assumes f: "(f \<longlongrightarrow> c) F" | 
| 63546 | 1829 | and g: "LIM x F. g x :> at_infinity" | 
| 61973 | 1830 | shows "((\<lambda>x. f x / g x) \<longlongrightarrow> 0) F" | 
| 63546 | 1831 | using tendsto_mult[OF f filterlim_compose[OF tendsto_inverse_0 g]] | 
| 1832 | by (simp add: divide_inverse) | |
| 50331 | 1833 | |
| 1834 | lemma linear_plus_1_le_power: | |
| 1835 | fixes x :: real | |
| 1836 | assumes x: "0 \<le> x" | |
| 1837 | shows "real n * x + 1 \<le> (x + 1) ^ n" | |
| 1838 | proof (induct n) | |
| 63546 | 1839 | case 0 | 
| 1840 | then show ?case by simp | |
| 1841 | next | |
| 50331 | 1842 | case (Suc n) | 
| 63546 | 1843 | from x have "real (Suc n) * x + 1 \<le> (x + 1) * (real n * x + 1)" | 
| 1844 | by (simp add: field_simps) | |
| 50331 | 1845 | also have "\<dots> \<le> (x + 1)^Suc n" | 
| 1846 | using Suc x by (simp add: mult_left_mono) | |
| 1847 | finally show ?case . | |
| 63546 | 1848 | qed | 
| 50331 | 1849 | |
| 1850 | lemma filterlim_realpow_sequentially_gt1: | |
| 1851 | fixes x :: "'a :: real_normed_div_algebra" | |
| 1852 | assumes x[arith]: "1 < norm x" | |
| 1853 | shows "LIM n sequentially. x ^ n :> at_infinity" | |
| 1854 | proof (intro filterlim_at_infinity[THEN iffD2] allI impI) | |
| 63546 | 1855 | fix y :: real | 
| 1856 | assume "0 < y" | |
| 50331 | 1857 | have "0 < norm x - 1" by simp | 
| 63546 | 1858 | then obtain N :: nat where "y < real N * (norm x - 1)" | 
| 1859 | by (blast dest: reals_Archimedean3) | |
| 1860 | also have "\<dots> \<le> real N * (norm x - 1) + 1" | |
| 1861 | by simp | |
| 1862 | also have "\<dots> \<le> (norm x - 1 + 1) ^ N" | |
| 1863 | by (rule linear_plus_1_le_power) simp | |
| 1864 | also have "\<dots> = norm x ^ N" | |
| 1865 | by simp | |
| 50331 | 1866 | finally have "\<forall>n\<ge>N. y \<le> norm x ^ n" | 
| 1867 | by (metis order_less_le_trans power_increasing order_less_imp_le x) | |
| 1868 | then show "eventually (\<lambda>n. y \<le> norm (x ^ n)) sequentially" | |
| 1869 | unfolding eventually_sequentially | |
| 1870 | by (auto simp: norm_power) | |
| 1871 | qed simp | |
| 1872 | ||
| 51471 | 1873 | |
| 66456 
621897f47fab
Various lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66447diff
changeset | 1874 | lemma filterlim_divide_at_infinity: | 
| 
621897f47fab
Various lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66447diff
changeset | 1875 | fixes f g :: "'a \<Rightarrow> 'a :: real_normed_field" | 
| 
621897f47fab
Various lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66447diff
changeset | 1876 | assumes "filterlim f (nhds c) F" "filterlim g (at 0) F" "c \<noteq> 0" | 
| 
621897f47fab
Various lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66447diff
changeset | 1877 | shows "filterlim (\<lambda>x. f x / g x) at_infinity F" | 
| 
621897f47fab
Various lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66447diff
changeset | 1878 | proof - | 
| 
621897f47fab
Various lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66447diff
changeset | 1879 | have "filterlim (\<lambda>x. f x * inverse (g x)) at_infinity F" | 
| 
621897f47fab
Various lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66447diff
changeset | 1880 | by (intro tendsto_mult_filterlim_at_infinity[OF assms(1,3)] | 
| 
621897f47fab
Various lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66447diff
changeset | 1881 | filterlim_compose [OF filterlim_inverse_at_infinity assms(2)]) | 
| 
621897f47fab
Various lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66447diff
changeset | 1882 | thus ?thesis by (simp add: field_simps) | 
| 
621897f47fab
Various lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66447diff
changeset | 1883 | qed | 
| 
621897f47fab
Various lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66447diff
changeset | 1884 | |
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1885 | subsection \<open>Floor and Ceiling\<close> | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1886 | |
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1887 | lemma eventually_floor_less: | 
| 63546 | 1888 |   fixes f :: "'a \<Rightarrow> 'b::{order_topology,floor_ceiling}"
 | 
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1889 | assumes f: "(f \<longlongrightarrow> l) F" | 
| 63546 | 1890 | and l: "l \<notin> \<int>" | 
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1891 | shows "\<forall>\<^sub>F x in F. of_int (floor l) < f x" | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1892 | by (intro order_tendstoD[OF f]) (metis Ints_of_int antisym_conv2 floor_correct l) | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1893 | |
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1894 | lemma eventually_less_ceiling: | 
| 63546 | 1895 |   fixes f :: "'a \<Rightarrow> 'b::{order_topology,floor_ceiling}"
 | 
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1896 | assumes f: "(f \<longlongrightarrow> l) F" | 
| 63546 | 1897 | and l: "l \<notin> \<int>" | 
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1898 | shows "\<forall>\<^sub>F x in F. f x < of_int (ceiling l)" | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1899 | by (intro order_tendstoD[OF f]) (metis Ints_of_int l le_of_int_ceiling less_le) | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1900 | |
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1901 | lemma eventually_floor_eq: | 
| 63546 | 1902 |   fixes f::"'a \<Rightarrow> 'b::{order_topology,floor_ceiling}"
 | 
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1903 | assumes f: "(f \<longlongrightarrow> l) F" | 
| 63546 | 1904 | and l: "l \<notin> \<int>" | 
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1905 | shows "\<forall>\<^sub>F x in F. floor (f x) = floor l" | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1906 | using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms] | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1907 | by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq) | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1908 | |
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1909 | lemma eventually_ceiling_eq: | 
| 63546 | 1910 |   fixes f::"'a \<Rightarrow> 'b::{order_topology,floor_ceiling}"
 | 
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1911 | assumes f: "(f \<longlongrightarrow> l) F" | 
| 63546 | 1912 | and l: "l \<notin> \<int>" | 
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1913 | shows "\<forall>\<^sub>F x in F. ceiling (f x) = ceiling l" | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1914 | using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms] | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1915 | by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq) | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1916 | |
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1917 | lemma tendsto_of_int_floor: | 
| 63546 | 1918 |   fixes f::"'a \<Rightarrow> 'b::{order_topology,floor_ceiling}"
 | 
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1919 | assumes "(f \<longlongrightarrow> l) F" | 
| 63546 | 1920 | and "l \<notin> \<int>" | 
| 1921 |   shows "((\<lambda>x. of_int (floor (f x)) :: 'c::{ring_1,topological_space}) \<longlongrightarrow> of_int (floor l)) F"
 | |
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1922 | using eventually_floor_eq[OF assms] | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1923 | by (simp add: eventually_mono topological_tendstoI) | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1924 | |
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1925 | lemma tendsto_of_int_ceiling: | 
| 63546 | 1926 |   fixes f::"'a \<Rightarrow> 'b::{order_topology,floor_ceiling}"
 | 
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1927 | assumes "(f \<longlongrightarrow> l) F" | 
| 63546 | 1928 | and "l \<notin> \<int>" | 
| 1929 |   shows "((\<lambda>x. of_int (ceiling (f x)):: 'c::{ring_1,topological_space}) \<longlongrightarrow> of_int (ceiling l)) F"
 | |
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1930 | using eventually_ceiling_eq[OF assms] | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1931 | by (simp add: eventually_mono topological_tendstoI) | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1932 | |
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1933 | lemma continuous_on_of_int_floor: | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1934 |   "continuous_on (UNIV - \<int>::'a::{order_topology, floor_ceiling} set)
 | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1935 |     (\<lambda>x. of_int (floor x)::'b::{ring_1, topological_space})"
 | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1936 | unfolding continuous_on_def | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1937 | by (auto intro!: tendsto_of_int_floor) | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1938 | |
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1939 | lemma continuous_on_of_int_ceiling: | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1940 |   "continuous_on (UNIV - \<int>::'a::{order_topology, floor_ceiling} set)
 | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1941 |     (\<lambda>x. of_int (ceiling x)::'b::{ring_1, topological_space})"
 | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1942 | unfolding continuous_on_def | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1943 | by (auto intro!: tendsto_of_int_ceiling) | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1944 | |
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63104diff
changeset | 1945 | |
| 60758 | 1946 | subsection \<open>Limits of Sequences\<close> | 
| 51526 | 1947 | |
| 62368 | 1948 | lemma [trans]: "X = Y \<Longrightarrow> Y \<longlonglongrightarrow> z \<Longrightarrow> X \<longlonglongrightarrow> z" | 
| 51526 | 1949 | by simp | 
| 1950 | ||
| 1951 | lemma LIMSEQ_iff: | |
| 1952 | fixes L :: "'a::real_normed_vector" | |
| 61969 | 1953 | shows "(X \<longlonglongrightarrow> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)" | 
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 1954 | unfolding lim_sequentially dist_norm .. | 
| 51526 | 1955 | |
| 63546 | 1956 | lemma LIMSEQ_I: "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X \<longlonglongrightarrow> L" | 
| 1957 | for L :: "'a::real_normed_vector" | |
| 1958 | by (simp add: LIMSEQ_iff) | |
| 1959 | ||
| 1960 | lemma LIMSEQ_D: "X \<longlonglongrightarrow> L \<Longrightarrow> 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r" | |
| 1961 | for L :: "'a::real_normed_vector" | |
| 1962 | by (simp add: LIMSEQ_iff) | |
| 1963 | ||
| 1964 | lemma LIMSEQ_linear: "X \<longlonglongrightarrow> x \<Longrightarrow> l > 0 \<Longrightarrow> (\<lambda> n. X (n * l)) \<longlonglongrightarrow> x" | |
| 51526 | 1965 | unfolding tendsto_def eventually_sequentially | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57447diff
changeset | 1966 | by (metis div_le_dividend div_mult_self1_is_m le_trans mult.commute) | 
| 51526 | 1967 | |
| 63546 | 1968 | |
| 1969 | text \<open>Transformation of limit.\<close> | |
| 1970 | ||
| 1971 | lemma Lim_transform: "(g \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> a) F" | |
| 1972 | for a b :: "'a::real_normed_vector" | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1973 | using tendsto_add [of g a F "\<lambda>x. f x - g x" 0] by simp | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1974 | |
| 63546 | 1975 | lemma Lim_transform2: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F \<Longrightarrow> (g \<longlongrightarrow> a) F" | 
| 1976 | for a b :: "'a::real_normed_vector" | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1977 | by (erule Lim_transform) (simp add: tendsto_minus_cancel) | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1978 | |
| 63546 | 1979 | proposition Lim_transform_eq: "((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> a) F \<longleftrightarrow> (g \<longlongrightarrow> a) F" | 
| 1980 | for a :: "'a::real_normed_vector" | |
| 1981 | using Lim_transform Lim_transform2 by blast | |
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62369diff
changeset | 1982 | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1983 | lemma Lim_transform_eventually: | 
| 61973 | 1984 | "eventually (\<lambda>x. f x = g x) net \<Longrightarrow> (f \<longlongrightarrow> l) net \<Longrightarrow> (g \<longlongrightarrow> l) net" | 
| 68615 | 1985 | using eventually_elim2 by (fastforce simp add: tendsto_def) | 
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1986 | |
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1987 | lemma Lim_transform_within: | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 1988 | assumes "(f \<longlongrightarrow> l) (at x within S)" | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 1989 | and "0 < d" | 
| 63546 | 1990 | and "\<And>x'. x'\<in>S \<Longrightarrow> 0 < dist x' x \<Longrightarrow> dist x' x < d \<Longrightarrow> f x' = g x'" | 
| 61973 | 1991 | shows "(g \<longlongrightarrow> l) (at x within S)" | 
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1992 | proof (rule Lim_transform_eventually) | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1993 | show "eventually (\<lambda>x. f x = g x) (at x within S)" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 1994 | using assms by (auto simp: eventually_at) | 
| 63546 | 1995 | show "(f \<longlongrightarrow> l) (at x within S)" | 
| 1996 | by fact | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1997 | qed | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1998 | |
| 67706 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67673diff
changeset | 1999 | lemma filterlim_transform_within: | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67673diff
changeset | 2000 | assumes "filterlim g G (at x within S)" | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67673diff
changeset | 2001 | assumes "G \<le> F" "0<d" "(\<And>x'. x' \<in> S \<Longrightarrow> 0 < dist x' x \<Longrightarrow> dist x' x < d \<Longrightarrow> f x' = g x') " | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67673diff
changeset | 2002 | shows "filterlim f F (at x within S)" | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67673diff
changeset | 2003 | using assms | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67673diff
changeset | 2004 | apply (elim filterlim_mono_eventually) | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67673diff
changeset | 2005 | unfolding eventually_at by auto | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67673diff
changeset | 2006 | |
| 63546 | 2007 | text \<open>Common case assuming being away from some crucial point like 0.\<close> | 
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2008 | lemma Lim_transform_away_within: | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2009 | fixes a b :: "'a::t1_space" | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2010 | assumes "a \<noteq> b" | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2011 | and "\<forall>x\<in>S. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x" | 
| 61973 | 2012 | and "(f \<longlongrightarrow> l) (at a within S)" | 
| 2013 | shows "(g \<longlongrightarrow> l) (at a within S)" | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2014 | proof (rule Lim_transform_eventually) | 
| 63546 | 2015 | show "(f \<longlongrightarrow> l) (at a within S)" | 
| 2016 | by fact | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2017 | show "eventually (\<lambda>x. f x = g x) (at a within S)" | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2018 | unfolding eventually_at_topological | 
| 63546 | 2019 |     by (rule exI [where x="- {b}"]) (simp add: open_Compl assms)
 | 
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2020 | qed | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2021 | |
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2022 | lemma Lim_transform_away_at: | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2023 | fixes a b :: "'a::t1_space" | 
| 63546 | 2024 | assumes ab: "a \<noteq> b" | 
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2025 | and fg: "\<forall>x. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x" | 
| 61973 | 2026 | and fl: "(f \<longlongrightarrow> l) (at a)" | 
| 2027 | shows "(g \<longlongrightarrow> l) (at a)" | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2028 | using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl by simp | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2029 | |
| 63546 | 2030 | text \<open>Alternatively, within an open set.\<close> | 
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2031 | lemma Lim_transform_within_open: | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 2032 | assumes "(f \<longlongrightarrow> l) (at a within T)" | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 2033 | and "open s" and "a \<in> s" | 
| 63546 | 2034 | and "\<And>x. x\<in>s \<Longrightarrow> x \<noteq> a \<Longrightarrow> f x = g x" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 2035 | shows "(g \<longlongrightarrow> l) (at a within T)" | 
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2036 | proof (rule Lim_transform_eventually) | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 2037 | show "eventually (\<lambda>x. f x = g x) (at a within T)" | 
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2038 | unfolding eventually_at_topological | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 2039 | using assms by auto | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 2040 | show "(f \<longlongrightarrow> l) (at a within T)" by fact | 
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2041 | qed | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2042 | |
| 63546 | 2043 | |
| 2044 | text \<open>A congruence rule allowing us to transform limits assuming not at point.\<close> | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2045 | |
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2046 | (* FIXME: Only one congruence rule for tendsto can be used at a time! *) | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2047 | |
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2048 | lemma Lim_cong_within(*[cong add]*): | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2049 | assumes "a = b" | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2050 | and "x = y" | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2051 | and "S = T" | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2052 | and "\<And>x. x \<noteq> b \<Longrightarrow> x \<in> T \<Longrightarrow> f x = g x" | 
| 61973 | 2053 | shows "(f \<longlongrightarrow> x) (at a within S) \<longleftrightarrow> (g \<longlongrightarrow> y) (at b within T)" | 
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2054 | unfolding tendsto_def eventually_at_topological | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2055 | using assms by simp | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2056 | |
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2057 | lemma Lim_cong_at(*[cong add]*): | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2058 | assumes "a = b" "x = y" | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2059 | and "\<And>x. x \<noteq> a \<Longrightarrow> f x = g x" | 
| 61973 | 2060 | shows "((\<lambda>x. f x) \<longlongrightarrow> x) (at a) \<longleftrightarrow> ((g \<longlongrightarrow> y) (at a))" | 
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2061 | unfolding tendsto_def eventually_at_topological | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2062 | using assms by simp | 
| 63546 | 2063 | |
| 2064 | text \<open>An unbounded sequence's inverse tends to 0.\<close> | |
| 65578 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2065 | lemma LIMSEQ_inverse_zero: | 
| 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2066 | assumes "\<And>r::real. \<exists>N. \<forall>n\<ge>N. r < X n" | 
| 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2067 | shows "(\<lambda>n. inverse (X n)) \<longlonglongrightarrow> 0" | 
| 51526 | 2068 | apply (rule filterlim_compose[OF tendsto_inverse_0]) | 
| 68615 | 2069 | by (metis assms eventually_at_top_linorderI filterlim_at_top_dense filterlim_at_top_imp_at_infinity) | 
| 51526 | 2070 | |
| 69593 | 2071 | text \<open>The sequence \<^term>\<open>1/n\<close> tends to 0 as \<^term>\<open>n\<close> tends to infinity.\<close> | 
| 63546 | 2072 | lemma LIMSEQ_inverse_real_of_nat: "(\<lambda>n. inverse (real (Suc n))) \<longlonglongrightarrow> 0" | 
| 51526 | 2073 | by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc | 
| 63546 | 2074 | filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity) | 
| 2075 | ||
| 2076 | text \<open> | |
| 69593 | 2077 | The sequence \<^term>\<open>r + 1/n\<close> tends to \<^term>\<open>r\<close> as \<^term>\<open>n\<close> tends to | 
| 63546 | 2078 | infinity is now easily proved. | 
| 2079 | \<close> | |
| 2080 | ||
| 2081 | lemma LIMSEQ_inverse_real_of_nat_add: "(\<lambda>n. r + inverse (real (Suc n))) \<longlonglongrightarrow> r" | |
| 51526 | 2082 | using tendsto_add [OF tendsto_const LIMSEQ_inverse_real_of_nat] by auto | 
| 2083 | ||
| 63546 | 2084 | lemma LIMSEQ_inverse_real_of_nat_add_minus: "(\<lambda>n. r + -inverse (real (Suc n))) \<longlonglongrightarrow> r" | 
| 51526 | 2085 | using tendsto_add [OF tendsto_const tendsto_minus [OF LIMSEQ_inverse_real_of_nat]] | 
| 2086 | by auto | |
| 2087 | ||
| 63546 | 2088 | lemma LIMSEQ_inverse_real_of_nat_add_minus_mult: "(\<lambda>n. r * (1 + - inverse (real (Suc n)))) \<longlonglongrightarrow> r" | 
| 51526 | 2089 | using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]] | 
| 2090 | by auto | |
| 2091 | ||
| 61973 | 2092 | lemma lim_inverse_n: "((\<lambda>n. inverse(of_nat n)) \<longlongrightarrow> (0::'a::real_normed_field)) sequentially" | 
| 61524 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 2093 | using lim_1_over_n by (simp add: inverse_eq_divide) | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 2094 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2095 | lemma lim_inverse_n': "((\<lambda>n. 1 / n) \<longlongrightarrow> 0) sequentially" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2096 | using lim_inverse_n | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2097 | by (simp add: inverse_eq_divide) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2098 | |
| 61969 | 2099 | lemma LIMSEQ_Suc_n_over_n: "(\<lambda>n. of_nat (Suc n) / of_nat n :: 'a :: real_normed_field) \<longlonglongrightarrow> 1" | 
| 61524 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 2100 | proof (rule Lim_transform_eventually) | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 2101 | show "eventually (\<lambda>n. 1 + inverse (of_nat n :: 'a) = of_nat (Suc n) / of_nat n) sequentially" | 
| 63546 | 2102 | using eventually_gt_at_top[of "0::nat"] | 
| 2103 | by eventually_elim (simp add: field_simps) | |
| 61969 | 2104 | have "(\<lambda>n. 1 + inverse (of_nat n) :: 'a) \<longlonglongrightarrow> 1 + 0" | 
| 61524 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 2105 | by (intro tendsto_add tendsto_const lim_inverse_n) | 
| 63546 | 2106 | then show "(\<lambda>n. 1 + inverse (of_nat n) :: 'a) \<longlonglongrightarrow> 1" | 
| 2107 | by simp | |
| 61524 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 2108 | qed | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 2109 | |
| 61969 | 2110 | lemma LIMSEQ_n_over_Suc_n: "(\<lambda>n. of_nat n / of_nat (Suc n) :: 'a :: real_normed_field) \<longlonglongrightarrow> 1" | 
| 61524 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 2111 | proof (rule Lim_transform_eventually) | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 2112 | show "eventually (\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a) = | 
| 63546 | 2113 | of_nat n / of_nat (Suc n)) sequentially" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 2114 | using eventually_gt_at_top[of "0::nat"] | 
| 61524 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 2115 | by eventually_elim (simp add: field_simps del: of_nat_Suc) | 
| 61969 | 2116 | have "(\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \<longlonglongrightarrow> inverse 1" | 
| 61524 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 2117 | by (intro tendsto_inverse LIMSEQ_Suc_n_over_n) simp_all | 
| 63546 | 2118 | then show "(\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \<longlonglongrightarrow> 1" | 
| 2119 | by simp | |
| 61524 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 2120 | qed | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 2121 | |
| 63546 | 2122 | |
| 60758 | 2123 | subsection \<open>Convergence on sequences\<close> | 
| 51526 | 2124 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2125 | lemma convergent_cong: | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2126 | assumes "eventually (\<lambda>x. f x = g x) sequentially" | 
| 63546 | 2127 | shows "convergent f \<longleftrightarrow> convergent g" | 
| 2128 | unfolding convergent_def | |
| 2129 | by (subst filterlim_cong[OF refl refl assms]) (rule refl) | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2130 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2131 | lemma convergent_Suc_iff: "convergent (\<lambda>n. f (Suc n)) \<longleftrightarrow> convergent f" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2132 | by (auto simp: convergent_def LIMSEQ_Suc_iff) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2133 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2134 | lemma convergent_ignore_initial_segment: "convergent (\<lambda>n. f (n + m)) = convergent f" | 
| 63546 | 2135 | proof (induct m arbitrary: f) | 
| 2136 | case 0 | |
| 2137 | then show ?case by simp | |
| 2138 | next | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2139 | case (Suc m) | 
| 63546 | 2140 | have "convergent (\<lambda>n. f (n + Suc m)) \<longleftrightarrow> convergent (\<lambda>n. f (Suc n + m))" | 
| 2141 | by simp | |
| 2142 | also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. f (n + m))" | |
| 2143 | by (rule convergent_Suc_iff) | |
| 2144 | also have "\<dots> \<longleftrightarrow> convergent f" | |
| 2145 | by (rule Suc) | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2146 | finally show ?case . | 
| 63546 | 2147 | qed | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2148 | |
| 51526 | 2149 | lemma convergent_add: | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 2150 | fixes X Y :: "nat \<Rightarrow> 'a::topological_monoid_add" | 
| 51526 | 2151 | assumes "convergent (\<lambda>n. X n)" | 
| 63546 | 2152 | and "convergent (\<lambda>n. Y n)" | 
| 51526 | 2153 | shows "convergent (\<lambda>n. X n + Y n)" | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 2154 | using assms unfolding convergent_def by (blast intro: tendsto_add) | 
| 51526 | 2155 | |
| 64267 | 2156 | lemma convergent_sum: | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 2157 | fixes X :: "'a \<Rightarrow> nat \<Rightarrow> 'b::topological_comm_monoid_add" | 
| 63915 | 2158 | shows "(\<And>i. i \<in> A \<Longrightarrow> convergent (\<lambda>n. X i n)) \<Longrightarrow> convergent (\<lambda>n. \<Sum>i\<in>A. X i n)" | 
| 2159 | by (induct A rule: infinite_finite_induct) (simp_all add: convergent_const convergent_add) | |
| 51526 | 2160 | |
| 2161 | lemma (in bounded_linear) convergent: | |
| 2162 | assumes "convergent (\<lambda>n. X n)" | |
| 2163 | shows "convergent (\<lambda>n. f (X n))" | |
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 2164 | using assms unfolding convergent_def by (blast intro: tendsto) | 
| 51526 | 2165 | |
| 2166 | lemma (in bounded_bilinear) convergent: | |
| 63546 | 2167 | assumes "convergent (\<lambda>n. X n)" | 
| 2168 | and "convergent (\<lambda>n. Y n)" | |
| 51526 | 2169 | shows "convergent (\<lambda>n. X n ** Y n)" | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 2170 | using assms unfolding convergent_def by (blast intro: tendsto) | 
| 51526 | 2171 | |
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 2172 | lemma convergent_minus_iff: | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 2173 | fixes X :: "nat \<Rightarrow> 'a::topological_group_add" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 2174 | shows "convergent X \<longleftrightarrow> convergent (\<lambda>n. - X n)" | 
| 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 2175 | unfolding convergent_def by (force dest: tendsto_minus) | 
| 51526 | 2176 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2177 | lemma convergent_diff: | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 2178 | fixes X Y :: "nat \<Rightarrow> 'a::topological_group_add" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2179 | assumes "convergent (\<lambda>n. X n)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2180 | assumes "convergent (\<lambda>n. Y n)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2181 | shows "convergent (\<lambda>n. X n - Y n)" | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 2182 | using assms unfolding convergent_def by (blast intro: tendsto_diff) | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2183 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2184 | lemma convergent_norm: | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2185 | assumes "convergent f" | 
| 63546 | 2186 | shows "convergent (\<lambda>n. norm (f n))" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2187 | proof - | 
| 63546 | 2188 | from assms have "f \<longlonglongrightarrow> lim f" | 
| 2189 | by (simp add: convergent_LIMSEQ_iff) | |
| 2190 | then have "(\<lambda>n. norm (f n)) \<longlonglongrightarrow> norm (lim f)" | |
| 2191 | by (rule tendsto_norm) | |
| 2192 | then show ?thesis | |
| 2193 | by (auto simp: convergent_def) | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2194 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2195 | |
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 2196 | lemma convergent_of_real: | 
| 63546 | 2197 | "convergent f \<Longrightarrow> convergent (\<lambda>n. of_real (f n) :: 'a::real_normed_algebra_1)" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2198 | unfolding convergent_def by (blast intro!: tendsto_of_real) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2199 | |
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 2200 | lemma convergent_add_const_iff: | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 2201 | "convergent (\<lambda>n. c + f n :: 'a::topological_ab_group_add) \<longleftrightarrow> convergent f" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2202 | proof | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2203 | assume "convergent (\<lambda>n. c + f n)" | 
| 63546 | 2204 | from convergent_diff[OF this convergent_const[of c]] show "convergent f" | 
| 2205 | by simp | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2206 | next | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2207 | assume "convergent f" | 
| 63546 | 2208 | from convergent_add[OF convergent_const[of c] this] show "convergent (\<lambda>n. c + f n)" | 
| 2209 | by simp | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2210 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2211 | |
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 2212 | lemma convergent_add_const_right_iff: | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 2213 | "convergent (\<lambda>n. f n + c :: 'a::topological_ab_group_add) \<longleftrightarrow> convergent f" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2214 | using convergent_add_const_iff[of c f] by (simp add: add_ac) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2215 | |
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 2216 | lemma convergent_diff_const_right_iff: | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 2217 | "convergent (\<lambda>n. f n - c :: 'a::topological_ab_group_add) \<longleftrightarrow> convergent f" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2218 | using convergent_add_const_right_iff[of f "-c"] by (simp add: add_ac) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2219 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2220 | lemma convergent_mult: | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 2221 | fixes X Y :: "nat \<Rightarrow> 'a::topological_semigroup_mult" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2222 | assumes "convergent (\<lambda>n. X n)" | 
| 63546 | 2223 | and "convergent (\<lambda>n. Y n)" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2224 | shows "convergent (\<lambda>n. X n * Y n)" | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 2225 | using assms unfolding convergent_def by (blast intro: tendsto_mult) | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2226 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2227 | lemma convergent_mult_const_iff: | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2228 | assumes "c \<noteq> 0" | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 2229 |   shows "convergent (\<lambda>n. c * f n :: 'a::{field,topological_semigroup_mult}) \<longleftrightarrow> convergent f"
 | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2230 | proof | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2231 | assume "convergent (\<lambda>n. c * f n)" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61976diff
changeset | 2232 | from assms convergent_mult[OF this convergent_const[of "inverse c"]] | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2233 | show "convergent f" by (simp add: field_simps) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2234 | next | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2235 | assume "convergent f" | 
| 63546 | 2236 | from convergent_mult[OF convergent_const[of c] this] show "convergent (\<lambda>n. c * f n)" | 
| 2237 | by simp | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2238 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2239 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2240 | lemma convergent_mult_const_right_iff: | 
| 68064 
b249fab48c76
type class generalisations; some work on infinite products
 paulson <lp15@cam.ac.uk> parents: 
67958diff
changeset | 2241 |   fixes c :: "'a::{field,topological_semigroup_mult}"
 | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2242 | assumes "c \<noteq> 0" | 
| 63546 | 2243 | shows "convergent (\<lambda>n. f n * c) \<longleftrightarrow> convergent f" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2244 | using convergent_mult_const_iff[OF assms, of f] by (simp add: mult_ac) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2245 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2246 | lemma convergent_imp_Bseq: "convergent f \<Longrightarrow> Bseq f" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2247 | by (simp add: Cauchy_Bseq convergent_Cauchy) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2248 | |
| 51526 | 2249 | |
| 60758 | 2250 | text \<open>A monotone sequence converges to its least upper bound.\<close> | 
| 51526 | 2251 | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 2252 | lemma LIMSEQ_incseq_SUP: | 
| 63546 | 2253 |   fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder,linorder_topology}"
 | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 2254 | assumes u: "bdd_above (range X)" | 
| 63546 | 2255 | and X: "incseq X" | 
| 61969 | 2256 | shows "X \<longlonglongrightarrow> (SUP i. X i)" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 2257 | by (rule order_tendstoI) | 
| 63546 | 2258 | (auto simp: eventually_sequentially u less_cSUP_iff | 
| 2259 | intro: X[THEN incseqD] less_le_trans cSUP_lessD[OF u]) | |
| 51526 | 2260 | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 2261 | lemma LIMSEQ_decseq_INF: | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 2262 |   fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology}"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 2263 | assumes u: "bdd_below (range X)" | 
| 63546 | 2264 | and X: "decseq X" | 
| 61969 | 2265 | shows "X \<longlonglongrightarrow> (INF i. X i)" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 2266 | by (rule order_tendstoI) | 
| 63546 | 2267 | (auto simp: eventually_sequentially u cINF_less_iff | 
| 2268 | intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u]) | |
| 2269 | ||
| 2270 | text \<open>Main monotonicity theorem.\<close> | |
| 2271 | ||
| 2272 | lemma Bseq_monoseq_convergent: "Bseq X \<Longrightarrow> monoseq X \<Longrightarrow> convergent X" | |
| 2273 | for X :: "nat \<Rightarrow> real" | |
| 2274 | by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP | |
| 2275 | dest: Bseq_bdd_above Bseq_bdd_below) | |
| 2276 | ||
| 2277 | lemma Bseq_mono_convergent: "Bseq X \<Longrightarrow> (\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n) \<Longrightarrow> convergent X" | |
| 2278 | for X :: "nat \<Rightarrow> real" | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 2279 | by (auto intro!: Bseq_monoseq_convergent incseq_imp_monoseq simp: incseq_def) | 
| 51526 | 2280 | |
| 63546 | 2281 | lemma monoseq_imp_convergent_iff_Bseq: "monoseq f \<Longrightarrow> convergent f \<longleftrightarrow> Bseq f" | 
| 2282 | for f :: "nat \<Rightarrow> real" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2283 | using Bseq_monoseq_convergent[of f] convergent_imp_Bseq[of f] by blast | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2284 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2285 | lemma Bseq_monoseq_convergent'_inc: | 
| 63546 | 2286 | fixes f :: "nat \<Rightarrow> real" | 
| 2287 | shows "Bseq (\<lambda>n. f (n + M)) \<Longrightarrow> (\<And>m n. M \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m \<le> f n) \<Longrightarrow> convergent f" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2288 | by (subst convergent_ignore_initial_segment [symmetric, of _ M]) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2289 | (auto intro!: Bseq_monoseq_convergent simp: monoseq_def) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2290 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2291 | lemma Bseq_monoseq_convergent'_dec: | 
| 63546 | 2292 | fixes f :: "nat \<Rightarrow> real" | 
| 2293 | shows "Bseq (\<lambda>n. f (n + M)) \<Longrightarrow> (\<And>m n. M \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m \<ge> f n) \<Longrightarrow> convergent f" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2294 | by (subst convergent_ignore_initial_segment [symmetric, of _ M]) | 
| 63546 | 2295 | (auto intro!: Bseq_monoseq_convergent simp: monoseq_def) | 
| 2296 | ||
| 2297 | lemma Cauchy_iff: "Cauchy X \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e)" | |
| 2298 | for X :: "nat \<Rightarrow> 'a::real_normed_vector" | |
| 51526 | 2299 | unfolding Cauchy_def dist_norm .. | 
| 2300 | ||
| 63546 | 2301 | lemma CauchyI: "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X" | 
| 2302 | for X :: "nat \<Rightarrow> 'a::real_normed_vector" | |
| 2303 | by (simp add: Cauchy_iff) | |
| 2304 | ||
| 2305 | lemma CauchyD: "Cauchy X \<Longrightarrow> 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e" | |
| 2306 | for X :: "nat \<Rightarrow> 'a::real_normed_vector" | |
| 2307 | by (simp add: Cauchy_iff) | |
| 51526 | 2308 | |
| 2309 | lemma incseq_convergent: | |
| 2310 | fixes X :: "nat \<Rightarrow> real" | |
| 63546 | 2311 | assumes "incseq X" | 
| 2312 | and "\<forall>i. X i \<le> B" | |
| 61969 | 2313 | obtains L where "X \<longlonglongrightarrow> L" "\<forall>i. X i \<le> L" | 
| 51526 | 2314 | proof atomize_elim | 
| 60758 | 2315 | from incseq_bounded[OF assms] \<open>incseq X\<close> Bseq_monoseq_convergent[of X] | 
| 61969 | 2316 | obtain L where "X \<longlonglongrightarrow> L" | 
| 51526 | 2317 | by (auto simp: convergent_def monoseq_def incseq_def) | 
| 61969 | 2318 | with \<open>incseq X\<close> show "\<exists>L. X \<longlonglongrightarrow> L \<and> (\<forall>i. X i \<le> L)" | 
| 51526 | 2319 | by (auto intro!: exI[of _ L] incseq_le) | 
| 2320 | qed | |
| 2321 | ||
| 2322 | lemma decseq_convergent: | |
| 2323 | fixes X :: "nat \<Rightarrow> real" | |
| 63546 | 2324 | assumes "decseq X" | 
| 2325 | and "\<forall>i. B \<le> X i" | |
| 61969 | 2326 | obtains L where "X \<longlonglongrightarrow> L" "\<forall>i. L \<le> X i" | 
| 51526 | 2327 | proof atomize_elim | 
| 60758 | 2328 | from decseq_bounded[OF assms] \<open>decseq X\<close> Bseq_monoseq_convergent[of X] | 
| 61969 | 2329 | obtain L where "X \<longlonglongrightarrow> L" | 
| 51526 | 2330 | by (auto simp: convergent_def monoseq_def decseq_def) | 
| 61969 | 2331 | with \<open>decseq X\<close> show "\<exists>L. X \<longlonglongrightarrow> L \<and> (\<forall>i. L \<le> X i)" | 
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68296diff
changeset | 2332 | by (auto intro!: exI[of _ L] decseq_ge) | 
| 51526 | 2333 | qed | 
| 2334 | ||
| 63546 | 2335 | |
| 60758 | 2336 | subsection \<open>Power Sequences\<close> | 
| 51526 | 2337 | |
| 63546 | 2338 | lemma Bseq_realpow: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> Bseq (\<lambda>n. x ^ n)" | 
| 2339 | for x :: real | |
| 68615 | 2340 | by (metis decseq_bounded decseq_def power_decreasing zero_le_power) | 
| 63546 | 2341 | |
| 2342 | lemma monoseq_realpow: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> monoseq (\<lambda>n. x ^ n)" | |
| 2343 | for x :: real | |
| 68615 | 2344 | using monoseq_def power_decreasing by blast | 
| 63546 | 2345 | |
| 2346 | lemma convergent_realpow: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> convergent (\<lambda>n. x ^ n)" | |
| 2347 | for x :: real | |
| 2348 | by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow) | |
| 2349 | ||
| 2350 | lemma LIMSEQ_inverse_realpow_zero: "1 < x \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) \<longlonglongrightarrow> 0" | |
| 2351 | for x :: real | |
| 51526 | 2352 | by (rule filterlim_compose[OF tendsto_inverse_0 filterlim_realpow_sequentially_gt1]) simp | 
| 2353 | ||
| 2354 | lemma LIMSEQ_realpow_zero: | |
| 63546 | 2355 | fixes x :: real | 
| 2356 | assumes "0 \<le> x" "x < 1" | |
| 2357 | shows "(\<lambda>n. x ^ n) \<longlonglongrightarrow> 0" | |
| 2358 | proof (cases "x = 0") | |
| 2359 | case False | |
| 2360 | with \<open>0 \<le> x\<close> have x0: "0 < x" by simp | |
| 2361 | then have "1 < inverse x" | |
| 2362 | using \<open>x < 1\<close> by (rule one_less_inverse) | |
| 2363 | then have "(\<lambda>n. inverse (inverse x ^ n)) \<longlonglongrightarrow> 0" | |
| 51526 | 2364 | by (rule LIMSEQ_inverse_realpow_zero) | 
| 63546 | 2365 | then show ?thesis by (simp add: power_inverse) | 
| 2366 | next | |
| 2367 | case True | |
| 2368 | show ?thesis | |
| 2369 | by (rule LIMSEQ_imp_Suc) (simp add: True) | |
| 2370 | qed | |
| 2371 | ||
| 2372 | lemma LIMSEQ_power_zero: "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) \<longlonglongrightarrow> 0" | |
| 2373 | for x :: "'a::real_normed_algebra_1" | |
| 2374 | apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero]) | |
| 68615 | 2375 | by (simp add: Zfun_le norm_power_ineq tendsto_Zfun_iff) | 
| 51526 | 2376 | |
| 61969 | 2377 | lemma LIMSEQ_divide_realpow_zero: "1 < x \<Longrightarrow> (\<lambda>n. a / (x ^ n) :: real) \<longlonglongrightarrow> 0" | 
| 51526 | 2378 | by (rule tendsto_divide_0 [OF tendsto_const filterlim_realpow_sequentially_gt1]) simp | 
| 2379 | ||
| 63556 | 2380 | lemma | 
| 2381 | tendsto_power_zero: | |
| 2382 | fixes x::"'a::real_normed_algebra_1" | |
| 2383 | assumes "filterlim f at_top F" | |
| 2384 | assumes "norm x < 1" | |
| 2385 | shows "((\<lambda>y. x ^ (f y)) \<longlongrightarrow> 0) F" | |
| 2386 | proof (rule tendstoI) | |
| 2387 | fix e::real assume "0 < e" | |
| 2388 | from tendstoD[OF LIMSEQ_power_zero[OF \<open>norm x < 1\<close>] \<open>0 < e\<close>] | |
| 2389 | have "\<forall>\<^sub>F xa in sequentially. norm (x ^ xa) < e" | |
| 2390 | by simp | |
| 2391 | then obtain N where N: "norm (x ^ n) < e" if "n \<ge> N" for n | |
| 2392 | by (auto simp: eventually_sequentially) | |
| 2393 | have "\<forall>\<^sub>F i in F. f i \<ge> N" | |
| 2394 | using \<open>filterlim f sequentially F\<close> | |
| 2395 | by (simp add: filterlim_at_top) | |
| 2396 | then show "\<forall>\<^sub>F i in F. dist (x ^ f i) 0 < e" | |
| 68615 | 2397 | by eventually_elim (auto simp: N) | 
| 63556 | 2398 | qed | 
| 2399 | ||
| 69593 | 2400 | text \<open>Limit of \<^term>\<open>c^n\<close> for \<^term>\<open>\<bar>c\<bar> < 1\<close>.\<close> | 
| 51526 | 2401 | |
| 68614 | 2402 | lemma LIMSEQ_abs_realpow_zero: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. \<bar>c\<bar> ^ n :: real) \<longlonglongrightarrow> 0" | 
| 51526 | 2403 | by (rule LIMSEQ_realpow_zero [OF abs_ge_zero]) | 
| 2404 | ||
| 68614 | 2405 | lemma LIMSEQ_abs_realpow_zero2: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. c ^ n :: real) \<longlonglongrightarrow> 0" | 
| 51526 | 2406 | by (rule LIMSEQ_power_zero) simp | 
| 2407 | ||
| 2408 | ||
| 60758 | 2409 | subsection \<open>Limits of Functions\<close> | 
| 51526 | 2410 | |
| 63546 | 2411 | lemma LIM_eq: "f \<midarrow>a\<rightarrow> L = (\<forall>r>0. \<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - L) < r)" | 
| 2412 | for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" | |
| 2413 | by (simp add: LIM_def dist_norm) | |
| 51526 | 2414 | |
| 2415 | lemma LIM_I: | |
| 63546 | 2416 | "(\<And>r. 0 < r \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - L) < r) \<Longrightarrow> f \<midarrow>a\<rightarrow> L" | 
| 2417 | for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" | |
| 2418 | by (simp add: LIM_eq) | |
| 2419 | ||
| 2420 | lemma LIM_D: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> 0 < r \<Longrightarrow> \<exists>s>0.\<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - L) < r" | |
| 2421 | for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" | |
| 2422 | by (simp add: LIM_eq) | |
| 2423 | ||
| 2424 | lemma LIM_offset: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. f (x + k)) \<midarrow>(a - k)\<rightarrow> L" | |
| 2425 | for a :: "'a::real_normed_vector" | |
| 2426 | by (simp add: filtermap_at_shift[symmetric, of a k] filterlim_def filtermap_filtermap) | |
| 2427 | ||
| 2428 | lemma LIM_offset_zero: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L" | |
| 2429 | for a :: "'a::real_normed_vector" | |
| 2430 | by (drule LIM_offset [where k = a]) (simp add: add.commute) | |
| 2431 | ||
| 2432 | lemma LIM_offset_zero_cancel: "(\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L \<Longrightarrow> f \<midarrow>a\<rightarrow> L" | |
| 2433 | for a :: "'a::real_normed_vector" | |
| 2434 | by (drule LIM_offset [where k = "- a"]) simp | |
| 2435 | ||
| 2436 | lemma LIM_offset_zero_iff: "f \<midarrow>a\<rightarrow> L \<longleftrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L" | |
| 2437 | for f :: "'a :: real_normed_vector \<Rightarrow> _" | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 2438 | using LIM_offset_zero_cancel[of f a L] LIM_offset_zero[of f L a] by auto | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 2439 | |
| 63546 | 2440 | lemma LIM_zero: "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. f x - l) \<longlongrightarrow> 0) F" | 
| 65578 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2441 | for f :: "'a \<Rightarrow> 'b::real_normed_vector" | 
| 63546 | 2442 | unfolding tendsto_iff dist_norm by simp | 
| 51526 | 2443 | |
| 2444 | lemma LIM_zero_cancel: | |
| 65578 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2445 | fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" | 
| 61973 | 2446 | shows "((\<lambda>x. f x - l) \<longlongrightarrow> 0) F \<Longrightarrow> (f \<longlongrightarrow> l) F" | 
| 51526 | 2447 | unfolding tendsto_iff dist_norm by simp | 
| 2448 | ||
| 63546 | 2449 | lemma LIM_zero_iff: "((\<lambda>x. f x - l) \<longlongrightarrow> 0) F = (f \<longlongrightarrow> l) F" | 
| 65578 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2450 | for f :: "'a \<Rightarrow> 'b::real_normed_vector" | 
| 63546 | 2451 | unfolding tendsto_iff dist_norm by simp | 
| 51526 | 2452 | |
| 2453 | lemma LIM_imp_LIM: | |
| 2454 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" | |
| 2455 | fixes g :: "'a::topological_space \<Rightarrow> 'c::real_normed_vector" | |
| 61976 | 2456 | assumes f: "f \<midarrow>a\<rightarrow> l" | 
| 63546 | 2457 | and le: "\<And>x. x \<noteq> a \<Longrightarrow> norm (g x - m) \<le> norm (f x - l)" | 
| 61976 | 2458 | shows "g \<midarrow>a\<rightarrow> m" | 
| 63546 | 2459 | by (rule metric_LIM_imp_LIM [OF f]) (simp add: dist_norm le) | 
| 51526 | 2460 | |
| 2461 | lemma LIM_equal2: | |
| 2462 | fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space" | |
| 63546 | 2463 | assumes "0 < R" | 
| 2464 | and "\<And>x. x \<noteq> a \<Longrightarrow> norm (x - a) < R \<Longrightarrow> f x = g x" | |
| 61976 | 2465 | shows "g \<midarrow>a\<rightarrow> l \<Longrightarrow> f \<midarrow>a\<rightarrow> l" | 
| 68594 | 2466 | by (rule metric_LIM_equal2 [OF _ assms]) (simp_all add: dist_norm) | 
| 51526 | 2467 | |
| 2468 | lemma LIM_compose2: | |
| 2469 | fixes a :: "'a::real_normed_vector" | |
| 61976 | 2470 | assumes f: "f \<midarrow>a\<rightarrow> b" | 
| 63546 | 2471 | and g: "g \<midarrow>b\<rightarrow> c" | 
| 2472 | and inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b" | |
| 61976 | 2473 | shows "(\<lambda>x. g (f x)) \<midarrow>a\<rightarrow> c" | 
| 63546 | 2474 | by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]]) | 
| 51526 | 2475 | |
| 2476 | lemma real_LIM_sandwich_zero: | |
| 2477 | fixes f g :: "'a::topological_space \<Rightarrow> real" | |
| 61976 | 2478 | assumes f: "f \<midarrow>a\<rightarrow> 0" | 
| 63546 | 2479 | and 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x" | 
| 2480 | and 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x" | |
| 61976 | 2481 | shows "g \<midarrow>a\<rightarrow> 0" | 
| 51526 | 2482 | proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *) | 
| 63546 | 2483 | fix x | 
| 2484 | assume x: "x \<noteq> a" | |
| 2485 | with 1 have "norm (g x - 0) = g x" by simp | |
| 51526 | 2486 | also have "g x \<le> f x" by (rule 2 [OF x]) | 
| 2487 | also have "f x \<le> \<bar>f x\<bar>" by (rule abs_ge_self) | |
| 2488 | also have "\<bar>f x\<bar> = norm (f x - 0)" by simp | |
| 2489 | finally show "norm (g x - 0) \<le> norm (f x - 0)" . | |
| 2490 | qed | |
| 2491 | ||
| 2492 | ||
| 60758 | 2493 | subsection \<open>Continuity\<close> | 
| 51526 | 2494 | |
| 63546 | 2495 | lemma LIM_isCont_iff: "(f \<midarrow>a\<rightarrow> f a) = ((\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> f a)" | 
| 2496 | for f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space" | |
| 2497 | by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel]) | |
| 2498 | ||
| 2499 | lemma isCont_iff: "isCont f x = (\<lambda>h. f (x + h)) \<midarrow>0\<rightarrow> f x" | |
| 2500 | for f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space" | |
| 2501 | by (simp add: isCont_def LIM_isCont_iff) | |
| 51526 | 2502 | |
| 2503 | lemma isCont_LIM_compose2: | |
| 2504 | fixes a :: "'a::real_normed_vector" | |
| 2505 | assumes f [unfolded isCont_def]: "isCont f a" | |
| 63546 | 2506 | and g: "g \<midarrow>f a\<rightarrow> l" | 
| 2507 | and inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a" | |
| 61976 | 2508 | shows "(\<lambda>x. g (f x)) \<midarrow>a\<rightarrow> l" | 
| 63546 | 2509 | by (rule LIM_compose2 [OF f g inj]) | 
| 2510 | ||
| 2511 | lemma isCont_norm [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a" | |
| 2512 | for f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" | |
| 51526 | 2513 | by (fact continuous_norm) | 
| 2514 | ||
| 63546 | 2515 | lemma isCont_rabs [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a" | 
| 2516 | for f :: "'a::t2_space \<Rightarrow> real" | |
| 51526 | 2517 | by (fact continuous_rabs) | 
| 2518 | ||
| 63546 | 2519 | lemma isCont_add [simp]: "isCont f a \<Longrightarrow> isCont g a \<Longrightarrow> isCont (\<lambda>x. f x + g x) a" | 
| 2520 | for f :: "'a::t2_space \<Rightarrow> 'b::topological_monoid_add" | |
| 51526 | 2521 | by (fact continuous_add) | 
| 2522 | ||
| 63546 | 2523 | lemma isCont_minus [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a" | 
| 2524 | for f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" | |
| 51526 | 2525 | by (fact continuous_minus) | 
| 2526 | ||
| 63546 | 2527 | lemma isCont_diff [simp]: "isCont f a \<Longrightarrow> isCont g a \<Longrightarrow> isCont (\<lambda>x. f x - g x) a" | 
| 2528 | for f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" | |
| 51526 | 2529 | by (fact continuous_diff) | 
| 2530 | ||
| 63546 | 2531 | lemma isCont_mult [simp]: "isCont f a \<Longrightarrow> isCont g a \<Longrightarrow> isCont (\<lambda>x. f x * g x) a" | 
| 2532 | for f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra" | |
| 51526 | 2533 | by (fact continuous_mult) | 
| 2534 | ||
| 63546 | 2535 | lemma (in bounded_linear) isCont: "isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a" | 
| 51526 | 2536 | by (fact continuous) | 
| 2537 | ||
| 63546 | 2538 | lemma (in bounded_bilinear) isCont: "isCont f a \<Longrightarrow> isCont g a \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a" | 
| 51526 | 2539 | by (fact continuous) | 
| 2540 | ||
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2541 | lemmas isCont_scaleR [simp] = | 
| 51526 | 2542 | bounded_bilinear.isCont [OF bounded_bilinear_scaleR] | 
| 2543 | ||
| 2544 | lemmas isCont_of_real [simp] = | |
| 2545 | bounded_linear.isCont [OF bounded_linear_of_real] | |
| 2546 | ||
| 63546 | 2547 | lemma isCont_power [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a" | 
| 2548 |   for f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
 | |
| 51526 | 2549 | by (fact continuous_power) | 
| 2550 | ||
| 64267 | 2551 | lemma isCont_sum [simp]: "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a" | 
| 63546 | 2552 | for f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::topological_comm_monoid_add" | 
| 64267 | 2553 | by (auto intro: continuous_sum) | 
| 51526 | 2554 | |
| 63546 | 2555 | |
| 60758 | 2556 | subsection \<open>Uniform Continuity\<close> | 
| 51526 | 2557 | |
| 63104 | 2558 | lemma uniformly_continuous_on_def: | 
| 2559 | fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" | |
| 2560 | shows "uniformly_continuous_on s f \<longleftrightarrow> | |
| 2561 | (\<forall>e>0. \<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e)" | |
| 2562 | unfolding uniformly_continuous_on_uniformity | |
| 2563 | uniformity_dist filterlim_INF filterlim_principal eventually_inf_principal | |
| 2564 | by (force simp: Ball_def uniformity_dist[symmetric] eventually_uniformity_metric) | |
| 2565 | ||
| 63546 | 2566 | abbreviation isUCont :: "['a::metric_space \<Rightarrow> 'b::metric_space] \<Rightarrow> bool" | 
| 2567 | where "isUCont f \<equiv> uniformly_continuous_on UNIV f" | |
| 2568 | ||
| 2569 | lemma isUCont_def: "isUCont f \<longleftrightarrow> (\<forall>r>0. \<exists>s>0. \<forall>x y. dist x y < s \<longrightarrow> dist (f x) (f y) < r)" | |
| 63104 | 2570 | by (auto simp: uniformly_continuous_on_def dist_commute) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 2571 | |
| 63546 | 2572 | lemma isUCont_isCont: "isUCont f \<Longrightarrow> isCont f x" | 
| 63104 | 2573 | by (drule uniformly_continuous_imp_continuous) (simp add: continuous_on_eq_continuous_at) | 
| 2574 | ||
| 2575 | lemma uniformly_continuous_on_Cauchy: | |
| 63546 | 2576 | fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" | 
| 63104 | 2577 | assumes "uniformly_continuous_on S f" "Cauchy X" "\<And>n. X n \<in> S" | 
| 2578 | shows "Cauchy (\<lambda>n. f (X n))" | |
| 2579 | using assms | |
| 68594 | 2580 | unfolding uniformly_continuous_on_def by (meson Cauchy_def) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 2581 | |
| 63546 | 2582 | lemma isUCont_Cauchy: "isUCont f \<Longrightarrow> Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))" | 
| 63104 | 2583 | by (rule uniformly_continuous_on_Cauchy[where S=UNIV and f=f]) simp_all | 
| 68611 | 2584 | |
| 64287 | 2585 | lemma uniformly_continuous_imp_Cauchy_continuous: | 
| 2586 | fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" | |
| 67091 | 2587 | shows "\<lbrakk>uniformly_continuous_on S f; Cauchy \<sigma>; \<And>n. (\<sigma> n) \<in> S\<rbrakk> \<Longrightarrow> Cauchy(f \<circ> \<sigma>)" | 
| 64287 | 2588 | by (simp add: uniformly_continuous_on_def Cauchy_def) meson | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 2589 | |
| 51526 | 2590 | lemma (in bounded_linear) isUCont: "isUCont f" | 
| 63546 | 2591 | unfolding isUCont_def dist_norm | 
| 51526 | 2592 | proof (intro allI impI) | 
| 63546 | 2593 | fix r :: real | 
| 2594 | assume r: "0 < r" | |
| 2595 | obtain K where K: "0 < K" and norm_le: "norm (f x) \<le> norm x * K" for x | |
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 2596 | using pos_bounded by blast | 
| 51526 | 2597 | show "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r" | 
| 2598 | proof (rule exI, safe) | |
| 56541 | 2599 | from r K show "0 < r / K" by simp | 
| 51526 | 2600 | next | 
| 2601 | fix x y :: 'a | |
| 2602 | assume xy: "norm (x - y) < r / K" | |
| 2603 | have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff) | |
| 2604 | also have "\<dots> \<le> norm (x - y) * K" by (rule norm_le) | |
| 2605 | also from K xy have "\<dots> < r" by (simp only: pos_less_divide_eq) | |
| 2606 | finally show "norm (f x - f y) < r" . | |
| 2607 | qed | |
| 2608 | qed | |
| 2609 | ||
| 2610 | lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))" | |
| 63546 | 2611 | by (rule isUCont [THEN isUCont_Cauchy]) | 
| 51526 | 2612 | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2613 | lemma LIM_less_bound: | 
| 51526 | 2614 | fixes f :: "real \<Rightarrow> real" | 
| 2615 |   assumes ev: "b < x" "\<forall> x' \<in> { b <..< x}. 0 \<le> f x'" and "isCont f x"
 | |
| 2616 | shows "0 \<le> f x" | |
| 63952 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 2617 | proof (rule tendsto_lowerbound) | 
| 61973 | 2618 | show "(f \<longlongrightarrow> f x) (at_left x)" | 
| 60758 | 2619 | using \<open>isCont f x\<close> by (simp add: filterlim_at_split isCont_def) | 
| 51526 | 2620 | show "eventually (\<lambda>x. 0 \<le> f x) (at_left x)" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51531diff
changeset | 2621 | using ev by (auto simp: eventually_at dist_real_def intro!: exI[of _ "x - b"]) | 
| 51526 | 2622 | qed simp | 
| 51471 | 2623 | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2624 | |
| 60758 | 2625 | subsection \<open>Nested Intervals and Bisection -- Needed for Compactness\<close> | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2626 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2627 | lemma nested_sequence_unique: | 
| 61969 | 2628 | assumes "\<forall>n. f n \<le> f (Suc n)" "\<forall>n. g (Suc n) \<le> g n" "\<forall>n. f n \<le> g n" "(\<lambda>n. f n - g n) \<longlonglongrightarrow> 0" | 
| 2629 | shows "\<exists>l::real. ((\<forall>n. f n \<le> l) \<and> f \<longlonglongrightarrow> l) \<and> ((\<forall>n. l \<le> g n) \<and> g \<longlonglongrightarrow> l)" | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2630 | proof - | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2631 | have "incseq f" unfolding incseq_Suc_iff by fact | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2632 | have "decseq g" unfolding decseq_Suc_iff by fact | 
| 63546 | 2633 | have "f n \<le> g 0" for n | 
| 2634 | proof - | |
| 2635 | from \<open>decseq g\<close> have "g n \<le> g 0" | |
| 2636 | by (rule decseqD) simp | |
| 2637 | with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] show ?thesis | |
| 2638 | by auto | |
| 2639 | qed | |
| 61969 | 2640 | then obtain u where "f \<longlonglongrightarrow> u" "\<forall>i. f i \<le> u" | 
| 60758 | 2641 | using incseq_convergent[OF \<open>incseq f\<close>] by auto | 
| 63546 | 2642 | moreover have "f 0 \<le> g n" for n | 
| 2643 | proof - | |
| 60758 | 2644 | from \<open>incseq f\<close> have "f 0 \<le> f n" by (rule incseqD) simp | 
| 63546 | 2645 | with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] show ?thesis | 
| 2646 | by simp | |
| 2647 | qed | |
| 61969 | 2648 | then obtain l where "g \<longlonglongrightarrow> l" "\<forall>i. l \<le> g i" | 
| 60758 | 2649 | using decseq_convergent[OF \<open>decseq g\<close>] by auto | 
| 61969 | 2650 | moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF \<open>f \<longlonglongrightarrow> u\<close> \<open>g \<longlonglongrightarrow> l\<close>]] | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2651 | ultimately show ?thesis by auto | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2652 | qed | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2653 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2654 | lemma Bolzano[consumes 1, case_names trans local]: | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2655 | fixes P :: "real \<Rightarrow> real \<Rightarrow> bool" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2656 | assumes [arith]: "a \<le> b" | 
| 63546 | 2657 | and trans: "\<And>a b c. P a b \<Longrightarrow> P b c \<Longrightarrow> a \<le> b \<Longrightarrow> b \<le> c \<Longrightarrow> P a c" | 
| 2658 | and local: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<exists>d>0. \<forall>a b. a \<le> x \<and> x \<le> b \<and> b - a < d \<longrightarrow> P a b" | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2659 | shows "P a b" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2660 | proof - | 
| 63040 | 2661 | define bisect where "bisect = | 
| 2662 | rec_nat (a, b) (\<lambda>n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))" | |
| 2663 | define l u where "l n = fst (bisect n)" and "u n = snd (bisect n)" for n | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2664 | have l[simp]: "l 0 = a" "\<And>n. l (Suc n) = (if P (l n) ((l n + u n) / 2) then (l n + u n) / 2 else l n)" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2665 | and u[simp]: "u 0 = b" "\<And>n. u (Suc n) = (if P (l n) ((l n + u n) / 2) then u n else (l n + u n) / 2)" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2666 | by (simp_all add: l_def u_def bisect_def split: prod.split) | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2667 | |
| 63546 | 2668 | have [simp]: "l n \<le> u n" for n by (induct n) auto | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2669 | |
| 61969 | 2670 | have "\<exists>x. ((\<forall>n. l n \<le> x) \<and> l \<longlonglongrightarrow> x) \<and> ((\<forall>n. x \<le> u n) \<and> u \<longlonglongrightarrow> x)" | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2671 | proof (safe intro!: nested_sequence_unique) | 
| 63546 | 2672 | show "l n \<le> l (Suc n)" "u (Suc n) \<le> u n" for n | 
| 2673 | by (induct n) auto | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2674 | next | 
| 63546 | 2675 | have "l n - u n = (a - b) / 2^n" for n | 
| 2676 | by (induct n) (auto simp: field_simps) | |
| 2677 | then show "(\<lambda>n. l n - u n) \<longlonglongrightarrow> 0" | |
| 2678 | by (simp add: LIMSEQ_divide_realpow_zero) | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2679 | qed fact | 
| 63546 | 2680 | then obtain x where x: "\<And>n. l n \<le> x" "\<And>n. x \<le> u n" and "l \<longlonglongrightarrow> x" "u \<longlonglongrightarrow> x" | 
| 2681 | by auto | |
| 2682 | obtain d where "0 < d" and d: "a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> b - a < d \<Longrightarrow> P a b" for a b | |
| 60758 | 2683 | using \<open>l 0 \<le> x\<close> \<open>x \<le> u 0\<close> local[of x] by auto | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2684 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2685 | show "P a b" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2686 | proof (rule ccontr) | 
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2687 | assume "\<not> P a b" | 
| 63546 | 2688 | have "\<not> P (l n) (u n)" for n | 
| 2689 | proof (induct n) | |
| 2690 | case 0 | |
| 2691 | then show ?case | |
| 2692 | by (simp add: \<open>\<not> P a b\<close>) | |
| 2693 | next | |
| 2694 | case (Suc n) | |
| 2695 | with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case | |
| 2696 | by auto | |
| 2697 | qed | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2698 | moreover | 
| 63546 | 2699 |     {
 | 
| 2700 | have "eventually (\<lambda>n. x - d / 2 < l n) sequentially" | |
| 61969 | 2701 | using \<open>0 < d\<close> \<open>l \<longlonglongrightarrow> x\<close> by (intro order_tendstoD[of _ x]) auto | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2702 | moreover have "eventually (\<lambda>n. u n < x + d / 2) sequentially" | 
| 61969 | 2703 | using \<open>0 < d\<close> \<open>u \<longlonglongrightarrow> x\<close> by (intro order_tendstoD[of _ x]) auto | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2704 | ultimately have "eventually (\<lambda>n. P (l n) (u n)) sequentially" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2705 | proof eventually_elim | 
| 63546 | 2706 | case (elim n) | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2707 | from add_strict_mono[OF this] have "u n - l n < d" by simp | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2708 | with x show "P (l n) (u n)" by (rule d) | 
| 63546 | 2709 | qed | 
| 2710 | } | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2711 | ultimately show False by simp | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2712 | qed | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2713 | qed | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2714 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2715 | lemma compact_Icc[simp, intro]: "compact {a .. b::real}"
 | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2716 | proof (cases "a \<le> b", rule compactI) | 
| 63546 | 2717 | fix C | 
| 2718 |   assume C: "a \<le> b" "\<forall>t\<in>C. open t" "{a..b} \<subseteq> \<Union>C"
 | |
| 63040 | 2719 |   define T where "T = {a .. b}"
 | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2720 |   from C(1,3) show "\<exists>C'\<subseteq>C. finite C' \<and> {a..b} \<subseteq> \<Union>C'"
 | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2721 | proof (induct rule: Bolzano) | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2722 | case (trans a b c) | 
| 63546 | 2723 |     then have *: "{a..c} = {a..b} \<union> {b..c}"
 | 
| 2724 | by auto | |
| 2725 | with trans obtain C1 C2 | |
| 2726 |       where "C1\<subseteq>C" "finite C1" "{a..b} \<subseteq> \<Union>C1" "C2\<subseteq>C" "finite C2" "{b..c} \<subseteq> \<Union>C2"
 | |
| 2727 | by auto | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2728 | with trans show ?case | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2729 | unfolding * by (intro exI[of _ "C1 \<union> C2"]) auto | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2730 | next | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2731 | case (local x) | 
| 63546 | 2732 | with C have "x \<in> \<Union>C" by auto | 
| 2733 | with C(2) obtain c where "x \<in> c" "open c" "c \<in> C" | |
| 2734 | by auto | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2735 |     then obtain e where "0 < e" "{x - e <..< x + e} \<subseteq> c"
 | 
| 62101 | 2736 | by (auto simp: open_dist dist_real_def subset_eq Ball_def abs_less_iff) | 
| 60758 | 2737 | with \<open>c \<in> C\<close> show ?case | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2738 |       by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto
 | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2739 | qed | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2740 | qed simp | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2741 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2742 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2743 | lemma continuous_image_closed_interval: | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2744 | fixes a b and f :: "real \<Rightarrow> real" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2745 |   defines "S \<equiv> {a..b}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2746 | assumes "a \<le> b" and f: "continuous_on S f" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2747 |   shows "\<exists>c d. f`S = {c..d} \<and> c \<le> d"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2748 | proof - | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2749 |   have S: "compact S" "S \<noteq> {}"
 | 
| 60758 | 2750 | using \<open>a \<le> b\<close> by (auto simp: S_def) | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2751 | obtain c where "c \<in> S" "\<forall>d\<in>S. f d \<le> f c" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2752 | using continuous_attains_sup[OF S f] by auto | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2753 | moreover obtain d where "d \<in> S" "\<forall>c\<in>S. f d \<le> f c" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2754 | using continuous_attains_inf[OF S f] by auto | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2755 | moreover have "connected (f`S)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2756 | using connected_continuous_image[OF f] connected_Icc by (auto simp: S_def) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2757 |   ultimately have "f ` S = {f d .. f c} \<and> f d \<le> f c"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2758 | by (auto simp: connected_iff_interval) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2759 | then show ?thesis | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2760 | by auto | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2761 | qed | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57276diff
changeset | 2762 | |
| 60974 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 paulson <lp15@cam.ac.uk> parents: 
60758diff
changeset | 2763 | lemma open_Collect_positive: | 
| 67958 
732c0b059463
tuned proofs and generalized some lemmas about limits
 huffman parents: 
67950diff
changeset | 2764 | fixes f :: "'a::topological_space \<Rightarrow> real" | 
| 63546 | 2765 | assumes f: "continuous_on s f" | 
| 2766 |   shows "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. 0 < f x}"
 | |
| 2767 |   using continuous_on_open_invariant[THEN iffD1, OF f, rule_format, of "{0 <..}"]
 | |
| 2768 | by (auto simp: Int_def field_simps) | |
| 60974 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 paulson <lp15@cam.ac.uk> parents: 
60758diff
changeset | 2769 | |
| 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 paulson <lp15@cam.ac.uk> parents: 
60758diff
changeset | 2770 | lemma open_Collect_less_Int: | 
| 67958 
732c0b059463
tuned proofs and generalized some lemmas about limits
 huffman parents: 
67950diff
changeset | 2771 | fixes f g :: "'a::topological_space \<Rightarrow> real" | 
| 63546 | 2772 | assumes f: "continuous_on s f" | 
| 2773 | and g: "continuous_on s g" | |
| 2774 |   shows "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. f x < g x}"
 | |
| 2775 | using open_Collect_positive[OF continuous_on_diff[OF g f]] by (simp add: field_simps) | |
| 60974 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 paulson <lp15@cam.ac.uk> parents: 
60758diff
changeset | 2776 | |
| 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 paulson <lp15@cam.ac.uk> parents: 
60758diff
changeset | 2777 | |
| 60758 | 2778 | subsection \<open>Boundedness of continuous functions\<close> | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2779 | |
| 60758 | 2780 | text\<open>By bisection, function continuous on closed interval is bounded above\<close> | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2781 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2782 | lemma isCont_eq_Ub: | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2783 | fixes f :: "real \<Rightarrow> 'a::linorder_topology" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2784 | shows "a \<le> b \<Longrightarrow> \<forall>x::real. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow> | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2785 | \<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)" | 
| 63546 | 2786 |   using continuous_attains_sup[of "{a..b}" f]
 | 
| 68615 | 2787 | by (auto simp: continuous_at_imp_continuous_on Ball_def Bex_def) | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2788 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2789 | lemma isCont_eq_Lb: | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2790 | fixes f :: "real \<Rightarrow> 'a::linorder_topology" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2791 | shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow> | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2792 | \<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> M \<le> f x) \<and> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)" | 
| 63546 | 2793 |   using continuous_attains_inf[of "{a..b}" f]
 | 
| 68615 | 2794 | by (auto simp: continuous_at_imp_continuous_on Ball_def Bex_def) | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2795 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2796 | lemma isCont_bounded: | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2797 | fixes f :: "real \<Rightarrow> 'a::linorder_topology" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2798 | shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow> \<exists>M. \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2799 | using isCont_eq_Ub[of a b f] by auto | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2800 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2801 | lemma isCont_has_Ub: | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2802 | fixes f :: "real \<Rightarrow> 'a::linorder_topology" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2803 | shows "a \<le> b \<Longrightarrow> \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x \<Longrightarrow> | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2804 | \<exists>M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M) \<and> (\<forall>N. N < M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> N < f x))" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2805 | using isCont_eq_Ub[of a b f] by auto | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2806 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2807 | (*HOL style here: object-level formulations*) | 
| 63546 | 2808 | lemma IVT_objl: | 
| 2809 | "(f a \<le> y \<and> y \<le> f b \<and> a \<le> b \<and> (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x)) \<longrightarrow> | |
| 2810 | (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = y)" | |
| 2811 | for a y :: real | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2812 | by (blast intro: IVT) | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2813 | |
| 63546 | 2814 | lemma IVT2_objl: | 
| 2815 | "(f b \<le> y \<and> y \<le> f a \<and> a \<le> b \<and> (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x)) \<longrightarrow> | |
| 2816 | (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = y)" | |
| 2817 | for b y :: real | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2818 | by (blast intro: IVT2) | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2819 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2820 | lemma isCont_Lb_Ub: | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2821 | fixes f :: "real \<Rightarrow> real" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2822 | assumes "a \<le> b" "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x" | 
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 2823 | shows "\<exists>L M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> L \<le> f x \<and> f x \<le> M) \<and> | 
| 63546 | 2824 | (\<forall>y. L \<le> y \<and> y \<le> M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> (f x = y)))" | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2825 | proof - | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2826 | obtain M where M: "a \<le> M" "M \<le> b" "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> f M" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2827 | using isCont_eq_Ub[OF assms] by auto | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2828 | obtain L where L: "a \<le> L" "L \<le> b" "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f L \<le> f x" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2829 | using isCont_eq_Lb[OF assms] by auto | 
| 68615 | 2830 | have "(\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f L \<le> f x \<and> f x \<le> f M)" | 
| 2831 | using M L by simp | |
| 2832 | moreover | |
| 2833 | have "(\<forall>y. f L \<le> y \<and> y \<le> f M \<longrightarrow> (\<exists>x\<ge>a. x \<le> b \<and> f x = y))" | |
| 2834 | proof (cases "L \<le> M") | |
| 2835 | case True then show ?thesis | |
| 2836 | using IVT[of f L _ M] M L assms by (metis order.trans) | |
| 2837 | next | |
| 2838 | case False then show ?thesis | |
| 2839 | using IVT2[of f L _ M] | |
| 2840 | by (metis L(2) M(1) assms(2) le_cases order.trans) | |
| 2841 | qed | |
| 2842 | ultimately show ?thesis | |
| 2843 | by blast | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2844 | qed | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2845 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2846 | |
| 63546 | 2847 | text \<open>Continuity of inverse function.\<close> | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2848 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2849 | lemma isCont_inverse_function: | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2850 | fixes f g :: "real \<Rightarrow> real" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2851 | assumes d: "0 < d" | 
| 68611 | 2852 | and inj: "\<And>z. \<bar>z-x\<bar> \<le> d \<Longrightarrow> g (f z) = z" | 
| 2853 | and cont: "\<And>z. \<bar>z-x\<bar> \<le> d \<Longrightarrow> isCont f z" | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2854 | shows "isCont g (f x)" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2855 | proof - | 
| 63546 | 2856 | let ?A = "f (x - d)" | 
| 2857 | let ?B = "f (x + d)" | |
| 2858 |   let ?D = "{x - d..x + d}"
 | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2859 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2860 | have f: "continuous_on ?D f" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2861 | using cont by (intro continuous_at_imp_continuous_on ballI) auto | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2862 | then have g: "continuous_on (f`?D) g" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2863 | using inj by (intro continuous_on_inv) auto | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2864 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2865 |   from d f have "{min ?A ?B <..< max ?A ?B} \<subseteq> f ` ?D"
 | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2866 | by (intro connected_contains_Ioo connected_continuous_image) (auto split: split_min split_max) | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2867 |   with g have "continuous_on {min ?A ?B <..< max ?A ?B} g"
 | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2868 | by (rule continuous_on_subset) | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2869 | moreover | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2870 | have "(?A < f x \<and> f x < ?B) \<or> (?B < f x \<and> f x < ?A)" | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2871 | using d inj by (intro continuous_inj_imp_mono[OF _ _ f] inj_on_imageI2[of g, OF inj_onI]) auto | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2872 |   then have "f x \<in> {min ?A ?B <..< max ?A ?B}"
 | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2873 | by auto | 
| 68615 | 2874 | ultimately show ?thesis | 
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2875 | by (simp add: continuous_on_eq_continuous_at) | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2876 | qed | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2877 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2878 | lemma isCont_inverse_function2: | 
| 63546 | 2879 | fixes f g :: "real \<Rightarrow> real" | 
| 2880 | shows | |
| 68611 | 2881 | "\<lbrakk>a < x; x < b; | 
| 2882 | \<And>z. \<lbrakk>a \<le> z; z \<le> b\<rbrakk> \<Longrightarrow> g (f z) = z; | |
| 2883 | \<And>z. \<lbrakk>a \<le> z; z \<le> b\<rbrakk> \<Longrightarrow> isCont f z\<rbrakk> \<Longrightarrow> isCont g (f x)" | |
| 63546 | 2884 | apply (rule isCont_inverse_function [where f=f and d="min (x - a) (b - x)"]) | 
| 2885 | apply (simp_all add: abs_le_iff) | |
| 2886 | done | |
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 2887 | |
| 63546 | 2888 | text \<open>Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110.\<close> | 
| 2889 | lemma LIM_fun_gt_zero: "f \<midarrow>c\<rightarrow> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> 0 < f x)" | |
| 2890 | for f :: "real \<Rightarrow> real" | |
| 68615 | 2891 | by (force simp: dest: LIM_D) | 
| 63546 | 2892 | |
| 2893 | lemma LIM_fun_less_zero: "f \<midarrow>c\<rightarrow> l \<Longrightarrow> l < 0 \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> f x < 0)" | |
| 2894 | for f :: "real \<Rightarrow> real" | |
| 68615 | 2895 | by (drule LIM_D [where r="-l"]) force+ | 
| 63546 | 2896 | |
| 2897 | lemma LIM_fun_not_zero: "f \<midarrow>c\<rightarrow> l \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> f x \<noteq> 0)" | |
| 2898 | for f :: "real \<Rightarrow> real" | |
| 68615 | 2899 | using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp: neq_iff) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51529diff
changeset | 2900 | |
| 31349 
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
 huffman parents: diff
changeset | 2901 | end |