| author | haftmann |
| Sat, 23 Mar 2013 17:11:06 +0100 | |
| changeset 51487 | f4bfdee99304 |
| parent 51478 | 270b21f3ae0a |
| permissions | -rw-r--r-- |
| 10751 | 1 |
(* Title : Lim.thy |
2 |
Author : Jacques D. Fleuriot |
|
3 |
Copyright : 1998 University of Cambridge |
|
| 14477 | 4 |
Conversion to Isar and new proofs by Lawrence C Paulson, 2004 |
| 10751 | 5 |
*) |
6 |
||
|
21165
8fb49f668511
moved DERIV stuff from Lim.thy to new Deriv.thy; cleaned up LIMSEQ_SEQ proofs
huffman
parents:
21141
diff
changeset
|
7 |
header{* Limits and Continuity *}
|
| 10751 | 8 |
|
| 15131 | 9 |
theory Lim |
|
29197
6d4cb27ed19c
adapted HOL source structure to distribution layout
haftmann
parents:
28952
diff
changeset
|
10 |
imports SEQ |
| 15131 | 11 |
begin |
| 14477 | 12 |
|
| 31392 | 13 |
subsection {* Limits of Functions *}
|
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
31338
diff
changeset
|
14 |
|
| 14477 | 15 |
lemma LIM_eq: |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
16 |
fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" |
|
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
17 |
shows "f -- a --> L = |
|
20561
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
huffman
parents:
20552
diff
changeset
|
18 |
(\<forall>r>0.\<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r)" |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
19 |
by (simp add: LIM_def dist_norm) |
| 14477 | 20 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20432
diff
changeset
|
21 |
lemma LIM_I: |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
22 |
fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" |
|
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
23 |
shows "(!!r. 0<r ==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r) |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20432
diff
changeset
|
24 |
==> f -- a --> L" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20432
diff
changeset
|
25 |
by (simp add: LIM_eq) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20432
diff
changeset
|
26 |
|
| 14477 | 27 |
lemma LIM_D: |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
28 |
fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" |
|
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
29 |
shows "[| f -- a --> L; 0<r |] |
|
20561
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
huffman
parents:
20552
diff
changeset
|
30 |
==> \<exists>s>0.\<forall>x. x \<noteq> a & norm (x-a) < s --> norm (f x - L) < r" |
| 14477 | 31 |
by (simp add: LIM_eq) |
32 |
||
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
33 |
lemma LIM_offset: |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
34 |
fixes a :: "'a::real_normed_vector" |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
35 |
shows "f -- a --> L \<Longrightarrow> (\<lambda>x. f (x + k)) -- a - k --> L" |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
36 |
apply (rule topological_tendstoI) |
|
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
37 |
apply (drule (2) topological_tendstoD) |
|
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
38 |
apply (simp only: eventually_at dist_norm) |
|
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
39 |
apply (clarify, rule_tac x=d in exI, safe) |
| 20756 | 40 |
apply (drule_tac x="x + k" in spec) |
| 29667 | 41 |
apply (simp add: algebra_simps) |
| 20756 | 42 |
done |
43 |
||
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
44 |
lemma LIM_offset_zero: |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
45 |
fixes a :: "'a::real_normed_vector" |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
46 |
shows "f -- a --> L \<Longrightarrow> (\<lambda>h. f (a + h)) -- 0 --> L" |
| 21239 | 47 |
by (drule_tac k="a" in LIM_offset, simp add: add_commute) |
48 |
||
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
49 |
lemma LIM_offset_zero_cancel: |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
50 |
fixes a :: "'a::real_normed_vector" |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
51 |
shows "(\<lambda>h. f (a + h)) -- 0 --> L \<Longrightarrow> f -- a --> L" |
| 21239 | 52 |
by (drule_tac k="- a" in LIM_offset, simp) |
53 |
||
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
54 |
lemma LIM_zero: |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
55 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" |
| 44570 | 56 |
shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. f x - l) ---> 0) F" |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
57 |
unfolding tendsto_iff dist_norm by simp |
| 21239 | 58 |
|
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
59 |
lemma LIM_zero_cancel: |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
60 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" |
| 44570 | 61 |
shows "((\<lambda>x. f x - l) ---> 0) F \<Longrightarrow> (f ---> l) F" |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
62 |
unfolding tendsto_iff dist_norm by simp |
| 21239 | 63 |
|
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
64 |
lemma LIM_zero_iff: |
|
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
65 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" |
| 44570 | 66 |
shows "((\<lambda>x. f x - l) ---> 0) F = (f ---> l) F" |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
67 |
unfolding tendsto_iff dist_norm by simp |
| 21399 | 68 |
|
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
69 |
lemma LIM_imp_LIM: |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
70 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" |
|
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
71 |
fixes g :: "'a::topological_space \<Rightarrow> 'c::real_normed_vector" |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
72 |
assumes f: "f -- a --> l" |
|
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
73 |
assumes le: "\<And>x. x \<noteq> a \<Longrightarrow> norm (g x - m) \<le> norm (f x - l)" |
|
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
74 |
shows "g -- a --> m" |
| 44251 | 75 |
by (rule metric_LIM_imp_LIM [OF f], |
76 |
simp add: dist_norm le) |
|
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
77 |
|
|
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
78 |
lemma LIM_equal2: |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
79 |
fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space" |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
80 |
assumes 1: "0 < R" |
|
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
81 |
assumes 2: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < R\<rbrakk> \<Longrightarrow> f x = g x" |
|
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
82 |
shows "g -- a --> l \<Longrightarrow> f -- a --> l" |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
83 |
by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm) |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
84 |
|
| 23040 | 85 |
lemma LIM_compose2: |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
86 |
fixes a :: "'a::real_normed_vector" |
| 23040 | 87 |
assumes f: "f -- a --> b" |
88 |
assumes g: "g -- b --> c" |
|
89 |
assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> b" |
|
90 |
shows "(\<lambda>x. g (f x)) -- a --> c" |
|
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
91 |
by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]]) |
| 23040 | 92 |
|
|
21282
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
93 |
lemma real_LIM_sandwich_zero: |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36661
diff
changeset
|
94 |
fixes f g :: "'a::topological_space \<Rightarrow> real" |
|
21282
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
95 |
assumes f: "f -- a --> 0" |
|
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
96 |
assumes 1: "\<And>x. x \<noteq> a \<Longrightarrow> 0 \<le> g x" |
|
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
97 |
assumes 2: "\<And>x. x \<noteq> a \<Longrightarrow> g x \<le> f x" |
|
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
98 |
shows "g -- a --> 0" |
| 51471 | 99 |
proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *) |
|
21282
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
100 |
fix x assume x: "x \<noteq> a" |
|
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
101 |
have "norm (g x - 0) = g x" by (simp add: 1 x) |
|
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
102 |
also have "g x \<le> f x" by (rule 2 [OF x]) |
|
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
103 |
also have "f x \<le> \<bar>f x\<bar>" by (rule abs_ge_self) |
|
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
104 |
also have "\<bar>f x\<bar> = norm (f x - 0)" by simp |
|
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
105 |
finally show "norm (g x - 0) \<le> norm (f x - 0)" . |
|
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
106 |
qed |
|
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
107 |
|
| 14477 | 108 |
|
| 20755 | 109 |
subsection {* Continuity *}
|
| 14477 | 110 |
|
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
111 |
lemma LIM_isCont_iff: |
| 36665 | 112 |
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space" |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
113 |
shows "(f -- a --> f a) = ((\<lambda>h. f (a + h)) -- 0 --> f a)" |
| 21239 | 114 |
by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel]) |
115 |
||
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
116 |
lemma isCont_iff: |
| 36665 | 117 |
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space" |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
118 |
shows "isCont f x = (\<lambda>h. f (x + h)) -- 0 --> f x" |
| 21239 | 119 |
by (simp add: isCont_def LIM_isCont_iff) |
120 |
||
| 23040 | 121 |
lemma isCont_LIM_compose2: |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
122 |
fixes a :: "'a::real_normed_vector" |
| 23040 | 123 |
assumes f [unfolded isCont_def]: "isCont f a" |
124 |
assumes g: "g -- f a --> l" |
|
125 |
assumes inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> f x \<noteq> f a" |
|
126 |
shows "(\<lambda>x. g (f x)) -- a --> l" |
|
127 |
by (rule LIM_compose2 [OF f g inj]) |
|
128 |
||
|
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:
51473
diff
changeset
|
129 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51473
diff
changeset
|
130 |
lemma isCont_norm [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:
51473
diff
changeset
|
131 |
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:
51473
diff
changeset
|
132 |
shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. norm (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:
51473
diff
changeset
|
133 |
by (fact 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:
51473
diff
changeset
|
134 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51473
diff
changeset
|
135 |
lemma isCont_rabs [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:
51473
diff
changeset
|
136 |
fixes f :: "'a::t2_space \<Rightarrow> 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:
51473
diff
changeset
|
137 |
shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) 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:
51473
diff
changeset
|
138 |
by (fact continuous_rabs) |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51473
diff
changeset
|
139 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51473
diff
changeset
|
140 |
lemma isCont_add [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:
51473
diff
changeset
|
141 |
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:
51473
diff
changeset
|
142 |
shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> 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:
51473
diff
changeset
|
143 |
by (fact continuous_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:
51473
diff
changeset
|
144 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51473
diff
changeset
|
145 |
lemma isCont_minus [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:
51473
diff
changeset
|
146 |
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:
51473
diff
changeset
|
147 |
shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. - 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:
51473
diff
changeset
|
148 |
by (fact continuous_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:
51473
diff
changeset
|
149 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51473
diff
changeset
|
150 |
lemma isCont_diff [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:
51473
diff
changeset
|
151 |
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:
51473
diff
changeset
|
152 |
shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> 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:
51473
diff
changeset
|
153 |
by (fact continuous_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:
51473
diff
changeset
|
154 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51473
diff
changeset
|
155 |
lemma isCont_mult [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:
51473
diff
changeset
|
156 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::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:
51473
diff
changeset
|
157 |
shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> 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:
51473
diff
changeset
|
158 |
by (fact continuous_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:
51473
diff
changeset
|
159 |
|
| 44233 | 160 |
lemma (in bounded_linear) isCont: |
161 |
"isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a" |
|
|
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:
51473
diff
changeset
|
162 |
by (fact continuous) |
|
21282
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
163 |
|
|
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
164 |
lemma (in bounded_bilinear) isCont: |
|
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
165 |
"\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x ** g x) a" |
|
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:
51473
diff
changeset
|
166 |
by (fact continuous) |
|
21282
dd647b4d7952
added bounded_linear and bounded_bilinear locales
huffman
parents:
21257
diff
changeset
|
167 |
|
|
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:
51473
diff
changeset
|
168 |
lemmas isCont_scaleR [simp] = |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44254
diff
changeset
|
169 |
bounded_bilinear.isCont [OF bounded_bilinear_scaleR] |
| 21239 | 170 |
|
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44254
diff
changeset
|
171 |
lemmas isCont_of_real [simp] = |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44254
diff
changeset
|
172 |
bounded_linear.isCont [OF bounded_linear_of_real] |
|
22627
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
huffman
parents:
22613
diff
changeset
|
173 |
|
| 44233 | 174 |
lemma isCont_power [simp]: |
|
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:
51473
diff
changeset
|
175 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
|
|
22627
2b093ba973bc
new LIM/isCont lemmas for abs, of_real, and power
huffman
parents:
22613
diff
changeset
|
176 |
shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x ^ n) a" |
|
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:
51473
diff
changeset
|
177 |
by (fact continuous_power) |
| 29885 | 178 |
|
| 44233 | 179 |
lemma isCont_setsum [simp]: |
|
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:
51473
diff
changeset
|
180 |
fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::real_normed_vector" |
| 44233 | 181 |
shows "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a" |
|
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:
51473
diff
changeset
|
182 |
by (auto intro: continuous_setsum) |
| 15228 | 183 |
|
| 44233 | 184 |
lemmas isCont_intros = |
185 |
isCont_ident isCont_const isCont_norm isCont_rabs isCont_add isCont_minus |
|
186 |
isCont_diff isCont_mult isCont_inverse isCont_divide isCont_scaleR |
|
187 |
isCont_of_real isCont_power isCont_sgn isCont_setsum |
|
|
29803
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents:
29667
diff
changeset
|
188 |
|
| 20755 | 189 |
subsection {* Uniform Continuity *}
|
190 |
||
| 23118 | 191 |
lemma (in bounded_linear) isUCont: "isUCont f" |
|
31338
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
huffman
parents:
31336
diff
changeset
|
192 |
unfolding isUCont_def dist_norm |
| 23118 | 193 |
proof (intro allI impI) |
194 |
fix r::real assume r: "0 < r" |
|
195 |
obtain K where K: "0 < K" and norm_le: "\<And>x. norm (f x) \<le> norm x * K" |
|
196 |
using pos_bounded by fast |
|
197 |
show "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r" |
|
198 |
proof (rule exI, safe) |
|
199 |
from r K show "0 < r / K" by (rule divide_pos_pos) |
|
200 |
next |
|
201 |
fix x y :: 'a |
|
202 |
assume xy: "norm (x - y) < r / K" |
|
203 |
have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff) |
|
204 |
also have "\<dots> \<le> norm (x - y) * K" by (rule norm_le) |
|
205 |
also from K xy have "\<dots> < r" by (simp only: pos_less_divide_eq) |
|
206 |
finally show "norm (f x - f y) < r" . |
|
207 |
qed |
|
208 |
qed |
|
209 |
||
210 |
lemma (in bounded_linear) Cauchy: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. f (X n))" |
|
211 |
by (rule isUCont [THEN isUCont_Cauchy]) |
|
212 |
||
| 14477 | 213 |
|
| 50331 | 214 |
lemma LIM_less_bound: |
215 |
fixes f :: "real \<Rightarrow> real" |
|
216 |
assumes ev: "b < x" "\<forall> x' \<in> { b <..< x}. 0 \<le> f x'" and "isCont f x"
|
|
217 |
shows "0 \<le> f x" |
|
218 |
proof (rule tendsto_le_const) |
|
219 |
show "(f ---> f x) (at_left x)" |
|
220 |
using `isCont f x` by (simp add: filterlim_at_split isCont_def) |
|
221 |
show "eventually (\<lambda>x. 0 \<le> f x) (at_left x)" |
|
222 |
using ev by (auto simp: eventually_within_less dist_real_def intro!: exI[of _ "x - b"]) |
|
223 |
qed simp |
|
224 |
||
| 10751 | 225 |
end |