author | nipkow |
Thu, 12 Jul 2018 11:23:46 +0200 | |
changeset 68617 | 75129a73aca3 |
parent 68611 | 4bc4b5c0ccfc |
child 68623 | b942da0962c2 |
permissions | -rw-r--r-- |
63971
da89140186e2
HOL-Analysis: move Product_Vector and Inner_Product from Library
hoelzl
parents:
63886
diff
changeset
|
1 |
(* Title: HOL/Analysis/Inner_Product.thy |
41959 | 2 |
Author: Brian Huffman |
29993 | 3 |
*) |
4 |
||
60500 | 5 |
section \<open>Inner Product Spaces and the 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> |
62101 | 14 |
Temporarily relax type constraints for @{term "open"}, @{term "uniformity"}, |
31492
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
15 |
@{term dist}, and @{term norm}. |
60500 | 16 |
\<close> |
31492
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
17 |
|
60500 | 18 |
setup \<open>Sign.add_const_constraint |
19 |
(@{const_name "open"}, SOME @{typ "'a::open set \<Rightarrow> bool"})\<close> |
|
31446 | 20 |
|
60500 | 21 |
setup \<open>Sign.add_const_constraint |
22 |
(@{const_name dist}, SOME @{typ "'a::dist \<Rightarrow> 'a \<Rightarrow> real"})\<close> |
|
31446 | 23 |
|
60500 | 24 |
setup \<open>Sign.add_const_constraint |
62101 | 25 |
(@{const_name uniformity}, SOME @{typ "('a::uniformity \<times> 'a) filter"})\<close> |
26 |
||
27 |
setup \<open>Sign.add_const_constraint |
|
60500 | 28 |
(@{const_name norm}, SOME @{typ "'a::norm \<Rightarrow> real"})\<close> |
31446 | 29 |
|
67962 | 30 |
class%important 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:
31492
diff
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:
31492
diff
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:
31492
diff
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:
31492
diff
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:
53938
diff
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:
44233
diff
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:
44233
diff
changeset
|
51 |
|
67986
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
paulson <lp15@cam.ac.uk>
parents:
67962
diff
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:
67962
diff
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:
67962
diff
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:
31492
diff
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:
31492
diff
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:
31492
diff
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:
44233
diff
changeset
|
74 |
|
31590
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents:
31492
diff
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:
31492
diff
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:
31492
diff
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:
31492
diff
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:
31492
diff
changeset
|
80 |
lemmas inner_left_distrib = inner_add_left |
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents:
31492
diff
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:
51642
diff
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:
60679
diff
changeset
|
90 |
text \<open>Identities involving real multiplication and division.\<close> |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
changeset
|
91 |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
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:
60679
diff
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:
60679
diff
changeset
|
94 |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
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:
60679
diff
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:
60679
diff
changeset
|
97 |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
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:
60679
diff
changeset
|
99 |
by (simp add: of_real_def) |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
changeset
|
100 |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
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:
60679
diff
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:
60679
diff
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:
51642
diff
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:
31492
diff
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:
51642
diff
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:
51642
diff
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:
51642
diff
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:
51642
diff
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:
51642
diff
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:
51642
diff
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:
51642
diff
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:
31492
diff
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:
51642
diff
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:
56381
diff
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:
67986
diff
changeset
|
164 |
lemma square_bound_lemma: |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
165 |
fixes x :: real |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
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:
67986
diff
changeset
|
167 |
proof - |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
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:
67986
diff
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:
67986
diff
changeset
|
170 |
then show ?thesis |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
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:
67986
diff
changeset
|
172 |
qed |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
173 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
174 |
lemma square_continuous: |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
175 |
fixes e :: real |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
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:
67986
diff
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:
67986
diff
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:
67986
diff
changeset
|
179 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
180 |
lemma norm_triangle_sub: |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
181 |
fixes x y :: "'a::real_normed_vector" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
182 |
shows "norm x \<le> norm y + norm (x - y)" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
183 |
using norm_triangle_ineq[of "y" "x - y"] by (simp add: field_simps) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
184 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
185 |
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:
67986
diff
changeset
|
186 |
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:
67986
diff
changeset
|
187 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
188 |
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:
67986
diff
changeset
|
189 |
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:
67986
diff
changeset
|
190 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
191 |
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:
67986
diff
changeset
|
192 |
apply (subst order_eq_iff) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
193 |
apply (auto simp: norm_le) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
194 |
done |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
195 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67986
diff
changeset
|
196 |
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:
67986
diff
changeset
|
197 |
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:
67986
diff
changeset
|
198 |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
changeset
|
199 |
lemma inner_divide_left: |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
changeset
|
200 |
fixes a :: "'a :: {real_inner,real_div_algebra}" |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
changeset
|
201 |
shows "inner (a / of_real m) b = (inner a b) / m" |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
changeset
|
202 |
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:
60679
diff
changeset
|
203 |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
changeset
|
204 |
lemma inner_divide_right: |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
changeset
|
205 |
fixes a :: "'a :: {real_inner,real_div_algebra}" |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
changeset
|
206 |
shows "inner a (b / of_real m) = (inner a b) / m" |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
changeset
|
207 |
by (metis inner_commute inner_divide_left) |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
60679
diff
changeset
|
208 |
|
60500 | 209 |
text \<open> |
62101 | 210 |
Re-enable constraints for @{term "open"}, @{term "uniformity"}, |
31492
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
211 |
@{term dist}, and @{term norm}. |
60500 | 212 |
\<close> |
31492
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
213 |
|
60500 | 214 |
setup \<open>Sign.add_const_constraint |
215 |
(@{const_name "open"}, SOME @{typ "'a::topological_space set \<Rightarrow> bool"})\<close> |
|
31446 | 216 |
|
60500 | 217 |
setup \<open>Sign.add_const_constraint |
62101 | 218 |
(@{const_name uniformity}, SOME @{typ "('a::uniform_space \<times> 'a) filter"})\<close> |
219 |
||
220 |
setup \<open>Sign.add_const_constraint |
|
60500 | 221 |
(@{const_name dist}, SOME @{typ "'a::metric_space \<Rightarrow> 'a \<Rightarrow> real"})\<close> |
31446 | 222 |
|
60500 | 223 |
setup \<open>Sign.add_const_constraint |
224 |
(@{const_name norm}, SOME @{typ "'a::real_normed_vector \<Rightarrow> real"})\<close> |
|
31446 | 225 |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
226 |
lemma bounded_bilinear_inner: |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
227 |
"bounded_bilinear (inner::'a::real_inner \<Rightarrow> 'a \<Rightarrow> real)" |
29993 | 228 |
proof |
229 |
fix x y z :: 'a and r :: real |
|
230 |
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:
31492
diff
changeset
|
231 |
by (rule inner_add_left) |
29993 | 232 |
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:
31492
diff
changeset
|
233 |
by (rule inner_add_right) |
29993 | 234 |
show "inner (scaleR r x) y = scaleR r (inner x y)" |
235 |
unfolding real_scaleR_def by (rule inner_scaleR_left) |
|
236 |
show "inner x (scaleR r y) = scaleR r (inner x y)" |
|
237 |
unfolding real_scaleR_def by (rule inner_scaleR_right) |
|
238 |
show "\<exists>K. \<forall>x y::'a. norm (inner x y) \<le> norm x * norm y * K" |
|
239 |
proof |
|
240 |
show "\<forall>x y::'a. norm (inner x y) \<le> norm x * norm y * 1" |
|
30046 | 241 |
by (simp add: Cauchy_Schwarz_ineq2) |
29993 | 242 |
qed |
243 |
qed |
|
244 |
||
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
245 |
lemmas tendsto_inner [tendsto_intros] = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
246 |
bounded_bilinear.tendsto [OF bounded_bilinear_inner] |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
247 |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
248 |
lemmas isCont_inner [simp] = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
249 |
bounded_bilinear.isCont [OF bounded_bilinear_inner] |
29993 | 250 |
|
56381
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
hoelzl
parents:
56181
diff
changeset
|
251 |
lemmas has_derivative_inner [derivative_intros] = |
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
252 |
bounded_bilinear.FDERIV [OF bounded_bilinear_inner] |
29993 | 253 |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
254 |
lemmas bounded_linear_inner_left = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
255 |
bounded_bilinear.bounded_linear_left [OF bounded_bilinear_inner] |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
256 |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
257 |
lemmas bounded_linear_inner_right = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
258 |
bounded_bilinear.bounded_linear_right [OF bounded_bilinear_inner] |
44233 | 259 |
|
61915
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61518
diff
changeset
|
260 |
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:
61518
diff
changeset
|
261 |
|
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61518
diff
changeset
|
262 |
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:
61518
diff
changeset
|
263 |
|
56381
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
hoelzl
parents:
56181
diff
changeset
|
264 |
lemmas has_derivative_inner_right [derivative_intros] = |
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
265 |
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:
51002
diff
changeset
|
266 |
|
56381
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
hoelzl
parents:
56181
diff
changeset
|
267 |
lemmas has_derivative_inner_left [derivative_intros] = |
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
268 |
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:
51002
diff
changeset
|
269 |
|
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:
51002
diff
changeset
|
270 |
lemma differentiable_inner [simp]: |
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
271 |
"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:
54230
diff
changeset
|
272 |
unfolding differentiable_def by (blast intro: has_derivative_inner) |
29993 | 273 |
|
60679 | 274 |
|
60500 | 275 |
subsection \<open>Class instances\<close> |
29993 | 276 |
|
68617 | 277 |
instantiation real :: real_inner |
29993 | 278 |
begin |
279 |
||
67399 | 280 |
definition inner_real_def [simp]: "inner = ( * )" |
29993 | 281 |
|
60679 | 282 |
instance |
283 |
proof |
|
29993 | 284 |
fix x y z r :: real |
285 |
show "inner x y = inner y x" |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
56381
diff
changeset
|
286 |
unfolding inner_real_def by (rule mult.commute) |
29993 | 287 |
show "inner (x + y) z = inner x z + inner y z" |
49962
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents:
44902
diff
changeset
|
288 |
unfolding inner_real_def by (rule distrib_right) |
29993 | 289 |
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:
56381
diff
changeset
|
290 |
unfolding inner_real_def real_scaleR_def by (rule mult.assoc) |
29993 | 291 |
show "0 \<le> inner x x" |
292 |
unfolding inner_real_def by simp |
|
293 |
show "inner x x = 0 \<longleftrightarrow> x = 0" |
|
294 |
unfolding inner_real_def by simp |
|
295 |
show "norm x = sqrt (inner x x)" |
|
296 |
unfolding inner_real_def by simp |
|
297 |
qed |
|
298 |
||
299 |
end |
|
300 |
||
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63589
diff
changeset
|
301 |
lemma |
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63589
diff
changeset
|
302 |
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:
63589
diff
changeset
|
303 |
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:
63589
diff
changeset
|
304 |
by simp_all |
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63589
diff
changeset
|
305 |
|
68617 | 306 |
instantiation complex :: real_inner |
29993 | 307 |
begin |
308 |
||
309 |
definition inner_complex_def: |
|
310 |
"inner x y = Re x * Re y + Im x * Im y" |
|
311 |
||
60679 | 312 |
instance |
313 |
proof |
|
29993 | 314 |
fix x y z :: complex and r :: real |
315 |
show "inner x y = inner y x" |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
56381
diff
changeset
|
316 |
unfolding inner_complex_def by (simp add: mult.commute) |
29993 | 317 |
show "inner (x + y) z = inner x z + inner y z" |
49962
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents:
44902
diff
changeset
|
318 |
unfolding inner_complex_def by (simp add: distrib_right) |
29993 | 319 |
show "inner (scaleR r x) y = r * inner x y" |
49962
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents:
44902
diff
changeset
|
320 |
unfolding inner_complex_def by (simp add: distrib_left) |
29993 | 321 |
show "0 \<le> inner x x" |
44126 | 322 |
unfolding inner_complex_def by simp |
29993 | 323 |
show "inner x x = 0 \<longleftrightarrow> x = 0" |
324 |
unfolding inner_complex_def |
|
68499
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents:
68073
diff
changeset
|
325 |
by (simp add: add_nonneg_eq_0_iff complex_eq_iff) |
29993 | 326 |
show "norm x = sqrt (inner x x)" |
68499
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents:
68073
diff
changeset
|
327 |
unfolding inner_complex_def norm_complex_def |
29993 | 328 |
by (simp add: power2_eq_square) |
329 |
qed |
|
330 |
||
331 |
end |
|
332 |
||
44902
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents:
44282
diff
changeset
|
333 |
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:
44282
diff
changeset
|
334 |
unfolding inner_complex_def by simp |
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents:
44282
diff
changeset
|
335 |
|
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents:
44282
diff
changeset
|
336 |
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:
44282
diff
changeset
|
337 |
unfolding inner_complex_def by simp |
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents:
44282
diff
changeset
|
338 |
|
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:
64267
diff
changeset
|
339 |
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:
44282
diff
changeset
|
340 |
unfolding inner_complex_def by simp |
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents:
44282
diff
changeset
|
341 |
|
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:
64267
diff
changeset
|
342 |
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:
44282
diff
changeset
|
343 |
unfolding inner_complex_def by simp |
9ba11d41cd1f
move lemmas about complex number 'i' to Complex.thy and Library/Inner_Product.thy
huffman
parents:
44282
diff
changeset
|
344 |
|
29993 | 345 |
|
66486
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
346 |
lemma dot_square_norm: "inner x x = (norm x)\<^sup>2" |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
347 |
by (simp only: power2_norm_eq_inner) (* TODO: move? *) |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
348 |
|
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
349 |
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:
65513
diff
changeset
|
350 |
by (auto simp add: norm_eq_sqrt_inner) |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
351 |
|
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
352 |
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:
65513
diff
changeset
|
353 |
apply (simp add: dot_square_norm abs_le_square_iff[symmetric]) |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
354 |
using norm_ge_zero[of x] |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
355 |
apply arith |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
356 |
done |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
357 |
|
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
358 |
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:
65513
diff
changeset
|
359 |
apply (simp add: dot_square_norm abs_le_square_iff[symmetric]) |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
360 |
using norm_ge_zero[of x] |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
361 |
apply arith |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
362 |
done |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
363 |
|
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
364 |
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:
65513
diff
changeset
|
365 |
by (metis not_le norm_ge_square) |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
366 |
|
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
367 |
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:
65513
diff
changeset
|
368 |
by (metis norm_le_square not_less) |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
369 |
|
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
370 |
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:
65513
diff
changeset
|
371 |
|
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
372 |
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:
65513
diff
changeset
|
373 |
inner_scaleR_left inner_scaleR_right |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
374 |
|
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
375 |
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:
65513
diff
changeset
|
376 |
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:
65513
diff
changeset
|
377 |
|
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
378 |
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:
65513
diff
changeset
|
379 |
by (simp only: power2_norm_eq_inner inner_simps inner_commute) |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
380 |
(auto simp add: algebra_simps) |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
381 |
|
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
382 |
lemma of_real_inner_1 [simp]: |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
383 |
"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:
65513
diff
changeset
|
384 |
by (simp add: of_real_def dot_square_norm) |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
385 |
|
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
386 |
lemma summable_of_real_iff: |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
387 |
"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:
65513
diff
changeset
|
388 |
proof |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
389 |
assume *: "summable (\<lambda>x. of_real (f x) :: 'a)" |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
390 |
interpret bounded_linear "\<lambda>x::'a. inner x 1" |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
391 |
by (rule bounded_linear_inner_left) |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
392 |
from summable [OF *] show "summable f" by simp |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
393 |
qed (auto intro: summable_of_real) |
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
394 |
|
ffaaa83543b2
Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents:
65513
diff
changeset
|
395 |
|
60500 | 396 |
subsection \<open>Gradient derivative\<close> |
29993 | 397 |
|
67962 | 398 |
definition%important |
29993 | 399 |
gderiv :: |
400 |
"['a::real_inner \<Rightarrow> real, 'a, 'a] \<Rightarrow> bool" |
|
401 |
("(GDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) |
|
402 |
where |
|
403 |
"GDERIV f x :> D \<longleftrightarrow> FDERIV f x :> (\<lambda>h. inner h D)" |
|
404 |
||
405 |
lemma gderiv_deriv [simp]: "GDERIV f x :> D \<longleftrightarrow> DERIV f x :> D" |
|
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
406 |
by (simp only: gderiv_def has_field_derivative_def inner_real_def mult_commute_abs) |
29993 | 407 |
|
408 |
lemma GDERIV_DERIV_compose: |
|
409 |
"\<lbrakk>GDERIV f x :> df; DERIV g (f x) :> dg\<rbrakk> |
|
410 |
\<Longrightarrow> GDERIV (\<lambda>x. g (f x)) x :> scaleR dg df" |
|
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
411 |
unfolding gderiv_def has_field_derivative_def |
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
412 |
apply (drule (1) has_derivative_compose) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
413 |
apply (simp add: ac_simps) |
29993 | 414 |
done |
415 |
||
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
416 |
lemma has_derivative_subst: "\<lbrakk>FDERIV f x :> df; df = d\<rbrakk> \<Longrightarrow> FDERIV f x :> d" |
29993 | 417 |
by simp |
418 |
||
419 |
lemma GDERIV_subst: "\<lbrakk>GDERIV f x :> df; df = d\<rbrakk> \<Longrightarrow> GDERIV f x :> d" |
|
420 |
by simp |
|
421 |
||
422 |
lemma GDERIV_const: "GDERIV (\<lambda>x. k) x :> 0" |
|
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
423 |
unfolding gderiv_def inner_zero_right by (rule has_derivative_const) |
29993 | 424 |
|
425 |
lemma GDERIV_add: |
|
426 |
"\<lbrakk>GDERIV f x :> df; GDERIV g x :> dg\<rbrakk> |
|
427 |
\<Longrightarrow> GDERIV (\<lambda>x. f x + g x) x :> df + dg" |
|
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
428 |
unfolding gderiv_def inner_add_right by (rule has_derivative_add) |
29993 | 429 |
|
430 |
lemma GDERIV_minus: |
|
431 |
"GDERIV f x :> df \<Longrightarrow> GDERIV (\<lambda>x. - f x) x :> - df" |
|
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
432 |
unfolding gderiv_def inner_minus_right by (rule has_derivative_minus) |
29993 | 433 |
|
434 |
lemma GDERIV_diff: |
|
435 |
"\<lbrakk>GDERIV f x :> df; GDERIV g x :> dg\<rbrakk> |
|
436 |
\<Longrightarrow> GDERIV (\<lambda>x. f x - g x) x :> df - dg" |
|
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
437 |
unfolding gderiv_def inner_diff_right by (rule has_derivative_diff) |
29993 | 438 |
|
439 |
lemma GDERIV_scaleR: |
|
440 |
"\<lbrakk>DERIV f x :> df; GDERIV g x :> dg\<rbrakk> |
|
441 |
\<Longrightarrow> GDERIV (\<lambda>x. scaleR (f x) (g x)) x |
|
442 |
:> (scaleR (f x) dg + scaleR df (g x))" |
|
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
443 |
unfolding gderiv_def has_field_derivative_def inner_add_right inner_scaleR_right |
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
444 |
apply (rule has_derivative_subst) |
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
445 |
apply (erule (1) has_derivative_scaleR) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
446 |
apply (simp add: ac_simps) |
29993 | 447 |
done |
448 |
||
449 |
lemma GDERIV_mult: |
|
450 |
"\<lbrakk>GDERIV f x :> df; GDERIV g x :> dg\<rbrakk> |
|
451 |
\<Longrightarrow> GDERIV (\<lambda>x. f x * g x) x :> scaleR (f x) dg + scaleR (g x) df" |
|
452 |
unfolding gderiv_def |
|
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
453 |
apply (rule has_derivative_subst) |
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
454 |
apply (erule (1) has_derivative_mult) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
455 |
apply (simp add: inner_add ac_simps) |
29993 | 456 |
done |
457 |
||
458 |
lemma GDERIV_inverse: |
|
459 |
"\<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:
51642
diff
changeset
|
460 |
\<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:
67962
diff
changeset
|
461 |
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:
67962
diff
changeset
|
462 |
|
29993 | 463 |
lemma GDERIV_norm: |
464 |
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:
67962
diff
changeset
|
465 |
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:
67962
diff
changeset
|
466 |
by (rule derivative_eq_intros | force simp add: inner_commute sgn_div_norm norm_eq_sqrt_inner assms)+ |
29993 | 467 |
|
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
54230
diff
changeset
|
468 |
lemmas has_derivative_norm = GDERIV_norm [unfolded gderiv_def] |
29993 | 469 |
|
470 |
end |