| author | huffman | 
| Wed, 24 Aug 2011 11:56:57 -0700 | |
| changeset 44514 | d02b01e5ab8f | 
| parent 44465 | fa56622bb7bc | 
| child 44515 | fcaacc214a36 | 
| permissions | -rw-r--r-- | 
| 44133 | 1  | 
(* Title: HOL/Multivariate_Analysis/Linear_Algebra.thy  | 
2  | 
Author: Amine Chaieb, University of Cambridge  | 
|
3  | 
*)  | 
|
4  | 
||
5  | 
header {* Elementary linear algebra on Euclidean spaces *}
 | 
|
6  | 
||
7  | 
theory Linear_Algebra  | 
|
8  | 
imports  | 
|
9  | 
Euclidean_Space  | 
|
10  | 
"~~/src/HOL/Library/Infinite_Set"  | 
|
11  | 
L2_Norm  | 
|
12  | 
"~~/src/HOL/Library/Convex"  | 
|
| 
44287
 
598ed12b9bee
import Library/Sum_of_Squares instead of reloading positivstellensatz.ML
 
huffman 
parents: 
44286 
diff
changeset
 | 
13  | 
"~~/src/HOL/Library/Sum_of_Squares"  | 
| 44133 | 14  | 
uses  | 
15  | 
  ("normarith.ML")
 | 
|
16  | 
begin  | 
|
17  | 
||
18  | 
lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)"  | 
|
19  | 
by auto  | 
|
20  | 
||
21  | 
notation inner (infix "\<bullet>" 70)  | 
|
22  | 
||
23  | 
lemma square_bound_lemma: "(x::real) < (1 + x) * (1 + x)"  | 
|
24  | 
proof-  | 
|
25  | 
have "(x + 1/2)^2 + 3/4 > 0" using zero_le_power2[of "x+1/2"] by arith  | 
|
26  | 
thus ?thesis by (simp add: field_simps power2_eq_square)  | 
|
27  | 
qed  | 
|
28  | 
||
29  | 
lemma square_continuous: "0 < (e::real) ==> \<exists>d. 0 < d \<and> (\<forall>y. abs(y - x) < d \<longrightarrow> abs(y * y - x * x) < e)"  | 
|
30  | 
using isCont_power[OF isCont_ident, of 2, unfolded isCont_def LIM_eq, rule_format, of e x] apply (auto simp add: power2_eq_square)  | 
|
31  | 
apply (rule_tac x="s" in exI)  | 
|
32  | 
apply auto  | 
|
33  | 
apply (erule_tac x=y in allE)  | 
|
34  | 
apply auto  | 
|
35  | 
done  | 
|
36  | 
||
37  | 
lemma real_le_lsqrt: "0 <= x \<Longrightarrow> 0 <= y \<Longrightarrow> x <= y^2 ==> sqrt x <= y"  | 
|
38  | 
using real_sqrt_le_iff[of x "y^2"] by simp  | 
|
39  | 
||
40  | 
lemma real_le_rsqrt: "x^2 \<le> y \<Longrightarrow> x \<le> sqrt y"  | 
|
41  | 
using real_sqrt_le_mono[of "x^2" y] by simp  | 
|
42  | 
||
43  | 
lemma real_less_rsqrt: "x^2 < y \<Longrightarrow> x < sqrt y"  | 
|
44  | 
using real_sqrt_less_mono[of "x^2" y] by simp  | 
|
45  | 
||
46  | 
lemma sqrt_even_pow2: assumes n: "even n"  | 
|
47  | 
shows "sqrt(2 ^ n) = 2 ^ (n div 2)"  | 
|
48  | 
proof-  | 
|
49  | 
from n obtain m where m: "n = 2*m" unfolding even_mult_two_ex ..  | 
|
50  | 
from m have "sqrt(2 ^ n) = sqrt ((2 ^ m) ^ 2)"  | 
|
51  | 
by (simp only: power_mult[symmetric] mult_commute)  | 
|
52  | 
then show ?thesis using m by simp  | 
|
53  | 
qed  | 
|
54  | 
||
55  | 
lemma real_div_sqrt: "0 <= x ==> x / sqrt(x) = sqrt(x)"  | 
|
56  | 
apply (cases "x = 0", simp_all)  | 
|
57  | 
using sqrt_divide_self_eq[of x]  | 
|
58  | 
apply (simp add: inverse_eq_divide field_simps)  | 
|
59  | 
done  | 
|
60  | 
||
61  | 
text{* Hence derive more interesting properties of the norm. *}
 | 
|
62  | 
||
63  | 
(* FIXME: same as norm_scaleR  | 
|
64  | 
lemma norm_mul[simp]: "norm(a *\<^sub>R x) = abs(a) * norm x"  | 
|
65  | 
by (simp add: norm_vector_def setL2_right_distrib abs_mult)  | 
|
66  | 
*)  | 
|
67  | 
||
68  | 
lemma norm_eq_0_dot: "(norm x = 0) \<longleftrightarrow> (inner x x = (0::real))"  | 
|
69  | 
by (simp add: setL2_def power2_eq_square)  | 
|
70  | 
||
71  | 
lemma norm_cauchy_schwarz:  | 
|
72  | 
shows "inner x y <= norm x * norm y"  | 
|
73  | 
using Cauchy_Schwarz_ineq2[of x y] by auto  | 
|
74  | 
||
75  | 
lemma norm_cauchy_schwarz_abs:  | 
|
76  | 
shows "\<bar>inner x y\<bar> \<le> norm x * norm y"  | 
|
77  | 
by (rule Cauchy_Schwarz_ineq2)  | 
|
78  | 
||
79  | 
lemma norm_triangle_sub:  | 
|
80  | 
fixes x y :: "'a::real_normed_vector"  | 
|
81  | 
shows "norm x \<le> norm y + norm (x - y)"  | 
|
82  | 
using norm_triangle_ineq[of "y" "x - y"] by (simp add: field_simps)  | 
|
83  | 
||
84  | 
lemma real_abs_norm: "\<bar>norm x\<bar> = norm x"  | 
|
85  | 
by (rule abs_norm_cancel)  | 
|
86  | 
lemma real_abs_sub_norm: "\<bar>norm x - norm y\<bar> <= norm(x - y)"  | 
|
87  | 
by (rule norm_triangle_ineq3)  | 
|
88  | 
lemma norm_le: "norm(x) <= norm(y) \<longleftrightarrow> x \<bullet> x <= y \<bullet> y"  | 
|
89  | 
by (simp add: norm_eq_sqrt_inner)  | 
|
90  | 
lemma norm_lt: "norm(x) < norm(y) \<longleftrightarrow> x \<bullet> x < y \<bullet> y"  | 
|
91  | 
by (simp add: norm_eq_sqrt_inner)  | 
|
92  | 
lemma norm_eq: "norm(x) = norm (y) \<longleftrightarrow> x \<bullet> x = y \<bullet> y"  | 
|
93  | 
apply(subst order_eq_iff) unfolding norm_le by auto  | 
|
94  | 
lemma norm_eq_1: "norm(x) = 1 \<longleftrightarrow> x \<bullet> x = 1"  | 
|
95  | 
unfolding norm_eq_sqrt_inner by auto  | 
|
96  | 
||
97  | 
text{* Squaring equations and inequalities involving norms.  *}
 | 
|
98  | 
||
99  | 
lemma dot_square_norm: "x \<bullet> x = norm(x)^2"  | 
|
100  | 
by (simp add: norm_eq_sqrt_inner)  | 
|
101  | 
||
102  | 
lemma norm_eq_square: "norm(x) = a \<longleftrightarrow> 0 <= a \<and> x \<bullet> x = a^2"  | 
|
103  | 
by (auto simp add: norm_eq_sqrt_inner)  | 
|
104  | 
||
105  | 
lemma real_abs_le_square_iff: "\<bar>x\<bar> \<le> \<bar>y\<bar> \<longleftrightarrow> (x::real)^2 \<le> y^2"  | 
|
106  | 
proof  | 
|
107  | 
assume "\<bar>x\<bar> \<le> \<bar>y\<bar>"  | 
|
108  | 
then have "\<bar>x\<bar>\<twosuperior> \<le> \<bar>y\<bar>\<twosuperior>" by (rule power_mono, simp)  | 
|
109  | 
then show "x\<twosuperior> \<le> y\<twosuperior>" by simp  | 
|
110  | 
next  | 
|
111  | 
assume "x\<twosuperior> \<le> y\<twosuperior>"  | 
|
112  | 
then have "sqrt (x\<twosuperior>) \<le> sqrt (y\<twosuperior>)" by (rule real_sqrt_le_mono)  | 
|
113  | 
then show "\<bar>x\<bar> \<le> \<bar>y\<bar>" by simp  | 
|
114  | 
qed  | 
|
115  | 
||
116  | 
lemma norm_le_square: "norm(x) <= a \<longleftrightarrow> 0 <= a \<and> x \<bullet> x <= a^2"  | 
|
117  | 
apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric])  | 
|
118  | 
using norm_ge_zero[of x]  | 
|
119  | 
apply arith  | 
|
120  | 
done  | 
|
121  | 
||
122  | 
lemma norm_ge_square: "norm(x) >= a \<longleftrightarrow> a <= 0 \<or> x \<bullet> x >= a ^ 2"  | 
|
123  | 
apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric])  | 
|
124  | 
using norm_ge_zero[of x]  | 
|
125  | 
apply arith  | 
|
126  | 
done  | 
|
127  | 
||
128  | 
lemma norm_lt_square: "norm(x) < a \<longleftrightarrow> 0 < a \<and> x \<bullet> x < a^2"  | 
|
129  | 
by (metis not_le norm_ge_square)  | 
|
130  | 
lemma norm_gt_square: "norm(x) > a \<longleftrightarrow> a < 0 \<or> x \<bullet> x > a^2"  | 
|
131  | 
by (metis norm_le_square not_less)  | 
|
132  | 
||
133  | 
text{* Dot product in terms of the norm rather than conversely. *}
 | 
|
134  | 
||
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
44176 
diff
changeset
 | 
135  | 
lemmas inner_simps = inner_add_left inner_add_right inner_diff_right inner_diff_left  | 
| 
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
44176 
diff
changeset
 | 
136  | 
inner_scaleR_left inner_scaleR_right  | 
| 44133 | 137  | 
|
138  | 
lemma dot_norm: "x \<bullet> y = (norm(x + y) ^2 - norm x ^ 2 - norm y ^ 2) / 2"  | 
|
139  | 
unfolding power2_norm_eq_inner inner_simps inner_commute by auto  | 
|
140  | 
||
141  | 
lemma dot_norm_neg: "x \<bullet> y = ((norm x ^ 2 + norm y ^ 2) - norm(x - y) ^ 2) / 2"  | 
|
142  | 
unfolding power2_norm_eq_inner inner_simps inner_commute by(auto simp add:algebra_simps)  | 
|
143  | 
||
144  | 
text{* Equality of vectors in terms of @{term "op \<bullet>"} products.    *}
 | 
|
145  | 
||
146  | 
lemma vector_eq: "x = y \<longleftrightarrow> x \<bullet> x = x \<bullet> y \<and> y \<bullet> y = x \<bullet> x" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
147  | 
proof  | 
|
148  | 
assume ?lhs then show ?rhs by simp  | 
|
149  | 
next  | 
|
150  | 
assume ?rhs  | 
|
151  | 
then have "x \<bullet> x - x \<bullet> y = 0 \<and> x \<bullet> y - y \<bullet> y = 0" by simp  | 
|
| 44454 | 152  | 
hence "x \<bullet> (x - y) = 0 \<and> y \<bullet> (x - y) = 0" by (simp add: inner_diff inner_commute)  | 
153  | 
then have "(x - y) \<bullet> (x - y) = 0" by (simp add: field_simps inner_diff inner_commute)  | 
|
| 44133 | 154  | 
then show "x = y" by (simp)  | 
155  | 
qed  | 
|
156  | 
||
157  | 
subsection{* General linear decision procedure for normed spaces. *}
 | 
