author | paulson <lp15@cam.ac.uk> |
Fri, 13 Nov 2015 12:27:13 +0000 | |
changeset 61649 | 268d88ec9087 |
parent 61609 | 77b453bd616f |
child 61694 | 6571c78c9667 |
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 |
|
60758 | 8 |
section \<open>Limits on Real Vector Spaces\<close> |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
9 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
10 |
theory Limits |
51524 | 11 |
imports Real_Vector_Spaces |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
12 |
begin |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
13 |
|
60758 | 14 |
subsection \<open>Filter going to infinity norm\<close> |
51526 | 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: |
61076 | 25 |
"(at_infinity :: 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 |
|
59613
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
45 |
lemma lim_infinity_imp_sequentially: |
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
46 |
"(f ---> l) at_infinity \<Longrightarrow> ((\<lambda>n. f(n)) ---> l) sequentially" |
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
47 |
by (simp add: filterlim_at_top_imp_at_infinity filterlim_compose filterlim_real_sequentially) |
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
48 |
|
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
49 |
|
60758 | 50 |
subsubsection \<open>Boundedness\<close> |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
51 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
52 |
definition Bfun :: "('a \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a filter \<Rightarrow> bool" where |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
53 |
Bfun_metric_def: "Bfun f F = (\<exists>y. \<exists>K>0. eventually (\<lambda>x. dist (f x) y \<le> K) F)" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
54 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
55 |
abbreviation Bseq :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
56 |
"Bseq X \<equiv> Bfun X sequentially" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
57 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
58 |
lemma Bseq_conv_Bfun: "Bseq X \<longleftrightarrow> Bfun X sequentially" .. |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
59 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
60 |
lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
61 |
unfolding Bfun_metric_def by (subst eventually_sequentially_seg) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
62 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
63 |
lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
64 |
unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg) |
31355 | 65 |
|
51474
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
66 |
lemma Bfun_def: |
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
67 |
"Bfun f F \<longleftrightarrow> (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) F)" |
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
68 |
unfolding Bfun_metric_def norm_conv_dist |
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
69 |
proof safe |
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
70 |
fix y K assume "0 < K" and *: "eventually (\<lambda>x. dist (f x) y \<le> K) F" |
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
71 |
moreover have "eventually (\<lambda>x. dist (f x) 0 \<le> dist (f x) y + dist 0 y) F" |
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
72 |
by (intro always_eventually) (metis dist_commute dist_triangle) |
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
73 |
with * have "eventually (\<lambda>x. dist (f x) 0 \<le> K + dist 0 y) F" |
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
74 |
by eventually_elim auto |
60758 | 75 |
with \<open>0 < K\<close> show "\<exists>K>0. eventually (\<lambda>x. dist (f x) 0 \<le> K) F" |
51474
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
76 |
by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto |
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
hoelzl
parents:
51472
diff
changeset
|
77 |
qed auto |
31355 | 78 |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
79 |
lemma BfunI: |
44195 | 80 |
assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) F" shows "Bfun f F" |
31355 | 81 |
unfolding Bfun_def |
82 |
proof (intro exI conjI allI) |
|
83 |
show "0 < max K 1" by simp |
|
84 |
next |
|
44195 | 85 |
show "eventually (\<lambda>x. norm (f x) \<le> max K 1) F" |
31355 | 86 |
using K by (rule eventually_elim1, simp) |
87 |
qed |
|
88 |
||
89 |
lemma BfunE: |
|
44195 | 90 |
assumes "Bfun f F" |
91 |
obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) F" |
|
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
92 |
using assms unfolding Bfun_def by blast |
31355 | 93 |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
94 |
lemma Cauchy_Bseq: "Cauchy X \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
95 |
unfolding Cauchy_def Bfun_metric_def eventually_sequentially |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
96 |
apply (erule_tac x=1 in allE) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
97 |
apply simp |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
98 |
apply safe |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
99 |
apply (rule_tac x="X M" in exI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
100 |
apply (rule_tac x=1 in exI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
101 |
apply (erule_tac x=M in allE) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
102 |
apply simp |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
103 |
apply (rule_tac x=M in exI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
104 |
apply (auto simp: dist_commute) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
105 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
106 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
107 |
|
60758 | 108 |
subsubsection \<open>Bounded Sequences\<close> |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
109 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
110 |
lemma BseqI': "(\<And>n. norm (X n) \<le> K) \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
111 |
by (intro BfunI) (auto simp: eventually_sequentially) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
112 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
113 |
lemma BseqI2': "\<forall>n\<ge>N. norm (X n) \<le> K \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
114 |
by (intro BfunI) (auto simp: eventually_sequentially) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
115 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
116 |
lemma Bseq_def: "Bseq X \<longleftrightarrow> (\<exists>K>0. \<forall>n. norm (X n) \<le> K)" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
117 |
unfolding Bfun_def eventually_sequentially |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
118 |
proof safe |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
119 |
fix N K assume "0 < K" "\<forall>n\<ge>N. norm (X n) \<le> K" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
120 |
then show "\<exists>K>0. \<forall>n. norm (X n) \<le> K" |
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54263
diff
changeset
|
121 |
by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] max.strict_coboundedI2) |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
122 |
(auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
123 |
qed auto |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
124 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
125 |
lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
126 |
unfolding Bseq_def by auto |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
127 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
128 |
lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
129 |
by (simp add: Bseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
130 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
131 |
lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
132 |
by (auto simp add: Bseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
133 |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
134 |
lemma Bseq_bdd_above: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_above (range X)" |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
135 |
proof (elim BseqE, intro bdd_aboveI2) |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
136 |
fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "X n \<le> K" |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
137 |
by (auto elim!: allE[of _ n]) |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
138 |
qed |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
139 |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
140 |
lemma Bseq_bdd_above': |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
141 |
"Bseq (X::nat \<Rightarrow> 'a :: real_normed_vector) \<Longrightarrow> bdd_above (range (\<lambda>n. norm (X n)))" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
142 |
proof (elim BseqE, intro bdd_aboveI2) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
143 |
fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "norm (X n) \<le> K" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
144 |
by (auto elim!: allE[of _ n]) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
145 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
146 |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
147 |
lemma Bseq_bdd_below: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_below (range X)" |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
148 |
proof (elim BseqE, intro bdd_belowI2) |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
149 |
fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "- K \<le> X n" |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
150 |
by (auto elim!: allE[of _ n]) |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
151 |
qed |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54230
diff
changeset
|
152 |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
153 |
lemma Bseq_eventually_mono: |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
154 |
assumes "eventually (\<lambda>n. norm (f n) \<le> norm (g n)) sequentially" "Bseq g" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
155 |
shows "Bseq f" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
156 |
proof - |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
157 |
from assms(1) obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> norm (f n) \<le> norm (g n)" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
158 |
by (auto simp: eventually_at_top_linorder) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
159 |
moreover from assms(2) obtain K where K: "\<And>n. norm (g n) \<le> K" by (blast elim!: BseqE) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
160 |
ultimately have "norm (f n) \<le> max K (Max {norm (f n) |n. n < N})" for n |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
161 |
apply (cases "n < N") |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
162 |
apply (rule max.coboundedI2, rule Max.coboundedI, auto) [] |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
163 |
apply (rule max.coboundedI1, force intro: order.trans[OF N K]) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
164 |
done |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
165 |
thus ?thesis by (blast intro: BseqI') |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
166 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
167 |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
168 |
lemma lemma_NBseq_def: |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
169 |
"(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
170 |
proof safe |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
171 |
fix K :: real |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
172 |
from reals_Archimedean2 obtain n :: nat where "K < real n" .. |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
173 |
then have "K \<le> real (Suc n)" by auto |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
174 |
moreover assume "\<forall>m. norm (X m) \<le> K" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
175 |
ultimately have "\<forall>m. norm (X m) \<le> real (Suc n)" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
176 |
by (blast intro: order_trans) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
177 |
then show "\<exists>N. \<forall>n. norm (X n) \<le> real (Suc N)" .. |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
178 |
next |
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
179 |
show "\<And>N. \<forall>n. norm (X n) \<le> real (Suc N) \<Longrightarrow> \<exists>K>0. \<forall>n. norm (X n) \<le> K" |
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
180 |
using of_nat_0_less_iff by blast |
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
181 |
qed |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
182 |
|
60758 | 183 |
text\<open>alternative definition for Bseq\<close> |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
184 |
lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
185 |
apply (simp add: Bseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
186 |
apply (simp (no_asm) add: lemma_NBseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
187 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
188 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
189 |
lemma lemma_NBseq_def2: |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
190 |
"(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
191 |
apply (subst lemma_NBseq_def, auto) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
192 |
apply (rule_tac x = "Suc N" in exI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
193 |
apply (rule_tac [2] x = N in exI) |
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
194 |
apply (auto simp add: of_nat_Suc) |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
195 |
prefer 2 apply (blast intro: order_less_imp_le) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
196 |
apply (drule_tac x = n in spec, simp) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
197 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
198 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
199 |
(* yet another definition for Bseq *) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
200 |
lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
201 |
by (simp add: Bseq_def lemma_NBseq_def2) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
202 |
|
60758 | 203 |
subsubsection\<open>A Few More Equivalence Theorems for Boundedness\<close> |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
204 |
|
60758 | 205 |
text\<open>alternative formulation for boundedness\<close> |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
206 |
lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
207 |
apply (unfold Bseq_def, safe) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
208 |
apply (rule_tac [2] x = "k + norm x" in exI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
209 |
apply (rule_tac x = K in exI, simp) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
210 |
apply (rule exI [where x = 0], auto) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
211 |
apply (erule order_less_le_trans, simp) |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
212 |
apply (drule_tac x=n in spec) |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
213 |
apply (drule order_trans [OF norm_triangle_ineq2]) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
214 |
apply simp |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
215 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
216 |
|
60758 | 217 |
text\<open>alternative formulation for boundedness\<close> |
53602 | 218 |
lemma Bseq_iff3: |
219 |
"Bseq X \<longleftrightarrow> (\<exists>k>0. \<exists>N. \<forall>n. norm (X n + - X N) \<le> k)" (is "?P \<longleftrightarrow> ?Q") |
|
220 |
proof |
|
221 |
assume ?P |
|
222 |
then obtain K |
|
223 |
where *: "0 < K" and **: "\<And>n. norm (X n) \<le> K" by (auto simp add: Bseq_def) |
|
224 |
from * have "0 < K + norm (X 0)" by (rule order_less_le_trans) simp |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
225 |
from ** have "\<forall>n. norm (X n - X 0) \<le> K + norm (X 0)" |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
226 |
by (auto intro: order_trans norm_triangle_ineq4) |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
227 |
then have "\<forall>n. norm (X n + - X 0) \<le> K + norm (X 0)" |
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
228 |
by simp |
60758 | 229 |
with \<open>0 < K + norm (X 0)\<close> show ?Q by blast |
53602 | 230 |
next |
231 |
assume ?Q then show ?P by (auto simp add: Bseq_iff2) |
|
232 |
qed |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
233 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
234 |
lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
235 |
apply (simp add: Bseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
236 |
apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
237 |
apply (drule_tac x = n in spec, arith) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
238 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
239 |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
240 |
|
60758 | 241 |
subsubsection\<open>Upper Bounds and Lubs of Bounded Sequences\<close> |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
242 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
243 |
lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
244 |
by (simp add: Bseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
245 |
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
246 |
lemma Bseq_add: |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
247 |
assumes "Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector)" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
248 |
shows "Bseq (\<lambda>x. f x + c)" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
249 |
proof - |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
250 |
from assms obtain K where K: "\<And>x. norm (f x) \<le> K" unfolding Bseq_def by blast |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
251 |
{ |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
252 |
fix x :: nat |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
253 |
have "norm (f x + c) \<le> norm (f x) + norm c" by (rule norm_triangle_ineq) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
254 |
also have "norm (f x) \<le> K" by (rule K) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
255 |
finally have "norm (f x + c) \<le> K + norm c" by simp |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
256 |
} |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
257 |
thus ?thesis by (rule BseqI') |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
258 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
259 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
260 |
lemma Bseq_add_iff: "Bseq (\<lambda>x. f x + c) \<longleftrightarrow> Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector)" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
261 |
using Bseq_add[of f c] Bseq_add[of "\<lambda>x. f x + c" "-c"] by auto |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
262 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
263 |
lemma Bseq_mult: |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
264 |
assumes "Bseq (f :: nat \<Rightarrow> 'a :: real_normed_field)" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
265 |
assumes "Bseq (g :: nat \<Rightarrow> 'a :: real_normed_field)" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
266 |
shows "Bseq (\<lambda>x. f x * g x)" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
267 |
proof - |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
268 |
from assms obtain K1 K2 where K: "\<And>x. norm (f x) \<le> K1" "K1 > 0" "\<And>x. norm (g x) \<le> K2" "K2 > 0" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
269 |
unfolding Bseq_def by blast |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
270 |
hence "\<And>x. norm (f x * g x) \<le> K1 * K2" by (auto simp: norm_mult intro!: mult_mono) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
271 |
thus ?thesis by (rule BseqI') |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
272 |
qed |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
273 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
274 |
lemma Bfun_const [simp]: "Bfun (\<lambda>_. c) F" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
275 |
unfolding Bfun_metric_def by (auto intro!: exI[of _ c] exI[of _ "1::real"]) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
276 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
277 |
lemma Bseq_cmult_iff: "(c :: 'a :: real_normed_field) \<noteq> 0 \<Longrightarrow> Bseq (\<lambda>x. c * f x) \<longleftrightarrow> Bseq f" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
278 |
proof |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
279 |
assume "c \<noteq> 0" "Bseq (\<lambda>x. c * f x)" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
280 |
find_theorems "Bfun (\<lambda>_. ?c) _" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
281 |
from Bfun_const this(2) have "Bseq (\<lambda>x. inverse c * (c * f x))" by (rule Bseq_mult) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
282 |
with `c \<noteq> 0` show "Bseq f" by (simp add: divide_simps) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
283 |
qed (intro Bseq_mult Bfun_const) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
284 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
285 |
lemma Bseq_subseq: "Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector) \<Longrightarrow> Bseq (\<lambda>x. f (g x))" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
286 |
unfolding Bseq_def by auto |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
287 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
288 |
lemma Bseq_Suc_iff: "Bseq (\<lambda>n. f (Suc n)) \<longleftrightarrow> Bseq (f :: nat \<Rightarrow> 'a :: real_normed_vector)" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
289 |
using Bseq_offset[of f 1] by (auto intro: Bseq_subseq) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
290 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
291 |
lemma increasing_Bseq_subseq_iff: |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
292 |
assumes "\<And>x y. x \<le> y \<Longrightarrow> norm (f x :: 'a :: real_normed_vector) \<le> norm (f y)" "subseq g" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
293 |
shows "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
294 |
proof |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
295 |
assume "Bseq (\<lambda>x. f (g x))" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
296 |
then obtain K where K: "\<And>x. norm (f (g x)) \<le> K" unfolding Bseq_def by auto |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
297 |
{ |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
298 |
fix x :: nat |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
299 |
from filterlim_subseq[OF assms(2)] obtain y where "g y \<ge> x" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
300 |
by (auto simp: filterlim_at_top eventually_at_top_linorder) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
301 |
hence "norm (f x) \<le> norm (f (g y))" using assms(1) by blast |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
302 |
also have "norm (f (g y)) \<le> K" by (rule K) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
303 |
finally have "norm (f x) \<le> K" . |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
304 |
} |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
305 |
thus "Bseq f" by (rule BseqI') |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
306 |
qed (insert Bseq_subseq[of f g], simp_all) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
307 |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
308 |
lemma nonneg_incseq_Bseq_subseq_iff: |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
309 |
assumes "\<And>x. f x \<ge> 0" "incseq (f :: nat \<Rightarrow> real)" "subseq g" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
310 |
shows "Bseq (\<lambda>x. f (g x)) \<longleftrightarrow> Bseq f" |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
311 |
using assms by (intro increasing_Bseq_subseq_iff) (auto simp: incseq_def) |
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
312 |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
313 |
lemma Bseq_eq_bounded: "range f \<subseteq> {a .. b::real} \<Longrightarrow> Bseq f" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
314 |
apply (simp add: subset_eq) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
315 |
apply (rule BseqI'[where K="max (norm a) (norm b)"]) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
316 |
apply (erule_tac x=n in allE) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
317 |
apply auto |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
318 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
319 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
320 |
lemma incseq_bounded: "incseq X \<Longrightarrow> \<forall>i. X i \<le> (B::real) \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
321 |
by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
322 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
323 |
lemma decseq_bounded: "decseq X \<Longrightarrow> \<forall>i. (B::real) \<le> X i \<Longrightarrow> Bseq X" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
324 |
by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
325 |
|
60758 | 326 |
subsection \<open>Bounded Monotonic Sequences\<close> |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
327 |
|
60758 | 328 |
subsubsection\<open>A Bounded and Monotonic Sequence Converges\<close> |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
329 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
330 |
(* TODO: delete *) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
331 |
(* FIXME: one use in NSA/HSEQ.thy *) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
332 |
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
|
333 |
apply (rule_tac x="X m" in exI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
334 |
apply (rule filterlim_cong[THEN iffD2, OF refl refl _ tendsto_const]) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
335 |
unfolding eventually_sequentially |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
336 |
apply blast |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
337 |
done |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
338 |
|
60758 | 339 |
subsection \<open>Convergence to Zero\<close> |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
340 |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
341 |
definition Zfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool" |
44195 | 342 |
where "Zfun f F = (\<forall>r>0. eventually (\<lambda>x. norm (f x) < r) F)" |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
343 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
344 |
lemma ZfunI: |
44195 | 345 |
"(\<And>r. 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F) \<Longrightarrow> Zfun f F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
346 |
unfolding Zfun_def by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
347 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
348 |
lemma ZfunD: |
44195 | 349 |
"\<lbrakk>Zfun f F; 0 < r\<rbrakk> \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
350 |
unfolding Zfun_def by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
351 |
|
31355 | 352 |
lemma Zfun_ssubst: |
44195 | 353 |
"eventually (\<lambda>x. f x = g x) F \<Longrightarrow> Zfun g F \<Longrightarrow> Zfun f F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
354 |
unfolding Zfun_def by (auto elim!: eventually_rev_mp) |
31355 | 355 |
|
44195 | 356 |
lemma Zfun_zero: "Zfun (\<lambda>x. 0) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
357 |
unfolding Zfun_def by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
358 |
|
44195 | 359 |
lemma Zfun_norm_iff: "Zfun (\<lambda>x. norm (f x)) F = Zfun (\<lambda>x. f x) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
360 |
unfolding Zfun_def by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
361 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
362 |
lemma Zfun_imp_Zfun: |
44195 | 363 |
assumes f: "Zfun f F" |
364 |
assumes g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F" |
|
365 |
shows "Zfun (\<lambda>x. g x) F" |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
366 |
proof (cases) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
367 |
assume K: "0 < K" |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
368 |
show ?thesis |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
369 |
proof (rule ZfunI) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
370 |
fix r::real assume "0 < r" |
56541 | 371 |
hence "0 < r / K" using K by simp |
44195 | 372 |
then have "eventually (\<lambda>x. norm (f x) < r / K) F" |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
373 |
using ZfunD [OF f] by blast |
44195 | 374 |
with g show "eventually (\<lambda>x. norm (g x) < r) F" |
46887 | 375 |
proof eventually_elim |
376 |
case (elim x) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
377 |
hence "norm (f x) * K < r" |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
378 |
by (simp add: pos_less_divide_eq K) |
46887 | 379 |
thus ?case |
380 |
by (simp add: order_le_less_trans [OF elim(1)]) |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
381 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
382 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
383 |
next |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
384 |
assume "\<not> 0 < K" |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
385 |
hence K: "K \<le> 0" by (simp only: not_less) |
31355 | 386 |
show ?thesis |
387 |
proof (rule ZfunI) |
|
388 |
fix r :: real |
|
389 |
assume "0 < r" |
|
44195 | 390 |
from g show "eventually (\<lambda>x. norm (g x) < r) F" |
46887 | 391 |
proof eventually_elim |
392 |
case (elim x) |
|
393 |
also have "norm (f x) * K \<le> norm (f x) * 0" |
|
31355 | 394 |
using K norm_ge_zero by (rule mult_left_mono) |
46887 | 395 |
finally show ?case |
60758 | 396 |
using \<open>0 < r\<close> by simp |
31355 | 397 |
qed |
398 |
qed |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
399 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
400 |
|
44195 | 401 |
lemma Zfun_le: "\<lbrakk>Zfun g F; \<forall>x. norm (f x) \<le> norm (g x)\<rbrakk> \<Longrightarrow> Zfun f F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
402 |
by (erule_tac K="1" in Zfun_imp_Zfun, simp) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
403 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
404 |
lemma Zfun_add: |
44195 | 405 |
assumes f: "Zfun f F" and g: "Zfun g F" |
406 |
shows "Zfun (\<lambda>x. f x + g x) F" |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
407 |
proof (rule ZfunI) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
408 |
fix r::real assume "0 < r" |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
409 |
hence r: "0 < r / 2" by simp |
44195 | 410 |
have "eventually (\<lambda>x. norm (f x) < r/2) F" |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
411 |
using f r by (rule ZfunD) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
412 |
moreover |
44195 | 413 |
have "eventually (\<lambda>x. norm (g x) < r/2) F" |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
414 |
using g r by (rule ZfunD) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
415 |
ultimately |
44195 | 416 |
show "eventually (\<lambda>x. norm (f x + g x) < r) F" |
46887 | 417 |
proof eventually_elim |
418 |
case (elim x) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
419 |
have "norm (f x + g x) \<le> norm (f x) + norm (g x)" |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
420 |
by (rule norm_triangle_ineq) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
421 |
also have "\<dots> < r/2 + r/2" |
46887 | 422 |
using elim by (rule add_strict_mono) |
423 |
finally show ?case |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
424 |
by simp |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
425 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
426 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
427 |
|
44195 | 428 |
lemma Zfun_minus: "Zfun f F \<Longrightarrow> Zfun (\<lambda>x. - f x) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
429 |
unfolding Zfun_def by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
430 |
|
44195 | 431 |
lemma Zfun_diff: "\<lbrakk>Zfun f F; Zfun g F\<rbrakk> \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F" |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
432 |
using Zfun_add [of f F "\<lambda>x. - g x"] by (simp add: Zfun_minus) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
433 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
434 |
lemma (in bounded_linear) Zfun: |
44195 | 435 |
assumes g: "Zfun g F" |
436 |
shows "Zfun (\<lambda>x. f (g x)) F" |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
437 |
proof - |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
438 |
obtain K where "\<And>x. norm (f x) \<le> norm x * K" |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
439 |
using bounded by blast |
44195 | 440 |
then have "eventually (\<lambda>x. norm (f (g x)) \<le> norm (g x) * K) F" |
31355 | 441 |
by simp |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
442 |
with g show ?thesis |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
443 |
by (rule Zfun_imp_Zfun) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
444 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
445 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
446 |
lemma (in bounded_bilinear) Zfun: |
44195 | 447 |
assumes f: "Zfun f F" |
448 |
assumes g: "Zfun g F" |
|
449 |
shows "Zfun (\<lambda>x. f x ** g x) F" |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
450 |
proof (rule ZfunI) |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
451 |
fix r::real assume r: "0 < r" |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
452 |
obtain K where K: "0 < K" |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
453 |
and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K" |
61649
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
454 |
using pos_bounded by blast |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
455 |
from K have K': "0 < inverse K" |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
456 |
by (rule positive_imp_inverse_positive) |
44195 | 457 |
have "eventually (\<lambda>x. norm (f x) < r) F" |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
458 |
using f r by (rule ZfunD) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
459 |
moreover |
44195 | 460 |
have "eventually (\<lambda>x. norm (g x) < inverse K) F" |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
461 |
using g K' by (rule ZfunD) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
462 |
ultimately |
44195 | 463 |
show "eventually (\<lambda>x. norm (f x ** g x) < r) F" |
46887 | 464 |
proof eventually_elim |
465 |
case (elim x) |
|
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
466 |
have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K" |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
467 |
by (rule norm_le) |
31487
93938cafc0e6
put syntax for tendsto in Limits.thy; rename variables
huffman
parents:
31447
diff
changeset
|
468 |
also have "norm (f x) * norm (g x) * K < r * inverse K * K" |
46887 | 469 |
by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero elim K) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
470 |
also from K have "r * inverse K * K = r" |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
471 |
by simp |
46887 | 472 |
finally show ?case . |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
473 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
474 |
qed |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
475 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
476 |
lemma (in bounded_bilinear) Zfun_left: |
44195 | 477 |
"Zfun f F \<Longrightarrow> Zfun (\<lambda>x. f x ** a) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
478 |
by (rule bounded_linear_left [THEN bounded_linear.Zfun]) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
479 |
|
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
480 |
lemma (in bounded_bilinear) Zfun_right: |
44195 | 481 |
"Zfun f F \<Longrightarrow> Zfun (\<lambda>x. a ** f x) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
482 |
by (rule bounded_linear_right [THEN bounded_linear.Zfun]) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
483 |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
484 |
lemmas Zfun_mult = bounded_bilinear.Zfun [OF bounded_bilinear_mult] |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
485 |
lemmas Zfun_mult_right = bounded_bilinear.Zfun_right [OF bounded_bilinear_mult] |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
486 |
lemmas Zfun_mult_left = bounded_bilinear.Zfun_left [OF bounded_bilinear_mult] |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
487 |
|
44195 | 488 |
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
|
489 |
by (simp only: tendsto_iff Zfun_def dist_norm) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
490 |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60017
diff
changeset
|
491 |
lemma tendsto_0_le: "\<lbrakk>(f ---> 0) F; eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) F\<rbrakk> |
56366 | 492 |
\<Longrightarrow> (g ---> 0) F" |
493 |
by (simp add: Zfun_imp_Zfun tendsto_Zfun_iff) |
|
494 |
||
60758 | 495 |
subsubsection \<open>Distance and norms\<close> |
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
496 |
|
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
497 |
lemma tendsto_dist [tendsto_intros]: |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
498 |
fixes l m :: "'a :: metric_space" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
499 |
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
|
500 |
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
|
501 |
proof (rule tendstoI) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
502 |
fix e :: real assume "0 < e" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
503 |
hence e2: "0 < e/2" by simp |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
504 |
from tendstoD [OF f e2] tendstoD [OF g e2] |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
505 |
show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) F" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
506 |
proof (eventually_elim) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
507 |
case (elim x) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
508 |
then show "dist (dist (f x) (g x)) (dist l m) < e" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
509 |
unfolding dist_real_def |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
510 |
using dist_triangle2 [of "f x" "g x" "l"] |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
511 |
using dist_triangle2 [of "g x" "l" "m"] |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
512 |
using dist_triangle3 [of "l" "m" "f x"] |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
513 |
using dist_triangle [of "f x" "m" "g x"] |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
514 |
by arith |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
515 |
qed |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
516 |
qed |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
517 |
|
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
518 |
lemma continuous_dist[continuous_intros]: |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
519 |
fixes f g :: "_ \<Rightarrow> 'a :: metric_space" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
520 |
shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. dist (f x) (g x))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
521 |
unfolding continuous_def by (rule tendsto_dist) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
522 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
523 |
lemma continuous_on_dist[continuous_intros]: |
51531
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
524 |
fixes f g :: "_ \<Rightarrow> 'a :: metric_space" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
525 |
shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. dist (f x) (g x))" |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
526 |
unfolding continuous_on_def by (auto intro: tendsto_dist) |
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
hoelzl
parents:
51529
diff
changeset
|
527 |
|
31565 | 528 |
lemma tendsto_norm [tendsto_intros]: |
44195 | 529 |
"(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
|
530 |
unfolding norm_conv_dist by (intro tendsto_intros) |
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
531 |
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
532 |
lemma continuous_norm [continuous_intros]: |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
533 |
"continuous F f \<Longrightarrow> continuous F (\<lambda>x. norm (f x))" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
534 |
unfolding continuous_def by (rule tendsto_norm) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
535 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
536 |
lemma continuous_on_norm [continuous_intros]: |
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
537 |
"continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. norm (f x))" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
538 |
unfolding continuous_on_def by (auto intro: tendsto_norm) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
539 |
|
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
540 |
lemma tendsto_norm_zero: |
44195 | 541 |
"(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
|
542 |
by (drule tendsto_norm, simp) |
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
543 |
|
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
544 |
lemma tendsto_norm_zero_cancel: |
44195 | 545 |
"((\<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
|
546 |
unfolding tendsto_iff dist_norm by simp |
36662
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
547 |
|
621122eeb138
generalize types of LIMSEQ and LIM; generalize many lemmas
huffman
parents:
36656
diff
changeset
|
548 |
lemma tendsto_norm_zero_iff: |
44195 | 549 |
"((\<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
|
550 |
unfolding tendsto_iff dist_norm by simp |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
551 |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
552 |
lemma tendsto_rabs [tendsto_intros]: |
44195 | 553 |
"(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
|
554 |
by (fold real_norm_def, rule tendsto_norm) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
555 |
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
556 |
lemma continuous_rabs [continuous_intros]: |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
557 |
"continuous F f \<Longrightarrow> continuous F (\<lambda>x. \<bar>f x :: real\<bar>)" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
558 |
unfolding real_norm_def[symmetric] by (rule continuous_norm) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
559 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
560 |
lemma continuous_on_rabs [continuous_intros]: |
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
561 |
"continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. \<bar>f x :: real\<bar>)" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
562 |
unfolding real_norm_def[symmetric] by (rule continuous_on_norm) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
563 |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
564 |
lemma tendsto_rabs_zero: |
44195 | 565 |
"(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
|
566 |
by (fold real_norm_def, rule tendsto_norm_zero) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
567 |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
568 |
lemma tendsto_rabs_zero_cancel: |
44195 | 569 |
"((\<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
|
570 |
by (fold real_norm_def, rule tendsto_norm_zero_cancel) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
571 |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
572 |
lemma tendsto_rabs_zero_iff: |
44195 | 573 |
"((\<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
|
574 |
by (fold real_norm_def, rule tendsto_norm_zero_iff) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
575 |
|
60758 | 576 |
subsubsection \<open>Addition and subtraction\<close> |
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
577 |
|
31565 | 578 |
lemma tendsto_add [tendsto_intros]: |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
579 |
fixes a b :: "'a::real_normed_vector" |
44195 | 580 |
shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> a + b) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
581 |
by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
582 |
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
583 |
lemma continuous_add [continuous_intros]: |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
584 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
585 |
shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x + g x)" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
586 |
unfolding continuous_def by (rule tendsto_add) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
587 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
588 |
lemma continuous_on_add [continuous_intros]: |
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
589 |
fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
590 |
shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f x + g x)" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
591 |
unfolding continuous_on_def by (auto intro: tendsto_add) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
592 |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
593 |
lemma tendsto_add_zero: |
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
594 |
fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector" |
44195 | 595 |
shows "\<lbrakk>(f ---> 0) F; (g ---> 0) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> 0) F" |
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
596 |
by (drule (1) tendsto_add, simp) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
597 |
|
31565 | 598 |
lemma tendsto_minus [tendsto_intros]: |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
599 |
fixes a :: "'a::real_normed_vector" |
44195 | 600 |
shows "(f ---> a) F \<Longrightarrow> ((\<lambda>x. - f x) ---> - a) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
601 |
by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
602 |
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
603 |
lemma continuous_minus [continuous_intros]: |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
604 |
fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
605 |
shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. - f x)" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
606 |
unfolding continuous_def by (rule tendsto_minus) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
607 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
608 |
lemma continuous_on_minus [continuous_intros]: |
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
609 |
fixes f :: "_ \<Rightarrow> 'b::real_normed_vector" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
610 |
shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. - f x)" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
611 |
unfolding continuous_on_def by (auto intro: tendsto_minus) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
612 |
|
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
613 |
lemma tendsto_minus_cancel: |
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
614 |
fixes a :: "'a::real_normed_vector" |
44195 | 615 |
shows "((\<lambda>x. - f x) ---> - a) F \<Longrightarrow> (f ---> a) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
616 |
by (drule tendsto_minus, simp) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
617 |
|
50330
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents:
50327
diff
changeset
|
618 |
lemma tendsto_minus_cancel_left: |
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents:
50327
diff
changeset
|
619 |
"(f ---> - (y::_::real_normed_vector)) F \<longleftrightarrow> ((\<lambda>x. - f x) ---> y) F" |
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents:
50327
diff
changeset
|
620 |
using tendsto_minus_cancel[of f "- y" F] tendsto_minus[of f "- y" F] |
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents:
50327
diff
changeset
|
621 |
by auto |
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents:
50327
diff
changeset
|
622 |
|
31565 | 623 |
lemma tendsto_diff [tendsto_intros]: |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
624 |
fixes a b :: "'a::real_normed_vector" |
44195 | 625 |
shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x - g x) ---> a - b) F" |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53602
diff
changeset
|
626 |
using tendsto_add [of f a F "\<lambda>x. - g x" "- b"] by (simp add: tendsto_minus) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
627 |
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
628 |
lemma continuous_diff [continuous_intros]: |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
629 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
630 |
shows "continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x - g x)" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
631 |
unfolding continuous_def by (rule tendsto_diff) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
632 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56366
diff
changeset
|
633 |
lemma continuous_on_diff [continuous_intros]: |
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
|
634 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
635 |
shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f x - g x)" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
636 |
unfolding continuous_on_def by (auto intro: tendsto_diff) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
637 |
|
31588 | 638 |
lemma tendsto_setsum [tendsto_intros]: |
639 |
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::real_normed_vector" |
|
44195 | 640 |
assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> a i) F" |
641 |
shows "((\<lambda>x. \<Sum>i\<in>S. f i x) ---> (\<Sum>i\<in>S. a i)) F" |
|
31588 | 642 |
proof (cases "finite S") |
643 |
assume "finite S" thus ?thesis using assms |
|
58729
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents:
57512
diff
changeset
|
644 |
by (induct, simp, simp add: tendsto_add) |
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents:
57512
diff
changeset
|
645 |
qed simp |
31588 | 646 |
|
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
|
647 |
lemma continuous_setsum [continuous_intros]: |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
648 |
fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::real_normed_vector" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
649 |
shows "(\<And>i. i \<in> S \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Sum>i\<in>S. f i x)" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
650 |
unfolding continuous_def by (rule tendsto_setsum) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
651 |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
652 |
lemma continuous_on_setsum [continuous_intros]: |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
653 |
fixes f :: "'a \<Rightarrow> _ \<Rightarrow> 'c::real_normed_vector" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
654 |
shows "(\<And>i. i \<in> S \<Longrightarrow> continuous_on s (f i)) \<Longrightarrow> continuous_on s (\<lambda>x. \<Sum>i\<in>S. f i x)" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
655 |
unfolding continuous_on_def by (auto intro: tendsto_setsum) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
656 |
|
50999 | 657 |
lemmas real_tendsto_sandwich = tendsto_sandwich[where 'b=real] |
658 |
||
60758 | 659 |
subsubsection \<open>Linear operators and multiplication\<close> |
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
660 |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
661 |
lemma (in bounded_linear) tendsto: |
44195 | 662 |
"(g ---> a) F \<Longrightarrow> ((\<lambda>x. f (g x)) ---> f a) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
663 |
by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
664 |
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
665 |
lemma (in bounded_linear) continuous: |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
666 |
"continuous F g \<Longrightarrow> continuous F (\<lambda>x. f (g x))" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
667 |
using tendsto[of g _ F] by (auto simp: continuous_def) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
668 |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
669 |
lemma (in bounded_linear) continuous_on: |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
670 |
"continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f (g x))" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
671 |
using tendsto[of g] by (auto simp: continuous_on_def) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
672 |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
673 |
lemma (in bounded_linear) tendsto_zero: |
44195 | 674 |
"(g ---> 0) F \<Longrightarrow> ((\<lambda>x. f (g x)) ---> 0) F" |
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
675 |
by (drule tendsto, simp only: zero) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
676 |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
677 |
lemma (in bounded_bilinear) tendsto: |
44195 | 678 |
"\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x ** g x) ---> a ** b) F" |
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
679 |
by (simp only: tendsto_Zfun_iff prod_diff_prod |
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44079
diff
changeset
|
680 |
Zfun_add Zfun Zfun_left Zfun_right) |
31349
2261c8781f73
new theory of filters and limits; prove LIMSEQ and LIM lemmas using filters
huffman
parents:
diff
changeset
|
681 |
|
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
|
682 |
lemma (in bounded_bilinear) continuous: |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
683 |
"continuous F f \<Longrightarrow> continuous F g \<Longrightarrow> continuous F (\<lambda>x. f x ** g x)" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
684 |
using tendsto[of f _ F g] by (auto simp: continuous_def) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
685 |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
686 |
lemma (in bounded_bilinear) continuous_on: |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
687 |
"continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f x ** g x)" |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
688 |
using tendsto[of f _ _ g] by (auto simp: continuous_on_def) |
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51474
diff
changeset
|
689 |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
690 |
lemma (in bounded_bilinear) tendsto_zero: |
44195 | 691 |
assumes f: "(f ---> 0) F" |
692 |
assumes g: "(g ---> 0) F" |
|
693 |
shows "((\<lambda>x. f x ** g x) ---> 0) F" |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
694 |
using tendsto [OF f g] by (simp add: zero_left) |
31355 | 695 |
|
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
696 |
lemma (in bounded_bilinear) tendsto_left_zero: |
44195 | 697 |
"(f ---> 0) F \<Longrightarrow> ((\<lambda>x. f x ** c) ---> 0) F" |
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
698 |
by (rule bounded_linear.tendsto_zero [OF bounded_linear_left]) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
699 |
|
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
700 |
lemma (in bounded_bilinear) tendsto_right_zero: |
44195 | 701 |
"(f ---> 0) F \<Longrightarrow> ((\<lambda>x. c ** f x) ---> 0) F" |
44194
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
702 |
by (rule bounded_linear.tendsto_zero [OF bounded_linear_right]) |
0639898074ae
generalize lemmas about LIM and LIMSEQ to tendsto
huffman
parents:
44081
diff
changeset
|
703 |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
704 |
lemmas tendsto_of_real [tendsto_intros] = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
705 |
bounded_linear.tendsto [OF bounded_linear_of_real] |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
706 |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
707 |
lemmas tendsto_scaleR [tendsto_intros] = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
708 |
bounded_bilinear.tendsto [OF bounded_bilinear_scaleR] |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
709 |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
710 |
lemmas tendsto_mult [tendsto_intros] = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44253
diff
changeset
|
711 |
bounded_bilinear.tendsto [OF bounded_bilinear_mult] |
44194 |