author | huffman |
Wed, 03 Jun 2009 09:58:11 -0700 | |
changeset 31417 | c12b25b7f015 |
parent 31415 | 80686a815b59 |
child 31491 | f7310185481d |
permissions | -rw-r--r-- |
30019
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
1 |
(* Title: HOL/Library/Product_Vector.thy |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
2 |
Author: Brian Huffman |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
3 |
*) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
4 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
5 |
header {* Cartesian Products as Vector Spaces *} |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
6 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
7 |
theory Product_Vector |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
8 |
imports Inner_Product Product_plus |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
9 |
begin |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
10 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
11 |
subsection {* Product is a real vector space *} |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
12 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
13 |
instantiation "*" :: (real_vector, real_vector) real_vector |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
14 |
begin |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
15 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
16 |
definition scaleR_prod_def: |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
17 |
"scaleR r A = (scaleR r (fst A), scaleR r (snd A))" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
18 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
19 |
lemma fst_scaleR [simp]: "fst (scaleR r A) = scaleR r (fst A)" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
20 |
unfolding scaleR_prod_def by simp |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
21 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
22 |
lemma snd_scaleR [simp]: "snd (scaleR r A) = scaleR r (snd A)" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
23 |
unfolding scaleR_prod_def by simp |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
24 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
25 |
lemma scaleR_Pair [simp]: "scaleR r (a, b) = (scaleR r a, scaleR r b)" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
26 |
unfolding scaleR_prod_def by simp |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
27 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
28 |
instance proof |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
29 |
fix a b :: real and x y :: "'a \<times> 'b" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
30 |
show "scaleR a (x + y) = scaleR a x + scaleR a y" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
31 |
by (simp add: expand_prod_eq scaleR_right_distrib) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
32 |
show "scaleR (a + b) x = scaleR a x + scaleR b x" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
33 |
by (simp add: expand_prod_eq scaleR_left_distrib) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
34 |
show "scaleR a (scaleR b x) = scaleR (a * b) x" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
35 |
by (simp add: expand_prod_eq) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
36 |
show "scaleR 1 x = x" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
37 |
by (simp add: expand_prod_eq) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
38 |
qed |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
39 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
40 |
end |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
41 |
|
31415 | 42 |
subsection {* Product is a topological space *} |
43 |
||
44 |
instantiation |
|
45 |
"*" :: (topological_space, topological_space) topological_space |
|
46 |
begin |
|
47 |
||
31417 | 48 |
definition topo_prod_def: |
49 |
"topo = {S. \<forall>x\<in>S. \<exists>A\<in>topo. \<exists>B\<in>topo. x \<in> A \<times> B \<and> A \<times> B \<subseteq> S}" |
|
31415 | 50 |
|
51 |
instance proof |
|
31417 | 52 |
show "(UNIV :: ('a \<times> 'b) set) \<in> topo" |
53 |
unfolding topo_prod_def by (auto intro: topo_UNIV) |
|
31415 | 54 |
next |
55 |
fix S T :: "('a \<times> 'b) set" |
|
31417 | 56 |
assume "S \<in> topo" "T \<in> topo" thus "S \<inter> T \<in> topo" |
57 |
unfolding topo_prod_def |
|
31415 | 58 |
apply clarify |
59 |
apply (drule (1) bspec)+ |
|
60 |
apply (clarify, rename_tac Sa Ta Sb Tb) |
|
31417 | 61 |
apply (rule_tac x="Sa \<inter> Ta" in rev_bexI) |
62 |
apply (simp add: topo_Int) |
|
63 |
apply (rule_tac x="Sb \<inter> Tb" in rev_bexI) |
|
64 |
apply (simp add: topo_Int) |
|
31415 | 65 |
apply fast |
66 |
done |
|
67 |
next |
|
68 |
fix T :: "('a \<times> 'b) set set" |
|
31417 | 69 |
assume "T \<subseteq> topo" thus "\<Union>T \<in> topo" |
70 |
unfolding topo_prod_def Bex_def by fast |
|
31415 | 71 |
qed |
72 |
||
73 |
end |
|
74 |
||
31339
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
75 |
subsection {* Product is a metric space *} |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
76 |
|
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
77 |
instantiation |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
78 |
"*" :: (metric_space, metric_space) metric_space |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
79 |
begin |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
80 |
|
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
81 |
definition dist_prod_def: |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
82 |
"dist (x::'a \<times> 'b) y = sqrt ((dist (fst x) (fst y))\<twosuperior> + (dist (snd x) (snd y))\<twosuperior>)" |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
83 |
|
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
84 |
lemma dist_Pair_Pair: "dist (a, b) (c, d) = sqrt ((dist a c)\<twosuperior> + (dist b d)\<twosuperior>)" |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
85 |
unfolding dist_prod_def by simp |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
86 |
|
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
87 |
instance proof |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
88 |
fix x y :: "'a \<times> 'b" |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
89 |
show "dist x y = 0 \<longleftrightarrow> x = y" |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
90 |
unfolding dist_prod_def |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
91 |
by (simp add: expand_prod_eq) |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
92 |
next |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
93 |
fix x y z :: "'a \<times> 'b" |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
94 |
show "dist x y \<le> dist x z + dist y z" |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
95 |
unfolding dist_prod_def |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
96 |
apply (rule order_trans [OF _ real_sqrt_sum_squares_triangle_ineq]) |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
97 |
apply (rule real_sqrt_le_mono) |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
98 |
apply (rule order_trans [OF add_mono]) |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
99 |
apply (rule power_mono [OF dist_triangle2 [of _ _ "fst z"] zero_le_dist]) |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
100 |
apply (rule power_mono [OF dist_triangle2 [of _ _ "snd z"] zero_le_dist]) |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
101 |
apply (simp only: real_sum_squared_expand) |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
102 |
done |
31415 | 103 |
next |
104 |
(* FIXME: long proof! *) |
|
105 |
(* Maybe it would be easier to define topological spaces *) |
|
106 |
(* in terms of neighborhoods instead of open sets? *) |
|
31417 | 107 |
show "topo = {S::('a \<times> 'b) set. \<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S}" |
108 |
unfolding topo_prod_def topo_dist |
|
109 |
apply (safe, rename_tac S a b) |
|
31415 | 110 |
apply (drule (1) bspec) |
111 |
apply clarify |
|
112 |
apply (drule (1) bspec)+ |
|
113 |
apply (clarify, rename_tac r s) |
|
114 |
apply (rule_tac x="min r s" in exI, simp) |
|
115 |
apply (clarify, rename_tac c d) |
|
116 |
apply (erule subsetD) |
|
117 |
apply (simp add: dist_Pair_Pair) |
|
118 |
apply (rule conjI) |
|
119 |
apply (drule spec, erule mp) |
|
120 |
apply (erule le_less_trans [OF real_sqrt_sum_squares_ge1]) |
|
121 |
apply (drule spec, erule mp) |
|
122 |
apply (erule le_less_trans [OF real_sqrt_sum_squares_ge2]) |
|
123 |
||
31417 | 124 |
apply (rename_tac S a b) |
31415 | 125 |
apply (drule (1) bspec) |
126 |
apply clarify |
|
127 |
apply (subgoal_tac "\<exists>r>0. \<exists>s>0. e = sqrt (r\<twosuperior> + s\<twosuperior>)") |
|
128 |
apply clarify |
|
31417 | 129 |
apply (rule_tac x="{y. dist y a < r}" in rev_bexI) |
31415 | 130 |
apply clarify |
131 |
apply (rule_tac x="r - dist x a" in exI, rule conjI, simp) |
|
132 |
apply clarify |
|
133 |
apply (rule le_less_trans [OF dist_triangle]) |
|
134 |
apply (erule less_le_trans [OF add_strict_right_mono], simp) |
|
31417 | 135 |
apply (rule_tac x="{y. dist y b < s}" in rev_bexI) |
31415 | 136 |
apply clarify |
137 |
apply (rule_tac x="s - dist x b" in exI, rule conjI, simp) |
|
138 |
apply clarify |
|
139 |
apply (rule le_less_trans [OF dist_triangle]) |
|
140 |
apply (erule less_le_trans [OF add_strict_right_mono], simp) |
|
141 |
apply (rule conjI) |
|
142 |
apply simp |
|
143 |
apply (clarify, rename_tac c d) |
|
144 |
apply (drule spec, erule mp) |
|
145 |
apply (simp add: dist_Pair_Pair add_strict_mono power_strict_mono) |
|
146 |
apply (rule_tac x="e / sqrt 2" in exI, simp add: divide_pos_pos) |
|
147 |
apply (rule_tac x="e / sqrt 2" in exI, simp add: divide_pos_pos) |
|
148 |
apply (simp add: power_divide) |
|
149 |
done |
|
31339
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
150 |
qed |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
151 |
|
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
152 |
end |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
153 |
|
31405
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
154 |
subsection {* Continuity of operations *} |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
155 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
156 |
lemma dist_fst_le: "dist (fst x) (fst y) \<le> dist x y" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
157 |
unfolding dist_prod_def by simp |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
158 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
159 |
lemma dist_snd_le: "dist (snd x) (snd y) \<le> dist x y" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
160 |
unfolding dist_prod_def by simp |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
161 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
162 |
lemma tendsto_fst: |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
163 |
assumes "tendsto f a net" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
164 |
shows "tendsto (\<lambda>x. fst (f x)) (fst a) net" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
165 |
proof (rule tendstoI) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
166 |
fix r :: real assume "0 < r" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
167 |
have "eventually (\<lambda>x. dist (f x) a < r) net" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
168 |
using `tendsto f a net` `0 < r` by (rule tendstoD) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
169 |
thus "eventually (\<lambda>x. dist (fst (f x)) (fst a) < r) net" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
170 |
by (rule eventually_elim1) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
171 |
(rule le_less_trans [OF dist_fst_le]) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
172 |
qed |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
173 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
174 |
lemma tendsto_snd: |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
175 |
assumes "tendsto f a net" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
176 |
shows "tendsto (\<lambda>x. snd (f x)) (snd a) net" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
177 |
proof (rule tendstoI) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
178 |
fix r :: real assume "0 < r" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
179 |
have "eventually (\<lambda>x. dist (f x) a < r) net" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
180 |
using `tendsto f a net` `0 < r` by (rule tendstoD) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
181 |
thus "eventually (\<lambda>x. dist (snd (f x)) (snd a) < r) net" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
182 |
by (rule eventually_elim1) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
183 |
(rule le_less_trans [OF dist_snd_le]) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
184 |
qed |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
185 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
186 |
lemma tendsto_Pair: |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
187 |
assumes "tendsto X a net" and "tendsto Y b net" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
188 |
shows "tendsto (\<lambda>i. (X i, Y i)) (a, b) net" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
189 |
proof (rule tendstoI) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
190 |
fix r :: real assume "0 < r" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
191 |
then have "0 < r / sqrt 2" (is "0 < ?s") |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
192 |
by (simp add: divide_pos_pos) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
193 |
have "eventually (\<lambda>i. dist (X i) a < ?s) net" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
194 |
using `tendsto X a net` `0 < ?s` by (rule tendstoD) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
195 |
moreover |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
196 |
have "eventually (\<lambda>i. dist (Y i) b < ?s) net" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
197 |
using `tendsto Y b net` `0 < ?s` by (rule tendstoD) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
198 |
ultimately |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
199 |
show "eventually (\<lambda>i. dist (X i, Y i) (a, b) < r) net" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
200 |
by (rule eventually_elim2) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
201 |
(simp add: real_sqrt_sum_squares_less dist_Pair_Pair) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
202 |
qed |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
203 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
204 |
lemma LIMSEQ_fst: "(X ----> a) \<Longrightarrow> (\<lambda>n. fst (X n)) ----> fst a" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
205 |
unfolding LIMSEQ_conv_tendsto by (rule tendsto_fst) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
206 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
207 |
lemma LIMSEQ_snd: "(X ----> a) \<Longrightarrow> (\<lambda>n. snd (X n)) ----> snd a" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
208 |
unfolding LIMSEQ_conv_tendsto by (rule tendsto_snd) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
209 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
210 |
lemma LIMSEQ_Pair: |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
211 |
assumes "X ----> a" and "Y ----> b" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
212 |
shows "(\<lambda>n. (X n, Y n)) ----> (a, b)" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
213 |
using assms unfolding LIMSEQ_conv_tendsto |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
214 |
by (rule tendsto_Pair) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
215 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
216 |
lemma LIM_fst: "f -- x --> a \<Longrightarrow> (\<lambda>x. fst (f x)) -- x --> fst a" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
217 |
unfolding LIM_conv_tendsto by (rule tendsto_fst) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
218 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
219 |
lemma LIM_snd: "f -- x --> a \<Longrightarrow> (\<lambda>x. snd (f x)) -- x --> snd a" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
220 |
unfolding LIM_conv_tendsto by (rule tendsto_snd) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
221 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
222 |
lemma LIM_Pair: |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
223 |
assumes "f -- x --> a" and "g -- x --> b" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
224 |
shows "(\<lambda>x. (f x, g x)) -- x --> (a, b)" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
225 |
using assms unfolding LIM_conv_tendsto |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
226 |
by (rule tendsto_Pair) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
227 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
228 |
lemma Cauchy_fst: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. fst (X n))" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
229 |
unfolding Cauchy_def by (fast elim: le_less_trans [OF dist_fst_le]) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
230 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
231 |
lemma Cauchy_snd: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. snd (X n))" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
232 |
unfolding Cauchy_def by (fast elim: le_less_trans [OF dist_snd_le]) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
233 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
234 |
lemma Cauchy_Pair: |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
235 |
assumes "Cauchy X" and "Cauchy Y" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
236 |
shows "Cauchy (\<lambda>n. (X n, Y n))" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
237 |
proof (rule metric_CauchyI) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
238 |
fix r :: real assume "0 < r" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
239 |
then have "0 < r / sqrt 2" (is "0 < ?s") |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
240 |
by (simp add: divide_pos_pos) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
241 |
obtain M where M: "\<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < ?s" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
242 |
using metric_CauchyD [OF `Cauchy X` `0 < ?s`] .. |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
243 |
obtain N where N: "\<forall>m\<ge>N. \<forall>n\<ge>N. dist (Y m) (Y n) < ?s" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
244 |
using metric_CauchyD [OF `Cauchy Y` `0 < ?s`] .. |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
245 |
have "\<forall>m\<ge>max M N. \<forall>n\<ge>max M N. dist (X m, Y m) (X n, Y n) < r" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
246 |
using M N by (simp add: real_sqrt_sum_squares_less dist_Pair_Pair) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
247 |
then show "\<exists>n0. \<forall>m\<ge>n0. \<forall>n\<ge>n0. dist (X m, Y m) (X n, Y n) < r" .. |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
248 |
qed |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
249 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
250 |
lemma isCont_Pair [simp]: |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
251 |
"\<lbrakk>isCont f x; isCont g x\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. (f x, g x)) x" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
252 |
unfolding isCont_def by (rule LIM_Pair) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
253 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
254 |
subsection {* Product is a complete metric space *} |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
255 |
|
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
256 |
instance "*" :: (complete_space, complete_space) complete_space |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
257 |
proof |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
258 |
fix X :: "nat \<Rightarrow> 'a \<times> 'b" assume "Cauchy X" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
259 |
have 1: "(\<lambda>n. fst (X n)) ----> lim (\<lambda>n. fst (X n))" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
260 |
using Cauchy_fst [OF `Cauchy X`] |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
261 |
by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
262 |
have 2: "(\<lambda>n. snd (X n)) ----> lim (\<lambda>n. snd (X n))" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
263 |
using Cauchy_snd [OF `Cauchy X`] |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
264 |
by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
265 |
have "X ----> (lim (\<lambda>n. fst (X n)), lim (\<lambda>n. snd (X n)))" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
266 |
using LIMSEQ_Pair [OF 1 2] by simp |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
267 |
then show "convergent X" |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
268 |
by (rule convergentI) |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
269 |
qed |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
270 |
|
30019
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
271 |
subsection {* Product is a normed vector space *} |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
272 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
273 |
instantiation |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
274 |
"*" :: (real_normed_vector, real_normed_vector) real_normed_vector |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
275 |
begin |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
276 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
277 |
definition norm_prod_def: |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
278 |
"norm x = sqrt ((norm (fst x))\<twosuperior> + (norm (snd x))\<twosuperior>)" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
279 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
280 |
definition sgn_prod_def: |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
281 |
"sgn (x::'a \<times> 'b) = scaleR (inverse (norm x)) x" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
282 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
283 |
lemma norm_Pair: "norm (a, b) = sqrt ((norm a)\<twosuperior> + (norm b)\<twosuperior>)" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
284 |
unfolding norm_prod_def by simp |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
285 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
286 |
instance proof |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
287 |
fix r :: real and x y :: "'a \<times> 'b" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
288 |
show "0 \<le> norm x" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
289 |
unfolding norm_prod_def by simp |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
290 |
show "norm x = 0 \<longleftrightarrow> x = 0" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
291 |
unfolding norm_prod_def |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
292 |
by (simp add: expand_prod_eq) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
293 |
show "norm (x + y) \<le> norm x + norm y" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
294 |
unfolding norm_prod_def |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
295 |
apply (rule order_trans [OF _ real_sqrt_sum_squares_triangle_ineq]) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
296 |
apply (simp add: add_mono power_mono norm_triangle_ineq) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
297 |
done |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
298 |
show "norm (scaleR r x) = \<bar>r\<bar> * norm x" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
299 |
unfolding norm_prod_def |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
300 |
apply (simp add: norm_scaleR power_mult_distrib) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
301 |
apply (simp add: right_distrib [symmetric]) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
302 |
apply (simp add: real_sqrt_mult_distrib) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
303 |
done |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
304 |
show "sgn x = scaleR (inverse (norm x)) x" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
305 |
by (rule sgn_prod_def) |
31290 | 306 |
show "dist x y = norm (x - y)" |
31339
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
307 |
unfolding dist_prod_def norm_prod_def |
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
huffman
parents:
31290
diff
changeset
|
308 |
by (simp add: dist_norm) |
30019
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
309 |
qed |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
310 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
311 |
end |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
312 |
|
31405
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
313 |
instance "*" :: (banach, banach) banach .. |
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
314 |
|
30019
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
315 |
subsection {* Product is an inner product space *} |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
316 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
317 |
instantiation "*" :: (real_inner, real_inner) real_inner |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
318 |
begin |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
319 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
320 |
definition inner_prod_def: |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
321 |
"inner x y = inner (fst x) (fst y) + inner (snd x) (snd y)" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
322 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
323 |
lemma inner_Pair [simp]: "inner (a, b) (c, d) = inner a c + inner b d" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
324 |
unfolding inner_prod_def by simp |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
325 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
326 |
instance proof |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
327 |
fix r :: real |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
328 |
fix x y z :: "'a::real_inner * 'b::real_inner" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
329 |
show "inner x y = inner y x" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
330 |
unfolding inner_prod_def |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
331 |
by (simp add: inner_commute) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
332 |
show "inner (x + y) z = inner x z + inner y z" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
333 |
unfolding inner_prod_def |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
334 |
by (simp add: inner_left_distrib) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
335 |
show "inner (scaleR r x) y = r * inner x y" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
336 |
unfolding inner_prod_def |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
337 |
by (simp add: inner_scaleR_left right_distrib) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
338 |
show "0 \<le> inner x x" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
339 |
unfolding inner_prod_def |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
340 |
by (intro add_nonneg_nonneg inner_ge_zero) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
341 |
show "inner x x = 0 \<longleftrightarrow> x = 0" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
342 |
unfolding inner_prod_def expand_prod_eq |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
343 |
by (simp add: add_nonneg_eq_0_iff) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
344 |
show "norm x = sqrt (inner x x)" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
345 |
unfolding norm_prod_def inner_prod_def |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
346 |
by (simp add: power2_norm_eq_inner) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
347 |
qed |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
348 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
349 |
end |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
350 |
|
31405
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
huffman
parents:
31388
diff
changeset
|
351 |
subsection {* Pair operations are linear *} |
30019
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
352 |
|
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
30019
diff
changeset
|
353 |
interpretation fst: bounded_linear fst |
30019
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
354 |
apply (unfold_locales) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
355 |
apply (rule fst_add) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
356 |
apply (rule fst_scaleR) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
357 |
apply (rule_tac x="1" in exI, simp add: norm_Pair) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
358 |
done |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
359 |
|
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
30019
diff
changeset
|
360 |
interpretation snd: bounded_linear snd |
30019
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
361 |
apply (unfold_locales) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
362 |
apply (rule snd_add) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
363 |
apply (rule snd_scaleR) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
364 |
apply (rule_tac x="1" in exI, simp add: norm_Pair) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
365 |
done |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
366 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
367 |
text {* TODO: move to NthRoot *} |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
368 |
lemma sqrt_add_le_add_sqrt: |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
369 |
assumes x: "0 \<le> x" and y: "0 \<le> y" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
370 |
shows "sqrt (x + y) \<le> sqrt x + sqrt y" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
371 |
apply (rule power2_le_imp_le) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
372 |
apply (simp add: real_sum_squared_expand add_nonneg_nonneg x y) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
373 |
apply (simp add: mult_nonneg_nonneg x y) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
374 |
apply (simp add: add_nonneg_nonneg x y) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
375 |
done |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
376 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
377 |
lemma bounded_linear_Pair: |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
378 |
assumes f: "bounded_linear f" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
379 |
assumes g: "bounded_linear g" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
380 |
shows "bounded_linear (\<lambda>x. (f x, g x))" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
381 |
proof |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
382 |
interpret f: bounded_linear f by fact |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
383 |
interpret g: bounded_linear g by fact |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
384 |
fix x y and r :: real |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
385 |
show "(f (x + y), g (x + y)) = (f x, g x) + (f y, g y)" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
386 |
by (simp add: f.add g.add) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
387 |
show "(f (r *\<^sub>R x), g (r *\<^sub>R x)) = r *\<^sub>R (f x, g x)" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
388 |
by (simp add: f.scaleR g.scaleR) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
389 |
obtain Kf where "0 < Kf" and norm_f: "\<And>x. norm (f x) \<le> norm x * Kf" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
390 |
using f.pos_bounded by fast |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
391 |
obtain Kg where "0 < Kg" and norm_g: "\<And>x. norm (g x) \<le> norm x * Kg" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
392 |
using g.pos_bounded by fast |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
393 |
have "\<forall>x. norm (f x, g x) \<le> norm x * (Kf + Kg)" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
394 |
apply (rule allI) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
395 |
apply (simp add: norm_Pair) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
396 |
apply (rule order_trans [OF sqrt_add_le_add_sqrt], simp, simp) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
397 |
apply (simp add: right_distrib) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
398 |
apply (rule add_mono [OF norm_f norm_g]) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
399 |
done |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
400 |
then show "\<exists>K. \<forall>x. norm (f x, g x) \<le> norm x * K" .. |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
401 |
qed |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
402 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
403 |
subsection {* Frechet derivatives involving pairs *} |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
404 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
405 |
lemma FDERIV_Pair: |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
406 |
assumes f: "FDERIV f x :> f'" and g: "FDERIV g x :> g'" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
407 |
shows "FDERIV (\<lambda>x. (f x, g x)) x :> (\<lambda>h. (f' h, g' h))" |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
408 |
apply (rule FDERIV_I) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
409 |
apply (rule bounded_linear_Pair) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
410 |
apply (rule FDERIV_bounded_linear [OF f]) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
411 |
apply (rule FDERIV_bounded_linear [OF g]) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
412 |
apply (simp add: norm_Pair) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
413 |
apply (rule real_LIM_sandwich_zero) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
414 |
apply (rule LIM_add_zero) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
415 |
apply (rule FDERIV_D [OF f]) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
416 |
apply (rule FDERIV_D [OF g]) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
417 |
apply (rename_tac h) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
418 |
apply (simp add: divide_nonneg_pos) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
419 |
apply (rename_tac h) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
420 |
apply (subst add_divide_distrib [symmetric]) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
421 |
apply (rule divide_right_mono [OF _ norm_ge_zero]) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
422 |
apply (rule order_trans [OF sqrt_add_le_add_sqrt]) |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
423 |
apply simp |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
424 |
apply simp |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
425 |
apply simp |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
426 |
done |
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
427 |
|
a2f19e0a28b2
add theory of products as real vector spaces to Library
huffman
parents:
diff
changeset
|
428 |
end |