|
158  | 
||
159  | 
lemma norm_cmul_rule_thm:  | 
|
160  | 
fixes x :: "'a::real_normed_vector"  | 
|
161  | 
shows "b >= norm(x) ==> \<bar>c\<bar> * b >= norm(scaleR c x)"  | 
|
162  | 
unfolding norm_scaleR  | 
|
163  | 
apply (erule mult_left_mono)  | 
|
164  | 
apply simp  | 
|
165  | 
done  | 
|
166  | 
||
167  | 
(* FIXME: Move all these theorems into the ML code using lemma antiquotation *)  | 
|
168  | 
lemma norm_add_rule_thm:  | 
|
169  | 
fixes x1 x2 :: "'a::real_normed_vector"  | 
|
170  | 
shows "norm x1 \<le> b1 \<Longrightarrow> norm x2 \<le> b2 \<Longrightarrow> norm (x1 + x2) \<le> b1 + b2"  | 
|
171  | 
by (rule order_trans [OF norm_triangle_ineq add_mono])  | 
|
172  | 
||
173  | 
lemma ge_iff_diff_ge_0: "(a::'a::linordered_ring) \<ge> b == a - b \<ge> 0"  | 
|
174  | 
by (simp add: field_simps)  | 
|
175  | 
||
176  | 
lemma pth_1:  | 
|
177  | 
fixes x :: "'a::real_normed_vector"  | 
|
178  | 
shows "x == scaleR 1 x" by simp  | 
|
179  | 
||
180  | 
lemma pth_2:  | 
|
181  | 
fixes x :: "'a::real_normed_vector"  | 
|
182  | 
shows "x - y == x + -y" by (atomize (full)) simp  | 
|
183  | 
||
184  | 
lemma pth_3:  | 
|
185  | 
fixes x :: "'a::real_normed_vector"  | 
|
186  | 
shows "- x == scaleR (-1) x" by simp  | 
|
187  | 
||
188  | 
lemma pth_4:  | 
|
189  | 
fixes x :: "'a::real_normed_vector"  | 
|
190  | 
shows "scaleR 0 x == 0" and "scaleR c 0 = (0::'a)" by simp_all  | 
|
191  | 
||
192  | 
lemma pth_5:  | 
|
193  | 
fixes x :: "'a::real_normed_vector"  | 
|
194  | 
shows "scaleR c (scaleR d x) == scaleR (c * d) x" by simp  | 
|
195  | 
||
196  | 
lemma pth_6:  | 
|
197  | 
fixes x :: "'a::real_normed_vector"  | 
|
198  | 
shows "scaleR c (x + y) == scaleR c x + scaleR c y"  | 
|
199  | 
by (simp add: scaleR_right_distrib)  | 
|
200  | 
||
201  | 
lemma pth_7:  | 
|
202  | 
fixes x :: "'a::real_normed_vector"  | 
|
203  | 
shows "0 + x == x" and "x + 0 == x" by simp_all  | 
|
204  | 
||
205  | 
lemma pth_8:  | 
|
206  | 
fixes x :: "'a::real_normed_vector"  | 
|
207  | 
shows "scaleR c x + scaleR d x == scaleR (c + d) x"  | 
|
208  | 
by (simp add: scaleR_left_distrib)  | 
|
209  | 
||
210  | 
lemma pth_9:  | 
|
211  | 
fixes x :: "'a::real_normed_vector" shows  | 
|
212  | 
"(scaleR c x + z) + scaleR d x == scaleR (c + d) x + z"  | 
|
213  | 
"scaleR c x + (scaleR d x + z) == scaleR (c + d) x + z"  | 
|
214  | 
"(scaleR c x + w) + (scaleR d x + z) == scaleR (c + d) x + (w + z)"  | 
|
215  | 
by (simp_all add: algebra_simps)  | 
|
216  | 
||
217  | 
lemma pth_a:  | 
|
218  | 
fixes x :: "'a::real_normed_vector"  | 
|
219  | 
shows "scaleR 0 x + y == y" by simp  | 
|
220  | 
||
221  | 
lemma pth_b:  | 
|
222  | 
fixes x :: "'a::real_normed_vector" shows  | 
|
223  | 
"scaleR c x + scaleR d y == scaleR c x + scaleR d y"  | 
|
224  | 
"(scaleR c x + z) + scaleR d y == scaleR c x + (z + scaleR d y)"  | 
|
225  | 
"scaleR c x + (scaleR d y + z) == scaleR c x + (scaleR d y + z)"  | 
|
226  | 
"(scaleR c x + w) + (scaleR d y + z) == scaleR c x + (w + (scaleR d y + z))"  | 
|
227  | 
by (simp_all add: algebra_simps)  | 
|
228  | 
||
229  | 
lemma pth_c:  | 
|
230  | 
fixes x :: "'a::real_normed_vector" shows  | 
|
231  | 
"scaleR c x + scaleR d y == scaleR d y + scaleR c x"  | 
|
232  | 
"(scaleR c x + z) + scaleR d y == scaleR d y + (scaleR c x + z)"  | 
|
233  | 
"scaleR c x + (scaleR d y + z) == scaleR d y + (scaleR c x + z)"  | 
|
234  | 
"(scaleR c x + w) + (scaleR d y + z) == scaleR d y + ((scaleR c x + w) + z)"  | 
|
235  | 
by (simp_all add: algebra_simps)  | 
|
236  | 
||
237  | 
lemma pth_d:  | 
|
238  | 
fixes x :: "'a::real_normed_vector"  | 
|
239  | 
shows "x + 0 == x" by simp  | 
|
240  | 
||
241  | 
lemma norm_imp_pos_and_ge:  | 
|
242  | 
fixes x :: "'a::real_normed_vector"  | 
|
243  | 
shows "norm x == n \<Longrightarrow> norm x \<ge> 0 \<and> n \<ge> norm x"  | 
|
244  | 
by atomize auto  | 
|
245  | 
||
246  | 
lemma real_eq_0_iff_le_ge_0: "(x::real) = 0 == x \<ge> 0 \<and> -x \<ge> 0" by arith  | 
|
247  | 
||
248  | 
lemma norm_pths:  | 
|
249  | 
fixes x :: "'a::real_normed_vector" shows  | 
|
250  | 
"x = y \<longleftrightarrow> norm (x - y) \<le> 0"  | 
|
251  | 
"x \<noteq> y \<longleftrightarrow> \<not> (norm (x - y) \<le> 0)"  | 
|
252  | 
using norm_ge_zero[of "x - y"] by auto  | 
|
253  | 
||
254  | 
use "normarith.ML"  | 
|
255  | 
||
256  | 
method_setup norm = {* Scan.succeed (SIMPLE_METHOD' o NormArith.norm_arith_tac)
 | 
|
257  | 
*} "prove simple linear statements about vector norms"  | 
|
258  | 
||
259  | 
||
260  | 
text{* Hence more metric properties. *}
 | 
|
261  | 
||
262  | 
lemma norm_triangle_half_r:  | 
|
263  | 
shows "norm (y - x1) < e / 2 \<Longrightarrow> norm (y - x2) < e / 2 \<Longrightarrow> norm (x1 - x2) < e"  | 
|
264  | 
using dist_triangle_half_r unfolding dist_norm[THEN sym] by auto  | 
|
265  | 
||
266  | 
lemma norm_triangle_half_l: assumes "norm (x - y) < e / 2" "norm (x' - (y)) < e / 2"  | 
|
267  | 
shows "norm (x - x') < e"  | 
|
268  | 
using dist_triangle_half_l[OF assms[unfolded dist_norm[THEN sym]]]  | 
|
269  | 
unfolding dist_norm[THEN sym] .  | 
|
270  | 
||
271  | 
lemma norm_triangle_le: "norm(x) + norm y <= e ==> norm(x + y) <= e"  | 
|
272  | 
by (metis order_trans norm_triangle_ineq)  | 
|
273  | 
||
274  | 
lemma norm_triangle_lt: "norm(x) + norm(y) < e ==> norm(x + y) < e"  | 
|
275  | 
by (metis basic_trans_rules(21) norm_triangle_ineq)  | 
|
276  | 
||
277  | 
lemma dist_triangle_add:  | 
|
278  | 
fixes x y x' y' :: "'a::real_normed_vector"  | 
|
279  | 
shows "dist (x + y) (x' + y') <= dist x x' + dist y y'"  | 
|
280  | 
by norm  | 
|
281  | 
||
282  | 
lemma dist_triangle_add_half:  | 
|
283  | 
fixes x x' y y' :: "'a::real_normed_vector"  | 
|
284  | 
shows "dist x x' < e / 2 \<Longrightarrow> dist y y' < e / 2 \<Longrightarrow> dist(x + y) (x' + y') < e"  | 
|
285  | 
by norm  | 
|
286  | 
||
287  | 
lemma setsum_clauses:  | 
|
288  | 
  shows "setsum f {} = 0"
 | 
|
289  | 
and "finite S \<Longrightarrow> setsum f (insert x S) =  | 
|
290  | 
(if x \<in> S then setsum f S else f x + setsum f S)"  | 
|
291  | 
by (auto simp add: insert_absorb)  | 
|
292  | 
||
293  | 
lemma setsum_norm_le:  | 
|
294  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"  | 
|
| 
44176
 
eda112e9cdee
remove redundant lemma setsum_norm in favor of norm_setsum;
 
huffman 
parents: 
44170 
diff
changeset
 | 
295  | 
assumes fg: "\<forall>x \<in> S. norm (f x) \<le> g x"  | 
| 44133 | 296  | 
shows "norm (setsum f S) \<le> setsum g S"  | 
| 
44176
 
eda112e9cdee
remove redundant lemma setsum_norm in favor of norm_setsum;
 
huffman 
parents: 
44170 
diff
changeset
 | 
297  | 
by (rule order_trans [OF norm_setsum setsum_mono], simp add: fg)  | 
| 44133 | 298  | 
|
299  | 
lemma setsum_norm_bound:  | 
|
300  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"  | 
|
301  | 
assumes fS: "finite S"  | 
|
302  | 
and K: "\<forall>x \<in> S. norm (f x) \<le> K"  | 
|
303  | 
shows "norm (setsum f S) \<le> of_nat (card S) * K"  | 
|
| 
44176
 
eda112e9cdee
remove redundant lemma setsum_norm in favor of norm_setsum;
 
huffman 
parents: 
44170 
diff
changeset
 | 
304  | 
using setsum_norm_le[OF K] setsum_constant[symmetric]  | 
| 44133 | 305  | 
by simp  | 
306  | 
||
307  | 
lemma setsum_group:  | 
|
308  | 
assumes fS: "finite S" and fT: "finite T" and fST: "f ` S \<subseteq> T"  | 
|
309  | 
  shows "setsum (\<lambda>y. setsum g {x. x\<in> S \<and> f x = y}) T = setsum g S"
 | 
|
310  | 
apply (subst setsum_image_gen[OF fS, of g f])  | 
|
311  | 
apply (rule setsum_mono_zero_right[OF fT fST])  | 
|
312  | 
by (auto intro: setsum_0')  | 
|
313  | 
||
314  | 
lemma dot_lsum: "finite S \<Longrightarrow> setsum f S \<bullet> y = setsum (\<lambda>x. f x \<bullet> y) S "  | 
|
| 44454 | 315  | 
apply(induct rule: finite_induct) by(auto simp add: inner_add)  | 
| 44133 | 316  | 
|
317  | 
lemma dot_rsum: "finite S \<Longrightarrow> y \<bullet> setsum f S = setsum (\<lambda>x. y \<bullet> f x) S "  | 
|
| 44454 | 318  | 
apply(induct rule: finite_induct) by(auto simp add: inner_add)  | 
| 44133 | 319  | 
|
320  | 
lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = z"  | 
|
321  | 
proof  | 
|
322  | 
assume "\<forall>x. x \<bullet> y = x \<bullet> z"  | 
|
| 44454 | 323  | 
hence "\<forall>x. x \<bullet> (y - z) = 0" by (simp add: inner_diff)  | 
| 44133 | 324  | 
hence "(y - z) \<bullet> (y - z) = 0" ..  | 
325  | 
thus "y = z" by simp  | 
|
326  | 
qed simp  | 
|
327  | 
||
328  | 
lemma vector_eq_rdot: "(\<forall>z. x \<bullet> z = y \<bullet> z) \<longleftrightarrow> x = y"  | 
|
329  | 
proof  | 
|
330  | 
assume "\<forall>z. x \<bullet> z = y \<bullet> z"  | 
|
| 44454 | 331  | 
hence "\<forall>z. (x - y) \<bullet> z = 0" by (simp add: inner_diff)  | 
| 44133 | 332  | 
hence "(x - y) \<bullet> (x - y) = 0" ..  | 
333  | 
thus "x = y" by simp  | 
|
334  | 
qed simp  | 
|
335  | 
||
336  | 
subsection{* Orthogonality. *}
 | 
|
337  | 
||
338  | 
context real_inner  | 
|
339  | 
begin  | 
|
340  | 
||
341  | 
definition "orthogonal x y \<longleftrightarrow> (x \<bullet> y = 0)"  | 
|
342  | 
||
343  | 
lemma orthogonal_clauses:  | 
|
344  | 
"orthogonal a 0"  | 
|
345  | 
"orthogonal a x \<Longrightarrow> orthogonal a (c *\<^sub>R x)"  | 
|
346  | 
"orthogonal a x \<Longrightarrow> orthogonal a (-x)"  | 
|
347  | 
"orthogonal a x \<Longrightarrow> orthogonal a y \<Longrightarrow> orthogonal a (x + y)"  | 
|
348  | 
"orthogonal a x \<Longrightarrow> orthogonal a y \<Longrightarrow> orthogonal a (x - y)"  | 
|
349  | 
"orthogonal 0 a"  | 
|
350  | 
"orthogonal x a \<Longrightarrow> orthogonal (c *\<^sub>R x) a"  | 
|
351  | 
"orthogonal x a \<Longrightarrow> orthogonal (-x) a"  | 
|
352  | 
"orthogonal x a \<Longrightarrow> orthogonal y a \<Longrightarrow> orthogonal (x + y) a"  | 
|
353  | 
"orthogonal x a \<Longrightarrow> orthogonal y a \<Longrightarrow> orthogonal (x - y) a"  | 
|
354  | 
unfolding orthogonal_def inner_simps inner_add_left inner_add_right inner_diff_left inner_diff_right (*FIXME*) by auto  | 
|
355  | 
||
356  | 
end  | 
|
357  | 
||
358  | 
lemma orthogonal_commute: "orthogonal x y \<longleftrightarrow> orthogonal y x"  | 
|
359  | 
by (simp add: orthogonal_def inner_commute)  | 
|
360  | 
||
361  | 
subsection{* Linear functions. *}
 | 
|
362  | 
||
363  | 
definition  | 
|
364  | 
  linear :: "('a::real_vector \<Rightarrow> 'b::real_vector) \<Rightarrow> bool" where
 | 
|
365  | 
"linear f \<longleftrightarrow> (\<forall>x y. f(x + y) = f x + f y) \<and> (\<forall>c x. f(c *\<^sub>R x) = c *\<^sub>R f x)"  | 
|
366  | 
||
367  | 
lemma linearI: assumes "\<And>x y. f (x + y) = f x + f y" "\<And>c x. f (c *\<^sub>R x) = c *\<^sub>R f x"  | 
|
368  | 
shows "linear f" using assms unfolding linear_def by auto  | 
|
369  | 
||
370  | 
lemma linear_compose_cmul: "linear f ==> linear (\<lambda>x. c *\<^sub>R f x)"  | 
|
371  | 
by (simp add: linear_def algebra_simps)  | 
|
372  | 
||
373  | 
lemma linear_compose_neg: "linear f ==> linear (\<lambda>x. -(f(x)))"  | 
|
374  | 
by (simp add: linear_def)  | 
|
375  | 
||
376  | 
lemma linear_compose_add: "linear f \<Longrightarrow> linear g ==> linear (\<lambda>x. f(x) + g(x))"  | 
|
377  | 
by (simp add: linear_def algebra_simps)  | 
|
378  | 
||
379  | 
lemma linear_compose_sub: "linear f \<Longrightarrow> linear g ==> linear (\<lambda>x. f x - g x)"  | 
|
380  | 
by (simp add: linear_def algebra_simps)  | 
|
381  | 
||
382  | 
lemma linear_compose: "linear f \<Longrightarrow> linear g ==> linear (g o f)"  | 
|
383  | 
by (simp add: linear_def)  | 
|
384  | 
||
385  | 
lemma linear_id: "linear id" by (simp add: linear_def id_def)  | 
|
386  | 
||
387  | 
lemma linear_zero: "linear (\<lambda>x. 0)" by (simp add: linear_def)  | 
|
388  | 
||
389  | 
lemma linear_compose_setsum:  | 
|
390  | 
assumes fS: "finite S" and lS: "\<forall>a \<in> S. linear (f a)"  | 
|
391  | 
shows "linear(\<lambda>x. setsum (\<lambda>a. f a x) S)"  | 
|
392  | 
using lS  | 
|
393  | 
apply (induct rule: finite_induct[OF fS])  | 
|
394  | 
by (auto simp add: linear_zero intro: linear_compose_add)  | 
|
395  | 
||
396  | 
lemma linear_0: "linear f \<Longrightarrow> f 0 = 0"  | 
|
397  | 
unfolding linear_def  | 
|
398  | 
apply clarsimp  | 
|
399  | 
apply (erule allE[where x="0::'a"])  | 
|
400  | 
apply simp  | 
|
401  | 
done  | 
|
402  | 
||
403  | 
lemma linear_cmul: "linear f ==> f(c *\<^sub>R x) = c *\<^sub>R f x" by (simp add: linear_def)  | 
|
404  | 
||
405  | 
lemma linear_neg: "linear f ==> f (-x) = - f x"  | 
|
406  | 
using linear_cmul [where c="-1"] by simp  | 
|
407  | 
||
408  | 
lemma linear_add: "linear f ==> f(x + y) = f x + f y" by (metis linear_def)  | 
|
409  | 
||
410  | 
lemma linear_sub: "linear f ==> f(x - y) = f x - f y"  | 
|
411  | 
by (simp add: diff_minus linear_add linear_neg)  | 
|
412  | 
||
413  | 
lemma linear_setsum:  | 
|
414  | 
assumes lf: "linear f" and fS: "finite S"  | 
|
415  | 
shows "f (setsum g S) = setsum (f o g) S"  | 
|
416  | 
proof (induct rule: finite_induct[OF fS])  | 
|
417  | 
case 1 thus ?case by (simp add: linear_0[OF lf])  | 
|
418  | 
next  | 
|
419  | 
case (2 x F)  | 
|
420  | 
have "f (setsum g (insert x F)) = f (g x + setsum g F)" using "2.hyps"  | 
|
421  | 
by simp  | 
|
422  | 
also have "\<dots> = f (g x) + f (setsum g F)" using linear_add[OF lf] by simp  | 
|
423  | 
also have "\<dots> = setsum (f o g) (insert x F)" using "2.hyps" by simp  | 
|
424  | 
finally show ?case .  | 
|
425  | 
qed  | 
|
426  | 
||
427  | 
lemma linear_setsum_mul:  | 
|
428  | 
assumes lf: "linear f" and fS: "finite S"  | 
|
429  | 
shows "f (setsum (\<lambda>i. c i *\<^sub>R v i) S) = setsum (\<lambda>i. c i *\<^sub>R f (v i)) S"  | 
|
430  | 
using linear_setsum[OF lf fS, of "\<lambda>i. c i *\<^sub>R v i" , unfolded o_def]  | 
|
431  | 
linear_cmul[OF lf] by simp  | 
|
432  | 
||
433  | 
lemma linear_injective_0:  | 
|
434  | 
assumes lf: "linear f"  | 
|
435  | 
shows "inj f \<longleftrightarrow> (\<forall>x. f x = 0 \<longrightarrow> x = 0)"  | 
|
436  | 
proof-  | 
|
437  | 
have "inj f \<longleftrightarrow> (\<forall> x y. f x = f y \<longrightarrow> x = y)" by (simp add: inj_on_def)  | 
|
438  | 
also have "\<dots> \<longleftrightarrow> (\<forall> x y. f x - f y = 0 \<longrightarrow> x - y = 0)" by simp  | 
|
439  | 
also have "\<dots> \<longleftrightarrow> (\<forall> x y. f (x - y) = 0 \<longrightarrow> x - y = 0)"  | 
|
440  | 
by (simp add: linear_sub[OF lf])  | 
|
441  | 
also have "\<dots> \<longleftrightarrow> (\<forall> x. f x = 0 \<longrightarrow> x = 0)" by auto  | 
|
442  | 
finally show ?thesis .  | 
|
443  | 
qed  | 
|
444  | 
||
445  | 
subsection{* Bilinear functions. *}
 | 
|
446  | 
||
447  | 
definition "bilinear f \<longleftrightarrow> (\<forall>x. linear(\<lambda>y. f x y)) \<and> (\<forall>y. linear(\<lambda>x. f x y))"  | 
|
448  | 
||
449  | 
lemma bilinear_ladd: "bilinear h ==> h (x + y) z = (h x z) + (h y z)"  | 
|
450  | 
by (simp add: bilinear_def linear_def)  | 
|
451  | 
lemma bilinear_radd: "bilinear h ==> h x (y + z) = (h x y) + (h x z)"  | 
|
452  | 
by (simp add: bilinear_def linear_def)  | 
|
453  | 
||
454  | 
lemma bilinear_lmul: "bilinear h ==> h (c *\<^sub>R x) y = c *\<^sub>R (h x y)"  | 
|
455  | 
by (simp add: bilinear_def linear_def)  | 
|
456  | 
||
457  | 
lemma bilinear_rmul: "bilinear h ==> h x (c *\<^sub>R y) = c *\<^sub>R (h x y)"  | 
|
458  | 
by (simp add: bilinear_def linear_def)  | 
|
459  | 
||
460  | 
lemma bilinear_lneg: "bilinear h ==> h (- x) y = -(h x y)"  | 
|
461  | 
by (simp only: scaleR_minus1_left [symmetric] bilinear_lmul)  | 
|
462  | 
||
463  | 
lemma bilinear_rneg: "bilinear h ==> h x (- y) = - h x y"  | 
|
464  | 
by (simp only: scaleR_minus1_left [symmetric] bilinear_rmul)  | 
|
465  | 
||
466  | 
lemma (in ab_group_add) eq_add_iff: "x = x + y \<longleftrightarrow> y = 0"  | 
|
467  | 
using add_imp_eq[of x y 0] by auto  | 
|
468  | 
||
469  | 
lemma bilinear_lzero:  | 
|
470  | 
assumes bh: "bilinear h" shows "h 0 x = 0"  | 
|
471  | 
using bilinear_ladd[OF bh, of 0 0 x]  | 
|
472  | 
by (simp add: eq_add_iff field_simps)  | 
|
473  | 
||
474  | 
lemma bilinear_rzero:  | 
|
475  | 
assumes bh: "bilinear h" shows "h x 0 = 0"  | 
|
476  | 
using bilinear_radd[OF bh, of x 0 0 ]  | 
|
477  | 
by (simp add: eq_add_iff field_simps)  | 
|
478  | 
||
479  | 
lemma bilinear_lsub: "bilinear h ==> h (x - y) z = h x z - h y z"  | 
|
480  | 
by (simp add: diff_minus bilinear_ladd bilinear_lneg)  | 
|
481  | 
||
482  | 
lemma bilinear_rsub: "bilinear h ==> h z (x - y) = h z x - h z y"  | 
|
483  | 
by (simp add: diff_minus bilinear_radd bilinear_rneg)  | 
|
484  | 
||
485  | 
lemma bilinear_setsum:  | 
|
486  | 
assumes bh: "bilinear h" and fS: "finite S" and fT: "finite T"  | 
|
487  | 
shows "h (setsum f S) (setsum g T) = setsum (\<lambda>(i,j). h (f i) (g j)) (S \<times> T) "  | 
|
488  | 
proof-  | 
|
489  | 
have "h (setsum f S) (setsum g T) = setsum (\<lambda>x. h (f x) (setsum g T)) S"  | 
|
490  | 
apply (rule linear_setsum[unfolded o_def])  | 
|
491  | 
using bh fS by (auto simp add: bilinear_def)  | 
|
492  | 
also have "\<dots> = setsum (\<lambda>x. setsum (\<lambda>y. h (f x) (g y)) T) S"  | 
|
493  | 
apply (rule setsum_cong, simp)  | 
|
494  | 
apply (rule linear_setsum[unfolded o_def])  | 
|
495  | 
using bh fT by (auto simp add: bilinear_def)  | 
|
496  | 
finally show ?thesis unfolding setsum_cartesian_product .  | 
|
497  | 
qed  | 
|
498  | 
||
499  | 
subsection{* Adjoints. *}
 | 
|
500  | 
||
501  | 
definition "adjoint f = (SOME f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y)"  | 
|
502  | 
||
503  | 
lemma adjoint_unique:  | 
|
504  | 
assumes "\<forall>x y. inner (f x) y = inner x (g y)"  | 
|
505  | 
shows "adjoint f = g"  | 
|
506  | 
unfolding adjoint_def  | 
|
507  | 
proof (rule some_equality)  | 
|
508  | 
show "\<forall>x y. inner (f x) y = inner x (g y)" using assms .  | 
|
509  | 
next  | 
|
510  | 
fix h assume "\<forall>x y. inner (f x) y = inner x (h y)"  | 
|
511  | 
hence "\<forall>x y. inner x (g y) = inner x (h y)" using assms by simp  | 
|
512  | 
hence "\<forall>x y. inner x (g y - h y) = 0" by (simp add: inner_diff_right)  | 
|
513  | 
hence "\<forall>y. inner (g y - h y) (g y - h y) = 0" by simp  | 
|
514  | 
hence "\<forall>y. h y = g y" by simp  | 
|
515  | 
thus "h = g" by (simp add: ext)  | 
|
516  | 
qed  | 
|
517  | 
||
518  | 
lemma choice_iff: "(\<forall>x. \<exists>y. P x y) \<longleftrightarrow> (\<exists>f. \<forall>x. P x (f x))" by metis  | 
|
519  | 
||
520  | 
subsection{* Interlude: Some properties of real sets *}
 | 
|
521  | 
||
522  | 
lemma seq_mono_lemma: assumes "\<forall>(n::nat) \<ge> m. (d n :: real) < e n" and "\<forall>n \<ge> m. e n <= e m"  | 
|
523  | 
shows "\<forall>n \<ge> m. d n < e m"  | 
|
524  | 
using assms apply auto  | 
|
525  | 
apply (erule_tac x="n" in allE)  | 
|
526  | 
apply (erule_tac x="n" in allE)  | 
|
527  | 
apply auto  | 
|
528  | 
done  | 
|
529  | 
||
530  | 
||
531  | 
lemma infinite_enumerate: assumes fS: "infinite S"  | 
|
532  | 
shows "\<exists>r. subseq r \<and> (\<forall>n. r n \<in> S)"  | 
|
533  | 
unfolding subseq_def  | 
|
534  | 
using enumerate_in_set[OF fS] enumerate_mono[of _ _ S] fS by auto  | 
|
535  | 
||
536  | 
lemma approachable_lt_le: "(\<exists>(d::real)>0. \<forall>x. f x < d \<longrightarrow> P x) \<longleftrightarrow> (\<exists>d>0. \<forall>x. f x \<le> d \<longrightarrow> P x)"  | 
|
537  | 
apply auto  | 
|
538  | 
apply (rule_tac x="d/2" in exI)  | 
|
539  | 
apply auto  | 
|
540  | 
done  | 
|
541  | 
||
542  | 
||
543  | 
lemma triangle_lemma:  | 
|
544  | 
assumes x: "0 <= (x::real)" and y:"0 <= y" and z: "0 <= z" and xy: "x^2 <= y^2 + z^2"  | 
|
545  | 
shows "x <= y + z"  | 
|
546  | 
proof-  | 
|
547  | 
have "y^2 + z^2 \<le> y^2 + 2*y*z + z^2" using z y by (simp add: mult_nonneg_nonneg)  | 
|
548  | 
with xy have th: "x ^2 \<le> (y+z)^2" by (simp add: power2_eq_square field_simps)  | 
|
549  | 
from y z have yz: "y + z \<ge> 0" by arith  | 
|
550  | 
from power2_le_imp_le[OF th yz] show ?thesis .  | 
|
551  | 
qed  | 
|
552  | 
||
553  | 
text {* TODO: move to NthRoot *}
 | 
|
554  | 
lemma sqrt_add_le_add_sqrt:  | 
|
555  | 
assumes x: "0 \<le> x" and y: "0 \<le> y"  | 
|
556  | 
shows "sqrt (x + y) \<le> sqrt x + sqrt y"  | 
|
557  | 
apply (rule power2_le_imp_le)  | 
|
| 44142 | 558  | 
apply (simp add: real_sum_squared_expand x y)  | 
| 44133 | 559  | 
apply (simp add: mult_nonneg_nonneg x y)  | 
| 44142 | 560  | 
apply (simp add: x y)  | 
| 44133 | 561  | 
done  | 
562  | 
||
563  | 
subsection {* A generic notion of "hull" (convex, affine, conic hull and closure). *}
 | 
|
564  | 
||
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
565  | 
definition hull :: "('a set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "hull" 75) where
 | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
566  | 
  "S hull s = Inter {t. S t \<and> s \<subseteq> t}"
 | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
567  | 
|
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
568  | 
lemma hull_same: "S s \<Longrightarrow> S hull s = s"  | 
| 44133 | 569  | 
unfolding hull_def by auto  | 
570  | 
||
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
571  | 
lemma hull_in: "(\<And>T. Ball T S ==> S (Inter T)) ==> S (S hull s)"  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
572  | 
unfolding hull_def Ball_def by auto  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
573  | 
|
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
574  | 
lemma hull_eq: "(\<And>T. Ball T S ==> S (Inter T)) ==> (S hull s) = s \<longleftrightarrow> S s"  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
575  | 
using hull_same[of S s] hull_in[of S s] by metis  | 
| 44133 | 576  | 
|
577  | 
||
578  | 
lemma hull_hull: "S hull (S hull s) = S hull s"  | 
|
579  | 
unfolding hull_def by blast  | 
|
580  | 
||
581  | 
lemma hull_subset[intro]: "s \<subseteq> (S hull s)"  | 
|
582  | 
unfolding hull_def by blast  | 
|
583  | 
||
584  | 
lemma hull_mono: " s \<subseteq> t ==> (S hull s) \<subseteq> (S hull t)"  | 
|
585  | 
unfolding hull_def by blast  | 
|
586  | 
||
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
587  | 
lemma hull_antimono: "\<forall>x. S x \<longrightarrow> T x ==> (T hull s) \<subseteq> (S hull s)"  | 
| 44133 | 588  | 
unfolding hull_def by blast  | 
589  | 
||
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
590  | 
lemma hull_minimal: "s \<subseteq> t \<Longrightarrow> S t ==> (S hull s) \<subseteq> t"  | 
| 44133 | 591  | 
unfolding hull_def by blast  | 
592  | 
||
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
593  | 
lemma subset_hull: "S t ==> S hull s \<subseteq> t \<longleftrightarrow> s \<subseteq> t"  | 
| 44133 | 594  | 
unfolding hull_def by blast  | 
595  | 
||
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
596  | 
lemma hull_unique: "s \<subseteq> t \<Longrightarrow> S t \<Longrightarrow> (\<And>t'. s \<subseteq> t' \<Longrightarrow> S t' ==> t \<subseteq> t')  | 
| 44133 | 597  | 
==> (S hull s = t)"  | 
598  | 
unfolding hull_def by auto  | 
|
599  | 
||
600  | 
lemma hull_induct: "(\<And>x. x\<in> S \<Longrightarrow> P x) \<Longrightarrow> Q {x. P x} \<Longrightarrow> \<forall>x\<in> Q hull S. P x"
 | 
|
601  | 
  using hull_minimal[of S "{x. P x}" Q]
 | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
602  | 
by (auto simp add: subset_eq)  | 
| 44133 | 603  | 
|
604  | 
lemma hull_inc: "x \<in> S \<Longrightarrow> x \<in> P hull S" by (metis hull_subset subset_eq)  | 
|
605  | 
||
606  | 
lemma hull_union_subset: "(S hull s) \<union> (S hull t) \<subseteq> (S hull (s \<union> t))"  | 
|
607  | 
unfolding Un_subset_iff by (metis hull_mono Un_upper1 Un_upper2)  | 
|
608  | 
||
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
609  | 
lemma hull_union: assumes T: "\<And>T. Ball T S ==> S (Inter T)"  | 
| 44133 | 610  | 
shows "S hull (s \<union> t) = S hull (S hull s \<union> S hull t)"  | 
611  | 
apply rule  | 
|
612  | 
apply (rule hull_mono)  | 
|
613  | 
unfolding Un_subset_iff  | 
|
614  | 
apply (metis hull_subset Un_upper1 Un_upper2 subset_trans)  | 
|
615  | 
apply (rule hull_minimal)  | 
|
616  | 
apply (metis hull_union_subset)  | 
|
617  | 
apply (metis hull_in T)  | 
|
618  | 
done  | 
|
619  | 
||
620  | 
lemma hull_redundant_eq: "a \<in> (S hull s) \<longleftrightarrow> (S hull (insert a s) = S hull s)"  | 
|
621  | 
unfolding hull_def by blast  | 
|
622  | 
||
623  | 
lemma hull_redundant: "a \<in> (S hull s) ==> (S hull (insert a s) = S hull s)"  | 
|
624  | 
by (metis hull_redundant_eq)  | 
|
625  | 
||
626  | 
text{* Archimedian properties and useful consequences. *}
 | 
|
627  | 
||
628  | 
lemma real_arch_simple: "\<exists>n. x <= real (n::nat)"  | 
|
629  | 
using reals_Archimedean2[of x] apply auto by (rule_tac x="Suc n" in exI, auto)  | 
|
630  | 
lemmas real_arch_lt = reals_Archimedean2  | 
|
631  | 
||
632  | 
lemmas real_arch = reals_Archimedean3  | 
|
633  | 
||
634  | 
lemma real_arch_inv: "0 < e \<longleftrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> 0 < inverse (real n) \<and> inverse (real n) < e)"  | 
|
635  | 
using reals_Archimedean  | 
|
636  | 
apply (auto simp add: field_simps)  | 
|
637  | 
apply (subgoal_tac "inverse (real n) > 0")  | 
|
638  | 
apply arith  | 
|
639  | 
apply simp  | 
|
640  | 
done  | 
|
641  | 
||
642  | 
lemma real_pow_lbound: "0 <= x ==> 1 + real n * x <= (1 + x) ^ n"  | 
|
643  | 
proof(induct n)  | 
|
644  | 
case 0 thus ?case by simp  | 
|
645  | 
next  | 
|
646  | 
case (Suc n)  | 
|
647  | 
hence h: "1 + real n * x \<le> (1 + x) ^ n" by simp  | 
|
648  | 
from h have p: "1 \<le> (1 + x) ^ n" using Suc.prems by simp  | 
|
649  | 
from h have "1 + real n * x + x \<le> (1 + x) ^ n + x" by simp  | 
|
650  | 
also have "\<dots> \<le> (1 + x) ^ Suc n" apply (subst diff_le_0_iff_le[symmetric])  | 
|
651  | 
apply (simp add: field_simps)  | 
|
652  | 
using mult_left_mono[OF p Suc.prems] by simp  | 
|
653  | 
finally show ?case by (simp add: real_of_nat_Suc field_simps)  | 
|
654  | 
qed  | 
|
655  | 
||
656  | 
lemma real_arch_pow: assumes x: "1 < (x::real)" shows "\<exists>n. y < x^n"  | 
|
657  | 
proof-  | 
|
658  | 
from x have x0: "x - 1 > 0" by arith  | 
|
659  | 
from real_arch[OF x0, rule_format, of y]  | 
|
660  | 
obtain n::nat where n:"y < real n * (x - 1)" by metis  | 
|
661  | 
from x0 have x00: "x- 1 \<ge> 0" by arith  | 
|
662  | 
from real_pow_lbound[OF x00, of n] n  | 
|
663  | 
have "y < x^n" by auto  | 
|
664  | 
then show ?thesis by metis  | 
|
665  | 
qed  | 
|
666  | 
||
667  | 
lemma real_arch_pow2: "\<exists>n. (x::real) < 2^ n"  | 
|
668  | 
using real_arch_pow[of 2 x] by simp  | 
|
669  | 
||
670  | 
lemma real_arch_pow_inv: assumes y: "(y::real) > 0" and x1: "x < 1"  | 
|
671  | 
shows "\<exists>n. x^n < y"  | 
|
672  | 
proof-  | 
|
673  | 
  {assume x0: "x > 0"
 | 
|
674  | 
from x0 x1 have ix: "1 < 1/x" by (simp add: field_simps)  | 
|
675  | 
from real_arch_pow[OF ix, of "1/y"]  | 
|
676  | 
obtain n where n: "1/y < (1/x)^n" by blast  | 
|
677  | 
then  | 
|
678  | 
have ?thesis using y x0 by (auto simp add: field_simps power_divide) }  | 
|
679  | 
moreover  | 
|
680  | 
  {assume "\<not> x > 0" with y x1 have ?thesis apply auto by (rule exI[where x=1], auto)}
 | 
|
681  | 
ultimately show ?thesis by metis  | 
|
682  | 
qed  | 
|
683  | 
||
684  | 
lemma forall_pos_mono: "(\<And>d e::real. d < e \<Longrightarrow> P d ==> P e) \<Longrightarrow> (\<And>n::nat. n \<noteq> 0 ==> P(inverse(real n))) \<Longrightarrow> (\<And>e. 0 < e ==> P e)"  | 
|
685  | 
by (metis real_arch_inv)  | 
|
686  | 
||
687  | 
lemma forall_pos_mono_1: "(\<And>d e::real. d < e \<Longrightarrow> P d ==> P e) \<Longrightarrow> (\<And>n. P(inverse(real (Suc n)))) ==> 0 < e ==> P e"  | 
|
688  | 
apply (rule forall_pos_mono)  | 
|
689  | 
apply auto  | 
|
690  | 
apply (atomize)  | 
|
691  | 
apply (erule_tac x="n - 1" in allE)  | 
|
692  | 
apply auto  | 
|
693  | 
done  | 
|
694  | 
||
695  | 
lemma real_archimedian_rdiv_eq_0: assumes x0: "x \<ge> 0" and c: "c \<ge> 0" and xc: "\<forall>(m::nat)>0. real m * x \<le> c"  | 
|
696  | 
shows "x = 0"  | 
|
697  | 
proof-  | 
|
698  | 
  {assume "x \<noteq> 0" with x0 have xp: "x > 0" by arith
 | 
|
699  | 
from real_arch[OF xp, rule_format, of c] obtain n::nat where n: "c < real n * x" by blast  | 
|
700  | 
with xc[rule_format, of n] have "n = 0" by arith  | 
|
701  | 
with n c have False by simp}  | 
|
702  | 
then show ?thesis by blast  | 
|
703  | 
qed  | 
|
704  | 
||
705  | 
subsection{* A bit of linear algebra. *}
 | 
|
706  | 
||
707  | 
definition (in real_vector)  | 
|
708  | 
subspace :: "'a set \<Rightarrow> bool" where  | 
|
709  | 
"subspace S \<longleftrightarrow> 0 \<in> S \<and> (\<forall>x\<in> S. \<forall>y \<in>S. x + y \<in> S) \<and> (\<forall>c. \<forall>x \<in>S. c *\<^sub>R x \<in>S )"  | 
|
710  | 
||
711  | 
definition (in real_vector) "span S = (subspace hull S)"  | 
|
712  | 
definition (in real_vector) "dependent S \<longleftrightarrow> (\<exists>a \<in> S. a \<in> span(S - {a}))"
 | 
|
713  | 
abbreviation (in real_vector) "independent s == ~(dependent s)"  | 
|
714  | 
||
715  | 
text {* Closure properties of subspaces. *}
 | 
|
716  | 
||
717  | 
lemma subspace_UNIV[simp]: "subspace(UNIV)" by (simp add: subspace_def)  | 
|
718  | 
||
719  | 
lemma (in real_vector) subspace_0: "subspace S ==> 0 \<in> S" by (metis subspace_def)  | 
|
720  | 
||
721  | 
lemma (in real_vector) subspace_add: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S ==> x + y \<in> S"  | 
|
722  | 
by (metis subspace_def)  | 
|
723  | 
||
724  | 
lemma (in real_vector) subspace_mul: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> c *\<^sub>R x \<in> S"  | 
|
725  | 
by (metis subspace_def)  | 
|
726  | 
||
727  | 
lemma subspace_neg: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> - x \<in> S"  | 
|
728  | 
by (metis scaleR_minus1_left subspace_mul)  | 
|
729  | 
||
730  | 
lemma subspace_sub: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x - y \<in> S"  | 
|
731  | 
by (metis diff_minus subspace_add subspace_neg)  | 
|
732  | 
||
733  | 
lemma (in real_vector) subspace_setsum:  | 
|
734  | 
assumes sA: "subspace A" and fB: "finite B"  | 
|
735  | 
and f: "\<forall>x\<in> B. f x \<in> A"  | 
|
736  | 
shows "setsum f B \<in> A"  | 
|
737  | 
using fB f sA  | 
|
738  | 
apply(induct rule: finite_induct[OF fB])  | 
|
739  | 
by (simp add: subspace_def sA, auto simp add: sA subspace_add)  | 
|
740  | 
||
741  | 
lemma subspace_linear_image:  | 
|
742  | 
assumes lf: "linear f" and sS: "subspace S"  | 
|
743  | 
shows "subspace(f ` S)"  | 
|
744  | 
using lf sS linear_0[OF lf]  | 
|
745  | 
unfolding linear_def subspace_def  | 
|
746  | 
apply (auto simp add: image_iff)  | 
|
747  | 
apply (rule_tac x="x + y" in bexI, auto)  | 
|
748  | 
apply (rule_tac x="c *\<^sub>R x" in bexI, auto)  | 
|
749  | 
done  | 
|
750  | 
||
751  | 
lemma subspace_linear_preimage: "linear f ==> subspace S ==> subspace {x. f x \<in> S}"
 | 
|
752  | 
by (auto simp add: subspace_def linear_def linear_0[of f])  | 
|
753  | 
||
754  | 
lemma subspace_trivial: "subspace {0}"
 | 
|
755  | 
by (simp add: subspace_def)  | 
|
756  | 
||
757  | 
lemma (in real_vector) subspace_inter: "subspace A \<Longrightarrow> subspace B ==> subspace (A \<inter> B)"  | 
|
758  | 
by (simp add: subspace_def)  | 
|
759  | 
||
760  | 
lemma (in real_vector) span_mono: "A \<subseteq> B ==> span A \<subseteq> span B"  | 
|
761  | 
by (metis span_def hull_mono)  | 
|
762  | 
||
763  | 
lemma (in real_vector) subspace_span: "subspace(span S)"  | 
|
764  | 
unfolding span_def  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
765  | 
apply (rule hull_in)  | 
| 44133 | 766  | 
apply (simp only: subspace_def Inter_iff Int_iff subset_eq)  | 
767  | 
apply auto  | 
|
768  | 
done  | 
|
769  | 
||
770  | 
lemma (in real_vector) span_clauses:  | 
|
771  | 
"a \<in> S ==> a \<in> span S"  | 
|
772  | 
"0 \<in> span S"  | 
|
773  | 
"x\<in> span S \<Longrightarrow> y \<in> span S ==> x + y \<in> span S"  | 
|
774  | 
"x \<in> span S \<Longrightarrow> c *\<^sub>R x \<in> span S"  | 
|
775  | 
by (metis span_def hull_subset subset_eq)  | 
|
776  | 
(metis subspace_span subspace_def)+  | 
|
777  | 
||
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
778  | 
lemma (in real_vector) span_induct: assumes SP: "\<And>x. x \<in> S ==> x \<in> P"  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
779  | 
and P: "subspace P" and x: "x \<in> span S" shows "x \<in> P"  | 
| 44133 | 780  | 
proof-  | 
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
781  | 
from SP have SP': "S \<subseteq> P" by (simp add: subset_eq)  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
782  | 
from x hull_minimal[where S=subspace, OF SP' P, unfolded span_def[symmetric]]  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
783  | 
show "x \<in> P" by (metis subset_eq)  | 
| 44133 | 784  | 
qed  | 
785  | 
||
786  | 
lemma span_empty[simp]: "span {} = {0}"
 | 
|
787  | 
apply (simp add: span_def)  | 
|
788  | 
apply (rule hull_unique)  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
789  | 
apply (auto simp add: subspace_def)  | 
| 44133 | 790  | 
done  | 
791  | 
||
792  | 
lemma (in real_vector) independent_empty[intro]: "independent {}"
 | 
|
793  | 
by (simp add: dependent_def)  | 
|
794  | 
||
795  | 
lemma dependent_single[simp]:  | 
|
796  | 
  "dependent {x} \<longleftrightarrow> x = 0"
 | 
|
797  | 
unfolding dependent_def by auto  | 
|
798  | 
||
799  | 
lemma (in real_vector) independent_mono: "independent A \<Longrightarrow> B \<subseteq> A ==> independent B"  | 
|
800  | 
apply (clarsimp simp add: dependent_def span_mono)  | 
|
801  | 
  apply (subgoal_tac "span (B - {a}) \<le> span (A - {a})")
 | 
|
802  | 
apply force  | 
|
803  | 
apply (rule span_mono)  | 
|
804  | 
apply auto  | 
|
805  | 
done  | 
|
806  | 
||
807  | 
lemma (in real_vector) span_subspace: "A \<subseteq> B \<Longrightarrow> B \<le> span A \<Longrightarrow> subspace B \<Longrightarrow> span A = B"  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
808  | 
by (metis order_antisym span_def hull_minimal)  | 
| 44133 | 809  | 
|
810  | 
lemma (in real_vector) span_induct': assumes SP: "\<forall>x \<in> S. P x"  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
811  | 
  and P: "subspace {x. P x}" shows "\<forall>x \<in> span S. P x"
 | 
| 44133 | 812  | 
using span_induct SP P by blast  | 
813  | 
||
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
814  | 
inductive_set (in real_vector) span_induct_alt_help for S:: "'a set"  | 
| 44133 | 815  | 
where  | 
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
816  | 
span_induct_alt_help_0: "0 \<in> span_induct_alt_help S"  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
817  | 
| span_induct_alt_help_S: "x \<in> S \<Longrightarrow> z \<in> span_induct_alt_help S \<Longrightarrow> (c *\<^sub>R x + z) \<in> span_induct_alt_help S"  | 
| 44133 | 818  | 
|
819  | 
lemma span_induct_alt':  | 
|
820  | 
assumes h0: "h 0" and hS: "\<And>c x y. x \<in> S \<Longrightarrow> h y \<Longrightarrow> h (c *\<^sub>R x + y)" shows "\<forall>x \<in> span S. h x"  | 
|
821  | 
proof-  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
822  | 
  {fix x:: "'a" assume x: "x \<in> span_induct_alt_help S"
 | 
| 44133 | 823  | 
have "h x"  | 
824  | 
apply (rule span_induct_alt_help.induct[OF x])  | 
|
825  | 
apply (rule h0)  | 
|
826  | 
apply (rule hS, assumption, assumption)  | 
|
827  | 
done}  | 
|
828  | 
note th0 = this  | 
|
829  | 
  {fix x assume x: "x \<in> span S"
 | 
|
830  | 
||
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
831  | 
have "x \<in> span_induct_alt_help S"  | 
| 44133 | 832  | 
proof(rule span_induct[where x=x and S=S])  | 
833  | 
show "x \<in> span S" using x .  | 
|
834  | 
next  | 
|
835  | 
fix x assume xS : "x \<in> S"  | 
|
836  | 
from span_induct_alt_help_S[OF xS span_induct_alt_help_0, of 1]  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
837  | 
show "x \<in> span_induct_alt_help S" by simp  | 
| 44133 | 838  | 
next  | 
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
839  | 
have "0 \<in> span_induct_alt_help S" by (rule span_induct_alt_help_0)  | 
| 44133 | 840  | 
moreover  | 
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
841  | 
        {fix x y assume h: "x \<in> span_induct_alt_help S" "y \<in> span_induct_alt_help S"
 | 
| 44133 | 842  | 
from h  | 
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
843  | 
have "(x + y) \<in> span_induct_alt_help S"  | 
| 44133 | 844  | 
apply (induct rule: span_induct_alt_help.induct)  | 
845  | 
apply simp  | 
|
846  | 
unfolding add_assoc  | 
|
847  | 
apply (rule span_induct_alt_help_S)  | 
|
848  | 
apply assumption  | 
|
849  | 
apply simp  | 
|
850  | 
done}  | 
|
851  | 
moreover  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
852  | 
        {fix c x assume xt: "x \<in> span_induct_alt_help S"
 | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
853  | 
then have "(c *\<^sub>R x) \<in> span_induct_alt_help S"  | 
| 44133 | 854  | 
apply (induct rule: span_induct_alt_help.induct)  | 
855  | 
apply (simp add: span_induct_alt_help_0)  | 
|
856  | 
apply (simp add: scaleR_right_distrib)  | 
|
857  | 
apply (rule span_induct_alt_help_S)  | 
|
858  | 
apply assumption  | 
|
859  | 
apply simp  | 
|
860  | 
done  | 
|
861  | 
}  | 
|
862  | 
ultimately show "subspace (span_induct_alt_help S)"  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
863  | 
unfolding subspace_def Ball_def by blast  | 
| 44133 | 864  | 
qed}  | 
865  | 
with th0 show ?thesis by blast  | 
|
866  | 
qed  | 
|
867  | 
||
868  | 
lemma span_induct_alt:  | 
|
869  | 
assumes h0: "h 0" and hS: "\<And>c x y. x \<in> S \<Longrightarrow> h y \<Longrightarrow> h (c *\<^sub>R x + y)" and x: "x \<in> span S"  | 
|
870  | 
shows "h x"  | 
|
871  | 
using span_induct_alt'[of h S] h0 hS x by blast  | 
|
872  | 
||
873  | 
text {* Individual closure properties. *}
 | 
|
874  | 
||
875  | 
lemma span_span: "span (span A) = span A"  | 
|
876  | 
unfolding span_def hull_hull ..  | 
|
877  | 
||
878  | 
lemma (in real_vector) span_superset: "x \<in> S ==> x \<in> span S" by (metis span_clauses(1))  | 
|
879  | 
||
880  | 
lemma (in real_vector) span_0: "0 \<in> span S" by (metis subspace_span subspace_0)  | 
|
881  | 
||
882  | 
lemma span_inc: "S \<subseteq> span S"  | 
|
883  | 
by (metis subset_eq span_superset)  | 
|
884  | 
||
885  | 
lemma (in real_vector) dependent_0: assumes "0\<in>A" shows "dependent A"  | 
|
886  | 
unfolding dependent_def apply(rule_tac x=0 in bexI)  | 
|
887  | 
using assms span_0 by auto  | 
|
888  | 
||
889  | 
lemma (in real_vector) span_add: "x \<in> span S \<Longrightarrow> y \<in> span S ==> x + y \<in> span S"  | 
|
890  | 
by (metis subspace_add subspace_span)  | 
|
891  | 
||
892  | 
lemma (in real_vector) span_mul: "x \<in> span S ==> (c *\<^sub>R x) \<in> span S"  | 
|
893  | 
by (metis subspace_span subspace_mul)  | 
|
894  | 
||
895  | 
lemma span_neg: "x \<in> span S ==> - x \<in> span S"  | 
|
896  | 
by (metis subspace_neg subspace_span)  | 
|
897  | 
||
898  | 
lemma span_sub: "x \<in> span S \<Longrightarrow> y \<in> span S ==> x - y \<in> span S"  | 
|
899  | 
by (metis subspace_span subspace_sub)  | 
|
900  | 
||
901  | 
lemma (in real_vector) span_setsum: "finite A \<Longrightarrow> \<forall>x \<in> A. f x \<in> span S ==> setsum f A \<in> span S"  | 
|
902  | 
by (rule subspace_setsum, rule subspace_span)  | 
|
903  | 
||
904  | 
lemma span_add_eq: "x \<in> span S \<Longrightarrow> x + y \<in> span S \<longleftrightarrow> y \<in> span S"  | 
|
905  | 
apply (auto simp only: span_add span_sub)  | 
|
906  | 
apply (subgoal_tac "(x + y) - x \<in> span S", simp)  | 
|
907  | 
by (simp only: span_add span_sub)  | 
|
908  | 
||
909  | 
text {* Mapping under linear image. *}
 | 
|
910  | 
||
911  | 
lemma span_linear_image: assumes lf: "linear f"  | 
|
912  | 
shows "span (f ` S) = f ` (span S)"  | 
|
913  | 
proof-  | 
|
914  | 
  {fix x
 | 
|
915  | 
assume x: "x \<in> span (f ` S)"  | 
|
916  | 
have "x \<in> f ` span S"  | 
|
917  | 
apply (rule span_induct[where x=x and S = "f ` S"])  | 
|
918  | 
apply (clarsimp simp add: image_iff)  | 
|
919  | 
apply (frule span_superset)  | 
|
920  | 
apply blast  | 
|
921  | 
apply (rule subspace_linear_image[OF lf])  | 
|
922  | 
apply (rule subspace_span)  | 
|
923  | 
apply (rule x)  | 
|
924  | 
done}  | 
|
925  | 
moreover  | 
|
926  | 
  {fix x assume x: "x \<in> span S"
 | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
927  | 
    have "x \<in> {x. f x \<in> span (f ` S)}"
 | 
| 44133 | 928  | 
apply (rule span_induct[where S=S])  | 
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
929  | 
apply simp  | 
| 44133 | 930  | 
apply (rule span_superset)  | 
931  | 
apply simp  | 
|
932  | 
apply (rule subspace_linear_preimage[OF lf subspace_span, of "f ` S"])  | 
|
933  | 
apply (rule x)  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
934  | 
done  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
935  | 
hence "f x \<in> span (f ` S)" by simp  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
936  | 
}  | 
| 44133 | 937  | 
ultimately show ?thesis by blast  | 
938  | 
qed  | 
|
939  | 
||
940  | 
text {* The key breakdown property. *}
 | 
|
941  | 
||
942  | 
lemma span_breakdown:  | 
|
943  | 
assumes bS: "b \<in> S" and aS: "a \<in> span S"  | 
|
944  | 
  shows "\<exists>k. a - k *\<^sub>R b \<in> span (S - {b})" (is "?P a")
 | 
|
945  | 
proof-  | 
|
946  | 
  {fix x assume xS: "x \<in> S"
 | 
|
947  | 
    {assume ab: "x = b"
 | 
|
948  | 
then have "?P x"  | 
|
949  | 
apply simp  | 
|
950  | 
apply (rule exI[where x="1"], simp)  | 
|
951  | 
by (rule span_0)}  | 
|
952  | 
moreover  | 
|
953  | 
    {assume ab: "x \<noteq> b"
 | 
|
954  | 
then have "?P x" using xS  | 
|
955  | 
apply -  | 
|
956  | 
apply (rule exI[where x=0])  | 
|
957  | 
apply (rule span_superset)  | 
|
958  | 
by simp}  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
959  | 
ultimately have "x \<in> Collect ?P" by blast}  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
960  | 
moreover have "subspace (Collect ?P)"  | 
| 44133 | 961  | 
unfolding subspace_def  | 
962  | 
apply auto  | 
|
963  | 
apply (rule exI[where x=0])  | 
|
964  | 
    using span_0[of "S - {b}"]
 | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
965  | 
apply simp  | 
| 44133 | 966  | 
apply (rule_tac x="k + ka" in exI)  | 
967  | 
apply (subgoal_tac "x + y - (k + ka) *\<^sub>R b = (x - k*\<^sub>R b) + (y - ka *\<^sub>R b)")  | 
|
968  | 
apply (simp only: )  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
969  | 
apply (rule span_add)  | 
| 44133 | 970  | 
apply assumption+  | 
971  | 
apply (simp add: algebra_simps)  | 
|
972  | 
apply (rule_tac x= "c*k" in exI)  | 
|
973  | 
apply (subgoal_tac "c *\<^sub>R x - (c * k) *\<^sub>R b = c*\<^sub>R (x - k*\<^sub>R b)")  | 
|
974  | 
apply (simp only: )  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
975  | 
apply (rule span_mul)  | 
| 44133 | 976  | 
apply assumption  | 
977  | 
by (simp add: algebra_simps)  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
978  | 
ultimately have "a \<in> Collect ?P" using aS by (rule span_induct)  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
979  | 
thus "?P a" by simp  | 
| 44133 | 980  | 
qed  | 
981  | 
||
982  | 
lemma span_breakdown_eq:  | 
|
983  | 
"x \<in> span (insert a S) \<longleftrightarrow> (\<exists>k. (x - k *\<^sub>R a) \<in> span S)" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
984  | 
proof-  | 
|
985  | 
  {assume x: "x \<in> span (insert a S)"
 | 
|
986  | 
from x span_breakdown[of "a" "insert a S" "x"]  | 
|
987  | 
have ?rhs apply clarsimp  | 
|
988  | 
apply (rule_tac x= "k" in exI)  | 
|
989  | 
      apply (rule set_rev_mp[of _ "span (S - {a})" _])
 | 
|
990  | 
apply assumption  | 
|
991  | 
apply (rule span_mono)  | 
|
992  | 
apply blast  | 
|
993  | 
done}  | 
|
994  | 
moreover  | 
|
995  | 
  { fix k assume k: "x - k *\<^sub>R a \<in> span S"
 | 
|
996  | 
have eq: "x = (x - k *\<^sub>R a) + k *\<^sub>R a" by simp  | 
|
997  | 
have "(x - k *\<^sub>R a) + k *\<^sub>R a \<in> span (insert a S)"  | 
|
998  | 
apply (rule span_add)  | 
|
999  | 
apply (rule set_rev_mp[of _ "span S" _])  | 
|
1000  | 
apply (rule k)  | 
|
1001  | 
apply (rule span_mono)  | 
|
1002  | 
apply blast  | 
|
1003  | 
apply (rule span_mul)  | 
|
1004  | 
apply (rule span_superset)  | 
|
1005  | 
apply blast  | 
|
1006  | 
done  | 
|
1007  | 
then have ?lhs using eq by metis}  | 
|
1008  | 
ultimately show ?thesis by blast  | 
|
1009  | 
qed  | 
|
1010  | 
||
1011  | 
text {* Hence some "reversal" results. *}
 | 
|
1012  | 
||
1013  | 
lemma in_span_insert:  | 
|
1014  | 
assumes a: "a \<in> span (insert b S)" and na: "a \<notin> span S"  | 
|
1015  | 
shows "b \<in> span (insert a S)"  | 
|
1016  | 
proof-  | 
|
1017  | 
from span_breakdown[of b "insert b S" a, OF insertI1 a]  | 
|
1018  | 
  obtain k where k: "a - k*\<^sub>R b \<in> span (S - {b})" by auto
 | 
|
1019  | 
  {assume k0: "k = 0"
 | 
|
1020  | 
with k have "a \<in> span S"  | 
|
1021  | 
apply (simp)  | 
|
1022  | 
apply (rule set_rev_mp)  | 
|
1023  | 
apply assumption  | 
|
1024  | 
apply (rule span_mono)  | 
|
1025  | 
apply blast  | 
|
1026  | 
done  | 
|
1027  | 
with na have ?thesis by blast}  | 
|
1028  | 
moreover  | 
|
1029  | 
  {assume k0: "k \<noteq> 0"
 | 
|
1030  | 
have eq: "b = (1/k) *\<^sub>R a - ((1/k) *\<^sub>R a - b)" by simp  | 
|
1031  | 
from k0 have eq': "(1/k) *\<^sub>R (a - k*\<^sub>R b) = (1/k) *\<^sub>R a - b"  | 
|
1032  | 
by (simp add: algebra_simps)  | 
|
1033  | 
    from k have "(1/k) *\<^sub>R (a - k*\<^sub>R b) \<in> span (S - {b})"
 | 
|
1034  | 
by (rule span_mul)  | 
|
1035  | 
    hence th: "(1/k) *\<^sub>R a - b \<in> span (S - {b})"
 | 
|
1036  | 
unfolding eq' .  | 
|
1037  | 
||
1038  | 
from k  | 
|
1039  | 
have ?thesis  | 
|
1040  | 
apply (subst eq)  | 
|
1041  | 
apply (rule span_sub)  | 
|
1042  | 
apply (rule span_mul)  | 
|
1043  | 
apply (rule span_superset)  | 
|
1044  | 
apply blast  | 
|
1045  | 
apply (rule set_rev_mp)  | 
|
1046  | 
apply (rule th)  | 
|
1047  | 
apply (rule span_mono)  | 
|
1048  | 
using na by blast}  | 
|
1049  | 
ultimately show ?thesis by blast  | 
|
1050  | 
qed  | 
|
1051  | 
||
1052  | 
lemma in_span_delete:  | 
|
1053  | 
assumes a: "a \<in> span S"  | 
|
1054  | 
  and na: "a \<notin> span (S-{b})"
 | 
|
1055  | 
  shows "b \<in> span (insert a (S - {b}))"
 | 
|
1056  | 
apply (rule in_span_insert)  | 
|
1057  | 
apply (rule set_rev_mp)  | 
|
1058  | 
apply (rule a)  | 
|
1059  | 
apply (rule span_mono)  | 
|
1060  | 
apply blast  | 
|
1061  | 
apply (rule na)  | 
|
1062  | 
done  | 
|
1063  | 
||
1064  | 
text {* Transitivity property. *}
 | 
|
1065  | 
||
1066  | 
lemma span_trans:  | 
|
1067  | 
assumes x: "x \<in> span S" and y: "y \<in> span (insert x S)"  | 
|
1068  | 
shows "y \<in> span S"  | 
|
1069  | 
proof-  | 
|
1070  | 
from span_breakdown[of x "insert x S" y, OF insertI1 y]  | 
|
1071  | 
  obtain k where k: "y -k*\<^sub>R x \<in> span (S - {x})" by auto
 | 
|
1072  | 
have eq: "y = (y - k *\<^sub>R x) + k *\<^sub>R x" by simp  | 
|
1073  | 
show ?thesis  | 
|
1074  | 
apply (subst eq)  | 
|
1075  | 
apply (rule span_add)  | 
|
1076  | 
apply (rule set_rev_mp)  | 
|
1077  | 
apply (rule k)  | 
|
1078  | 
apply (rule span_mono)  | 
|
1079  | 
apply blast  | 
|
1080  | 
apply (rule span_mul)  | 
|
1081  | 
by (rule x)  | 
|
1082  | 
qed  | 
|
1083  | 
||
1084  | 
lemma span_insert_0[simp]: "span (insert 0 S) = span S"  | 
|
1085  | 
using span_mono[of S "insert 0 S"] by (auto intro: span_trans span_0)  | 
|
1086  | 
||
1087  | 
text {* An explicit expansion is sometimes needed. *}
 | 
|
1088  | 
||
1089  | 
lemma span_explicit:  | 
|
1090  | 
  "span P = {y. \<exists>S u. finite S \<and> S \<subseteq> P \<and> setsum (\<lambda>v. u v *\<^sub>R v) S = y}"
 | 
|
1091  | 
  (is "_ = ?E" is "_ = {y. ?h y}" is "_ = {y. \<exists>S u. ?Q S u y}")
 | 
|
1092  | 
proof-  | 
|
1093  | 
  {fix x assume x: "x \<in> ?E"
 | 
|
1094  | 
then obtain S u where fS: "finite S" and SP: "S\<subseteq>P" and u: "setsum (\<lambda>v. u v *\<^sub>R v) S = x"  | 
|
1095  | 
by blast  | 
|
1096  | 
have "x \<in> span P"  | 
|
1097  | 
unfolding u[symmetric]  | 
|
1098  | 
apply (rule span_setsum[OF fS])  | 
|
1099  | 
using span_mono[OF SP]  | 
|
1100  | 
by (auto intro: span_superset span_mul)}  | 
|
1101  | 
moreover  | 
|
1102  | 
have "\<forall>x \<in> span P. x \<in> ?E"  | 
|
1103  | 
proof(rule span_induct_alt')  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
1104  | 
show "0 \<in> Collect ?h"  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
1105  | 
unfolding mem_Collect_eq  | 
| 44133 | 1106  | 
      apply (rule exI[where x="{}"]) by simp
 | 
1107  | 
next  | 
|
1108  | 
fix c x y  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
1109  | 
assume x: "x \<in> P" and hy: "y \<in> Collect ?h"  | 
| 44133 | 1110  | 
from hy obtain S u where fS: "finite S" and SP: "S\<subseteq>P"  | 
1111  | 
and u: "setsum (\<lambda>v. u v *\<^sub>R v) S = y" by blast  | 
|
1112  | 
let ?S = "insert x S"  | 
|
1113  | 
let ?u = "\<lambda>y. if y = x then (if x \<in> S then u y + c else c)  | 
|
1114  | 
else u y"  | 
|
1115  | 
from fS SP x have th0: "finite (insert x S)" "insert x S \<subseteq> P" by blast+  | 
|
1116  | 
    {assume xS: "x \<in> S"
 | 
|
1117  | 
      have S1: "S = (S - {x}) \<union> {x}"
 | 
|
1118  | 
        and Sss:"finite (S - {x})" "finite {x}" "(S -{x}) \<inter> {x} = {}" using xS fS by auto
 | 
|
1119  | 
      have "setsum (\<lambda>v. ?u v *\<^sub>R v) ?S =(\<Sum>v\<in>S - {x}. u v *\<^sub>R v) + (u x + c) *\<^sub>R x"
 | 
|
1120  | 
using xS  | 
|
1121  | 
by (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]]  | 
|
1122  | 
setsum_clauses(2)[OF fS] cong del: if_weak_cong)  | 
|
1123  | 
also have "\<dots> = (\<Sum>v\<in>S. u v *\<^sub>R v) + c *\<^sub>R x"  | 
|
1124  | 
apply (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]])  | 
|
1125  | 
by (simp add: algebra_simps)  | 
|
1126  | 
also have "\<dots> = c*\<^sub>R x + y"  | 
|
1127  | 
by (simp add: add_commute u)  | 
|
1128  | 
finally have "setsum (\<lambda>v. ?u v *\<^sub>R v) ?S = c*\<^sub>R x + y" .  | 
|
1129  | 
then have "?Q ?S ?u (c*\<^sub>R x + y)" using th0 by blast}  | 
|
1130  | 
moreover  | 
|
1131  | 
  {assume xS: "x \<notin> S"
 | 
|
1132  | 
have th00: "(\<Sum>v\<in>S. (if v = x then c else u v) *\<^sub>R v) = y"  | 
|
1133  | 
unfolding u[symmetric]  | 
|
1134  | 
apply (rule setsum_cong2)  | 
|
1135  | 
using xS by auto  | 
|
1136  | 
have "?Q ?S ?u (c*\<^sub>R x + y)" using fS xS th0  | 
|
1137  | 
by (simp add: th00 setsum_clauses add_commute cong del: if_weak_cong)}  | 
|
1138  | 
ultimately have "?Q ?S ?u (c*\<^sub>R x + y)"  | 
|
1139  | 
by (cases "x \<in> S", simp, simp)  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
1140  | 
then show "(c*\<^sub>R x + y) \<in> Collect ?h"  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
1141  | 
unfolding mem_Collect_eq  | 
| 44133 | 1142  | 
apply -  | 
1143  | 
apply (rule exI[where x="?S"])  | 
|
1144  | 
apply (rule exI[where x="?u"]) by metis  | 
|
1145  | 
qed  | 
|
1146  | 
ultimately show ?thesis by blast  | 
|
1147  | 
qed  | 
|
1148  | 
||
1149  | 
lemma dependent_explicit:  | 
|
1150  | 
"dependent P \<longleftrightarrow> (\<exists>S u. finite S \<and> S \<subseteq> P \<and> (\<exists>v\<in>S. u v \<noteq> 0 \<and> setsum (\<lambda>v. u v *\<^sub>R v) S = 0))" (is "?lhs = ?rhs")  | 
|
1151  | 
proof-  | 
|
1152  | 
  {assume dP: "dependent P"
 | 
|
1153  | 
then obtain a S u where aP: "a \<in> P" and fS: "finite S"  | 
|
1154  | 
      and SP: "S \<subseteq> P - {a}" and ua: "setsum (\<lambda>v. u v *\<^sub>R v) S = a"
 | 
|
1155  | 
unfolding dependent_def span_explicit by blast  | 
|
1156  | 
let ?S = "insert a S"  | 
|
1157  | 
let ?u = "\<lambda>y. if y = a then - 1 else u y"  | 
|
1158  | 
let ?v = a  | 
|
1159  | 
from aP SP have aS: "a \<notin> S" by blast  | 
|
1160  | 
from fS SP aP have th0: "finite ?S" "?S \<subseteq> P" "?v \<in> ?S" "?u ?v \<noteq> 0" by auto  | 
|
1161  | 
have s0: "setsum (\<lambda>v. ?u v *\<^sub>R v) ?S = 0"  | 
|
1162  | 
using fS aS  | 
|
1163  | 
apply (simp add: setsum_clauses field_simps)  | 
|
1164  | 
apply (subst (2) ua[symmetric])  | 
|
1165  | 
apply (rule setsum_cong2)  | 
|
1166  | 
by auto  | 
|
1167  | 
with th0 have ?rhs  | 
|
1168  | 
apply -  | 
|
1169  | 
apply (rule exI[where x= "?S"])  | 
|
1170  | 
apply (rule exI[where x= "?u"])  | 
|
1171  | 
by clarsimp}  | 
|
1172  | 
moreover  | 
|
1173  | 
  {fix S u v assume fS: "finite S"
 | 
|
1174  | 
and SP: "S \<subseteq> P" and vS: "v \<in> S" and uv: "u v \<noteq> 0"  | 
|
1175  | 
and u: "setsum (\<lambda>v. u v *\<^sub>R v) S = 0"  | 
|
1176  | 
let ?a = v  | 
|
1177  | 
    let ?S = "S - {v}"
 | 
|
1178  | 
let ?u = "\<lambda>i. (- u i) / u v"  | 
|
1179  | 
have th0: "?a \<in> P" "finite ?S" "?S \<subseteq> P" using fS SP vS by auto  | 
|
1180  | 
have "setsum (\<lambda>v. ?u v *\<^sub>R v) ?S = setsum (\<lambda>v. (- (inverse (u ?a))) *\<^sub>R (u v *\<^sub>R v)) S - ?u v *\<^sub>R v"  | 
|
1181  | 
using fS vS uv  | 
|
1182  | 
by (simp add: setsum_diff1 divide_inverse field_simps)  | 
|
1183  | 
also have "\<dots> = ?a"  | 
|
1184  | 
unfolding scaleR_right.setsum [symmetric] u  | 
|
1185  | 
using uv by simp  | 
|
1186  | 
finally have "setsum (\<lambda>v. ?u v *\<^sub>R v) ?S = ?a" .  | 
|
1187  | 
with th0 have ?lhs  | 
|
1188  | 
unfolding dependent_def span_explicit  | 
|
1189  | 
apply -  | 
|
1190  | 
apply (rule bexI[where x= "?a"])  | 
|
1191  | 
apply (simp_all del: scaleR_minus_left)  | 
|
1192  | 
apply (rule exI[where x= "?S"])  | 
|
1193  | 
by (auto simp del: scaleR_minus_left)}  | 
|
1194  | 
ultimately show ?thesis by blast  | 
|
1195  | 
qed  | 
|
1196  | 
||
1197  | 
||
1198  | 
lemma span_finite:  | 
|
1199  | 
assumes fS: "finite S"  | 
|
1200  | 
  shows "span S = {y. \<exists>u. setsum (\<lambda>v. u v *\<^sub>R v) S = y}"
 | 
