| author | wenzelm | 
| Thu, 09 Jan 2020 15:45:31 +0100 | |
| changeset 71360 | fcf5ee85743d | 
| parent 71174 | 7fac205dd737 | 
| child 74475 | 409ca22dee4c | 
| permissions | -rw-r--r-- | 
| 63971 
da89140186e2
HOL-Analysis: move Product_Vector and Inner_Product from Library
 hoelzl parents: 
63040diff
changeset | 1 | (* Title: HOL/Analysis/Product_Vector.thy | 
| 30019 
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 | |
| 60500 | 5 | section \<open>Cartesian Products as Vector Spaces\<close> | 
| 30019 
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 | 
| 69511 | 8 | imports | 
| 9 | Complex_Main | |
| 10 | "HOL-Library.Product_Plus" | |
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 11 | begin | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 12 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 13 | lemma Times_eq_image_sum: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 14 | fixes S :: "'a :: comm_monoid_add set" and T :: "'b :: comm_monoid_add set" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 15 |   shows "S \<times> T = {u + v |u v. u \<in> (\<lambda>x. (x, 0)) ` S \<and> v \<in> Pair 0 ` T}"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 16 | by force | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 17 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 18 | |
| 69541 | 19 | subsection \<open>Product is a Module\<close> | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 20 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 21 | locale module_prod = module_pair begin | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 22 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 23 | definition scale :: "'a \<Rightarrow> 'b \<times> 'c \<Rightarrow> 'b \<times> 'c" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 24 | where "scale a v = (s1 a (fst v), s2 a (snd v))" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 25 | |
| 70136 | 26 | lemma\<^marker>\<open>tag important\<close> scale_prod: "scale x (a, b) = (s1 x a, s2 x b)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 27 | by (auto simp: scale_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 28 | |
| 70136 | 29 | sublocale\<^marker>\<open>tag important\<close> p: module scale | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 30 | proof qed (simp_all add: scale_def | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 31 | m1.scale_left_distrib m1.scale_right_distrib m2.scale_left_distrib m2.scale_right_distrib) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 32 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 33 | lemma subspace_Times: "m1.subspace A \<Longrightarrow> m2.subspace B \<Longrightarrow> p.subspace (A \<times> B)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 34 | unfolding m1.subspace_def m2.subspace_def p.subspace_def | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 35 | by (auto simp: zero_prod_def scale_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 36 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 37 | lemma module_hom_fst: "module_hom scale s1 fst" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 38 | by unfold_locales (auto simp: scale_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 39 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 40 | lemma module_hom_snd: "module_hom scale s2 snd" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 41 | by unfold_locales (auto simp: scale_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 42 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 43 | end | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 44 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 45 | locale vector_space_prod = vector_space_pair begin | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 46 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 47 | sublocale module_prod s1 s2 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 48 | rewrites "module_hom = Vector_Spaces.linear" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 49 | by unfold_locales (fact module_hom_eq_linear) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 50 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 51 | sublocale p: vector_space scale by unfold_locales (auto simp: algebra_simps) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 52 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 53 | lemmas linear_fst = module_hom_fst | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 54 | and linear_snd = module_hom_snd | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 55 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 56 | end | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 57 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 58 | |
| 69541 | 59 | subsection \<open>Product is a Real Vector Space\<close> | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 60 | |
| 68617 | 61 | instantiation prod :: (real_vector, real_vector) real_vector | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 62 | begin | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 63 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 64 | definition scaleR_prod_def: | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 65 | "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 | 66 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 67 | 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 | 68 | unfolding scaleR_prod_def by simp | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 69 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 70 | 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 | 71 | unfolding scaleR_prod_def by simp | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 72 | |
| 69541 | 73 | proposition scaleR_Pair [simp]: "scaleR r (a, b) = (scaleR r a, scaleR r b)" | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 74 | unfolding scaleR_prod_def by simp | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 75 | |
| 60679 | 76 | instance | 
| 77 | proof | |
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 78 | 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 | 79 | show "scaleR a (x + y) = scaleR a x + scaleR a y" | 
| 44066 
d74182c93f04
rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
 huffman parents: 
37678diff
changeset | 80 | by (simp add: prod_eq_iff scaleR_right_distrib) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 81 | show "scaleR (a + b) x = scaleR a x + scaleR b x" | 
| 44066 
d74182c93f04
rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
 huffman parents: 
37678diff
changeset | 82 | by (simp add: prod_eq_iff scaleR_left_distrib) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 83 | show "scaleR a (scaleR b x) = scaleR (a * b) x" | 
| 44066 
d74182c93f04
rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
 huffman parents: 
37678diff
changeset | 84 | by (simp add: prod_eq_iff) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 85 | show "scaleR 1 x = x" | 
| 44066 
d74182c93f04
rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
 huffman parents: 
37678diff
changeset | 86 | by (simp add: prod_eq_iff) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 87 | qed | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 88 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 89 | end | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 90 | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68617diff
changeset | 91 | lemma module_prod_scale_eq_scaleR: "module_prod.scale (*\<^sub>R) (*\<^sub>R) = scaleR" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 92 | apply (rule ext) apply (rule ext) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 93 | apply (subst module_prod.scale_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 94 | subgoal by unfold_locales | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 95 | by (simp add: scaleR_prod_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 96 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 97 | interpretation real_vector?: vector_space_prod "scaleR::_\<Rightarrow>_\<Rightarrow>'a::real_vector" "scaleR::_\<Rightarrow>_\<Rightarrow>'b::real_vector" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68617diff
changeset | 98 |   rewrites "scale = ((*\<^sub>R)::_\<Rightarrow>_\<Rightarrow>('a \<times> 'b))"
 | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68617diff
changeset | 99 | and "module.dependent (*\<^sub>R) = dependent" | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68617diff
changeset | 100 | and "module.representation (*\<^sub>R) = representation" | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68617diff
changeset | 101 | and "module.subspace (*\<^sub>R) = subspace" | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68617diff
changeset | 102 | and "module.span (*\<^sub>R) = span" | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68617diff
changeset | 103 | and "vector_space.extend_basis (*\<^sub>R) = extend_basis" | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68617diff
changeset | 104 | and "vector_space.dim (*\<^sub>R) = dim" | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68617diff
changeset | 105 | and "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 106 | subgoal by unfold_locales | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 107 | subgoal by (fact module_prod_scale_eq_scaleR) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 108 | unfolding dependent_raw_def representation_raw_def subspace_raw_def span_raw_def | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 109 | extend_basis_raw_def dim_raw_def linear_def | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 110 | by (rule refl)+ | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 111 | |
| 69541 | 112 | subsection \<open>Product is a Metric Space\<close> | 
| 31339 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 113 | |
| 62101 | 114 | (* TODO: Product of uniform spaces and compatibility with metric_spaces! *) | 
| 115 | ||
| 70136 | 116 | instantiation\<^marker>\<open>tag unimportant\<close> prod :: (metric_space, metric_space) dist | 
| 31339 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 117 | begin | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 118 | |
| 69541 | 119 | definition dist_prod_def[code del]: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51644diff
changeset | 120 | "dist x y = sqrt ((dist (fst x) (fst y))\<^sup>2 + (dist (snd x) (snd y))\<^sup>2)" | 
| 31339 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 121 | |
| 62101 | 122 | instance .. | 
| 123 | end | |
| 124 | ||
| 70136 | 125 | instantiation\<^marker>\<open>tag unimportant\<close> prod :: (metric_space, metric_space) uniformity_dist | 
| 62101 | 126 | begin | 
| 127 | ||
| 128 | definition [code del]: | |
| 62102 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 129 |   "(uniformity :: (('a \<times> 'b) \<times> ('a \<times> 'b)) filter) =
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69064diff
changeset | 130 |     (INF e\<in>{0 <..}. principal {(x, y). dist x y < e})"
 | 
| 62101 | 131 | |
| 62102 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 132 | instance | 
| 62101 | 133 | by standard (rule uniformity_prod_def) | 
| 134 | end | |
| 135 | ||
| 62102 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 136 | declare uniformity_Abort[where 'a="'a :: metric_space \<times> 'b :: metric_space", code] | 
| 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 137 | |
| 68617 | 138 | instantiation prod :: (metric_space, metric_space) metric_space | 
| 62101 | 139 | begin | 
| 140 | ||
| 69541 | 141 | proposition dist_Pair_Pair: "dist (a, b) (c, d) = sqrt ((dist a c)\<^sup>2 + (dist b d)\<^sup>2)" | 
| 31339 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 142 | unfolding dist_prod_def by simp | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 143 | |
| 36332 | 144 | lemma dist_fst_le: "dist (fst x) (fst y) \<le> dist x y" | 
| 53930 | 145 | unfolding dist_prod_def by (rule real_sqrt_sum_squares_ge1) | 
| 36332 | 146 | |
| 147 | lemma dist_snd_le: "dist (snd x) (snd y) \<le> dist x y" | |
| 53930 | 148 | unfolding dist_prod_def by (rule real_sqrt_sum_squares_ge2) | 
| 36332 | 149 | |
| 60679 | 150 | instance | 
| 151 | proof | |
| 31339 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 152 | fix x y :: "'a \<times> 'b" | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 153 | show "dist x y = 0 \<longleftrightarrow> x = y" | 
| 44066 
d74182c93f04
rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
 huffman parents: 
37678diff
changeset | 154 | unfolding dist_prod_def prod_eq_iff by simp | 
| 31339 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 155 | next | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 156 | fix x y z :: "'a \<times> 'b" | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 157 | 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: 
31290diff
changeset | 158 | unfolding dist_prod_def | 
| 31563 | 159 | by (intro order_trans [OF _ real_sqrt_sum_squares_triangle_ineq] | 
| 160 | real_sqrt_le_mono add_mono power_mono dist_triangle2 zero_le_dist) | |
| 31415 | 161 | next | 
| 31492 
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
 huffman parents: 
31491diff
changeset | 162 |   fix S :: "('a \<times> 'b) set"
 | 
| 62101 | 163 | have *: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)" | 
| 31563 | 164 | proof | 
| 36332 | 165 | assume "open S" show "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" | 
| 166 | proof | |
| 167 | fix x assume "x \<in> S" | |
| 168 | obtain A B where "open A" "open B" "x \<in> A \<times> B" "A \<times> B \<subseteq> S" | |
| 60500 | 169 | using \<open>open S\<close> and \<open>x \<in> S\<close> by (rule open_prod_elim) | 
| 36332 | 170 | obtain r where r: "0 < r" "\<forall>y. dist y (fst x) < r \<longrightarrow> y \<in> A" | 
| 60500 | 171 | using \<open>open A\<close> and \<open>x \<in> A \<times> B\<close> unfolding open_dist by auto | 
| 36332 | 172 | obtain s where s: "0 < s" "\<forall>y. dist y (snd x) < s \<longrightarrow> y \<in> B" | 
| 60500 | 173 | using \<open>open B\<close> and \<open>x \<in> A \<times> B\<close> unfolding open_dist by auto | 
| 36332 | 174 | let ?e = "min r s" | 
| 175 | have "0 < ?e \<and> (\<forall>y. dist y x < ?e \<longrightarrow> y \<in> S)" | |
| 176 | proof (intro allI impI conjI) | |
| 177 | show "0 < min r s" by (simp add: r(1) s(1)) | |
| 178 | next | |
| 179 | fix y assume "dist y x < min r s" | |
| 180 | hence "dist y x < r" and "dist y x < s" | |
| 181 | by simp_all | |
| 182 | hence "dist (fst y) (fst x) < r" and "dist (snd y) (snd x) < s" | |
| 183 | by (auto intro: le_less_trans dist_fst_le dist_snd_le) | |
| 184 | hence "fst y \<in> A" and "snd y \<in> B" | |
| 185 | by (simp_all add: r(2) s(2)) | |
| 186 | hence "y \<in> A \<times> B" by (induct y, simp) | |
| 60500 | 187 | with \<open>A \<times> B \<subseteq> S\<close> show "y \<in> S" .. | 
| 36332 | 188 | qed | 
| 189 | thus "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" .. | |
| 190 | qed | |
| 31563 | 191 | next | 
| 44575 | 192 | assume *: "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" show "open S" | 
| 193 | proof (rule open_prod_intro) | |
| 194 | fix x assume "x \<in> S" | |
| 195 | then obtain e where "0 < e" and S: "\<forall>y. dist y x < e \<longrightarrow> y \<in> S" | |
| 196 | using * by fast | |
| 63040 | 197 | define r where "r = e / sqrt 2" | 
| 198 | define s where "s = e / sqrt 2" | |
| 60500 | 199 | from \<open>0 < e\<close> have "0 < r" and "0 < s" | 
| 56541 | 200 | unfolding r_def s_def by simp_all | 
| 60500 | 201 | from \<open>0 < e\<close> have "e = sqrt (r\<^sup>2 + s\<^sup>2)" | 
| 44575 | 202 | unfolding r_def s_def by (simp add: power_divide) | 
| 63040 | 203 |       define A where "A = {y. dist (fst x) y < r}"
 | 
| 204 |       define B where "B = {y. dist (snd x) y < s}"
 | |
| 44575 | 205 | have "open A" and "open B" | 
| 206 | unfolding A_def B_def by (simp_all add: open_ball) | |
| 207 | moreover have "x \<in> A \<times> B" | |
| 208 | unfolding A_def B_def mem_Times_iff | |
| 60500 | 209 | using \<open>0 < r\<close> and \<open>0 < s\<close> by simp | 
| 44575 | 210 | moreover have "A \<times> B \<subseteq> S" | 
| 211 | proof (clarify) | |
| 212 | fix a b assume "a \<in> A" and "b \<in> B" | |
| 213 | hence "dist a (fst x) < r" and "dist b (snd x) < s" | |
| 214 | unfolding A_def B_def by (simp_all add: dist_commute) | |
| 215 | hence "dist (a, b) x < e" | |
| 60500 | 216 | unfolding dist_prod_def \<open>e = sqrt (r\<^sup>2 + s\<^sup>2)\<close> | 
| 44575 | 217 | by (simp add: add_strict_mono power_strict_mono) | 
| 218 | thus "(a, b) \<in> S" | |
| 219 | by (simp add: S) | |
| 220 | qed | |
| 221 | ultimately show "\<exists>A B. open A \<and> open B \<and> x \<in> A \<times> B \<and> A \<times> B \<subseteq> S" by fast | |
| 222 | qed | |
| 31563 | 223 | qed | 
| 62101 | 224 | show "open S = (\<forall>x\<in>S. \<forall>\<^sub>F (x', y) in uniformity. x' = x \<longrightarrow> y \<in> S)" | 
| 225 | unfolding * eventually_uniformity_metric | |
| 226 | by (simp del: split_paired_All add: dist_prod_def dist_commute) | |
| 31339 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 227 | qed | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 228 | |
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 229 | end | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 230 | |
| 54890 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54779diff
changeset | 231 | declare [[code abort: "dist::('a::metric_space*'b::metric_space)\<Rightarrow>('a*'b) \<Rightarrow> real"]]
 | 
| 54779 
d9edb711ef31
pragmatic executability of instance prod::{open,dist,norm}
 immler parents: 
53930diff
changeset | 232 | |
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 233 | lemma Cauchy_fst: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. fst (X n))" | 
| 53930 | 234 | unfolding Cauchy_def by (fast elim: le_less_trans [OF dist_fst_le]) | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 235 | |
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 236 | lemma Cauchy_snd: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. snd (X n))" | 
| 53930 | 237 | unfolding Cauchy_def by (fast elim: le_less_trans [OF dist_snd_le]) | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 238 | |
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 239 | lemma Cauchy_Pair: | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 240 | assumes "Cauchy X" and "Cauchy Y" | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 241 | shows "Cauchy (\<lambda>n. (X n, Y n))" | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 242 | proof (rule metric_CauchyI) | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 243 | fix r :: real assume "0 < r" | 
| 56541 | 244 | hence "0 < r / sqrt 2" (is "0 < ?s") by simp | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 245 | obtain M where M: "\<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < ?s" | 
| 60500 | 246 | using metric_CauchyD [OF \<open>Cauchy X\<close> \<open>0 < ?s\<close>] .. | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 247 | obtain N where N: "\<forall>m\<ge>N. \<forall>n\<ge>N. dist (Y m) (Y n) < ?s" | 
| 60500 | 248 | using metric_CauchyD [OF \<open>Cauchy Y\<close> \<open>0 < ?s\<close>] .. | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 249 | 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: 
31388diff
changeset | 250 | 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: 
31388diff
changeset | 251 | 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: 
31388diff
changeset | 252 | qed | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 253 | |
| 69541 | 254 | subsection \<open>Product is a Complete Metric Space\<close> | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 255 | |
| 70136 | 256 | instance\<^marker>\<open>tag important\<close> prod :: (complete_space, complete_space) complete_space | 
| 257 | proof | |
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 258 | fix X :: "nat \<Rightarrow> 'a \<times> 'b" assume "Cauchy X" | 
| 61969 | 259 | have 1: "(\<lambda>n. fst (X n)) \<longlonglongrightarrow> lim (\<lambda>n. fst (X n))" | 
| 60500 | 260 | using Cauchy_fst [OF \<open>Cauchy X\<close>] | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 261 | by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) | 
| 61969 | 262 | have 2: "(\<lambda>n. snd (X n)) \<longlonglongrightarrow> lim (\<lambda>n. snd (X n))" | 
| 60500 | 263 | using Cauchy_snd [OF \<open>Cauchy X\<close>] | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 264 | by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) | 
| 61969 | 265 | have "X \<longlonglongrightarrow> (lim (\<lambda>n. fst (X n)), lim (\<lambda>n. snd (X n)))" | 
| 36660 
1cc4ab4b7ff7
make (X ----> L) an abbreviation for (X ---> L) sequentially
 huffman parents: 
