src/HOL/Limits.thy
author Manuel Eberl <eberlm@in.tum.de>
Tue, 19 Jan 2016 11:19:25 +0100
changeset 62201 eca7b38c8ee5
parent 62101 26c0a70f78a3
child 62368 106569399cd6
permissions -rw-r--r--
Added approximation of powr to NEWS/CONTRIBUTORS
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
51524
7cb5ac44ca9e rename RealVector.thy to Real_Vector_Spaces.thy
hoelzl
parents: 51478
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
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
    14
subsection \<open>Filter going to infinity norm\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
    15
50324
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    16
definition at_infinity :: "'a::real_normed_vector filter" where
57276
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
    17
  "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
    18
57276
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
    19
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
    20
  unfolding at_infinity_def
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
    21
  by (subst eventually_INF_base)
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
    22
     (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
    23
50325
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    24
lemma at_infinity_eq_at_top_bot:
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60974
diff changeset
    25
  "(at_infinity :: real filter) = sup at_top at_bot"
57276
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
    26
  apply (simp add: filter_eq_iff eventually_sup eventually_at_infinity
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
    27
                   eventually_at_top_linorder eventually_at_bot_linorder)
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
    28
  apply safe
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
    29
  apply (rule_tac x="b" in exI, simp)
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
    30
  apply (rule_tac x="- b" in exI, simp)
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
    31
  apply (rule_tac x="max (- Na) N" in exI, auto simp: abs_real_def)
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
    32
  done
50325
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    33
57276
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
    34
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
    35
  unfolding at_infinity_eq_at_top_bot by simp
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    36
57276
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
    37
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
    38
  unfolding at_infinity_eq_at_top_bot by simp
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    39
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 56541
diff changeset
    40
lemma filterlim_at_top_imp_at_infinity:
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 56541
diff changeset
    41
  fixes f :: "_ \<Rightarrow> real"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 56541
diff changeset
    42
  shows "filterlim f at_top F \<Longrightarrow> filterlim f at_infinity F"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 56541
diff changeset
    43
  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
    44
59613
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    45
lemma lim_infinity_imp_sequentially:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
    46
  "(f \<longlongrightarrow> l) at_infinity \<Longrightarrow> ((\<lambda>n. f(n)) \<longlongrightarrow> l) sequentially"
59613
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    47
by (simp add: filterlim_at_top_imp_at_infinity filterlim_compose filterlim_real_sequentially)
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    48
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    49
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
    50
subsubsection \<open>Boundedness\<close>
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    51
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    52
definition Bfun :: "('a \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a filter \<Rightarrow> bool" where
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    53
  Bfun_metric_def: "Bfun f F = (\<exists>y. \<exists>K>0. eventually (\<lambda>x. dist (f x) y \<le> K) F)"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    54
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    55
abbreviation Bseq :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    56
  "Bseq X \<equiv> Bfun X sequentially"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    57
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    58
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
    59
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    60
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
    61
  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
    62
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    63
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
    64
  unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    65
51474
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    66
lemma Bfun_def:
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    67
  "Bfun f F \<longleftrightarrow> (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) F)"
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    68
  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
    69
proof safe
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    70
  fix y K assume "0 < K" and *: "eventually (\<lambda>x. dist (f x) y \<le> K) F"
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    71
  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
    72
    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
    73
  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
    74
    by eventually_elim auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
    75
  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
    76
    by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    77
qed auto
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    78
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
    79
lemma BfunI:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
    80
  assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) F" shows "Bfun f F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    81
unfolding Bfun_def
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    82
proof (intro exI conjI allI)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    83
  show "0 < max K 1" by simp
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    84
next
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
    85
  show "eventually (\<lambda>x. norm (f x) \<le> max K 1) F"
61810
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61806
diff changeset
    86
    using K by (rule eventually_mono, simp)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    87
qed
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    88
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    89
lemma BfunE:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
    90
  assumes "Bfun f F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
    91
  obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) 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
    92
