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