36332diff
changeset | 266 | using tendsto_Pair [OF 1 2] by simp | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 267 | then show "convergent X" | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 268 | by (rule convergentI) | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 269 | qed | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 270 | |
| 69541 | 271 | subsection \<open>Product is a Normed Vector Space\<close> | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 272 | |
| 68617 | 273 | instantiation prod :: (real_normed_vector, real_normed_vector) real_normed_vector | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 274 | begin | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 275 | |
| 54779 
d9edb711ef31
pragmatic executability of instance prod::{open,dist,norm}
 immler parents: 
53930diff
changeset | 276 | definition norm_prod_def[code del]: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51644diff
changeset | 277 | "norm x = sqrt ((norm (fst x))\<^sup>2 + (norm (snd x))\<^sup>2)" | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 278 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 279 | definition sgn_prod_def: | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 280 | "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 | 281 | |
| 69541 | 282 | proposition norm_Pair: "norm (a, b) = sqrt ((norm a)\<^sup>2 + (norm b)\<^sup>2)" | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 283 | unfolding norm_prod_def by simp | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 284 | |
| 60679 | 285 | instance | 
| 286 | proof | |
| 30019 
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 "norm x = 0 \<longleftrightarrow> x = 0" | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 289 | unfolding norm_prod_def | 
| 44066 
d74182c93f04
rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
 huffman parents: 