|
1201  | 
(is "_ = ?rhs")  | 
|
1202  | 
proof-  | 
|
1203  | 
  {fix y assume y: "y \<in> span S"
 | 
|
1204  | 
from y obtain S' u where fS': "finite S'" and SS': "S' \<subseteq> S" and  | 
|
1205  | 
u: "setsum (\<lambda>v. u v *\<^sub>R v) S' = y" unfolding span_explicit by blast  | 
|
1206  | 
let ?u = "\<lambda>x. if x \<in> S' then u x else 0"  | 
|
1207  | 
have "setsum (\<lambda>v. ?u v *\<^sub>R v) S = setsum (\<lambda>v. u v *\<^sub>R v) S'"  | 
|
1208  | 
using SS' fS by (auto intro!: setsum_mono_zero_cong_right)  | 
|
1209  | 
hence "setsum (\<lambda>v. ?u v *\<^sub>R v) S = y" by (metis u)  | 
|
1210  | 
hence "y \<in> ?rhs" by auto}  | 
|
1211  | 
moreover  | 
|
1212  | 
  {fix y u assume u: "setsum (\<lambda>v. u v *\<^sub>R v) S = y"
 | 
|
1213  | 
then have "y \<in> span S" using fS unfolding span_explicit by auto}  | 
|
1214  | 
ultimately show ?thesis by blast  | 
|
1215  | 
qed  | 
|
1216  | 
||
1217  | 
lemma Int_Un_cancel: "(A \<union> B) \<inter> A = A" "(A \<union> B) \<inter> B = B" by auto  | 
|
1218  | 
||
1219  | 
lemma span_union: "span (A \<union> B) = (\<lambda>(a, b). a + b) ` (span A \<times> span B)"  | 
|
1220  | 
proof safe  | 
|
1221  | 
fix x assume "x \<in> span (A \<union> B)"  | 
|
1222  | 
then obtain S u where S: "finite S" "S \<subseteq> A \<union> B" and x: "x = (\<Sum>v\<in>S. u v *\<^sub>R v)"  | 
|
1223  | 
unfolding span_explicit by auto  | 
|
1224  | 
||
1225  | 
let ?Sa = "\<Sum>v\<in>S\<inter>A. u v *\<^sub>R v"  | 
|
1226  | 
let ?Sb = "(\<Sum>v\<in>S\<inter>(B - A). u v *\<^sub>R v)"  | 
|
1227  | 
show "x \<in> (\<lambda>(a, b). a + b) ` (span A \<times> span B)"  | 
|
1228  | 
proof  | 
|
1229  | 
show "x = (case (?Sa, ?Sb) of (a, b) \<Rightarrow> a + b)"  | 
|
1230  | 
unfolding x using S  | 
|
1231  | 
by (simp, subst setsum_Un_disjoint[symmetric]) (auto intro!: setsum_cong)  | 
|
1232  | 
||
1233  | 
from S have "?Sa \<in> span A" unfolding span_explicit  | 
|
1234  | 
by (auto intro!: exI[of _ "S \<inter> A"])  | 
|
1235  | 
moreover from S have "?Sb \<in> span B" unfolding span_explicit  | 
|
1236  | 
by (auto intro!: exI[of _ "S \<inter> (B - A)"])  | 
|
1237  | 
ultimately show "(?Sa, ?Sb) \<in> span A \<times> span B" by simp  | 
|
1238  | 
qed  | 
|
1239  | 
next  | 
|
1240  | 
fix a b assume "a \<in> span A" and "b \<in> span B"  | 
|
1241  | 
then obtain Sa ua Sb ub where span:  | 
|
1242  | 
"finite Sa" "Sa \<subseteq> A" "a = (\<Sum>v\<in>Sa. ua v *\<^sub>R v)"  | 
|
1243  | 
"finite Sb" "Sb \<subseteq> B" "b = (\<Sum>v\<in>Sb. ub v *\<^sub>R v)"  | 
|
1244  | 
unfolding span_explicit by auto  | 
|
1245  | 
let "?u v" = "(if v \<in> Sa then ua v else 0) + (if v \<in> Sb then ub v else 0)"  | 
|
1246  | 
from span have "finite (Sa \<union> Sb)" "Sa \<union> Sb \<subseteq> A \<union> B"  | 
|
1247  | 
and "a + b = (\<Sum>v\<in>(Sa\<union>Sb). ?u v *\<^sub>R v)"  | 
|
1248  | 
unfolding setsum_addf scaleR_left_distrib  | 
|
1249  | 
by (auto simp add: if_distrib cond_application_beta setsum_cases Int_Un_cancel)  | 
|
1250  | 
thus "a + b \<in> span (A \<union> B)"  | 
|
1251  | 
unfolding span_explicit by (auto intro!: exI[of _ ?u])  | 
|
1252  | 
qed  | 
|
1253  | 
||
1254  | 
text {* This is useful for building a basis step-by-step. *}
 | 
