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