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