|
1255  | 
||
1256  | 
lemma independent_insert:  | 
|
1257  | 
"independent(insert a S) \<longleftrightarrow>  | 
|
1258  | 
(if a \<in> S then independent S  | 
|
1259  | 
else independent S \<and> a \<notin> span S)" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
1260  | 
proof-  | 
|
1261  | 
  {assume aS: "a \<in> S"
 | 
|
1262  | 
hence ?thesis using insert_absorb[OF aS] by simp}  | 
|
1263  | 
moreover  | 
|
1264  | 
  {assume aS: "a \<notin> S"
 | 
|
1265  | 
    {assume i: ?lhs
 | 
|
1266  | 
then have ?rhs using aS  | 
|
1267  | 
apply simp  | 
|
1268  | 
apply (rule conjI)  | 
|
1269  | 
apply (rule independent_mono)  | 
|
1270  | 
apply assumption  | 
|
1271  | 
apply blast  | 
|
1272  | 
by (simp add: dependent_def)}  | 
|
1273  | 
moreover  | 
|
1274  | 
    {assume i: ?rhs
 | 
|
1275  | 
have ?lhs using i aS  | 
|
1276  | 
apply simp  | 
|
1277  | 
apply (auto simp add: dependent_def)  | 
|
1278  | 
apply (case_tac "aa = a", auto)  | 
|
1279  | 
        apply (subgoal_tac "insert a S - {aa} = insert a (S - {aa})")
 | 
|
1280  | 
apply simp  | 
|
1281  | 
        apply (subgoal_tac "a \<in> span (insert aa (S - {aa}))")
 | 
|
1282  | 
        apply (subgoal_tac "insert aa (S - {aa}) = S")
 | 
|
1283  | 
apply simp  | 
|
1284  | 
apply blast  | 
|
1285  | 
apply (rule in_span_insert)  | 
|
1286  | 
apply assumption  | 
|
1287  | 
apply blast  | 
|
1288  | 
apply blast  | 
|
1289  | 
done}  | 
|
1290  | 
ultimately have ?thesis by blast}  | 
|
1291  | 
ultimately show ?thesis by blast  | 
|
1292  | 
qed  | 
|
1293  | 
||
1294  | 
text {* The degenerate case of the Exchange Lemma. *}
 | 
|
1295  | 
||
1296  | 
lemma mem_delete: "x \<in> (A - {a}) \<longleftrightarrow> x \<noteq> a \<and> x \<in> A"
 | 
|
1297  | 
by blast  | 
|
1298  | 
||
1299  | 
lemma spanning_subset_independent:  | 
|
1300  | 
assumes BA: "B \<subseteq> A" and iA: "independent A"  | 
|
1301  | 
and AsB: "A \<subseteq> span B"  | 
|
1302  | 
shows "A = B"  | 
|
1303  | 
proof  | 
|
1304  | 
from BA show "B \<subseteq> A" .  | 
|
1305  | 
next  | 
|
1306  | 
from span_mono[OF BA] span_mono[OF AsB]  | 
|
1307  | 
have sAB: "span A = span B" unfolding span_span by blast  | 
|
1308  | 
||
1309  | 
  {fix x assume x: "x \<in> A"
 | 
|
1310  | 
    from iA have th0: "x \<notin> span (A - {x})"
 | 
|
1311  | 
unfolding dependent_def using x by blast  | 
|
1312  | 
from x have xsA: "x \<in> span A" by (blast intro: span_superset)  | 
|
1313  | 
    have "A - {x} \<subseteq> A" by blast
 | 
|
1314  | 
    hence th1:"span (A - {x}) \<subseteq> span A" by (metis span_mono)
 | 
|
1315  | 
    {assume xB: "x \<notin> B"
 | 
|
1316  | 
      from xB BA have "B \<subseteq> A -{x}" by blast
 | 
|
1317  | 
      hence "span B \<subseteq> span (A - {x})" by (metis span_mono)
 | 
|
1318  | 
with th1 th0 sAB have "x \<notin> span A" by blast  | 
|
1319  | 
with x have False by (metis span_superset)}  | 
|
1320  | 
then have "x \<in> B" by blast}  | 
|
1321  | 
then show "A \<subseteq> B" by blast  | 
|
1322  | 
qed  | 
|
1323  | 
||
1324  | 
text {* The general case of the Exchange Lemma, the key to what follows. *}
 | 
|
1325  | 
||
1326  | 
lemma exchange_lemma:  | 
|
1327  | 
assumes f:"finite t" and i: "independent s"  | 
|
1328  | 
and sp:"s \<subseteq> span t"  | 
|
1329  | 
shows "\<exists>t'. (card t' = card t) \<and> finite t' \<and> s \<subseteq> t' \<and> t' \<subseteq> s \<union> t \<and> s \<subseteq> span t'"  | 
|
1330  | 
using f i sp  | 
|
1331  | 
proof(induct "card (t - s)" arbitrary: s t rule: less_induct)  | 
|
1332  | 
case less  | 
|
1333  | 
note ft = `finite t` and s = `independent s` and sp = `s \<subseteq> span t`  | 
|
1334  | 
let ?P = "\<lambda>t'. (card t' = card t) \<and> finite t' \<and> s \<subseteq> t' \<and> t' \<subseteq> s \<union> t \<and> s \<subseteq> span t'"  | 
|
1335  | 
let ?ths = "\<exists>t'. ?P t'"  | 
|
1336  | 
  {assume st: "s \<subseteq> t"
 | 
|
1337  | 
from st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t])  | 
|
1338  | 
by (auto intro: span_superset)}  | 
|
1339  | 
moreover  | 
|
1340  | 
  {assume st: "t \<subseteq> s"
 | 
|
1341  | 
||
1342  | 
from spanning_subset_independent[OF st s sp]  | 
|
1343  | 
st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t])  | 
|
1344  | 
by (auto intro: span_superset)}  | 
|
1345  | 
moreover  | 
|
1346  | 
  {assume st: "\<not> s \<subseteq> t" "\<not> t \<subseteq> s"
 | 
|
1347  | 
from st(2) obtain b where b: "b \<in> t" "b \<notin> s" by blast  | 
|
1348  | 
      from b have "t - {b} - s \<subset> t - s" by blast
 | 
|
1349  | 
      then have cardlt: "card (t - {b} - s) < card (t - s)" using ft
 | 
|
1350  | 
by (auto intro: psubset_card_mono)  | 
|
1351  | 
from b ft have ct0: "card t \<noteq> 0" by auto  | 
|
1352  | 
    {assume stb: "s \<subseteq> span(t -{b})"
 | 
|
1353  | 
      from ft have ftb: "finite (t -{b})" by auto
 | 
|
1354  | 
from less(1)[OF cardlt ftb s stb]  | 
|
1355  | 
      obtain u where u: "card u = card (t-{b})" "s \<subseteq> u" "u \<subseteq> s \<union> (t - {b})" "s \<subseteq> span u" and fu: "finite u" by blast
 | 
|
1356  | 
let ?w = "insert b u"  | 
|
1357  | 
have th0: "s \<subseteq> insert b u" using u by blast  | 
|
1358  | 
from u(3) b have "u \<subseteq> s \<union> t" by blast  | 
|
1359  | 
then have th1: "insert b u \<subseteq> s \<union> t" using u b by blast  | 
|
1360  | 
have bu: "b \<notin> u" using b u by blast  | 
|
1361  | 
from u(1) ft b have "card u = (card t - 1)" by auto  | 
|
1362  | 
then  | 
|
1363  | 
have th2: "card (insert b u) = card t"  | 
|
1364  | 
using card_insert_disjoint[OF fu bu] ct0 by auto  | 
|
1365  | 
from u(4) have "s \<subseteq> span u" .  | 
|
1366  | 
also have "\<dots> \<subseteq> span (insert b u)" apply (rule span_mono) by blast  | 
|
1367  | 
finally have th3: "s \<subseteq> span (insert b u)" .  | 
|
1368  | 
from th0 th1 th2 th3 fu have th: "?P ?w" by blast  | 
|
1369  | 
from th have ?ths by blast}  | 
|
1370  | 
moreover  | 
|
1371  | 
    {assume stb: "\<not> s \<subseteq> span(t -{b})"
 | 
|
1372  | 
      from stb obtain a where a: "a \<in> s" "a \<notin> span (t - {b})" by blast
 | 
|
1373  | 
have ab: "a \<noteq> b" using a b by blast  | 
|
1374  | 
      have at: "a \<notin> t" using a ab span_superset[of a "t- {b}"] by auto
 | 
|
1375  | 
      have mlt: "card ((insert a (t - {b})) - s) < card (t - s)"
 | 
|
1376  | 
using cardlt ft a b by auto  | 
|
1377  | 
      have ft': "finite (insert a (t - {b}))" using ft by auto
 | 
|
1378  | 
      {fix x assume xs: "x \<in> s"
 | 
|
1379  | 
        have t: "t \<subseteq> (insert b (insert a (t -{b})))" using b by auto
 | 
|
1380  | 
from b(1) have "b \<in> span t" by (simp add: span_superset)  | 
|
1381  | 
        have bs: "b \<in> span (insert a (t - {b}))" apply(rule in_span_delete)
 | 
|
1382  | 
using a sp unfolding subset_eq by auto  | 
|
1383  | 
from xs sp have "x \<in> span t" by blast  | 
|
1384  | 
with span_mono[OF t]  | 
|
1385  | 
        have x: "x \<in> span (insert b (insert a (t - {b})))" ..
 | 
|
1386  | 
        from span_trans[OF bs x] have "x \<in> span (insert a (t - {b}))"  .}
 | 
|
1387  | 
      then have sp': "s \<subseteq> span (insert a (t - {b}))" by blast
 | 
|
1388  | 
||
1389  | 
from less(1)[OF mlt ft' s sp'] obtain u where  | 
|
1390  | 
        u: "card u = card (insert a (t -{b}))" "finite u" "s \<subseteq> u" "u \<subseteq> s \<union> insert a (t -{b})"
 | 
|
1391  | 
"s \<subseteq> span u" by blast  | 
|
1392  | 
from u a b ft at ct0 have "?P u" by auto  | 
|
1393  | 
then have ?ths by blast }  | 
|
1394  | 
ultimately have ?ths by blast  | 
|
1395  | 
}  | 
|
1396  | 
ultimately  | 
|
1397  | 
show ?ths by blast  | 
|
1398  | 
qed  | 
|
1399  | 
||
1400  | 
text {* This implies corresponding size bounds. *}
 | 
