author | wenzelm |
Wed, 23 Jul 2014 21:01:28 +0200 | |
changeset 57625 | 2a9d8dcea893 |
parent 57512 | cc97b347b301 |
child 58729 | e8ecc79aee43 |
permissions | -rw-r--r-- |
52265 | 1 |
(* Title: HOL/Limits.thy |
51526 | 2 |
Author: Brian Huffman |
3 |
Author: Jacques D. Fleuriot, University of Cambridge |
|
4 |
Author: Lawrence C Paulson |
|
5 |
Author: Jeremy Avigad |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
6 |
*) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
7 |
|
51526 | 8 |
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
|
9 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
10 |
theory Limits |
51524 | 11 |
imports Real_Vector_Spaces |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
12 |
begin |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
13 |
|
51526 | 14 |
subsection {* Filter going to infinity norm *} |
15 |
||
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
16 |
definition at_infinity :: "'a::real_normed_vector filter" where |
57276 | 17 |
"at_infinity = (INF r. principal {x. r \<le> norm x})" |
50324
0a1242d5e7d4
add filterlim rules for diverging multiplication and addition; move at_infinity to the HOL image
hoelzl
parents:
50323
diff
changeset
|
18 |
|
57276 | 19 |
lemma eventually_at_infinity: "eventually P at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. b \<le> norm x \<longrightarrow> P x)" |
20 |
unfolding at_infinity_def |
|
21 |
by (subst eventually_INF_base) |
|
22 |
(auto simp: subset_eq eventually_principal intro!: exI[of _ "max a b" for a b]) |
|
31392 | 23 |
|
50325 | 24 |
lemma at_infinity_eq_at_top_bot: |
25 |
"(at_infinity \<Colon> real filter) = sup at_top at_bot" |
|
57276 | 26 |
apply (simp add: filter_eq_iff eventually_sup eventually_at_infinity |
27 |
eventually_at_top_linorder eventually_at_bot_linorder) |
|
28 |
apply safe |
|
29 |
apply (rule_tac x="b" in exI, simp) |
|
30 |
apply (rule_tac x="- b" in exI, simp) |
|
31 |
apply (rule_tac x="max (- Na) N" in exI, auto simp: abs_real_def) |
|
32 |
done |
|
50325 | 33 |
|
57276 | 34 |
lemma at_top_le_at_infinity: "at_top \<le> (at_infinity :: real filter)" |
50325 | 35 |
unfolding at_infinity_eq_at_top_bot by simp |
36 |
||
57276 | 37 |
lemma at_bot_le_at_infinity: "at_bot \<le> (at_infinity :: real filter)" |
50325 | 38 |
unfolding at_infinity_eq_at_top_bot by simp |
39 |
||
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
56541
diff
changeset
|
40 |
lemma filterlim_at_top_imp_at_infinity: |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
56541
diff
changeset
|
41 |
fixes f :: "_ \<Rightarrow> real" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
56541
diff
changeset
|
42 |
shows "filterlim f at_top F \<Longrightarrow> filterlim f at_infinity F" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
56541
diff
changeset
|
43 |
by (rule filterlim_mono[OF _ at_top_le_at_infinity order_refl]) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
56541
diff
changeset
|
44 |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
45 |
subsubsection {* Boundedness *} |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
46 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
47 |
definition Bfun :: "('a \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a filter \<Rightarrow> bool" where |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
48 |
Bfun_metric_def: "Bfun f F = (\<exists>y. \<exists>K>0. eventually (\<lambda>x. dist (f x) y \<le> K) F)" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
49 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
50 |
abbreviation Bseq :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
51 |
"Bseq X \<equiv> Bfun X sequentially" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
52 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
53 |
lemma Bseq_conv_Bfun: "Bseq X \<longleftrightarrow> Bfun X sequentially" .. |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
54 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
55 |
lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
56 |
unfolding Bfun_metric_def by (subst eventually_sequentially_seg) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
57 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
58 |
lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
59 |
unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg) |
31355 | 60 |
|
51474
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
61 |
lemma Bfun_def: |
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
62 |
"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
|
63 |
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
|
64 |
proof safe |
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
65 |
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
|
66 |
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
|
67 |
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
|
68 |
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
|
69 |
by eventually_elim auto |
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
70 |
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
|
71 |
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
|
72 |
qed auto |
31355 | 73 |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
74 |
lemma BfunI: |
44195 | 75 |
assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) F" shows "Bfun f F" |
31355 | 76 |
unfolding Bfun_def |
77 |
proof (intro exI conjI allI) |
|
78 |
show "0 < max K 1" by simp |
|
79 |
next |
|
44195 | 80 |
show "eventually (\<lambda>x. norm (f x) \<le> max K 1) F" |
31355 | 81 |
using K by (rule eventually_elim1, simp) |
82 |
qed |
|
83 |
||
84 |
lemma BfunE: |
|
44195 | 85 |
assumes "Bfun f F" |
86 |
obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) F" |
|
31355 | 87 |
using assms unfolding Bfun_def by fast |
88 |
||
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
89 |
lemma Cauchy_Bseq: "Cauchy X \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
90 |
unfolding Cauchy_def Bfun_metric_def eventually_sequentially |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
91 |
apply (erule_tac x=1 in allE) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
92 |
apply simp |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
93 |
apply safe |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
94 |
apply (rule_tac x="X M" in exI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
95 |
apply (rule_tac x=1 in exI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
96 |
apply (erule_tac x=M in allE) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
97 |
apply simp |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
98 |
apply (rule_tac x=M in exI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
99 |
apply (auto simp: dist_commute) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
100 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
101 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
102 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
103 |
subsubsection {* Bounded Sequences *} |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
104 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
105 |
lemma BseqI': "(\<And>n. norm (X n) \<le> K) \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
106 |
by (intro BfunI) (auto simp: eventually_sequentially) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
107 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
108 |
lemma BseqI2': "\<forall>n\<ge>N. norm (X n) \<le> K \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
109 |
by (intro BfunI) (auto simp: eventually_sequentially) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
110 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
111 |
lemma Bseq_def: "Bseq X \<longleftrightarrow> (\<exists>K>0. \<forall>n. norm (X n) \<le> K)" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
112 |
unfolding Bfun_def eventually_sequentially |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
113 |
proof safe |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
114 |
fix N K assume "0 < K" "\<forall>n\<ge>N. norm (X n) \<le> K" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
115 |
then show "\<exists>K>0. \<forall>n. norm (X n) \<le> K" |
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54263
diff
changeset
|
116 |
by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] max.strict_coboundedI2) |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
117 |
(auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
118 |
qed auto |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
119 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
120 |
lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
121 |
unfolding Bseq_def by auto |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
122 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
123 |
lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
124 |
by (simp add: Bseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
125 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
126 |
lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
127 |
by (auto simp add: Bseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
128 |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
129 |
lemma Bseq_bdd_above: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_above (range X)" |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
130 |
proof (elim BseqE, intro bdd_aboveI2) |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
131 |
fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "X n \<le> K" |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
132 |
by (auto elim!: allE[of _ n]) |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
133 |
qed |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
134 |
|
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
135 |
lemma Bseq_bdd_below: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_below (range X)" |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
136 |
proof (elim BseqE, intro bdd_belowI2) |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
137 |
fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "- K \<le> X n" |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
138 |
by (auto elim!: allE[of _ n]) |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
139 |
qed |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
140 |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
141 |
lemma lemma_NBseq_def: |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
142 |
"(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
143 |
proof safe |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
144 |
fix K :: real |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
145 |
from reals_Archimedean2 obtain n :: nat where "K < real n" .. |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
146 |
then have "K \<le> real (Suc n)" by auto |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
147 |
moreover assume "\<forall>m. norm (X m) \<le> K" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
148 |
ultimately have "\<forall>m. norm (X m) \<le> real (Suc n)" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
149 |
by (blast intro: order_trans) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
150 |
then show "\<exists>N. \<forall>n. norm (X n) \<le> real (Suc N)" .. |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
151 |
qed (force simp add: real_of_nat_Suc) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
152 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
153 |
text{* alternative definition for Bseq *} |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
154 |
lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
155 |
apply (simp add: Bseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
156 |
apply (simp (no_asm) add: lemma_NBseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
157 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
158 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
159 |
lemma lemma_NBseq_def2: |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
160 |
"(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
161 |
apply (subst lemma_NBseq_def, auto) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
162 |
apply (rule_tac x = "Suc N" in exI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
163 |
apply (rule_tac [2] x = N in exI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
164 |
apply (auto simp add: real_of_nat_Suc) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
165 |
prefer 2 apply (blast intro: order_less_imp_le) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
166 |
apply (drule_tac x = n in spec, simp) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
167 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
168 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
169 |
(* yet another definition for Bseq *) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
170 |
lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
171 |
by (simp add: Bseq_def lemma_NBseq_def2) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
172 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
173 |
subsubsection{*A Few More Equivalence Theorems for Boundedness*} |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
174 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
175 |
text{*alternative formulation for boundedness*} |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
176 |
lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
177 |
apply (unfold Bseq_def, safe) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
178 |
apply (rule_tac [2] x = "k + norm x" in exI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
179 |
apply (rule_tac x = K in exI, simp) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
180 |
apply (rule exI [where x = 0], auto) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
181 |
apply (erule order_less_le_trans, simp) |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
182 |
apply (drule_tac x=n in spec) |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
183 |
apply (drule order_trans [OF norm_triangle_ineq2]) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
184 |
apply simp |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
185 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
186 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
187 |
text{*alternative formulation for boundedness*} |
53602 | 188 |
lemma Bseq_iff3: |
189 |
"Bseq X \<longleftrightarrow> (\<exists>k>0. \<exists>N. \<forall>n. norm (X n + - X N) \<le> k)" (is "?P \<longleftrightarrow> ?Q") |
|
190 |
proof |
|
191 |
assume ?P |
|
192 |
then obtain K |
|
193 |
where *: "0 < K" and **: "\<And>n. norm (X n) \<le> K" by (auto simp add: Bseq_def) |
|
194 |
from * have "0 < K + norm (X 0)" by (rule order_less_le_trans) simp |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
195 |
from ** have "\<forall>n. norm (X n - X 0) \<le> K + norm (X 0)" |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
196 |
by (auto intro: order_trans norm_triangle_ineq4) |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
197 |
then have "\<forall>n. norm (X n + - X 0) \<le> K + norm (X 0)" |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
198 |
by simp |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
199 |
with `0 < K + norm (X 0)` show ?Q by blast |
53602 | 200 |
next |
201 |
assume ?Q then show ?P by (auto simp add: Bseq_iff2) |
|
202 |
qed |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
203 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
204 |
lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
205 |
apply (simp add: Bseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
206 |
apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
207 |
apply (drule_tac x = n in spec, arith) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
208 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
209 |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
210 |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
211 |
subsubsection{*Upper Bounds and Lubs of Bounded Sequences*} |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
212 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
213 |
lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
214 |
by (simp add: Bseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
215 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
216 |
lemma Bseq_eq_bounded: "range f \<subseteq> {a .. b::real} \<Longrightarrow> Bseq f" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
217 |
apply (simp add: subset_eq) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
218 |
apply (rule BseqI'[where K="max (norm a) (norm b)"]) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
219 |
apply (erule_tac x=n in allE) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
220 |
apply auto |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
221 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
222 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
223 |
lemma incseq_bounded: "incseq X \<Longrightarrow> \<forall>i. X i \<le> (B::real) \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
224 |
by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
225 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
226 |
lemma decseq_bounded: "decseq X \<Longrightarrow> \<forall>i. (B::real) \<le> X i \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
227 |
by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
228 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
229 |
subsection {* Bounded Monotonic Sequences *} |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
230 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
231 |
subsubsection{*A Bounded and Monotonic Sequence Converges*} |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
232 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
233 |
(* TODO: delete *) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
234 |
(* FIXME: one use in NSA/HSEQ.thy *) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
235 |
lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
236 |
apply (rule_tac x="X m" in exI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
237 |
apply (rule filterlim_cong[THEN iffD2, OF refl refl _ tendsto_const]) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
238 |
unfolding eventually_sequentially |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
239 |
apply blast |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
240 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
241 |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
242 |
subsection {* Convergence to Zero *} |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
243 |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
244 |
definition Zfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool" |
44195 | 245 |
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
|
246 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
247 |
lemma ZfunI: |
44195 | 248 |
"(\<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
|
249 |
unfolding Zfun_def by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
250 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
251 |
lemma ZfunD: |
44195 | 252 |
"\<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
|
253 |
unfolding Zfun_def by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
254 |
|
31355 | 255 |
lemma Zfun_ssubst: |
44195 | 256 |
"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
|
257 |
unfolding Zfun_def by (auto elim!: eventually_rev_mp) |
31355 | 258 |
|
44195 | 259 |
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
|
260 |
unfolding Zfun_def by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
261 |
|
44195 | 262 |
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
|
263 |
unfolding Zfun_def by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
264 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
265 |
lemma Zfun_imp_Zfun: |
44195 | 266 |
assumes f: "Zfun f F" |
267 |
assumes g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F" |
|
268 |
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
|
269 |
proof (cases) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
270 |
assume K: "0 < K" |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
271 |
show ?thesis |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
272 |
proof (rule ZfunI) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
273 |
fix r::real assume "0 < r" |
56541 | 274 |
hence "0 < r / K" using K by simp |
44195 | 275 |
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
|
276 |
using ZfunD [OF f] by fast |
44195 | 277 |
with g show "eventually (\<lambda>x. norm (g x) < r) F" |
46887 | 278 |
proof eventually_elim |
279 |
case (elim x) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
280 |
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
|
281 |
by (simp add: pos_less_divide_eq K) |
46887 | 282 |
thus ?case |
283 |
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
|
284 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
285 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
286 |
next |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
287 |
assume "\<not> 0 < K" |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
288 |
hence K: "K \<le> 0" by (simp only: not_less) |
31355 | 289 |
show ?thesis |
290 |
proof (rule ZfunI) |
|
291 |
fix r :: real |
|
292 |
assume "0 < r" |
|
44195 | 293 |
from g show "eventually (\<lambda>x. norm (g x) < r) F" |
46887 | 294 |
proof eventually_elim |
295 |
case (elim x) |
|
296 |
also have "norm (f x) * K \<le> norm (f x) * 0" |
|
31355 | 297 |
using K norm_ge_zero by (rule mult_left_mono) |
46887 | 298 |
finally show ?case |
31355 | 299 |
using `0 < r` by simp |
300 |
qed |
|
301 |
qed |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
302 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
303 |
|
44195 | 304 |
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
|
305 |
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
|
306 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
307 |
lemma Zfun_add: |
44195 | 308 |
assumes f: "Zfun f F" and g: "Zfun g F" |
309 |
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
|
310 |
proof (rule ZfunI) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
311 |
fix r::real assume "0 < r" |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
312 |
hence r: "0 < r / 2" by simp |
44195 | 313 |
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
|
314 |
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
|
315 |
moreover |
44195 | 316 |
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
|
317 |
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
|
318 |
ultimately |
44195 | 319 |
show "eventually (\<lambda>x. norm (f x + g x) < r) F" |
46887 | 320 |
proof eventually_elim |
321 |
case (elim x) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
322 |
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
|
323 |
by (rule norm_triangle_ineq) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
324 |
also have "\<dots> < r/2 + r/2" |
46887 | 325 |
using elim by (rule add_strict_mono) |
326 |
finally show ?case |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
327 |
by simp |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
328 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
329 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
330 |
|
44195 | 331 |
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
|
332 |
unfolding Zfun_def by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
333 |
|
44195 | 334 |
lemma Zfun_diff: "\<lbrakk>Zfun f F; Zfun g F\<rbrakk> \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F" |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
335 |
using Zfun_add [of f F "\<lambda>x. - g x"] by (simp add: Zfun_minus) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
336 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
337 |
lemma (in bounded_linear) Zfun: |
44195 | 338 |
assumes g: "Zfun g F" |
339 |
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
|
340 |
proof - |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
341 |
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
|
342 |
using bounded by fast |
44195 | 343 |
then have "eventually (\<lambda>x. norm (f (g x)) \<le> norm (g x) * K) F" |
31355 | 344 |
by simp |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
345 |
with g show ?thesis |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
346 |
by (rule Zfun_imp_Zfun) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
347 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
348 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
349 |
lemma (in bounded_bilinear) Zfun: |
44195 | 350 |
assumes f: "Zfun f F" |
351 |
assumes g: "Zfun g F" |
|
352 |
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
|
353 |
proof (rule ZfunI) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
354 |
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
|
355 |
obtain K where K: "0 < K" |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
356 |
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
|
357 |
using pos_bounded by fast |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
358 |
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
|
359 |
by (rule positive_imp_inverse_positive) |
44195 | 360 |
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
|
361 |
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
|
362 |
moreover |
44195 | 363 |
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
|
364 |
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
|
365 |
ultimately |
44195 | 366 |
show "eventually (\<lambda>x. norm (f x ** g x) < r) F" |
46887 | 367 |
proof eventually_elim |
368 |
case (elim x) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
369 |
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
|
370 |
by (rule norm_le) |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
371 |
also have "norm (f x) * norm (g x) * K < r * inverse K * K" |
46887 | 372 |
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
|
373 |
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
|
374 |
by simp |
46887 | 375 |
finally show ?case . |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
376 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
377 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
378 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
379 |
lemma (in bounded_bilinear) Zfun_left: |
44195 | 380 |
"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
|
381 |
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
|
382 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
383 |
lemma (in bounded_bilinear) Zfun_right: |
44195 | 384 |
"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
|
385 |
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
|
386 |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
387 |
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
|
388 |
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
|
389 |
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
|
390 |
|
44195 | 391 |
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
|
392 |
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
|
393 |
|
56366 | 394 |
lemma tendsto_0_le: "\<lbrakk>(f ---> 0) F; eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F\<rbrakk> |
395 |
\<Longrightarrow> (g ---> 0) F" |
|
396 |
by (simp add: Zfun_imp_Zfun tendsto_Zfun_iff) |
|
397 |
||
44205
18da2a87421c
generalize constant 'lim' and limit uniqueness theorems to class t2_space
huffman
parents:
44195
diff
changeset
|
398 |
subsubsection {* Distance and norms *} |
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
399 |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
400 |
lemma tendsto_dist [tendsto_intros]: |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
401 |
fixes l m :: "'a :: metric_space" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
402 |
assumes f: "(f ---> l) F" and g: "(g ---> m) F" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
403 |
shows "((\<lambda>x. dist (f x) (g x)) ---> dist l m) F" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
404 |
proof (rule tendstoI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
405 |
fix e :: real assume "0 < e" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
406 |
hence e2: "0 < e/2" by simp |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
407 |
from tendstoD [OF f e2] tendstoD [OF g e2] |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
408 |
show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) F" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
409 |
proof (eventually_elim) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
410 |
case (elim x) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
411 |
then show "dist (dist (f x) (g x)) (dist l m) < e" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
412 |
unfolding dist_real_def |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
413 |
using dist_triangle2 [of "f x" "g x" "l"] |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
414 |
using dist_triangle2 [of "g x" "l" "m"] |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
415 |
using dist_triangle3 [of "l" "m" "f x"] |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
416 |
using dist_triangle [of "f x" "m" "g x"] |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
417 |
by arith |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
418 |
qed |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
419 |
qed |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
420 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
421 |
lemma continuous_dist[continuous_intros]: |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
422 |
fixes f g :: "_ \<Rightarrow> 'a :: metric_space" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
423 |
shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. dist (f x) (g x))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
424 |
unfolding continuous_def by (rule tendsto_dist) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
425 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
426 |
lemma continuous_on_dist[continuous_intros]: |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
427 |
fixes f g :: "_ \<Rightarrow> 'a :: metric_space" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
428 |
shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. dist (f x) (g x))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
429 |
unfolding continuous_on_def by (auto intro: tendsto_dist) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
430 |
|
31565 | 431 |
lemma tendsto_norm [tendsto_intros]: |
44195 | 432 |
"(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
|
433 |
unfolding norm_conv_dist by (intro tendsto_intros) |
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
434 |
|
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
|
435 |
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
|
436 |
"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
|
437 |
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
|
438 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
439 |
lemma continuous_on_norm [continuous_intros]: |
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
440 |
"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
|
441 |
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
|
442 |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
443 |
lemma tendsto_norm_zero: |
44195 | 444 |
"(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
|
445 |
by (drule tendsto_norm, simp) |
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
446 |
|
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
447 |
lemma tendsto_norm_zero_cancel: |
44195 | 448 |
"((\<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
|
449 |
unfolding tendsto_iff dist_norm by simp |
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
450 |
|
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
451 |
lemma tendsto_norm_zero_iff: |
44195 | 452 |
"((\<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
|
453 |
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
|
454 |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
455 |
lemma tendsto_rabs [tendsto_intros]: |
44195 | 456 |
"(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
|
457 |
by (fold real_norm_def, rule tendsto_norm) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
458 |
|
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
|
459 |
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
|
460 |
"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
|
461 |
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
|
462 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
463 |
lemma continuous_on_rabs [continuous_intros]: |
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
464 |
"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
|
465 |
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
|
466 |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
467 |
lemma tendsto_rabs_zero: |
44195 | 468 |
"(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
|
469 |
by (fold real_norm_def, rule tendsto_norm_zero) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
470 |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
471 |
lemma tendsto_rabs_zero_cancel: |
44195 | 472 |
"((\<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
|
473 |
by (fold real_norm_def, rule tendsto_norm_zero_cancel) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
474 |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
475 |
lemma tendsto_rabs_zero_iff: |
44195 | 476 |
"((\<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
|
477 |
by (fold real_norm_def, rule tendsto_norm_zero_iff) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
478 |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
479 |
subsubsection {* Addition and subtraction *} |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
480 |
|
31565 | 481 |
lemma tendsto_add [tendsto_intros]: |
31349 |