| author | wenzelm | 
| Sun, 12 Jan 2025 13:27:47 +0100 | |
| changeset 81776 | c6d8db03dfdc | 
| parent 81116 | 0fb1e2dd4122 | 
| child 82489 | d35d355f7330 | 
| permissions | -rw-r--r-- | 
| 63971 
da89140186e2
HOL-Analysis: move Product_Vector and Inner_Product from Library
 hoelzl parents: 
63886diff
changeset | 1 | (* Title: HOL/Analysis/Inner_Product.thy | 
| 41959 | 2 | Author: Brian Huffman | 
| 29993 | 3 | *) | 
| 4 | ||
| 69600 | 5 | section \<open>Inner Product Spaces and Gradient Derivative\<close> | 
| 29993 | 6 | |
| 7 | theory Inner_Product | |
| 65513 | 8 | imports Complex_Main | 
| 29993 | 9 | begin | 
| 10 | ||
| 60500 | 11 | subsection \<open>Real inner product spaces\<close> | 
| 29993 | 12 | |
| 60500 | 13 | text \<open> | 
| 69597 | 14 | Temporarily relax type constraints for \<^term>\<open>open\<close>, \<^term>\<open>uniformity\<close>, | 
| 15 | \<^term>\<open>dist\<close>, and \<^term>\<open>norm\<close>. | |
| 60500 | 16 | \<close> | 
| 31492 
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
 huffman parents: 