|
1401  | 
||
1402  | 
lemma independent_span_bound:  | 
|
1403  | 
assumes f: "finite t" and i: "independent s" and sp:"s \<subseteq> span t"  | 
|
1404  | 
shows "finite s \<and> card s \<le> card t"  | 
|
1405  | 
by (metis exchange_lemma[OF f i sp] finite_subset card_mono)  | 
|
1406  | 
||
1407  | 
||
1408  | 
lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x\<in> (UNIV::'a::finite set)}"
 | 
|
1409  | 
proof-  | 
|
1410  | 
  have eq: "{f x |x. x\<in> UNIV} = f ` UNIV" by auto
 | 
|
1411  | 
show ?thesis unfolding eq  | 
|
1412  | 
apply (rule finite_imageI)  | 
|
1413  | 
apply (rule finite)  | 
|
1414  | 
done  | 
|
1415  | 
qed  | 
|
1416  | 
||
1417  | 
subsection{* Euclidean Spaces as Typeclass*}
 | 
|
1418  | 
||
1419  | 
lemma independent_eq_inj_on:  | 
|
1420  | 
  fixes D :: nat and f :: "nat \<Rightarrow> 'c::real_vector" assumes *: "inj_on f {..<D}"
 | 
|
1421  | 
  shows "independent (f ` {..<D}) \<longleftrightarrow> (\<forall>a u. a < D \<longrightarrow> (\<Sum>i\<in>{..<D}-{a}. u (f i) *\<^sub>R f i) \<noteq> f a)"
 | 
|
1422  | 
proof -  | 
|
1423  | 
  from * have eq: "\<And>i. i < D \<Longrightarrow> f ` {..<D} - {f i} = f`({..<D} - {i})"
 | 
|
1424  | 
    and inj: "\<And>i. inj_on f ({..<D} - {i})"
 | 
|
1425  | 
by (auto simp: inj_on_def)  | 
|
1426  | 
  have *: "\<And>i. finite (f ` {..<D} - {i})" by simp
 | 
|
1427  | 
show ?thesis unfolding dependent_def span_finite[OF *]  | 
|
1428  | 
by (auto simp: eq setsum_reindex[OF inj])  | 
|
1429  | 
qed  | 
|
1430  | 
||
1431  | 
lemma independent_basis:  | 
|
1432  | 
  "independent (basis ` {..<DIM('a)} :: 'a::euclidean_space set)"
 | 
|
1433  | 
unfolding independent_eq_inj_on [OF basis_inj]  | 
|
1434  | 
apply clarify  | 
|
1435  | 
apply (drule_tac f="inner (basis a)" in arg_cong)  | 
|
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
44176 
diff
changeset
 | 
1436  | 
apply (simp add: inner_setsum_right dot_basis)  | 
| 44133 | 1437  | 
done  | 
1438  | 
||
1439  | 
lemma dimensionI:  | 
|
1440  | 
  assumes "\<And>d. \<lbrakk> 0 < d; basis ` {d..} = {0::'a::euclidean_space};
 | 
|
1441  | 
    independent (basis ` {..<d} :: 'a set);
 | 
|
1442  | 
    inj_on (basis :: nat \<Rightarrow> 'a) {..<d} \<rbrakk> \<Longrightarrow> P d"
 | 
|
1443  | 
  shows "P DIM('a::euclidean_space)"
 | 
|
1444  | 
using DIM_positive basis_finite independent_basis basis_inj  | 
|
1445  | 
by (rule assms)  | 
|
1446  | 
||
1447  | 
lemma (in euclidean_space) dimension_eq:  | 
|
1448  | 
assumes "\<And>i. i < d \<Longrightarrow> basis i \<noteq> 0"  | 
|
1449  | 
assumes "\<And>i. d \<le> i \<Longrightarrow> basis i = 0"  | 
|
1450  | 
  shows "DIM('a) = d"
 | 
|
1451  | 
proof (rule linorder_cases [of "DIM('a)" d])
 | 
|
1452  | 
  assume "DIM('a) < d"
 | 
|
1453  | 
  hence "basis DIM('a) \<noteq> 0" by (rule assms)
 | 
|
1454  | 
thus ?thesis by simp  | 
|
1455  | 
next  | 
|
1456  | 
  assume "d < DIM('a)"
 | 
|
1457  | 
hence "basis d \<noteq> 0" by simp  | 
|
1458  | 
thus ?thesis by (simp add: assms)  | 
|
1459  | 
next  | 
|
1460  | 
  assume "DIM('a) = d" thus ?thesis .
 | 
|
1461  | 
qed  | 
|
1462  | 
||
1463  | 
lemma (in euclidean_space) range_basis:  | 
|
1464  | 
    "range basis = insert 0 (basis ` {..<DIM('a)})"
 | 
|
1465  | 
proof -  | 
|
1466  | 
  have *: "UNIV = {..<DIM('a)} \<union> {DIM('a)..}" by auto
 | 
|
1467  | 
show ?thesis unfolding * image_Un basis_finite by auto  | 
|
1468  | 
qed  | 
|
1469  | 
||
1470  | 
lemma (in euclidean_space) range_basis_finite[intro]:  | 
|
1471  | 
"finite (range basis)"  | 
|
1472  | 
unfolding range_basis by auto  | 
|
1473  | 
||
1474  | 
lemma span_basis: "span (range basis) = (UNIV :: 'a::euclidean_space set)"  | 
|
1475  | 
proof -  | 
|
1476  | 
  { fix x :: 'a
 | 
|
1477  | 
    have "(\<Sum>i<DIM('a). (x $$ i) *\<^sub>R basis i) \<in> span (range basis :: 'a set)"
 | 
|
1478  | 
by (simp add: span_setsum span_mul span_superset)  | 
|
1479  | 
hence "x \<in> span (range basis)"  | 
|
1480  | 
by (simp only: euclidean_representation [symmetric])  | 
|
1481  | 
} thus ?thesis by auto  | 
|
1482  | 
qed  | 
|
1483  | 
||
1484  | 
lemma basis_representation:  | 
|
1485  | 
  "\<exists>u. x = (\<Sum>v\<in>basis ` {..<DIM('a)}. u v *\<^sub>R (v\<Colon>'a\<Colon>euclidean_space))"
 | 
|
1486  | 
proof -  | 
|
1487  | 
have "x\<in>UNIV" by auto from this[unfolded span_basis[THEN sym]]  | 
|
1488  | 
  have "\<exists>u. (\<Sum>v\<in>basis ` {..<DIM('a)}. u v *\<^sub>R v) = x"
 | 
|
1489  | 
unfolding range_basis span_insert_0 apply(subst (asm) span_finite) by auto  | 
|
1490  | 
thus ?thesis by fastsimp  | 
|
1491  | 
qed  | 
|
1492  | 
||
1493  | 
lemma span_basis'[simp]:"span ((basis::nat=>'a) ` {..<DIM('a::euclidean_space)}) = UNIV"
 | 
|
1494  | 
apply(subst span_basis[symmetric]) unfolding range_basis by auto  | 
|
1495  | 
||
1496  | 
lemma card_basis[simp]:"card ((basis::nat=>'a) ` {..<DIM('a::euclidean_space)}) = DIM('a)"
 | 
|
1497  | 
apply(subst card_image) using basis_inj by auto  | 
|
1498  | 
||
1499  | 
lemma in_span_basis: "(x::'a::euclidean_space) \<in> span (basis ` {..<DIM('a)})"
 | 
|
1500  | 
unfolding span_basis' ..  | 
|
1501  | 
||
1502  | 
lemma norm_bound_component_le: "norm (x::'a::euclidean_space) \<le> e \<Longrightarrow> \<bar>x$$i\<bar> <= e"  | 
|
1503  | 
by (metis component_le_norm order_trans)  | 
|
1504  | 
||
1505  | 
lemma norm_bound_component_lt: "norm (x::'a::euclidean_space) < e \<Longrightarrow> \<bar>x$$i\<bar> < e"  | 
|
1506  | 
by (metis component_le_norm basic_trans_rules(21))  | 
|
1507  | 
||
1508  | 
lemma norm_le_l1: "norm (x::'a::euclidean_space) \<le> (\<Sum>i<DIM('a). \<bar>x $$ i\<bar>)"
 | 
|
1509  | 
apply (subst euclidean_representation[of x])  | 
|
| 
44176
 
eda112e9cdee
remove redundant lemma setsum_norm in favor of norm_setsum;
 
huffman 
parents: 
44170 
diff
changeset
 | 
1510  | 
apply (rule order_trans[OF norm_setsum])  | 
| 44133 | 1511  | 
by (auto intro!: setsum_mono)  | 
1512  | 
||
1513  | 
lemma setsum_norm_allsubsets_bound:  | 
|
1514  | 
fixes f:: "'a \<Rightarrow> 'n::euclidean_space"  | 
|
1515  | 
assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (setsum f Q) \<le> e"  | 
|
1516  | 
  shows "setsum (\<lambda>x. norm (f x)) P \<le> 2 * real DIM('n) *  e"
 | 
|
1517  | 
proof-  | 
|
1518  | 
  let ?d = "real DIM('n)"
 | 
|
1519  | 
let ?nf = "\<lambda>x. norm (f x)"  | 
|
1520  | 
  let ?U = "{..<DIM('n)}"
 | 
|
1521  | 
have th0: "setsum (\<lambda>x. setsum (\<lambda>i. \<bar>f x $$ i\<bar>) ?U) P = setsum (\<lambda>i. setsum (\<lambda>x. \<bar>f x $$ i\<bar>) P) ?U"  | 
|
1522  | 
by (rule setsum_commute)  | 
|
1523  | 
have th1: "2 * ?d * e = of_nat (card ?U) * (2 * e)" by (simp add: real_of_nat_def)  | 
|
1524  | 
have "setsum ?nf P \<le> setsum (\<lambda>x. setsum (\<lambda>i. \<bar>f x $$ i\<bar>) ?U) P"  | 
|
1525  | 
apply (rule setsum_mono) by (rule norm_le_l1)  | 
|
1526  | 
also have "\<dots> \<le> 2 * ?d * e"  | 
|
1527  | 
unfolding th0 th1  | 
|
1528  | 
proof(rule setsum_bounded)  | 
|
1529  | 
fix i assume i: "i \<in> ?U"  | 
|
1530  | 
    let ?Pp = "{x. x\<in> P \<and> f x $$ i \<ge> 0}"
 | 
|
1531  | 
    let ?Pn = "{x. x \<in> P \<and> f x $$ i < 0}"
 | 
|
1532  | 
have thp: "P = ?Pp \<union> ?Pn" by auto  | 
|
1533  | 
    have thp0: "?Pp \<inter> ?Pn ={}" by auto
 | 
|
1534  | 
have PpP: "?Pp \<subseteq> P" and PnP: "?Pn \<subseteq> P" by blast+  | 
|
1535  | 
have Ppe:"setsum (\<lambda>x. \<bar>f x $$ i\<bar>) ?Pp \<le> e"  | 
|
1536  | 
using component_le_norm[of "setsum (\<lambda>x. f x) ?Pp" i] fPs[OF PpP]  | 
|
| 
44457
 
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
 
huffman 
parents: 
44454 
diff
changeset
 | 
1537  | 
by(auto intro: abs_le_D1)  | 
| 44133 | 1538  | 
have Pne: "setsum (\<lambda>x. \<bar>f x $$ i\<bar>) ?Pn \<le> e"  | 
1539  | 
using component_le_norm[of "setsum (\<lambda>x. - f x) ?Pn" i] fPs[OF PnP]  | 
|
1540  | 
by(auto simp add: setsum_negf intro: abs_le_D1)  | 
|
1541  | 
have "setsum (\<lambda>x. \<bar>f x $$ i\<bar>) P = setsum (\<lambda>x. \<bar>f x $$ i\<bar>) ?Pp + setsum (\<lambda>x. \<bar>f x $$ i\<bar>) ?Pn"  | 
|
1542  | 
apply (subst thp)  | 
|
1543  | 
apply (rule setsum_Un_zero)  | 
|
1544  | 
using fP thp0 by auto  | 
|
1545  | 
also have "\<dots> \<le> 2*e" using Pne Ppe by arith  | 
|
1546  | 
finally show "setsum (\<lambda>x. \<bar>f x $$ i\<bar>) P \<le> 2*e" .  | 
|
1547  | 
qed  | 
|
1548  | 
finally show ?thesis .  | 
|
1549  | 
qed  | 
|
1550  | 
||
1551  | 
lemma choice_iff': "(\<forall>x<d. \<exists>y. P x y) \<longleftrightarrow> (\<exists>f. \<forall>x<d. P x (f x))" by metis  | 
|
1552  | 
||
1553  | 
lemma lambda_skolem': "(\<forall>i<DIM('a::euclidean_space). \<exists>x. P i x) \<longleftrightarrow>
 | 
|
1554  | 
   (\<exists>x::'a. \<forall>i<DIM('a). P i (x$$i))" (is "?lhs \<longleftrightarrow> ?rhs")
 | 
|
1555  | 
proof-  | 
|
1556  | 
  let ?S = "{..<DIM('a)}"
 | 
|
1557  | 
  {assume H: "?rhs"
 | 
|
1558  | 
then have ?lhs by auto}  | 
|
1559  | 
moreover  | 
|
1560  | 
  {assume H: "?lhs"
 | 
|
1561  | 
    then obtain f where f:"\<forall>i<DIM('a). P i (f i)" unfolding choice_iff' by metis
 | 
|
1562  | 
let ?x = "(\<chi>\<chi> i. (f i)) :: 'a"  | 
|
1563  | 
    {fix i assume i:"i<DIM('a)"
 | 
|
1564  | 
with f have "P i (f i)" by metis  | 
|
1565  | 
then have "P i (?x$$i)" using i by auto  | 
|
1566  | 
}  | 
|
1567  | 
    hence "\<forall>i<DIM('a). P i (?x$$i)" by metis
 | 
|
1568  | 
hence ?rhs by metis }  | 
|
1569  | 
ultimately show ?thesis by metis  | 
|
1570  | 
qed  | 
|
1571  | 
||
1572  | 
subsection {* An ordering on euclidean spaces that will allow us to talk about intervals *}
 | 
|
1573  | 
||
1574  | 
class ordered_euclidean_space = ord + euclidean_space +  | 
|
1575  | 
  assumes eucl_le: "x \<le> y \<longleftrightarrow> (\<forall>i < DIM('a). x $$ i \<le> y $$ i)"
 | 
|
1576  | 
  and eucl_less: "x < y \<longleftrightarrow> (\<forall>i < DIM('a). x $$ i < y $$ i)"
 | 
|
1577  | 
||
1578  | 
lemma eucl_less_not_refl[simp, intro!]: "\<not> x < (x::'a::ordered_euclidean_space)"  | 
|
1579  | 
unfolding eucl_less[where 'a='a] by auto  | 
|
1580  | 
||
1581  | 
lemma euclidean_trans[trans]:  | 
|
1582  | 
fixes x y z :: "'a::ordered_euclidean_space"  | 
|
1583  | 
shows "x < y \<Longrightarrow> y < z \<Longrightarrow> x < z"  | 
|
1584  | 
and "x \<le> y \<Longrightarrow> y < z \<Longrightarrow> x < z"  | 
|
1585  | 
and "x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z"  | 
|
1586  | 
by (force simp: eucl_less[where 'a='a] eucl_le[where 'a='a])+  | 
|
1587  | 
||
1588  | 
subsection {* Linearity and Bilinearity continued *}
 | 
|
1589  | 
||
1590  | 
lemma linear_bounded:  | 
|
1591  | 
fixes f:: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"  | 
|
1592  | 
assumes lf: "linear f"  | 
|
1593  | 
shows "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x"  | 
|
1594  | 
proof-  | 
|
1595  | 
  let ?S = "{..<DIM('a)}"
 | 
|
1596  | 
let ?B = "setsum (\<lambda>i. norm(f(basis i))) ?S"  | 
|
1597  | 
have fS: "finite ?S" by simp  | 
|
1598  | 
  {fix x:: "'a"
 | 
|
1599  | 
let ?g = "(\<lambda> i. (x$$i) *\<^sub>R (basis i) :: 'a)"  | 
|
1600  | 
have "norm (f x) = norm (f (setsum (\<lambda>i. (x$$i) *\<^sub>R (basis i)) ?S))"  | 
|
1601  | 
apply(subst euclidean_representation[of x]) ..  | 
|
1602  | 
also have "\<dots> = norm (setsum (\<lambda> i. (x$$i) *\<^sub>R f (basis i)) ?S)"  | 
|
1603  | 
using linear_setsum[OF lf fS, of ?g, unfolded o_def] linear_cmul[OF lf] by auto  | 
|
1604  | 
finally have th0: "norm (f x) = norm (setsum (\<lambda>i. (x$$i) *\<^sub>R f (basis i))?S)" .  | 
|
1605  | 
    {fix i assume i: "i \<in> ?S"
 | 
|
1606  | 
from component_le_norm[of x i]  | 
|
1607  | 
have "norm ((x$$i) *\<^sub>R f (basis i :: 'a)) \<le> norm (f (basis i)) * norm x"  | 
|
1608  | 
unfolding norm_scaleR  | 
|
1609  | 
apply (simp only: mult_commute)  | 
|
1610  | 
apply (rule mult_mono)  | 
|
1611  | 
by (auto simp add: field_simps) }  | 
|
1612  | 
then have th: "\<forall>i\<in> ?S. norm ((x$$i) *\<^sub>R f (basis i :: 'a)) \<le> norm (f (basis i)) * norm x" by metis  | 
|
| 
44176
 
eda112e9cdee
remove redundant lemma setsum_norm in favor of norm_setsum;
 
huffman 
parents: 
44170 
diff
changeset
 | 
1613  | 
from setsum_norm_le[of _ "\<lambda>i. (x$$i) *\<^sub>R (f (basis i))", OF th]  | 
| 44133 | 1614  | 
have "norm (f x) \<le> ?B * norm x" unfolding th0 setsum_left_distrib by metis}  | 
1615  | 
then show ?thesis by blast  | 
|
1616  | 
qed  | 
|
1617  | 
||
1618  | 
lemma linear_bounded_pos:  | 
|
1619  | 
fixes f:: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"  | 
|
1620  | 
assumes lf: "linear f"  | 
|
1621  | 
shows "\<exists>B > 0. \<forall>x. norm (f x) \<le> B * norm x"  | 
|
1622  | 
proof-  | 
|
1623  | 
from linear_bounded[OF lf] obtain B where  | 
|
1624  | 
B: "\<forall>x. norm (f x) \<le> B * norm x" by blast  | 
|
1625  | 
let ?K = "\<bar>B\<bar> + 1"  | 
|
1626  | 
have Kp: "?K > 0" by arith  | 
|
1627  | 
    { assume C: "B < 0"
 | 
|
1628  | 
have "((\<chi>\<chi> i. 1)::'a) \<noteq> 0" unfolding euclidean_eq[where 'a='a]  | 
|
| 
44286
 
8766839efb1b
declare euclidean_component_zero[simp] at the point it is proved
 
huffman 
parents: 
44282 
diff
changeset
 | 
1629  | 
by(auto intro!:exI[where x=0])  | 
| 44133 | 1630  | 
hence "norm ((\<chi>\<chi> i. 1)::'a) > 0" by auto  | 
1631  | 
with C have "B * norm ((\<chi>\<chi> i. 1)::'a) < 0"  | 
|
1632  | 
by (simp add: mult_less_0_iff)  | 
|
1633  | 
with B[rule_format, of "(\<chi>\<chi> i. 1)::'a"] norm_ge_zero[of "f ((\<chi>\<chi> i. 1)::'a)"] have False by simp  | 
|
1634  | 
}  | 
|
1635  | 
then have Bp: "B \<ge> 0" by (metis not_leE)  | 
|
1636  | 
    {fix x::"'a"
 | 
|
1637  | 
have "norm (f x) \<le> ?K * norm x"  | 
|
1638  | 
using B[rule_format, of x] norm_ge_zero[of x] norm_ge_zero[of "f x"] Bp  | 
|
1639  | 
apply (auto simp add: field_simps split add: abs_split)  | 
|
1640  | 
apply (erule order_trans, simp)  | 
|
1641  | 
done  | 
|
1642  | 
}  | 
|
1643  | 
then show ?thesis using Kp by blast  | 
|
1644  | 
qed  | 
|
1645  | 
||
1646  | 
lemma linear_conv_bounded_linear:  | 
|
1647  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"  | 
|
1648  | 
shows "linear f \<longleftrightarrow> bounded_linear f"  | 
|
1649  | 
proof  | 
|
1650  | 
assume "linear f"  | 
|
1651  | 
show "bounded_linear f"  | 
|
1652  | 
proof  | 
|
1653  | 
fix x y show "f (x + y) = f x + f y"  | 
|
1654  | 
using `linear f` unfolding linear_def by simp  | 
|
1655  | 
next  | 
|
1656  | 
fix r x show "f (scaleR r x) = scaleR r (f x)"  | 
|
1657  | 
using `linear f` unfolding linear_def by simp  | 
|
1658  | 
next  | 
|
1659  | 
have "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x"  | 
|
1660  | 
using `linear f` by (rule linear_bounded)  | 
|
1661  | 
thus "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"  | 
|
1662  | 
by (simp add: mult_commute)  | 
|
1663  | 
qed  | 
|
1664  | 
next  | 
|
1665  | 
assume "bounded_linear f"  | 
|
1666  | 
then interpret f: bounded_linear f .  | 
|
1667  | 
show "linear f"  | 
|
1668  | 
by (simp add: f.add f.scaleR linear_def)  | 
|
1669  | 
qed  | 
|
1670  | 
||
1671  | 
lemma bounded_linearI': fixes f::"'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"  | 
|
1672  | 
assumes "\<And>x y. f (x + y) = f x + f y" "\<And>c x. f (c *\<^sub>R x) = c *\<^sub>R f x"  | 
|
1673  | 
shows "bounded_linear f" unfolding linear_conv_bounded_linear[THEN sym]  | 
|
1674  | 
by(rule linearI[OF assms])  | 
|
1675  | 
||
1676  | 
||
1677  | 
lemma bilinear_bounded:  | 
|
1678  | 
fixes h:: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'k::real_normed_vector"  | 
|
1679  | 
assumes bh: "bilinear h"  | 
|
1680  | 
shows "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"  | 
|
1681  | 
proof-  | 
|
1682  | 
  let ?M = "{..<DIM('m)}"
 | 
|
1683  | 
  let ?N = "{..<DIM('n)}"
 | 
|
1684  | 
let ?B = "setsum (\<lambda>(i,j). norm (h (basis i) (basis j))) (?M \<times> ?N)"  | 
|
1685  | 
have fM: "finite ?M" and fN: "finite ?N" by simp_all  | 
|
1686  | 
  {fix x:: "'m" and  y :: "'n"
 | 
|
1687  | 
have "norm (h x y) = norm (h (setsum (\<lambda>i. (x$$i) *\<^sub>R basis i) ?M) (setsum (\<lambda>i. (y$$i) *\<^sub>R basis i) ?N))"  | 
|
1688  | 
apply(subst euclidean_representation[where 'a='m])  | 
|
1689  | 
apply(subst euclidean_representation[where 'a='n]) ..  | 
|
1690  | 
also have "\<dots> = norm (setsum (\<lambda> (i,j). h ((x$$i) *\<^sub>R basis i) ((y$$j) *\<^sub>R basis j)) (?M \<times> ?N))"  | 
|
1691  | 
unfolding bilinear_setsum[OF bh fM fN] ..  | 
|
1692  | 
finally have th: "norm (h x y) = \<dots>" .  | 
|
1693  | 
have "norm (h x y) \<le> ?B * norm x * norm y"  | 
|
1694  | 
apply (simp add: setsum_left_distrib th)  | 
|
1695  | 
apply (rule setsum_norm_le)  | 
|
1696  | 
using fN fM  | 
|
1697  | 
apply simp  | 
|
1698  | 
apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] field_simps simp del: scaleR_scaleR)  | 
|
1699  | 
apply (rule mult_mono)  | 
|
1700  | 
apply (auto simp add: zero_le_mult_iff component_le_norm)  | 
|
1701  | 
apply (rule mult_mono)  | 
|
1702  | 
apply (auto simp add: zero_le_mult_iff component_le_norm)  | 
|
1703  | 
done}  | 
|
1704  | 
then show ?thesis by metis  | 
|
1705  | 
qed  | 
|
1706  | 
||
1707  | 
lemma bilinear_bounded_pos:  | 
|
1708  | 
fixes h:: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::real_normed_vector"  | 
|
1709  | 
assumes bh: "bilinear h"  | 
|
1710  | 
shows "\<exists>B > 0. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"  | 
|
1711  | 
proof-  | 
|
1712  | 
from bilinear_bounded[OF bh] obtain B where  | 
|
1713  | 
B: "\<forall>x y. norm (h x y) \<le> B * norm x * norm y" by blast  | 
|
1714  | 
let ?K = "\<bar>B\<bar> + 1"  | 
|
1715  | 
have Kp: "?K > 0" by arith  | 
|
1716  | 
have KB: "B < ?K" by arith  | 
|
1717  | 
  {fix x::'a and y::'b
 | 
|
1718  | 
from KB Kp  | 
|
1719  | 
have "B * norm x * norm y \<le> ?K * norm x * norm y"  | 
|
1720  | 
apply -  | 
|
1721  | 
apply (rule mult_right_mono, rule mult_right_mono)  | 
|
1722  | 
by auto  | 
|
1723  | 
then have "norm (h x y) \<le> ?K * norm x * norm y"  | 
|
1724  | 
using B[rule_format, of x y] by simp}  | 
|
1725  | 
with Kp show ?thesis by blast  | 
|
1726  | 
qed  | 
|
1727  | 
||
1728  | 
lemma bilinear_conv_bounded_bilinear:  | 
|
1729  | 
fixes h :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::real_normed_vector"  | 
|
1730  | 
shows "bilinear h \<longleftrightarrow> bounded_bilinear h"  | 
|
1731  | 
proof  | 
|
1732  | 
assume "bilinear h"  | 
|
1733  | 
show "bounded_bilinear h"  | 
|
1734  | 
proof  | 
|
1735  | 
fix x y z show "h (x + y) z = h x z + h y z"  | 
|
1736  | 
using `bilinear h` unfolding bilinear_def linear_def by simp  | 
|
1737  | 
next  | 
|
1738  | 
fix x y z show "h x (y + z) = h x y + h x z"  | 
|
1739  | 
using `bilinear h` unfolding bilinear_def linear_def by simp  | 
|
1740  | 
next  | 
|
1741  | 
fix r x y show "h (scaleR r x) y = scaleR r (h x y)"  | 
|
1742  | 
using `bilinear h` unfolding bilinear_def linear_def  | 
|
1743  | 
by simp  | 
|
1744  | 
next  | 
|
1745  | 
fix r x y show "h x (scaleR r y) = scaleR r (h x y)"  | 
|
1746  | 
using `bilinear h` unfolding bilinear_def linear_def  | 
|
1747  | 
by simp  | 
|
1748  | 
next  | 
|
1749  | 
have "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"  | 
|
1750  | 
using `bilinear h` by (rule bilinear_bounded)  | 
|
1751  | 
thus "\<exists>K. \<forall>x y. norm (h x y) \<le> norm x * norm y * K"  | 
|
1752  | 
by (simp add: mult_ac)  | 
|
1753  | 
qed  | 
|
1754  | 
next  | 
|
1755  | 
assume "bounded_bilinear h"  | 
|
1756  | 
then interpret h: bounded_bilinear h .  | 
|
1757  | 
show "bilinear h"  | 
|
1758  | 
unfolding bilinear_def linear_conv_bounded_linear  | 
|
1759  | 
using h.bounded_linear_left h.bounded_linear_right  | 
|
1760  | 
by simp  | 
|
1761  | 
qed  | 
|
1762  | 
||
1763  | 
subsection {* We continue. *}
 | 
|
1764  | 
||
1765  | 
lemma independent_bound:  | 
|
1766  | 
  fixes S:: "('a::euclidean_space) set"
 | 
|
1767  | 
  shows "independent S \<Longrightarrow> finite S \<and> card S <= DIM('a::euclidean_space)"
 | 
|
1768  | 
  using independent_span_bound[of "(basis::nat=>'a) ` {..<DIM('a)}" S] by auto
 | 
|
1769  | 
||
1770  | 
lemma dependent_biggerset: "(finite (S::('a::euclidean_space) set) ==> card S > DIM('a)) ==> dependent S"
 | 
|
1771  | 
by (metis independent_bound not_less)  | 
|
1772  | 
||
1773  | 
text {* Hence we can create a maximal independent subset. *}
 | 
|
1774  | 
||
1775  | 
lemma maximal_independent_subset_extend:  | 
|
1776  | 
  assumes sv: "(S::('a::euclidean_space) set) \<subseteq> V" and iS: "independent S"
 | 
|
1777  | 
shows "\<exists>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B"  | 
|
1778  | 
using sv iS  | 
|
1779  | 
proof(induct "DIM('a) - card S" arbitrary: S rule: less_induct)
 | 
|
1780  | 
case less  | 
|
1781  | 
note sv = `S \<subseteq> V` and i = `independent S`  | 
|
1782  | 
let ?P = "\<lambda>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B"  | 
|
1783  | 
let ?ths = "\<exists>x. ?P x"  | 
|
1784  | 
  let ?d = "DIM('a)"
 | 
|
1785  | 
  {assume "V \<subseteq> span S"
 | 
|
1786  | 
then have ?ths using sv i by blast }  | 
|
1787  | 
moreover  | 
|
1788  | 
  {assume VS: "\<not> V \<subseteq> span S"
 | 
|
1789  | 
from VS obtain a where a: "a \<in> V" "a \<notin> span S" by blast  | 
|
1790  | 
from a have aS: "a \<notin> S" by (auto simp add: span_superset)  | 
|
1791  | 
have th0: "insert a S \<subseteq> V" using a sv by blast  | 
|
1792  | 
from independent_insert[of a S] i a  | 
|
1793  | 
have th1: "independent (insert a S)" by auto  | 
|
1794  | 
have mlt: "?d - card (insert a S) < ?d - card S"  | 
|
1795  | 
using aS a independent_bound[OF th1]  | 
|
1796  | 
by auto  | 
|
1797  | 
||
1798  | 
from less(1)[OF mlt th0 th1]  | 
|
1799  | 
obtain B where B: "insert a S \<subseteq> B" "B \<subseteq> V" "independent B" " V \<subseteq> span B"  | 
|
1800  | 
by blast  | 
|
1801  | 
from B have "?P B" by auto  | 
|
1802  | 
then have ?ths by blast}  | 
|
1803  | 
ultimately show ?ths by blast  | 
|
1804  | 
qed  | 
|
1805  | 
||
1806  | 
lemma maximal_independent_subset:  | 
|
1807  | 
  "\<exists>(B:: ('a::euclidean_space) set). B\<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
 | 
|
1808  | 
  by (metis maximal_independent_subset_extend[of "{}:: ('a::euclidean_space) set"] empty_subsetI independent_empty)
 | 
|
1809  | 
||
1810  | 
||
1811  | 
text {* Notion of dimension. *}
 | 
|
1812  | 
||
1813  | 
definition "dim V = (SOME n. \<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (card B = n))"  | 
|
1814  | 
||
1815  | 
lemma basis_exists:  "\<exists>B. (B :: ('a::euclidean_space) set) \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (card B = dim V)"
 | 
|
1816  | 
unfolding dim_def some_eq_ex[of "\<lambda>n. \<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (card B = n)"]  | 
|
1817  | 
using maximal_independent_subset[of V] independent_bound  | 
|
1818  | 
by auto  | 
|
1819  | 
||
1820  | 
text {* Consequences of independence or spanning for cardinality. *}
 | 
|
1821  | 
||
1822  | 
lemma independent_card_le_dim:  | 
|
1823  | 
  assumes "(B::('a::euclidean_space) set) \<subseteq> V" and "independent B" shows "card B \<le> dim V"
 | 
