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