| author | wenzelm | 
| Sun, 07 Jan 2018 16:55:45 +0100 | |
| changeset 67362 | 221612c942de | 
| parent 66453 | cc19f7ca2ed6 | 
| child 67685 | bdff8bf0a75b | 
| 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 | 
| 63971 
da89140186e2
HOL-Analysis: move Product_Vector and Inner_Product from Library
 hoelzl parents: 
63040diff
changeset | 8 | imports | 
| 
da89140186e2
HOL-Analysis: move Product_Vector and Inner_Product from Library
 hoelzl parents: 
63040diff
changeset | 9 | Inner_Product | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
63972diff
changeset | 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 | |
| 60500 | 13 | 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 | 14 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
36661diff
changeset | 15 | instantiation prod :: (real_vector, real_vector) real_vector | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 16 | begin | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 17 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 18 | definition scaleR_prod_def: | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 19 | "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 | 20 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 21 | 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 | 22 | unfolding scaleR_prod_def by simp | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 23 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 24 | 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 | 25 | unfolding scaleR_prod_def by simp | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 26 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 27 | 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 | 28 | unfolding scaleR_prod_def by simp | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 29 | |
| 60679 | 30 | instance | 
| 31 | proof | |
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 32 | 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 | 33 | 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 | 34 | 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 | 35 | 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 | 36 | 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 | 37 | 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 | 38 | by (simp add: prod_eq_iff) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 39 | show "scaleR 1 x = x" | 
| 44066 
d74182c93f04
rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
 huffman parents: 
37678diff
changeset | 40 | by (simp add: prod_eq_iff) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 41 | qed | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 42 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 43 | end | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 44 | |
| 60500 | 45 | 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 | 46 | |
| 62101 | 47 | (* TODO: Product of uniform spaces and compatibility with metric_spaces! *) | 
| 48 | ||
| 49 | instantiation 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 | 50 | begin | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 51 | |
| 54779 
d9edb711ef31
pragmatic executability of instance prod::{open,dist,norm}
 immler parents: 
53930diff
changeset | 52 | 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 | 53 | "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 | 54 | |
| 62101 | 55 | instance .. | 
| 56 | end | |
| 57 | ||
| 58 | instantiation prod :: (metric_space, metric_space) uniformity_dist | |
| 59 | begin | |
| 60 | ||
| 61 | definition [code del]: | |
| 62102 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 62 |   "(uniformity :: (('a \<times> 'b) \<times> ('a \<times> 'b)) filter) =
 | 
| 62101 | 63 |     (INF e:{0 <..}. principal {(x, y). dist x y < e})"
 | 
| 64 | ||
| 62102 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 65 | instance | 
| 62101 | 66 | by standard (rule uniformity_prod_def) | 
| 67 | end | |
| 68 | ||
| 62102 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 69 | 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 | 70 | |
| 62101 | 71 | instantiation prod :: (metric_space, metric_space) metric_space | 
| 72 | begin | |
| 73 | ||
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51644diff
changeset | 74 | lemma 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 | 75 | unfolding dist_prod_def by simp | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 76 | |
| 36332 | 77 | lemma dist_fst_le: "dist (fst x) (fst y) \<le> dist x y" | 
| 53930 | 78 | unfolding dist_prod_def by (rule real_sqrt_sum_squares_ge1) | 
| 36332 | 79 | |
| 80 | lemma dist_snd_le: "dist (snd x) (snd y) \<le> dist x y" | |
| 53930 | 81 | unfolding dist_prod_def by (rule real_sqrt_sum_squares_ge2) | 
| 36332 | 82 | |
| 60679 | 83 | instance | 
| 84 | proof | |
| 31339 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 85 | fix x y :: "'a \<times> 'b" | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 86 | 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 | 87 | 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 | 88 | next | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 89 | 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 | 90 | 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 | 91 | unfolding dist_prod_def | 
| 31563 | 92 | by (intro order_trans [OF _ real_sqrt_sum_squares_triangle_ineq] | 
| 93 | real_sqrt_le_mono add_mono power_mono dist_triangle2 zero_le_dist) | |
| 31415 | 94 | next | 
| 31492 
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
 huffman parents: 
31491diff
changeset | 95 |   fix S :: "('a \<times> 'b) set"
 | 