|
1824  | 
proof -  | 
|
1825  | 
from basis_exists[of V] `B \<subseteq> V`  | 
|
1826  | 
obtain B' where "independent B'" and "B \<subseteq> span B'" and "card B' = dim V" by blast  | 
|
1827  | 
with independent_span_bound[OF _ `independent B` `B \<subseteq> span B'`] independent_bound[of B']  | 
|
1828  | 
show ?thesis by auto  | 
|
1829  | 
qed  | 
|
1830  | 
||
1831  | 
lemma span_card_ge_dim:  "(B::('a::euclidean_space) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> finite B \<Longrightarrow> dim V \<le> card B"
 | 
|
1832  | 
by (metis basis_exists[of V] independent_span_bound subset_trans)  | 
|
1833  | 
||
1834  | 
lemma basis_card_eq_dim:  | 
|
1835  | 
  "B \<subseteq> (V:: ('a::euclidean_space) set) \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> finite B \<and> card B = dim V"
 | 
|
1836  | 
by (metis order_eq_iff independent_card_le_dim span_card_ge_dim independent_bound)  | 
|
1837  | 
||
1838  | 
lemma dim_unique: "(B::('a::euclidean_space) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> card B = n \<Longrightarrow> dim V = n"
 | 
|
1839  | 
by (metis basis_card_eq_dim)  | 
|
1840  | 
||
1841  | 
text {* More lemmas about dimension. *}
 | 
|
1842  | 
||
1843  | 
lemma dim_UNIV: "dim (UNIV :: ('a::euclidean_space) set) = DIM('a)"
 | 
|
1844  | 
  apply (rule dim_unique[of "(basis::nat=>'a) ` {..<DIM('a)}"])
 | 
|
1845  | 
using independent_basis by auto  | 
|
1846  | 
||
1847  | 
lemma dim_subset:  | 
|
1848  | 
  "(S:: ('a::euclidean_space) set) \<subseteq> T \<Longrightarrow> dim S \<le> dim T"
 | 
|
1849  | 
using basis_exists[of T] basis_exists[of S]  | 
|
1850  | 
by (metis independent_card_le_dim subset_trans)  | 
|
1851  | 
||
1852  | 
lemma dim_subset_UNIV: "dim (S:: ('a::euclidean_space) set) \<le> DIM('a)"
 | 
|
1853  | 
by (metis dim_subset subset_UNIV dim_UNIV)  | 
|
1854  | 
||
1855  | 
text {* Converses to those. *}
 | 
|
1856  | 
||
1857  | 
lemma card_ge_dim_independent:  | 
|
1858  | 
  assumes BV:"(B::('a::euclidean_space) set) \<subseteq> V" and iB:"independent B" and dVB:"dim V \<le> card B"
 | 
|
1859  | 
shows "V \<subseteq> span B"  | 
|
1860  | 
proof-  | 
|
1861  | 
  {fix a assume aV: "a \<in> V"
 | 
|
1862  | 
    {assume aB: "a \<notin> span B"
 | 
|
1863  | 
then have iaB: "independent (insert a B)" using iB aV BV by (simp add: independent_insert)  | 
|
1864  | 
from aV BV have th0: "insert a B \<subseteq> V" by blast  | 
|
1865  | 
from aB have "a \<notin>B" by (auto simp add: span_superset)  | 
|
1866  | 
with independent_card_le_dim[OF th0 iaB] dVB independent_bound[OF iB] have False by auto }  | 
|
1867  | 
then have "a \<in> span B" by blast}  | 
|
1868  | 
then show ?thesis by blast  | 
|
1869  | 
qed  | 
|
1870  | 
||
1871  | 
lemma card_le_dim_spanning:  | 
|
1872  | 
  assumes BV: "(B:: ('a::euclidean_space) set) \<subseteq> V" and VB: "V \<subseteq> span B"
 | 
|
1873  | 
and fB: "finite B" and dVB: "dim V \<ge> card B"  | 
|
1874  | 
shows "independent B"  | 
|
1875  | 
proof-  | 
|
1876  | 
  {fix a assume a: "a \<in> B" "a \<in> span (B -{a})"
 | 
|
1877  | 
from a fB have c0: "card B \<noteq> 0" by auto  | 
|
1878  | 
    from a fB have cb: "card (B -{a}) = card B - 1" by auto
 | 
|
1879  | 
    from BV a have th0: "B -{a} \<subseteq> V" by blast
 | 
|
1880  | 
    {fix x assume x: "x \<in> V"
 | 
|
1881  | 
      from a have eq: "insert a (B -{a}) = B" by blast
 | 
|
1882  | 
from x VB have x': "x \<in> span B" by blast  | 
|
1883  | 
from span_trans[OF a(2), unfolded eq, OF x']  | 
|
1884  | 
      have "x \<in> span (B -{a})" . }
 | 
|
1885  | 
    then have th1: "V \<subseteq> span (B -{a})" by blast
 | 
|
1886  | 
    have th2: "finite (B -{a})" using fB by auto
 | 
|
1887  | 
from span_card_ge_dim[OF th0 th1 th2]  | 
|
1888  | 
    have c: "dim V \<le> card (B -{a})" .
 | 
|
1889  | 
from c c0 dVB cb have False by simp}  | 
|
1890  | 
then show ?thesis unfolding dependent_def by blast  | 
|
1891  | 
qed  | 
|
1892  | 
||
1893  | 
lemma card_eq_dim: "(B:: ('a::euclidean_space) set) \<subseteq> V \<Longrightarrow> card B = dim V \<Longrightarrow> finite B \<Longrightarrow> independent B \<longleftrightarrow> V \<subseteq> span B"
 | 
|
1894  | 
by (metis order_eq_iff card_le_dim_spanning  | 
|
1895  | 
card_ge_dim_independent)  | 
|
1896  | 
||
1897  | 
text {* More general size bound lemmas. *}
 | 
|
1898  | 
||
1899  | 
lemma independent_bound_general:  | 
|
1900  | 
  "independent (S:: ('a::euclidean_space) set) \<Longrightarrow> finite S \<and> card S \<le> dim S"
 | 
|
1901  | 
by (metis independent_card_le_dim independent_bound subset_refl)  | 
|
1902  | 
||
1903  | 
lemma dependent_biggerset_general: "(finite (S:: ('a::euclidean_space) set) \<Longrightarrow> card S > dim S) \<Longrightarrow> dependent S"
 | 
|
1904  | 
using independent_bound_general[of S] by (metis linorder_not_le)  | 
|
1905  | 
||
1906  | 
lemma dim_span: "dim (span (S:: ('a::euclidean_space) set)) = dim S"
 | 
|
1907  | 
proof-  | 
|
1908  | 
have th0: "dim S \<le> dim (span S)"  | 
|
1909  | 
by (auto simp add: subset_eq intro: dim_subset span_superset)  | 
|
1910  | 
from basis_exists[of S]  | 
|
1911  | 
obtain B where B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" by blast  | 
|
1912  | 
from B have fB: "finite B" "card B = dim S" using independent_bound by blast+  | 
|
1913  | 
have bSS: "B \<subseteq> span S" using B(1) by (metis subset_eq span_inc)  | 
|
1914  | 
have sssB: "span S \<subseteq> span B" using span_mono[OF B(3)] by (simp add: span_span)  | 
|
1915  | 
from span_card_ge_dim[OF bSS sssB fB(1)] th0 show ?thesis  | 
|
1916  | 
using fB(2) by arith  | 
|
1917  | 
qed  | 
|
1918  | 
||
1919  | 
lemma subset_le_dim: "(S:: ('a::euclidean_space) set) \<subseteq> span T \<Longrightarrow> dim S \<le> dim T"
 | 
|
1920  | 
by (metis dim_span dim_subset)  | 
|
1921  | 
||
1922  | 
lemma span_eq_dim: "span (S:: ('a::euclidean_space) set) = span T ==> dim S = dim T"
 | 
|
1923  | 
by (metis dim_span)  | 
|
1924  | 
||
1925  | 
lemma spans_image:  | 
|
1926  | 
assumes lf: "linear f" and VB: "V \<subseteq> span B"  | 
|
1927  | 
shows "f ` V \<subseteq> span (f ` B)"  | 
|
1928  | 
unfolding span_linear_image[OF lf]  | 
|
1929  | 
by (metis VB image_mono)  | 
|
1930  | 
||
1931  | 
lemma dim_image_le:  | 
|
1932  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
1933  | 
assumes lf: "linear f" shows "dim (f ` S) \<le> dim (S)"  | 
|
1934  | 
proof-  | 
|
1935  | 
from basis_exists[of S] obtain B where  | 
|
1936  | 
B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" by blast  | 
|
1937  | 
from B have fB: "finite B" "card B = dim S" using independent_bound by blast+  | 
|
1938  | 
have "dim (f ` S) \<le> card (f ` B)"  | 
|
1939  | 
apply (rule span_card_ge_dim)  | 
|
1940  | 
using lf B fB by (auto simp add: span_linear_image spans_image subset_image_iff)  | 
|
1941  | 
also have "\<dots> \<le> dim S" using card_image_le[OF fB(1)] fB by simp  | 
|
1942  | 
finally show ?thesis .  | 
|
1943  | 
qed  | 
|
1944  | 
||
1945  | 
text {* Relation between bases and injectivity/surjectivity of map. *}
 | 
|
1946  | 
||
1947  | 
lemma spanning_surjective_image:  | 
|
1948  | 
assumes us: "UNIV \<subseteq> span S"  | 
|
1949  | 
and lf: "linear f" and sf: "surj f"  | 
|
1950  | 
shows "UNIV \<subseteq> span (f ` S)"  | 
|
1951  | 
proof-  | 
|
1952  | 
have "UNIV \<subseteq> f ` UNIV" using sf by (auto simp add: surj_def)  | 
|
1953  | 
also have " \<dots> \<subseteq> span (f ` S)" using spans_image[OF lf us] .  | 
|
1954  | 
finally show ?thesis .  | 
|
1955  | 
qed  | 
|
1956  | 
||
1957  | 
lemma independent_injective_image:  | 
|
1958  | 
assumes iS: "independent S" and lf: "linear f" and fi: "inj f"  | 
|
1959  | 
shows "independent (f ` S)"  | 
|
1960  | 
proof-  | 
|
1961  | 
  {fix a assume a: "a \<in> S" "f a \<in> span (f ` S - {f a})"
 | 
|
1962  | 
    have eq: "f ` S - {f a} = f ` (S - {a})" using fi
 | 
|
1963  | 
by (auto simp add: inj_on_def)  | 
|
1964  | 
    from a have "f a \<in> f ` span (S -{a})"
 | 
|
1965  | 
      unfolding eq span_linear_image[OF lf, of "S - {a}"]  by blast
 | 
|
1966  | 
    hence "a \<in> span (S -{a})" using fi by (auto simp add: inj_on_def)
 | 
|
1967  | 
with a(1) iS have False by (simp add: dependent_def) }  | 
|
1968  | 
then show ?thesis unfolding dependent_def by blast  | 
|
1969  | 
qed  | 
|
1970  | 
||
1971  | 
text {* Picking an orthogonal replacement for a spanning set. *}
 | 
|
1972  | 
||
1973  | 
(* FIXME : Move to some general theory ?*)  | 
|
1974  | 
definition "pairwise R S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y\<in> S. x\<noteq>y \<longrightarrow> R x y)"  | 
|
1975  | 
||
1976  | 
lemma vector_sub_project_orthogonal: "(b::'a::euclidean_space) \<bullet> (x - ((b \<bullet> x) / (b \<bullet> b)) *\<^sub>R b) = 0"  | 
|
1977  | 
unfolding inner_simps by auto  | 
|
1978  | 
||
1979  | 
lemma basis_orthogonal:  | 
|
1980  | 
  fixes B :: "('a::euclidean_space) set"
 | 
|
1981  | 
assumes fB: "finite B"  | 
|
1982  | 
shows "\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C"  | 
|
1983  | 
(is " \<exists>C. ?P B C")  | 
|
1984  | 
proof(induct rule: finite_induct[OF fB])  | 
|
1985  | 
  case 1 thus ?case apply (rule exI[where x="{}"]) by (auto simp add: pairwise_def)
 | 
|
1986  | 
next  | 
|
1987  | 
case (2 a B)  | 
|
1988  | 
note fB = `finite B` and aB = `a \<notin> B`  | 
|
1989  | 
from `\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C`  | 
|
1990  | 
obtain C where C: "finite C" "card C \<le> card B"  | 
|
1991  | 
"span C = span B" "pairwise orthogonal C" by blast  | 
|
1992  | 
let ?a = "a - setsum (\<lambda>x. (x \<bullet> a / (x \<bullet> x)) *\<^sub>R x) C"  | 
|
1993  | 
let ?C = "insert ?a C"  | 
|
1994  | 
from C(1) have fC: "finite ?C" by simp  | 
|
1995  | 
from fB aB C(1,2) have cC: "card ?C \<le> card (insert a B)" by (simp add: card_insert_if)  | 
|
1996  | 
  {fix x k
 | 
|
1997  | 
have th0: "\<And>(a::'a) b c. a - (b - c) = c + (a - b)" by (simp add: field_simps)  | 
|
1998  | 
have "x - k *\<^sub>R (a - (\<Sum>x\<in>C. (x \<bullet> a / (x \<bullet> x)) *\<^sub>R x)) \<in> span C \<longleftrightarrow> x - k *\<^sub>R a \<in> span C"  | 
|
1999  | 
apply (simp only: scaleR_right_diff_distrib th0)  | 
|
2000  | 
apply (rule span_add_eq)  | 
|
2001  | 
apply (rule span_mul)  | 
|
2002  | 
apply (rule span_setsum[OF C(1)])  | 
|
2003  | 
apply clarify  | 
|
2004  | 
apply (rule span_mul)  | 
|
2005  | 
by (rule span_superset)}  | 
|
2006  | 
then have SC: "span ?C = span (insert a B)"  | 
|
2007  | 
unfolding set_eq_iff span_breakdown_eq C(3)[symmetric] by auto  | 
|
2008  | 
thm pairwise_def  | 
|
2009  | 
  {fix x y assume xC: "x \<in> ?C" and yC: "y \<in> ?C" and xy: "x \<noteq> y"
 | 
|
2010  | 
    {assume xa: "x = ?a" and ya: "y = ?a"
 | 
|
2011  | 
have "orthogonal x y" using xa ya xy by blast}  | 
|
2012  | 
moreover  | 
|
2013  | 
    {assume xa: "x = ?a" and ya: "y \<noteq> ?a" "y \<in> C"
 | 
|
2014  | 
      from ya have Cy: "C = insert y (C - {y})" by blast
 | 
|
2015  | 
      have fth: "finite (C - {y})" using C by simp
 | 
|
2016  | 
have "orthogonal x y"  | 
|
2017  | 
using xa ya  | 
|
2018  | 
unfolding orthogonal_def xa inner_simps diff_eq_0_iff_eq  | 
|
2019  | 
apply simp  | 
|
2020  | 
apply (subst Cy)  | 
|
2021  | 
using C(1) fth  | 
|
2022  | 
apply (simp only: setsum_clauses)  | 
|
| 44454 | 2023  | 
apply (auto simp add: inner_add inner_commute[of y a] dot_lsum[OF fth])  | 
| 44133 | 2024  | 
apply (rule setsum_0')  | 
2025  | 
apply clarsimp  | 
|
2026  | 
apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])  | 
|
2027  | 
by auto}  | 
|
2028  | 
moreover  | 
|
2029  | 
    {assume xa: "x \<noteq> ?a" "x \<in> C" and ya: "y = ?a"
 | 
|
2030  | 
      from xa have Cx: "C = insert x (C - {x})" by blast
 | 
|
2031  | 
      have fth: "finite (C - {x})" using C by simp
 | 
|
2032  | 
have "orthogonal x y"  | 
|
2033  | 
using xa ya  | 
|
2034  | 
unfolding orthogonal_def ya inner_simps diff_eq_0_iff_eq  | 
|
2035  | 
apply simp  | 
|
2036  | 
apply (subst Cx)  | 
|
2037  | 
using C(1) fth  | 
|
2038  | 
apply (simp only: setsum_clauses)  | 
|
2039  | 
apply (subst inner_commute[of x])  | 
|
| 44454 | 2040  | 
apply (auto simp add: inner_add inner_commute[of x a] dot_rsum[OF fth])  | 
| 44133 | 2041  | 
apply (rule setsum_0')  | 
2042  | 
apply clarsimp  | 
|
2043  | 
apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])  | 
|
2044  | 
by auto}  | 
|
2045  | 
moreover  | 
|
2046  | 
    {assume xa: "x \<in> C" and ya: "y \<in> C"
 | 
|
2047  | 
have "orthogonal x y" using xa ya xy C(4) unfolding pairwise_def by blast}  | 
|
2048  | 
ultimately have "orthogonal x y" using xC yC by blast}  | 
|
2049  | 
then have CPO: "pairwise orthogonal ?C" unfolding pairwise_def by blast  | 
|
2050  | 
from fC cC SC CPO have "?P (insert a B) ?C" by blast  | 
|
2051  | 
then show ?case by blast  | 
|
2052  | 
qed  | 
|
2053  | 
||
2054  | 
lemma orthogonal_basis_exists:  | 
|
2055  | 
  fixes V :: "('a::euclidean_space) set"
 | 
|
2056  | 
shows "\<exists>B. independent B \<and> B \<subseteq> span V \<and> V \<subseteq> span B \<and> (card B = dim V) \<and> pairwise orthogonal B"  | 
|
2057  | 
proof-  | 
|
2058  | 
from basis_exists[of V] obtain B where B: "B \<subseteq> V" "independent B" "V \<subseteq> span B" "card B = dim V" by blast  | 
|
2059  | 
from B have fB: "finite B" "card B = dim V" using independent_bound by auto  | 
|
2060  | 
from basis_orthogonal[OF fB(1)] obtain C where  | 
|
2061  | 
C: "finite C" "card C \<le> card B" "span C = span B" "pairwise orthogonal C" by blast  | 
|
2062  | 
from C B  | 
|
2063  | 
have CSV: "C \<subseteq> span V" by (metis span_inc span_mono subset_trans)  | 
|
2064  | 
from span_mono[OF B(3)] C have SVC: "span V \<subseteq> span C" by (simp add: span_span)  | 
|
2065  | 
from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB  | 
|
2066  | 
have iC: "independent C" by (simp add: dim_span)  | 
|
2067  | 
from C fB have "card C \<le> dim V" by simp  | 
|
2068  | 
moreover have "dim V \<le> card C" using span_card_ge_dim[OF CSV SVC C(1)]  | 
|
2069  | 
by (simp add: dim_span)  | 
|
2070  | 
ultimately have CdV: "card C = dim V" using C(1) by simp  | 
|
2071  | 
from C B CSV CdV iC show ?thesis by auto  | 
|
2072  | 
qed  | 
|
2073  | 
||
2074  | 
lemma span_eq: "span S = span T \<longleftrightarrow> S \<subseteq> span T \<and> T \<subseteq> span S"  | 
|
2075  | 
using span_inc[unfolded subset_eq] using span_mono[of T "span S"] span_mono[of S "span T"]  | 
|
2076  | 
by(auto simp add: span_span)  | 
|
2077  | 
||
2078  | 
text {* Low-dimensional subset is in a hyperplane (weak orthogonal complement). *}
 | 
|
2079  | 
||
2080  | 
lemma span_not_univ_orthogonal: fixes S::"('a::euclidean_space) set"
 | 
