author | huffman |
Fri, 02 Sep 2011 20:58:31 -0700 | |
changeset 44678 | 21eb31192850 |
parent 44282 | f0de18b62d63 |
child 44902 | 9ba11d41cd1f |
permissions | -rw-r--r-- |
41959 | 1 |
(* Title: HOL/Library/Inner_Product.thy |
2 |
Author: Brian Huffman |
|
29993 | 3 |
*) |
4 |
||
5 |
header {* Inner Product Spaces and the Gradient Derivative *} |
|
6 |
||
7 |
theory Inner_Product |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
8 |
imports FrechetDeriv |
29993 | 9 |
begin |
10 |
||
11 |
subsection {* Real inner product spaces *} |
|
12 |
||
31492
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
13 |
text {* |
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
14 |
Temporarily relax type constraints for @{term "open"}, |
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
15 |
@{term dist}, and @{term norm}. |
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
16 |
*} |
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
17 |
|
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
18 |
setup {* Sign.add_const_constraint |
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
19 |
(@{const_name "open"}, SOME @{typ "'a::open set \<Rightarrow> bool"}) *} |
31446 | 20 |
|
21 |
setup {* Sign.add_const_constraint |
|
22 |
(@{const_name dist}, SOME @{typ "'a::dist \<Rightarrow> 'a \<Rightarrow> real"}) *} |
|
23 |
||
24 |
setup {* Sign.add_const_constraint |
|
25 |
(@{const_name norm}, SOME @{typ "'a::norm \<Rightarrow> real"}) *} |
|
26 |
||
31492
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
27 |
class real_inner = real_vector + sgn_div_norm + dist_norm + open_dist + |
29993 | 28 |
fixes inner :: "'a \<Rightarrow> 'a \<Rightarrow> real" |
29 |
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
|
30 |
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
|
31 |
and inner_scaleR_left [simp]: "inner (scaleR r x) y = r * (inner x y)" |
29993 | 32 |
and inner_ge_zero [simp]: "0 \<le> inner x x" |
33 |
and inner_eq_zero_iff [simp]: "inner x x = 0 \<longleftrightarrow> x = 0" |
|
34 |
and norm_eq_sqrt_inner: "norm x = sqrt (inner x x)" |
|
35 |
begin |
|
36 |
||
37 |
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
|
38 |
using inner_add_left [of 0 0 x] by simp |
29993 | 39 |
|
40 |
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
|
41 |
using inner_add_left [of x "- x" y] by simp |
29993 | 42 |
|
43 |
lemma inner_diff_left: "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
|
44 |
by (simp add: diff_minus inner_add_left) |
29993 | 45 |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
46 |
lemma inner_setsum_left: "inner (\<Sum>x\<in>A. f x) y = (\<Sum>x\<in>A. inner (f x) y)" |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
47 |
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
|
48 |
|
29993 | 49 |
text {* Transfer distributivity rules to right argument. *} |
50 |
||
31590
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents:
31492
diff
changeset
|
51 |
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
|
52 |
using inner_add_left [of y z x] by (simp only: inner_commute) |
29993 | 53 |
|
31590
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents:
31492
diff
changeset
|
54 |
lemma inner_scaleR_right [simp]: "inner x (scaleR r y) = r * (inner x y)" |
29993 | 55 |
using inner_scaleR_left [of r y x] by (simp only: inner_commute) |
56 |
||
57 |
lemma inner_zero_right [simp]: "inner x 0 = 0" |
|
58 |
using inner_zero_left [of x] by (simp only: inner_commute) |
|
59 |
||
60 |
lemma inner_minus_right [simp]: "inner x (- y) = - inner x y" |
|
61 |
using inner_minus_left [of y x] by (simp only: inner_commute) |
|
62 |
||
63 |
lemma inner_diff_right: "inner x (y - z) = inner x y - inner x z" |
|
64 |
using inner_diff_left [of y z x] by (simp only: inner_commute) |
|
65 |
||
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
66 |
lemma inner_setsum_right: "inner x (\<Sum>y\<in>A. f y) = (\<Sum>y\<in>A. inner x (f y))" |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
67 |
using inner_setsum_left [of f A x] by (simp only: inner_commute) |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
68 |
|
31590
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents:
31492
diff
changeset
|
69 |
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
|
70 |
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
|
71 |
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
|
72 |
|
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents:
31492
diff
changeset
|
73 |
text {* Legacy theorem names *} |
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents:
31492
diff
changeset
|
74 |
lemmas inner_left_distrib = inner_add_left |
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents:
31492
diff
changeset
|
75 |
lemmas inner_right_distrib = inner_add_right |
29993 | 76 |
lemmas inner_distrib = inner_left_distrib inner_right_distrib |
77 |
||
78 |
lemma inner_gt_zero_iff [simp]: "0 < inner x x \<longleftrightarrow> x \<noteq> 0" |
|
79 |
by (simp add: order_less_le) |
|
80 |
||
81 |
lemma power2_norm_eq_inner: "(norm x)\<twosuperior> = inner x x" |
|
82 |
by (simp add: norm_eq_sqrt_inner) |
|
83 |
||
30046 | 84 |
lemma Cauchy_Schwarz_ineq: |
29993 | 85 |
"(inner x y)\<twosuperior> \<le> inner x x * inner y y" |
86 |
proof (cases) |
|
87 |
assume "y = 0" |
|
88 |
thus ?thesis by simp |
|
89 |
next |
|
90 |
assume y: "y \<noteq> 0" |
|
91 |
let ?r = "inner x y / inner y y" |
|
92 |
have "0 \<le> inner (x - scaleR ?r y) (x - scaleR ?r y)" |
|
93 |
by (rule inner_ge_zero) |
|
94 |
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
|
95 |
by (simp add: inner_diff) |
29993 | 96 |
also have "\<dots> = inner x x - (inner x y)\<twosuperior> / inner y y" |
97 |
by (simp add: power2_eq_square inner_commute) |
|
98 |
finally have "0 \<le> inner x x - (inner x y)\<twosuperior> / inner y y" . |
|
99 |
hence "(inner x y)\<twosuperior> / inner y y \<le> inner x x" |
|
100 |
by (simp add: le_diff_eq) |
|
101 |
thus "(inner x y)\<twosuperior> \<le> inner x x * inner y y" |
|
102 |
by (simp add: pos_divide_le_eq y) |
|
103 |
qed |
|
104 |
||
30046 | 105 |
lemma Cauchy_Schwarz_ineq2: |
29993 | 106 |
"\<bar>inner x y\<bar> \<le> norm x * norm y" |
107 |
proof (rule power2_le_imp_le) |
|
108 |
have "(inner x y)\<twosuperior> \<le> inner x x * inner y y" |
|
30046 | 109 |
using Cauchy_Schwarz_ineq . |
29993 | 110 |
thus "\<bar>inner x y\<bar>\<twosuperior> \<le> (norm x * norm y)\<twosuperior>" |
111 |
by (simp add: power_mult_distrib power2_norm_eq_inner) |
|
112 |
show "0 \<le> norm x * norm y" |
|
113 |
unfolding norm_eq_sqrt_inner |
|
114 |
by (intro mult_nonneg_nonneg real_sqrt_ge_zero inner_ge_zero) |
|
115 |
qed |
|
116 |
||
117 |
subclass real_normed_vector |
|
118 |
proof |
|
119 |
fix a :: real and x y :: 'a |
|
120 |
show "0 \<le> norm x" |
|
121 |
unfolding norm_eq_sqrt_inner by simp |
|
122 |
show "norm x = 0 \<longleftrightarrow> x = 0" |
|
123 |
unfolding norm_eq_sqrt_inner by simp |
|
124 |
show "norm (x + y) \<le> norm x + norm y" |
|
125 |
proof (rule power2_le_imp_le) |
|
126 |
have "inner x y \<le> norm x * norm y" |
|
30046 | 127 |
by (rule order_trans [OF abs_ge_self Cauchy_Schwarz_ineq2]) |
29993 | 128 |
thus "(norm (x + y))\<twosuperior> \<le> (norm x + norm y)\<twosuperior>" |
129 |
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
|
130 |
by (simp add: inner_add inner_commute) |
29993 | 131 |
show "0 \<le> norm x + norm y" |
44126 | 132 |
unfolding norm_eq_sqrt_inner by simp |
29993 | 133 |
qed |
134 |
have "sqrt (a\<twosuperior> * inner x x) = \<bar>a\<bar> * sqrt (inner x x)" |
|
135 |
by (simp add: real_sqrt_mult_distrib) |
|
136 |
then show "norm (a *\<^sub>R x) = \<bar>a\<bar> * norm x" |
|
137 |
unfolding norm_eq_sqrt_inner |
|
31590
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents:
31492
diff
changeset
|
138 |
by (simp add: power2_eq_square mult_assoc) |
29993 | 139 |
qed |
140 |
||
141 |
end |
|
142 |
||
31492
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
143 |
text {* |
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
144 |
Re-enable constraints for @{term "open"}, |
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
145 |
@{term dist}, and @{term norm}. |
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
146 |
*} |
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
147 |
|
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
148 |
setup {* Sign.add_const_constraint |
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
huffman
parents:
31446
diff
changeset
|
149 |
(@{const_name "open"}, SOME @{typ "'a::topological_space set \<Rightarrow> bool"}) *} |
31446 | 150 |
|
151 |
setup {* Sign.add_const_constraint |
|
152 |
(@{const_name dist}, SOME @{typ "'a::metric_space \<Rightarrow> 'a \<Rightarrow> real"}) *} |
|
153 |
||
154 |
setup {* Sign.add_const_constraint |
|
155 |
(@{const_name norm}, SOME @{typ "'a::real_normed_vector \<Rightarrow> real"}) *} |
|
156 |
||
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
157 |
lemma bounded_bilinear_inner: |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
158 |
"bounded_bilinear (inner::'a::real_inner \<Rightarrow> 'a \<Rightarrow> real)" |
29993 | 159 |
proof |
160 |
fix x y z :: 'a and r :: real |
|
161 |
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
|
162 |
by (rule inner_add_left) |
29993 | 163 |
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
|
164 |
by (rule inner_add_right) |
29993 | 165 |
show "inner (scaleR r x) y = scaleR r (inner x y)" |
166 |
unfolding real_scaleR_def by (rule inner_scaleR_left) |
|
167 |
show "inner x (scaleR r y) = scaleR r (inner x y)" |
|
168 |
unfolding real_scaleR_def by (rule inner_scaleR_right) |
|
169 |
show "\<exists>K. \<forall>x y::'a. norm (inner x y) \<le> norm x * norm y * K" |
|
170 |
proof |
|
171 |
show "\<forall>x y::'a. norm (inner x y) \<le> norm x * norm y * 1" |
|
30046 | 172 |
by (simp add: Cauchy_Schwarz_ineq2) |
29993 | 173 |
qed |
174 |
qed |
|
175 |
||
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
176 |
lemmas tendsto_inner [tendsto_intros] = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
177 |
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
|
178 |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
179 |
lemmas isCont_inner [simp] = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
180 |
bounded_bilinear.isCont [OF bounded_bilinear_inner] |
29993 | 181 |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
182 |
lemmas FDERIV_inner = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
183 |
bounded_bilinear.FDERIV [OF bounded_bilinear_inner] |
29993 | 184 |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
185 |
lemmas bounded_linear_inner_left = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
186 |
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
|
187 |
|
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
188 |
lemmas bounded_linear_inner_right = |
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
189 |
bounded_bilinear.bounded_linear_right [OF bounded_bilinear_inner] |
44233 | 190 |
|
29993 | 191 |
|
192 |
subsection {* Class instances *} |
|
193 |
||
194 |
instantiation real :: real_inner |
|
195 |
begin |
|
196 |
||
197 |
definition inner_real_def [simp]: "inner = op *" |
|
198 |
||
199 |
instance proof |
|
200 |
fix x y z r :: real |
|
201 |
show "inner x y = inner y x" |
|
202 |
unfolding inner_real_def by (rule mult_commute) |
|
203 |
show "inner (x + y) z = inner x z + inner y z" |
|
204 |
unfolding inner_real_def by (rule left_distrib) |
|
205 |
show "inner (scaleR r x) y = r * inner x y" |
|
206 |
unfolding inner_real_def real_scaleR_def by (rule mult_assoc) |
|
207 |
show "0 \<le> inner x x" |
|
208 |
unfolding inner_real_def by simp |
|
209 |
show "inner x x = 0 \<longleftrightarrow> x = 0" |
|
210 |
unfolding inner_real_def by simp |
|
211 |
show "norm x = sqrt (inner x x)" |
|
212 |
unfolding inner_real_def by simp |
|
213 |
qed |
|
214 |
||
215 |
end |
|
216 |
||
217 |
instantiation complex :: real_inner |
|
218 |
begin |
|
219 |
||
220 |
definition inner_complex_def: |
|
221 |
"inner x y = Re x * Re y + Im x * Im y" |
|
222 |
||
223 |
instance proof |
|
224 |
fix x y z :: complex and r :: real |
|
225 |
show "inner x y = inner y x" |
|
226 |
unfolding inner_complex_def by (simp add: mult_commute) |
|
227 |
show "inner (x + y) z = inner x z + inner y z" |
|
228 |
unfolding inner_complex_def by (simp add: left_distrib) |
|
229 |
show "inner (scaleR r x) y = r * inner x y" |
|
230 |
unfolding inner_complex_def by (simp add: right_distrib) |
|
231 |
show "0 \<le> inner x x" |
|
44126 | 232 |
unfolding inner_complex_def by simp |
29993 | 233 |
show "inner x x = 0 \<longleftrightarrow> x = 0" |
234 |
unfolding inner_complex_def |
|
235 |
by (simp add: add_nonneg_eq_0_iff complex_Re_Im_cancel_iff) |
|
236 |
show "norm x = sqrt (inner x x)" |
|
237 |
unfolding inner_complex_def complex_norm_def |
|
238 |
by (simp add: power2_eq_square) |
|
239 |
qed |
|
240 |
||
241 |
end |
|
242 |
||
243 |
||
244 |
subsection {* Gradient derivative *} |
|
245 |
||
246 |
definition |
|
247 |
gderiv :: |
|
248 |
"['a::real_inner \<Rightarrow> real, 'a, 'a] \<Rightarrow> bool" |
|
249 |
("(GDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) |
|
250 |
where |
|
251 |
"GDERIV f x :> D \<longleftrightarrow> FDERIV f x :> (\<lambda>h. inner h D)" |
|
252 |
||
253 |
lemma deriv_fderiv: "DERIV f x :> D \<longleftrightarrow> FDERIV f x :> (\<lambda>h. h * D)" |
|
254 |
by (simp only: deriv_def field_fderiv_def) |
|
255 |
||
256 |
lemma gderiv_deriv [simp]: "GDERIV f x :> D \<longleftrightarrow> DERIV f x :> D" |
|
257 |
by (simp only: gderiv_def deriv_fderiv inner_real_def) |
|
258 |
||
259 |
lemma GDERIV_DERIV_compose: |
|
260 |
"\<lbrakk>GDERIV f x :> df; DERIV g (f x) :> dg\<rbrakk> |
|
261 |
\<Longrightarrow> GDERIV (\<lambda>x. g (f x)) x :> scaleR dg df" |
|
262 |
unfolding gderiv_def deriv_fderiv |
|
263 |
apply (drule (1) FDERIV_compose) |
|
31590
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents:
31492
diff
changeset
|
264 |
apply (simp add: mult_ac) |
29993 | 265 |
done |
266 |
||
267 |
lemma FDERIV_subst: "\<lbrakk>FDERIV f x :> df; df = d\<rbrakk> \<Longrightarrow> FDERIV f x :> d" |
|
268 |
by simp |
|
269 |
||
270 |
lemma GDERIV_subst: "\<lbrakk>GDERIV f x :> df; df = d\<rbrakk> \<Longrightarrow> GDERIV f x :> d" |
|
271 |
by simp |
|
272 |
||
273 |
lemma GDERIV_const: "GDERIV (\<lambda>x. k) x :> 0" |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
274 |
unfolding gderiv_def inner_zero_right by (rule FDERIV_const) |
29993 | 275 |
|
276 |
lemma GDERIV_add: |
|
277 |
"\<lbrakk>GDERIV f x :> df; GDERIV g x :> dg\<rbrakk> |
|
278 |
\<Longrightarrow> GDERIV (\<lambda>x. f x + g x) x :> df + dg" |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
279 |
unfolding gderiv_def inner_add_right by (rule FDERIV_add) |
29993 | 280 |
|
281 |
lemma GDERIV_minus: |
|
282 |
"GDERIV f x :> df \<Longrightarrow> GDERIV (\<lambda>x. - f x) x :> - df" |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
283 |
unfolding gderiv_def inner_minus_right by (rule FDERIV_minus) |
29993 | 284 |
|
285 |
lemma GDERIV_diff: |
|
286 |
"\<lbrakk>GDERIV f x :> df; GDERIV g x :> dg\<rbrakk> |
|
287 |
\<Longrightarrow> GDERIV (\<lambda>x. f x - g x) x :> df - dg" |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
288 |
unfolding gderiv_def inner_diff_right by (rule FDERIV_diff) |
29993 | 289 |
|
290 |
lemma GDERIV_scaleR: |
|
291 |
"\<lbrakk>DERIV f x :> df; GDERIV g x :> dg\<rbrakk> |
|
292 |
\<Longrightarrow> GDERIV (\<lambda>x. scaleR (f x) (g x)) x |
|
293 |
:> (scaleR (f x) dg + scaleR df (g x))" |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
294 |
unfolding gderiv_def deriv_fderiv inner_add_right inner_scaleR_right |
29993 | 295 |
apply (rule FDERIV_subst) |
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
296 |
apply (erule (1) FDERIV_scaleR) |
29993 | 297 |
apply (simp add: mult_ac) |
298 |
done |
|
299 |
||
300 |
lemma GDERIV_mult: |
|
301 |
"\<lbrakk>GDERIV f x :> df; GDERIV g x :> dg\<rbrakk> |
|
302 |
\<Longrightarrow> GDERIV (\<lambda>x. f x * g x) x :> scaleR (f x) dg + scaleR (g x) df" |
|
303 |
unfolding gderiv_def |
|
304 |
apply (rule FDERIV_subst) |
|
305 |
apply (erule (1) FDERIV_mult) |
|
31590
776d6a4c1327
declare inner_add, inner_diff [algebra_simps]; declare inner_scaleR [simp]
huffman
parents:
31492
diff
changeset
|
306 |
apply (simp add: inner_add mult_ac) |
29993 | 307 |
done |
308 |
||
309 |
lemma GDERIV_inverse: |
|
310 |
"\<lbrakk>GDERIV f x :> df; f x \<noteq> 0\<rbrakk> |
|
311 |
\<Longrightarrow> GDERIV (\<lambda>x. inverse (f x)) x :> - (inverse (f x))\<twosuperior> *\<^sub>R df" |
|
312 |
apply (erule GDERIV_DERIV_compose) |
|
313 |
apply (erule DERIV_inverse [folded numeral_2_eq_2]) |
|
314 |
done |
|
315 |
||
316 |
lemma GDERIV_norm: |
|
317 |
assumes "x \<noteq> 0" shows "GDERIV (\<lambda>x. norm x) x :> sgn x" |
|
318 |
proof - |
|
319 |
have 1: "FDERIV (\<lambda>x. inner x x) x :> (\<lambda>h. inner x h + inner h x)" |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44233
diff
changeset
|
320 |
by (intro FDERIV_inner FDERIV_ident) |
29993 | 321 |
have 2: "(\<lambda>h. inner x h + inner h x) = (\<lambda>h. inner h (scaleR 2 x))" |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
322 |
by (simp add: fun_eq_iff inner_commute) |
29993 | 323 |
have "0 < inner x x" using `x \<noteq> 0` by simp |
324 |
then have 3: "DERIV sqrt (inner x x) :> (inverse (sqrt (inner x x)) / 2)" |
|
325 |
by (rule DERIV_real_sqrt) |
|
326 |
have 4: "(inverse (sqrt (inner x x)) / 2) *\<^sub>R 2 *\<^sub>R x = sgn x" |
|
327 |
by (simp add: sgn_div_norm norm_eq_sqrt_inner) |
|
328 |
show ?thesis |
|
329 |
unfolding norm_eq_sqrt_inner |
|
330 |
apply (rule GDERIV_subst [OF _ 4]) |
|
331 |
apply (rule GDERIV_DERIV_compose [where g=sqrt and df="scaleR 2 x"]) |
|
332 |
apply (subst gderiv_def) |
|
333 |
apply (rule FDERIV_subst [OF _ 2]) |
|
334 |
apply (rule 1) |
|
335 |
apply (rule 3) |
|
336 |
done |
|
337 |
qed |
|
338 |
||
339 |
lemmas FDERIV_norm = GDERIV_norm [unfolded gderiv_def] |
|
340 |
||
341 |
end |