src/HOL/Limits.thy
author hoelzl
Tue, 26 Mar 2013 12:20:58 +0100
changeset 51526 155263089e7b
parent 51524 7cb5ac44ca9e
child 51529 2d2f59e6055a
permissions -rw-r--r--
move SEQ.thy and Lim.thy to Limits.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
     1
(*  Title:      Limits.thy
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
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
     6
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
     7
*)
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
     8
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
     9
header {* Limits on Real Vector Spaces *}
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
    10
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
    11
theory Limits
51524
7cb5ac44ca9e rename RealVector.thy to Real_Vector_Spaces.thy
hoelzl
parents: 51478
diff changeset
    12
imports Real_Vector_Spaces
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
    13
begin
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
    14
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
    15
(* Unfortunately eventually_within was overwritten by Multivariate_Analysis.
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
    16
   Hence it was references as Limits.eventually_within, but now it is Basic_Topology.eventually_within *)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
    17
lemmas eventually_within = eventually_within
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
    18
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
    19
subsection {* Filter going to infinity norm *}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
    20
50324
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    21
definition at_infinity :: "'a::real_normed_vector filter" where
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    22
  "at_infinity = Abs_filter (\<lambda>P. \<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x)"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    23
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    24
lemma eventually_at_infinity:
50325
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    25
  "eventually P at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. b \<le> norm x \<longrightarrow> P x)"
50324
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    26
unfolding at_infinity_def
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    27
proof (rule eventually_Abs_filter, rule is_filter.intro)
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    28
  fix P Q :: "'a \<Rightarrow> bool"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    29
  assume "\<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x" and "\<exists>s. \<forall>x. s \<le> norm x \<longrightarrow> Q x"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    30
  then obtain r s where
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    31
    "\<forall>x. r \<le> norm x \<longrightarrow> P x" and "\<forall>x. s \<le> norm x \<longrightarrow> Q x" by auto
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    32
  then have "\<forall>x. max r s \<le> norm x \<longrightarrow> P x \<and> Q x" by simp
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    33
  then show "\<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x \<and> Q x" ..
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
    34
qed auto
31392
69570155ddf8 replace filters with filter bases
huffman
parents: 31357
diff changeset
    35
50325
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    36
lemma at_infinity_eq_at_top_bot:
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    37
  "(at_infinity \<Colon> real filter) = sup at_top at_bot"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    38
  unfolding sup_filter_def at_infinity_def eventually_at_top_linorder eventually_at_bot_linorder
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    39
proof (intro arg_cong[where f=Abs_filter] ext iffI)
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    40
  fix P :: "real \<Rightarrow> bool" assume "\<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    41
  then guess r ..
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    42
  then have "(\<forall>x\<ge>r. P x) \<and> (\<forall>x\<le>-r. P x)" by auto
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    43
  then show "(\<exists>r. \<forall>x\<ge>r. P x) \<and> (\<exists>r. \<forall>x\<le>r. P x)" by auto
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    44
next
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    45
  fix P :: "real \<Rightarrow> bool" assume "(\<exists>r. \<forall>x\<ge>r. P x) \<and> (\<exists>r. \<forall>x\<le>r. P x)"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    46
  then obtain p q where "\<forall>x\<ge>p. P x" "\<forall>x\<le>q. P x" by auto
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    47
  then show "\<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    48
    by (intro exI[of _ "max p (-q)"])
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    49
       (auto simp: abs_real_def)
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    50
qed
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    51
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    52
lemma at_top_le_at_infinity:
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    53
  "at_top \<le> (at_infinity :: real filter)"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    54
  unfolding at_infinity_eq_at_top_bot by simp
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    55
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    56
lemma at_bot_le_at_infinity:
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    57
  "at_bot \<le> (at_infinity :: real filter)"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    58
  unfolding at_infinity_eq_at_top_bot by simp
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
    59
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    60
subsection {* Boundedness *}
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    61
51474
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    62
lemma Bfun_def:
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    63
  "Bfun f F \<longleftrightarrow> (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) F)"
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    64
  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
    65
proof safe
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    66
  fix y K assume "0 < K" and *: "eventually (\<lambda>x. dist (f x) y \<le> K) F"
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    67
  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
    68
    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
    69
  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
    70
    by eventually_elim auto
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    71
  with `0 < K` show "\<exists>K>0. eventually (\<lambda>x. dist (f x) 0 \<le> K) F"
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    72
    by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto
1e9e68247ad1 generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents: 51472
diff changeset
    73
qed auto
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    74
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
    75
lemma BfunI:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
    76
  assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) F" shows "Bfun f F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    77
unfolding Bfun_def
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    78
proof (intro exI conjI allI)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    79
  show "0 < max K 1" by simp
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    80
next
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
    81
  show "eventually (\<lambda>x. norm (f x) \<le> max K 1) F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    82
    using K by (rule eventually_elim1, simp)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    83
qed
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    84
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    85
lemma BfunE:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
    86
  assumes "Bfun f F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
    87
  obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    88
using assms unfolding Bfun_def by fast
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
    89
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
    90
subsection {* Convergence to Zero *}
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
    91
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
    92
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
    93
  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
    94
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
    95
lemma ZfunI:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
    96
  "(\<And>r. 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F) \<Longrightarrow> Zfun f F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
    97
  unfolding Zfun_def by simp
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
    98
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
    99
lemma ZfunD:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   100
  "\<lbrakk>Zfun f F; 0 < r\<rbrakk> \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   101
  unfolding Zfun_def by simp
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   102
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   103
lemma Zfun_ssubst:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   104
  "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
   105
  unfolding Zfun_def by (auto elim!: eventually_rev_mp)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   106
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   107
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
   108
  unfolding Zfun_def by simp
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   109
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   110
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
   111
  unfolding Zfun_def by simp
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   112
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   113
lemma Zfun_imp_Zfun:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   114
  assumes f: "Zfun f F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   115
  assumes g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   116
  shows "Zfun (\<lambda>x. g x) F"
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   117
proof (cases)
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   118
  assume K: "0 < K"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   119
  show ?thesis
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   120
  proof (rule ZfunI)
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   121
    fix r::real assume "0 < r"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   122
    hence "0 < r / K"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   123
      using K by (rule divide_pos_pos)
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   124
    then have "eventually (\<lambda>x. norm (f x) < r / K) F"
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   125
      using ZfunD [OF f] by fast
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   126
    with g show "eventually (\<lambda>x. norm (g x) < r) F"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   127
    proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   128
      case (elim x)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   129
      hence "norm (f x) * K < r"
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   130
        by (simp add: pos_less_divide_eq K)
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   131
      thus ?case
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   132
        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
   133
    qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   134
  qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   135
next
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   136
  assume "\<not> 0 < K"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   137
  hence K: "K \<le> 0" by (simp only: not_less)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   138
  show ?thesis
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   139
  proof (rule ZfunI)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   140
    fix r :: real
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   141
    assume "0 < r"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   142
    from g show "eventually (\<lambda>x. norm (g x) < r) F"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   143
    proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   144
      case (elim x)
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   145
      also have "norm (f x) * K \<le> norm (f x) * 0"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   146
        using K norm_ge_zero by (rule mult_left_mono)
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   147
      finally show ?case
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   148
        using `0 < r` by simp
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   149
    qed
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   150
  qed
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   151
qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   152
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   153
lemma Zfun_le: "\<lbrakk>Zfun g F; \<forall>x. norm (f x) \<le> norm (g x)\<rbrakk> \<Longrightarrow> Zfun f F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   154
  by (erule_tac K="1" in Zfun_imp_Zfun, simp)
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   155
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   156
lemma Zfun_add:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   157
  assumes f: "Zfun f F" and g: "Zfun g F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   158
  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
   159
proof (rule ZfunI)
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   160
  fix r::real assume "0 < r"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   161
  hence r: "0 < r / 2" by simp
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   162
  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
   163
    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
   164
  moreover
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   165
  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
   166
    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
   167
  ultimately
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   168
  show "eventually (\<lambda>x. norm (f x + g x) < r) F"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   169
  proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   170
    case (elim x)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   171
    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
   172
      by (rule norm_triangle_ineq)
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   173
    also have "\<dots> < r/2 + r/2"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   174
      using elim by (rule add_strict_mono)
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   175
    finally show ?case
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   176
      by simp
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   177
  qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   178
qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   179
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   180
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
   181
  unfolding Zfun_def by simp
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   182
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   183
lemma Zfun_diff: "\<lbrakk>Zfun f F; Zfun g F\<rbrakk> \<Longrightarrow> 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
   184
  by (simp only: diff_minus Zfun_add Zfun_minus)
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   185
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   186
lemma (in bounded_linear) Zfun:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   187
  assumes g: "Zfun g F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   188
  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
   189
proof -
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   190
  obtain K where "\<And>x. norm (f x) \<le> norm x * K"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   191
    using bounded by fast
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   192
  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
   193
    by simp
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   194
  with g show ?thesis
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   195
    by (rule Zfun_imp_Zfun)
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   196
qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   197
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   198
lemma (in bounded_bilinear) Zfun:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   199
  assumes f: "Zfun f F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   200
  assumes g: "Zfun g F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   201
  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
   202
proof (rule ZfunI)
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   203
  fix r::real assume r: "0 < r"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   204
  obtain K where K: "0 < K"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   205
    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   206
    using pos_bounded by fast
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   207
  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
   208
    by (rule positive_imp_inverse_positive)
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   209
  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
   210
    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
   211
  moreover
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   212
  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
   213
    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
   214
  ultimately
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   215
  show "eventually (\<lambda>x. norm (f x ** g x) < r) F"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   216
  proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   217
    case (elim x)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   218
    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
   219
      by (rule norm_le)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   220
    also have "norm (f x) * norm (g x) * K < r * inverse K * K"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   221
      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
   222
    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
   223
      by simp
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   224
    finally show ?case .
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   225
  qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   226
