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