37678diff
changeset | 290 | by (simp add: prod_eq_iff) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 291 | show "norm (x + y) \<le> norm x + norm y" | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 292 | unfolding norm_prod_def | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 293 | 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 | 294 | 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 | 295 | done | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 296 | 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 | 297 | unfolding norm_prod_def | 
| 31587 | 298 | apply (simp add: power_mult_distrib) | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
44749diff
changeset | 299 | apply (simp add: distrib_left [symmetric]) | 
| 68611 | 300 | apply (simp add: real_sqrt_mult) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 301 | done | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 302 | show "sgn x = scaleR (inverse (norm x)) x" | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 303 | by (rule sgn_prod_def) | 
| 31290 | 304 | show "dist x y = norm (x - y)" | 
| 31339 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 305 | unfolding dist_prod_def norm_prod_def | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 306 | by (simp add: dist_norm) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 307 | qed | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 308 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 309 | end | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 310 | |
| 54890 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54779diff
changeset | 311 | declare [[code abort: "norm::('a::real_normed_vector*'b::real_normed_vector) \<Rightarrow> real"]]
 | 
| 54779 
d9edb711ef31
pragmatic executability of instance prod::{open,dist,norm}
 immler parents: 
53930diff
changeset | 312 | |
| 70136 | 313 | instance\<^marker>\<open>tag important\<close> prod :: (banach, banach) banach .. | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 314 | |
| 70136 | 315 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Pair operations are linear\<close> | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 316 | |
| 69541 | 317 | lemma bounded_linear_fst: "bounded_linear fst" | 
| 44127 | 318 | using fst_add fst_scaleR | 
| 319 | by (rule bounded_linear_intro [where K=1], simp add: norm_prod_def) | |
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 320 | |
| 69541 | 321 | lemma bounded_linear_snd: "bounded_linear snd" | 
| 44127 | 322 | using snd_add snd_scaleR | 
| 323 | by (rule bounded_linear_intro [where K=1], simp add: norm_prod_def) | |
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 324 | |
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
60679diff
changeset | 325 | lemmas bounded_linear_fst_comp = bounded_linear_fst[THEN bounded_linear_compose] | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
60679diff
changeset | 326 | |
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
60679diff
changeset | 327 | lemmas bounded_linear_snd_comp = bounded_linear_snd[THEN bounded_linear_compose] | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
60679diff
changeset | 328 | |
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 329 | lemma bounded_linear_Pair: | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 330 | assumes f: "bounded_linear f" | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 331 | assumes g: "bounded_linear g" | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 332 | shows "bounded_linear (\<lambda>x. (f x, g x))" | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 333 | proof | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 334 | interpret f: bounded_linear f by fact | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 335 | interpret g: bounded_linear g by fact | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 336 | fix x y and r :: real | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 337 | 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 | 338 | by (simp add: f.add g.add) | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 339 | show "(f (r *\<^sub>R x), g (r *\<^sub>R x)) = r *\<^sub>R (f x, g x)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 340 | by (simp add: f.scale g.scale) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 341 | 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 | 342 | using f.pos_bounded by fast | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 343 | 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 | 344 | using g.pos_bounded by fast | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 345 | 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 | 346 | apply (rule allI) | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 347 | apply (simp add: norm_Pair) | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 348 | apply (rule order_trans [OF sqrt_add_le_add_sqrt], simp, simp) | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
44749diff
changeset | 349 | apply (simp add: distrib_left) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 350 | apply (rule add_mono [OF norm_f norm_g]) | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 351 | done | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 352 | 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 | 353 | qed | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 354 | |
| 70136 | 355 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Frechet derivatives involving pairs\<close> | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 356 | |
| 70137 | 357 | text\<^marker>\<open>tag important\<close> \<open>%whitespace\<close> | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68072diff
changeset | 358 | proposition has_derivative_Pair [derivative_intros]: | 
| 69541 | 359 | assumes f: "(f has_derivative f') (at x within s)" | 
| 360 | and g: "(g has_derivative g') (at x within s)" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54890diff
changeset | 361 | shows "((\<lambda>x. (f x, g x)) has_derivative (\<lambda>h. (f' h, g' h))) (at x within s)" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68072diff
changeset | 362 | proof (rule has_derivativeI_sandwich[of 1]) | 
| 44575 | 363 | show "bounded_linear (\<lambda>h. (f' h, g' h))" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54890diff
changeset | 364 | using f g by (intro bounded_linear_Pair has_derivative_bounded_linear) | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 365 | let ?Rf = "\<lambda>y. f y - f x - f' (y - x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 366 | let ?Rg = "\<lambda>y. g y - g x - g' (y - x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 367 | let ?R = "\<lambda>y. ((f y, g y) - (f x, g x) - (f' (y - x), g' (y - x)))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 368 | |
| 61973 | 369 | show "((\<lambda>y. norm (?Rf y) / norm (y - x) + norm (?Rg y) / norm (y - x)) \<longlongrightarrow> 0) (at x within s)" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54890diff
changeset | 370 | using f g by (intro tendsto_add_zero) (auto simp: has_derivative_iff_norm) | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 371 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 372 | fix y :: 'a assume "y \<noteq> x" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 373 | show "norm (?R y) / norm (y - x) \<le> norm (?Rf y) / norm (y - x) + norm (?Rg y) / norm (y - x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 374 | unfolding add_divide_distrib [symmetric] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 375 | by (simp add: norm_Pair divide_right_mono order_trans [OF sqrt_add_le_add_sqrt]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 376 | qed simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 377 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 378 | lemma differentiable_Pair [simp, derivative_intros]: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 379 | "f differentiable at x within s \<Longrightarrow> g differentiable at x within s \<Longrightarrow> | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 380 | (\<lambda>x. (f x, g x)) differentiable at x within s" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 381 | unfolding differentiable_def by (blast intro: has_derivative_Pair) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 382 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 383 | lemmas has_derivative_fst [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_fst] | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 384 | lemmas has_derivative_snd [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_snd] | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 385 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 386 | lemma has_derivative_split [derivative_intros]: | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 387 | "((\<lambda>p. f (fst p) (snd p)) has_derivative f') F \<Longrightarrow> ((\<lambda>(a, b). f a b) has_derivative f') F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51478diff
changeset | 388 | unfolding split_beta' . | 
| 44575 | 389 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 390 | |
| 70136 | 391 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Vector derivatives involving pairs\<close> | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 392 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 393 | lemma has_vector_derivative_Pair[derivative_intros]: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 394 | assumes "(f has_vector_derivative f') (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 395 | "(g has_vector_derivative g') (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 396 | shows "((\<lambda>x. (f x, g x)) has_vector_derivative (f', g')) (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 397 | using assms | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 398 | by (auto simp: has_vector_derivative_def intro!: derivative_eq_intros) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 399 | |
| 62102 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 400 | lemma | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60500diff
changeset | 401 | fixes x :: "'a::real_normed_vector" | 
| 62102 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 402 | shows norm_Pair1 [simp]: "norm (0,x) = norm x" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60500diff
changeset | 403 | and norm_Pair2 [simp]: "norm (x,0) = norm x" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60500diff
changeset | 404 | by (auto simp: norm_Pair) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60500diff
changeset | 405 | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 406 | lemma norm_commute: "norm (x,y) = norm (y,x)" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 407 | by (simp add: norm_Pair) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 408 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 409 | lemma norm_fst_le: "norm x \<le> norm (x,y)" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 410 | by (metis dist_fst_le fst_conv fst_zero norm_conv_dist) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 411 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 412 | lemma norm_snd_le: "norm y \<le> norm (x,y)" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 413 | by (metis dist_snd_le snd_conv snd_zero norm_conv_dist) | 
| 59425 | 414 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 415 | lemma norm_Pair_le: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 416 | shows "norm (x, y) \<le> norm x + norm y" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 417 | unfolding norm_Pair | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 418 | by (metis norm_ge_zero sqrt_sum_squares_le_sum) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66453diff
changeset | 419 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 420 | lemma (in vector_space_prod) span_Times_sing1: "p.span ({0} \<times> B) = {0} \<times> vs2.span B"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 421 | apply (rule p.span_unique) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 422 | subgoal by (auto intro!: vs1.span_base vs2.span_base) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 423 | subgoal using vs1.subspace_single_0 vs2.subspace_span by (rule subspace_Times) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 424 | subgoal for T | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 425 | proof safe | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 426 | fix b | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 427 |     assume subset_T: "{0} \<times> B \<subseteq> T" and subspace: "p.subspace T" and b_span: "b \<in> vs2.span B"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 428 | then obtain t r where b: "b = (\<Sum>a\<in>t. r a *b a)" and t: "finite t" "t \<subseteq> B" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 429 | by (auto simp: vs2.span_explicit) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 430 | have "(0, b) = (\<Sum>b\<in>t. scale (r b) (0, b))" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 431 | unfolding b scale_prod sum_prod | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 432 | by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 433 | also have "\<dots> \<in> T" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 434 | using \<open>t \<subseteq> B\<close> subset_T | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 435 | by (auto intro!: p.subspace_sum p.subspace_scale subspace) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 436 | finally show "(0, b) \<in> T" . | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 437 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 438 | done | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 439 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 440 | lemma (in vector_space_prod) span_Times_sing2: "p.span (A \<times> {0}) = vs1.span A \<times> {0}"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 441 | apply (rule p.span_unique) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 442 | subgoal by (auto intro!: vs1.span_base vs2.span_base) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 443 | subgoal using vs1.subspace_span vs2.subspace_single_0 by (rule subspace_Times) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 444 | subgoal for T | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 445 | proof safe | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 446 | fix a | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 447 |     assume subset_T: "A \<times> {0} \<subseteq> T" and subspace: "p.subspace T" and a_span: "a \<in> vs1.span A"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 448 | then obtain t r where a: "a = (\<Sum>a\<in>t. r a *a a)" and t: "finite t" "t \<subseteq> A" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 449 | by (auto simp: vs1.span_explicit) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 450 | have "(a, 0) = (\<Sum>a\<in>t. scale (r a) (a, 0))" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 451 | unfolding a scale_prod sum_prod | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 452 | by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 453 | also have "\<dots> \<in> T" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 454 | using \<open>t \<subseteq> A\<close> subset_T | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 455 | by (auto intro!: p.subspace_sum p.subspace_scale subspace) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 456 | finally show "(a, 0) \<in> T" . | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 457 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 458 | done | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 459 | |
| 69541 | 460 | subsection \<open>Product is Finite Dimensional\<close> | 
| 461 | ||
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 462 | lemma (in finite_dimensional_vector_space) zero_not_in_Basis[simp]: "0 \<notin> Basis" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 463 | using dependent_zero local.independent_Basis by blast | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 464 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 465 | locale finite_dimensional_vector_space_prod = vector_space_prod + finite_dimensional_vector_space_pair begin | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 466 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 467 | definition "Basis_pair = B1 \<times> {0} \<union> {0} \<times> B2"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 468 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 469 | sublocale p: finite_dimensional_vector_space scale Basis_pair | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 470 | proof unfold_locales | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 471 | show "finite Basis_pair" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 472 | by (auto intro!: finite_cartesian_product vs1.finite_Basis vs2.finite_Basis simp: Basis_pair_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 473 | show "p.independent Basis_pair" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 474 | unfolding p.dependent_def Basis_pair_def | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 475 | proof safe | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 476 | fix a | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 477 | assume a: "a \<in> B1" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 478 |     assume "(a, 0) \<in> p.span (B1 \<times> {0} \<union> {0} \<times> B2 - {(a, 0)})"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 479 |     also have "B1 \<times> {0} \<union> {0} \<times> B2 - {(a, 0)} = (B1 - {a}) \<times> {0} \<union> {0} \<times> B2"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 480 | by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 481 | finally show False | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 482 | using a vs1.dependent_def vs1.independent_Basis | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 483 | by (auto simp: p.span_Un span_Times_sing1 span_Times_sing2) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 484 | next | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 485 | fix b | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 486 | assume b: "b \<in> B2" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 487 |     assume "(0, b) \<in> p.span (B1 \<times> {0} \<union> {0} \<times> B2 - {(0, b)})"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 488 |     also have "(B1 \<times> {0} \<union> {0} \<times> B2 - {(0, b)}) = B1 \<times> {0} \<union> {0} \<times> (B2 - {b})"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 489 | by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 490 | finally show False | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 491 | using b vs2.dependent_def vs2.independent_Basis | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 492 | by (auto simp: p.span_Un span_Times_sing1 span_Times_sing2) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 493 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 494 | show "p.span Basis_pair = UNIV" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 495 | by (auto simp: p.span_Un span_Times_sing2 span_Times_sing1 vs1.span_Basis vs2.span_Basis | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 496 | Basis_pair_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 497 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 498 | |
| 69541 | 499 | proposition dim_Times: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 500 | assumes "vs1.subspace S" "vs2.subspace T" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 501 | shows "p.dim(S \<times> T) = vs1.dim S + vs2.dim T" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 502 | proof - | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 503 | interpret p1: Vector_Spaces.linear s1 scale "(\<lambda>x. (x, 0))" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 504 | by unfold_locales (auto simp: scale_def) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68617diff
changeset | 505 | interpret pair1: finite_dimensional_vector_space_pair "(*a)" B1 scale Basis_pair | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 506 | by unfold_locales | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 507 | interpret p2: Vector_Spaces.linear s2 scale "(\<lambda>x. (0, x))" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 508 | by unfold_locales (auto simp: scale_def) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68617diff
changeset | 509 | interpret pair2: finite_dimensional_vector_space_pair "(*b)" B2 scale Basis_pair | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 510 | by unfold_locales | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 511 | have ss: "p.subspace ((\<lambda>x. (x, 0)) ` S)" "p.subspace (Pair 0 ` T)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 512 | by (rule p1.subspace_image p2.subspace_image assms)+ | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 513 |   have "p.dim(S \<times> T) = p.dim({u + v |u v. u \<in> (\<lambda>x. (x, 0)) ` S \<and> v \<in> Pair 0 ` T})"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 514 | by (simp add: Times_eq_image_sum) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 515 | moreover have "p.dim ((\<lambda>x. (x, 0::'c)) ` S) = vs1.dim S" "p.dim (Pair (0::'b) ` T) = vs2.dim T" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 516 | by (simp_all add: inj_on_def p1.linear_axioms pair1.dim_image_eq p2.linear_axioms pair2.dim_image_eq) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 517 | moreover have "p.dim ((\<lambda>x. (x, 0)) ` S \<inter> Pair 0 ` T) = 0" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 518 | by (subst p.dim_eq_0) auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 519 | ultimately show ?thesis | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 520 | using p.dim_sums_Int [OF ss] by linarith | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 521 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 522 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 523 | lemma dimension_pair: "p.dimension = vs1.dimension + vs2.dimension" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 524 | using dim_Times[OF vs1.subspace_UNIV vs2.subspace_UNIV] | 
| 71174 | 525 | by (auto simp: p.dimension_def vs1.dimension_def vs2.dimension_def) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 526 | |
| 44575 | 527 | end | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 528 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67962diff
changeset | 529 | end |