qed
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   227
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   228
lemma (in bounded_bilinear) Zfun_left:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   229
  "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
   230
  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
   231
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   232
lemma (in bounded_bilinear) Zfun_right:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   233
  "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
   234
  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
   235
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   236
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
   237
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
   238
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
   239
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   240
lemma tendsto_Zfun_iff: "(f ---> 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
   241
  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
   242
44205
18da2a87421c generalize constant 'lim' and limit uniqueness theorems to class t2_space
huffman
parents: 44195
diff changeset
   243
subsubsection {* Distance and norms *}
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   244
31565
da5a5589418e theorem attribute [tendsto_intros]
huffman
parents: 31492
diff changeset
   245
lemma tendsto_norm [tendsto_intros]:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   246
  "(f ---> a) F \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> norm a) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   247
  unfolding norm_conv_dist by (intro tendsto_intros)
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   248
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
   249
lemma continuous_norm [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   250
  "continuous F f \<Longrightarrow> continuous F (\<lambda>x. norm (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   251
  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
   252
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   253
lemma continuous_on_norm [continuous_on_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
   254
  "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
   255
  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
   256
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   257
lemma tendsto_norm_zero:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   258
  "(f ---> 0) F \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> 0) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   259
  by (drule tendsto_norm, simp)
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   260
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   261
lemma tendsto_norm_zero_cancel:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   262
  "((\<lambda>x. norm (f x)) ---> 0) F \<Longrightarrow> (f ---> 0) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   263
  unfolding tendsto_iff dist_norm by simp
36662
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   264
621122eeb138 generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents: 36656
diff changeset
   265
lemma tendsto_norm_zero_iff:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   266
  "((\<lambda>x. norm (f x)) ---> 0) F \<longleftrightarrow> (f ---> 0) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   267
  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
   268
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   269
lemma tendsto_rabs [tendsto_intros]:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   270
  "(f ---> (l::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> \<bar>l\<bar>) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   271
  by (fold real_norm_def, rule tendsto_norm)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   272
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
   273
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
   274
  "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
   275
  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
   276
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   277
lemma continuous_on_rabs [continuous_on_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
   278
  "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
   279
  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
   280
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   281
lemma tendsto_rabs_zero:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   282
  "(f ---> (0::real)) F \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> 0) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   283
  by (fold real_norm_def, rule tendsto_norm_zero)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   284
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   285
lemma tendsto_rabs_zero_cancel:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   286
  "((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) F \<Longrightarrow> (f ---> 0) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   287
  by (fold real_norm_def, rule tendsto_norm_zero_cancel)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   288
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   289
lemma tendsto_rabs_zero_iff:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   290
  "((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) F \<longleftrightarrow> (f ---> 0) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   291
  by (fold real_norm_def, rule tendsto_norm_zero_iff)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   292
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   293
subsubsection {* Addition and subtraction *}
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   294
31565
da5a5589418e theorem attribute [tendsto_intros]
huffman
parents: 31492
diff changeset
   295
lemma tendsto_add [tendsto_intros]:
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   296
  fixes a b :: "'a::real_normed_vector"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   297
  shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> a + b) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   298
  by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add)
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   299
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
   300
lemma continuous_add [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   301
  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   302
  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
   303
  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
   304
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   305
lemma continuous_on_add [continuous_on_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
   306
  fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   307
  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
   308
  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
   309
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   310
lemma tendsto_add_zero:
51478
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   311
  fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   312
  shows "\<lbrakk>(f ---> 0) F; (g ---> 0) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> 0) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   313
  by (drule (1) tendsto_add, simp)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   314
31565
da5a5589418e theorem attribute [tendsto_intros]
huffman
parents: 31492
diff changeset
   315
lemma tendsto_minus [tendsto_intros]:
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   316
  fixes a :: "'a::real_normed_vector"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   317
  shows "(f ---> a) F \<Longrightarrow> ((\<lambda>x. - f x) ---> - a) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   318
  by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus)
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   319
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
   320
lemma continuous_minus [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   321
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   322
  shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. - f x)"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   323
  unfolding continuous_def by (rule tendsto_minus)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   324
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   325
lemma continuous_on_minus [continuous_on_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
   326
  fixes f :: "_ \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   327
  shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. - f x)"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   328
  unfolding continuous_on_def by (auto intro: tendsto_minus)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   329
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   330
lemma tendsto_minus_cancel:
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   331
  fixes a :: "'a::real_normed_vector"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   332
  shows "((\<lambda>x. - f x) ---> - a) F \<Longrightarrow> (f ---> a) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   333
  by (drule tendsto_minus, simp)
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   334
50330
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50327
diff changeset
   335
lemma tendsto_minus_cancel_left:
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50327
diff changeset
   336
    "(f ---> - (y::_::real_normed_vector)) F \<longleftrightarrow> ((\<lambda>x. - f x) ---> y) F"
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50327
diff changeset
   337
  using tendsto_minus_cancel[of f "- y" F]  tendsto_minus[of f "- y" F]
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50327
diff changeset
   338
  by auto
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50327
diff changeset
   339
31565
da5a5589418e theorem attribute [tendsto_intros]
huffman
parents: 31492
diff changeset
   340
lemma tendsto_diff [tendsto_intros]:
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   341
  fixes a b :: "'a::real_normed_vector"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   342
  shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x - g x) ---> a - b) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   343
  by (simp add: diff_minus tendsto_add tendsto_minus)
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
   344
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
   345