| 62101 | 96 | have *: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)" | 
| 31563 | 97 | proof | 
| 36332 | 98 | assume "open S" show "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" | 
| 99 | proof | |
| 100 | fix x assume "x \<in> S" | |
| 101 | obtain A B where "open A" "open B" "x \<in> A \<times> B" "A \<times> B \<subseteq> S" | |
| 60500 | 102 | using \<open>open S\<close> and \<open>x \<in> S\<close> by (rule open_prod_elim) | 
| 36332 | 103 | obtain r where r: "0 < r" "\<forall>y. dist y (fst x) < r \<longrightarrow> y \<in> A" | 
| 60500 | 104 | using \<open>open A\<close> and \<open>x \<in> A \<times> B\<close> unfolding open_dist by auto | 
| 36332 | 105 | obtain s where s: "0 < s" "\<forall>y. dist y (snd x) < s \<longrightarrow> y \<in> B" | 
| 60500 | 106 | using \<open>open B\<close> and \<open>x \<in> A \<times> B\<close> unfolding open_dist by auto | 
| 36332 | 107 | let ?e = "min r s" | 
| 108 | have "0 < ?e \<and> (\<forall>y. dist y x < ?e \<longrightarrow> y \<in> S)" | |
| 109 | proof (intro allI impI conjI) | |
| 110 | show "0 < min r s" by (simp add: r(1) s(1)) | |
| 111 | next | |
| 112 | fix y assume "dist y x < min r s" | |
| 113 | hence "dist y x < r" and "dist y x < s" | |
| 114 | by simp_all | |
| 115 | hence "dist (fst y) (fst x) < r" and "dist (snd y) (snd x) < s" | |
| 116 | by (auto intro: le_less_trans dist_fst_le dist_snd_le) | |
| 117 | hence "fst y \<in> A" and "snd y \<in> B" | |
| 118 | by (simp_all add: r(2) s(2)) | |
| 119 | hence "y \<in> A \<times> B" by (induct y, simp) | |
| 60500 | 120 | with \<open>A \<times> B \<subseteq> S\<close> show "y \<in> S" .. | 
| 36332 | 121 | qed | 
| 122 | thus "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" .. | |
| 123 | qed | |
| 31563 | 124 | next | 
| 44575 | 125 | assume *: "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" show "open S" | 
| 126 | proof (rule open_prod_intro) | |
| 127 | fix x assume "x \<in> S" | |
| 128 | then obtain e where "0 < e" and S: "\<forall>y. dist y x < e \<longrightarrow> y \<in> S" | |
| 129 | using * by fast | |
| 63040 | 130 | define r where "r = e / sqrt 2" | 
| 131 | define s where "s = e / sqrt 2" | |
| 60500 | 132 | from \<open>0 < e\<close> have "0 < r" and "0 < s" | 
| 56541 | 133 | unfolding r_def s_def by simp_all | 
| 60500 | 134 | from \<open>0 < e\<close> have "e = sqrt (r\<^sup>2 + s\<^sup>2)" | 
| 44575 | 135 | unfolding r_def s_def by (simp add: power_divide) | 
| 63040 | 136 |       define A where "A = {y. dist (fst x) y < r}"
 | 
| 137 |       define B where "B = {y. dist (snd x) y < s}"
 | |
| 44575 | 138 | have "open A" and "open B" | 
| 139 | unfolding A_def B_def by (simp_all add: open_ball) | |
| 140 | moreover have "x \<in> A \<times> B" | |
| 141 | unfolding A_def B_def mem_Times_iff | |
| 60500 | 142 | using \<open>0 < r\<close> and \<open>0 < s\<close> by simp | 
| 44575 | 143 | moreover have "A \<times> B \<subseteq> S" | 
| 144 | proof (clarify) | |
| 145 | fix a b assume "a \<in> A" and "b \<in> B" | |
| 146 | hence "dist a (fst x) < r" and "dist b (snd x) < s" | |
| 147 | unfolding A_def B_def by (simp_all add: dist_commute) | |
| 148 | hence "dist (a, b) x < e" | |
| 60500 | 149 | unfolding dist_prod_def \<open>e = sqrt (r\<^sup>2 + s\<^sup>2)\<close> | 
| 44575 | 150 | by (simp add: add_strict_mono power_strict_mono) | 
| 151 | thus "(a, b) \<in> S" | |
| 152 | by (simp add: S) | |
| 153 | qed | |
| 154 | ultimately show "\<exists>A B. open A \<and> open B \<and> x \<in> A \<times> B \<and> A \<times> B \<subseteq> S" by fast | |
| 155 | qed | |
| 31563 | 156 | qed | 
| 62101 | 157 | show "open S = (\<forall>x\<in>S. \<forall>\<^sub>F (x', y) in uniformity. x' = x \<longrightarrow> y \<in> S)" | 
| 158 | unfolding * eventually_uniformity_metric | |
| 159 | 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 | 160 | qed | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 161 | |
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 162 | end | 
| 
b4660351e8e7
instance * :: (metric_space, metric_space) metric_space; generalize lemmas to class metric_space
 huffman parents: 