|
2081  | 
assumes sU: "span S \<noteq> UNIV"  | 
|
2082  | 
shows "\<exists>(a::'a). a \<noteq>0 \<and> (\<forall>x \<in> span S. a \<bullet> x = 0)"  | 
|
2083  | 
proof-  | 
|
2084  | 
from sU obtain a where a: "a \<notin> span S" by blast  | 
|
2085  | 
from orthogonal_basis_exists obtain B where  | 
|
2086  | 
B: "independent B" "B \<subseteq> span S" "S \<subseteq> span B" "card B = dim S" "pairwise orthogonal B"  | 
|
2087  | 
by blast  | 
|
2088  | 
from B have fB: "finite B" "card B = dim S" using independent_bound by auto  | 
|
2089  | 
from span_mono[OF B(2)] span_mono[OF B(3)]  | 
|
2090  | 
have sSB: "span S = span B" by (simp add: span_span)  | 
|
2091  | 
let ?a = "a - setsum (\<lambda>b. (a \<bullet> b / (b \<bullet> b)) *\<^sub>R b) B"  | 
|
2092  | 
have "setsum (\<lambda>b. (a \<bullet> b / (b \<bullet> b)) *\<^sub>R b) B \<in> span S"  | 
|
2093  | 
unfolding sSB  | 
|
2094  | 
apply (rule span_setsum[OF fB(1)])  | 
|
2095  | 
apply clarsimp  | 
|
2096  | 
apply (rule span_mul)  | 
|
2097  | 
by (rule span_superset)  | 
|
2098  | 
with a have a0:"?a \<noteq> 0" by auto  | 
|
2099  | 
have "\<forall>x\<in>span B. ?a \<bullet> x = 0"  | 
|
2100  | 
proof(rule span_induct')  | 
|
| 44454 | 2101  | 
    show "subspace {x. ?a \<bullet> x = 0}" by (auto simp add: subspace_def inner_add)
 | 
| 44133 | 2102  | 
next  | 
2103  | 
    {fix x assume x: "x \<in> B"
 | 
|
2104  | 
      from x have B': "B = insert x (B - {x})" by blast
 | 
|
2105  | 
      have fth: "finite (B - {x})" using fB by simp
 | 
|
2106  | 
have "?a \<bullet> x = 0"  | 
|
2107  | 
apply (subst B') using fB fth  | 
|
2108  | 
unfolding setsum_clauses(2)[OF fth]  | 
|
2109  | 
apply simp unfolding inner_simps  | 
|
| 44454 | 2110  | 
apply (clarsimp simp add: inner_add dot_lsum)  | 
| 44133 | 2111  | 
apply (rule setsum_0', rule ballI)  | 
2112  | 
unfolding inner_commute  | 
|
2113  | 
by (auto simp add: x field_simps intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])}  | 
|
2114  | 
then show "\<forall>x \<in> B. ?a \<bullet> x = 0" by blast  | 
|
2115  | 
qed  | 
|
2116  | 
with a0 show ?thesis unfolding sSB by (auto intro: exI[where x="?a"])  | 
|
2117  | 
qed  | 
|
2118  | 
||
2119  | 
lemma span_not_univ_subset_hyperplane:  | 
|
2120  | 
  assumes SU: "span S \<noteq> (UNIV ::('a::euclidean_space) set)"
 | 
|
2121  | 
  shows "\<exists> a. a \<noteq>0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
 | 
|
2122  | 
using span_not_univ_orthogonal[OF SU] by auto  | 
|
2123  | 
||
2124  | 
lemma lowdim_subset_hyperplane: fixes S::"('a::euclidean_space) set"
 | 
|
2125  | 
  assumes d: "dim S < DIM('a)"
 | 
|
2126  | 
  shows "\<exists>(a::'a). a  \<noteq> 0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
 | 
|
2127  | 
proof-  | 
|
2128  | 
  {assume "span S = UNIV"
 | 
|
2129  | 
    hence "dim (span S) = dim (UNIV :: ('a) set)" by simp
 | 
|
2130  | 
    hence "dim S = DIM('a)" by (simp add: dim_span dim_UNIV)
 | 
|
2131  | 
with d have False by arith}  | 
|
2132  | 
hence th: "span S \<noteq> UNIV" by blast  | 
|
2133  | 
from span_not_univ_subset_hyperplane[OF th] show ?thesis .  | 
|
2134  | 
qed  | 
|
2135  | 
||
2136  | 
text {* We can extend a linear basis-basis injection to the whole set. *}
 | 
|
2137  | 
||
2138  | 
lemma linear_indep_image_lemma:  | 
|
2139  | 
assumes lf: "linear f" and fB: "finite B"  | 
|
2140  | 
and ifB: "independent (f ` B)"  | 
|
2141  | 
and fi: "inj_on f B" and xsB: "x \<in> span B"  | 
|
2142  | 
and fx: "f x = 0"  | 
|
2143  | 
shows "x = 0"  | 
|
2144  | 
using fB ifB fi xsB fx  | 
|
2145  | 
proof(induct arbitrary: x rule: finite_induct[OF fB])  | 
|
| 44142 | 2146  | 
case 1 thus ?case by auto  | 
| 44133 | 2147  | 
next  | 
2148  | 
case (2 a b x)  | 
|
2149  | 
have fb: "finite b" using "2.prems" by simp  | 
|
2150  | 
have th0: "f ` b \<subseteq> f ` (insert a b)"  | 
|
2151  | 
apply (rule image_mono) by blast  | 
|
2152  | 
from independent_mono[ OF "2.prems"(2) th0]  | 
|
2153  | 
have ifb: "independent (f ` b)" .  | 
|
2154  | 
have fib: "inj_on f b"  | 
|
2155  | 
apply (rule subset_inj_on [OF "2.prems"(3)])  | 
|
2156  | 
by blast  | 
|
2157  | 
from span_breakdown[of a "insert a b", simplified, OF "2.prems"(4)]  | 
|
2158  | 
  obtain k where k: "x - k*\<^sub>R a \<in> span (b -{a})" by blast
 | 
|
2159  | 
have "f (x - k*\<^sub>R a) \<in> span (f ` b)"  | 
|
2160  | 
unfolding span_linear_image[OF lf]  | 
|
2161  | 
apply (rule imageI)  | 
|
2162  | 
    using k span_mono[of "b-{a}" b] by blast
 | 
|
2163  | 
hence "f x - k*\<^sub>R f a \<in> span (f ` b)"  | 
|
2164  | 
by (simp add: linear_sub[OF lf] linear_cmul[OF lf])  | 
|
2165  | 
hence th: "-k *\<^sub>R f a \<in> span (f ` b)"  | 
|
2166  | 
using "2.prems"(5) by simp  | 
|
2167  | 
  {assume k0: "k = 0"
 | 
|
2168  | 
    from k0 k have "x \<in> span (b -{a})" by simp
 | 
|
2169  | 
    then have "x \<in> span b" using span_mono[of "b-{a}" b]
 | 
|
2170  | 
by blast}  | 
|
2171  | 
moreover  | 
|
2172  | 
  {assume k0: "k \<noteq> 0"
 | 
|
2173  | 
from span_mul[OF th, of "- 1/ k"] k0  | 
|
2174  | 
have th1: "f a \<in> span (f ` b)"  | 
|
2175  | 
by auto  | 
|
2176  | 
    from inj_on_image_set_diff[OF "2.prems"(3), of "insert a b " "{a}", symmetric]
 | 
|
2177  | 
    have tha: "f ` insert a b - f ` {a} = f ` (insert a b - {a})" by blast
 | 
|
2178  | 
from "2.prems"(2) [unfolded dependent_def bex_simps(8), rule_format, of "f a"]  | 
|
2179  | 
have "f a \<notin> span (f ` b)" using tha  | 
|
2180  | 
using "2.hyps"(2)  | 
|
2181  | 
"2.prems"(3) by auto  | 
|
2182  | 
with th1 have False by blast  | 
|
2183  | 
then have "x \<in> span b" by blast}  | 
|
2184  | 
ultimately have xsb: "x \<in> span b" by blast  | 
|
2185  | 
from "2.hyps"(3)[OF fb ifb fib xsb "2.prems"(5)]  | 
|
2186  | 
show "x = 0" .  | 
|
2187  | 
qed  | 
|
2188  | 
||
2189  | 
text {* We can extend a linear mapping from basis. *}
 | 
|
2190  | 
||
2191  | 
lemma linear_independent_extend_lemma:  | 
|
2192  | 
fixes f :: "'a::real_vector \<Rightarrow> 'b::real_vector"  | 
|
2193  | 
assumes fi: "finite B" and ib: "independent B"  | 
|
2194  | 
shows "\<exists>g. (\<forall>x\<in> span B. \<forall>y\<in> span B. g (x + y) = g x + g y)  | 
|
2195  | 
\<and> (\<forall>x\<in> span B. \<forall>c. g (c*\<^sub>R x) = c *\<^sub>R g x)  | 
|
2196  | 
\<and> (\<forall>x\<in> B. g x = f x)"  | 
|
2197  | 
using ib fi  | 
|
2198  | 
proof(induct rule: finite_induct[OF fi])  | 
|
| 44142 | 2199  | 
case 1 thus ?case by auto  | 
| 44133 | 2200  | 
next  | 
2201  | 
case (2 a b)  | 
|
2202  | 
from "2.prems" "2.hyps" have ibf: "independent b" "finite b"  | 
|
2203  | 
by (simp_all add: independent_insert)  | 
|
2204  | 
from "2.hyps"(3)[OF ibf] obtain g where  | 
|
2205  | 
g: "\<forall>x\<in>span b. \<forall>y\<in>span b. g (x + y) = g x + g y"  | 
|
2206  | 
"\<forall>x\<in>span b. \<forall>c. g (c *\<^sub>R x) = c *\<^sub>R g x" "\<forall>x\<in>b. g x = f x" by blast  | 
|
2207  | 
let ?h = "\<lambda>z. SOME k. (z - k *\<^sub>R a) \<in> span b"  | 
|
2208  | 
  {fix z assume z: "z \<in> span (insert a b)"
 | 
|
2209  | 
have th0: "z - ?h z *\<^sub>R a \<in> span b"  | 
|
2210  | 
apply (rule someI_ex)  | 
|
2211  | 
unfolding span_breakdown_eq[symmetric]  | 
|
2212  | 
using z .  | 
|
2213  | 
    {fix k assume k: "z - k *\<^sub>R a \<in> span b"
 | 
|
2214  | 
have eq: "z - ?h z *\<^sub>R a - (z - k*\<^sub>R a) = (k - ?h z) *\<^sub>R a"  | 
|
2215  | 
by (simp add: field_simps scaleR_left_distrib [symmetric])  | 
|
2216  | 
from span_sub[OF th0 k]  | 
|
2217  | 
have khz: "(k - ?h z) *\<^sub>R a \<in> span b" by (simp add: eq)  | 
|
2218  | 
      {assume "k \<noteq> ?h z" hence k0: "k - ?h z \<noteq> 0" by simp
 | 
|
2219  | 
from k0 span_mul[OF khz, of "1 /(k - ?h z)"]  | 
|
2220  | 
have "a \<in> span b" by simp  | 
|
2221  | 
with "2.prems"(1) "2.hyps"(2) have False  | 
|
2222  | 
by (auto simp add: dependent_def)}  | 
|
2223  | 
then have "k = ?h z" by blast}  | 
|
2224  | 
with th0 have "z - ?h z *\<^sub>R a \<in> span b \<and> (\<forall>k. z - k *\<^sub>R a \<in> span b \<longrightarrow> k = ?h z)" by blast}  | 
|
2225  | 
note h = this  | 
|
2226  | 
let ?g = "\<lambda>z. ?h z *\<^sub>R f a + g (z - ?h z *\<^sub>R a)"  | 
|
2227  | 
  {fix x y assume x: "x \<in> span (insert a b)" and y: "y \<in> span (insert a b)"
 | 
|
2228  | 
have tha: "\<And>(x::'a) y a k l. (x + y) - (k + l) *\<^sub>R a = (x - k *\<^sub>R a) + (y - l *\<^sub>R a)"  | 
|
2229  | 
by (simp add: algebra_simps)  | 
|
2230  | 
have addh: "?h (x + y) = ?h x + ?h y"  | 
|
2231  | 
apply (rule conjunct2[OF h, rule_format, symmetric])  | 
|
2232  | 
apply (rule span_add[OF x y])  | 
|
2233  | 
unfolding tha  | 
|
2234  | 
by (metis span_add x y conjunct1[OF h, rule_format])  | 
|
2235  | 
have "?g (x + y) = ?g x + ?g y"  | 
|
2236  | 
unfolding addh tha  | 
|
2237  | 
g(1)[rule_format,OF conjunct1[OF h, OF x] conjunct1[OF h, OF y]]  | 
|
2238  | 
by (simp add: scaleR_left_distrib)}  | 
|
2239  | 
moreover  | 
|
2240  | 
  {fix x:: "'a" and c:: real  assume x: "x \<in> span (insert a b)"
 | 
|
2241  | 
have tha: "\<And>(x::'a) c k a. c *\<^sub>R x - (c * k) *\<^sub>R a = c *\<^sub>R (x - k *\<^sub>R a)"  | 
|
2242  | 
by (simp add: algebra_simps)  | 
|
2243  | 
have hc: "?h (c *\<^sub>R x) = c * ?h x"  | 
|
2244  | 
apply (rule conjunct2[OF h, rule_format, symmetric])  | 
|
2245  | 
apply (metis span_mul x)  | 
|
2246  | 
by (metis tha span_mul x conjunct1[OF h])  | 
|
2247  | 
have "?g (c *\<^sub>R x) = c*\<^sub>R ?g x"  | 
|
2248  | 
unfolding hc tha g(2)[rule_format, OF conjunct1[OF h, OF x]]  | 
|
2249  | 
by (simp add: algebra_simps)}  | 
|
2250  | 
moreover  | 
|
2251  | 
  {fix x assume x: "x \<in> (insert a b)"
 | 
|
2252  | 
    {assume xa: "x = a"
 | 
|
2253  | 
have ha1: "1 = ?h a"  | 
|
2254  | 
apply (rule conjunct2[OF h, rule_format])  | 
|
2255  | 
apply (metis span_superset insertI1)  | 
|
2256  | 
using conjunct1[OF h, OF span_superset, OF insertI1]  | 
|
2257  | 
by (auto simp add: span_0)  | 
|
2258  | 
||
2259  | 
from xa ha1[symmetric] have "?g x = f x"  | 
|
2260  | 
apply simp  | 
|
2261  | 
using g(2)[rule_format, OF span_0, of 0]  | 
|
2262  | 
by simp}  | 
|
2263  | 
moreover  | 
|
2264  | 
    {assume xb: "x \<in> b"
 | 
|
2265  | 
have h0: "0 = ?h x"  | 
|
2266  | 
apply (rule conjunct2[OF h, rule_format])  | 
|
2267  | 
apply (metis span_superset x)  | 
|
2268  | 
apply simp  | 
|
2269  | 
apply (metis span_superset xb)  | 
|
2270  | 
done  | 
|
2271  | 
have "?g x = f x"  | 
|
2272  | 
by (simp add: h0[symmetric] g(3)[rule_format, OF xb])}  | 
|
2273  | 
ultimately have "?g x = f x" using x by blast }  | 
|
2274  | 
ultimately show ?case apply - apply (rule exI[where x="?g"]) by blast  | 
|
2275  | 
qed  | 
|
2276  | 
||
2277  | 
lemma linear_independent_extend:  | 
|
2278  | 
  assumes iB: "independent (B:: ('a::euclidean_space) set)"
 | 
|
2279  | 
shows "\<exists>g. linear g \<and> (\<forall>x\<in>B. g x = f x)"  | 
|
2280  | 
proof-  | 
|
2281  | 
from maximal_independent_subset_extend[of B UNIV] iB  | 
|
2282  | 
obtain C where C: "B \<subseteq> C" "independent C" "\<And>x. x \<in> span C" by auto  | 
|
2283  | 
||
2284  | 
from C(2) independent_bound[of C] linear_independent_extend_lemma[of C f]  | 
|
2285  | 
obtain g where g: "(\<forall>x\<in> span C. \<forall>y\<in> span C. g (x + y) = g x + g y)  | 
|
2286  | 
\<and> (\<forall>x\<in> span C. \<forall>c. g (c*\<^sub>R x) = c *\<^sub>R g x)  | 
|
2287  | 
\<and> (\<forall>x\<in> C. g x = f x)" by blast  | 
|
2288  | 
from g show ?thesis unfolding linear_def using C  | 
|
2289  | 
apply clarsimp by blast  | 
|
2290  | 
qed  | 
|
2291  | 
||
2292  | 
text {* Can construct an isomorphism between spaces of same dimension. *}
 | 
|
2293  | 
||
2294  | 
lemma card_le_inj: assumes fA: "finite A" and fB: "finite B"  | 
|
2295  | 
and c: "card A \<le> card B" shows "(\<exists>f. f ` A \<subseteq> B \<and> inj_on f A)"  | 
|
2296  | 
using fB c  | 
|
2297  | 
proof(induct arbitrary: B rule: finite_induct[OF fA])  | 
|
2298  | 
case 1 thus ?case by simp  | 
|
2299  | 
next  | 
|
2300  | 
case (2 x s t)  | 
|
2301  | 
thus ?case  | 
|
2302  | 
proof(induct rule: finite_induct[OF "2.prems"(1)])  | 
|
2303  | 
case 1 then show ?case by simp  | 
|
2304  | 
next  | 
|
2305  | 
case (2 y t)  | 
|
2306  | 
from "2.prems"(1,2,5) "2.hyps"(1,2) have cst:"card s \<le> card t" by simp  | 
|
2307  | 
from "2.prems"(3) [OF "2.hyps"(1) cst] obtain f where  | 
|
2308  | 
f: "f ` s \<subseteq> t \<and> inj_on f s" by blast  | 
|
2309  | 
from f "2.prems"(2) "2.hyps"(2) show ?case  | 
|
2310  | 
apply -  | 
|
2311  | 
apply (rule exI[where x = "\<lambda>z. if z = x then y else f z"])  | 
|
2312  | 
by (auto simp add: inj_on_def)  | 
|
2313  | 
qed  | 
|
2314  | 
qed  | 
|
2315  | 
||
2316  | 
lemma card_subset_eq: assumes fB: "finite B" and AB: "A \<subseteq> B" and  | 
|
2317  | 
c: "card A = card B"  | 
|
2318  | 
shows "A = B"  | 
|
2319  | 
proof-  | 
|
2320  | 
from fB AB have fA: "finite A" by (auto intro: finite_subset)  | 
|
2321  | 
from fA fB have fBA: "finite (B - A)" by auto  | 
|
2322  | 
  have e: "A \<inter> (B - A) = {}" by blast
 | 
|
2323  | 
have eq: "A \<union> (B - A) = B" using AB by blast  | 
|
2324  | 
from card_Un_disjoint[OF fA fBA e, unfolded eq c]  | 
|
2325  | 
have "card (B - A) = 0" by arith  | 
|
2326  | 
  hence "B - A = {}" unfolding card_eq_0_iff using fA fB by simp
 | 
|
2327  | 
with AB show "A = B" by blast  | 
|
2328  | 
qed  | 
|
2329  | 
||
2330  | 
lemma subspace_isomorphism:  | 
|
2331  | 
  assumes s: "subspace (S:: ('a::euclidean_space) set)"
 | 
|
2332  | 
  and t: "subspace (T :: ('b::euclidean_space) set)"
 | 
|
2333  | 
and d: "dim S = dim T"  | 
|
2334  | 
shows "\<exists>f. linear f \<and> f ` S = T \<and> inj_on f S"  | 
|
2335  | 
proof-  | 
|
2336  | 
from basis_exists[of S] independent_bound obtain B where  | 
|
2337  | 
B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" and fB: "finite B" by blast  | 
|
2338  | 
from basis_exists[of T] independent_bound obtain C where  | 
|
2339  | 
C: "C \<subseteq> T" "independent C" "T \<subseteq> span C" "card C = dim T" and fC: "finite C" by blast  | 
|
2340  | 
from B(4) C(4) card_le_inj[of B C] d obtain f where  | 
|
2341  | 
f: "f ` B \<subseteq> C" "inj_on f B" using `finite B` `finite C` by auto  | 
|
2342  | 
from linear_independent_extend[OF B(2)] obtain g where  | 
|
2343  | 
g: "linear g" "\<forall>x\<in> B. g x = f x" by blast  | 
|
2344  | 
from inj_on_iff_eq_card[OF fB, of f] f(2)  | 
|
2345  | 
have "card (f ` B) = card B" by simp  | 
|
2346  | 
with B(4) C(4) have ceq: "card (f ` B) = card C" using d  | 
|
2347  | 
by simp  | 
|
2348  | 
have "g ` B = f ` B" using g(2)  | 
|
2349  | 
by (auto simp add: image_iff)  | 
|
2350  | 
also have "\<dots> = C" using card_subset_eq[OF fC f(1) ceq] .  | 
|
2351  | 
finally have gBC: "g ` B = C" .  | 
|
2352  | 
have gi: "inj_on g B" using f(2) g(2)  | 
|
2353  | 
by (auto simp add: inj_on_def)  | 
|
2354  | 
note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi]  | 
|
2355  | 
  {fix x y assume x: "x \<in> S" and y: "y \<in> S" and gxy:"g x = g y"
 | 
|
2356  | 
from B(3) x y have x': "x \<in> span B" and y': "y \<in> span B" by blast+  | 
|
2357  | 
from gxy have th0: "g (x - y) = 0" by (simp add: linear_sub[OF g(1)])  | 
|
2358  | 
have th1: "x - y \<in> span B" using x' y' by (metis span_sub)  | 
|
2359  | 
have "x=y" using g0[OF th1 th0] by simp }  | 
|
2360  | 
then have giS: "inj_on g S"  | 
|
2361  | 
unfolding inj_on_def by blast  | 
|
2362  | 
from span_subspace[OF B(1,3) s]  | 
|
2363  | 
have "g ` S = span (g ` B)" by (simp add: span_linear_image[OF g(1)])  | 
|
2364  | 
also have "\<dots> = span C" unfolding gBC ..  | 
|
2365  | 
also have "\<dots> = T" using span_subspace[OF C(1,3) t] .  | 
|
2366  | 
finally have gS: "g ` S = T" .  | 
|
2367  | 
from g(1) gS giS show ?thesis by blast  | 
|
2368  | 
qed  | 
|
2369  | 
||
2370  | 
text {* Linear functions are equal on a subspace if they are on a spanning set. *}
 | 
|
2371  | 
||
2372  | 
lemma subspace_kernel:  | 
|
2373  | 
assumes lf: "linear f"  | 
|
2374  | 
  shows "subspace {x. f x = 0}"
 | 
|
2375  | 
apply (simp add: subspace_def)  | 
|
2376  | 
by (simp add: linear_add[OF lf] linear_cmul[OF lf] linear_0[OF lf])  | 
|
2377  | 
||
2378  | 
lemma linear_eq_0_span:  | 
|
2379  | 
assumes lf: "linear f" and f0: "\<forall>x\<in>B. f x = 0"  | 
|
2380  | 
shows "\<forall>x \<in> span B. f x = 0"  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
2381  | 
using f0 subspace_kernel[OF lf]  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
2382  | 
by (rule span_induct')  | 
| 44133 | 2383  | 
|
2384  | 
lemma linear_eq_0:  | 
|
2385  | 
assumes lf: "linear f" and SB: "S \<subseteq> span B" and f0: "\<forall>x\<in>B. f x = 0"  | 
|
2386  | 
shows "\<forall>x \<in> S. f x = 0"  | 
|
2387  | 
by (metis linear_eq_0_span[OF lf] subset_eq SB f0)  | 
|
2388  | 
||
2389  | 
lemma linear_eq:  | 
|
2390  | 
assumes lf: "linear f" and lg: "linear g" and S: "S \<subseteq> span B"  | 
|
2391  | 
and fg: "\<forall> x\<in> B. f x = g x"  | 
|
2392  | 
shows "\<forall>x\<in> S. f x = g x"  | 
|
2393  | 
proof-  | 
|
2394  | 
let ?h = "\<lambda>x. f x - g x"  | 
|
2395  | 
from fg have fg': "\<forall>x\<in> B. ?h x = 0" by simp  | 
|
2396  | 
from linear_eq_0[OF linear_compose_sub[OF lf lg] S fg']  | 
|
2397  | 
show ?thesis by simp  | 
|
2398  | 
qed  | 
|
2399  | 
||
2400  | 
lemma linear_eq_stdbasis:  | 
|
2401  | 
assumes lf: "linear (f::'a::euclidean_space \<Rightarrow> _)" and lg: "linear g"  | 
|
2402  | 
  and fg: "\<forall>i<DIM('a::euclidean_space). f (basis i) = g(basis i)"
 | 
|
2403  | 
shows "f = g"  | 
|
2404  | 
proof-  | 
|
2405  | 
  let ?U = "{..<DIM('a)}"
 | 
|
2406  | 
  let ?I = "(basis::nat=>'a) ` {..<DIM('a)}"
 | 
|
2407  | 
  {fix x assume x: "x \<in> (UNIV :: 'a set)"
 | 
|
2408  | 
from equalityD2[OF span_basis'[where 'a='a]]  | 
|
2409  | 
have IU: " (UNIV :: 'a set) \<subseteq> span ?I" by blast  | 
|
2410  | 
have "f x = g x" apply(rule linear_eq[OF lf lg IU,rule_format]) using fg x by auto }  | 
|
| 44454 | 2411  | 
then show ?thesis by auto  | 
| 44133 | 2412  | 
qed  | 
2413  | 
||
2414  | 
text {* Similar results for bilinear functions. *}
 | 
|
2415  | 
||
2416  | 
lemma bilinear_eq:  | 
|
2417  | 
assumes bf: "bilinear f"  | 
|
2418  | 
and bg: "bilinear g"  | 
|
2419  | 
and SB: "S \<subseteq> span B" and TC: "T \<subseteq> span C"  | 
|
2420  | 
and fg: "\<forall>x\<in> B. \<forall>y\<in> C. f x y = g x y"  | 
|
2421  | 
shows "\<forall>x\<in>S. \<forall>y\<in>T. f x y = g x y "  | 
|
2422  | 
proof-  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
2423  | 
  let ?P = "{x. \<forall>y\<in> span C. f x y = g x y}"
 | 
| 44133 | 2424  | 
from bf bg have sp: "subspace ?P"  | 
2425  | 
unfolding bilinear_def linear_def subspace_def bf bg  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
2426  | 
by(auto simp add: span_0 bilinear_lzero[OF bf] bilinear_lzero[OF bg] span_add Ball_def intro: bilinear_ladd[OF bf])  | 
| 44133 | 2427  | 
|
2428  | 
have "\<forall>x \<in> span B. \<forall>y\<in> span C. f x y = g x y"  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
2429  | 
apply (rule span_induct' [OF _ sp])  | 
| 44133 | 2430  | 
apply (rule ballI)  | 
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
2431  | 
apply (rule span_induct')  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
2432  | 
apply (simp add: fg)  | 
| 44133 | 2433  | 
apply (auto simp add: subspace_def)  | 
2434  | 
using bf bg unfolding bilinear_def linear_def  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44166 
diff
changeset
 | 
2435  | 
by(auto simp add: span_0 bilinear_rzero[OF bf] bilinear_rzero[OF bg] span_add Ball_def intro: bilinear_ladd[OF bf])  | 
| 44454 | 2436  | 
then show ?thesis using SB TC by auto  | 
| 44133 | 2437  | 
qed  | 
2438  | 
||
2439  | 
lemma bilinear_eq_stdbasis: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> _"  | 
|
2440  | 
assumes bf: "bilinear f"  | 
|
2441  | 
and bg: "bilinear g"  | 
|
2442  | 
  and fg: "\<forall>i<DIM('a). \<forall>j<DIM('b). f (basis i) (basis j) = g (basis i) (basis j)"
 | 
|
2443  | 
shows "f = g"  | 
|
2444  | 
proof-  | 
|
2445  | 
  from fg have th: "\<forall>x \<in> (basis ` {..<DIM('a)}). \<forall>y\<in> (basis ` {..<DIM('b)}). f x y = g x y" by blast
 | 
|
2446  | 
from bilinear_eq[OF bf bg equalityD2[OF span_basis'] equalityD2[OF span_basis'] th]  | 
|
| 44454 | 2447  | 
show ?thesis by blast  | 
| 44133 | 2448  | 
qed  | 
2449  | 
||
2450  | 
text {* Detailed theorems about left and right invertibility in general case. *}
 | 
|
2451  | 
||
2452  | 
lemma linear_injective_left_inverse: fixes f::"'a::euclidean_space => 'b::euclidean_space"  | 
|
2453  | 
assumes lf: "linear f" and fi: "inj f"  | 
|
2454  | 
shows "\<exists>g. linear g \<and> g o f = id"  | 
|
2455  | 
proof-  | 
|
2456  | 
from linear_independent_extend[OF independent_injective_image, OF independent_basis, OF lf fi]  | 
|
2457  | 
obtain h:: "'b => 'a" where h: "linear h"  | 
|
2458  | 
    " \<forall>x \<in> f ` basis ` {..<DIM('a)}. h x = inv f x" by blast
 | 
|
2459  | 
from h(2)  | 
|
2460  | 
  have th: "\<forall>i<DIM('a). (h \<circ> f) (basis i) = id (basis i)"
 | 
|
2461  | 
using inv_o_cancel[OF fi, unfolded fun_eq_iff id_def o_def]  | 
|
2462  | 
by auto  | 
|
2463  | 
||
2464  | 
from linear_eq_stdbasis[OF linear_compose[OF lf h(1)] linear_id th]  | 
|
2465  | 
have "h o f = id" .  | 
|
2466  | 
then show ?thesis using h(1) by blast  | 
|
2467  | 
qed  | 
|
2468  | 
||
2469  | 
lemma linear_surjective_right_inverse: fixes f::"'a::euclidean_space => 'b::euclidean_space"  | 
|
2470  | 
assumes lf: "linear f" and sf: "surj f"  | 
|
2471  | 
shows "\<exists>g. linear g \<and> f o g = id"  | 
|
2472  | 
proof-  | 
|
2473  | 
from linear_independent_extend[OF independent_basis[where 'a='b],of "inv f"]  | 
|
2474  | 
obtain h:: "'b \<Rightarrow> 'a" where  | 
|
2475  | 
    h: "linear h" "\<forall> x\<in> basis ` {..<DIM('b)}. h x = inv f x" by blast
 | 
|
2476  | 
from h(2)  | 
|
2477  | 
  have th: "\<forall>i<DIM('b). (f o h) (basis i) = id (basis i)"
 | 
|
2478  | 
using sf by(auto simp add: surj_iff_all)  | 
|
2479  | 
from linear_eq_stdbasis[OF linear_compose[OF h(1) lf] linear_id th]  | 
|
2480  | 
have "f o h = id" .  | 
|
2481  | 
then show ?thesis using h(1) by blast  | 
|
2482  | 
qed  | 
|
2483  | 
||
2484  | 
text {* An injective map @{typ "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"} is also surjective. *}
 | 
|
2485  | 
||
2486  | 
lemma linear_injective_imp_surjective: fixes f::"'a::euclidean_space => 'a::euclidean_space"  | 
|
2487  | 
assumes lf: "linear f" and fi: "inj f"  | 
|
2488  | 
shows "surj f"  | 
|
2489  | 
proof-  | 
|
2490  | 
let ?U = "UNIV :: 'a set"  | 
|
2491  | 
from basis_exists[of ?U] obtain B  | 
|
2492  | 
where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" "card B = dim ?U"  | 
|
2493  | 
by blast  | 
|
2494  | 
from B(4) have d: "dim ?U = card B" by simp  | 
|
2495  | 
have th: "?U \<subseteq> span (f ` B)"  | 
|
2496  | 
apply (rule card_ge_dim_independent)  | 
|
2497  | 
apply blast  | 
|
2498  | 
apply (rule independent_injective_image[OF B(2) lf fi])  | 
|
2499  | 
apply (rule order_eq_refl)  | 
|
2500  | 
apply (rule sym)  | 
|
2501  | 
unfolding d  | 
|
2502  | 
apply (rule card_image)  | 
|
2503  | 
apply (rule subset_inj_on[OF fi])  | 
|
2504  | 
by blast  | 
|
2505  | 
from th show ?thesis  | 
|
2506  | 
unfolding span_linear_image[OF lf] surj_def  | 
|
2507  | 
using B(3) by blast  | 
|
2508  | 
qed  | 
|
2509  | 
||
2510  | 
text {* And vice versa. *}
 | 
|
2511  | 
||
2512  | 
lemma surjective_iff_injective_gen:  | 
|
2513  | 
assumes fS: "finite S" and fT: "finite T" and c: "card S = card T"  | 
|
2514  | 
and ST: "f ` S \<subseteq> T"  | 
|
2515  | 
shows "(\<forall>y \<in> T. \<exists>x \<in> S. f x = y) \<longleftrightarrow> inj_on f S" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
2516  | 
proof-  | 
|
2517  | 
  {assume h: "?lhs"
 | 
|
2518  | 
    {fix x y assume x: "x \<in> S" and y: "y \<in> S" and f: "f x = f y"
 | 
|
2519  | 
from x fS have S0: "card S \<noteq> 0" by auto  | 
|
2520  | 
      {assume xy: "x \<noteq> y"
 | 
|
2521  | 
        have th: "card S \<le> card (f ` (S - {y}))"
 | 
|
2522  | 
unfolding c  | 
|
2523  | 
apply (rule card_mono)  | 
|
2524  | 
apply (rule finite_imageI)  | 
|
2525  | 
using fS apply simp  | 
|
2526  | 
using h xy x y f unfolding subset_eq image_iff  | 
|
2527  | 
apply auto  | 
|
2528  | 
apply (case_tac "xa = f x")  | 
|
2529  | 
apply (rule bexI[where x=x])  | 
|
2530  | 
apply auto  | 
|
2531  | 
done  | 
|
2532  | 
        also have " \<dots> \<le> card (S -{y})"
 | 
|
2533  | 
apply (rule card_image_le)  | 
|
2534  | 
using fS by simp  | 
|
2535  | 
also have "\<dots> \<le> card S - 1" using y fS by simp  | 
|
2536  | 
finally have False using S0 by arith }  | 
|
2537  | 
then have "x = y" by blast}  | 
|
2538  | 
then have ?rhs unfolding inj_on_def by blast}  | 
|
2539  | 
moreover  | 
|
2540  | 
  {assume h: ?rhs
 | 
|
2541  | 
have "f ` S = T"  | 
|
2542  | 
apply (rule card_subset_eq[OF fT ST])  | 
|
2543  | 
unfolding card_image[OF h] using c .  | 
|
2544  | 
then have ?lhs by blast}  | 
|
2545  | 
ultimately show ?thesis by blast  | 
|
2546  | 
qed  | 
|
2547  | 
||
2548  | 
lemma linear_surjective_imp_injective: fixes f::"'a::euclidean_space => 'a::euclidean_space"  | 
|
2549  | 
assumes lf: "linear f" and sf: "surj f"  | 
|
2550  | 
shows "inj f"  | 
|
2551  | 
proof-  | 
|
2552  | 
let ?U = "UNIV :: 'a set"  | 
|
2553  | 
from basis_exists[of ?U] obtain B  | 
|
2554  | 
where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" and d: "card B = dim ?U"  | 
|
2555  | 
by blast  | 
|
2556  | 
  {fix x assume x: "x \<in> span B" and fx: "f x = 0"
 | 
|
2557  | 
from B(2) have fB: "finite B" using independent_bound by auto  | 
|
2558  | 
have fBi: "independent (f ` B)"  | 
|
2559  | 
apply (rule card_le_dim_spanning[of "f ` B" ?U])  | 
|
2560  | 
apply blast  | 
|
2561  | 
using sf B(3)  | 
|
2562  | 
unfolding span_linear_image[OF lf] surj_def subset_eq image_iff  | 
|
2563  | 
apply blast  | 
|
2564  | 
using fB apply blast  | 
|
2565  | 
unfolding d[symmetric]  | 
|
2566  | 
apply (rule card_image_le)  | 
|
2567  | 
apply (rule fB)  | 
|
2568  | 
done  | 
|
2569  | 
have th0: "dim ?U \<le> card (f ` B)"  | 
|
2570  | 
apply (rule span_card_ge_dim)  | 
|
2571  | 
apply blast  | 
|
2572  | 
unfolding span_linear_image[OF lf]  | 
|
2573  | 
apply (rule subset_trans[where B = "f ` UNIV"])  | 
|
2574  | 
using sf unfolding surj_def apply blast  | 
|
2575  | 
apply (rule image_mono)  | 
|
2576  | 
apply (rule B(3))  | 
|
2577  | 
apply (metis finite_imageI fB)  | 
|
2578  | 
done  | 
|
2579  | 
||
2580  | 
moreover have "card (f ` B) \<le> card B"  | 
|
2581  | 
by (rule card_image_le, rule fB)  | 
|
2582  | 
ultimately have th1: "card B = card (f ` B)" unfolding d by arith  | 
|
2583  | 
have fiB: "inj_on f B"  | 
|
2584  | 
unfolding surjective_iff_injective_gen[OF fB finite_imageI[OF fB] th1 subset_refl, symmetric] by blast  | 
|
2585  | 
from linear_indep_image_lemma[OF lf fB fBi fiB x] fx  | 
|
2586  | 
have "x = 0" by blast}  | 
|
2587  | 
note th = this  | 
|
2588  | 
from th show ?thesis unfolding linear_injective_0[OF lf]  | 
|
2589  | 
using B(3) by blast  | 
|
2590  | 
qed  | 
|
2591  | 
||
2592  | 
text {* Hence either is enough for isomorphism. *}
 | 
|
2593  | 
||
2594  | 
lemma left_right_inverse_eq:  | 
|
2595  | 
assumes fg: "f o g = id" and gh: "g o h = id"  | 
|
2596  | 
shows "f = h"  | 
|
2597  | 
proof-  | 
|
2598  | 
have "f = f o (g o h)" unfolding gh by simp  | 
|
2599  | 
also have "\<dots> = (f o g) o h" by (simp add: o_assoc)  | 
|
2600  | 
finally show "f = h" unfolding fg by simp  | 
|
2601  | 
qed  | 
|
2602  | 
||
2603  | 
lemma isomorphism_expand:  | 
|
2604  | 
"f o g = id \<and> g o f = id \<longleftrightarrow> (\<forall>x. f(g x) = x) \<and> (\<forall>x. g(f x) = x)"  | 
|
2605  | 
by (simp add: fun_eq_iff o_def id_def)  | 
|
2606  | 
||
2607  | 
lemma linear_injective_isomorphism: fixes f::"'a::euclidean_space => 'a::euclidean_space"  | 
|
2608  | 
assumes lf: "linear f" and fi: "inj f"  | 
|
2609  | 
shows "\<exists>f'. linear f' \<and> (\<forall>x. f' (f x) = x) \<and> (\<forall>x. f (f' x) = x)"  | 
|
2610  | 
unfolding isomorphism_expand[symmetric]  | 
|
2611  | 
using linear_surjective_right_inverse[OF lf linear_injective_imp_surjective[OF lf fi]] linear_injective_left_inverse[OF lf fi]  | 
|
2612  | 
by (metis left_right_inverse_eq)  | 
|
2613  | 
||
2614  | 
lemma linear_surjective_isomorphism: fixes f::"'a::euclidean_space => 'a::euclidean_space"  | 
|
2615  | 
assumes lf: "linear f" and sf: "surj f"  | 
|
2616  | 
shows "\<exists>f'. linear f' \<and> (\<forall>x. f' (f x) = x) \<and> (\<forall>x. f (f' x) = x)"  | 
|
2617  | 
unfolding isomorphism_expand[symmetric]  | 
|
2618  | 
using linear_surjective_right_inverse[OF lf sf] linear_injective_left_inverse[OF lf linear_surjective_imp_injective[OF lf sf]]  | 
|
2619  | 
by (metis left_right_inverse_eq)  | 
|
2620  | 
||
2621  | 
text {* Left and right inverses are the same for @{typ "'a::euclidean_space => 'a::euclidean_space"}. *}
 | 
|
2622  | 
||
2623  | 
lemma linear_inverse_left: fixes f::"'a::euclidean_space => 'a::euclidean_space"  | 
|
2624  | 
assumes lf: "linear f" and lf': "linear f'"  | 
|
2625  | 
shows "f o f' = id \<longleftrightarrow> f' o f = id"  | 
|
2626  | 
proof-  | 
|
2627  | 
  {fix f f':: "'a => 'a"
 | 
|
2628  | 
assume lf: "linear f" "linear f'" and f: "f o f' = id"  | 
|
2629  | 
from f have sf: "surj f"  | 
|
2630  | 
apply (auto simp add: o_def id_def surj_def)  | 
|
2631  | 
by metis  | 
|
2632  | 
from linear_surjective_isomorphism[OF lf(1) sf] lf f  | 
|
2633  | 
have "f' o f = id" unfolding fun_eq_iff o_def id_def  | 
|
2634  | 
by metis}  | 
|
2635  | 
then show ?thesis using lf lf' by metis  | 
|
2636  | 
qed  | 
|
2637  | 
||
2638  | 
text {* Moreover, a one-sided inverse is automatically linear. *}
 | 
|
2639  | 
||
2640  | 
lemma left_inverse_linear: fixes f::"'a::euclidean_space => 'a::euclidean_space"  | 
|
2641  | 
assumes lf: "linear f" and gf: "g o f = id"  | 
|
2642  | 
shows "linear g"  | 
|
2643  | 
proof-  | 
|
2644  | 
from gf have fi: "inj f" apply (auto simp add: inj_on_def o_def id_def fun_eq_iff)  | 
|
2645  | 
by metis  | 
|
2646  | 
from linear_injective_isomorphism[OF lf fi]  | 
|
2647  | 
obtain h:: "'a \<Rightarrow> 'a" where  | 
|
2648  | 
h: "linear h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x" by blast  | 
|
2649  | 
have "h = g" apply (rule ext) using gf h(2,3)  | 
|
2650  | 
apply (simp add: o_def id_def fun_eq_iff)  | 
|
2651  | 
by metis  | 
|
2652  | 
with h(1) show ?thesis by blast  | 
|
2653  | 
qed  | 
|
2654  | 
||
2655  | 
subsection {* Infinity norm *}
 | 
|
2656  | 
||
2657  | 
definition "infnorm (x::'a::euclidean_space) = Sup {abs(x$$i) |i. i<DIM('a)}"
 | 
|
2658  | 
||
2659  | 
lemma numseg_dimindex_nonempty: "\<exists>i. i \<in> (UNIV :: 'n set)"  | 
|
2660  | 
by auto  | 
|
2661  | 
||
2662  | 
lemma infnorm_set_image:  | 
|
2663  | 
  "{abs((x::'a::euclidean_space)$$i) |i. i<DIM('a)} =
 | 
|
2664  | 
  (\<lambda>i. abs(x$$i)) ` {..<DIM('a)}" by blast
 | 
|
2665  | 
||
2666  | 
lemma infnorm_set_lemma:  | 
|
2667  | 
  shows "finite {abs((x::'a::euclidean_space)$$i) |i. i<DIM('a)}"
 | 
|
2668  | 
  and "{abs(x$$i) |i. i<DIM('a::euclidean_space)} \<noteq> {}"
 | 
|
2669  | 
unfolding infnorm_set_image  | 
|
2670  | 
by auto  | 
|
2671  | 
||
2672  | 
lemma infnorm_pos_le: "0 \<le> infnorm (x::'a::euclidean_space)"  | 
|
2673  | 
unfolding infnorm_def  | 
|
2674  | 
unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma]  | 
|
2675  | 
unfolding infnorm_set_image  | 
|
2676  | 
by auto  | 
|
2677  | 
||
2678  | 
lemma infnorm_triangle: "infnorm ((x::'a::euclidean_space) + y) \<le> infnorm x + infnorm y"  | 
|
2679  | 
proof-  | 
|
2680  | 
have th: "\<And>x y (z::real). x - y <= z \<longleftrightarrow> x - z <= y" by arith  | 
|
2681  | 
  have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
 | 
|
2682  | 
have th2: "\<And>x (y::real). abs(x + y) - abs(x) <= abs(y)" by arith  | 
|
2683  | 
  have *:"\<And>i. i \<in> {..<DIM('a)} \<longleftrightarrow> i <DIM('a)" by auto
 | 
|
2684  | 
show ?thesis  | 
|
2685  | 
unfolding infnorm_def unfolding Sup_finite_le_iff[ OF infnorm_set_lemma[where 'a='a]]  | 
|
2686  | 
apply (subst diff_le_eq[symmetric])  | 
|
2687  | 
unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma]  | 
|
2688  | 
unfolding infnorm_set_image bex_simps  | 
|
2689  | 
apply (subst th)  | 
|
2690  | 
unfolding th1 *  | 
|
2691  | 
unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma[where 'a='a]]  | 
|
2692  | 
unfolding infnorm_set_image ball_simps bex_simps  | 
|
2693  | 
unfolding euclidean_simps by (metis th2)  | 
|
2694  | 
qed  | 
|
2695  | 
||
2696  | 
lemma infnorm_eq_0: "infnorm x = 0 \<longleftrightarrow> (x::_::euclidean_space) = 0"  | 
|
2697  | 
proof-  | 
|
2698  | 
have "infnorm x <= 0 \<longleftrightarrow> x = 0"  | 
|
2699  | 
unfolding infnorm_def  | 
|
2700  | 
unfolding Sup_finite_le_iff[OF infnorm_set_lemma]  | 
|
2701  | 
unfolding infnorm_set_image ball_simps  | 
|
| 
44457
 
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
 
huffman 
parents: 
44454 
diff
changeset
 | 
2702  | 
apply(subst (1) euclidean_eq)  | 
| 44133 | 2703  | 
by auto  | 
2704  | 
then show ?thesis using infnorm_pos_le[of x] by simp  | 
|
2705  | 
qed  | 
|
2706  | 
||
2707  | 
lemma infnorm_0: "infnorm 0 = 0"  | 
|
2708  | 
by (simp add: infnorm_eq_0)  | 
|
2709  | 
||
2710  | 
lemma infnorm_neg: "infnorm (- x) = infnorm x"  | 
|
2711  | 
unfolding infnorm_def  | 
|
2712  | 
apply (rule cong[of "Sup" "Sup"])  | 
|
| 
44457
 
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
 
huffman 
parents: 
44454 
diff
changeset
 | 
2713  | 
apply blast by auto  | 
| 44133 | 2714  | 
|
2715  | 
lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)"  | 
|
2716  | 
proof-  | 
|
2717  | 
have "y - x = - (x - y)" by simp  | 
|
2718  | 
then show ?thesis by (metis infnorm_neg)  | 
|
2719  | 
qed  | 
|
2720  | 
||
2721  | 
lemma real_abs_sub_infnorm: "\<bar> infnorm x - infnorm y\<bar> \<le> infnorm (x - y)"  | 
|
2722  | 
proof-  | 
|
2723  | 
have th: "\<And>(nx::real) n ny. nx <= n + ny \<Longrightarrow> ny <= n + nx ==> \<bar>nx - ny\<bar> <= n"  | 
|
2724  | 
by arith  | 
|
2725  | 
from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"]  | 
|
2726  | 
have ths: "infnorm x \<le> infnorm (x - y) + infnorm y"  | 
|
2727  | 
"infnorm y \<le> infnorm (x - y) + infnorm x"  | 
|
| 44454 | 2728  | 
by (simp_all add: field_simps infnorm_neg)  | 
| 44133 | 2729  | 
from th[OF ths] show ?thesis .  | 
2730  | 
qed  | 
|
2731  | 
||
2732  | 
lemma real_abs_infnorm: " \<bar>infnorm x\<bar> = infnorm x"  | 
|
2733  | 
using infnorm_pos_le[of x] by arith  | 
|
2734  | 
||
2735  | 
lemma component_le_infnorm:  | 
|
2736  | 
shows "\<bar>x$$i\<bar> \<le> infnorm (x::'a::euclidean_space)"  | 
|
2737  | 
proof(cases "i<DIM('a)")
 | 
|
2738  | 
case False thus ?thesis using infnorm_pos_le by auto  | 
|
2739  | 
next case True  | 
|
2740  | 
  let ?U = "{..<DIM('a)}"
 | 
|
2741  | 
  let ?S = "{\<bar>x$$i\<bar> |i. i<DIM('a)}"
 | 
|
2742  | 
have fS: "finite ?S" unfolding image_Collect[symmetric]  | 
|
2743  | 
apply (rule finite_imageI) by simp  | 
|
2744  | 
  have S0: "?S \<noteq> {}" by blast
 | 
|
2745  | 
  have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
 | 
|
2746  | 
show ?thesis unfolding infnorm_def  | 
|
2747  | 
apply(subst Sup_finite_ge_iff) using Sup_finite_in[OF fS S0]  | 
|
2748  | 
using infnorm_set_image using True by auto  | 
|
2749  | 
qed  | 
|
2750  | 
||
2751  | 
lemma infnorm_mul_lemma: "infnorm(a *\<^sub>R x) <= \<bar>a\<bar> * infnorm x"  | 
|
2752  | 
apply (subst infnorm_def)  | 
|
2753  | 
unfolding Sup_finite_le_iff[OF infnorm_set_lemma]  | 
|
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
44176 
diff
changeset
 | 
2754  | 
unfolding infnorm_set_image ball_simps euclidean_component_scaleR abs_mult  | 
| 44133 | 2755  | 
using component_le_infnorm[of x] by(auto intro: mult_mono)  | 
2756  | 
||
2757  | 
lemma infnorm_mul: "infnorm(a *\<^sub>R x) = abs a * infnorm x"  | 
|
2758  | 
proof-  | 
|
2759  | 
  {assume a0: "a = 0" hence ?thesis by (simp add: infnorm_0) }
 | 
|
2760  | 
moreover  | 
|
2761  | 
  {assume a0: "a \<noteq> 0"
 | 
|
2762  | 
from a0 have th: "(1/a) *\<^sub>R (a *\<^sub>R x) = x" by simp  | 
|
2763  | 
from a0 have ap: "\<bar>a\<bar> > 0" by arith  | 
|
2764  | 
from infnorm_mul_lemma[of "1/a" "a *\<^sub>R x"]  | 
|
2765  | 
have "infnorm x \<le> 1/\<bar>a\<bar> * infnorm (a*\<^sub>R x)"  | 
|
2766  | 
unfolding th by simp  | 
|
2767  | 
with ap have "\<bar>a\<bar> * infnorm x \<le> \<bar>a\<bar> * (1/\<bar>a\<bar> * infnorm (a *\<^sub>R x))" by (simp add: field_simps)  | 
|
2768  | 
then have "\<bar>a\<bar> * infnorm x \<le> infnorm (a*\<^sub>R x)"  | 
|
2769  | 
using ap by (simp add: field_simps)  | 
|
2770  | 
with infnorm_mul_lemma[of a x] have ?thesis by arith }  | 
|
2771  | 
ultimately show ?thesis by blast  | 
|
2772  | 
qed  | 
|
2773  | 
||
2774  | 
lemma infnorm_pos_lt: "infnorm x > 0 \<longleftrightarrow> x \<noteq> 0"  | 
|
2775  | 
using infnorm_pos_le[of x] infnorm_eq_0[of x] by arith  | 
|
2776  | 
||
2777  | 
text {* Prove that it differs only up to a bound from Euclidean norm. *}
 | 
|
2778  | 
||
2779  | 
lemma infnorm_le_norm: "infnorm x \<le> norm x"  | 
|
2780  | 
unfolding infnorm_def Sup_finite_le_iff[OF infnorm_set_lemma]  | 
|
2781  | 
unfolding infnorm_set_image ball_simps  | 
|
2782  | 
by (metis component_le_norm)  | 
|
2783  | 
||
2784  | 
lemma card_enum: "card {1 .. n} = n" by auto
 | 
|
2785  | 
||
2786  | 
lemma norm_le_infnorm: "norm(x) <= sqrt(real DIM('a)) * infnorm(x::'a::euclidean_space)"
 | 
|
2787  | 
proof-  | 
|
2788  | 
  let ?d = "DIM('a)"
 | 
|
2789  | 
have "real ?d \<ge> 0" by simp  | 
|
2790  | 
hence d2: "(sqrt (real ?d))^2 = real ?d"  | 
|
2791  | 
by (auto intro: real_sqrt_pow2)  | 
|
2792  | 
have th: "sqrt (real ?d) * infnorm x \<ge> 0"  | 
|
2793  | 
by (simp add: zero_le_mult_iff infnorm_pos_le)  | 
|
2794  | 
have th1: "x \<bullet> x \<le> (sqrt (real ?d) * infnorm x)^2"  | 
|
2795  | 
unfolding power_mult_distrib d2  | 
|
2796  | 
unfolding real_of_nat_def apply(subst euclidean_inner)  | 
|
2797  | 
apply (subst power2_abs[symmetric])  | 
|
2798  | 
apply(rule order_trans[OF setsum_bounded[where K="\<bar>infnorm x\<bar>\<twosuperior>"]])  | 
|
2799  | 
apply(auto simp add: power2_eq_square[symmetric])  | 
|
2800  | 
apply (subst power2_abs[symmetric])  | 
|
2801  | 
apply (rule power_mono)  | 
|
2802  | 
unfolding infnorm_def Sup_finite_ge_iff[OF infnorm_set_lemma]  | 
|
2803  | 
unfolding infnorm_set_image bex_simps apply(rule_tac x=i in bexI) by auto  | 
|
2804  | 
from real_le_lsqrt[OF inner_ge_zero th th1]  | 
|
2805  | 
show ?thesis unfolding norm_eq_sqrt_inner id_def .  | 
|
2806  | 
qed  | 
|
2807  | 
||
2808  | 
text {* Equality in Cauchy-Schwarz and triangle inequalities. *}
 | 
|
2809  | 
||
2810  | 
lemma norm_cauchy_schwarz_eq: "x \<bullet> y = norm x * norm y \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
2811  | 
proof-  | 
|
2812  | 
  {assume h: "x = 0"
 | 
|
2813  | 
hence ?thesis by simp}  | 
|
2814  | 
moreover  | 
|
2815  | 
  {assume h: "y = 0"
 | 
|
2816  | 
hence ?thesis by simp}  | 
|
2817  | 
moreover  | 
|
2818  | 
  {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
 | 
|
2819  | 
from inner_eq_zero_iff[of "norm y *\<^sub>R x - norm x *\<^sub>R y"]  | 
|
2820  | 
have "?rhs \<longleftrightarrow> (norm y * (norm y * norm x * norm x - norm x * (x \<bullet> y)) - norm x * (norm y * (y \<bullet> x) - norm x * norm y * norm y) = 0)"  | 
|
2821  | 
using x y  | 
|
2822  | 
unfolding inner_simps  | 
|
2823  | 
unfolding power2_norm_eq_inner[symmetric] power2_eq_square diff_eq_0_iff_eq apply (simp add: inner_commute)  | 
|
2824  | 
apply (simp add: field_simps) by metis  | 
|
2825  | 
also have "\<dots> \<longleftrightarrow> (2 * norm x * norm y * (norm x * norm y - x \<bullet> y) = 0)" using x y  | 
|
2826  | 
by (simp add: field_simps inner_commute)  | 
|
2827  | 
also have "\<dots> \<longleftrightarrow> ?lhs" using x y  | 
|
2828  | 
apply simp  | 
|
2829  | 
by metis  | 
|
2830  | 
finally have ?thesis by blast}  | 
|
2831  | 
ultimately show ?thesis by blast  | 
|
2832  | 
qed  | 
|
2833  | 
||
2834  | 
lemma norm_cauchy_schwarz_abs_eq:  | 
|
2835  | 
shows "abs(x \<bullet> y) = norm x * norm y \<longleftrightarrow>  | 
|
2836  | 
norm x *\<^sub>R y = norm y *\<^sub>R x \<or> norm(x) *\<^sub>R y = - norm y *\<^sub>R x" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
2837  | 
proof-  | 
|
2838  | 
have th: "\<And>(x::real) a. a \<ge> 0 \<Longrightarrow> abs x = a \<longleftrightarrow> x = a \<or> x = - a" by arith  | 
|
2839  | 
have "?rhs \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x \<or> norm (- x) *\<^sub>R y = norm y *\<^sub>R (- x)"  | 
|
2840  | 
by simp  | 
|
2841  | 
also have "\<dots> \<longleftrightarrow>(x \<bullet> y = norm x * norm y \<or>  | 
|
2842  | 
(-x) \<bullet> y = norm x * norm y)"  | 
|
2843  | 
unfolding norm_cauchy_schwarz_eq[symmetric]  | 
|
2844  | 
unfolding norm_minus_cancel norm_scaleR ..  | 
|
2845  | 
also have "\<dots> \<longleftrightarrow> ?lhs"  | 
|
2846  | 
unfolding th[OF mult_nonneg_nonneg, OF norm_ge_zero[of x] norm_ge_zero[of y]] inner_simps by auto  | 
|
2847  | 
finally show ?thesis ..  | 
|
2848  | 
qed  | 
|
2849  | 
||
2850  | 
lemma norm_triangle_eq:  | 
|
2851  | 
fixes x y :: "'a::real_inner"  | 
|
2852  | 
shows "norm(x + y) = norm x + norm y \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x"  | 
|
2853  | 
proof-  | 
|
2854  | 
  {assume x: "x =0 \<or> y =0"
 | 
|
2855  | 
hence ?thesis by (cases "x=0", simp_all)}  | 
|
2856  | 
moreover  | 
|
2857  | 
  {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
 | 
|
2858  | 
hence "norm x \<noteq> 0" "norm y \<noteq> 0"  | 
|
2859  | 
by simp_all  | 
|
2860  | 
hence n: "norm x > 0" "norm y > 0"  | 
|
2861  | 
using norm_ge_zero[of x] norm_ge_zero[of y]  | 
|
2862  | 
by arith+  | 
|
2863  | 
have th: "\<And>(a::real) b c. a + b + c \<noteq> 0 ==> (a = b + c \<longleftrightarrow> a^2 = (b + c)^2)" by algebra  | 
|
2864  | 
have "norm(x + y) = norm x + norm y \<longleftrightarrow> norm(x + y)^ 2 = (norm x + norm y) ^2"  | 
|
2865  | 
apply (rule th) using n norm_ge_zero[of "x + y"]  | 
|
2866  | 
by arith  | 
|
2867  | 
also have "\<dots> \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x"  | 
|
2868  | 
unfolding norm_cauchy_schwarz_eq[symmetric]  | 
|
2869  | 
unfolding power2_norm_eq_inner inner_simps  | 
|
2870  | 
by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps)  | 
|
2871  | 
finally have ?thesis .}  | 
|
2872  | 
ultimately show ?thesis by blast  | 
|
2873  | 
qed  | 
|
2874  | 
||
2875  | 
subsection {* Collinearity *}
 | 
|
2876  | 
||
2877  | 
definition  | 
|
2878  | 
collinear :: "'a::real_vector set \<Rightarrow> bool" where  | 
|
2879  | 
"collinear S \<longleftrightarrow> (\<exists>u. \<forall>x \<in> S. \<forall> y \<in> S. \<exists>c. x - y = c *\<^sub>R u)"  | 
|
2880  | 
||
2881  | 
lemma collinear_empty:  "collinear {}" by (simp add: collinear_def)
 | 
|
2882  | 
||
2883  | 
lemma collinear_sing: "collinear {x}"
 | 
|
2884  | 
by (simp add: collinear_def)  | 
|
2885  | 
||
2886  | 
lemma collinear_2: "collinear {x, y}"
 | 
|
2887  | 
apply (simp add: collinear_def)  | 
|
2888  | 
apply (rule exI[where x="x - y"])  | 
|
2889  | 
apply auto  | 
|
2890  | 
apply (rule exI[where x=1], simp)  | 
|
2891  | 
apply (rule exI[where x="- 1"], simp)  | 
|
2892  | 
done  | 
|
2893  | 
||
2894  | 
lemma collinear_lemma: "collinear {0,x,y} \<longleftrightarrow> x = 0 \<or> y = 0 \<or> (\<exists>c. y = c *\<^sub>R x)" (is "?lhs \<longleftrightarrow> ?rhs")
 | 
|
2895  | 
proof-  | 
|
2896  | 
  {assume "x=0 \<or> y = 0" hence ?thesis
 | 
|
2897  | 
by (cases "x = 0", simp_all add: collinear_2 insert_commute)}  | 
|
2898  | 
moreover  | 
|
2899  | 
  {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
 | 
|
2900  | 
    {assume h: "?lhs"
 | 
|
2901  | 
      then obtain u where u: "\<forall> x\<in> {0,x,y}. \<forall>y\<in> {0,x,y}. \<exists>c. x - y = c *\<^sub>R u" unfolding collinear_def by blast
 | 
|
2902  | 
from u[rule_format, of x 0] u[rule_format, of y 0]  | 
|
2903  | 
obtain cx and cy where  | 
|
2904  | 
cx: "x = cx *\<^sub>R u" and cy: "y = cy *\<^sub>R u"  | 
|
2905  | 
by auto  | 
|
2906  | 
from cx x have cx0: "cx \<noteq> 0" by auto  | 
|
2907  | 
from cy y have cy0: "cy \<noteq> 0" by auto  | 
|
2908  | 
let ?d = "cy / cx"  | 
|
2909  | 
from cx cy cx0 have "y = ?d *\<^sub>R x"  | 
|
2910  | 
by simp  | 
|
2911  | 
hence ?rhs using x y by blast}  | 
|
2912  | 
moreover  | 
|
2913  | 
    {assume h: "?rhs"
 | 
|
2914  | 
then obtain c where c: "y = c *\<^sub>R x" using x y by blast  | 
|
2915  | 
have ?lhs unfolding collinear_def c  | 
|
2916  | 
apply (rule exI[where x=x])  | 
|
2917  | 
apply auto  | 
|
2918  | 
apply (rule exI[where x="- 1"], simp)  | 
|
2919  | 
apply (rule exI[where x= "-c"], simp)  | 
|
2920  | 
apply (rule exI[where x=1], simp)  | 
|
2921  | 
apply (rule exI[where x="1 - c"], simp add: scaleR_left_diff_distrib)  | 
|
2922  | 
apply (rule exI[where x="c - 1"], simp add: scaleR_left_diff_distrib)  | 
|
2923  | 
done}  | 
|
2924  | 
ultimately have ?thesis by blast}  | 
|
2925  | 
ultimately show ?thesis by blast  | 
|
2926  | 
qed  | 
|
2927  | 
||
2928  | 
lemma norm_cauchy_schwarz_equal:  | 
|
2929  | 
  shows "abs(x \<bullet> y) = norm x * norm y \<longleftrightarrow> collinear {0,x,y}"
 | 
|
2930  | 
unfolding norm_cauchy_schwarz_abs_eq  | 
|
2931  | 
apply (cases "x=0", simp_all add: collinear_2)  | 
|
2932  | 
apply (cases "y=0", simp_all add: collinear_2 insert_commute)  | 
|
2933  | 
unfolding collinear_lemma  | 
|
2934  | 
apply simp  | 
|
2935  | 
apply (subgoal_tac "norm x \<noteq> 0")  | 
|
2936  | 
apply (subgoal_tac "norm y \<noteq> 0")  | 
|
2937  | 
apply (rule iffI)  | 
|
2938  | 
apply (cases "norm x *\<^sub>R y = norm y *\<^sub>R x")  | 
|
2939  | 
apply (rule exI[where x="(1/norm x) * norm y"])  | 
|
2940  | 
apply (drule sym)  | 
|
2941  | 
unfolding scaleR_scaleR[symmetric]  | 
|
2942  | 
apply (simp add: field_simps)  | 
|
2943  | 
apply (rule exI[where x="(1/norm x) * - norm y"])  | 
|
2944  | 
apply clarify  | 
|
2945  | 
apply (drule sym)  | 
|
2946  | 
unfolding scaleR_scaleR[symmetric]  | 
|
2947  | 
apply (simp add: field_simps)  | 
|
2948  | 
apply (erule exE)  | 
|
2949  | 
apply (erule ssubst)  | 
|
2950  | 
unfolding scaleR_scaleR  | 
|
2951  | 
unfolding norm_scaleR  | 
|
2952  | 
apply (subgoal_tac "norm x * c = \<bar>c\<bar> * norm x \<or> norm x * c = - \<bar>c\<bar> * norm x")  | 
|
2953  | 
apply (case_tac "c <= 0", simp add: field_simps)  | 
|
2954  | 
apply (simp add: field_simps)  | 
|
2955  | 
apply (case_tac "c <= 0", simp add: field_simps)  | 
|
2956  | 
apply (simp add: field_simps)  | 
|
2957  | 
apply simp  | 
|
2958  | 
apply simp  | 
|
2959  | 
done  | 
|
2960  | 
||
2961  | 
subsection "Instantiate @{typ real} and @{typ complex} as typeclass @{text ordered_euclidean_space}."
 | 
|
2962  | 
||
2963  | 
lemma basis_real_range: "basis ` {..<1} = {1::real}" by auto
 | 
|
2964  | 
||
2965  | 
instance real::ordered_euclidean_space  | 
|
2966  | 
by default (auto simp add: euclidean_component_def)  | 
|
2967  | 
||
2968  | 
lemma Eucl_real_simps[simp]:  | 
|
2969  | 
"(x::real) $$ 0 = x"  | 
|
2970  | 
"(\<chi>\<chi> i. f i) = ((f 0)::real)"  | 
|
2971  | 
"\<And>i. i > 0 \<Longrightarrow> x $$ i = 0"  | 
|
2972  | 
defer apply(subst euclidean_eq) apply safe  | 
|
2973  | 
unfolding euclidean_lambda_beta'  | 
|
2974  | 
unfolding euclidean_component_def by auto  | 
|
2975  | 
||
2976  | 
lemma complex_basis[simp]:  | 
|
2977  | 
shows "basis 0 = (1::complex)" and "basis 1 = ii" and "basis (Suc 0) = ii"  | 
|
2978  | 
unfolding basis_complex_def by auto  | 
|
2979  | 
||
| 44360 | 2980  | 
subsection {* Products Spaces *}
 | 
| 44133 | 2981  | 
|
2982  | 
lemma DIM_prod[simp]: "DIM('a \<times> 'b) = DIM('b::euclidean_space) + DIM('a::euclidean_space)"
 | 
|
2983  | 
  (* FIXME: why this orientation? Why not "DIM('a) + DIM('b)" ? *)
 | 
|
2984  | 
unfolding dimension_prod_def by (rule add_commute)  | 
|
2985  | 
||
2986  | 
instantiation prod :: (ordered_euclidean_space, ordered_euclidean_space) ordered_euclidean_space  | 
|
2987  | 
begin  | 
|
2988  | 
||
2989  | 
definition "x \<le> (y::('a\<times>'b)) \<longleftrightarrow> (\<forall>i<DIM('a\<times>'b). x $$ i \<le> y $$ i)"
 | 
|
2990  | 
definition "x < (y::('a\<times>'b)) \<longleftrightarrow> (\<forall>i<DIM('a\<times>'b). x $$ i < y $$ i)"
 | 
|
2991  | 
||
2992  | 
instance proof qed (auto simp: less_prod_def less_eq_prod_def)  | 
|
2993  | 
end  | 
|
2994  | 
||
2995  | 
||
2996  | 
end  |