lemma continuous_diff [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   346
  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   347
  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
   348
  unfolding continuous_def by (rule tendsto_diff)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   349
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   350
lemma continuous_on_diff [continuous_on_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
   351
  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   352
  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
   353
  unfolding continuous_on_def by (auto intro: tendsto_diff)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   354
31588
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   355
lemma tendsto_setsum [tendsto_intros]:
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   356
  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::real_normed_vector"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   357
  assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> a i) F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   358
  shows "((\<lambda>x. \<Sum>i\<in>S. f i x) ---> (\<Sum>i\<in>S. a i)) F"
31588
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   359
proof (cases "finite S")
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   360
  assume "finite S" thus ?thesis using assms
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   361
    by (induct, simp add: tendsto_const, simp add: tendsto_add)
31588
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   362
next
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   363
  assume "\<not> finite S" thus ?thesis
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   364
    by (simp add: tendsto_const)
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   365
qed
2651f172c38b add lemma tendsto_setsum
huffman
parents: 31565
diff changeset
   366
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
   367
lemma continuous_setsum [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   368
  fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::real_normed_vector"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   369
  shows "(\<And>i. i \<in> S \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Sum>i\<in>S. f i x)"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   370
  unfolding continuous_def by (rule tendsto_setsum)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   371
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   372
lemma continuous_on_setsum [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   373
  fixes f :: "'a \<Rightarrow> _ \<Rightarrow> 'c::real_normed_vector"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   374
  shows "(\<And>i. i \<in> S \<Longrightarrow> continuous_on s (f i)) \<Longrightarrow> continuous_on s (\<lambda>x. \<Sum>i\<in>S. f i x)"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   375
  unfolding continuous_on_def by (auto intro: tendsto_setsum)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   376
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50880
diff changeset
   377
lemmas real_tendsto_sandwich = tendsto_sandwich[where 'b=real]
3de230ed0547 introduce order topology
hoelzl
parents: 50880
diff changeset
   378
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   379
subsubsection {* Linear operators and multiplication *}
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   380
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   381
lemma (in bounded_linear) tendsto:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   382
  "(g ---> a) F \<Longrightarrow> ((\<lambda>x. f (g x)) ---> f a) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   383
  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
   384
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
   385
lemma (in bounded_linear) continuous:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   386
  "continuous F g \<Longrightarrow> continuous F (\<lambda>x. f (g x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   387
  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
   388
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   389
lemma (in bounded_linear) continuous_on:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   390
  "continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f (g x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   391
  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
   392
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   393
lemma (in bounded_linear) tendsto_zero:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   394
  "(g ---> 0) F \<Longrightarrow> ((\<lambda>x. f (g x)) ---> 0) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   395
  by (drule tendsto, simp only: zero)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   396
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   397
lemma (in bounded_bilinear) tendsto:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   398
  "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x ** g x) ---> a ** b) F"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   399
  by (simp only: tendsto_Zfun_iff prod_diff_prod
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   400
                 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
   401
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
   402
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
   403
  "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
   404
  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
   405
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   406
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
   407
  "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
   408
  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
   409
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   410
lemma (in bounded_bilinear) tendsto_zero:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   411
  assumes f: "(f ---> 0) F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   412
  assumes g: "(g ---> 0) F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   413
  shows "((\<lambda>x. f x ** g x) ---> 0) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   414
  using tendsto [OF f g] by (simp add: zero_left)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   415
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   416
lemma (in bounded_bilinear) tendsto_left_zero:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   417
  "(f ---> 0) F \<Longrightarrow> ((\<lambda>x. f x ** c) ---> 0) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   418
  by (rule bounded_linear.tendsto_zero [OF bounded_linear_left])
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   419
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   420
lemma (in bounded_bilinear) tendsto_right_zero:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   421
  "(f ---> 0) F \<Longrightarrow> ((\<lambda>x. c ** f x) ---> 0) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   422
  by (rule bounded_linear.tendsto_zero [OF bounded_linear_right])
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   423
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   424
lemmas tendsto_of_real [tendsto_intros] =
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   425
  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
   426
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   427
lemmas tendsto_scaleR [tendsto_intros] =
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   428
  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
   429
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   430
lemmas tendsto_mult [tendsto_intros] =
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   431
  bounded_bilinear.tendsto [OF bounded_bilinear_mult]
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   432
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
   433
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
   434
  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
   435
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   436
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
   437
  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
   438
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   439
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
   440
  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
   441
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   442
lemmas continuous_on_of_real [continuous_on_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
   443
  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
   444
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   445
lemmas continuous_on_scaleR [continuous_on_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
   446
  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
   447
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   448
lemmas continuous_on_mult [continuous_on_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
   449
  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
   450
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44342
diff changeset
   451
lemmas tendsto_mult_zero =
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44342
diff changeset
   452
  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
   453
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44342
diff changeset
   454
lemmas tendsto_mult_left_zero =
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44342
diff changeset
   455
  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
   456
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44342
diff changeset
   457
lemmas tendsto_mult_right_zero =
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44342
diff changeset
   458
  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
   459
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   460
lemma tendsto_power [tendsto_intros]:
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   461
  fixes f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra}"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   462
  shows "(f ---> a) F \<Longrightarrow> ((\<lambda>x. f x ^ n) ---> a ^ n) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   463
  by (induct n) (simp_all add: tendsto_const tendsto_mult)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   464
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
   465
lemma continuous_power [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   466
  fixes f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   467
  shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. (f x)^n)"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   468
  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
   469
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   470
lemma continuous_on_power [continuous_on_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
   471
  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
   472
  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
   473
  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
   474
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   475
lemma tendsto_setprod [tendsto_intros]:
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   476
  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   477
  assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> L i) F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   478
  shows "((\<lambda>x. \<Prod>i\<in>S. f i x) ---> (\<Prod>i\<in>S. L i)) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   479
proof (cases "finite S")
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   480
  assume "finite S" thus ?thesis using assms
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   481
    by (induct, simp add: tendsto_const, simp add: tendsto_mult)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   482
next
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   483
  assume "\<not> finite S" thus ?thesis
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   484
    by (simp add: tendsto_const)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   485
qed
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   486
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
   487
lemma continuous_setprod [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   488
  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
   489
  shows "(\<And>i. i \<in> S \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Prod>i\<in>S. f i x)"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   490
  unfolding continuous_def by (rule tendsto_setprod)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   491
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   492
lemma continuous_on_setprod [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   493
  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
   494
  shows "(\<And>i. i \<in> S \<Longrightarrow> continuous_on s (f i)) \<Longrightarrow> continuous_on s (\<lambda>x. \<Prod>i\<in>S. f i x)"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   495
  unfolding continuous_on_def by (auto intro: tendsto_setprod)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   496
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   497
subsubsection {* Inverse and division *}
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   498
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   499
lemma (in bounded_bilinear) Zfun_prod_Bfun:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   500
  assumes f: "Zfun f F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   501
  assumes g: "Bfun g F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   502
  shows "Zfun (\<lambda>x. f x ** g x) F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   503
proof -
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   504
  obtain K where K: "0 \<le> K"
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   505
    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   506
    using nonneg_bounded by fast
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   507
  obtain B where B: "0 < B"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   508
    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
   509
    using g by (rule BfunE)
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   510
  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
   511
  using norm_g proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   512
    case (elim x)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   513
    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
   514
      by (rule norm_le)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   515
    also have "\<dots> \<le> norm (f x) * B * K"
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   516
      by (intro mult_mono' order_refl norm_g norm_ge_zero
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   517
                mult_nonneg_nonneg K elim)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   518
    also have "\<dots> = norm (f x) * (B * K)"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   519
      by (rule mult_assoc)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   520
    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
   521
  qed
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   522
  with f show ?thesis
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   523
    by (rule Zfun_imp_Zfun)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   524
qed
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   525
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   526
lemma (in bounded_bilinear) flip:
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   527
  "bounded_bilinear (\<lambda>x y. y ** x)"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   528
  apply default
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   529
  apply (rule add_right)
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   530
  apply (rule add_left)
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   531
  apply (rule scaleR_right)
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   532
  apply (rule scaleR_left)
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   533
  apply (subst mult_commute)
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   534
  using bounded by fast
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   535
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   536
lemma (in bounded_bilinear) Bfun_prod_Zfun:
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   537
  assumes f: "Bfun f F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   538
  assumes g: "Zfun g F"
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   539
  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
   540
  using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   541
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   542
lemma Bfun_inverse_lemma:
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   543
  fixes x :: "'a::real_normed_div_algebra"
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   544
  shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   545
  apply (subst nonzero_norm_inverse, clarsimp)
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   546
  apply (erule (1) le_imp_inverse_le)
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44079
diff changeset
   547
  done
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   548
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   549
lemma Bfun_inverse:
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   550
  fixes a :: "'a::real_normed_div_algebra"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   551
  assumes f: "(f ---> a) F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   552
  assumes a: "a \<noteq> 0"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   553
  shows "Bfun (\<lambda>x. inverse (f x)) F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   554
proof -
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   555
  from a have "0 < norm a" by simp
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   556
  hence "\<exists>r>0. r < norm a" by (rule dense)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   557
  then obtain r where r1: "0 < r" and r2: "r < norm a" by fast
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   558
  have "eventually (\<lambda>x. dist (f x) a < r) F"
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   559
    using tendstoD [OF f r1] by fast
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   560
  hence "eventually (\<lambda>x. norm (inverse (f x)) \<le> inverse (norm a - r)) F"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   561
  proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 46886
diff changeset
   562
    case (elim x)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   563
    hence 1: "norm (f x - a) < r"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   564
      by (simp add: dist_norm)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   565
    hence 2: "f x \<noteq> 0" using r2 by auto
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   566
    hence "norm (inverse (f x)) = inverse (norm (f x))"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   567
      by (rule nonzero_norm_inverse)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   568
    also have "\<dots> \<le> inverse (norm a - r)"
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   569
    proof (rule le_imp_inverse_le)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   570
      show "0 < norm a - r" using r2 by simp
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   571
    next
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   572
      have "norm a - norm (f x) \<le> norm (a - f x)"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   573
        by (rule norm_triangle_ineq2)
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   574
      also have "\<dots> = norm (f x - a)"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   575
        by (rule norm_minus_commute)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   576
      also have "\<dots> < r" using 1 .
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   577
      finally show "norm a - r \<le> norm (f x)" by simp
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   578
    qed
31487
93938cafc0e6 put syntax for tendsto in Limits.thy; rename variables
huffman
parents: 31447
diff changeset
   579
    finally show "norm (inverse (f x)) \<le> inverse (norm a - r)" .
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   580
  qed
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   581
  thus ?thesis by (rule BfunI)
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   582
qed
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   583
31565
da5a5589418e theorem attribute [tendsto_intros]
huffman
parents: 31492
diff changeset
   584
lemma tendsto_inverse [tendsto_intros]:
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   585
  fixes a :: "'a::real_normed_div_algebra"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   586
  assumes f: "(f ---> a) F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   587
  assumes a: "a \<noteq> 0"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   588
  shows "((\<lambda>x. inverse (f x)) ---> inverse a) F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   589
proof -
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   590
  from a have "0 < norm a" by simp
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   591
  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
   592
    by (rule tendstoD)
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   593
  then have "eventually (\<lambda>x. f x \<noteq> 0) F"
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   594
    unfolding dist_norm by (auto elim!: eventually_elim1)
44627
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   595
  with a have "eventually (\<lambda>x. inverse (f x) - inverse a =
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   596
    - (inverse (f x) * (f x - a) * inverse a)) F"
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   597
    by (auto elim!: eventually_elim1 simp: inverse_diff_inverse)
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   598
  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
   599
    by (intro Zfun_minus Zfun_mult_left
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   600
      bounded_bilinear.Bfun_prod_Zfun [OF bounded_bilinear_mult]
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   601
      Bfun_inverse [OF f a] f [unfolded tendsto_Zfun_iff])
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   602
  ultimately show ?thesis
134c06282ae6 convert to Isar-style proof
huffman
parents: 44571
diff changeset
   603
    unfolding tendsto_Zfun_iff by (rule Zfun_ssubst)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   604
qed
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   605
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
   606
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
   607
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   608
  assumes "continuous F f" and "f (Lim F (\<lambda>x. x)) \<noteq> 0"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   609
  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
   610
  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
   611
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   612
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
   613
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   614
  assumes "continuous (at a within s) f" and "f a \<noteq> 0"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   615
  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
   616
  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
   617
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   618
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
   619
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   620
  assumes "isCont f a" and "f a \<noteq> 0"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   621
  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
   622
  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
   623
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   624
lemma continuous_on_inverse[continuous_on_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
   625
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   626
  assumes "continuous_on s f" and "\<forall>x\<in>s. f x \<noteq> 0"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   627
  shows "continuous_on 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
   628
  using assms unfolding continuous_on_def by (fast intro: 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
   629
31565
da5a5589418e theorem attribute [tendsto_intros]
huffman
parents: 31492
diff changeset
   630
lemma tendsto_divide [tendsto_intros]:
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   631
  fixes a b :: "'a::real_normed_field"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   632
  shows "\<lbrakk>(f ---> a) F; (g ---> b) F; b \<noteq> 0\<rbrakk>
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   633
    \<Longrightarrow> ((\<lambda>x. f x / g x) ---> a / b) F"
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44253
diff changeset
   634
  by (simp add: tendsto_mult tendsto_inverse divide_inverse)
31355
3d18766ddc4b limits of inverse using filters
huffman
parents: 31353
diff changeset
   635
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
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
   637
  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
   638
  assumes "continuous F f" and "continuous F g" and "g (Lim F (\<lambda>x. x)) \<noteq> 0"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   639
  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
   640
  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
   641
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   642
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
   643
  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
   644
  assumes "continuous (at a within s) f" "continuous (at a within s) g" and "g a \<noteq> 0"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   645
  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
   646
  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
   647
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   648
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
   649
  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
   650
  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
   651
  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
   652
  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
   653
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   654
lemma continuous_on_divide[continuous_on_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
   655
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_field"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   656
  assumes "continuous_on s f" "continuous_on s g" and "\<forall>x\<in>s. g x \<noteq> 0"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   657
  shows "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
   658
  using assms unfolding continuous_on_def by (fast intro: 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
   659
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   660
lemma tendsto_sgn [tendsto_intros]:
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   661
  fixes l :: "'a::real_normed_vector"
44195
f5363511b212 consistently use variable name 'F' for filters
huffman
parents: 44194
diff changeset
   662
  shows "\<lbrakk>(f ---> l) F; l \<noteq> 0\<rbrakk> \<Longrightarrow> ((\<lambda>x. sgn (f x)) ---> sgn l) F"
44194
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   663
  unfolding sgn_div_norm by (simp add: tendsto_intros)
0639898074ae generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents: 44081
diff changeset
   664
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
   665
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
   666
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   667
  assumes "continuous F f" and "f (Lim F (\<lambda>x. x)) \<noteq> 0"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   668
  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
   669
  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
   670
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   671
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
   672
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   673
  assumes "continuous (at a within s) f" and "f a \<noteq> 0"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   674
  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
   675
  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
   676
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   677
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
   678
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   679
  assumes "isCont f a" and "f a \<noteq> 0"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   680
  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
   681
  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
   682
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   683
lemma continuous_on_sgn[continuous_on_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
   684
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   685
  assumes "continuous_on s f" and "\<forall>x\<in>s. f x \<noteq> 0"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51474
diff changeset
   686
  shows "continuous_on 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
   687
  using assms unfolding continuous_on_def by (fast intro: 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
   688
50325
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   689
lemma filterlim_at_infinity:
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   690
  fixes f :: "_ \<Rightarrow> 'a\<Colon>real_normed_vector"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   691
  assumes "0 \<le> c"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   692
  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
   693
  unfolding filterlim_iff eventually_at_infinity
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   694
proof safe
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   695
  fix P :: "'a \<Rightarrow> bool" and b
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   696
  assume *: "\<forall>r>c. eventually (\<lambda>x. r \<le> norm (f x)) F"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   697
    and P: "\<forall>x. b \<le> norm x \<longrightarrow> P x"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   698
  have "max b (c + 1) > c" by auto
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   699
  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
   700
    by auto
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   701
  then show "eventually (\<lambda>x. P (f x)) F"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   702
  proof eventually_elim
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   703
    fix x assume "max b (c + 1) \<le> norm (f x)"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   704
    with P show "P (f x)" by auto
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   705
  qed
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   706
qed force
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   707
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   708
subsection {* Relate @{const at}, @{const at_left} and @{const at_right} *}
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   709
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   710
text {*
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   711
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   712
This lemmas are useful for conversion between @{term "at x"} to @{term "at_left x"} and
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   713
@{term "at_right x"} and also @{term "at_right 0"}.
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   714
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   715
*}
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   716
51471
cad22a3cc09c move topological_space to its own theory
hoelzl
parents: 51360
diff changeset
   717
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
   718
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   719
lemma filtermap_nhds_shift: "filtermap (\<lambda>x. x - d) (nhds a) = nhds (a - d::real)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   720
  unfolding filter_eq_iff eventually_filtermap eventually_nhds_metric
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   721
  by (intro allI ex_cong) (auto simp: dist_real_def field_simps)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   722
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   723
lemma filtermap_nhds_minus: "filtermap (\<lambda>x. - x) (nhds a) = nhds (- a::real)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   724
  unfolding filter_eq_iff eventually_filtermap eventually_nhds_metric
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   725
  apply (intro allI ex_cong)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   726
  apply (auto simp: dist_real_def field_simps)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   727
  apply (erule_tac x="-x" in allE)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   728
  apply simp
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   729
  done
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   730
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   731
lemma filtermap_at_shift: "filtermap (\<lambda>x. x - d) (at a) = at (a - d::real)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   732
  unfolding at_def filtermap_nhds_shift[symmetric]
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   733
  by (simp add: filter_eq_iff eventually_filtermap eventually_within)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   734
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   735
lemma filtermap_at_right_shift: "filtermap (\<lambda>x. x - d) (at_right a) = at_right (a - d::real)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   736
  unfolding filtermap_at_shift[symmetric]
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   737
  by (simp add: filter_eq_iff eventually_filtermap eventually_within)
50323
3764d4620fb3 add filterlim rules for unary minus and inverse
hoelzl
parents: 50322
diff changeset
   738
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   739
lemma at_right_to_0: "at_right (a::real) = filtermap (\<lambda>x. x + a) (at_right 0)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   740
  using filtermap_at_right_shift[of "-a" 0] by simp
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   741
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   742
lemma filterlim_at_right_to_0:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   743
  "filterlim f F (at_right (a::real)) \<longleftrightarrow> filterlim (\<lambda>x. f (x + a)) F (at_right 0)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   744
  unfolding filterlim_def filtermap_filtermap at_right_to_0[of a] ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   745
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   746
lemma eventually_at_right_to_0:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   747
  "eventually P (at_right (a::real)) \<longleftrightarrow> eventually (\<lambda>x. P (x + a)) (at_right 0)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   748
  unfolding at_right_to_0[of a] by (simp add: eventually_filtermap)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   749
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   750
lemma filtermap_at_minus: "filtermap (\<lambda>x. - x) (at a) = at (- a::real)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   751
  unfolding at_def filtermap_nhds_minus[symmetric]
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   752
  by (simp add: filter_eq_iff eventually_filtermap eventually_within)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   753
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   754
lemma at_left_minus: "at_left (a::real) = filtermap (\<lambda>x. - x) (at_right (- a))"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   755
  by (simp add: filter_eq_iff eventually_filtermap eventually_within filtermap_at_minus[symmetric])
50323
3764d4620fb3 add filterlim rules for unary minus and inverse
hoelzl
parents: 50322
diff changeset
   756
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   757
lemma at_right_minus: "at_right (a::real) = filtermap (\<lambda>x. - x) (at_left (- a))"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   758
  by (simp add: filter_eq_iff eventually_filtermap eventually_within filtermap_at_minus[symmetric])
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   759
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   760
lemma filterlim_at_left_to_right:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   761
  "filterlim f F (at_left (a::real)) \<longleftrightarrow> filterlim (\<lambda>x. f (- x)) F (at_right (-a))"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   762
  unfolding filterlim_def filtermap_filtermap at_left_minus[of a] ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   763
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   764
lemma eventually_at_left_to_right:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   765
  "eventually P (at_left (a::real)) \<longleftrightarrow> eventually (\<lambda>x. P (- x)) (at_right (-a))"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   766
  unfolding at_left_minus[of a] by (simp add: eventually_filtermap)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   767
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
   768
lemma at_top_mirror: "at_top = filtermap uminus (at_bot :: 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
   769
  unfolding filter_eq_iff eventually_filtermap eventually_at_top_linorder eventually_at_bot_linorder
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
   770
  by (metis le_minus_iff minus_minus)
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
   771
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
   772
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
   773
  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
   774
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
   775
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
   776
  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
   777
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
   778
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
   779
  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
   780
50323
3764d4620fb3 add filterlim rules for unary minus and inverse
hoelzl
parents: 50322
diff changeset
   781
lemma filterlim_uminus_at_top_at_bot: "LIM x at_bot. - x :: real :> at_top"
3764d4620fb3 add filterlim rules for unary minus and inverse
hoelzl
parents: 50322
diff changeset
   782
  unfolding filterlim_at_top eventually_at_bot_dense
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
   783
  by (metis leI minus_less_iff order_less_asym)
50323
3764d4620fb3 add filterlim rules for unary minus and inverse
hoelzl
parents: 50322
diff changeset
   784
3764d4620fb3 add filterlim rules for unary minus and inverse
hoelzl
parents: 50322
diff changeset
   785
lemma filterlim_uminus_at_bot_at_top: "LIM x at_top. - x :: real :> at_bot"
3764d4620fb3 add filterlim rules for unary minus and inverse
hoelzl
parents: 50322
diff changeset
   786
  unfolding filterlim_at_bot eventually_at_top_dense
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
   787
  by (metis leI less_minus_iff order_less_asym)
50323
3764d4620fb3 add filterlim rules for unary minus and inverse
hoelzl
parents: 50322
diff changeset
   788
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
   789
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
   790
  using filterlim_compose[OF filterlim_uminus_at_bot_at_top, of f F]
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50331
diff changeset
   791
  using filterlim_compose[OF filterlim_uminus_at_top_at_bot, of "\<lambda>x. - f x" F]
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50331
diff changeset
   792
  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
   793
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
   794
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
   795
  unfolding filterlim_uminus_at_top by simp
50323
3764d4620fb3 add filterlim rules for unary minus and inverse
hoelzl
parents: 50322
diff changeset
   796
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   797
lemma filterlim_inverse_at_top_right: "LIM x at_right (0::real). inverse x :> at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   798
  unfolding filterlim_at_top_gt[where c=0] eventually_within at_def
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   799
proof safe
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   800
  fix Z :: real assume [arith]: "0 < Z"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   801
  then have "eventually (\<lambda>x. x < inverse Z) (nhds 0)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   802
    by (auto simp add: eventually_nhds_metric dist_real_def intro!: exI[of _ "\<bar>inverse Z\<bar>"])
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   803
  then show "eventually (\<lambda>x. x \<in> - {0} \<longrightarrow> x \<in> {0<..} \<longrightarrow> Z \<le> inverse x) (nhds 0)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   804
    by (auto elim!: eventually_elim1 simp: inverse_eq_divide field_simps)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   805
qed
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   806
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   807
lemma filterlim_inverse_at_top:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   808
  "(f ---> (0 :: real)) F \<Longrightarrow> eventually (\<lambda>x. 0 < f x) F \<Longrightarrow> LIM x F. inverse (f x) :> at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   809
  by (intro filterlim_compose[OF filterlim_inverse_at_top_right])
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   810
     (simp add: filterlim_def eventually_filtermap le_within_iff at_def eventually_elim1)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   811
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   812
lemma filterlim_inverse_at_bot_neg:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   813
  "LIM x (at_left (0::real)). inverse x :> at_bot"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   814
  by (simp add: filterlim_inverse_at_top_right filterlim_uminus_at_bot filterlim_at_left_to_right)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   815
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   816
lemma filterlim_inverse_at_bot:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   817
  "(f ---> (0 :: real)) F \<Longrightarrow> eventually (\<lambda>x. f x < 0) F \<Longrightarrow> LIM x F. inverse (f x) :> at_bot"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   818
  unfolding filterlim_uminus_at_bot inverse_minus_eq[symmetric]
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   819
  by (rule filterlim_inverse_at_top) (simp_all add: tendsto_minus_cancel_left[symmetric])
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   820
50325
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   821
lemma tendsto_inverse_0:
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   822
  fixes x :: "_ \<Rightarrow> 'a\<Colon>real_normed_div_algebra"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   823
  shows "(inverse ---> (0::'a)) at_infinity"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   824
  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
   825
proof safe
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   826
  fix r :: real assume "0 < r"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   827
  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
   828
  proof (intro exI[of _ "inverse (r / 2)"] allI impI)
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   829
    fix x :: 'a
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   830
    from `0 < r` have "0 < inverse (r / 2)" by simp
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   831
    also assume *: "inverse (r / 2) \<le> norm x"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   832
    finally show "norm (inverse x) < r"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   833
      using * `0 < r` by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps)
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   834
  qed
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   835
qed
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   836
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   837
lemma at_right_to_top: "(at_right (0::real)) = filtermap inverse at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   838
proof (rule antisym)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   839
  have "(inverse ---> (0::real)) at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   840
    by (metis tendsto_inverse_0 filterlim_mono at_top_le_at_infinity order_refl)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   841
  then show "filtermap inverse at_top \<le> at_right (0::real)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   842
    unfolding at_within_eq
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   843
    by (intro le_withinI) (simp_all add: eventually_filtermap eventually_gt_at_top filterlim_def)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   844
next
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   845
  have "filtermap inverse (filtermap inverse (at_right (0::real))) \<le> filtermap inverse at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   846
    using filterlim_inverse_at_top_right unfolding filterlim_def by (rule filtermap_mono)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   847
  then show "at_right (0::real) \<le> filtermap inverse at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   848
    by (simp add: filtermap_ident filtermap_filtermap)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   849
qed
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   850
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   851
lemma eventually_at_right_to_top:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   852
  "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
   853
  unfolding at_right_to_top eventually_filtermap ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   854
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   855
lemma filterlim_at_right_to_top:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   856
  "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
   857
  unfolding filterlim_def at_right_to_top filtermap_filtermap ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   858
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   859
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
   860
  unfolding at_right_to_top filtermap_filtermap inverse_inverse_eq filtermap_ident ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   861
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   862
lemma eventually_at_top_to_right:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   863
  "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
   864
  unfolding at_top_to_right eventually_filtermap ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   865
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   866
lemma filterlim_at_top_to_right:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   867
  "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
   868
  unfolding filterlim_def at_top_to_right filtermap_filtermap ..
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   869
50325
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   870
lemma filterlim_inverse_at_infinity:
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   871
  fixes x :: "_ \<Rightarrow> 'a\<Colon>{real_normed_div_algebra, division_ring_inverse_zero}"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   872
  shows "filterlim inverse at_infinity (at (0::'a))"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   873
  unfolding filterlim_at_infinity[OF order_refl]
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   874
proof safe
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   875
  fix r :: real assume "0 < r"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   876
  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
   877
    unfolding eventually_at norm_inverse
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   878
    by (intro exI[of _ "inverse r"])
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   879
       (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
   880
qed
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   881
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   882
lemma filterlim_inverse_at_iff:
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   883
  fixes g :: "'a \<Rightarrow> 'b\<Colon>{real_normed_div_algebra, division_ring_inverse_zero}"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   884
  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
   885
  unfolding filterlim_def filtermap_filtermap[symmetric]
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   886
proof
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   887
  assume "filtermap g F \<le> at_infinity"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   888
  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
   889
    by (rule filtermap_mono)
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   890
  also have "\<dots> \<le> at 0"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   891
    using tendsto_inverse_0
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   892
    by (auto intro!: le_withinI exI[of _ 1]
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   893
             simp: eventually_filtermap eventually_at_infinity filterlim_def at_def)
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   894
  finally show "filtermap inverse (filtermap g F) \<le> at 0" .
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   895
next
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   896
  assume "filtermap inverse (filtermap g F) \<le> at 0"
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   897
  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
   898
    by (rule filtermap_mono)
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   899
  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
   900
    by (auto intro: order_trans simp: filterlim_def filtermap_filtermap)
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   901
qed
5e40ad9f212a add filterlim rules for inverse and at_infinity
hoelzl
parents: 50324
diff changeset
   902
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
   903
lemma tendsto_inverse_0_at_top:
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
   904
  "LIM x F. f x :> at_top \<Longrightarrow> ((\<lambda>x. inverse (f x) :: real) ---> 0) F"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
   905
 by (metis at_top_le_at_infinity filterlim_at filterlim_inverse_at_iff filterlim_mono order_refl)
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
   906
50324
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   907
text {*
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   908
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   909
We only show rules for multiplication and addition when the functions are either against a real
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   910
value or against infinity. Further rules are easy to derive by using @{thm filterlim_uminus_at_top}.
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   911
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   912
*}
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   913
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   914
lemma filterlim_tendsto_pos_mult_at_top: 
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   915
  assumes f: "(f ---> c) F" and c: "0 < c"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   916
  assumes g: "LIM x F. g x :> at_top"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   917
  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
   918
  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
   919
proof safe
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   920
  fix Z :: real assume "0 < Z"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   921
  from f `0 < c` have "eventually (\<lambda>x. c / 2 < f x) F"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   922
    by (auto dest!: tendstoD[where e="c / 2"] elim!: eventually_elim1
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   923
             simp: dist_real_def abs_real_def split: split_if_asm)
50346
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50331
diff changeset
   924
  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
   925
    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
   926
  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
   927
  proof eventually_elim
50346
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50331
diff changeset
   928
    fix x assume "c / 2 < f x" "Z / c * 2 \<le> g x"
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
   929
    with `0 < Z` `0 < c` have "c / 2 * (Z / c * 2) \<le> f x * g x"
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
   930
      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
   931
    with `0 < c` 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
   932
       by simp
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   933
  qed
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   934
qed
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   935
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   936
lemma filterlim_at_top_mult_at_top: 
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   937
  assumes f: "LIM x F. f x :> at_top"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   938
  assumes g: "LIM x F. g x :> at_top"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   939
  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
   940
  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
   941
proof safe
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   942
  fix Z :: real assume "0 < Z"
50346
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50331
diff changeset
   943
  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
   944
    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
   945
  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
   946
    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
   947
  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
   948
  proof eventually_elim
50346
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50331
diff changeset
   949
    fix x assume "1 \<le> f x" "Z \<le> g x"
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
   950
    with `0 < Z` have "1 * Z \<le> f x * g x"
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
   951
      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
   952
    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
   953
       by simp
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   954
  qed
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   955
qed
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   956
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
   957
lemma filterlim_tendsto_pos_mult_at_bot:
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
   958
  assumes "(f ---> c) F" "0 < (c::real)" "filterlim g at_bot F"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
   959
  shows "LIM x F. f x * g x :> at_bot"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
   960
  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
   961
  unfolding filterlim_uminus_at_bot by simp
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50347
diff changeset
   962
50324
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   963
lemma filterlim_tendsto_add_at_top: 
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   964
  assumes f: "(f ---> c) F"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   965
  assumes g: "LIM x F. g x :> at_top"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   966
  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
   967
  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
   968
proof safe
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   969
  fix Z :: real assume "0 < Z"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   970
  from f have "eventually (\<lambda>x. c - 1 < f x) F"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   971
    by (auto dest!: tendstoD[where e=1] elim!: eventually_elim1 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
   972
  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
   973
    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
   974
  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
   975
    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
   976
qed
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   977
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   978
lemma LIM_at_top_divide:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   979
  fixes f g :: "'a \<Rightarrow> real"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   980
  assumes f: "(f ---> a) F" "0 < a"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   981
  assumes g: "(g ---> 0) F" "eventually (\<lambda>x. 0 < g x) F"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   982
  shows "LIM x F. f x / g x :> at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   983
  unfolding divide_inverse
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
   984
  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
   985
50324
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   986
lemma filterlim_at_top_add_at_top: 
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   987
  assumes f: "LIM x F. f x :> at_top"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   988
  assumes g: "LIM x F. g x :> at_top"
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   989
  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
   990
  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
   991
proof safe
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
   992
  fix Z :: real assume "0 < Z"
50346
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50331
diff changeset
   993
  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
   994
    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
   995
  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
   996
    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
   997
  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
   998
    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
   999
qed
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1000
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1001
lemma tendsto_divide_0:
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1002
  fixes f :: "_ \<Rightarrow> 'a\<Colon>{real_normed_div_algebra, division_ring_inverse_zero}"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1003
  assumes f: "(f ---> c) F"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1004
  assumes g: "LIM x F. g x :> at_infinity"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1005
  shows "((\<lambda>x. f x / g x) ---> 0) F"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1006
  using tendsto_mult[OF f filterlim_compose[OF tendsto_inverse_0 g]] by (simp add: divide_inverse)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1007
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1008
lemma linear_plus_1_le_power:
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1009
  fixes x :: real
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1010
  assumes x: "0 \<le> x"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1011
  shows "real n * x + 1 \<le> (x + 1) ^ n"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1012
proof (induct n)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1013
  case (Suc n)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1014
  have "real (Suc n) * x + 1 \<le> (x + 1) * (real n * x + 1)"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1015
    by (simp add: field_simps real_of_nat_Suc mult_nonneg_nonneg x)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1016
  also have "\<dots> \<le> (x + 1)^Suc n"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1017
    using Suc x by (simp add: mult_left_mono)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1018
  finally show ?case .
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1019
qed simp
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1020
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1021
lemma filterlim_realpow_sequentially_gt1:
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1022
  fixes x :: "'a :: real_normed_div_algebra"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1023
  assumes x[arith]: "1 < norm x"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1024
  shows "LIM n sequentially. x ^ n :> at_infinity"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1025
proof (intro filterlim_at_infinity[THEN iffD2] allI impI)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1026
  fix y :: real assume "0 < y"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1027
  have "0 < norm x - 1" by simp
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1028
  then obtain N::nat where "y < real N * (norm x - 1)" by (blast dest: reals_Archimedean3)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1029
  also have "\<dots> \<le> real N * (norm x - 1) + 1" by simp
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1030
  also have "\<dots> \<le> (norm x - 1 + 1) ^ N" by (rule linear_plus_1_le_power) simp
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1031
  also have "\<dots> = norm x ^ N" by simp
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1032
  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
  1033
    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
  1034
  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
  1035
    unfolding eventually_sequentially
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1036
    by (auto simp: norm_power)
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1037
qed simp
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
  1038
51471
cad22a3cc09c move topological_space to its own theory
hoelzl
parents: 51360
diff changeset
  1039
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1040
subsection {* Limits of Sequences *}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1041
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1042
lemma [trans]: "X=Y ==> Y ----> z ==> X ----> z"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1043
  by simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1044
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1045
lemma LIMSEQ_iff:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1046
  fixes L :: "'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1047
  shows "(X ----> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1048
unfolding LIMSEQ_def dist_norm ..
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1049
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1050
lemma LIMSEQ_I:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1051
  fixes L :: "'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1052
  shows "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X ----> L"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1053
by (simp add: LIMSEQ_iff)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1054
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1055
lemma LIMSEQ_D:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1056
  fixes L :: "'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1057
  shows "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1058
by (simp add: LIMSEQ_iff)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1059
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1060
lemma LIMSEQ_linear: "\<lbrakk> X ----> x ; l > 0 \<rbrakk> \<Longrightarrow> (\<lambda> n. X (n * l)) ----> x"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1061
  unfolding tendsto_def eventually_sequentially
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1062
  by (metis div_le_dividend div_mult_self1_is_m le_trans nat_mult_commute)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1063
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1064
lemma Bseq_inverse_lemma:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1065
  fixes x :: "'a::real_normed_div_algebra"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1066
  shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1067
apply (subst nonzero_norm_inverse, clarsimp)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1068
apply (erule (1) le_imp_inverse_le)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1069
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1070
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1071
lemma Bseq_inverse:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1072
  fixes a :: "'a::real_normed_div_algebra"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1073
  shows "\<lbrakk>X ----> a; a \<noteq> 0\<rbrakk> \<Longrightarrow> Bseq (\<lambda>n. inverse (X n))"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1074
  by (rule Bfun_inverse)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1075
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1076
lemma LIMSEQ_diff_approach_zero:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1077
  fixes L :: "'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1078
  shows "g ----> L ==> (%x. f x - g x) ----> 0 ==> f ----> L"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1079
  by (drule (1) tendsto_add, simp)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1080
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1081
lemma LIMSEQ_diff_approach_zero2:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1082
  fixes L :: "'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1083
  shows "f ----> L ==> (%x. f x - g x) ----> 0 ==> g ----> L"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1084
  by (drule (1) tendsto_diff, simp)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1085
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1086
text{*An unbounded sequence's inverse tends to 0*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1087
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1088
lemma LIMSEQ_inverse_zero:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1089
  "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1090
  apply (rule filterlim_compose[OF tendsto_inverse_0])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1091
  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
  1092
  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
  1093
  done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1094
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1095
text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1096
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1097
lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1098
  by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1099
            filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1100
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1101
text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1102
infinity is now easily proved*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1103
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1104
lemma LIMSEQ_inverse_real_of_nat_add:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1105
     "(%n. r + inverse(real(Suc n))) ----> r"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1106
  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
  1107
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1108
lemma LIMSEQ_inverse_real_of_nat_add_minus:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1109
     "(%n. r + -inverse(real(Suc n))) ----> r"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1110
  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
  1111
  by auto
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1112
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1113
lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1114
     "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1115
  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
  1116
  by auto
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1117
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1118
subsection {* Convergence on sequences *}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1119
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1120
lemma convergent_add:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1121
  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1122
  assumes "convergent (\<lambda>n. X n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1123
  assumes "convergent (\<lambda>n. Y n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1124
  shows "convergent (\<lambda>n. X n + Y n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1125
  using assms unfolding convergent_def by (fast intro: tendsto_add)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1126
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1127
lemma convergent_setsum:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1128
  fixes X :: "'a \<Rightarrow> nat \<Rightarrow> 'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1129
  assumes "\<And>i. i \<in> A \<Longrightarrow> convergent (\<lambda>n. X i n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1130
  shows "convergent (\<lambda>n. \<Sum>i\<in>A. X i n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1131
proof (cases "finite A")
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1132
  case True from this and assms show ?thesis
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1133
    by (induct A set: finite) (simp_all add: convergent_const convergent_add)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1134
qed (simp add: convergent_const)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1135
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1136
lemma (in bounded_linear) convergent:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1137
  assumes "convergent (\<lambda>n. X n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1138
  shows "convergent (\<lambda>n. f (X n))"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1139
  using assms unfolding convergent_def by (fast intro: tendsto)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1140
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1141
lemma (in bounded_bilinear) convergent:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1142
  assumes "convergent (\<lambda>n. X n)" and "convergent (\<lambda>n. Y n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1143
  shows "convergent (\<lambda>n. X n ** Y n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1144
  using assms unfolding convergent_def by (fast intro: tendsto)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1145
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1146
lemma convergent_minus_iff:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1147
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1148
  shows "convergent X \<longleftrightarrow> convergent (\<lambda>n. - X n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1149
apply (simp add: convergent_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1150
apply (auto dest: tendsto_minus)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1151
apply (drule tendsto_minus, auto)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1152
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1153
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1154
subsection {* Bounded Monotonic Sequences *}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1155
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1156
subsubsection {* Bounded Sequences *}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1157
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1158
lemma BseqI': "(\<And>n. norm (X n) \<le> K) \<Longrightarrow> Bseq X"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1159
  by (intro BfunI) (auto simp: eventually_sequentially)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1160
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1161
lemma BseqI2': "\<forall>n\<ge>N. norm (X n) \<le> K \<Longrightarrow> Bseq X"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1162
  by (intro BfunI) (auto simp: eventually_sequentially)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1163
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1164
lemma Bseq_def: "Bseq X \<longleftrightarrow> (\<exists>K>0. \<forall>n. norm (X n) \<le> K)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1165
  unfolding Bfun_def eventually_sequentially
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1166
proof safe
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1167
  fix N K assume "0 < K" "\<forall>n\<ge>N. norm (X n) \<le> K"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1168
  then show "\<exists>K>0. \<forall>n. norm (X n) \<le> K"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1169
    by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] min_max.less_supI2)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1170
       (auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1171
qed auto
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1172
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1173
lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1174
unfolding Bseq_def by auto
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1175
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1176
lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1177
by (simp add: Bseq_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1178
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1179
lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1180
by (auto simp add: Bseq_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1181
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1182
lemma lemma_NBseq_def:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1183
  "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1184
proof safe
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1185
  fix K :: real
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1186
  from reals_Archimedean2 obtain n :: nat where "K < real n" ..
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1187
  then have "K \<le> real (Suc n)" by auto
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1188
  moreover assume "\<forall>m. norm (X m) \<le> K"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1189
  ultimately have "\<forall>m. norm (X m) \<le> real (Suc n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1190
    by (blast intro: order_trans)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1191
  then show "\<exists>N. \<forall>n. norm (X n) \<le> real (Suc N)" ..
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1192
qed (force simp add: real_of_nat_Suc)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1193
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1194
text{* alternative definition for Bseq *}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1195
lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1196
apply (simp add: Bseq_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1197
apply (simp (no_asm) add: lemma_NBseq_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1198
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1199
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1200
lemma lemma_NBseq_def2:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1201
     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1202
apply (subst lemma_NBseq_def, auto)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1203
apply (rule_tac x = "Suc N" in exI)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1204
apply (rule_tac [2] x = N in exI)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1205
apply (auto simp add: real_of_nat_Suc)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1206
 prefer 2 apply (blast intro: order_less_imp_le)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1207
apply (drule_tac x = n in spec, simp)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1208
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1209
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1210
(* yet another definition for Bseq *)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1211
lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1212
by (simp add: Bseq_def lemma_NBseq_def2)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1213
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1214
subsubsection{*A Few More Equivalence Theorems for Boundedness*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1215
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1216
text{*alternative formulation for boundedness*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1217
lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1218
apply (unfold Bseq_def, safe)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1219
apply (rule_tac [2] x = "k + norm x" in exI)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1220
apply (rule_tac x = K in exI, simp)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1221
apply (rule exI [where x = 0], auto)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1222
apply (erule order_less_le_trans, simp)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1223
apply (drule_tac x=n in spec, fold diff_minus)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1224
apply (drule order_trans [OF norm_triangle_ineq2])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1225
apply simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1226
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1227
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1228
text{*alternative formulation for boundedness*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1229
lemma Bseq_iff3: "Bseq X = (\<exists>k > 0. \<exists>N. \<forall>n. norm(X(n) + -X(N)) \<le> k)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1230
apply safe
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1231
apply (simp add: Bseq_def, safe)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1232
apply (rule_tac x = "K + norm (X N)" in exI)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1233
apply auto
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1234
apply (erule order_less_le_trans, simp)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1235
apply (rule_tac x = N in exI, safe)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1236
apply (drule_tac x = n in spec)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1237
apply (rule order_trans [OF norm_triangle_ineq], simp)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1238
apply (auto simp add: Bseq_iff2)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1239
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1240
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1241
lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1242
apply (simp add: Bseq_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1243
apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1244
apply (drule_tac x = n in spec, arith)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1245
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1246
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1247
subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1248
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1249
lemma Bseq_isUb:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1250
  "!!(X::nat=>real). Bseq X ==> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1251
by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1252
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1253
text{* Use completeness of reals (supremum property)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1254
   to show that any bounded sequence has a least upper bound*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1255
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1256
lemma Bseq_isLub:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1257
  "!!(X::nat=>real). Bseq X ==>
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1258
   \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1259
by (blast intro: reals_complete Bseq_isUb)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1260
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1261
subsubsection{*A Bounded and Monotonic Sequence Converges*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1262
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1263
(* TODO: delete *)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1264
(* FIXME: one use in NSA/HSEQ.thy *)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1265
lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1266
  apply (rule_tac x="X m" in exI)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1267
  apply (rule filterlim_cong[THEN iffD2, OF refl refl _ tendsto_const])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1268
  unfolding eventually_sequentially
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1269
  apply blast
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1270
  done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1271
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1272
text {* A monotone sequence converges to its least upper bound. *}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1273
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1274
lemma isLub_mono_imp_LIMSEQ:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1275
  fixes X :: "nat \<Rightarrow> real"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1276
  assumes u: "isLub UNIV {x. \<exists>n. X n = x} u" (* FIXME: use 'range X' *)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1277
  assumes X: "\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1278
  shows "X ----> u"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1279
proof (rule LIMSEQ_I)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1280
  have 1: "\<forall>n. X n \<le> u"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1281
    using isLubD2 [OF u] by auto
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1282
  have "\<forall>y. (\<forall>n. X n \<le> y) \<longrightarrow> u \<le> y"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1283
    using isLub_le_isUb [OF u] by (auto simp add: isUb_def setle_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1284
  hence 2: "\<forall>y<u. \<exists>n. y < X n"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1285
    by (metis not_le)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1286
  fix r :: real assume "0 < r"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1287
  hence "u - r < u" by simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1288
  hence "\<exists>m. u - r < X m" using 2 by simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1289
  then obtain m where "u - r < X m" ..
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1290
  with X have "\<forall>n\<ge>m. u - r < X n"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1291
    by (fast intro: less_le_trans)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1292
  hence "\<exists>m. \<forall>n\<ge>m. u - r < X n" ..
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1293
  thus "\<exists>m. \<forall>n\<ge>m. norm (X n - u) < r"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1294
    using 1 by (simp add: diff_less_eq add_commute)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1295
qed
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1296
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1297
text{*A standard proof of the theorem for monotone increasing sequence*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1298
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1299
lemma Bseq_mono_convergent:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1300
   "Bseq X \<Longrightarrow> \<forall>m. \<forall>n \<ge> m. X m \<le> X n \<Longrightarrow> convergent (X::nat=>real)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1301
  by (metis Bseq_isLub isLub_mono_imp_LIMSEQ convergentI)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1302
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1303
lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1304
  by (simp add: Bseq_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1305
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1306
text{*Main monotonicity theorem*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1307
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1308
lemma Bseq_monoseq_convergent: "Bseq X \<Longrightarrow> monoseq X \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1309
  by (metis monoseq_iff incseq_def decseq_eq_incseq convergent_minus_iff Bseq_minus_iff
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1310
            Bseq_mono_convergent)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1311
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1312
lemma Cauchy_iff:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1313
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1314
  shows "Cauchy X \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1315
  unfolding Cauchy_def dist_norm ..
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1316
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1317
lemma CauchyI:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1318
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1319
  shows "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1320
by (simp add: Cauchy_iff)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1321
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1322
lemma CauchyD:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1323
  fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1324
  shows "\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1325
by (simp add: Cauchy_iff)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1326
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1327
lemma Bseq_eq_bounded: "range f \<subseteq> {a .. b::real} \<Longrightarrow> Bseq f"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1328
  apply (simp add: subset_eq)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1329
  apply (rule BseqI'[where K="max (norm a) (norm b)"])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1330
  apply (erule_tac x=n in allE)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1331
  apply auto
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1332
  done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1333
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1334
lemma incseq_bounded: "incseq X \<Longrightarrow> \<forall>i. X i \<le> (B::real) \<Longrightarrow> Bseq X"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1335
  by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1336
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1337
lemma decseq_bounded: "decseq X \<Longrightarrow> \<forall>i. (B::real) \<le> X i \<Longrightarrow> Bseq X"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1338
  by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1339
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1340
lemma incseq_convergent:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1341
  fixes X :: "nat \<Rightarrow> real"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1342
  assumes "incseq X" and "\<forall>i. X i \<le> B"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1343
  obtains L where "X ----> L" "\<forall>i. X i \<le> L"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1344
proof atomize_elim
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1345
  from incseq_bounded[OF assms] `incseq X` Bseq_monoseq_convergent[of X]
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1346
  obtain L where "X ----> L"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1347
    by (auto simp: convergent_def monoseq_def incseq_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1348
  with `incseq X` show "\<exists>L. X ----> L \<and> (\<forall>i. X i \<le> L)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1349
    by (auto intro!: exI[of _ L] incseq_le)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1350
qed
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1351
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1352
lemma decseq_convergent:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1353
  fixes X :: "nat \<Rightarrow> real"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1354
  assumes "decseq X" and "\<forall>i. B \<le> X i"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1355
  obtains L where "X ----> L" "\<forall>i. L \<le> X i"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1356
proof atomize_elim
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1357
  from decseq_bounded[OF assms] `decseq X` Bseq_monoseq_convergent[of X]
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1358
  obtain L where "X ----> L"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1359
    by (auto simp: convergent_def monoseq_def decseq_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1360
  with `decseq X` show "\<exists>L. X ----> L \<and> (\<forall>i. L \<le> X i)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1361
    by (auto intro!: exI[of _ L] decseq_le)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1362
qed
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1363
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1364
subsubsection {* Cauchy Sequences are Bounded *}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1365
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1366
text{*A Cauchy sequence is bounded -- this is the standard
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1367
  proof mechanization rather than the nonstandard proof*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1368
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1369
lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1370
          ==>  \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1371
apply (clarify, drule spec, drule (1) mp)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1372
apply (simp only: norm_minus_commute)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1373
apply (drule order_le_less_trans [OF norm_triangle_ineq2])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1374
apply simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1375
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1376
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1377
class banach = real_normed_vector + complete_space
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1378
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1379
instance real :: banach by default
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1380
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1381
subsection {* Power Sequences *}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1382
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1383
text{*The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1384
"x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1385
  also fact that bounded and monotonic sequence converges.*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1386
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1387
lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1388
apply (simp add: Bseq_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1389
apply (rule_tac x = 1 in exI)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1390
apply (simp add: power_abs)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1391
apply (auto dest: power_mono)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1392
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1393
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1394
lemma monoseq_realpow: fixes x :: real shows "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1395
apply (clarify intro!: mono_SucI2)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1396
apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1397
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1398
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1399
lemma convergent_realpow:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1400
  "[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1401
by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1402
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1403
lemma LIMSEQ_inverse_realpow_zero: "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) ----> 0"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1404
  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
  1405
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1406
lemma LIMSEQ_realpow_zero:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1407
  "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1408
proof cases
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1409
  assume "0 \<le> x" and "x \<noteq> 0"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1410
  hence x0: "0 < x" by simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1411
  assume x1: "x < 1"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1412
  from x0 x1 have "1 < inverse x"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1413
    by (rule one_less_inverse)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1414
  hence "(\<lambda>n. inverse (inverse x ^ n)) ----> 0"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1415
    by (rule LIMSEQ_inverse_realpow_zero)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1416
  thus ?thesis by (simp add: power_inverse)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1417
qed (rule LIMSEQ_imp_Suc, simp add: tendsto_const)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1418
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1419
lemma LIMSEQ_power_zero:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1420
  fixes x :: "'a::{real_normed_algebra_1}"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1421
  shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1422
apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1423
apply (simp only: tendsto_Zfun_iff, erule Zfun_le)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1424
apply (simp add: power_abs norm_power_ineq)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1425
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1426
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1427
lemma LIMSEQ_divide_realpow_zero: "1 < x \<Longrightarrow> (\<lambda>n. a / (x ^ n) :: real) ----> 0"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1428
  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
  1429
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1430
text{*Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}*}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1431
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1432
lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. \<bar>c\<bar> ^ n :: real) ----> 0"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1433
  by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1434
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1435
lemma LIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. c ^ n :: real) ----> 0"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1436
  by (rule LIMSEQ_power_zero) simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1437
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1438
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1439
subsection {* Limits of Functions *}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1440
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1441
lemma LIM_eq:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1442
  fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1443
  shows "f -- a --> L =
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1444
     (\<forall>r>0.\<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1445
by (simp add: LIM_def dist_norm)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1446
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1447
lemma LIM_I:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1448
  fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1449
  shows "(!!r. 0<r ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1450
      ==> f -- a --> L"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1451
by (simp add: LIM_eq)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1452
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1453
lemma LIM_D:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1454
  fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1455
  shows "[| f -- a --> L; 0<r |]
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1456
      ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1457
by (simp add: LIM_eq)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1458
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1459
lemma LIM_offset:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1460
  fixes a :: "'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1461
  shows "f -- a --> L \<Longrightarrow> (\<lambda>x. f (x + k)) -- a - k --> L"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1462
apply (rule topological_tendstoI)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1463
apply (drule (2) topological_tendstoD)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1464
apply (simp only: eventually_at dist_norm)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1465
apply (clarify, rule_tac x=d in exI, safe)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1466
apply (drule_tac x="x + k" in spec)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1467
apply (simp add: algebra_simps)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1468
done
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1469
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1470
lemma LIM_offset_zero:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1471
  fixes a :: "'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1472
  shows "f -- a --> L \<Longrightarrow> (\<lambda>h. f (a + h)) -- 0 --> L"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1473
by (drule_tac k="a" in LIM_offset, simp add: add_commute)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1474
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1475
lemma LIM_offset_zero_cancel:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1476
  fixes a :: "'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1477
  shows "(\<lambda>h. f (a + h)) -- 0 --> L \<Longrightarrow> f -- a --> L"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1478
by (drule_tac k="- a" in LIM_offset, simp)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1479
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1480
lemma LIM_zero:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1481
  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
  1482
  shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. f x - l) ---> 0) F"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1483
unfolding tendsto_iff dist_norm by simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1484
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1485
lemma LIM_zero_cancel:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1486
  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
  1487
  shows "((\<lambda>x. f x - l) ---> 0) F \<Longrightarrow> (f ---> l) F"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1488
unfolding tendsto_iff dist_norm by simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1489
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1490
lemma LIM_zero_iff:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1491
  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1492
  shows "((\<lambda>x. f x - l) ---> 0) F = (f ---> l) F"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1493
unfolding tendsto_iff dist_norm by simp
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1494
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1495
lemma LIM_imp_LIM:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1496
  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
  1497
  fixes g :: "'a::topological_space \<Rightarrow> 'c::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1498
  assumes f: "f -- a --> l"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1499
  assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> norm (g x - m) \<le> norm (f x - l)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1500
  shows "g -- a --> m"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1501
  by (rule metric_LIM_imp_LIM [OF f],
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1502
    simp add: dist_norm le)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1503
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1504
lemma LIM_equal2:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1505
  fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1506
  assumes 1: "0 < R"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1507
  assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < R\<rbrakk> \<Longrightarrow> f x = g x"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1508
  shows "g -- a --> l \<Longrightarrow> f -- a --> l"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1509
by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1510
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1511
lemma LIM_compose2:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1512
  fixes a :: "'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1513
  assumes f: "f -- a --> b"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1514
  assumes g: "g -- b --> c"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1515
  assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1516
  shows "(\<lambda>x. g (f x)) -- a --> c"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1517
by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1518
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1519
lemma real_LIM_sandwich_zero:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1520
  fixes f g :: "'a::topological_space \<Rightarrow> real"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1521
  assumes f: "f -- a --> 0"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1522
  assumes 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1523
  assumes 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1524
  shows "g -- a --> 0"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1525
proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1526
  fix x assume x: "x \<noteq> a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1527
  have "norm (g x - 0) = g x" by (simp add: 1 x)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1528
  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
  1529
  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
  1530
  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
  1531
  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
  1532
qed
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1533
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1534
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1535
subsection {* Continuity *}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1536
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1537
lemma LIM_isCont_iff:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1538
  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1539
  shows "(f -- a --> f a) = ((\<lambda>h. f (a + h)) -- 0 --> f a)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1540
by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1541
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1542
lemma isCont_iff:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1543
  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1544
  shows "isCont f x = (\<lambda>h. f (x + h)) -- 0 --> f x"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1545
by (simp add: isCont_def LIM_isCont_iff)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1546
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1547
lemma isCont_LIM_compose2:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1548
  fixes a :: "'a::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1549
  assumes f [unfolded isCont_def]: "isCont f a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1550
  assumes g: "g -- f a --> l"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1551
  assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1552
  shows "(\<lambda>x. g (f x)) -- a --> l"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1553
by (rule LIM_compose2 [OF f g inj])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1554
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1555
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1556
lemma isCont_norm [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1557
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1558
  shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (f x)) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1559
  by (fact continuous_norm)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1560
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1561
lemma isCont_rabs [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1562
  fixes f :: "'a::t2_space \<Rightarrow> real"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1563
  shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1564
  by (fact continuous_rabs)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1565
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1566
lemma isCont_add [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1567
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1568
  shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x + g x) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1569
  by (fact continuous_add)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1570
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1571
lemma isCont_minus [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1572
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1573
  shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. - f x) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1574
  by (fact continuous_minus)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1575
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1576
lemma isCont_diff [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1577
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1578
  shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x - g x) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1579
  by (fact continuous_diff)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1580
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1581
lemma isCont_mult [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1582
  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1583
  shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x * g x) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1584
  by (fact continuous_mult)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1585
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1586
lemma (in bounded_linear) isCont:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1587
  "isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1588
  by (fact continuous)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1589
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1590
lemma (in bounded_bilinear) isCont:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1591
  "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1592
  by (fact continuous)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1593
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1594
lemmas isCont_scaleR [simp] = 
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1595
  bounded_bilinear.isCont [OF bounded_bilinear_scaleR]
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1596
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1597
lemmas isCont_of_real [simp] =
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1598
  bounded_linear.isCont [OF bounded_linear_of_real]
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1599
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1600
lemma isCont_power [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1601
  fixes f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1602
  shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1603
  by (fact continuous_power)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1604
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1605
lemma isCont_setsum [simp]:
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1606
  fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::real_normed_vector"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1607
  shows "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1608
  by (auto intro: continuous_setsum)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1609
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1610
lemmas isCont_intros =
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1611
  isCont_ident isCont_const isCont_norm isCont_rabs isCont_add isCont_minus
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1612
  isCont_diff isCont_mult isCont_inverse isCont_divide isCont_scaleR
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1613
  isCont_of_real isCont_power isCont_sgn isCont_setsum
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1614
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1615
subsection {* Uniform Continuity *}
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1616
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1617
lemma (in bounded_linear) isUCont: "isUCont f"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1618
unfolding isUCont_def dist_norm
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1619
proof (intro allI impI)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1620
  fix r::real assume r: "0 < r"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1621
  obtain K where K: "0 < K" and norm_le: "\<And>x. norm (f x) \<le> norm x * K"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1622
    using pos_bounded by fast
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1623
  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
  1624
  proof (rule exI, safe)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1625
    from r K show "0 < r / K" by (rule divide_pos_pos)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1626
  next
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1627
    fix x y :: 'a
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1628
    assume xy: "norm (x - y) < r / K"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1629
    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
  1630
    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
  1631
    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
  1632
    finally show "norm (f x - f y) < r" .
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1633
  qed
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1634
qed
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1635
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1636
lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1637
by (rule isUCont [THEN isUCont_Cauchy])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1638
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1639
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1640
lemma LIM_less_bound: 
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1641
  fixes f :: "real \<Rightarrow> real"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1642
  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
  1643
  shows "0 \<le> f x"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1644
proof (rule tendsto_le_const)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1645
  show "(f ---> f x) (at_left x)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1646
    using `isCont f x` by (simp add: filterlim_at_split isCont_def)
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1647
  show "eventually (\<lambda>x. 0 \<le> f x) (at_left x)"
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1648
    using ev by (auto simp: eventually_within_less dist_real_def intro!: exI[of _ "x - b"])
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51524
diff changeset
  1649
qed simp
51471
cad22a3cc09c move topological_space to its own theory
hoelzl
parents: 51360
diff changeset
  1650
31349
2261c8781f73 new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff changeset
  1651
end
50324
0a1242d5e7d4 add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents: 50323
diff changeset
  1652