using assms unfolding Bfun_def by blast
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    93
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    94
lemma Cauchy_Bseq: "Cauchy X \<Longrightarrow> Bseq X"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    95
  unfolding Cauchy_def Bfun_metric_def eventually_sequentially
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    96
  apply (erule_tac x=1 in allE)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    97
  apply simp
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    98
  apply safe
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
    99
  apply (rule_tac x="X M" in exI)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   100
  apply (rule_tac x=1 in exI)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   101
  apply (erule_tac x=M in allE)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   102
  apply simp
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   103
  apply (rule_tac x=M in exI)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   104
  apply (auto simp: dist_commute)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   105
  done
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   106
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   107
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   108
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
   109
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   110
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
   111
  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
   112
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   113
lemma BseqI2': "\<forall>n\<ge>N. norm (X n) \<le> K \<Longrightarrow> Bseq X"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   114
  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
   115
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   116
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
   117
  unfolding Bfun_def eventually_sequentially
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   118
proof safe
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   119
  fix N K assume "0 < K" "\<forall>n\<ge>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
   120
  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
   121
    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
   122
       (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
   123
qed auto
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   124
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   125
lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   126
unfolding Bseq_def by auto
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   127
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   128
lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<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
   129
by (simp add: Bseq_def)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   130
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   131
lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   132
by (auto simp add: Bseq_def)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   133
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
   134
lemma Bseq_bdd_above: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_above (range X)"
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
   135
proof (elim BseqE, intro bdd_aboveI2)
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
   136
  fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "X n \<le> K"
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
   137
    by (auto elim!: allE[of _ n])
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
   138
qed
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
   139
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
   140
lemma Bseq_bdd_above':
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   141
  "Bseq (X::nat \<Rightarrow> 'a :: real_normed_vector) \<Longrightarrow> bdd_above (range (\<lambda>n. norm (X n)))"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   142
proof (elim BseqE, intro bdd_aboveI2)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   143
  fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "norm (X n) \<le> K"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   144
    by (auto elim!: allE[of _ n])
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   145
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   146
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
   147
lemma Bseq_bdd_below: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_below (range X)"
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
   148
proof (elim BseqE, intro bdd_belowI2)
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
   149
  fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "- K \<le> X n"
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
   150
    by (auto elim!: allE[of _ n])
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
   151
qed
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
   152
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   153
lemma Bseq_eventually_mono:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   154
  assumes "eventually (\<lambda>n. norm (f n) \<le> norm (g n)) sequentially" "Bseq g"
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
   155
  shows   "Bseq f"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   156
proof -
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   157
  from assms(1) obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> norm (f n) \<le> norm (g n)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   158
    by (auto simp: eventually_at_top_linorder)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   159
  moreover from assms(2) obtain K where K: "\<And>n. norm (g n) \<le> K" by (blast elim!: BseqE)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   160
  ultimately have "norm (f n) \<le> max K (Max {norm (f n) |n. n < N})" for n
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   161
    apply (cases "n < N")
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   162
    apply (rule max.coboundedI2, rule Max.coboundedI, auto) []
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   163
    apply (rule max.coboundedI1, force intro: order.trans[OF N K])
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   164
    done
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
   165
  thus ?thesis by (blast intro: BseqI')
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   166
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   167
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   168
lemma lemma_NBseq_def:
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   169
  "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   170
proof safe
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   171
  fix K :: real
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   172
  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
   173
  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
   174
  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
   175
  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
   176
    by (blast intro: order_trans)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   177
  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
   178
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
   179
  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
   180
    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
   181
qed
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   182
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   183
text\<open>alternative definition for Bseq\<close>
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   184
lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   185
apply (simp add: Bseq_def)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   186
apply (simp (no_asm) add: lemma_NBseq_def)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   187
done
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   188
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   189
lemma lemma_NBseq_def2:
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   190
     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   191
apply (subst lemma_NBseq_def, auto)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   192
apply (rule_tac x = "Suc N" in exI)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   193
apply (rule_tac [2] x = N in exI)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61552
diff changeset
   194
apply (auto simp add: of_nat_Suc)
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   195
 prefer 2 apply (blast intro: order_less_imp_le)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   196
apply (drule_tac x = n in spec, simp)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   197
done
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   198
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   199
(* yet another definition for Bseq *)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   200
lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   201
by (simp add: Bseq_def lemma_NBseq_def2)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   202
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   203
subsubsection\<open>A Few More Equivalence Theorems for Boundedness\<close>
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   204
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   205
text\<open>alternative formulation for boundedness\<close>
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   206
lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   207
apply (unfold Bseq_def, safe)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   208
apply (rule_tac [2] x = "k + norm x" in exI)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   209
apply (rule_tac x = K in exI, simp)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   210
apply (rule exI [where x = 0], auto)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   211
apply (erule order_less_le_trans, simp)
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53602
diff changeset
   212
apply (drule_tac x=n in spec)
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   213
apply (drule order_trans [OF norm_triangle_ineq2])
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   214
apply simp
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   215
done
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   216
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   217
text\<open>alternative formulation for boundedness\<close>
53602
0ae3db699a3e tuned proofs
haftmann
parents: 53381
diff changeset
   218
lemma Bseq_iff3:
0ae3db699a3e tuned proofs
haftmann
parents: 53381
diff changeset
   219
  "Bseq X \<longleftrightarrow> (\<exists>k>0. \<exists>N. \<forall>n. norm (X n + - X N) \<le> k)" (is "?P \<longleftrightarrow> ?Q")
0ae3db699a3e tuned proofs
haftmann
parents: 53381
diff changeset
   220
proof
0ae3db699a3e tuned proofs
haftmann
parents: 53381
diff changeset
   221
  assume ?P
0ae3db699a3e tuned proofs
haftmann
parents: 53381
diff changeset
   222
  then obtain K
0ae3db699a3e tuned proofs
haftmann
parents: 53381
diff changeset
   223
    where *: "0 < K" and **: "\<And>n. norm (X n) \<le> K" by (auto simp add: Bseq_def)
0ae3db699a3e tuned proofs
haftmann
parents: 53381
diff changeset
   224
  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
   225
  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
   226
    by (auto intro: order_trans norm_triangle_ineq4)
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53602
diff changeset
   227
  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
   228
    by simp
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   229
  with \<open>0 < K + norm (X 0)\<close> show ?Q by blast
53602
0ae3db699a3e tuned proofs
haftmann
parents: 53381
diff changeset
   230
next
0ae3db699a3e tuned proofs
haftmann
parents: 53381
diff changeset
   231
  assume ?Q then show ?P by (auto simp add: Bseq_iff2)
0ae3db699a3e tuned proofs
haftmann
parents: 53381
diff changeset
   232
qed
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   233
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   234
lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   235
apply (simp add: Bseq_def)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   236
apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   237
apply (drule_tac x = n in spec, arith)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   238
done
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   239
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53602
diff changeset
   240
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   241
subsubsection\<open>Upper Bounds and Lubs of Bounded Sequences\<close>
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   242
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   243
lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   244
  by (simp add: Bseq_def)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   245
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
   246
lemma Bseq_add:
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   247
  assumes "Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   248
  shows   "Bseq (\<lambda>x. f x + c)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   249
proof -
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   250
  from assms obtain K where K: "\<And>x. norm (f x) \<le> K" unfolding Bseq_def by blast
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   251
  {
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   252
    fix x :: nat
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   253
    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
   254
    also have "norm (f x) \<le> K" by (rule K)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   255
    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
   256
  }
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   257
  thus ?thesis by (rule BseqI')
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   258
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   259
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   260
lemma Bseq_add_iff: "Bseq (\<lambda>x. f x + c) \<longleftrightarrow> Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   261
  using Bseq_add[of f c] Bseq_add[of "\<lambda>x. f x + c" "-c"] by auto
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   262
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
   263
lemma Bseq_mult:
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   264
  assumes "Bseq (f :: nat \<Rightarrow> 'a :: real_normed_field)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   265
  assumes "Bseq (g :: nat \<Rightarrow> 'a :: real_normed_field)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   266
  shows   "Bseq (\<lambda>x. f x * g x)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   267
proof -
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
   268
  from assms obtain K1 K2 where K: "\<And>x. norm (f x) \<le> K1" "K1 > 0" "\<And>x. norm (g x) \<le> K2" "K2 > 0"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   269
    unfolding Bseq_def by blast
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   270
  hence "\<And>x. norm (f x * g x) \<le> K1 * K2" by (auto simp: norm_mult intro!: mult_mono)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   271
  thus ?thesis by (rule BseqI')
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   272
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   273
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   274
lemma Bfun_const [simp]: "Bfun (\<lambda>_. c) F"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   275
  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
   276
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   277
lemma Bseq_cmult_iff: "(c :: 'a :: real_normed_field) \<noteq> 0 \<Longrightarrow> Bseq (\<lambda>x. c * f x) \<longleftrightarrow> Bseq f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   278
proof
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   279
  assume "c \<noteq> 0" "Bseq (\<lambda>x. c * f x)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   280
  find_theorems "Bfun (\<lambda>_. ?c) _"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   281
  from Bfun_const this(2) have "Bseq (\<lambda>x. inverse c * (c * f x))" by (rule Bseq_mult)
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61738
diff changeset
   282
  with \<open>c \<noteq> 0\<close> show "Bseq f" by (simp add: divide_simps)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   283
qed (intro Bseq_mult Bfun_const)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   284
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   285
lemma Bseq_subseq: "Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector) \<Longrightarrow> Bseq (\<lambda>x. f (g x))"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   286
  unfolding Bseq_def by auto
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   287
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   288
lemma Bseq_Suc_iff: "Bseq (\<lambda>n. f (Suc n)) \<longleftrightarrow> Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   289
  using Bseq_offset[of f 1] by (auto intro: Bseq_subseq)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   290
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   291
lemma increasing_Bseq_subseq_iff:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   292
  assumes "\<And>x y. x \<le> y \<Longrightarrow> norm (f x :: 'a :: real_normed_vector) \<le> norm (f y)" "subseq g"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   293
  shows   "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   294
proof
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   295
  assume "Bseq (\<lambda>x. f (g x))"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   296
  then obtain K where K: "\<And>x. norm (f (g x)) \<le> K" unfolding Bseq_def by auto
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   297
  {
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   298
    fix x :: nat
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   299
    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
   300
      by (auto simp: filterlim_at_top eventually_at_top_linorder)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   301
    hence "norm (f x) \<le> norm (f (g y))" using assms(1) by blast
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   302
    also have "norm (f (g y)) \<le> K" by (rule K)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   303
    finally have "norm (f x) \<le> K" .
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   304
  }
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   305
  thus "Bseq f" by (rule BseqI')
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   306
qed (insert Bseq_subseq[of f g], simp_all)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   307
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   308
lemma nonneg_incseq_Bseq_subseq_iff:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   309
  assumes "\<And>x. f x \<ge> 0" "incseq (f :: nat \<Rightarrow> real)" "subseq g"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   310
  shows   "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   311
  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
   312
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   313
lemma Bseq_eq_bounded: "range f \<subseteq> {a .. b::real} \<Longrightarrow> Bseq f"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   314
  apply (simp add: subset_eq)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   315
  apply (rule BseqI'[where K="max (norm a) (norm b)"])
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   316
  apply (erule_tac x=n in allE)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   317
  apply auto
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   318
  done
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   319
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   320
lemma incseq_bounded: "incseq X \<Longrightarrow> \<forall>i. X i \<le> (B::real) \<Longrightarrow> Bseq X"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   321
  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
   322
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   323
lemma decseq_bounded: "decseq X \<Longrightarrow> \<forall>i. (B::real) \<le> X i \<Longrightarrow> Bseq X"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   324
  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
   325
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   326
subsection \<open>Bounded Monotonic Sequences\<close>
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   327
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   328
subsubsection\<open>A Bounded and Monotonic Sequence Converges\<close>
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   329
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   330
(* TODO: delete *)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   331
(* FIXME: one use in NSA/HSEQ.thy *)
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
   332
lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X \<longlonglongrightarrow> L)"
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   333
  apply (rule_tac x="X m" in exI)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   334
  apply (rule filterlim_cong[THEN iffD2, OF refl refl _ tendsto_const])
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   335
  unfolding eventually_sequentially
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   336
  apply blast
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   337
  done
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   338
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   339
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
   340
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   341
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
   342
  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
   343
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   344
lemma ZfunI:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   345
  "(\<And>r. 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F) \<Longrightarrow> Zfun f F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   346
  unfolding Zfun_def by simp
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   347
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   348
lemma ZfunD:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   349
  "\<lbrakk>Zfun f F; 0 < r\<rbrakk> \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   350
  unfolding Zfun_def by simp
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   351
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   352
lemma Zfun_ssubst:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   353
  "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
   354
  unfolding Zfun_def by (auto elim!: eventually_rev_mp)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   355
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   356
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
   357
  unfolding Zfun_def by simp
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   358
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   359
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
   360
  unfolding Zfun_def by simp
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   361
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   362
lemma Zfun_imp_Zfun:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   363
  assumes f: "Zfun f F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   364
  assumes g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   365
  shows "Zfun (\<lambda>x. g x) F"
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   366
proof (cases)
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   367
  assume K: "0 < K"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   368
  show ?thesis
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   369
  proof (rule ZfunI)
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   370
    fix r::real assume "0 < r"
56541
0e3abadbef39 made divide_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
   371
    hence "0 < r / K" using K by simp
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   372
    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
   373
      using ZfunD [OF f] by blast
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   374
    with g show "eventually (\<lambda>x. norm (g x) < r) F"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   375
    proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   376
      case (elim x)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   377
      hence "norm (f x) * K < r"
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   378
        by (simp add: pos_less_divide_eq K)
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   379
      thus ?case
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   380
        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
   381
    qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   382
  qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   383
next
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   384
  assume "\<not> 0 < K"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   385
  hence K: "K \<le> 0" by (simp only: not_less)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   386
  show ?thesis
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   387
  proof (rule ZfunI)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   388
    fix r :: real
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   389
    assume "0 < r"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   390
    from g show "eventually (\<lambda>x. norm (g x) < r) F"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   391
    proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   392
      case (elim x)
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   393
      also have "norm (f x) * K \<le> norm (f x) * 0"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   394
        using K norm_ge_zero by (rule mult_left_mono)
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   395
      finally show ?case
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   396
        using \<open>0 < r\<close> by simp
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   397
    qed
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   398
  qed
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   399
qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   400
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   401
lemma Zfun_le: "\<lbrakk>Zfun g F; \<forall>x. norm (f x) \<le> norm (g x)\<rbrakk> \<Longrightarrow> Zfun f F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   402
  by (erule_tac K="1" in Zfun_imp_Zfun, simp)
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   403
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   404
lemma Zfun_add:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   405
  assumes f: "Zfun f F" and g: "Zfun g F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   406
  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
   407
proof (rule ZfunI)
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   408
  fix r::real assume "0 < r"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   409
  hence r: "0 < r / 2" by simp
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   410
  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
   411
    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
   412
  moreover
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   413
  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
   414
    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
   415
  ultimately
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   416
  show "eventually (\<lambda>x. norm (f x + g x) < r) F"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   417
  proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   418
    case (elim x)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   419
    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
   420
      by (rule norm_triangle_ineq)
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   421
    also have "\<dots> < r/2 + r/2"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   422
      using elim by (rule add_strict_mono)
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   423
    finally show ?case
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   424
      by simp
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   425
  qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   426
qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   427
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   428
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
   429
  unfolding Zfun_def by simp
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   430
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   431
lemma Zfun_diff: "\<lbrakk>Zfun f F; Zfun g F\<rbrakk> \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53602
diff changeset
   432
  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
   433
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   434
lemma (in bounded_linear) Zfun:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   435
  assumes g: "Zfun g F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   436
  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
   437
proof -
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   438
  obtain K where "\<And>x. norm (f x) \<le> norm x * 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
   439
    using bounded by blast
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   440
  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
   441
    by simp
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   442
  with g show ?thesis
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   443
    by (rule Zfun_imp_Zfun)
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   444
qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   445
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   446
lemma (in bounded_bilinear) Zfun:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   447
  assumes f: "Zfun f F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   448
  assumes g: "Zfun g F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   449
  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
   450
proof (rule ZfunI)
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   451
  fix r::real assume r: "0 < r"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   452
  obtain K where K: "0 < K"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   453
    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
   454
    using pos_bounded by blast
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   455
  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
   456
    by (rule positive_imp_inverse_positive)
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   457
  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
   458
    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
   459
  moreover
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   460
  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
   461
    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
   462
  ultimately
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   463
  show "eventually (\<lambda>x. norm (f x ** g x) < r) F"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   464
  proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   465
    case (elim x)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   466
    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
   467
      by (rule norm_le)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   468
    also have "norm (f x) * norm (g x) * K < r * inverse K * K"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   469
      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
   470
    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
   471
      by simp
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   472
    finally show ?case .
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   473
  qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   474
qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   475
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   476
lemma (in bounded_bilinear) Zfun_left:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   477
  "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
   478
  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
   479
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   480
lemma (in bounded_bilinear) Zfun_right:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   481
  "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
   482
  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
   483
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   484
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
   485
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
   486
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
   487
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   488
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
   489
  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
   490
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   491
lemma tendsto_0_le: "\<lbrakk>(f \<longlongrightarrow> 0) F; eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F\<rbrakk>
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   492
                     \<Longrightarrow> (g \<longlongrightarrow> 0) F"
56366
0362c3bb4d02 new theorem about zero limits
paulson <lp15@cam.ac.uk>
parents: 56330
diff changeset
   493
  by (simp add: Zfun_imp_Zfun tendsto_Zfun_iff)
0362c3bb4d02 new theorem about zero limits
paulson <lp15@cam.ac.uk>
parents: 56330
diff changeset
   494
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   495
subsubsection \<open>Distance and norms\<close>
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   496
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   497
lemma tendsto_dist [tendsto_intros]:
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   498
  fixes l m :: "'a :: metric_space"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   499
  assumes f: "(f \<longlongrightarrow> l) F" and g: "(g \<longlongrightarrow> m) F"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   500
  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
   501
proof (rule tendstoI)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   502
  fix e :: real assume "0 < e"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   503
  hence e2: "0 < e/2" by simp
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   504
  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
   505
  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
   506
  proof (eventually_elim)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   507
    case (elim x)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   508
    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
   509
      unfolding dist_real_def
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   510
      using dist_triangle2 [of "f x" "g x" "l"]
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   511
      using dist_triangle2 [of "g x" "l" "m"]
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   512
      using dist_triangle3 [of "l" "m" "f x"]
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   513
      using dist_triangle [of "f x" "m" "g x"]
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   514
      by arith
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   515
  qed
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   516
qed
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   517
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   518
lemma continuous_dist[continuous_intros]:
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
   519
  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
   520
  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
   521
  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
   522
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56366
diff changeset
   523
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
   524
  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
   525
  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
   526
  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
   527
31565
da5a5589418e theorem attribute [tendsto_intros]
huffman
parents: 31492
diff changeset
   528
lemma tendsto_norm [tendsto_intros]:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   529
  "(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
   530
  unfolding norm_conv_dist by (intro tendsto_intros)
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   531
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
   532
lemma continuous_norm [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
   533
  "continuous F f \<Longrightarrow> continuous F (\<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
   534
  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
   535
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56366
diff changeset
   536
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
   537
  "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
   538
  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
   539
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   540
lemma tendsto_norm_zero:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   541
  "(f \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   542
  by (drule tendsto_norm, simp)
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   543
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   544
lemma tendsto_norm_zero_cancel:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   545
  "((\<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
   546
  unfolding tendsto_iff dist_norm by simp
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   547
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   548
lemma tendsto_norm_zero_iff:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   549
  "((\<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
   550
  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
   551
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   552
lemma tendsto_rabs [tendsto_intros]:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   553
  "(f \<longlongrightarrow> (l::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> \<bar>l\<bar>) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   554
  by (fold real_norm_def, rule tendsto_norm)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   555
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
   556
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
   557
  "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
   558
  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
   559
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56366
diff changeset
   560
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
   561
  "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
   562
  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
   563
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   564
lemma tendsto_rabs_zero:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   565
  "(f \<longlongrightarrow> (0::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> 0) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   566
  by (fold real_norm_def, rule tendsto_norm_zero)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   567
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   568
lemma tendsto_rabs_zero_cancel:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   569
  "((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> (0::real)) F \<Longrightarrow> (f \<longlongrightarrow> 0) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   570
  by (fold real_norm_def, rule tendsto_norm_zero_cancel)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   571
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   572
lemma tendsto_rabs_zero_iff:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   573
  "((\<lambda>x. \<bar>f x\<bar>) \<longlongrightarrow> (0::real)) F \<longleftrightarrow> (f \<longlongrightarrow> 0) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   574
  by (fold real_norm_def, rule tendsto_norm_zero_iff)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   575
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   576
subsubsection \<open>Addition and subtraction\<close>
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   577
31565
da5a5589418e theorem attribute [tendsto_intros]
huffman
parents: 31492
diff changeset
   578
lemma tendsto_add [tendsto_intros]:
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   579
  fixes a b :: "'a::real_normed_vector"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   580
  shows "\<lbrakk>(f \<longlongrightarrow> a) F; (g \<longlongrightarrow> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) \<longlongrightarrow> a + b) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   581
  by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add)
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   582
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
   583
lemma continuous_add [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
   584
  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and 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
   585
  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
   586
  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
   587
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56366
diff changeset
   588
lemma continuous_on_add [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
   589
  fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and 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
   590
  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
   591
  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
   592
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   593
lemma tendsto_add_zero:
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
   594
  fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   595
  shows "\<lbrakk>(f \<longlongrightarrow> 0) F; (g \<longlongrightarrow> 0) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) \<longlongrightarrow> 0) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   596
  by (drule (1) tendsto_add, simp)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   597
31565
da5a5589418e theorem attribute [tendsto_intros]
huffman
parents: 31492
diff changeset
   598
lemma tendsto_minus [tendsto_intros]:
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   599
  fixes a :: "'a::real_normed_vector"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   600
  shows "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. - f x) \<longlongrightarrow> - a) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   601
  by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus)
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   602
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
   603
lemma continuous_minus [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
   604
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and 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
   605
  shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. - 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
   606
  unfolding continuous_def by (rule tendsto_minus)
270b21f3ae0a move continuous and 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
   607
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56366
diff changeset
   608
lemma continuous_on_minus [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
   609
  fixes f :: "_ \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and 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
   610
  shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. - 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
   611
  unfolding continuous_on_def by (auto intro: tendsto_minus)
270b21f3ae0a move continuous and 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
   612
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   613
lemma tendsto_minus_cancel:
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   614
  fixes a :: "'a::real_normed_vector"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   615
  shows "((\<lambda>x. - f x) \<longlongrightarrow> - a) F \<Longrightarrow> (f \<longlongrightarrow> a) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   616
  by (drule tendsto_minus, simp)
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   617
50330
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50327
diff changeset
   618
lemma tendsto_minus_cancel_left:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   619
    "(f \<longlongrightarrow> - (y::_::real_normed_vector)) F \<longleftrightarrow> ((\<lambda>x. - f x) \<longlongrightarrow> y) F"
50330
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50327
diff changeset
   620
  using tendsto_minus_cancel[of f "- y" F]  tendsto_minus[of f "- y" F]
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50327
diff changeset
   621
  by auto
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50327
diff changeset
   622
31565
da5a5589418e theorem attribute [tendsto_intros]
huffman
parents: 31492
diff changeset
   623
lemma tendsto_diff [tendsto_intros]:
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   624
  fixes a b :: "'a::real_normed_vector"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   625
  shows "\<lbrakk>(f \<longlongrightarrow> a) F; (g \<longlongrightarrow> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x - g x) \<longlongrightarrow> a - b) F"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53602
diff changeset
   626
  using tendsto_add [of f a F "\<lambda>x. - g x" "- b"] by (simp add: tendsto_minus)
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   627
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
   628
lemma continuous_diff [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
   629
  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and 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
   630
  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
   631
  unfolding continuous_def by (rule tendsto_diff)
270b21f3ae0a move continuous and 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
   632
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56366
diff changeset
   633
lemma continuous_on_diff [continuous_intros]:
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
   634
  fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector"
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
   635
  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
   636
  unfolding continuous_on_def by (auto intro: tendsto_diff)
270b21f3ae0a move continuous and 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
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
   638
lemma continuous_on_op_minus: "continuous_on (s::'a::real_normed_vector set) (op - x)"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
   639
  by (rule continuous_intros | simp)+
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
   640
31588
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   641
lemma tendsto_setsum [tendsto_intros]:
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   642
  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::real_normed_vector"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   643
  assumes "\<And>i. i \<in> S \<Longrightarrow> (f i \<longlongrightarrow> a i) F"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   644
  shows "((\<lambda>x. \<Sum>i\<in>S. f i x) \<longlongrightarrow> (\<Sum>i\<in>S. a i)) F"
31588
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   645
proof (cases "finite S")
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   646
  assume "finite S" thus ?thesis using assms
58729
e8ecc79aee43 add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents: 57512
diff changeset
   647
    by (induct, simp, simp add: tendsto_add)
e8ecc79aee43 add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents: 57512
diff changeset
   648
qed simp
31588
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   649
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
   650
lemma continuous_setsum [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
   651
  fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::real_normed_vector"
270b21f3ae0a move continuous and 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
   652
  shows "(\<And>i. i \<in> S \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Sum>i\<in>S. f i 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
   653
  unfolding continuous_def by (rule tendsto_setsum)
270b21f3ae0a move continuous and 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
   654
270b21f3ae0a move continuous and 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
   655
lemma continuous_on_setsum [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
   656
  fixes f :: "'a \<Rightarrow> _ \<Rightarrow> 'c::real_normed_vector"
270b21f3ae0a move continuous and 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
   657
  shows "(\<And>i. i \<in> S \<Longrightarrow> continuous_on s (f i)) \<Longrightarrow> continuous_on s (\<lambda>x. \<Sum>i\<in>S. f i 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
   658
  unfolding continuous_on_def by (auto intro: tendsto_setsum)
270b21f3ae0a move continuous and 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
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50880
diff changeset
   660
lemmas real_tendsto_sandwich = tendsto_sandwich[where 'b=real]
3de230ed0547 introduce order topology
hoelzl
parents: 50880
diff changeset
   661
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   662
subsubsection \<open>Linear operators and multiplication\<close>
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   663
61806
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
   664
lemma linear_times:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
   665
  fixes c::"'a::real_algebra" shows "linear (\<lambda>x. c * x)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
   666
  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
   667
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   668
lemma (in bounded_linear) tendsto:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   669
  "(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
   670
  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
   671
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
   672
lemma (in bounded_linear) 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
   673
  "continuous F g \<Longrightarrow> continuous F (\<lambda>x. f (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
   674
  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
   675
270b21f3ae0a move continuous and 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
   676
lemma (in bounded_linear) 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
   677
  "continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f (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
   678
  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
   679
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   680
lemma (in bounded_linear) tendsto_zero:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   681
  "(g \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f (g x)) \<longlongrightarrow> 0) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   682
  by (drule tendsto, simp only: zero)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   683
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   684
lemma (in bounded_bilinear) tendsto:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   685
  "\<lbrakk>(f \<longlongrightarrow> a) F; (g \<longlongrightarrow> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x ** g x) \<longlongrightarrow> a ** b) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   686
  by (simp only: tendsto_Zfun_iff prod_diff_prod
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   687
                 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
   688
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
   689
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
   690
  "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
   691
  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
   692
270b21f3ae0a move continuous and 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
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
   694
  "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
   695
  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
   696
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   697
lemma (in bounded_bilinear) tendsto_zero:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   698
  assumes f: "(f \<longlongrightarrow> 0) F"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   699
  assumes g: "(g \<longlongrightarrow> 0) F"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   700
  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
   701
  using tendsto [OF f g] by (simp add: zero_left)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   702
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   703
lemma (in bounded_bilinear) tendsto_left_zero:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   704
  "(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
   705
  by (rule bounded_linear.tendsto_zero [OF bounded_linear_left])
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   706
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   707
lemma (in bounded_bilinear) tendsto_right_zero:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   708
  "(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
   709
  by (rule bounded_linear.tendsto_zero [OF bounded_linear_right])
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   710
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   711
lemmas tendsto_of_real [tendsto_intros] =
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   712
  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
   713
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   714
lemmas tendsto_scaleR [tendsto_intros] =
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   715
  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
   716
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   717
lemmas tendsto_mult [tendsto_intros] =
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   718
  bounded_bilinear.tendsto [OF bounded_bilinear_mult]
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   719
61806
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
   720
lemma tendsto_mult_left:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
   721
  fixes c::"'a::real_normed_algebra"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   722
  shows "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. c * (f x)) \<longlongrightarrow> c * l) F"
61806
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
   723
by (rule tendsto_mult [OF tendsto_const])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
   724
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
   725
lemma tendsto_mult_right:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
   726
  fixes c::"'a::real_normed_algebra"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   727
  shows "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. (f x) * c) \<longlongrightarrow> l * c) F"
61806
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
   728
by (rule tendsto_mult [OF _ tendsto_const])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
   729
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
   730
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
   731
  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
   732
270b21f3ae0a move continuous and 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
   733
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
   734
  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
   735
270b21f3ae0a move continuous and 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
   736
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
   737
  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
   738
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56366
diff changeset
   739
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
   740
  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
   741
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56366
diff changeset
   742
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
   743
  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
   744
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56366
diff changeset
   745
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
   746
  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
   747
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44342
diff changeset
   748
lemmas tendsto_mult_zero =
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44342
diff changeset
   749
  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
   750
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44342
diff changeset
   751
lemmas tendsto_mult_left_zero =
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44342
diff changeset
   752
  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
   753
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44342
diff changeset
   754
lemmas tendsto_mult_right_zero =
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44342
diff changeset
   755
  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
   756
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   757
lemma tendsto_power [tendsto_intros]:
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   758
  fixes f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra}"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   759
  shows "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f x ^ n) \<longlongrightarrow> a ^ n) F"
58729
e8ecc79aee43 add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents: 57512
diff changeset
   760
  by (induct n) (simp_all add: tendsto_mult)
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   761
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
   762
lemma continuous_power [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
   763
  fixes f :: "'a::t2_space \<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
   764
  shows "continuous F f \<Longrightarrow> continuous F (\<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
   765
  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
   766
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56366
diff changeset
   767
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
   768
  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
   769
  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
   770
  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
   771
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   772
lemma tendsto_setprod [tendsto_intros]:
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   773
  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   774
  assumes "\<And>i. i \<in> S \<Longrightarrow> (f i \<longlongrightarrow> L i) F"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   775
  shows "((\<lambda>x. \<Prod>i\<in>S. f i x) \<longlongrightarrow> (\<Prod>i\<in>S. L i)) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   776
proof (cases "finite S")
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   777
  assume "finite S" thus ?thesis using assms
58729
e8ecc79aee43 add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents: 57512
diff changeset
   778
    by (induct, simp, simp add: tendsto_mult)
e8ecc79aee43 add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents: 57512
diff changeset
   779
qed simp
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   780
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
   781
lemma continuous_setprod [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
   782
  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
   783
  shows "(\<And>i. i \<in> S \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Prod>i\<in>S. f i 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
   784
  unfolding continuous_def by (rule tendsto_setprod)
270b21f3ae0a move continuous and 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
   785
270b21f3ae0a move continuous and 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
   786
lemma continuous_on_setprod [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
   787
  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
   788
  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)"
270b21f3ae0a move continuous and 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
   789
  unfolding continuous_on_def by (auto intro: tendsto_setprod)
270b21f3ae0a move continuous and 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
   790
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   791
lemma tendsto_of_real_iff:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   792
  "((\<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
   793
  unfolding tendsto_iff by simp
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   794
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   795
lemma tendsto_add_const_iff:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   796
  "((\<lambda>x. c + f x :: 'a :: real_normed_vector) \<longlongrightarrow> c + d) F \<longleftrightarrow> (f \<longlongrightarrow> d) F"
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
   797
  using tendsto_add[OF tendsto_const[of c], of f d]
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   798
        tendsto_add[OF tendsto_const[of "-c"], of "\<lambda>x. c + f x" "c + d"] by auto
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   799
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
   800
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
   801
subsubsection \<open>Inverse and division\<close>
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   802
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   803
lemma (in bounded_bilinear) Zfun_prod_Bfun:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   804
  assumes f: "Zfun f F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   805
  assumes g: "Bfun g F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   806
  shows "Zfun (\<lambda>x. f x ** g x) F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   807
proof -
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   808
  obtain K where K: "0 \<le> K"
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   809
    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
   810
    using nonneg_bounded by blast
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   811
  obtain B where B: "0 < B"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   812
    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
   813
    using g by (rule BfunE)
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   814
  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
   815
  using norm_g proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   816
    case (elim x)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   817
    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
   818
      by (rule norm_le)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   819
    also have "\<dots> \<le> norm (f x) * B * K"
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   820
      by (intro mult_mono' order_refl norm_g norm_ge_zero
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   821
                mult_nonneg_nonneg K elim)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   822
    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
   823
      by (rule mult.assoc)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   824
    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
   825
  qed
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   826
  with f show ?thesis
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   827
    by (rule Zfun_imp_Zfun)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   828
qed
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   829
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   830
lemma (in bounded_bilinear) Bfun_prod_Zfun:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   831
  assumes f: "Bfun f F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   832
  assumes g: "Zfun g F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   833
  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
   834
  using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   835
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   836
lemma Bfun_inverse_lemma:
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   837
  fixes x :: "'a::real_normed_div_algebra"
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   838
  shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   839
  apply (subst nonzero_norm_inverse, clarsimp)
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   840
  apply (erule (1) le_imp_inverse_le)
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   841
  done
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   842
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   843
lemma Bfun_inverse:
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   844
  fixes a :: "'a::real_normed_div_algebra"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   845
  assumes f: "(f \<longlongrightarrow> a) F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   846
  assumes a: "a \<noteq> 0"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   847
  shows "Bfun (\<lambda>x. inverse (f x)) F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   848
proof -
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   849
  from a have "0 < norm a" by simp
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   850
  hence "\<exists>r>0. r < norm a" by (rule dense)
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
   851
  then obtain r where r1: "0 < r" and r2: "r < norm a" by blast
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   852
  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
   853
    using tendstoD [OF f r1] by blast
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   854
  hence "eventually (\<lambda>x. norm (inverse (f x)) \<le> inverse (norm a - r)) F"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   855
  proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   856
    case (elim x)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   857
    hence 1: "norm (f x - a) < r"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   858
      by (simp add: dist_norm)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   859
    hence 2: "f x \<noteq> 0" using r2 by auto
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   860
    hence "norm (inverse (f x)) = inverse (norm (f x))"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   861
      by (rule nonzero_norm_inverse)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   862
    also have "\<dots> \<le> inverse (norm a - r)"
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   863
    proof (rule le_imp_inverse_le)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   864
      show "0 < norm a - r" using r2 by simp
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   865
    next
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   866
      have "norm a - norm (f x) \<le> norm (a - f x)"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   867
        by (rule norm_triangle_ineq2)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   868
      also have "\<dots> = norm (f x - a)"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   869
        by (rule norm_minus_commute)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   870
      also have "\<dots> < r" using 1 .
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   871
      finally show "norm a - r \<le> norm (f x)" by simp
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   872
    qed
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   873
    finally show "norm (inverse (f x)) \<le> inverse (norm a - r)" .
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   874
  qed
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   875
  thus ?thesis by (rule BfunI)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   876
qed
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   877
31565
da5a5589418e theorem attribute [tendsto_intros]
huffman
parents: 31492
diff changeset
   878
lemma tendsto_inverse [tendsto_intros]:
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   879
  fixes a :: "'a::real_normed_div_algebra"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   880
  assumes f: "(f \<longlongrightarrow> a) F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   881
  assumes a: "a \<noteq> 0"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   882
  shows "((\<lambda>x. inverse (f x)) \<longlongrightarrow> inverse a) F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   883
proof -
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   884
  from a have "0 < norm a" by simp
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   885
  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
   886
    by (rule tendstoD)
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   887
  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
   888
    unfolding dist_norm by (auto elim!: eventually_mono)
44627
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   889
  with a have "eventually (\<lambda>x. inverse (f x) - inverse a =
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   890
    - (inverse (f x) * (f x - a) * inverse a)) F"
61810
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61806
diff changeset
   891
    by (auto elim!: eventually_mono simp: inverse_diff_inverse)
44627
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   892
  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
   893
    by (intro Zfun_minus Zfun_mult_left
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   894
      bounded_bilinear.Bfun_prod_Zfun [OF bounded_bilinear_mult]
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   895
      Bfun_inverse [OF f a] f [unfolded tendsto_Zfun_iff])
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   896
  ultimately show ?thesis
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   897
    unfolding tendsto_Zfun_iff by (rule Zfun_ssubst)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   898
qed
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   899
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
   900
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
   901
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_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
   902
  assumes "continuous F f" and "f (Lim F (\<lambda>x. x)) \<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
   903
  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
   904
  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
   905
270b21f3ae0a move continuous and 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
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
   907
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_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
   908
  assumes "continuous (at a within s) f" and "f 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
   909
  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
   910
  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
   911
270b21f3ae0a move continuous and 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
lemma isCont_inverse[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
   913
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_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
   914
  assumes "isCont f a" and "f 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
   915
  shows "isCont (\<lambda>x. inverse (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
   916
  using assms unfolding continuous_at 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
   917
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56366
diff changeset
   918
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
   919
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_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
   920
  assumes "continuous_on s f" and "\<forall>x\<in>s. f x \<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
   921
  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
   922
  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
   923
31565
da5a5589418e theorem attribute [tendsto_intros]
huffman
parents: 31492
diff changeset
   924
lemma tendsto_divide [tendsto_intros]:
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   925
  fixes a b :: "'a::real_normed_field"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   926
  shows "\<lbrakk>(f \<longlongrightarrow> a) F; (g \<longlongrightarrow> b) F; b \<noteq> 0\<rbrakk>
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   927
    \<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
   928
  by (simp add: tendsto_mult tendsto_inverse divide_inverse)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   929
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
   930
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
   931
  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
   932
  assumes "continuous F f" and "continuous F g" and "g (Lim F (\<lambda>x. x)) \<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
   933
  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
   934
  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
   935
270b21f3ae0a move continuous and 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
   936
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
   937
  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
   938
  assumes "continuous (at a within s) f" "continuous (at a within s) g" and "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
   939
  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
   940
  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
   941
270b21f3ae0a move continuous and 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
   942
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
   943
  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
   944
  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
   945
  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
   946
  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
   947
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56366
diff changeset
   948
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
   949
  fixes f :: "'a::topological_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
   950
  assumes "continuous_on s f" "continuous_on s g" and "\<forall>x\<in>s. g x \<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
   951
  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
   952
  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
   953
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   954
lemma tendsto_sgn [tendsto_intros]:
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   955
  fixes l :: "'a::real_normed_vector"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   956
  shows "\<lbrakk>(f \<longlongrightarrow> l) F; l \<noteq> 0\<rbrakk> \<Longrightarrow> ((\<lambda>x. sgn (f x)) \<longlongrightarrow> sgn l) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   957
  unfolding sgn_div_norm by (simp add: tendsto_intros)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   958
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
   959
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
   960
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and 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
   961
  assumes "continuous F f" and "f (Lim F (\<lambda>x. x)) \<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
   962
  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
   963
  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
   964
270b21f3ae0a move continuous and 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
   965
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
   966
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and 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
   967
  assumes "continuous (at a within s) f" and "f 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
   968
  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
   969
  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
   970
270b21f3ae0a move continuous and 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
   971
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
   972
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and 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
  assumes "isCont f a" and "f 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
   974
  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
   975
  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
   976
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56366
diff changeset
   977
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
   978
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and 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
  assumes "continuous_on s f" and "\<forall>x\<in>s. f x \<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
   980
  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
   981
  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
   982
50325
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   983
lemma filterlim_at_infinity:
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60974
diff changeset
   984
  fixes f :: "_ \<Rightarrow> 'a::real_normed_vector"
50325
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   985
  assumes "0 \<le> c"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   986
  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
   987
  unfolding filterlim_iff eventually_at_infinity
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   988
proof safe
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   989
  fix P :: "'a \<Rightarrow> bool" and b
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   990
  assume *: "\<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
   991
    and P: "\<forall>x. b \<le> norm x \<longrightarrow> P x"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   992
  have "max b (c + 1) > c" by auto
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   993
  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
   994
    by auto
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   995
  then show "eventually (\<lambda>x. P (f x)) F"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   996
  proof eventually_elim
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   997
    fix x assume "max b (c + 1) \<le> norm (f x)"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   998
    with P show "P (f x)" by auto
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   999
  qed
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1000
qed force
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1001
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1002
lemma not_tendsto_and_filterlim_at_infinity:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1003
  assumes "F \<noteq> bot"
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1004
  assumes "(f \<longlongrightarrow> (c :: 'a :: real_normed_vector)) F"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1005
  assumes "filterlim f at_infinity F"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1006
  shows   False
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1007
proof -
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1008
  from tendstoD[OF assms(2), of "1/2"]
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1009
    have "eventually (\<lambda>x. dist (f x) c < 1/2) F" by simp
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1010
  moreover from filterlim_at_infinity[of "norm c" f F] assms(3)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1011
    have "eventually (\<lambda>x. norm (f x) \<ge> norm c + 1) F" by simp
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1012
  ultimately have "eventually (\<lambda>x. False) F"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1013
  proof eventually_elim
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1014
    fix x assume A: "dist (f x) c < 1/2" and B: "norm (f x) \<ge> norm c + 1"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1015
    note B
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1016
    also have "norm (f x) = dist (f x) 0" by (simp add: norm_conv_dist)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1017
    also have "... \<le> dist (f x) c + dist c 0" by (rule dist_triangle)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1018
    also note A
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1019
    finally show False by (simp add: norm_conv_dist)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1020
  qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1021
  with assms show False by simp
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1022
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1023
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1024
lemma filterlim_at_infinity_imp_not_convergent:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1025
  assumes "filterlim f at_infinity sequentially"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1026
  shows   "\<not>convergent f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1027
  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
  1028
     (simp_all add: convergent_LIMSEQ_iff)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1029
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1030
lemma filterlim_at_infinity_imp_eventually_ne:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1031
  assumes "filterlim f at_infinity F"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1032
  shows   "eventually (\<lambda>z. f z \<noteq> c) F"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1033
proof -
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1034
  have "norm c + 1 > 0" by (intro add_nonneg_pos) simp_all
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1035
  with filterlim_at_infinity[OF order.refl, of f F] assms
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1036
    have "eventually (\<lambda>z. norm (f z) \<ge> norm c + 1) F" by blast
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1037
  thus ?thesis by eventually_elim auto
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1038
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1039
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1040
lemma tendsto_of_nat [tendsto_intros]:
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1041
  "filterlim (of_nat :: nat \<Rightarrow> 'a :: real_normed_algebra_1) at_infinity sequentially"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1042
proof (subst filterlim_at_infinity[OF order.refl], intro allI impI)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1043
  fix r :: real assume r: "r > 0"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1044
  def n \<equiv> "nat \<lceil>r\<rceil>"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1045
  from r have n: "\<forall>m\<ge>n. of_nat m \<ge> r" unfolding n_def by linarith
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1046
  from eventually_ge_at_top[of n] show "eventually (\<lambda>m. norm (of_nat m :: 'a) \<ge> r) sequentially"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1047
    by eventually_elim (insert n, simp_all)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1048
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1049
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1050
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1051
subsection \<open>Relate @{const at}, @{const at_left} and @{const at_right}\<close>
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1052
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1053
text \<open>
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1054
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1055
This lemmas are useful for conversion between @{term "at x"} to @{term "at_left x"} and
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1056
@{term "at_right x"} and also @{term "at_right 0"}.
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1057
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1058
\<close>
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1059
51471
cad22a3cc09c move topological_space to its own theory
hoelzl
parents: 51360
diff changeset
  1060
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
  1061
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
  1062
lemma filtermap_nhds_shift: "filtermap (\<lambda>x. x - d) (nhds a) = nhds (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
  1063
  by (rule filtermap_fun_inverse[where g="\<lambda>x. x + d"])
c1b7793c23a3 generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents: 60182
diff changeset
  1064
     (auto intro!: tendsto_eq_intros filterlim_ident)
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1065
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
  1066
lemma filtermap_nhds_minus: "filtermap (\<lambda>x. - x) (nhds a) = nhds (- 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
  1067
  by (rule filtermap_fun_inverse[where g=uminus])
c1b7793c23a3 generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents: 60182
diff changeset
  1068
     (auto intro!: tendsto_eq_intros filterlim_ident)
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
  1069
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
  1070
lemma filtermap_at_shift: "filtermap (\<lambda>x. x - d) (at a) = at (a - d::'a::real_normed_vector)"
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
  1071
  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
  1072
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1073
lemma filtermap_at_right_shift: "filtermap (\<lambda>x. x - d) (at_right a) = at_right (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
  1074
  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
  1075
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1076
lemma at_right_to_0: "at_right (a::real) = filtermap (\<lambda>x. x + a) (at_right 0)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1077
  using filtermap_at_right_shift[of "-a" 0] by simp
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1078
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1079
lemma filterlim_at_right_to_0:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1080
  "filterlim f F (at_right (a::real)) \<longleftrightarrow> filterlim (\<lambda>x. f (x + a)) F (at_right 0)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1081
  unfolding filterlim_def filtermap_filtermap at_right_to_0[of a] ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1082
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1083
lemma eventually_at_right_to_0:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1084
  "eventually P (at_right (a::real)) \<longleftrightarrow> eventually (\<lambda>x. P (x + a)) (at_right 0)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1085
  unfolding at_right_to_0[of a] by (simp add: eventually_filtermap)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1086
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
  1087
lemma filtermap_at_minus: "filtermap (\<lambda>x. - x) (at a) = at (- a::'a::real_normed_vector)"
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
  1088
  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
  1089
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1090
lemma at_left_minus: "at_left (a::real) = filtermap (\<lambda>x. - x) (at_right (- a))"
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
  1091
  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
  1092
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1093
lemma at_right_minus: "at_right (a::real) = filtermap (\<lambda>x. - x) (at_left (- a))"
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
  1094
  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
  1095
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1096
lemma filterlim_at_left_to_right:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1097
  "filterlim f F (at_left (a::real)) \<longleftrightarrow> filterlim (\<lambda>x. f (- x)) F (at_right (-a))"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1098
  unfolding filterlim_def filtermap_filtermap at_left_minus[of a] ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1099
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1100
lemma eventually_at_left_to_right:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1101
  "eventually P (at_left (a::real)) \<longleftrightarrow> eventually (\<lambda>x. P (- x)) (at_right (-a))"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1102
  unfolding at_left_minus[of a] by (simp add: eventually_filtermap)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1103
60721
c1b7793c23a3 generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents: 60182
diff changeset
  1104
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
  1105
  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
  1106
  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
  1107
c1b7793c23a3 generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents: 60182
diff changeset
  1108
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
  1109
  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
  1110
  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
  1111
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
  1112
lemma at_top_mirror: "at_top = filtermap uminus (at_bot :: real filter)"
60721
c1b7793c23a3 generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents: 60182
diff changeset
  1113
  by (rule filtermap_fun_inverse[symmetric, of uminus])
c1b7793c23a3 generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents: 60182
diff changeset
  1114
     (auto intro: filterlim_uminus_at_bot_at_top filterlim_uminus_at_top_at_bot)
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
  1115
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
  1116
lemma at_bot_mirror: "at_bot = filtermap uminus (at_top :: real filter)"
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
  1117
  unfolding at_top_mirror filtermap_filtermap by (simp add: filtermap_ident)
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
  1118
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
  1119
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
  1120
  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
  1121
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
  1122
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
  1123
  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
  1124
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
  1125
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
  1126
  using filterlim_compose[OF filterlim_uminus_at_bot_at_top, of f 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
  1127
  using filterlim_compose[OF filterlim_uminus_at_top_at_bot, of "\<lambda>x. - f x" 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
  1128
  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
  1129
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
  1130
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
  1131
  unfolding filterlim_uminus_at_top by simp
50323
3764d4620fb3 add filterlim rules for unary minus and inverse
hoelzl
parents: 50322
diff changeset
  1132
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1133
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
  1134
  unfolding filterlim_at_top_gt[where c=0] eventually_at_filter
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1135
proof safe
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1136
  fix Z :: real assume [arith]: "0 < Z"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1137
  then have "eventually (\<lambda>x. x < inverse Z) (nhds 0)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1138
    by (auto simp add: 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
  1139
  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
  1140
    by (auto elim!: eventually_mono simp: inverse_eq_divide field_simps)
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1141
qed
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1142
50325
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1143
lemma tendsto_inverse_0:
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60974
diff changeset
  1144
  fixes x :: "_ \<Rightarrow> 'a::real_normed_div_algebra"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1145
  shows "(inverse \<longlongrightarrow> (0::'a)) at_infinity"
50325
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1146
  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
  1147
proof safe
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1148
  fix r :: real assume "0 < r"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1149
  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
  1150
  proof (intro exI[of _ "inverse (r / 2)"] allI impI)
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1151
    fix x :: 'a
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1152
    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
  1153
    also assume *: "inverse (r / 2) \<le> norm x"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1154
    finally show "norm (inverse x) < r"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1155
      using * \<open>0 < r\<close> 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
  1156
  qed
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1157
qed
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1158
61552
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1159
lemma tendsto_add_filterlim_at_infinity:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1160
  assumes "(f \<longlongrightarrow> (c :: 'b :: real_normed_vector)) (F :: 'a filter)"
61552
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1161
  assumes "filterlim g at_infinity F"
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1162
  shows   "filterlim (\<lambda>x. f x + g x) at_infinity F"
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1163
proof (subst filterlim_at_infinity[OF order_refl], safe)
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1164
  fix r :: real assume r: "r > 0"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1165
  from assms(1) have "((\<lambda>x. norm (f x)) \<longlongrightarrow> norm c) F" by (rule tendsto_norm)
61552
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1166
  hence "eventually (\<lambda>x. norm (f x) < norm c + 1) F" by (rule order_tendstoD) simp_all
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1167
  moreover from r have "r + norm c + 1 > 0" 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
  1168
  with assms(2) have "eventually (\<lambda>x. norm (g x) \<ge> r + norm c + 1) F"
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1169
    unfolding filterlim_at_infinity[OF order_refl] by (elim allE[of _ "r + norm c + 1"]) simp_all
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1170
  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
  1171
  proof eventually_elim
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1172
    fix x :: 'a assume A: "norm (f x) < norm c + 1" and B: "r + norm c + 1 \<le> norm (g x)"
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1173
    from A B have "r \<le> norm (g x) - norm (f x)" by simp
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1174
    also have "norm (g x) - norm (f x) \<le> norm (g x + f x)" by (rule norm_diff_ineq)
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1175
    finally show "r \<le> norm (f x + g x)" by (simp add: add_ac)
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1176
  qed
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1177
qed
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1178
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1179
lemma tendsto_add_filterlim_at_infinity':
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1180
  assumes "filterlim f at_infinity F"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1181
  assumes "(g \<longlongrightarrow> (c :: 'b :: real_normed_vector)) (F :: 'a filter)"
61552
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1182
  shows   "filterlim (\<lambda>x. f x + g x) at_infinity F"
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
  1183
  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
  1184
60721
c1b7793c23a3 generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents: 60182
diff changeset
  1185
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
  1186
  unfolding filterlim_at
c1b7793c23a3 generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents: 60182
diff changeset
  1187
  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
  1188
     (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
  1189
c1b7793c23a3 generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents: 60182
diff changeset
  1190
lemma filterlim_inverse_at_top:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1191
  "(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
  1192
  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
  1193
     (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
  1194
c1b7793c23a3 generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents: 60182
diff changeset
  1195
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
  1196
  "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
  1197
  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
  1198
c1b7793c23a3 generalized filtermap_homeomorph to filtermap_fun_inverse; add eventually_at_top/bot_not_equal
hoelzl
parents: 60182
diff changeset
  1199
lemma filterlim_inverse_at_bot:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1200
  "(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
  1201
  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
  1202
  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
  1203
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1204
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
  1205
  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
  1206
     (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
  1207
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1208
lemma eventually_at_right_to_top:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1209
  "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
  1210
  unfolding at_right_to_top eventually_filtermap ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1211
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1212
lemma filterlim_at_right_to_top:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1213
  "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
  1214
  unfolding filterlim_def at_right_to_top filtermap_filtermap ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1215
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1216
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
  1217
  unfolding at_right_to_top filtermap_filtermap inverse_inverse_eq filtermap_ident ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1218
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1219
lemma eventually_at_top_to_right:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1220
  "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
  1221
  unfolding at_top_to_right eventually_filtermap ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1222
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1223
lemma filterlim_at_top_to_right:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1224
  "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
  1225
  unfolding filterlim_def at_top_to_right filtermap_filtermap ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1226
50325
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1227
lemma filterlim_inverse_at_infinity:
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60974
diff changeset
  1228
  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
  1229
  shows "filterlim inverse at_infinity (at (0::'a))"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1230
  unfolding filterlim_at_infinity[OF order_refl]
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1231
proof safe
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1232
  fix r :: real assume "0 < r"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1233
  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
  1234
    unfolding eventually_at norm_inverse
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1235
    by (intro exI[of _ "inverse r"])
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1236
       (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
  1237
qed
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1238
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1239
lemma filterlim_inverse_at_iff:
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60974
diff changeset
  1240
  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
  1241
  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
  1242
  unfolding filterlim_def filtermap_filtermap[symmetric]
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1243
proof
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1244
  assume "filtermap g F \<le> at_infinity"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1245
  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
  1246
    by (rule filtermap_mono)
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1247
  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
  1248
    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
  1249
    by (auto intro!: exI[of _ 1]
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
  1250
             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
  1251
  finally show "filtermap inverse (filtermap g F) \<le> at 0" .
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1252
next
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1253
  assume "filtermap inverse (filtermap g F) \<le> at 0"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1254
  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
  1255
    by (rule filtermap_mono)
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1256
  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
  1257
    by (auto intro: order_trans simp: filterlim_def filtermap_filtermap)
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1258
qed
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
  1259
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1260
lemma tendsto_mult_filterlim_at_infinity:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1261
  assumes "F \<noteq> bot" "(f \<longlongrightarrow> (c :: 'a :: real_normed_field)) F" "c \<noteq> 0"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1262
  assumes "filterlim g at_infinity F"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1263
  shows   "filterlim (\<lambda>x. f x * g x) at_infinity F"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1264
proof -
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1265
  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
  1266
    by (intro tendsto_mult tendsto_inverse assms filterlim_compose[OF tendsto_inverse_0])
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1267
  hence "filterlim (\<lambda>x. inverse (f x) * inverse (g x)) (at (inverse c * 0)) F"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1268
    unfolding filterlim_at using assms
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1269
    by (auto intro: filterlim_at_infinity_imp_eventually_ne tendsto_imp_eventually_ne eventually_conj)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1270
  thus ?thesis by (subst filterlim_inverse_at_iff[symmetric]) simp_all
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1271
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1272
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1273
lemma tendsto_inverse_0_at_top: "LIM x F. f x :> at_top \<Longrightarrow> ((\<lambda>x. inverse (f x) :: real) \<longlongrightarrow> 0) F"
51641
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51531
diff changeset
  1274
 by (metis filterlim_at filterlim_mono[OF _ at_top_le_at_infinity order_refl] filterlim_inverse_at_iff)
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
  1275
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1276
lemma mult_nat_left_at_top: "c > 0 \<Longrightarrow> filterlim (\<lambda>x. c * x :: nat) at_top sequentially"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1277
  by (rule filterlim_subseq) (auto simp: subseq_def)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1278
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1279
lemma mult_nat_right_at_top: "c > 0 \<Longrightarrow> filterlim (\<lambda>x. x * c :: nat) at_top sequentially"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1280
  by (rule filterlim_subseq) (auto simp: subseq_def)
59613
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1281
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1282
lemma at_to_infinity:
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60974
diff changeset
  1283
  fixes x :: "'a :: {real_normed_field,field}"
59613
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1284
  shows "(at (0::'a)) = 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
  1285
proof (rule antisym)
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1286
  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
  1287
    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
  1288
  then show "filtermap inverse at_infinity \<le> at (0::'a)"
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1289
    apply (simp add: le_principal eventually_filtermap eventually_at_infinity filterlim_def at_within_def)
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1290
    apply (rule_tac x="1" in exI, auto)
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1291
    done
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1292
next
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1293
  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
  1294
    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
  1295
    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
  1296
  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
  1297
    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
  1298
qed
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1299
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1300
lemma lim_at_infinity_0:
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59613
diff changeset
  1301
  fixes l :: "'a :: {real_normed_field,field}"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1302
  shows "(f \<longlongrightarrow> l) at_infinity \<longleftrightarrow> ((f o inverse) \<longlongrightarrow> l) (at (0::'a))"
59613
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1303
by (simp add: tendsto_compose_filtermap at_to_infinity filtermap_filtermap)
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1304
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1305
lemma lim_zero_infinity:
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59613
diff changeset
  1306
  fixes l :: "'a :: {real_normed_field,field}"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1307
  shows "((\<lambda>x. f(1 / x)) \<longlongrightarrow> l) (at (0::'a)) \<Longrightarrow> (f \<longlongrightarrow> l) 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
  1308
by (simp add: inverse_eq_divide lim_at_infinity_0 comp_def)
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1309
7103019278f0 The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
  1310
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1311
text \<open>
50324
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1312
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1313
We only show rules for multiplication and addition when the functions are either against a real
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1314
value or against infinity. Further rules are easy to derive by using @{thm filterlim_uminus_at_top}.
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1315
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1316
\<close>
50324
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1317
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1318
lemma filterlim_tendsto_pos_mult_at_top:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1319
  assumes f: "(f \<longlongrightarrow> c) F" and c: "0 < c"
50324
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1320
  assumes g: "LIM x F. g x :> at_top"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1321
  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
  1322
  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
  1323
proof safe
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1324
  fix Z :: real assume "0 < Z"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1325
  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
  1326
    by (auto dest!: tendstoD[where e="c / 2"] elim!: eventually_mono
50324
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1327
             simp: dist_real_def abs_real_def split: split_if_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
  1328
  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
  1329
    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
  1330
  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
  1331
  proof eventually_elim
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
  1332
    fix x assume "c / 2 < f x" "Z / c * 2 \<le> g x"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1333
    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
  1334
      by (intro mult_mono) (auto simp: zero_le_divide_iff)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1335
    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
  1336
       by simp
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1337
  qed
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1338
qed
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1339
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1340
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
  1341
  assumes f: "LIM x F. f x :> at_top"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1342
  assumes g: "LIM x F. g x :> at_top"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1343
  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
  1344
  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
  1345
proof safe
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1346
  fix Z :: real 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
  1347
  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
  1348
    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
  1349
  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
  1350
    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
  1351
  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
  1352
  proof eventually_elim
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
  1353
    fix x assume "1 \<le> f x" "Z \<le> g x"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1354
    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
  1355
      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
  1356
    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
  1357
       by simp
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1358
  qed
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1359
qed
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1360
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
  1361
lemma filterlim_tendsto_pos_mult_at_bot:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1362
  assumes "(f \<longlongrightarrow> c) F" "0 < (c::real)" "filterlim g at_bot F"
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
  1363
  shows "LIM x F. f x * g x :> at_bot"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
  1364
  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
  1365
  unfolding filterlim_uminus_at_bot by simp
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
  1366
60182
e1ea5a6379c9 generalized tends over powr; added DERIV rule for powr
hoelzl
parents: 60141
diff changeset
  1367
lemma filterlim_tendsto_neg_mult_at_bot:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1368
  assumes c: "(f \<longlongrightarrow> c) F" "(c::real) < 0" and g: "filterlim g at_top F"
60182
e1ea5a6379c9 generalized tends over powr; added DERIV rule for powr
hoelzl
parents: 60141
diff changeset
  1369
  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
  1370
  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
  1371
  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
  1372
56330
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1373
lemma filterlim_pow_at_top:
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1374
  fixes f :: "real \<Rightarrow> real"
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1375
  assumes "0 < n" and f: "LIM x F. f x :> at_top"
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1376
  shows "LIM x F. (f x)^n :: real :> at_top"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1377
using \<open>0 < n\<close> proof (induct n)
56330
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1378
  case (Suc n) with f show ?case
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1379
    by (cases "n = 0") (auto intro!: filterlim_at_top_mult_at_top)
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1380
qed simp
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1381
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1382
lemma filterlim_pow_at_bot_even:
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1383
  fixes f :: "real \<Rightarrow> real"
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1384
  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
  1385
  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
  1386
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1387
lemma filterlim_pow_at_bot_odd:
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1388
  fixes f :: "real \<Rightarrow> real"
5c4d3be7a6b0 add limits of power at top and bot
hoelzl
parents: 55415
diff changeset
  1389
  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
  1390
  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
  1391
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1392
lemma filterlim_tendsto_add_at_top:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1393
  assumes f: "(f \<longlongrightarrow> c) F"
50324
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1394
  assumes g: "LIM x F. g x :> at_top"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1395
  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
  1396
  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
  1397
proof safe
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1398
  fix Z :: real assume "0 < Z"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1399
  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
  1400
    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
  1401
  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
  1402
    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
  1403
  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
  1404
    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
  1405
qed
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1406
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1407
lemma LIM_at_top_divide:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1408
  fixes f g :: "'a \<Rightarrow> real"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1409
  assumes f: "(f \<longlongrightarrow> a) F" "0 < a"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1410
  assumes 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
  1411
  shows "LIM x F. f x / g x :> at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1412
  unfolding divide_inverse
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1413
  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
  1414
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1415
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
  1416
  assumes f: "LIM x F. f x :> at_top"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1417
  assumes g: "LIM x F. g x :> at_top"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1418
  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
  1419
  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
  1420
proof safe
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1421
  fix Z :: real 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
  1422
  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
  1423
    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
  1424
  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
  1425
    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
  1426
  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
  1427
    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
  1428
qed
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1429
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1430
lemma tendsto_divide_0:
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60974
diff changeset
  1431
  fixes f :: "_ \<Rightarrow> 'a::{real_normed_div_algebra, division_ring}"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1432
  assumes f: "(f \<longlongrightarrow> c) F"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1433
  assumes g: "LIM x F. g x :> at_infinity"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1434
  shows "((\<lambda>x. f x / g x) \<longlongrightarrow> 0) F"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1435
  using tendsto_mult[OF f filterlim_compose[OF tendsto_inverse_0 g]] by (simp add: divide_inverse)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1436
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1437
lemma linear_plus_1_le_power:
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1438
  fixes x :: real
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1439
  assumes x: "0 \<le> x"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1440
  shows "real n * x + 1 \<le> (x + 1) ^ n"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1441
proof (induct n)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1442
  case (Suc n)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1443
  have "real (Suc n) * x + 1 \<le> (x + 1) * (real n * x + 1)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61552
diff changeset
  1444
    by (simp add: field_simps of_nat_Suc x)
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1445
  also have "\<dots> \<le> (x + 1)^Suc n"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1446
    using Suc x by (simp add: mult_left_mono)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1447
  finally show ?case .
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1448
qed simp
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1449
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1450
lemma filterlim_realpow_sequentially_gt1:
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1451
  fixes x :: "'a :: real_normed_div_algebra"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1452
  assumes x[arith]: "1 < norm x"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1453
  shows "LIM n sequentially. x ^ n :> at_infinity"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1454
proof (intro filterlim_at_infinity[THEN iffD2] allI impI)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1455
  fix y :: real assume "0 < y"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1456
  have "0 < norm x - 1" by simp
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1457
  then obtain N::nat where "y < real N * (norm x - 1)" by (blast dest: reals_Archimedean3)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1458
  also have "\<dots> \<le> real N * (norm x - 1) + 1" by simp
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1459
  also have "\<dots> \<le> (norm x - 1 + 1) ^ N" by (rule linear_plus_1_le_power) simp
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1460
  also have "\<dots> = norm x ^ N" by simp
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1461
  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
  1462
    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
  1463
  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
  1464
    unfolding eventually_sequentially
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1465
    by (auto simp: norm_power)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1466
qed simp
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1467
51471
cad22a3cc09c move topological_space to its own theory
hoelzl
parents: 51360
diff changeset
  1468
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1469
subsection \<open>Limits of Sequences\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1470
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1471
lemma [trans]: "X=Y ==> Y \<longlonglongrightarrow> z ==> X \<longlonglongrightarrow> z"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1472
  by simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1473
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1474
lemma LIMSEQ_iff:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1475
  fixes L :: "'a::real_normed_vector"
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1476
  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
  1477
unfolding lim_sequentially dist_norm ..
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1478
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1479
lemma LIMSEQ_I:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1480
  fixes L :: "'a::real_normed_vector"
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1481
  shows "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X \<longlonglongrightarrow> L"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1482
by (simp add: LIMSEQ_iff)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1483
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1484
lemma LIMSEQ_D:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1485
  fixes L :: "'a::real_normed_vector"
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1486
  shows "\<lbrakk>X \<longlonglongrightarrow> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1487
by (simp add: LIMSEQ_iff)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1488
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1489
lemma LIMSEQ_linear: "\<lbrakk> X \<longlonglongrightarrow> x ; l > 0 \<rbrakk> \<Longrightarrow> (\<lambda> n. X (n * l)) \<longlonglongrightarrow> x"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1490
  unfolding tendsto_def eventually_sequentially
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57447
diff changeset
  1491
  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
  1492
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1493
lemma Bseq_inverse_lemma:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1494
  fixes x :: "'a::real_normed_div_algebra"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1495
  shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1496
apply (subst nonzero_norm_inverse, clarsimp)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1497
apply (erule (1) le_imp_inverse_le)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1498
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1499
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1500
lemma Bseq_inverse:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1501
  fixes a :: "'a::real_normed_div_algebra"
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1502
  shows "\<lbrakk>X \<longlonglongrightarrow> a; a \<noteq> 0\<rbrakk> \<Longrightarrow> Bseq (\<lambda>n. inverse (X n))"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1503
  by (rule Bfun_inverse)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1504
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1505
text\<open>Transformation of limit.\<close>
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1506
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1507
lemma eventually_at2:
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1508
  "eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1509
  unfolding eventually_at by auto
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1510
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1511
lemma Lim_transform:
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1512
  fixes a b :: "'a::real_normed_vector"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1513
  shows "\<lbrakk>(g \<longlongrightarrow> a) F; ((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F\<rbrakk> \<Longrightarrow> (f \<longlongrightarrow> a) F"
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1514
  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
  1515
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1516
lemma Lim_transform2:
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1517
  fixes a b :: "'a::real_normed_vector"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1518
  shows "\<lbrakk>(f \<longlongrightarrow> a) F; ((\<lambda>x. f x - g x) \<longlongrightarrow> 0) F\<rbrakk> \<Longrightarrow> (g \<longlongrightarrow> a) F"
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1519
  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
  1520
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1521
lemma Lim_transform_eventually:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1522
  "eventually (\<lambda>x. f x = g x) net \<Longrightarrow> (f \<longlongrightarrow> l) net \<Longrightarrow> (g \<longlongrightarrow> l) net"
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1523
  apply (rule topological_tendstoI)
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1524
  apply (drule (2) topological_tendstoD)
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1525
  apply (erule (1) eventually_elim2, simp)
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1526
  done
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1527
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1528
lemma Lim_transform_within:
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1529
  assumes "(f \<longlongrightarrow> l) (at x within S)"
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1530
    and "0 < d"
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1531
    and "\<And>x'. \<lbrakk>x'\<in>S; 0 < dist x' x; dist x' x < d\<rbrakk> \<Longrightarrow> f x' = g x'"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1532
  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
  1533
proof (rule Lim_transform_eventually)
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1534
  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
  1535
    using assms by (auto simp: eventually_at)
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1536
  show "(f \<longlongrightarrow> l) (at x within S)" by fact
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1537
qed
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1538
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1539
text\<open>Common case assuming being away from some crucial point like 0.\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1540
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1541
lemma Lim_transform_away_within:
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1542
  fixes a b :: "'a::t1_space"
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1543
  assumes "a \<noteq> b"
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1544
    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
  1545
    and "(f \<longlongrightarrow> l) (at a within S)"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1546
  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
  1547
proof (rule Lim_transform_eventually)
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1548
  show "(f \<longlongrightarrow> l) (at a within S)" by fact
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1549
  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
  1550
    unfolding eventually_at_topological
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1551
    by (rule exI [where x="- {b}"], simp add: open_Compl assms)
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1552
qed
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1553
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1554
lemma Lim_transform_away_at:
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1555
  fixes a b :: "'a::t1_space"
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1556
  assumes ab: "a\<noteq>b"
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1557
    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
  1558
    and fl: "(f \<longlongrightarrow> l) (at a)"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1559
  shows "(g \<longlongrightarrow> l) (at a)"
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1560
  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
  1561
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1562
text\<open>Alternatively, within an open set.\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1563
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1564
lemma Lim_transform_within_open:
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1565
  assumes "(f \<longlongrightarrow> l) (at a within T)"
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1566
    and "open s" and "a \<in> s"
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1567
    and "\<And>x. \<lbrakk>x\<in>s; x \<noteq> a\<rbrakk> \<Longrightarrow> f x = g x"
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1568
  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
  1569
proof (rule Lim_transform_eventually)
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1570
  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
  1571
    unfolding eventually_at_topological
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1572
    using assms by auto
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1573
  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
  1574
qed
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1575
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1576
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
  1577
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1578
(* FIXME: Only one congruence rule for tendsto can be used at a time! *)
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1579
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1580
lemma Lim_cong_within(*[cong add]*):
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1581
  assumes "a = b"
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1582
    and "x = y"
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1583
    and "S = T"
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1584
    and "\<And>x. x \<noteq> b \<Longrightarrow> x \<in> T \<Longrightarrow> f x = g x"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1585
  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
  1586
  unfolding tendsto_def eventually_at_topological
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1587
  using assms by simp
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1588
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1589
lemma Lim_cong_at(*[cong add]*):
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1590
  assumes "a = b" "x = y"
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1591
    and "\<And>x. x \<noteq> a \<Longrightarrow> f x = g x"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1592
  shows "((\<lambda>x. f x) \<longlongrightarrow> x) (at a) \<longleftrightarrow> ((g \<longlongrightarrow> y) (at a))"
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1593
  unfolding tendsto_def eventually_at_topological
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1594
  using assms by simp
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1595
text\<open>An unbounded sequence's inverse tends to 0\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1596
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1597
lemma LIMSEQ_inverse_zero:
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1598
  "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) \<longlonglongrightarrow> 0"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1599
  apply (rule filterlim_compose[OF tendsto_inverse_0])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1600
  apply (simp add: filterlim_at_infinity[OF order_refl] eventually_sequentially)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1601
  apply (metis abs_le_D1 linorder_le_cases linorder_not_le)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1602
  done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1603
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1604
text\<open>The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1605
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1606
lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) \<longlonglongrightarrow> 0"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1607
  by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1608
            filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1609
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1610
text\<open>The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1611
infinity is now easily proved\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1612
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1613
lemma LIMSEQ_inverse_real_of_nat_add:
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1614
     "(%n. r + inverse(real(Suc n))) \<longlonglongrightarrow> r"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1615
  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
  1616
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1617
lemma LIMSEQ_inverse_real_of_nat_add_minus:
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1618
     "(%n. r + -inverse(real(Suc n))) \<longlonglongrightarrow> r"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1619
  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
  1620
  by auto
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1621
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1622
lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1623
     "(%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
  1624
  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
  1625
  by auto
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1626
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1627
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
  1628
  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
  1629
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1630
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
  1631
proof (rule Lim_transform_eventually)
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61169
diff changeset
  1632
  show "eventually (\<lambda>n. 1 + inverse (of_nat n :: 'a) = of_nat (Suc n) / of_nat n) sequentially"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61169
diff changeset
  1633
    using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: field_simps)
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1634
  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
  1635
    by (intro tendsto_add tendsto_const lim_inverse_n)
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1636
  thus "(\<lambda>n. 1 + inverse (of_nat n) :: 'a) \<longlonglongrightarrow> 1" by simp
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61169
diff changeset
  1637
qed
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61169
diff changeset
  1638
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1639
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
  1640
proof (rule Lim_transform_eventually)
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1641
  show "eventually (\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a) =
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61169
diff changeset
  1642
                        of_nat n / of_nat (Suc n)) sequentially"
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1643
    using eventually_gt_at_top[of "0::nat"]
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61169
diff changeset
  1644
    by eventually_elim (simp add: field_simps del: of_nat_Suc)
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1645
  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
  1646
    by (intro tendsto_inverse LIMSEQ_Suc_n_over_n) simp_all
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1647
  thus "(\<lambda>n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \<longlonglongrightarrow> 1" by simp
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61169
diff changeset
  1648
qed
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61169
diff changeset
  1649
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1650
subsection \<open>Convergence on sequences\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1651
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1652
lemma convergent_cong:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1653
  assumes "eventually (\<lambda>x. f x = g x) sequentially"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1654
  shows   "convergent f \<longleftrightarrow> convergent g"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1655
  unfolding convergent_def by (subst filterlim_cong[OF refl refl assms]) (rule refl)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1656
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1657
lemma convergent_Suc_iff: "convergent (\<lambda>n. f (Suc n)) \<longleftrightarrow> convergent f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1658
  by (auto simp: convergent_def LIMSEQ_Suc_iff)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1659
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1660
lemma convergent_ignore_initial_segment: "convergent (\<lambda>n. f (n + m)) = convergent f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1661
proof (induction m arbitrary: f)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1662
  case (Suc m)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1663
  have "convergent (\<lambda>n. f (n + Suc m)) \<longleftrightarrow> convergent (\<lambda>n. f (Suc n + m))" by simp
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1664
  also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. f (n + m))" by (rule convergent_Suc_iff)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1665
  also have "\<dots> \<longleftrightarrow> convergent f" by (rule Suc)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1666
  finally show ?case .
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1667
qed simp_all
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1668
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1669
lemma convergent_add:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1670
  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1671
  assumes "convergent (\<lambda>n. X n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1672
  assumes "convergent (\<lambda>n. Y n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1673
  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
  1674
  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
  1675
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1676
lemma convergent_setsum:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1677
  fixes X :: "'a \<Rightarrow> nat \<Rightarrow> 'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1678
  assumes "\<And>i. i \<in> A \<Longrightarrow> convergent (\<lambda>n. X i n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1679
  shows "convergent (\<lambda>n. \<Sum>i\<in>A. X i n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1680
proof (cases "finite A")
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1681
  case True from this and assms show ?thesis
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1682
    by (induct A set: finite) (simp_all add: convergent_const convergent_add)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1683
qed (simp add: convergent_const)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1684
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1685
lemma (in bounded_linear) convergent:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1686
  assumes "convergent (\<lambda>n. X n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1687
  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
  1688
  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
  1689
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1690
lemma (in bounded_bilinear) convergent:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1691
  assumes "convergent (\<lambda>n. X n)" and "convergent (\<lambda>n. Y n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1692
  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
  1693
  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
  1694
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1695
lemma convergent_minus_iff:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1696
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1697
  shows "convergent X \<longleftrightarrow> convergent (\<lambda>n. - X n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1698
apply (simp add: convergent_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1699
apply (auto dest: tendsto_minus)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1700
apply (drule tendsto_minus, auto)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1701
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1702
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1703
lemma convergent_diff:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1704
  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1705
  assumes "convergent (\<lambda>n. X n)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1706
  assumes "convergent (\<lambda>n. Y n)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1707
  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
  1708
  using assms unfolding convergent_def by (blast intro: tendsto_diff)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1709
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1710
lemma convergent_norm:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1711
  assumes "convergent f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1712
  shows   "convergent (\<lambda>n. norm (f n))"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1713
proof -
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1714
  from assms have "f \<longlonglongrightarrow> lim f" by (simp add: convergent_LIMSEQ_iff)
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1715
  hence "(\<lambda>n. norm (f n)) \<longlonglongrightarrow> norm (lim f)" by (rule tendsto_norm)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1716
  thus ?thesis by (auto simp: convergent_def)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1717
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1718
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1719
lemma convergent_of_real:
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1720
  "convergent f \<Longrightarrow> convergent (\<lambda>n. of_real (f n) :: 'a :: real_normed_algebra_1)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1721
  unfolding convergent_def by (blast intro!: tendsto_of_real)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1722
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1723
lemma convergent_add_const_iff:
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1724
  "convergent (\<lambda>n. c + f n :: 'a :: real_normed_vector) \<longleftrightarrow> convergent f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1725
proof
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1726
  assume "convergent (\<lambda>n. c + f n)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1727
  from convergent_diff[OF this convergent_const[of c]] show "convergent f" by simp
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1728
next
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1729
  assume "convergent f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1730
  from convergent_add[OF convergent_const[of c] this] show "convergent (\<lambda>n. c + f n)" by simp
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1731
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1732
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1733
lemma convergent_add_const_right_iff:
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1734
  "convergent (\<lambda>n. f n + c :: 'a :: real_normed_vector) \<longleftrightarrow> convergent f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1735
  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
  1736
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1737
lemma convergent_diff_const_right_iff:
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1738
  "convergent (\<lambda>n. f n - c :: 'a :: real_normed_vector) \<longleftrightarrow> convergent f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1739
  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
  1740
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1741
lemma convergent_mult:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1742
  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1743
  assumes "convergent (\<lambda>n. X n)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1744
  assumes "convergent (\<lambda>n. Y n)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1745
  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
  1746
  using assms unfolding convergent_def by (blast intro: tendsto_mult)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1747
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1748
lemma convergent_mult_const_iff:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1749
  assumes "c \<noteq> 0"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1750
  shows   "convergent (\<lambda>n. c * f n :: 'a :: real_normed_field) \<longleftrightarrow> convergent f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1751
proof
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1752
  assume "convergent (\<lambda>n. c * f n)"
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61976
diff changeset
  1753
  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
  1754
    show "convergent f" by (simp add: field_simps)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1755
next
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1756
  assume "convergent f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1757
  from convergent_mult[OF convergent_const[of c] this] show "convergent (\<lambda>n. c * f n)" by simp
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1758
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1759
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1760
lemma convergent_mult_const_right_iff:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1761
  assumes "c \<noteq> 0"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1762
  shows   "convergent (\<lambda>n. (f n :: 'a :: real_normed_field) * c) \<longleftrightarrow> convergent f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1763
  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
  1764
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1765
lemma convergent_imp_Bseq: "convergent f \<Longrightarrow> Bseq f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1766
  by (simp add: Cauchy_Bseq convergent_Cauchy)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1767
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1768
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1769
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
  1770
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  1771
lemma LIMSEQ_incseq_SUP:
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  1772
  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
  1773
  assumes u: "bdd_above (range X)"
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  1774
  assumes X: "incseq X"
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1775
  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
  1776
  by (rule order_tendstoI)
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  1777
     (auto simp: eventually_sequentially u less_cSUP_iff 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
  1778
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  1779
lemma LIMSEQ_decseq_INF:
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  1780
  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
  1781
  assumes u: "bdd_below (range X)"
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  1782
  assumes X: "decseq X"
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1783
  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
  1784
  by (rule order_tendstoI)
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  1785
     (auto simp: eventually_sequentially u cINF_less_iff intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u])
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1786
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1787
text\<open>Main monotonicity theorem\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1788
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1789
lemma Bseq_monoseq_convergent: "Bseq X \<Longrightarrow> monoseq X \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  1790
  by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP dest: Bseq_bdd_above Bseq_bdd_below)
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  1791
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  1792
lemma Bseq_mono_convergent: "Bseq X \<Longrightarrow> (\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n) \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  1793
  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
  1794
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1795
lemma monoseq_imp_convergent_iff_Bseq: "monoseq (f :: nat \<Rightarrow> real) \<Longrightarrow> convergent f \<longleftrightarrow> Bseq f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1796
  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
  1797
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1798
lemma Bseq_monoseq_convergent'_inc:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1799
  "Bseq (\<lambda>n. f (n + M) :: real) \<Longrightarrow> (\<And>m n. M \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m \<le> f n) \<Longrightarrow> convergent f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1800
  by (subst convergent_ignore_initial_segment [symmetric, of _ M])
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1801
     (auto intro!: Bseq_monoseq_convergent simp: monoseq_def)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1802
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1803
lemma Bseq_monoseq_convergent'_dec:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1804
  "Bseq (\<lambda>n. f (n + M) :: real) \<Longrightarrow> (\<And>m n. M \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m \<ge> f n) \<Longrightarrow> convergent f"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1805
  by (subst convergent_ignore_initial_segment [symmetric, of _ M])
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1806
     (auto intro!: Bseq_monoseq_convergent simp: monoseq_def)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61524
diff changeset
  1807
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1808
lemma Cauchy_iff:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1809
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1810
  shows "Cauchy X \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1811
  unfolding Cauchy_def dist_norm ..
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1812
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1813
lemma CauchyI:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1814
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1815
  shows "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1816
by (simp add: Cauchy_iff)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1817
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1818
lemma CauchyD:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1819
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1820
  shows "\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1821
by (simp add: Cauchy_iff)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1822
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1823
lemma incseq_convergent:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1824
  fixes X :: "nat \<Rightarrow> real"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1825
  assumes "incseq X" and "\<forall>i. X i \<le> B"
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1826
  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
  1827
proof atomize_elim
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1828
  from incseq_bounded[OF assms] \<open>incseq X\<close> Bseq_monoseq_convergent[of X]
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1829
  obtain L where "X \<longlonglongrightarrow> L"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1830
    by (auto simp: convergent_def monoseq_def incseq_def)
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1831
  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
  1832
    by (auto intro!: exI[of _ L] incseq_le)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1833
qed
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1834
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1835
lemma decseq_convergent:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1836
  fixes X :: "nat \<Rightarrow> real"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1837
  assumes "decseq X" and "\<forall>i. B \<le> X i"
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1838
  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
  1839
proof atomize_elim
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1840
  from decseq_bounded[OF assms] \<open>decseq X\<close> Bseq_monoseq_convergent[of X]
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1841
  obtain L where "X \<longlonglongrightarrow> L"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1842
    by (auto simp: convergent_def monoseq_def decseq_def)
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1843
  with \<open>decseq X\<close> show "\<exists>L. X \<longlonglongrightarrow> L \<and> (\<forall>i. L \<le> X i)"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1844
    by (auto intro!: exI[of _ L] decseq_le)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1845
qed
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1846
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1847
subsubsection \<open>Cauchy Sequences are Bounded\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1848
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1849
text\<open>A Cauchy sequence is bounded -- this is the standard
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1850
  proof mechanization rather than the nonstandard proof\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1851
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1852
lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1853
          ==>  \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1854
apply (clarify, drule spec, drule (1) mp)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1855
apply (simp only: norm_minus_commute)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1856
apply (drule order_le_less_trans [OF norm_triangle_ineq2])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1857
apply simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1858
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1859
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1860
subsection \<open>Power Sequences\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1861
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1862
text\<open>The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1863
"x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1864
  also fact that bounded and monotonic sequence converges.\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1865
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1866
lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1867
apply (simp add: Bseq_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1868
apply (rule_tac x = 1 in exI)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1869
apply (simp add: power_abs)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1870
apply (auto dest: power_mono)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1871
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1872
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1873
lemma monoseq_realpow: fixes x :: real shows "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1874
apply (clarify intro!: mono_SucI2)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1875
apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1876
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1877
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1878
lemma convergent_realpow:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1879
  "[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1880
by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1881
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1882
lemma LIMSEQ_inverse_realpow_zero: "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) \<longlonglongrightarrow> 0"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1883
  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
  1884
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1885
lemma LIMSEQ_realpow_zero:
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1886
  "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) \<longlonglongrightarrow> 0"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1887
proof cases
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1888
  assume "0 \<le> x" and "x \<noteq> 0"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1889
  hence x0: "0 < x" by simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1890
  assume x1: "x < 1"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1891
  from x0 x1 have "1 < inverse x"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1892
    by (rule one_less_inverse)
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1893
  hence "(\<lambda>n. inverse (inverse x ^ n)) \<longlonglongrightarrow> 0"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1894
    by (rule LIMSEQ_inverse_realpow_zero)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1895
  thus ?thesis by (simp add: power_inverse)
58729
e8ecc79aee43 add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents: 57512
diff changeset
  1896
qed (rule LIMSEQ_imp_Suc, simp)
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1897
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1898
lemma LIMSEQ_power_zero:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1899
  fixes x :: "'a::{real_normed_algebra_1}"
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1900
  shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) \<longlonglongrightarrow> 0"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1901
apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1902
apply (simp only: tendsto_Zfun_iff, erule Zfun_le)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1903
apply (simp add: power_abs norm_power_ineq)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1904
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1905
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1906
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
  1907
  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
  1908
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1909
text\<open>Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1910
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1911
lemma LIMSEQ_rabs_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
  1912
  by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1913
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  1914
lemma LIMSEQ_rabs_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
  1915
  by (rule LIMSEQ_power_zero) simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1916
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1917
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  1918
subsection \<open>Limits of Functions\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1919
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1920
lemma LIM_eq:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1921
  fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  1922
  shows "f \<midarrow>a\<rightarrow> L =
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1923
     (\<forall>r>0.\<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1924
by (simp add: LIM_def dist_norm)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1925
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1926
lemma LIM_I:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1927
  fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1928
  shows "(!!r. 0<r ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  1929
      ==> f \<midarrow>a\<rightarrow> L"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1930
by (simp add: LIM_eq)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1931
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1932
lemma LIM_D:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1933
  fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  1934
  shows "[| f \<midarrow>a\<rightarrow> L; 0<r |]
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1935
      ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1936
by (simp add: LIM_eq)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1937
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1938
lemma LIM_offset:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1939
  fixes a :: "'a::real_normed_vector"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  1940
  shows "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. f (x + k)) \<midarrow>(a - k)\<rightarrow> L"
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
  1941
  unfolding filtermap_at_shift[symmetric, of a k] filterlim_def filtermap_filtermap by simp
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1942
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1943
lemma LIM_offset_zero:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1944
  fixes a :: "'a::real_normed_vector"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  1945
  shows "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57447
diff changeset
  1946
by (drule_tac k="a" in LIM_offset, simp add: add.commute)
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1947
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1948
lemma LIM_offset_zero_cancel:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1949
  fixes a :: "'a::real_normed_vector"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  1950
  shows "(\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L \<Longrightarrow> f \<midarrow>a\<rightarrow> L"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1951
by (drule_tac k="- a" in LIM_offset, simp)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1952
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
  1953
lemma LIM_offset_zero_iff:
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
  1954
  fixes f :: "'a :: real_normed_vector \<Rightarrow> _"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  1955
  shows  "f \<midarrow>a\<rightarrow> L \<longleftrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> L"
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
  1956
  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
  1957
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1958
lemma LIM_zero:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1959
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1960
  shows "(f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. f x - l) \<longlongrightarrow> 0) F"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1961
unfolding tendsto_iff dist_norm by simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1962
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1963
lemma LIM_zero_cancel:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1964
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1965
  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
  1966
unfolding tendsto_iff dist_norm by simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1967
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1968
lemma LIM_zero_iff:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1969
  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1970
  shows "((\<lambda>x. f x - l) \<longlongrightarrow> 0) F = (f \<longlongrightarrow> l) F"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1971
unfolding tendsto_iff dist_norm by simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1972
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1973
lemma LIM_imp_LIM:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1974
  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
  1975
  fixes g :: "'a::topological_space \<Rightarrow> 'c::real_normed_vector"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  1976
  assumes f: "f \<midarrow>a\<rightarrow> l"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1977
  assumes 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
  1978
  shows "g \<midarrow>a\<rightarrow> m"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1979
  by (rule metric_LIM_imp_LIM [OF f],
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1980
    simp add: dist_norm le)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1981
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1982
lemma LIM_equal2:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1983
  fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1984
  assumes 1: "0 < R"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1985
  assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < R\<rbrakk> \<Longrightarrow> f x = g x"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  1986
  shows "g \<midarrow>a\<rightarrow> l \<Longrightarrow> f \<midarrow>a\<rightarrow> l"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1987
by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1988
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1989
lemma LIM_compose2:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1990
  fixes a :: "'a::real_normed_vector"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  1991
  assumes f: "f \<midarrow>a\<rightarrow> b"
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  1992
  assumes g: "g \<midarrow>b\<rightarrow> c"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1993
  assumes 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
  1994
  shows "(\<lambda>x. g (f x)) \<midarrow>a\<rightarrow> c"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1995
by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1996
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1997
lemma real_LIM_sandwich_zero:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1998
  fixes f g :: "'a::topological_space \<Rightarrow> real"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  1999
  assumes f: "f \<midarrow>a\<rightarrow> 0"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2000
  assumes 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2001
  assumes 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  2002
  shows "g \<midarrow>a\<rightarrow> 0"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2003
proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2004
  fix x assume x: "x \<noteq> a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2005
  have "norm (g x - 0) = g x" by (simp add: 1 x)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2006
  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
  2007
  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
  2008
  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
  2009
  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
  2010
qed
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2011
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2012
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2013
subsection \<open>Continuity\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2014
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2015
lemma LIM_isCont_iff:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2016
  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  2017
  shows "(f \<midarrow>a\<rightarrow> f a) = ((\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow> f a)"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2018
by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2019
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2020
lemma isCont_iff:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2021
  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  2022
  shows "isCont f x = (\<lambda>h. f (x + h)) \<midarrow>0\<rightarrow> f x"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2023
by (simp add: isCont_def LIM_isCont_iff)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2024
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2025
lemma isCont_LIM_compose2:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2026
  fixes a :: "'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2027
  assumes f [unfolded isCont_def]: "isCont f a"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  2028
  assumes g: "g \<midarrow>f a\<rightarrow> l"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2029
  assumes 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
  2030
  shows "(\<lambda>x. g (f x)) \<midarrow>a\<rightarrow> l"
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2031
by (rule LIM_compose2 [OF f g inj])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2032
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2033
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2034
lemma isCont_norm [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2035
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2036
  shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2037
  by (fact continuous_norm)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2038
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2039
lemma isCont_rabs [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2040
  fixes f :: "'a::t2_space \<Rightarrow> real"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2041
  shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2042
  by (fact continuous_rabs)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2043
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2044
lemma isCont_add [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2045
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2046
  shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x + g x) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2047
  by (fact continuous_add)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2048
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2049
lemma isCont_minus [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2050
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2051
  shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2052
  by (fact continuous_minus)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2053
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2054
lemma isCont_diff [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2055
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2056
  shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x - g x) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2057
  by (fact continuous_diff)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2058
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2059
lemma isCont_mult [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2060
  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2061
  shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x * g x) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2062
  by (fact continuous_mult)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2063
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2064
lemma (in bounded_linear) isCont:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2065
  "isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2066
  by (fact continuous)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2067
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2068
lemma (in bounded_bilinear) isCont:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2069
  "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2070
  by (fact continuous)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2071
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2072
lemmas isCont_scaleR [simp] =
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2073
  bounded_bilinear.isCont [OF bounded_bilinear_scaleR]
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2074
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2075
lemmas isCont_of_real [simp] =
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2076
  bounded_linear.isCont [OF bounded_linear_of_real]
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2077
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2078
lemma isCont_power [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2079
  fixes f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2080
  shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2081
  by (fact continuous_power)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2082
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2083
lemma isCont_setsum [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2084
  fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2085
  shows "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2086
  by (auto intro: continuous_setsum)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2087
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2088
subsection \<open>Uniform Continuity\<close>
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2089
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2090
definition
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2091
  isUCont :: "['a::metric_space \<Rightarrow> 'b::metric_space] \<Rightarrow> bool" where
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2092
  "isUCont f = (\<forall>r>0. \<exists>s>0. \<forall>x y. dist x y < s \<longrightarrow> dist (f x) (f y) < r)"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2093
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2094
lemma isUCont_isCont: "isUCont f ==> isCont f x"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2095
by (simp add: isUCont_def isCont_def LIM_def, force)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2096
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2097
lemma isUCont_Cauchy:
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2098
  "\<lbrakk>isUCont f; Cauchy X\<rbrakk> \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2099
unfolding isUCont_def
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2100
apply (rule metric_CauchyI)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2101
apply (drule_tac x=e in spec, safe)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2102
apply (drule_tac e=s in metric_CauchyD, safe)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2103
apply (rule_tac x=M in exI, simp)
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2104
done
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2105
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2106
lemma (in bounded_linear) isUCont: "isUCont f"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2107
unfolding isUCont_def dist_norm
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2108
proof (intro allI impI)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2109
  fix r::real assume r: "0 < r"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2110
  obtain K where K: "0 < K" and norm_le: "\<And>x. norm (f x) \<le> norm x * 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
  2111
    using pos_bounded by blast
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2112
  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
  2113
  proof (rule exI, safe)
56541
0e3abadbef39 made divide_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  2114
    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
  2115
  next
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2116
    fix x y :: 'a
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2117
    assume xy: "norm (x - y) < r / K"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2118
    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
  2119
    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
  2120
    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
  2121
    finally show "norm (f x - f y) < r" .
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2122
  qed
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2123
qed
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2124
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2125
lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2126
by (rule isUCont [THEN isUCont_Cauchy])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2127
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2128
lemma LIM_less_bound:
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2129
  fixes f :: "real \<Rightarrow> real"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2130
  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
  2131
  shows "0 \<le> f x"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  2132
proof (rule tendsto_le_const)
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2133
  show "(f \<longlongrightarrow> f x) (at_left x)"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2134
    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
  2135
  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
  2136
    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
  2137
qed simp
51471
cad22a3cc09c move topological_space to its own theory
hoelzl
parents: 51360
diff changeset
  2138
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
  2139
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2140
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
  2141
2d2f59e6055a move theorems about 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
  2142
lemma nested_sequence_unique:
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  2143
  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
  2144
  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
  2145
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
  2146
  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
  2147
  have "decseq g" unfolding decseq_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
  2148
2d2f59e6055a move theorems about 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
  2149
  { fix n
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2150
    from \<open>decseq g\<close> have "g n \<le> g 0" by (rule decseqD) simp
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2151
    with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] have "f n \<le> g 0" by auto }
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  2152
  then obtain u where "f \<longlonglongrightarrow> u" "\<forall>i. f i \<le> u"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2153
    using incseq_convergent[OF \<open>incseq f\<close>] 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
  2154
  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
  2155
  { fix n
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2156
    from \<open>incseq f\<close> have "f 0 \<le> f n" by (rule incseqD) simp
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2157
    with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] have "f 0 \<le> g n" by simp }
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  2158
  then obtain l where "g \<longlonglongrightarrow> l" "\<forall>i. l \<le> g i"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2159
    using decseq_convergent[OF \<open>decseq g\<close>] by auto
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  2160
  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
  2161
  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
  2162
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
  2163
2d2f59e6055a move theorems about 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
  2164
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
  2165
  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
  2166
  assumes [arith]: "a \<le> 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
  2167
  assumes trans: "\<And>a b c. \<lbrakk>P a b; P b c; a \<le> b; b \<le> c\<rbrakk> \<Longrightarrow> P a 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
  2168
  assumes 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"
2d2f59e6055a move theorems about 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
  2169
  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
  2170
proof -
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 54863
diff changeset
  2171
  def bisect \<equiv> "rec_nat (a, b) (\<lambda>n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))"
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
  2172
  def l \<equiv> "\<lambda>n. fst (bisect n)" and u \<equiv> "\<lambda>n. snd (bisect 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
  2173
  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
  2174
    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
  2175
    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
  2176
2d2f59e6055a move theorems about 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
  2177
  { fix n have "l n \<le> u n" by (induct n) auto } note this[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
  2178
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  2179
  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
  2180
  proof (safe intro!: nested_sequence_unique)
2d2f59e6055a move theorems about 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
  2181
    fix n show "l n \<le> l (Suc n)" "u (Suc n) \<le> u n" by (induct n) 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
  2182
  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
  2183
    { fix n have "l n - u n = (a - b) / 2^n" by (induct n) (auto simp: field_simps) }
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  2184
    then show "(\<lambda>n. l n - u n) \<longlonglongrightarrow> 0" 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
  2185
  qed fact
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  2186
  then obtain x where x: "\<And>n. l n \<le> x" "\<And>n. x \<le> u n" and "l \<longlonglongrightarrow> x" "u \<longlonglongrightarrow> 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
  2187
  obtain d where "0 < d" and d: "\<And>a b. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> b - a < d \<Longrightarrow> P a b"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2188
    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
  2189
2d2f59e6055a move theorems about 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
  2190
  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
  2191
  proof (rule ccontr)
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2192
    assume "\<not> 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
  2193
    { fix n have "\<not> P (l n) (u 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
  2194
      proof (induct 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
  2195
        case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2196
      qed (simp add: \<open>\<not> P a b\<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
  2197
    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
  2198
    { have "eventually (\<lambda>n. x - d / 2 < l n) sequentially"
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  2199
        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
  2200
      moreover have "eventually (\<lambda>n. u n < x + d / 2) sequentially"
61969
e01015e49041 more symbols;
wenzelm
parents: 61916
diff changeset
  2201
        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
  2202
      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
  2203
      proof eventually_elim
2d2f59e6055a move theorems about 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
  2204
        fix n assume "x - d / 2 < l n" "u n < x + d / 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
  2205
        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
  2206
        with x show "P (l n) (u n)" by (rule 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
  2207
      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
  2208
    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
  2209
  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
  2210
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
  2211
2d2f59e6055a move theorems about 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
  2212
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
  2213
proof (cases "a \<le> b", rule compactI)
2d2f59e6055a move theorems about 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
  2214
  fix C assume C: "a \<le> b" "\<forall>t\<in>C. open t" "{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
  2215
  def T == "{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
  2216
  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
  2217
  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
  2218
    case (trans a b 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
  2219
    then have *: "{a .. c} = {a .. b} \<union> {b .. c}" 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
  2220
    from trans obtain C1 C2 where "C1\<subseteq>C \<and> finite C1 \<and> {a..b} \<subseteq> \<Union>C1" "C2\<subseteq>C \<and> finite C2 \<and> {b..c} \<subseteq> \<Union>C2"
2d2f59e6055a move theorems about 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
  2221
      by (auto 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
  2222
    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
  2223
      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
  2224
  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
  2225
    case (local 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
  2226
    then have "x \<in> \<Union>C" using C 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
  2227
    with C(2) obtain c where "x \<in> c" "open c" "c \<in> C" 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
  2228
    then obtain e where "0 < e" "{x - e <..< x + e} \<subseteq> c"
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 62087
diff changeset
  2229
      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
  2230
    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
  2231
      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
  2232
  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
  2233
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
  2234
2d2f59e6055a move theorems about 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
  2235
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57276
diff changeset
  2236
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
  2237
  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
  2238
  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
  2239
  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
  2240
  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
  2241
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57276
diff changeset
  2242
  have S: "compact S" "S \<noteq> {}"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2243
    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
  2244
  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
  2245
    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
  2246
  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
  2247
    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
  2248
  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
  2249
    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
  2250
  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
  2251
    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
  2252
  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
  2253
    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
  2254
qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57276
diff changeset
  2255
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2256
lemma open_Collect_positive:
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2257
 fixes f :: "'a::t2_space \<Rightarrow> real"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2258
 assumes f: "continuous_on s f"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2259
 shows "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. 0 < f x}"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2260
 using continuous_on_open_invariant[THEN iffD1, OF f, rule_format, of "{0 <..}"]
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2261
 by (auto simp: Int_def field_simps)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2262
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2263
lemma open_Collect_less_Int:
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2264
 fixes f g :: "'a::t2_space \<Rightarrow> real"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2265
 assumes f: "continuous_on s f" and g: "continuous_on s g"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2266
 shows "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. f x < g x}"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2267
 using open_Collect_positive[OF continuous_on_diff[OF g f]] by (simp add: field_simps)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2268
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60758
diff changeset
  2269
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2270
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
  2271
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2272
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
  2273
2d2f59e6055a move theorems about 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
  2274
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
  2275
  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
  2276
  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
  2277
    \<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)"
2d2f59e6055a move theorems about 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
  2278
  using continuous_attains_sup[of "{a .. b}" 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
  2279
  by (auto simp add: continuous_at_imp_continuous_on Ball_def Bex_def)
2d2f59e6055a move theorems about 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
  2280
2d2f59e6055a move theorems about 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
  2281
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
  2282
  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
  2283
  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
  2284
    \<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)"
2d2f59e6055a move theorems about 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
  2285
  using continuous_attains_inf[of "{a .. b}" 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
  2286
  by (auto simp add: continuous_at_imp_continuous_on Ball_def Bex_def)
2d2f59e6055a move theorems about 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
  2287
2d2f59e6055a move theorems about 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
  2288
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
  2289
  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
  2290
  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
  2291
  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
  2292
2d2f59e6055a move theorems about 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
  2293
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
  2294
  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
  2295
  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
  2296
    \<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
  2297
  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
  2298
2d2f59e6055a move theorems about 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
  2299
(*HOL style here: object-level formulations*)
2d2f59e6055a move theorems about 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
  2300
lemma IVT_objl: "(f(a::real) \<le> (y::real) & y \<le> f(b) & a \<le> 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
  2301
      (\<forall>x. a \<le> x & x \<le> b --> isCont 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
  2302
      --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
2d2f59e6055a move theorems about 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
  2303
  by (blast intro: IVT)
2d2f59e6055a move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents: 51526
diff changeset
  2304
2d2f59e6055a move theorems about 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
  2305
lemma IVT2_objl: "(f(b::real) \<le> (y::real) & y \<le> f(a) & a \<le> 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
  2306
      (\<forall>x. a \<le> x & x \<le> b --> isCont 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
  2307
      --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
2d2f59e6055a move theorems about 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
  2308
  by (blast intro: IVT2)
2d2f59e6055a move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents: 51526
diff changeset
  2309
2d2f59e6055a move theorems about 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
  2310
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
  2311
  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
  2312
  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
  2313
  shows "\<exists>L M. (\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> L \<le> f x \<and> f x \<le> M) \<and>
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
  2314
               (\<forall>y. L \<le> y \<and> y \<le> M \<longrightarrow> (\<exists>x. a \<le> x \<and> x \<le> b \<and> (f x = y)))"
2d2f59e6055a move theorems about 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
  2315
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
  2316
  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
  2317
    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
  2318
  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
  2319
    using isCont_eq_Lb[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
  2320
  show ?thesis
2d2f59e6055a move theorems about 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
  2321
    using IVT[of f L _ M] IVT2[of f L _ M] M L assms
2d2f59e6055a move theorems about 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
  2322
    apply (rule_tac x="f L" in exI)
2d2f59e6055a move theorems about 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
  2323
    apply (rule_tac x="f M" in exI)
2d2f59e6055a move theorems about 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
  2324
    apply (cases "L \<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
  2325
    apply (simp, metis order_trans)
2d2f59e6055a move theorems about 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
  2326
    apply (simp, metis order_trans)
2d2f59e6055a move theorems about 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
  2327
    done
2d2f59e6055a move theorems about 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
  2328
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
  2329
2d2f59e6055a move theorems about 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
  2330
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2331
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
  2332
2d2f59e6055a move theorems about 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
  2333
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
  2334
  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
  2335
  assumes d: "0 < 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
  2336
      and inj: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> g (f z) = z"
2d2f59e6055a move theorems about 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
  2337
      and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d \<longrightarrow> isCont f z"
2d2f59e6055a move theorems about 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
  2338
  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
  2339
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
  2340
  let ?A = "f (x - d)" and ?B = "f (x + d)" and ?D = "{x - d..x + 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
  2341
2d2f59e6055a move theorems about 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
  2342
  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
  2343
    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
  2344
  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
  2345
    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
  2346
2d2f59e6055a move theorems about 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
  2347
  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
  2348
    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
  2349
  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
  2350
    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
  2351
  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
  2352
  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
  2353
    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
  2354
  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
  2355
    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
  2356
  ultimately
2d2f59e6055a move theorems about 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
  2357
  show ?thesis
2d2f59e6055a move theorems about 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
  2358
    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
  2359
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
  2360
2d2f59e6055a move theorems about 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
  2361
lemma isCont_inverse_function2:
2d2f59e6055a move theorems about 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
  2362
  fixes f g :: "real \<Rightarrow> real" shows
2d2f59e6055a move theorems about 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
  2363
  "\<lbrakk>a < x; x < 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
  2364
    \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> g (f z) = z;
2d2f59e6055a move theorems about 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
  2365
    \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> isCont f z\<rbrakk>
2d2f59e6055a move theorems about 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
  2366
   \<Longrightarrow> 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
  2367
apply (rule 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
  2368
       [where f=f and d="min (x - a) (b - 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
  2369
apply (simp_all add: abs_le_iff)
2d2f59e6055a move theorems about 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
  2370
done
2d2f59e6055a move theorems about 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
  2371
2d2f59e6055a move theorems about 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
  2372
(* need to rename second isCont_inverse *)
2d2f59e6055a move theorems about 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
  2373
2d2f59e6055a move theorems about 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
  2374
lemma isCont_inv_fun:
2d2f59e6055a move theorems about 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
  2375
  fixes f g :: "real \<Rightarrow> real"
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2376
  shows "[| 0 < d; \<forall>z. \<bar>z - x\<bar> \<le> d --> g(f(z)) = z;
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2377
         \<forall>z. \<bar>z - x\<bar> \<le> d --> 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
  2378
      ==> 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
  2379
by (rule 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
  2380
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60721
diff changeset
  2381
text\<open>Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110\<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
  2382
lemma LIM_fun_gt_zero:
2d2f59e6055a move theorems about 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
  2383
  fixes f :: "real \<Rightarrow> real"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  2384
  shows "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)"
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
  2385
apply (drule (1) LIM_D, clarify)
2d2f59e6055a move theorems about 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
  2386
apply (rule_tac x = s in exI)
2d2f59e6055a move theorems about 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
  2387
apply (simp add: abs_less_iff)
2d2f59e6055a move theorems about 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
  2388
done
2d2f59e6055a move theorems about 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
  2389
2d2f59e6055a move theorems about 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
  2390
lemma LIM_fun_less_zero:
2d2f59e6055a move theorems about 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
  2391
  fixes f :: "real \<Rightarrow> real"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  2392
  shows "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)"
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
  2393
apply (drule LIM_D [where r="-l"], simp, clarify)
2d2f59e6055a move theorems about 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
  2394
apply (rule_tac x = s in exI)
2d2f59e6055a move theorems about 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
  2395
apply (simp add: abs_less_iff)
2d2f59e6055a move theorems about 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
  2396
done
2d2f59e6055a move theorems about 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
  2397
2d2f59e6055a move theorems about 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
  2398
lemma LIM_fun_not_zero:
2d2f59e6055a move theorems about 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
  2399
  fixes f :: "real \<Rightarrow> real"
61976
3a27957ac658 more symbols;
wenzelm
parents: 61973
diff changeset
  2400
  shows "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)"
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
  2401
  using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp add: neq_iff)
51531
f415febf4234 remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents: 51529
diff changeset
  2402
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
  2403
end
50324
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  2404