31290diff
changeset | 163 | |
| 54890 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54779diff
changeset | 164 | 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 | 165 | |
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 166 | lemma Cauchy_fst: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. fst (X n))" | 
| 53930 | 167 | 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 | 168 | |
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 169 | lemma Cauchy_snd: "Cauchy X \<Longrightarrow> Cauchy (\<lambda>n. snd (X n))" | 
| 53930 | 170 | 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 | 171 | |
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 172 | lemma Cauchy_Pair: | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 173 | assumes "Cauchy X" and "Cauchy Y" | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 174 | shows "Cauchy (\<lambda>n. (X n, Y n))" | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 175 | proof (rule metric_CauchyI) | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 176 | fix r :: real assume "0 < r" | 
| 56541 | 177 | 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 | 178 | obtain M where M: "\<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < ?s" | 
| 60500 | 179 | 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 | 180 | obtain N where N: "\<forall>m\<ge>N. \<forall>n\<ge>N. dist (Y m) (Y n) < ?s" | 
| 60500 | 181 | 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 | 182 | 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 | 183 | 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 | 184 | 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 | 185 | qed | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 186 | |
| 60500 | 187 | 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 | 188 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
36661diff
changeset | 189 | instance prod :: (complete_space, complete_space) complete_space | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 190 | proof | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 191 | fix X :: "nat \<Rightarrow> 'a \<times> 'b" assume "Cauchy X" | 
| 61969 | 192 | have 1: "(\<lambda>n. fst (X n)) \<longlonglongrightarrow> lim (\<lambda>n. fst (X n))" | 
| 60500 | 193 | using Cauchy_fst [OF \<open>Cauchy X\<close>] | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 194 | by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) | 
| 61969 | 195 | have 2: "(\<lambda>n. snd (X n)) \<longlonglongrightarrow> lim (\<lambda>n. snd (X n))" | 
| 60500 | 196 | using Cauchy_snd [OF \<open>Cauchy X\<close>] | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 197 | by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) | 
| 61969 | 198 | 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 | 199 | using tendsto_Pair [OF 1 2] by simp | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 200 | then show "convergent X" | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 201 | by (rule convergentI) | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 202 | qed | 
| 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 203 | |
| 60500 | 204 | 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 | 205 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
36661diff
changeset | 206 | 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 | 207 | begin | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 208 | |
| 54779 
d9edb711ef31
pragmatic executability of instance prod::{open,dist,norm}
 immler parents: 
53930diff
changeset | 209 | 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 | 210 | "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 | 211 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 212 | definition sgn_prod_def: | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 213 | "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 | 214 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51644diff
changeset | 215 | lemma 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 | 216 | unfolding norm_prod_def by simp | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 217 | |
| 60679 | 218 | instance | 
| 219 | proof | |
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 220 | fix r :: real and x y :: "'a \<times> 'b" | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 221 | show "norm x = 0 \<longleftrightarrow> x = 0" | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 222 | unfolding norm_prod_def | 
| 44066 
d74182c93f04
rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
 huffman parents: 
37678diff
changeset | 223 | by (simp add: prod_eq_iff) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 224 | show "norm (x + y) \<le> norm x + norm y" | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 225 | unfolding norm_prod_def | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 226 | 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 | 227 | 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 | 228 | done | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 229 | 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 | 230 | unfolding norm_prod_def | 
| 31587 | 231 | apply (simp add: power_mult_distrib) | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
44749diff
changeset | 232 | apply (simp add: distrib_left [symmetric]) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 233 | apply (simp add: real_sqrt_mult_distrib) | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 234 | done | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 235 | show "sgn x = scaleR (inverse (norm x)) x" | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 236 | by (rule sgn_prod_def) | 
| 31290 | 237 | 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 | 238 | 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 | 239 | by (simp add: dist_norm) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 240 | qed | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 241 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 242 | end | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 243 | |
| 54890 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54779diff
changeset | 244 | 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 | 245 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
36661diff
changeset | 246 | instance prod :: (banach, banach) banach .. | 
| 31405 
1f72869f1a2e
instance * :: complete_space; generalize continuity lemmas for fst, snd, Pair
 huffman parents: 