31446diff
changeset | 17 | |
| 60500 | 18 | setup \<open>Sign.add_const_constraint | 
| 69597 | 19 | (\<^const_name>\<open>open\<close>, SOME \<^typ>\<open>'a::open set \<Rightarrow> bool\<close>)\<close> | 
| 31446 | 20 | |
| 60500 | 21 | setup \<open>Sign.add_const_constraint | 
| 69597 | 22 | (\<^const_name>\<open>dist\<close>, SOME \<^typ>\<open>'a::dist \<Rightarrow> 'a \<Rightarrow> real\<close>)\<close> | 
| 31446 | 23 | |
| 60500 | 24 | setup \<open>Sign.add_const_constraint | 
| 69597 | 25 |   (\<^const_name>\<open>uniformity\<close>, SOME \<^typ>\<open>('a::uniformity \<times> 'a) filter\<close>)\<close>
 | 
| 62101 | 26 | |
| 27 | setup \<open>Sign.add_const_constraint | |
| 69597 | 28 | (\<^const_name>\<open>norm\<close>, SOME \<^typ>\<open>'a::norm \<Rightarrow> real\<close>)\<close> | 
| 31446 | 29 | |
| 68623 | 30 | class real_inner = real_vector + sgn_div_norm + dist_norm + uniformity_dist + open_uniformity + | 
| 29993 | 31 | fixes inner :: "'a \<Rightarrow> 'a \<Rightarrow> real" | 
| 32 | assumes inner_commute: "inner x y = inner y x" | |
| 31590 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 33 | and inner_add_left: "inner (x + y) z = inner x z + inner y z" | 
| 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 34 | and inner_scaleR_left [simp]: "inner (scaleR r x) y = r * (inner x y)" | 
| 29993 | 35 | and inner_ge_zero [simp]: "0 \<le> inner x x" | 
| 36 | and inner_eq_zero_iff [simp]: "inner x x = 0 \<longleftrightarrow> x = 0" | |
| 37 | and norm_eq_sqrt_inner: "norm x = sqrt (inner x x)" | |
| 38 | begin | |
| 39 | ||
| 40 | lemma inner_zero_left [simp]: "inner 0 x = 0" | |
| 31590 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 41 | using inner_add_left [of 0 0 x] by simp | 
| 29993 | 42 | |
| 43 | lemma inner_minus_left [simp]: "inner (- x) y = - inner x y" | |
| 31590 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 44 | using inner_add_left [of x "- x" y] by simp | 
| 29993 | 45 | |
| 46 | lemma inner_diff_left: "inner (x - y) z = inner x z - inner y z" | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53938diff
changeset | 47 | using inner_add_left [of x "- y" z] by simp | 
| 29993 | 48 | |
| 64267 | 49 | lemma inner_sum_left: "inner (\<Sum>x\<in>A. f x) y = (\<Sum>x\<in>A. inner (f x) y)" | 
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 50 | by (cases "finite A", induct set: finite, simp_all add: inner_add_left) | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 51 | |
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 52 | lemma all_zero_iff [simp]: "(\<forall>u. inner x u = 0) \<longleftrightarrow> (x = 0)" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 53 | by auto (use inner_eq_zero_iff in blast) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 54 | |
| 60500 | 55 | text \<open>Transfer distributivity rules to right argument.\<close> | 
| 29993 | 56 | |
| 31590 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 57 | lemma inner_add_right: "inner x (y + z) = inner x y + inner x z" | 
| 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 58 | using inner_add_left [of y z x] by (simp only: inner_commute) | 
| 29993 | 59 | |
| 31590 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 60 | lemma inner_scaleR_right [simp]: "inner x (scaleR r y) = r * (inner x y)" | 
| 29993 | 61 | using inner_scaleR_left [of r y x] by (simp only: inner_commute) | 
| 62 | ||
| 63 | lemma inner_zero_right [simp]: "inner x 0 = 0" | |
| 64 | using inner_zero_left [of x] by (simp only: inner_commute) | |
| 65 | ||
| 66 | lemma inner_minus_right [simp]: "inner x (- y) = - inner x y" | |
| 67 | using inner_minus_left [of y x] by (simp only: inner_commute) | |
| 68 | ||
| 69 | lemma inner_diff_right: "inner x (y - z) = inner x y - inner x z" | |
| 70 | using inner_diff_left [of y z x] by (simp only: inner_commute) | |
| 71 | ||
| 64267 | 72 | lemma inner_sum_right: "inner x (\<Sum>y\<in>A. f y) = (\<Sum>y\<in>A. inner x (f y))" | 
| 73 | using inner_sum_left [of f A x] by (simp only: inner_commute) | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 74 | |
| 31590 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 75 | lemmas inner_add [algebra_simps] = inner_add_left inner_add_right | 
| 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 76 | lemmas inner_diff [algebra_simps] = inner_diff_left inner_diff_right | 
| 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 77 | lemmas inner_scaleR = inner_scaleR_left inner_scaleR_right | 
| 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 78 | |
| 60500 | 79 | text \<open>Legacy theorem names\<close> | 
| 31590 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 80 | lemmas inner_left_distrib = inner_add_left | 
| 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 81 | lemmas inner_right_distrib = inner_add_right | 
| 29993 | 82 | lemmas inner_distrib = inner_left_distrib inner_right_distrib | 
| 83 | ||
| 84 | lemma inner_gt_zero_iff [simp]: "0 < inner x x \<longleftrightarrow> x \<noteq> 0" | |
| 85 | by (simp add: order_less_le) | |
| 86 | ||
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51642diff
changeset | 87 | lemma power2_norm_eq_inner: "(norm x)\<^sup>2 = inner x x" | 
| 29993 | 88 | by (simp add: norm_eq_sqrt_inner) | 
| 89 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 90 | text \<open>Identities involving real multiplication and division.\<close> | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 91 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 92 | lemma inner_mult_left: "inner (of_real m * a) b = m * (inner a b)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 93 | by (metis real_inner_class.inner_scaleR_left scaleR_conv_of_real) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 94 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 95 | lemma inner_mult_right: "inner a (of_real m * b) = m * (inner a b)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 96 | by (metis real_inner_class.inner_scaleR_right scaleR_conv_of_real) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 97 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 98 | lemma inner_mult_left': "inner (a * of_real m) b = m * (inner a b)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 99 | by (simp add: of_real_def) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 100 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 101 | lemma inner_mult_right': "inner a (b * of_real m) = (inner a b) * m" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 102 | by (simp add: of_real_def real_inner_class.inner_scaleR_right) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 103 | |
| 30046 | 104 | lemma Cauchy_Schwarz_ineq: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51642diff
changeset | 105 | "(inner x y)\<^sup>2 \<le> inner x x * inner y y" | 
| 29993 | 106 | proof (cases) | 
| 107 | assume "y = 0" | |
| 108 | thus ?thesis by simp | |
| 109 | next | |
| 110 | assume y: "y \<noteq> 0" | |
| 111 | let ?r = "inner x y / inner y y" | |
| 112 | have "0 \<le> inner (x - scaleR ?r y) (x - scaleR ?r y)" | |
| 113 | by (rule inner_ge_zero) | |
| 114 | also have "\<dots> = inner x x - inner y x * ?r" | |
| 31590 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 115 | by (simp add: inner_diff) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51642diff
changeset | 116 | also have "\<dots> = inner x x - (inner x y)\<^sup>2 / inner y y" | 
| 29993 | 117 | by (simp add: power2_eq_square inner_commute) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51642diff
changeset | 118 | finally have "0 \<le> inner x x - (inner x y)\<^sup>2 / inner y y" . | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51642diff
changeset | 119 | hence "(inner x y)\<^sup>2 / inner y y \<le> inner x x" | 
| 29993 | 120 | by (simp add: le_diff_eq) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51642diff
changeset | 121 | thus "(inner x y)\<^sup>2 \<le> inner x x * inner y y" | 
| 29993 | 122 | by (simp add: pos_divide_le_eq y) | 
| 123 | qed | |
| 124 | ||
| 30046 | 125 | lemma Cauchy_Schwarz_ineq2: | 
| 29993 | 126 | "\<bar>inner x y\<bar> \<le> norm x * norm y" | 
| 127 | proof (rule power2_le_imp_le) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51642diff
changeset | 128 | have "(inner x y)\<^sup>2 \<le> inner x x * inner y y" | 
| 30046 | 129 | using Cauchy_Schwarz_ineq . | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51642diff
changeset | 130 | thus "\<bar>inner x y\<bar>\<^sup>2 \<le> (norm x * norm y)\<^sup>2" | 
| 29993 | 131 | by (simp add: power_mult_distrib power2_norm_eq_inner) | 
| 132 | show "0 \<le> norm x * norm y" | |
| 133 | unfolding norm_eq_sqrt_inner | |
| 134 | by (intro mult_nonneg_nonneg real_sqrt_ge_zero inner_ge_zero) | |
| 135 | qed | |
| 136 | ||
| 53938 | 137 | lemma norm_cauchy_schwarz: "inner x y \<le> norm x * norm y" | 
| 138 | using Cauchy_Schwarz_ineq2 [of x y] by auto | |
| 139 | ||
| 29993 | 140 | subclass real_normed_vector | 
| 141 | proof | |
| 142 | fix a :: real and x y :: 'a | |
| 143 | show "norm x = 0 \<longleftrightarrow> x = 0" | |
| 144 | unfolding norm_eq_sqrt_inner by simp | |
| 145 | show "norm (x + y) \<le> norm x + norm y" | |
| 146 | proof (rule power2_le_imp_le) | |
| 147 | have "inner x y \<le> norm x * norm y" | |
| 53938 | 148 | by (rule norm_cauchy_schwarz) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51642diff
changeset | 149 | thus "(norm (x + y))\<^sup>2 \<le> (norm x + norm y)\<^sup>2" | 
| 29993 | 150 | unfolding power2_sum power2_norm_eq_inner | 
| 31590 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 151 | by (simp add: inner_add inner_commute) | 
| 29993 | 152 | show "0 \<le> norm x + norm y" | 
| 44126 | 153 | unfolding norm_eq_sqrt_inner by simp | 
| 29993 | 154 | qed | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51642diff
changeset | 155 | have "sqrt (a\<^sup>2 * inner x x) = \<bar>a\<bar> * sqrt (inner x x)" | 
| 68611 | 156 | by (simp add: real_sqrt_mult) | 
| 29993 | 157 | then show "norm (a *\<^sub>R x) = \<bar>a\<bar> * norm x" | 
| 158 | unfolding norm_eq_sqrt_inner | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56381diff
changeset | 159 | by (simp add: power2_eq_square mult.assoc) | 
| 29993 | 160 | qed | 
| 161 | ||
| 162 | end | |
| 163 | ||
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 164 | lemma square_bound_lemma: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 165 | fixes x :: real | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 166 | shows "x < (1 + x) * (1 + x)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 167 | proof - | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 168 | have "(x + 1/2)\<^sup>2 + 3/4 > 0" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 169 | using zero_le_power2[of "x+1/2"] by arith | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 170 | then show ?thesis | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 171 | by (simp add: field_simps power2_eq_square) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 172 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 173 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 174 | lemma square_continuous: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 175 | fixes e :: real | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 176 | shows "e > 0 \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>y. \<bar>y - x\<bar> < d \<longrightarrow> \<bar>y * y - x * x\<bar> < e)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 177 | using isCont_power[OF continuous_ident, of x, unfolded isCont_def LIM_eq, rule_format, of e 2] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 178 | by (force simp add: power2_eq_square) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 179 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 180 | lemma norm_le: "norm x \<le> norm y \<longleftrightarrow> inner x x \<le> inner y y" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 181 | by (simp add: norm_eq_sqrt_inner) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 182 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 183 | lemma norm_lt: "norm x < norm y \<longleftrightarrow> inner x x < inner y y" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 184 | by (simp add: norm_eq_sqrt_inner) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 185 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 186 | lemma norm_eq: "norm x = norm y \<longleftrightarrow> inner x x = inner y y" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 187 | apply (subst order_eq_iff) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 188 | apply (auto simp: norm_le) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 189 | done | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 190 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 191 | lemma norm_eq_1: "norm x = 1 \<longleftrightarrow> inner x x = 1" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 192 | by (simp add: norm_eq_sqrt_inner) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 193 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 194 | lemma inner_divide_left: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 195 |   fixes a :: "'a :: {real_inner,real_div_algebra}"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 196 | shows "inner (a / of_real m) b = (inner a b) / m" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 197 | by (metis (no_types) divide_inverse inner_commute inner_scaleR_right mult.left_neutral mult.right_neutral mult_scaleR_right of_real_inverse scaleR_conv_of_real times_divide_eq_left) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 198 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 199 | lemma inner_divide_right: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 200 |   fixes a :: "'a :: {real_inner,real_div_algebra}"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 201 | shows "inner a (b / of_real m) = (inner a b) / m" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 202 | by (metis inner_commute inner_divide_left) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
60679diff
changeset | 203 | |
| 60500 | 204 | text \<open> | 
| 69597 | 205 | Re-enable constraints for \<^term>\<open>open\<close>, \<^term>\<open>uniformity\<close>, | 
| 206 | \<^term>\<open>dist\<close>, and \<^term>\<open>norm\<close>. | |
| 60500 | 207 | \<close> | 
| 31492 
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
 huffman parents: 
31446diff
changeset | 208 | |
| 60500 | 209 | setup \<open>Sign.add_const_constraint | 
| 69597 | 210 | (\<^const_name>\<open>open\<close>, SOME \<^typ>\<open>'a::topological_space set \<Rightarrow> bool\<close>)\<close> | 
| 31446 | 211 | |
| 60500 | 212 | setup \<open>Sign.add_const_constraint | 
| 69597 | 213 |   (\<^const_name>\<open>uniformity\<close>, SOME \<^typ>\<open>('a::uniform_space \<times> 'a) filter\<close>)\<close>
 | 
| 62101 | 214 | |
| 215 | setup \<open>Sign.add_const_constraint | |
| 69597 | 216 | (\<^const_name>\<open>dist\<close>, SOME \<^typ>\<open>'a::metric_space \<Rightarrow> 'a \<Rightarrow> real\<close>)\<close> | 
| 31446 | 217 | |
| 60500 | 218 | setup \<open>Sign.add_const_constraint | 
| 69597 | 219 | (\<^const_name>\<open>norm\<close>, SOME \<^typ>\<open>'a::real_normed_vector \<Rightarrow> real\<close>)\<close> | 
| 31446 | 220 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 221 | lemma bounded_bilinear_inner: | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 222 | "bounded_bilinear (inner::'a::real_inner \<Rightarrow> 'a \<Rightarrow> real)" | 
| 29993 | 223 | proof | 
| 224 | fix x y z :: 'a and r :: real | |
| 225 | show "inner (x + y) z = inner x z + inner y z" | |
| 31590 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 226 | by (rule inner_add_left) | 
| 29993 | 227 | show "inner x (y + z) = inner x y + inner x z" | 
| 31590 
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
 huffman parents: 
31492diff
changeset | 228 | by (rule inner_add_right) | 
| 29993 | 229 | show "inner (scaleR r x) y = scaleR r (inner x y)" | 
| 230 | unfolding real_scaleR_def by (rule inner_scaleR_left) | |
| 231 | show "inner x (scaleR r y) = scaleR r (inner x y)" | |
| 232 | unfolding real_scaleR_def by (rule inner_scaleR_right) | |
| 233 | show "\<exists>K. \<forall>x y::'a. norm (inner x y) \<le> norm x * norm y * K" | |
| 234 | proof | |
| 235 | show "\<forall>x y::'a. norm (inner x y) \<le> norm x * norm y * 1" | |
| 30046 | 236 | by (simp add: Cauchy_Schwarz_ineq2) | 
| 29993 | 237 | qed | 
| 238 | qed | |
| 239 | ||
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 240 | lemmas tendsto_inner [tendsto_intros] = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 241 | bounded_bilinear.tendsto [OF bounded_bilinear_inner] | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 242 | |
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 243 | lemmas isCont_inner [simp] = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 244 | bounded_bilinear.isCont [OF bounded_bilinear_inner] | 
| 29993 | 245 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56181diff
changeset | 246 | lemmas has_derivative_inner [derivative_intros] = | 
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 247 | bounded_bilinear.FDERIV [OF bounded_bilinear_inner] | 
| 29993 | 248 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 249 | lemmas bounded_linear_inner_left = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 250 | bounded_bilinear.bounded_linear_left [OF bounded_bilinear_inner] | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 251 | |
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 252 | lemmas bounded_linear_inner_right = | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 253 | bounded_bilinear.bounded_linear_right [OF bounded_bilinear_inner] | 
| 44233 | 254 | |
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61518diff
changeset | 255 | lemmas bounded_linear_inner_left_comp = bounded_linear_inner_left[THEN bounded_linear_compose] | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61518diff
changeset | 256 | |
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61518diff
changeset | 257 | lemmas bounded_linear_inner_right_comp = bounded_linear_inner_right[THEN bounded_linear_compose] | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61518diff
changeset | 258 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56181diff
changeset | 259 | lemmas has_derivative_inner_right [derivative_intros] = | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 260 | bounded_linear.has_derivative [OF bounded_linear_inner_right] | 
| 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: 
51002diff
changeset | 261 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56181diff
changeset | 262 | lemmas has_derivative_inner_left [derivative_intros] = | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 263 | bounded_linear.has_derivative [OF bounded_linear_inner_left] | 
| 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: 
51002diff
changeset | 264 | |
| 
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: 
51002diff
changeset | 265 | lemma differentiable_inner [simp]: | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 266 | "f differentiable (at x within s) \<Longrightarrow> g differentiable at x within s \<Longrightarrow> (\<lambda>x. inner (f x) (g x)) differentiable at x within s" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 267 | unfolding differentiable_def by (blast intro: has_derivative_inner) | 
| 29993 | 268 | |
| 60679 | 269 | |
| 60500 | 270 | subsection \<open>Class instances\<close> | 
| 29993 | 271 | |
| 68617 | 272 | instantiation real :: real_inner | 
| 29993 | 273 | begin | 
| 274 | ||
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68623diff
changeset | 275 | definition inner_real_def [simp]: "inner = (*)" | 
| 29993 | 276 | |
| 60679 | 277 | instance | 
| 278 | proof | |
| 29993 | 279 | fix x y z r :: real | 
| 280 | show "inner x y = inner y x" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56381diff
changeset | 281 | unfolding inner_real_def by (rule mult.commute) | 
| 29993 | 282 | show "inner (x + y) z = inner x z + inner y z" | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
44902diff
changeset | 283 | unfolding inner_real_def by (rule distrib_right) | 
| 29993 | 284 | show "inner (scaleR r x) y = r * inner x y" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56381diff
changeset | 285 | unfolding inner_real_def real_scaleR_def by (rule mult.assoc) | 
| 29993 | 286 | show "0 \<le> inner x x" | 
| 287 | unfolding inner_real_def by simp | |
| 288 | show "inner x x = 0 \<longleftrightarrow> x = 0" | |
| 289 | unfolding inner_real_def by simp | |
| 290 | show "norm x = sqrt (inner x x)" | |
| 291 | unfolding inner_real_def by simp | |
| 292 | qed | |
| 293 | ||
| 294 | end | |
| 295 | ||
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63589diff
changeset | 296 | lemma | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63589diff
changeset | 297 | shows real_inner_1_left[simp]: "inner 1 x = x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63589diff
changeset | 298 | and real_inner_1_right[simp]: "inner x 1 = x" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63589diff
changeset | 299 | by simp_all | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63589diff
changeset | 300 | |
| 68617 | 301 | instantiation complex :: real_inner | 
| 29993 | 302 | begin | 
| 303 | ||
| 304 | definition inner_complex_def: | |
| 305 | "inner x y = Re x * Re y + Im x * Im y" | |
| 306 | ||
| 60679 | 307 | instance | 
| 308 | proof | |
| 29993 | 309 | fix x y z :: complex and r :: real | 
| 310 | show "inner x y = inner y x" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56381diff
changeset | 311 | unfolding inner_complex_def by (simp add: mult.commute) | 
| 29993 | 312 | show "inner (x + y) z = inner x z + inner y z" | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
44902diff
changeset | 313 | unfolding inner_complex_def by (simp add: distrib_right) | 
| 29993 | 314 | show "inner (scaleR r x) y = r * inner x y" | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
44902diff
changeset | 315 | unfolding inner_complex_def by (simp add: distrib_left) | 
| 29993 | 316 | show "0 \<le> inner x x" | 
| 44126 | 317 | unfolding inner_complex_def by simp | 
| 29993 | 318 | show "inner x x = 0 \<longleftrightarrow> x = 0" | 
| 319 | unfolding inner_complex_def | |
| 68499 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 paulson <lp15@cam.ac.uk> parents: 
68073diff
changeset | 320 | by (simp add: add_nonneg_eq_0_iff complex_eq_iff) | 
| 29993 | 321 | show "norm x = sqrt (inner x x)" | 
| 68499 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 paulson <lp15@cam.ac.uk> parents: 
68073diff
changeset | 322 | unfolding inner_complex_def norm_complex_def | 
| 29993 | 323 | by (simp add: power2_eq_square) | 
| 324 | qed | |
| 325 | ||
| 326 | end | |
| 327 | ||
| 44902 
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
 huffman parents: 
44282diff
changeset | 328 | lemma complex_inner_1 [simp]: "inner 1 x = Re x" | 
| 
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
 huffman parents: 
44282diff
changeset | 329 | unfolding inner_complex_def by simp | 
| 
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
 huffman parents: 
44282diff
changeset | 330 | |
| 
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
 huffman parents: 
44282diff
changeset | 331 | lemma complex_inner_1_right [simp]: "inner x 1 = Re x" | 
| 
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
 huffman parents: 
44282diff
changeset | 332 | unfolding inner_complex_def by simp | 
| 
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
 huffman parents: 
44282diff
changeset | 333 | |
| 65064 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 paulson <lp15@cam.ac.uk> parents: 
64267diff
changeset | 334 | lemma complex_inner_i_left [simp]: "inner \<i> x = Im x" | 
| 44902 
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
 huffman parents: 
44282diff
changeset | 335 | unfolding inner_complex_def by simp | 
| 
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
 huffman parents: 
44282diff
changeset | 336 | |
| 65064 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 paulson <lp15@cam.ac.uk> parents: 
64267diff
changeset | 337 | lemma complex_inner_i_right [simp]: "inner x \<i> = Im x" | 
| 44902 
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
 huffman parents: 
44282diff
changeset | 338 | unfolding inner_complex_def by simp | 
| 
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
 huffman parents: 
44282diff
changeset | 339 | |
| 29993 | 340 | |
| 66486 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 341 | lemma dot_square_norm: "inner x x = (norm x)\<^sup>2" | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 342 | by (simp only: power2_norm_eq_inner) (* TODO: move? *) | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 343 | |
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 344 | lemma norm_eq_square: "norm x = a \<longleftrightarrow> 0 \<le> a \<and> inner x x = a\<^sup>2" | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 345 | by (auto simp add: norm_eq_sqrt_inner) | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 346 | |
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 347 | lemma norm_le_square: "norm x \<le> a \<longleftrightarrow> 0 \<le> a \<and> inner x x \<le> a\<^sup>2" | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 348 | apply (simp add: dot_square_norm abs_le_square_iff[symmetric]) | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 349 | using norm_ge_zero[of x] | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 350 | apply arith | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 351 | done | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 352 | |
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 353 | lemma norm_ge_square: "norm x \<ge> a \<longleftrightarrow> a \<le> 0 \<or> inner x x \<ge> a\<^sup>2" | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 354 | apply (simp add: dot_square_norm abs_le_square_iff[symmetric]) | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 355 | using norm_ge_zero[of x] | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 356 | apply arith | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 357 | done | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 358 | |
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 359 | lemma norm_lt_square: "norm x < a \<longleftrightarrow> 0 < a \<and> inner x x < a\<^sup>2" | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 360 | by (metis not_le norm_ge_square) | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 361 | |
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 362 | lemma norm_gt_square: "norm x > a \<longleftrightarrow> a < 0 \<or> inner x x > a\<^sup>2" | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 363 | by (metis norm_le_square not_less) | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 364 | |
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 365 | text\<open>Dot product in terms of the norm rather than conversely.\<close> | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 366 | |
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 367 | lemmas inner_simps = inner_add_left inner_add_right inner_diff_right inner_diff_left | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 368 | inner_scaleR_left inner_scaleR_right | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 369 | |
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 370 | lemma dot_norm: "inner x y = ((norm (x + y))\<^sup>2 - (norm x)\<^sup>2 - (norm y)\<^sup>2) / 2" | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 371 | by (simp only: power2_norm_eq_inner inner_simps inner_commute) auto | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 372 | |
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 373 | lemma dot_norm_neg: "inner x y = (((norm x)\<^sup>2 + (norm y)\<^sup>2) - (norm (x - y))\<^sup>2) / 2" | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 374 | by (simp only: power2_norm_eq_inner inner_simps inner_commute) | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 375 | (auto simp add: algebra_simps) | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 376 | |
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 377 | lemma of_real_inner_1 [simp]: | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 378 |   "inner (of_real x) (1 :: 'a :: {real_inner, real_normed_algebra_1}) = x"
 | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 379 | by (simp add: of_real_def dot_square_norm) | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 380 | |
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 381 | lemma summable_of_real_iff: | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 382 |   "summable (\<lambda>x. of_real (f x) :: 'a :: {real_normed_algebra_1,real_inner}) \<longleftrightarrow> summable f"
 | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 383 | proof | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 384 | assume *: "summable (\<lambda>x. of_real (f x) :: 'a)" | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 385 | interpret bounded_linear "\<lambda>x::'a. inner x 1" | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 386 | by (rule bounded_linear_inner_left) | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 387 | from summable [OF *] show "summable f" by simp | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 388 | qed (auto intro: summable_of_real) | 
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 389 | |
| 
ffaaa83543b2
Lemmas about analysis and permutations
 Manuel Eberl <eberlm@in.tum.de> parents: 
65513diff
changeset | 390 | |
| 60500 | 391 | subsection \<open>Gradient derivative\<close> | 
| 29993 | 392 | |
| 70136 | 393 | definition\<^marker>\<open>tag important\<close> | 
| 81097 | 394 | gderiv :: "['a::real_inner \<Rightarrow> real, 'a, 'a] \<Rightarrow> bool" | 
| 395 | (\<open>(\<open>notation=\<open>mixfix GDERIV\<close>\<close>GDERIV (_)/ (_)/ :> (_))\<close> [1000, 1000, 60] 60) | |
| 29993 | 396 | where | 
| 397 | "GDERIV f x :> D \<longleftrightarrow> FDERIV f x :> (\<lambda>h. inner h D)" | |
| 398 | ||
| 399 | lemma gderiv_deriv [simp]: "GDERIV f x :> D \<longleftrightarrow> DERIV f x :> D" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 400 | by (simp only: gderiv_def has_field_derivative_def inner_real_def mult_commute_abs) | 
| 29993 | 401 | |
| 402 | lemma GDERIV_DERIV_compose: | |
| 403 | "\<lbrakk>GDERIV f x :> df; DERIV g (f x) :> dg\<rbrakk> | |
| 404 | \<Longrightarrow> GDERIV (\<lambda>x. g (f x)) x :> scaleR dg df" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 405 | unfolding gderiv_def has_field_derivative_def | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 406 | apply (drule (1) has_derivative_compose) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 407 | apply (simp add: ac_simps) | 
| 29993 | 408 | done | 
| 409 | ||
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 410 | lemma has_derivative_subst: "\<lbrakk>FDERIV f x :> df; df = d\<rbrakk> \<Longrightarrow> FDERIV f x :> d" | 
| 29993 | 411 | by simp | 
| 412 | ||
| 413 | lemma GDERIV_subst: "\<lbrakk>GDERIV f x :> df; df = d\<rbrakk> \<Longrightarrow> GDERIV f x :> d" | |
| 414 | by simp | |
| 415 | ||
| 416 | lemma GDERIV_const: "GDERIV (\<lambda>x. k) x :> 0" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 417 | unfolding gderiv_def inner_zero_right by (rule has_derivative_const) | 
| 29993 | 418 | |
| 419 | lemma GDERIV_add: | |
| 420 | "\<lbrakk>GDERIV f x :> df; GDERIV g x :> dg\<rbrakk> | |
| 421 | \<Longrightarrow> GDERIV (\<lambda>x. f x + g x) x :> df + dg" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 422 | unfolding gderiv_def inner_add_right by (rule has_derivative_add) | 
| 29993 | 423 | |
| 424 | lemma GDERIV_minus: | |
| 425 | "GDERIV f x :> df \<Longrightarrow> GDERIV (\<lambda>x. - f x) x :> - df" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 426 | unfolding gderiv_def inner_minus_right by (rule has_derivative_minus) | 
| 29993 | 427 | |
| 428 | lemma GDERIV_diff: | |
| 429 | "\<lbrakk>GDERIV f x :> df; GDERIV g x :> dg\<rbrakk> | |
| 430 | \<Longrightarrow> GDERIV (\<lambda>x. f x - g x) x :> df - dg" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 431 | unfolding gderiv_def inner_diff_right by (rule has_derivative_diff) | 
| 29993 | 432 | |
| 433 | lemma GDERIV_scaleR: | |
| 434 | "\<lbrakk>DERIV f x :> df; GDERIV g x :> dg\<rbrakk> | |
| 435 | \<Longrightarrow> GDERIV (\<lambda>x. scaleR (f x) (g x)) x | |
| 436 | :> (scaleR (f x) dg + scaleR df (g x))" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 437 | unfolding gderiv_def has_field_derivative_def inner_add_right inner_scaleR_right | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 438 | apply (rule has_derivative_subst) | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 439 | apply (erule (1) has_derivative_scaleR) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 440 | apply (simp add: ac_simps) | 
| 29993 | 441 | done | 
| 442 | ||
| 443 | lemma GDERIV_mult: | |
| 444 | "\<lbrakk>GDERIV f x :> df; GDERIV g x :> dg\<rbrakk> | |
| 445 | \<Longrightarrow> GDERIV (\<lambda>x. f x * g x) x :> scaleR (f x) dg + scaleR (g x) df" | |
| 446 | unfolding gderiv_def | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 447 | apply (rule has_derivative_subst) | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 448 | apply (erule (1) has_derivative_mult) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 449 | apply (simp add: inner_add ac_simps) | 
| 29993 | 450 | done | 
| 451 | ||
| 452 | lemma GDERIV_inverse: | |
| 453 | "\<lbrakk>GDERIV f x :> df; f x \<noteq> 0\<rbrakk> | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51642diff
changeset | 454 | \<Longrightarrow> GDERIV (\<lambda>x. inverse (f x)) x :> - (inverse (f x))\<^sup>2 *\<^sub>R df" | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 455 | by (metis DERIV_inverse GDERIV_DERIV_compose numerals(2)) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 456 | |
| 29993 | 457 | lemma GDERIV_norm: | 
| 458 | assumes "x \<noteq> 0" shows "GDERIV (\<lambda>x. norm x) x :> sgn x" | |
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 459 | unfolding gderiv_def norm_eq_sqrt_inner | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 460 | by (rule derivative_eq_intros | force simp add: inner_commute sgn_div_norm norm_eq_sqrt_inner assms)+ | 
| 29993 | 461 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
54230diff
changeset | 462 | lemmas has_derivative_norm = GDERIV_norm [unfolded gderiv_def] | 
| 29993 | 463 | |
| 81100 | 464 | bundle inner_syntax | 
| 465 | begin | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
70136diff
changeset | 466 | notation inner (infix \<open>\<bullet>\<close> 70) | 
| 29993 | 467 | end | 
| 69674 | 468 | |
| 469 | end |