31388diff
changeset | 247 | |
| 60500 | 248 | subsubsection \<open>Pair operations are linear\<close> | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 249 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 250 | lemma bounded_linear_fst: "bounded_linear fst" | 
| 44127 | 251 | using fst_add fst_scaleR | 
| 252 | 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 | 253 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 254 | lemma bounded_linear_snd: "bounded_linear snd" | 
| 44127 | 255 | using snd_add snd_scaleR | 
| 256 | 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 | 257 | |
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
60679diff
changeset | 258 | 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 | 259 | |
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
60679diff
changeset | 260 | 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 | 261 | |
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 262 | lemma bounded_linear_Pair: | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 263 | assumes f: "bounded_linear f" | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 264 | assumes g: "bounded_linear g" | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 265 | shows "bounded_linear (\<lambda>x. (f x, g x))" | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 266 | proof | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 267 | interpret f: bounded_linear f by fact | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 268 | interpret g: bounded_linear g by fact | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 269 | fix x y and r :: real | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 270 | 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 | 271 | by (simp add: f.add g.add) | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 272 | 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 | 273 | by (simp add: f.scaleR g.scaleR) | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 274 | 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 | 275 | using f.pos_bounded by fast | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 276 | 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 | 277 | using g.pos_bounded by fast | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 278 | 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 | 279 | apply (rule allI) | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 280 | apply (simp add: norm_Pair) | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 281 | 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 | 282 | apply (simp add: distrib_left) | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 283 | apply (rule add_mono [OF norm_f norm_g]) | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 284 | done | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 285 | 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 | 286 | qed | 
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 287 | |
| 60500 | 288 | subsubsection \<open>Frechet derivatives involving pairs\<close> | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 289 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 290 | lemma has_derivative_Pair [derivative_intros]: | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54890diff
changeset | 291 | assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at x within s)" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54890diff
changeset | 292 | shows "((\<lambda>x. (f x, g x)) has_derivative (\<lambda>h. (f' h, g' h))) (at x within s)" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54890diff
changeset | 293 | proof (rule has_derivativeI_sandwich[of 1]) | 
| 44575 | 294 | show "bounded_linear (\<lambda>h. (f' h, g' h))" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54890diff
changeset | 295 | 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 | 296 | 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 | 297 | 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 | 298 | 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 | 299 | |
| 61973 | 300 | 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 | 301 | 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 | 302 | |
| 
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 | 303 | 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 | 304 | 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 | 305 | 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 | 306 | 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 | 307 | 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 | 308 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 309 | 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 | 310 | 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 | 311 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 312 | 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 | 313 | "((\<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 | 314 | unfolding split_beta' . | 
| 44575 | 315 | |
| 60500 | 316 | subsection \<open>Product is an inner product space\<close> | 
| 44575 | 317 | |
| 318 | instantiation prod :: (real_inner, real_inner) real_inner | |
| 319 | begin | |
| 320 | ||
| 321 | definition inner_prod_def: | |
| 322 | "inner x y = inner (fst x) (fst y) + inner (snd x) (snd y)" | |
| 323 | ||
| 324 | lemma inner_Pair [simp]: "inner (a, b) (c, d) = inner a c + inner b d" | |
| 325 | unfolding inner_prod_def by simp | |
| 326 | ||
| 60679 | 327 | instance | 
| 328 | proof | |
| 44575 | 329 | fix r :: real | 
| 330 | fix x y z :: "'a::real_inner \<times> 'b::real_inner" | |
| 331 | show "inner x y = inner y x" | |
| 332 | unfolding inner_prod_def | |
| 333 | by (simp add: inner_commute) | |
| 334 | show "inner (x + y) z = inner x z + inner y z" | |
| 335 | unfolding inner_prod_def | |
| 336 | by (simp add: inner_add_left) | |
| 337 | show "inner (scaleR r x) y = r * inner x y" | |
| 338 | unfolding inner_prod_def | |
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
44749diff
changeset | 339 | by (simp add: distrib_left) | 
| 44575 | 340 | show "0 \<le> inner x x" | 
| 341 | unfolding inner_prod_def | |
| 342 | by (intro add_nonneg_nonneg inner_ge_zero) | |
| 343 | show "inner x x = 0 \<longleftrightarrow> x = 0" | |
| 344 | unfolding inner_prod_def prod_eq_iff | |
| 345 | by (simp add: add_nonneg_eq_0_iff) | |
| 346 | show "norm x = sqrt (inner x x)" | |
| 347 | unfolding norm_prod_def inner_prod_def | |
| 348 | by (simp add: power2_norm_eq_inner) | |
| 349 | qed | |
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 350 | |
| 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: diff
changeset | 351 | end | 
| 44575 | 352 | |
| 59425 | 353 | lemma inner_Pair_0: "inner x (0, b) = inner (snd x) b" "inner x (a, 0) = inner (fst x) a" | 
| 354 | by (cases x, simp)+ | |
| 355 | ||
| 62102 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 356 | lemma | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60500diff
changeset | 357 | fixes x :: "'a::real_normed_vector" | 
| 62102 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 358 | 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 | 359 | 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 | 360 | 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 | 361 | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 362 | lemma norm_commute: "norm (x,y) = norm (y,x)" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 363 | by (simp add: norm_Pair) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 364 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 365 | lemma norm_fst_le: "norm x \<le> norm (x,y)" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 366 | 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 | 367 | |
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 368 | lemma norm_snd_le: "norm y \<le> norm (x,y)" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 369 | by (metis dist_snd_le snd_conv snd_zero norm_conv_dist) | 
| 59425 | 370 | |
| 44575 | 371 | end |