| author | haftmann | 
| Wed, 24 Feb 2010 14:19:52 +0100 | |
| changeset 35366 | 6d474096698c | 
| parent 35253 | 68dd8b51c6b8 | 
| child 35540 | 3d073a3e1c61 | 
| permissions | -rw-r--r-- | 
| 35253 | 1  | 
(* Title: Library/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  | 
|
9  | 
Complex_Main "~~/src/HOL/Decision_Procs/Dense_Linear_Order"  | 
|
10  | 
Finite_Cartesian_Product Glbs Infinite_Set Numeral_Type  | 
|
11  | 
Inner_Product  | 
|
12  | 
uses "positivstellensatz.ML" ("normarith.ML")
 | 
|
13  | 
begin  | 
|
14  | 
||
15  | 
text{* Some common special cases.*}
 | 
|
16  | 
||
| 34964 | 17  | 
lemma forall_1[simp]: "(\<forall>i::1. P i) \<longleftrightarrow> P 1"  | 
| 33175 | 18  | 
by (metis num1_eq_iff)  | 
19  | 
||
| 34964 | 20  | 
lemma ex_1[simp]: "(\<exists>x::1. P x) \<longleftrightarrow> P 1"  | 
21  | 
by auto (metis num1_eq_iff)  | 
|
22  | 
||
| 33175 | 23  | 
lemma exhaust_2:  | 
24  | 
fixes x :: 2 shows "x = 1 \<or> x = 2"  | 
|
25  | 
proof (induct x)  | 
|
26  | 
case (of_int z)  | 
|
27  | 
then have "0 <= z" and "z < 2" by simp_all  | 
|
28  | 
then have "z = 0 | z = 1" by arith  | 
|
29  | 
then show ?case by auto  | 
|
30  | 
qed  | 
|
31  | 
||
32  | 
lemma forall_2: "(\<forall>i::2. P i) \<longleftrightarrow> P 1 \<and> P 2"  | 
|
33  | 
by (metis exhaust_2)  | 
|
34  | 
||
35  | 
lemma exhaust_3:  | 
|
36  | 
fixes x :: 3 shows "x = 1 \<or> x = 2 \<or> x = 3"  | 
|
37  | 
proof (induct x)  | 
|
38  | 
case (of_int z)  | 
|
39  | 
then have "0 <= z" and "z < 3" by simp_all  | 
|
40  | 
then have "z = 0 \<or> z = 1 \<or> z = 2" by arith  | 
|
41  | 
then show ?case by auto  | 
|
42  | 
qed  | 
|
43  | 
||
44  | 
lemma forall_3: "(\<forall>i::3. P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3"  | 
|
45  | 
by (metis exhaust_3)  | 
|
46  | 
||
47  | 
lemma UNIV_1: "UNIV = {1::1}"
 | 
|
48  | 
by (auto simp add: num1_eq_iff)  | 
|
49  | 
||
50  | 
lemma UNIV_2: "UNIV = {1::2, 2::2}"
 | 
|
51  | 
using exhaust_2 by auto  | 
|
52  | 
||
53  | 
lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}"
 | 
|
54  | 
using exhaust_3 by auto  | 
|
55  | 
||
56  | 
lemma setsum_1: "setsum f (UNIV::1 set) = f 1"  | 
|
57  | 
unfolding UNIV_1 by simp  | 
|
58  | 
||
59  | 
lemma setsum_2: "setsum f (UNIV::2 set) = f 1 + f 2"  | 
|
60  | 
unfolding UNIV_2 by simp  | 
|
61  | 
||
62  | 
lemma setsum_3: "setsum f (UNIV::3 set) = f 1 + f 2 + f 3"  | 
|
63  | 
unfolding UNIV_3 by (simp add: add_ac)  | 
|
64  | 
||
65  | 
subsection{* Basic componentwise operations on vectors. *}
 | 
|
66  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
67  | 
instantiation cart :: (plus,finite) plus  | 
| 33175 | 68  | 
begin  | 
| 35253 | 69  | 
definition vector_add_def : "op + \<equiv> (\<lambda> x y. (\<chi> i. (x$i) + (y$i)))"  | 
70  | 
instance ..  | 
|
| 33175 | 71  | 
end  | 
72  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
73  | 
instantiation cart :: (times,finite) times  | 
| 33175 | 74  | 
begin  | 
75  | 
definition vector_mult_def : "op * \<equiv> (\<lambda> x y. (\<chi> i. (x$i) * (y$i)))"  | 
|
76  | 
instance ..  | 
|
77  | 
end  | 
|
78  | 
||
| 35253 | 79  | 
instantiation cart :: (minus,finite) minus  | 
80  | 
begin  | 
|
| 33175 | 81  | 
definition vector_minus_def : "op - \<equiv> (\<lambda> x y. (\<chi> i. (x$i) - (y$i)))"  | 
| 35253 | 82  | 
instance ..  | 
| 33175 | 83  | 
end  | 
84  | 
||
| 35253 | 85  | 
instantiation cart :: (uminus,finite) uminus  | 
86  | 
begin  | 
|
| 33175 | 87  | 
definition vector_uminus_def : "uminus \<equiv> (\<lambda> x. (\<chi> i. - (x$i)))"  | 
| 35253 | 88  | 
instance ..  | 
| 33175 | 89  | 
end  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
90  | 
|
| 35253 | 91  | 
instantiation cart :: (zero,finite) zero  | 
92  | 
begin  | 
|
| 33175 | 93  | 
definition vector_zero_def : "0 \<equiv> (\<chi> i. 0)"  | 
| 35253 | 94  | 
instance ..  | 
| 33175 | 95  | 
end  | 
96  | 
||
| 35253 | 97  | 
instantiation cart :: (one,finite) one  | 
98  | 
begin  | 
|
| 33175 | 99  | 
definition vector_one_def : "1 \<equiv> (\<chi> i. 1)"  | 
| 35253 | 100  | 
instance ..  | 
| 33175 | 101  | 
end  | 
102  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
103  | 
instantiation cart :: (ord,finite) ord  | 
| 35253 | 104  | 
begin  | 
105  | 
definition vector_le_def:  | 
|
106  | 
"less_eq (x :: 'a ^'b) y = (ALL i. x$i <= y$i)"  | 
|
107  | 
definition vector_less_def: "less (x :: 'a ^'b) y = (ALL i. x$i < y$i)"  | 
|
108  | 
instance by (intro_classes)  | 
|
| 33175 | 109  | 
end  | 
110  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
111  | 
instantiation cart :: (scaleR, finite) scaleR  | 
| 33175 | 112  | 
begin  | 
| 35253 | 113  | 
definition vector_scaleR_def: "scaleR = (\<lambda> r x. (\<chi> i. scaleR r (x$i)))"  | 
114  | 
instance ..  | 
|
| 33175 | 115  | 
end  | 
116  | 
||
117  | 
text{* Also the scalar-vector multiplication. *}
 | 
|
118  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
119  | 
definition vector_scalar_mult:: "'a::times \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a ^ 'n" (infixl "*s" 70)  | 
| 33175 | 120  | 
where "c *s x = (\<chi> i. c * (x$i))"  | 
121  | 
||
122  | 
text{* Constant Vectors *} 
 | 
|
123  | 
||
124  | 
definition "vec x = (\<chi> i. x)"  | 
|
125  | 
||
126  | 
text{* Dot products. *}
 | 
|
127  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
128  | 
definition dot :: "'a::{comm_monoid_add, times} ^ 'n \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a" (infix "\<bullet>" 70) where
 | 
| 33175 | 129  | 
"x \<bullet> y = setsum (\<lambda>i. x$i * y$i) UNIV"  | 
130  | 
||
131  | 
lemma dot_1[simp]: "(x::'a::{comm_monoid_add, times}^1) \<bullet> y = (x$1) * (y$1)"
 | 
|
132  | 
by (simp add: dot_def setsum_1)  | 
|
133  | 
||
134  | 
lemma dot_2[simp]: "(x::'a::{comm_monoid_add, times}^2) \<bullet> y = (x$1) * (y$1) + (x$2) * (y$2)"
 | 
|
135  | 
by (simp add: dot_def setsum_2)  | 
|
136  | 
||
137  | 
lemma dot_3[simp]: "(x::'a::{comm_monoid_add, times}^3) \<bullet> y = (x$1) * (y$1) + (x$2) * (y$2) + (x$3) * (y$3)"
 | 
|
138  | 
by (simp add: dot_def setsum_3)  | 
|
139  | 
||
140  | 
subsection {* A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space. *}
 | 
|
141  | 
||
142  | 
method_setup vector = {*
 | 
|
143  | 
let  | 
|
144  | 
  val ss1 = HOL_basic_ss addsimps [@{thm dot_def}, @{thm setsum_addf} RS sym,
 | 
|
145  | 
  @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib},
 | 
|
146  | 
  @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym]
 | 
|
147  | 
  val ss2 = @{simpset} addsimps
 | 
|
148  | 
             [@{thm vector_add_def}, @{thm vector_mult_def},
 | 
|
149  | 
              @{thm vector_minus_def}, @{thm vector_uminus_def},
 | 
|
150  | 
              @{thm vector_one_def}, @{thm vector_zero_def}, @{thm vec_def},
 | 
|
151  | 
              @{thm vector_scaleR_def},
 | 
|
152  | 
              @{thm Cart_lambda_beta}, @{thm vector_scalar_mult_def}]
 | 
|
153  | 
fun vector_arith_tac ths =  | 
|
154  | 
simp_tac ss1  | 
|
155  | 
   THEN' (fn i => rtac @{thm setsum_cong2} i
 | 
|
156  | 
         ORELSE rtac @{thm setsum_0'} i
 | 
|
157  | 
         ORELSE simp_tac (HOL_basic_ss addsimps [@{thm "Cart_eq"}]) i)
 | 
|
158  | 
   (* THEN' TRY o clarify_tac HOL_cs  THEN' (TRY o rtac @{thm iffI}) *)
 | 
|
159  | 
THEN' asm_full_simp_tac (ss2 addsimps ths)  | 
|
160  | 
in  | 
|
161  | 
Attrib.thms >> (fn ths => K (SIMPLE_METHOD' (vector_arith_tac ths)))  | 
|
162  | 
end  | 
|
163  | 
*} "Lifts trivial vector statements to real arith statements"  | 
|
164  | 
||
165  | 
lemma vec_0[simp]: "vec 0 = 0" by (vector vector_zero_def)  | 
|
166  | 
lemma vec_1[simp]: "vec 1 = 1" by (vector vector_one_def)  | 
|
167  | 
||
168  | 
||
169  | 
||
170  | 
text{* Obvious "component-pushing". *}
 | 
|
171  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
172  | 
lemma vec_component [simp]: "vec x $ i = x"  | 
| 33175 | 173  | 
by (vector vec_def)  | 
174  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
175  | 
lemma vector_add_component [simp]: "(x + y)$i = x$i + y$i"  | 
| 33175 | 176  | 
by vector  | 
177  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
178  | 
lemma vector_minus_component [simp]: "(x - y)$i = x$i - y$i"  | 
| 33175 | 179  | 
by vector  | 
180  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
181  | 
lemma vector_mult_component [simp]: "(x * y)$i = x$i * y$i"  | 
| 33175 | 182  | 
by vector  | 
183  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
184  | 
lemma vector_smult_component [simp]: "(c *s y)$i = c * (y$i)"  | 
| 33175 | 185  | 
by vector  | 
186  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
187  | 
lemma vector_uminus_component [simp]: "(- x)$i = - (x$i)"  | 
| 33175 | 188  | 
by vector  | 
189  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
190  | 
lemma vector_scaleR_component [simp]: "(scaleR r x)$i = scaleR r (x$i)"  | 
| 33175 | 191  | 
by vector  | 
192  | 
||
193  | 
lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector  | 
|
194  | 
||
195  | 
lemmas vector_component =  | 
|
196  | 
vec_component vector_add_component vector_mult_component  | 
|
197  | 
vector_smult_component vector_minus_component vector_uminus_component  | 
|
198  | 
vector_scaleR_component cond_component  | 
|
199  | 
||
200  | 
subsection {* Some frequently useful arithmetic lemmas over vectors. *}
 | 
|
201  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
202  | 
instance cart :: (semigroup_add,finite) semigroup_add  | 
| 33175 | 203  | 
apply (intro_classes) by (vector add_assoc)  | 
204  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
205  | 
instance cart :: (monoid_add,finite) monoid_add  | 
| 33175 | 206  | 
apply (intro_classes) by vector+  | 
207  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
208  | 
instance cart :: (group_add,finite) group_add  | 
| 33175 | 209  | 
apply (intro_classes) by (vector algebra_simps)+  | 
210  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
211  | 
instance cart :: (ab_semigroup_add,finite) ab_semigroup_add  | 
| 33175 | 212  | 
apply (intro_classes) by (vector add_commute)  | 
213  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
214  | 
instance cart :: (comm_monoid_add,finite) comm_monoid_add  | 
| 33175 | 215  | 
apply (intro_classes) by vector  | 
216  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
217  | 
instance cart :: (ab_group_add,finite) ab_group_add  | 
| 33175 | 218  | 
apply (intro_classes) by vector+  | 
219  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
220  | 
instance cart :: (cancel_semigroup_add,finite) cancel_semigroup_add  | 
| 33175 | 221  | 
apply (intro_classes)  | 
222  | 
by (vector Cart_eq)+  | 
|
223  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
224  | 
instance cart :: (cancel_ab_semigroup_add,finite) cancel_ab_semigroup_add  | 
| 33175 | 225  | 
apply (intro_classes)  | 
226  | 
by (vector Cart_eq)  | 
|
227  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
228  | 
instance cart :: (real_vector, finite) real_vector  | 
| 33175 | 229  | 
by default (vector scaleR_left_distrib scaleR_right_distrib)+  | 
230  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
231  | 
instance cart :: (semigroup_mult,finite) semigroup_mult  | 
| 33175 | 232  | 
apply (intro_classes) by (vector mult_assoc)  | 
233  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
234  | 
instance cart :: (monoid_mult,finite) monoid_mult  | 
| 33175 | 235  | 
apply (intro_classes) by vector+  | 
236  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
237  | 
instance cart :: (ab_semigroup_mult,finite) ab_semigroup_mult  | 
| 33175 | 238  | 
apply (intro_classes) by (vector mult_commute)  | 
239  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
240  | 
instance cart :: (ab_semigroup_idem_mult,finite) ab_semigroup_idem_mult  | 
| 33175 | 241  | 
apply (intro_classes) by (vector mult_idem)  | 
242  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
243  | 
instance cart :: (comm_monoid_mult,finite) comm_monoid_mult  | 
| 33175 | 244  | 
apply (intro_classes) by vector  | 
245  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
246  | 
fun vector_power where  | 
| 33175 | 247  | 
"vector_power x 0 = 1"  | 
248  | 
| "vector_power x (Suc n) = x * vector_power x n"  | 
|
249  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
250  | 
instance cart :: (semiring,finite) semiring  | 
| 33175 | 251  | 
apply (intro_classes) by (vector ring_simps)+  | 
252  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
253  | 
instance cart :: (semiring_0,finite) semiring_0  | 
| 33175 | 254  | 
apply (intro_classes) by (vector ring_simps)+  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
255  | 
instance cart :: (semiring_1,finite) semiring_1  | 
| 33175 | 256  | 
apply (intro_classes) by vector  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
257  | 
instance cart :: (comm_semiring,finite) comm_semiring  | 
| 33175 | 258  | 
apply (intro_classes) by (vector ring_simps)+  | 
259  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
260  | 
instance cart :: (comm_semiring_0,finite) comm_semiring_0 by (intro_classes)  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
261  | 
instance cart :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add ..  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
262  | 
instance cart :: (semiring_0_cancel,finite) semiring_0_cancel by (intro_classes)  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
263  | 
instance cart :: (comm_semiring_0_cancel,finite) comm_semiring_0_cancel by (intro_classes)  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
264  | 
instance cart :: (ring,finite) ring by (intro_classes)  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
265  | 
instance cart :: (semiring_1_cancel,finite) semiring_1_cancel by (intro_classes)  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
266  | 
instance cart :: (comm_semiring_1,finite) comm_semiring_1 by (intro_classes)  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
267  | 
|
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
268  | 
instance cart :: (ring_1,finite) ring_1 ..  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
269  | 
|
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
270  | 
instance cart :: (real_algebra,finite) real_algebra  | 
| 33175 | 271  | 
apply intro_classes  | 
272  | 
apply (simp_all add: vector_scaleR_def ring_simps)  | 
|
273  | 
apply vector  | 
|
274  | 
apply vector  | 
|
275  | 
done  | 
|
276  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
277  | 
instance cart :: (real_algebra_1,finite) real_algebra_1 ..  | 
| 33175 | 278  | 
|
279  | 
lemma of_nat_index:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
280  | 
"(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n"  | 
| 33175 | 281  | 
apply (induct n)  | 
282  | 
apply vector  | 
|
283  | 
apply vector  | 
|
284  | 
done  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
285  | 
|
| 33175 | 286  | 
lemma zero_index[simp]:  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
287  | 
"(0 :: 'a::zero ^'n)$i = 0" by vector  | 
| 33175 | 288  | 
|
289  | 
lemma one_index[simp]:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
290  | 
"(1 :: 'a::one ^'n)$i = 1" by vector  | 
| 33175 | 291  | 
|
292  | 
lemma one_plus_of_nat_neq_0: "(1::'a::semiring_char_0) + of_nat n \<noteq> 0"  | 
|
293  | 
proof-  | 
|
294  | 
have "(1::'a) + of_nat n = 0 \<longleftrightarrow> of_nat 1 + of_nat n = (of_nat 0 :: 'a)" by simp  | 
|
295  | 
also have "\<dots> \<longleftrightarrow> 1 + n = 0" by (simp only: of_nat_add[symmetric] of_nat_eq_iff)  | 
|
296  | 
finally show ?thesis by simp  | 
|
297  | 
qed  | 
|
298  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
299  | 
instance cart :: (semiring_char_0,finite) semiring_char_0  | 
| 33175 | 300  | 
proof (intro_classes)  | 
301  | 
fix m n ::nat  | 
|
302  | 
show "(of_nat m :: 'a^'b) = of_nat n \<longleftrightarrow> m = n"  | 
|
303  | 
by (simp add: Cart_eq of_nat_index)  | 
|
304  | 
qed  | 
|
305  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
306  | 
instance cart :: (comm_ring_1,finite) comm_ring_1 by intro_classes  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
307  | 
instance cart :: (ring_char_0,finite) ring_char_0 by intro_classes  | 
| 33175 | 308  | 
|
309  | 
lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x"  | 
|
310  | 
by (vector mult_assoc)  | 
|
311  | 
lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x"  | 
|
312  | 
by (vector ring_simps)  | 
|
313  | 
lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y"  | 
|
314  | 
by (vector ring_simps)  | 
|
315  | 
lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector  | 
|
316  | 
lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector  | 
|
317  | 
lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y"  | 
|
318  | 
by (vector ring_simps)  | 
|
319  | 
lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector  | 
|
320  | 
lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector  | 
|
321  | 
lemma vector_sneg_minus1: "-x = (- (1::'a::ring_1)) *s x" by vector  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
322  | 
lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector  | 
| 33175 | 323  | 
lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x"  | 
324  | 
by (vector ring_simps)  | 
|
325  | 
||
326  | 
lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)"  | 
|
327  | 
by (simp add: Cart_eq)  | 
|
328  | 
||
329  | 
subsection {* Topological space *}
 | 
|
330  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
331  | 
instantiation cart :: (topological_space, finite) topological_space  | 
| 33175 | 332  | 
begin  | 
333  | 
||
334  | 
definition open_vector_def:  | 
|
335  | 
  "open (S :: ('a ^ 'b) set) \<longleftrightarrow>
 | 
|
336  | 
(\<forall>x\<in>S. \<exists>A. (\<forall>i. open (A i) \<and> x$i \<in> A i) \<and>  | 
|
337  | 
(\<forall>y. (\<forall>i. y$i \<in> A i) \<longrightarrow> y \<in> S))"  | 
|
338  | 
||
339  | 
instance proof  | 
|
340  | 
  show "open (UNIV :: ('a ^ 'b) set)"
 | 
|
341  | 
unfolding open_vector_def by auto  | 
|
342  | 
next  | 
|
343  | 
  fix S T :: "('a ^ 'b) set"
 | 
|
344  | 
assume "open S" "open T" thus "open (S \<inter> T)"  | 
|
345  | 
unfolding open_vector_def  | 
|
346  | 
apply clarify  | 
|
347  | 
apply (drule (1) bspec)+  | 
|
348  | 
apply (clarify, rename_tac Sa Ta)  | 
|
349  | 
apply (rule_tac x="\<lambda>i. Sa i \<inter> Ta i" in exI)  | 
|
350  | 
apply (simp add: open_Int)  | 
|
351  | 
done  | 
|
352  | 
next  | 
|
353  | 
  fix K :: "('a ^ 'b) set set"
 | 
|
354  | 
assume "\<forall>S\<in>K. open S" thus "open (\<Union>K)"  | 
|
355  | 
unfolding open_vector_def  | 
|
356  | 
apply clarify  | 
|
357  | 
apply (drule (1) bspec)  | 
|
358  | 
apply (drule (1) bspec)  | 
|
359  | 
apply clarify  | 
|
360  | 
apply (rule_tac x=A in exI)  | 
|
361  | 
apply fast  | 
|
362  | 
done  | 
|
363  | 
qed  | 
|
364  | 
||
365  | 
end  | 
|
366  | 
||
367  | 
lemma open_vector_box: "\<forall>i. open (S i) \<Longrightarrow> open {x. \<forall>i. x $ i \<in> S i}"
 | 
|
368  | 
unfolding open_vector_def by auto  | 
|
369  | 
||
370  | 
lemma open_vimage_Cart_nth: "open S \<Longrightarrow> open ((\<lambda>x. x $ i) -` S)"  | 
|
371  | 
unfolding open_vector_def  | 
|
372  | 
apply clarify  | 
|
373  | 
apply (rule_tac x="\<lambda>k. if k = i then S else UNIV" in exI, simp)  | 
|
374  | 
done  | 
|
375  | 
||
376  | 
lemma closed_vimage_Cart_nth: "closed S \<Longrightarrow> closed ((\<lambda>x. x $ i) -` S)"  | 
|
377  | 
unfolding closed_open vimage_Compl [symmetric]  | 
|
378  | 
by (rule open_vimage_Cart_nth)  | 
|
379  | 
||
380  | 
lemma closed_vector_box: "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}"
 | 
|
381  | 
proof -  | 
|
382  | 
  have "{x. \<forall>i. x $ i \<in> S i} = (\<Inter>i. (\<lambda>x. x $ i) -` S i)" by auto
 | 
|
383  | 
  thus "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}"
 | 
|
384  | 
by (simp add: closed_INT closed_vimage_Cart_nth)  | 
|
385  | 
qed  | 
|
386  | 
||
387  | 
lemma tendsto_Cart_nth [tendsto_intros]:  | 
|
388  | 
assumes "((\<lambda>x. f x) ---> a) net"  | 
|
389  | 
shows "((\<lambda>x. f x $ i) ---> a $ i) net"  | 
|
390  | 
proof (rule topological_tendstoI)  | 
|
391  | 
fix S assume "open S" "a $ i \<in> S"  | 
|
392  | 
then have "open ((\<lambda>y. y $ i) -` S)" "a \<in> ((\<lambda>y. y $ i) -` S)"  | 
|
393  | 
by (simp_all add: open_vimage_Cart_nth)  | 
|
394  | 
with assms have "eventually (\<lambda>x. f x \<in> (\<lambda>y. y $ i) -` S) net"  | 
|
395  | 
by (rule topological_tendstoD)  | 
|
396  | 
then show "eventually (\<lambda>x. f x $ i \<in> S) net"  | 
|
397  | 
by simp  | 
|
398  | 
qed  | 
|
399  | 
||
400  | 
subsection {* Square root of sum of squares *}
 | 
|
401  | 
||
402  | 
definition  | 
|
403  | 
"setL2 f A = sqrt (\<Sum>i\<in>A. (f i)\<twosuperior>)"  | 
|
404  | 
||
405  | 
lemma setL2_cong:  | 
|
406  | 
"\<lbrakk>A = B; \<And>x. x \<in> B \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"  | 
|
407  | 
unfolding setL2_def by simp  | 
|
408  | 
||
409  | 
lemma strong_setL2_cong:  | 
|
410  | 
"\<lbrakk>A = B; \<And>x. x \<in> B =simp=> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"  | 
|
411  | 
unfolding setL2_def simp_implies_def by simp  | 
|
412  | 
||
413  | 
lemma setL2_infinite [simp]: "\<not> finite A \<Longrightarrow> setL2 f A = 0"  | 
|
414  | 
unfolding setL2_def by simp  | 
|
415  | 
||
416  | 
lemma setL2_empty [simp]: "setL2 f {} = 0"
 | 
|
417  | 
unfolding setL2_def by simp  | 
|
418  | 
||
419  | 
lemma setL2_insert [simp]:  | 
|
420  | 
"\<lbrakk>finite F; a \<notin> F\<rbrakk> \<Longrightarrow>  | 
|
421  | 
setL2 f (insert a F) = sqrt ((f a)\<twosuperior> + (setL2 f F)\<twosuperior>)"  | 
|
422  | 
unfolding setL2_def by (simp add: setsum_nonneg)  | 
|
423  | 
||
424  | 
lemma setL2_nonneg [simp]: "0 \<le> setL2 f A"  | 
|
425  | 
unfolding setL2_def by (simp add: setsum_nonneg)  | 
|
426  | 
||
427  | 
lemma setL2_0': "\<forall>a\<in>A. f a = 0 \<Longrightarrow> setL2 f A = 0"  | 
|
428  | 
unfolding setL2_def by simp  | 
|
429  | 
||
430  | 
lemma setL2_constant: "setL2 (\<lambda>x. y) A = sqrt (of_nat (card A)) * \<bar>y\<bar>"  | 
|
431  | 
unfolding setL2_def by (simp add: real_sqrt_mult)  | 
|
432  | 
||
433  | 
lemma setL2_mono:  | 
|
434  | 
assumes "\<And>i. i \<in> K \<Longrightarrow> f i \<le> g i"  | 
|
435  | 
assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"  | 
|
436  | 
shows "setL2 f K \<le> setL2 g K"  | 
|
437  | 
unfolding setL2_def  | 
|
438  | 
by (simp add: setsum_nonneg setsum_mono power_mono prems)  | 
|
439  | 
||
440  | 
lemma setL2_strict_mono:  | 
|
441  | 
  assumes "finite K" and "K \<noteq> {}"
 | 
|
442  | 
assumes "\<And>i. i \<in> K \<Longrightarrow> f i < g i"  | 
|
443  | 
assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"  | 
|
444  | 
shows "setL2 f K < setL2 g K"  | 
|
445  | 
unfolding setL2_def  | 
|
446  | 
by (simp add: setsum_strict_mono power_strict_mono assms)  | 
|
447  | 
||
448  | 
lemma setL2_right_distrib:  | 
|
449  | 
"0 \<le> r \<Longrightarrow> r * setL2 f A = setL2 (\<lambda>x. r * f x) A"  | 
|
450  | 
unfolding setL2_def  | 
|
451  | 
apply (simp add: power_mult_distrib)  | 
|
452  | 
apply (simp add: setsum_right_distrib [symmetric])  | 
|
453  | 
apply (simp add: real_sqrt_mult setsum_nonneg)  | 
|
454  | 
done  | 
|
455  | 
||
456  | 
lemma setL2_left_distrib:  | 
|
457  | 
"0 \<le> r \<Longrightarrow> setL2 f A * r = setL2 (\<lambda>x. f x * r) A"  | 
|
458  | 
unfolding setL2_def  | 
|
459  | 
apply (simp add: power_mult_distrib)  | 
|
460  | 
apply (simp add: setsum_left_distrib [symmetric])  | 
|
461  | 
apply (simp add: real_sqrt_mult setsum_nonneg)  | 
|
462  | 
done  | 
|
463  | 
||
464  | 
lemma setsum_nonneg_eq_0_iff:  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34964 
diff
changeset
 | 
465  | 
fixes f :: "'a \<Rightarrow> 'b::ordered_ab_group_add"  | 
| 33175 | 466  | 
shows "\<lbrakk>finite A; \<forall>x\<in>A. 0 \<le> f x\<rbrakk> \<Longrightarrow> setsum f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"  | 
467  | 
apply (induct set: finite, simp)  | 
|
468  | 
apply (simp add: add_nonneg_eq_0_iff setsum_nonneg)  | 
|
469  | 
done  | 
|
470  | 
||
471  | 
lemma setL2_eq_0_iff: "finite A \<Longrightarrow> setL2 f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"  | 
|
472  | 
unfolding setL2_def  | 
|
473  | 
by (simp add: setsum_nonneg setsum_nonneg_eq_0_iff)  | 
|
474  | 
||
475  | 
lemma setL2_triangle_ineq:  | 
|
476  | 
shows "setL2 (\<lambda>i. f i + g i) A \<le> setL2 f A + setL2 g A"  | 
|
477  | 
proof (cases "finite A")  | 
|
478  | 
case False  | 
|
479  | 
thus ?thesis by simp  | 
|
480  | 
next  | 
|
481  | 
case True  | 
|
482  | 
thus ?thesis  | 
|
483  | 
proof (induct set: finite)  | 
|
484  | 
case empty  | 
|
485  | 
show ?case by simp  | 
|
486  | 
next  | 
|
487  | 
case (insert x F)  | 
|
488  | 
hence "sqrt ((f x + g x)\<twosuperior> + (setL2 (\<lambda>i. f i + g i) F)\<twosuperior>) \<le>  | 
|
489  | 
sqrt ((f x + g x)\<twosuperior> + (setL2 f F + setL2 g F)\<twosuperior>)"  | 
|
490  | 
by (intro real_sqrt_le_mono add_left_mono power_mono insert  | 
|
491  | 
setL2_nonneg add_increasing zero_le_power2)  | 
|
492  | 
also have  | 
|
493  | 
"\<dots> \<le> sqrt ((f x)\<twosuperior> + (setL2 f F)\<twosuperior>) + sqrt ((g x)\<twosuperior> + (setL2 g F)\<twosuperior>)"  | 
|
494  | 
by (rule real_sqrt_sum_squares_triangle_ineq)  | 
|
495  | 
finally show ?case  | 
|
496  | 
using insert by simp  | 
|
497  | 
qed  | 
|
498  | 
qed  | 
|
499  | 
||
500  | 
lemma sqrt_sum_squares_le_sum:  | 
|
501  | 
"\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> sqrt (x\<twosuperior> + y\<twosuperior>) \<le> x + y"  | 
|
502  | 
apply (rule power2_le_imp_le)  | 
|
503  | 
apply (simp add: power2_sum)  | 
|
504  | 
apply (simp add: mult_nonneg_nonneg)  | 
|
505  | 
apply (simp add: add_nonneg_nonneg)  | 
|
506  | 
done  | 
|
507  | 
||
508  | 
lemma setL2_le_setsum [rule_format]:  | 
|
509  | 
"(\<forall>i\<in>A. 0 \<le> f i) \<longrightarrow> setL2 f A \<le> setsum f A"  | 
|
510  | 
apply (cases "finite A")  | 
|
511  | 
apply (induct set: finite)  | 
|
512  | 
apply simp  | 
|
513  | 
apply clarsimp  | 
|
514  | 
apply (erule order_trans [OF sqrt_sum_squares_le_sum])  | 
|
515  | 
apply simp  | 
|
516  | 
apply simp  | 
|
517  | 
apply simp  | 
|
518  | 
done  | 
|
519  | 
||
520  | 
lemma sqrt_sum_squares_le_sum_abs: "sqrt (x\<twosuperior> + y\<twosuperior>) \<le> \<bar>x\<bar> + \<bar>y\<bar>"  | 
|
521  | 
apply (rule power2_le_imp_le)  | 
|
522  | 
apply (simp add: power2_sum)  | 
|
523  | 
apply (simp add: mult_nonneg_nonneg)  | 
|
524  | 
apply (simp add: add_nonneg_nonneg)  | 
|
525  | 
done  | 
|
526  | 
||
527  | 
lemma setL2_le_setsum_abs: "setL2 f A \<le> (\<Sum>i\<in>A. \<bar>f i\<bar>)"  | 
|
528  | 
apply (cases "finite A")  | 
|
529  | 
apply (induct set: finite)  | 
|
530  | 
apply simp  | 
|
531  | 
apply simp  | 
|
532  | 
apply (rule order_trans [OF sqrt_sum_squares_le_sum_abs])  | 
|
533  | 
apply simp  | 
|
534  | 
apply simp  | 
|
535  | 
done  | 
|
536  | 
||
537  | 
lemma setL2_mult_ineq_lemma:  | 
|
538  | 
fixes a b c d :: real  | 
|
539  | 
shows "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"  | 
|
540  | 
proof -  | 
|
541  | 
have "0 \<le> (a * d - b * c)\<twosuperior>" by simp  | 
|
542  | 
also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * d) * (b * c)"  | 
|
543  | 
by (simp only: power2_diff power_mult_distrib)  | 
|
544  | 
also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * c) * (b * d)"  | 
|
545  | 
by simp  | 
|
546  | 
finally show "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"  | 
|
547  | 
by simp  | 
|
548  | 
qed  | 
|
549  | 
||
550  | 
lemma setL2_mult_ineq: "(\<Sum>i\<in>A. \<bar>f i\<bar> * \<bar>g i\<bar>) \<le> setL2 f A * setL2 g A"  | 
|
551  | 
apply (cases "finite A")  | 
|
552  | 
apply (induct set: finite)  | 
|
553  | 
apply simp  | 
|
554  | 
apply (rule power2_le_imp_le, simp)  | 
|
555  | 
apply (rule order_trans)  | 
|
556  | 
apply (rule power_mono)  | 
|
557  | 
apply (erule add_left_mono)  | 
|
558  | 
apply (simp add: add_nonneg_nonneg mult_nonneg_nonneg setsum_nonneg)  | 
|
559  | 
apply (simp add: power2_sum)  | 
|
560  | 
apply (simp add: power_mult_distrib)  | 
|
561  | 
apply (simp add: right_distrib left_distrib)  | 
|
562  | 
apply (rule ord_le_eq_trans)  | 
|
563  | 
apply (rule setL2_mult_ineq_lemma)  | 
|
564  | 
apply simp  | 
|
565  | 
apply (intro mult_nonneg_nonneg setL2_nonneg)  | 
|
566  | 
apply simp  | 
|
567  | 
done  | 
|
568  | 
||
569  | 
lemma member_le_setL2: "\<lbrakk>finite A; i \<in> A\<rbrakk> \<Longrightarrow> f i \<le> setL2 f A"  | 
|
570  | 
  apply (rule_tac s="insert i (A - {i})" and t="A" in subst)
 | 
|
571  | 
apply fast  | 
|
572  | 
apply (subst setL2_insert)  | 
|
573  | 
apply simp  | 
|
574  | 
apply simp  | 
|
575  | 
apply simp  | 
|
576  | 
done  | 
|
577  | 
||
578  | 
subsection {* Metric *}
 | 
|
579  | 
||
580  | 
(* TODO: move somewhere else *)  | 
|
581  | 
lemma finite_choice: "finite A \<Longrightarrow> \<forall>x\<in>A. \<exists>y. P x y \<Longrightarrow> \<exists>f. \<forall>x\<in>A. P x (f x)"  | 
|
582  | 
apply (induct set: finite, simp_all)  | 
|
583  | 
apply (clarify, rename_tac y)  | 
|
584  | 
apply (rule_tac x="f(x:=y)" in exI, simp)  | 
|
585  | 
done  | 
|
586  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
587  | 
instantiation cart :: (metric_space, finite) metric_space  | 
| 33175 | 588  | 
begin  | 
589  | 
||
590  | 
definition dist_vector_def:  | 
|
591  | 
"dist (x::'a^'b) (y::'a^'b) = setL2 (\<lambda>i. dist (x$i) (y$i)) UNIV"  | 
|
592  | 
||
593  | 
lemma dist_nth_le: "dist (x $ i) (y $ i) \<le> dist x y"  | 
|
594  | 
unfolding dist_vector_def  | 
|
595  | 
by (rule member_le_setL2) simp_all  | 
|
596  | 
||
597  | 
instance proof  | 
|
598  | 
fix x y :: "'a ^ 'b"  | 
|
599  | 
show "dist x y = 0 \<longleftrightarrow> x = y"  | 
|
600  | 
unfolding dist_vector_def  | 
|
601  | 
by (simp add: setL2_eq_0_iff Cart_eq)  | 
|
602  | 
next  | 
|
603  | 
fix x y z :: "'a ^ 'b"  | 
|
604  | 
show "dist x y \<le> dist x z + dist y z"  | 
|
605  | 
unfolding dist_vector_def  | 
|
606  | 
apply (rule order_trans [OF _ setL2_triangle_ineq])  | 
|
607  | 
apply (simp add: setL2_mono dist_triangle2)  | 
|
608  | 
done  | 
|
609  | 
next  | 
|
610  | 
(* FIXME: long proof! *)  | 
|
611  | 
  fix S :: "('a ^ 'b) set"
 | 
|
612  | 
show "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"  | 
|
613  | 
unfolding open_vector_def open_dist  | 
|
614  | 
apply safe  | 
|
615  | 
apply (drule (1) bspec)  | 
|
616  | 
apply clarify  | 
|
617  | 
apply (subgoal_tac "\<exists>e>0. \<forall>i y. dist y (x$i) < e \<longrightarrow> y \<in> A i")  | 
|
618  | 
apply clarify  | 
|
619  | 
apply (rule_tac x=e in exI, clarify)  | 
|
620  | 
apply (drule spec, erule mp, clarify)  | 
|
621  | 
apply (drule spec, drule spec, erule mp)  | 
|
622  | 
apply (erule le_less_trans [OF dist_nth_le])  | 
|
623  | 
apply (subgoal_tac "\<forall>i\<in>UNIV. \<exists>e>0. \<forall>y. dist y (x$i) < e \<longrightarrow> y \<in> A i")  | 
|
624  | 
apply (drule finite_choice [OF finite], clarify)  | 
|
625  | 
apply (rule_tac x="Min (range f)" in exI, simp)  | 
|
626  | 
apply clarify  | 
|
627  | 
apply (drule_tac x=i in spec, clarify)  | 
|
628  | 
apply (erule (1) bspec)  | 
|
629  | 
apply (drule (1) bspec, clarify)  | 
|
630  | 
apply (subgoal_tac "\<exists>r. (\<forall>i::'b. 0 < r i) \<and> e = setL2 r UNIV")  | 
|
631  | 
apply clarify  | 
|
632  | 
     apply (rule_tac x="\<lambda>i. {y. dist y (x$i) < r i}" in exI)
 | 
|
633  | 
apply (rule conjI)  | 
|
634  | 
apply clarify  | 
|
635  | 
apply (rule conjI)  | 
|
636  | 
apply (clarify, rename_tac y)  | 
|
637  | 
apply (rule_tac x="r i - dist y (x$i)" in exI, rule conjI, simp)  | 
|
638  | 
apply clarify  | 
|
639  | 
apply (simp only: less_diff_eq)  | 
|
640  | 
apply (erule le_less_trans [OF dist_triangle])  | 
|
641  | 
apply simp  | 
|
642  | 
apply clarify  | 
|
643  | 
apply (drule spec, erule mp)  | 
|
644  | 
apply (simp add: dist_vector_def setL2_strict_mono)  | 
|
645  | 
    apply (rule_tac x="\<lambda>i. e / sqrt (of_nat CARD('b))" in exI)
 | 
|
646  | 
apply (simp add: divide_pos_pos setL2_constant)  | 
|
647  | 
done  | 
|
648  | 
qed  | 
|
649  | 
||
650  | 
end  | 
|
651  | 
||
652  | 
lemma LIMSEQ_Cart_nth:  | 
|
653  | 
"(X ----> a) \<Longrightarrow> (\<lambda>n. X n $ i) ----> a $ i"  | 
|
654  | 
unfolding LIMSEQ_conv_tendsto by (rule tendsto_Cart_nth)  | 
|
655  | 
||
656  | 
lemma LIM_Cart_nth:  | 
|
657  | 
"(f -- x --> y) \<Longrightarrow> (\<lambda>x. f x $ i) -- x --> y $ i"  | 
|
658  | 
unfolding LIM_conv_tendsto by (rule tendsto_Cart_nth)  | 
|
659  | 
||
660  | 
lemma Cauchy_Cart_nth:  | 
|
661  | 
"Cauchy (\<lambda>n. X n) \<Longrightarrow> Cauchy (\<lambda>n. X n $ i)"  | 
|
662  | 
unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_nth_le])  | 
|
663  | 
||
664  | 
lemma LIMSEQ_vector:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
665  | 
fixes X :: "nat \<Rightarrow> 'a::metric_space ^ 'n"  | 
| 33175 | 666  | 
assumes X: "\<And>i. (\<lambda>n. X n $ i) ----> (a $ i)"  | 
667  | 
shows "X ----> a"  | 
|
668  | 
proof (rule metric_LIMSEQ_I)  | 
|
669  | 
fix r :: real assume "0 < r"  | 
|
670  | 
  then have "0 < r / of_nat CARD('n)" (is "0 < ?s")
 | 
|
671  | 
by (simp add: divide_pos_pos)  | 
|
672  | 
def N \<equiv> "\<lambda>i. LEAST N. \<forall>n\<ge>N. dist (X n $ i) (a $ i) < ?s"  | 
|
673  | 
def M \<equiv> "Max (range N)"  | 
|
674  | 
have "\<And>i. \<exists>N. \<forall>n\<ge>N. dist (X n $ i) (a $ i) < ?s"  | 
|
675  | 
using X `0 < ?s` by (rule metric_LIMSEQ_D)  | 
|
676  | 
hence "\<And>i. \<forall>n\<ge>N i. dist (X n $ i) (a $ i) < ?s"  | 
|
677  | 
unfolding N_def by (rule LeastI_ex)  | 
|
678  | 
hence M: "\<And>i. \<forall>n\<ge>M. dist (X n $ i) (a $ i) < ?s"  | 
|
679  | 
unfolding M_def by simp  | 
|
680  | 
  {
 | 
|
681  | 
fix n :: nat assume "M \<le> n"  | 
|
682  | 
have "dist (X n) a = setL2 (\<lambda>i. dist (X n $ i) (a $ i)) UNIV"  | 
|
683  | 
unfolding dist_vector_def ..  | 
|
684  | 
also have "\<dots> \<le> setsum (\<lambda>i. dist (X n $ i) (a $ i)) UNIV"  | 
|
685  | 
by (rule setL2_le_setsum [OF zero_le_dist])  | 
|
686  | 
also have "\<dots> < setsum (\<lambda>i::'n. ?s) UNIV"  | 
|
687  | 
by (rule setsum_strict_mono, simp_all add: M `M \<le> n`)  | 
|
688  | 
also have "\<dots> = r"  | 
|
689  | 
by simp  | 
|
690  | 
finally have "dist (X n) a < r" .  | 
|
691  | 
}  | 
|
692  | 
hence "\<forall>n\<ge>M. dist (X n) a < r"  | 
|
693  | 
by simp  | 
|
694  | 
then show "\<exists>M. \<forall>n\<ge>M. dist (X n) a < r" ..  | 
|
695  | 
qed  | 
|
696  | 
||
697  | 
lemma Cauchy_vector:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
698  | 
fixes X :: "nat \<Rightarrow> 'a::metric_space ^ 'n"  | 
| 33175 | 699  | 
assumes X: "\<And>i. Cauchy (\<lambda>n. X n $ i)"  | 
700  | 
shows "Cauchy (\<lambda>n. X n)"  | 
|
701  | 
proof (rule metric_CauchyI)  | 
|
702  | 
fix r :: real assume "0 < r"  | 
|
703  | 
  then have "0 < r / of_nat CARD('n)" (is "0 < ?s")
 | 
|
704  | 
by (simp add: divide_pos_pos)  | 
|
705  | 
def N \<equiv> "\<lambda>i. LEAST N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s"  | 
|
706  | 
def M \<equiv> "Max (range N)"  | 
|
707  | 
have "\<And>i. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s"  | 
|
708  | 
using X `0 < ?s` by (rule metric_CauchyD)  | 
|
709  | 
hence "\<And>i. \<forall>m\<ge>N i. \<forall>n\<ge>N i. dist (X m $ i) (X n $ i) < ?s"  | 
|
710  | 
unfolding N_def by (rule LeastI_ex)  | 
|
711  | 
hence M: "\<And>i. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m $ i) (X n $ i) < ?s"  | 
|
712  | 
unfolding M_def by simp  | 
|
713  | 
  {
 | 
|
714  | 
fix m n :: nat  | 
|
715  | 
assume "M \<le> m" "M \<le> n"  | 
|
716  | 
have "dist (X m) (X n) = setL2 (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV"  | 
|
717  | 
unfolding dist_vector_def ..  | 
|
718  | 
also have "\<dots> \<le> setsum (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV"  | 
|
719  | 
by (rule setL2_le_setsum [OF zero_le_dist])  | 
|
720  | 
also have "\<dots> < setsum (\<lambda>i::'n. ?s) UNIV"  | 
|
721  | 
by (rule setsum_strict_mono, simp_all add: M `M \<le> m` `M \<le> n`)  | 
|
722  | 
also have "\<dots> = r"  | 
|
723  | 
by simp  | 
|
724  | 
finally have "dist (X m) (X n) < r" .  | 
|
725  | 
}  | 
|
726  | 
hence "\<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < r"  | 
|
727  | 
by simp  | 
|
728  | 
then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < r" ..  | 
|
729  | 
qed  | 
|
730  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
731  | 
instance cart :: (complete_space, finite) complete_space  | 
| 33175 | 732  | 
proof  | 
733  | 
fix X :: "nat \<Rightarrow> 'a ^ 'b" assume "Cauchy X"  | 
|
734  | 
have "\<And>i. (\<lambda>n. X n $ i) ----> lim (\<lambda>n. X n $ i)"  | 
|
735  | 
using Cauchy_Cart_nth [OF `Cauchy X`]  | 
|
736  | 
by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff)  | 
|
737  | 
hence "X ----> Cart_lambda (\<lambda>i. lim (\<lambda>n. X n $ i))"  | 
|
738  | 
by (simp add: LIMSEQ_vector)  | 
|
739  | 
then show "convergent X"  | 
|
740  | 
by (rule convergentI)  | 
|
741  | 
qed  | 
|
742  | 
||
743  | 
subsection {* Norms *}
 | 
|
744  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
745  | 
instantiation cart :: (real_normed_vector, finite) real_normed_vector  | 
| 33175 | 746  | 
begin  | 
747  | 
||
748  | 
definition norm_vector_def:  | 
|
749  | 
"norm (x::'a^'b) = setL2 (\<lambda>i. norm (x$i)) UNIV"  | 
|
750  | 
||
751  | 
definition vector_sgn_def:  | 
|
752  | 
"sgn (x::'a^'b) = scaleR (inverse (norm x)) x"  | 
|
753  | 
||
754  | 
instance proof  | 
|
755  | 
fix a :: real and x y :: "'a ^ 'b"  | 
|
756  | 
show "0 \<le> norm x"  | 
|
757  | 
unfolding norm_vector_def  | 
|
758  | 
by (rule setL2_nonneg)  | 
|
759  | 
show "norm x = 0 \<longleftrightarrow> x = 0"  | 
|
760  | 
unfolding norm_vector_def  | 
|
761  | 
by (simp add: setL2_eq_0_iff Cart_eq)  | 
|
762  | 
show "norm (x + y) \<le> norm x + norm y"  | 
|
763  | 
unfolding norm_vector_def  | 
|
764  | 
apply (rule order_trans [OF _ setL2_triangle_ineq])  | 
|
765  | 
apply (simp add: setL2_mono norm_triangle_ineq)  | 
|
766  | 
done  | 
|
767  | 
show "norm (scaleR a x) = \<bar>a\<bar> * norm x"  | 
|
768  | 
unfolding norm_vector_def  | 
|
769  | 
by (simp add: setL2_right_distrib)  | 
|
770  | 
show "sgn x = scaleR (inverse (norm x)) x"  | 
|
771  | 
by (rule vector_sgn_def)  | 
|
772  | 
show "dist x y = norm (x - y)"  | 
|
773  | 
unfolding dist_vector_def norm_vector_def  | 
|
774  | 
by (simp add: dist_norm)  | 
|
775  | 
qed  | 
|
776  | 
||
777  | 
end  | 
|
778  | 
||
779  | 
lemma norm_nth_le: "norm (x $ i) \<le> norm x"  | 
|
780  | 
unfolding norm_vector_def  | 
|
781  | 
by (rule member_le_setL2) simp_all  | 
|
782  | 
||
783  | 
interpretation Cart_nth: bounded_linear "\<lambda>x. x $ i"  | 
|
784  | 
apply default  | 
|
785  | 
apply (rule vector_add_component)  | 
|
786  | 
apply (rule vector_scaleR_component)  | 
|
787  | 
apply (rule_tac x="1" in exI, simp add: norm_nth_le)  | 
|
788  | 
done  | 
|
789  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
790  | 
instance cart :: (banach, finite) banach ..  | 
| 33175 | 791  | 
|
792  | 
subsection {* Inner products *}
 | 
|
793  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
794  | 
instantiation cart :: (real_inner, finite) real_inner  | 
| 33175 | 795  | 
begin  | 
796  | 
||
797  | 
definition inner_vector_def:  | 
|
798  | 
"inner x y = setsum (\<lambda>i. inner (x$i) (y$i)) UNIV"  | 
|
799  | 
||
800  | 
instance proof  | 
|
801  | 
fix r :: real and x y z :: "'a ^ 'b"  | 
|
802  | 
show "inner x y = inner y x"  | 
|
803  | 
unfolding inner_vector_def  | 
|
804  | 
by (simp add: inner_commute)  | 
|
805  | 
show "inner (x + y) z = inner x z + inner y z"  | 
|
806  | 
unfolding inner_vector_def  | 
|
807  | 
by (simp add: inner_add_left setsum_addf)  | 
|
808  | 
show "inner (scaleR r x) y = r * inner x y"  | 
|
809  | 
unfolding inner_vector_def  | 
|
810  | 
by (simp add: setsum_right_distrib)  | 
|
811  | 
show "0 \<le> inner x x"  | 
|
812  | 
unfolding inner_vector_def  | 
|
813  | 
by (simp add: setsum_nonneg)  | 
|
814  | 
show "inner x x = 0 \<longleftrightarrow> x = 0"  | 
|
815  | 
unfolding inner_vector_def  | 
|
816  | 
by (simp add: Cart_eq setsum_nonneg_eq_0_iff)  | 
|
817  | 
show "norm x = sqrt (inner x x)"  | 
|
818  | 
unfolding inner_vector_def norm_vector_def setL2_def  | 
|
819  | 
by (simp add: power2_norm_eq_inner)  | 
|
820  | 
qed  | 
|
821  | 
||
822  | 
end  | 
|
823  | 
||
824  | 
subsection{* Properties of the dot product.  *}
 | 
|
825  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
826  | 
lemma dot_sym: "(x::'a:: {comm_monoid_add, ab_semigroup_mult} ^ 'n) \<bullet> y = y \<bullet> x"
 | 
| 33175 | 827  | 
by (vector mult_commute)  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
828  | 
lemma dot_ladd: "((x::'a::ring ^ 'n) + y) \<bullet> z = (x \<bullet> z) + (y \<bullet> z)"  | 
| 33175 | 829  | 
by (vector ring_simps)  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
830  | 
lemma dot_radd: "x \<bullet> (y + (z::'a::ring ^ 'n)) = (x \<bullet> y) + (x \<bullet> z)"  | 
| 33175 | 831  | 
by (vector ring_simps)  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
832  | 
lemma dot_lsub: "((x::'a::ring ^ 'n) - y) \<bullet> z = (x \<bullet> z) - (y \<bullet> z)"  | 
| 33175 | 833  | 
by (vector ring_simps)  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
834  | 
lemma dot_rsub: "(x::'a::ring ^ 'n) \<bullet> (y - z) = (x \<bullet> y) - (x \<bullet> z)"  | 
| 33175 | 835  | 
by (vector ring_simps)  | 
836  | 
lemma dot_lmult: "(c *s x) \<bullet> y = (c::'a::ring) * (x \<bullet> y)" by (vector ring_simps)  | 
|
837  | 
lemma dot_rmult: "x \<bullet> (c *s y) = (c::'a::comm_ring) * (x \<bullet> y)" by (vector ring_simps)  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
838  | 
lemma dot_lneg: "(-x) \<bullet> (y::'a::ring ^ 'n) = -(x \<bullet> y)" by vector  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
839  | 
lemma dot_rneg: "(x::'a::ring ^ 'n) \<bullet> (-y) = -(x \<bullet> y)" by vector  | 
| 33175 | 840  | 
lemma dot_lzero[simp]: "0 \<bullet> x = (0::'a::{comm_monoid_add, mult_zero})" by vector
 | 
841  | 
lemma dot_rzero[simp]: "x \<bullet> 0 = (0::'a::{comm_monoid_add, mult_zero})" by vector
 | 
|
| 
35043
 
07dbdf60d5ad
dropped accidental duplication of "lin" prefix from cs. 108662d50512
 
haftmann 
parents: 
35028 
diff
changeset
 | 
842  | 
lemma dot_pos_le[simp]: "(0::'a\<Colon>linordered_ring_strict) <= x \<bullet> x"  | 
| 33175 | 843  | 
by (simp add: dot_def setsum_nonneg)  | 
844  | 
||
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34964 
diff
changeset
 | 
845  | 
lemma setsum_squares_eq_0_iff: assumes fS: "finite F" and fp: "\<forall>x \<in> F. f x \<ge> (0 ::'a::ordered_ab_group_add)" shows "setsum f F = 0 \<longleftrightarrow> (ALL x:F. f x = 0)"  | 
| 33175 | 846  | 
using fS fp setsum_nonneg[OF fp]  | 
847  | 
proof (induct set: finite)  | 
|
848  | 
case empty thus ?case by simp  | 
|
849  | 
next  | 
|
850  | 
case (insert x F)  | 
|
851  | 
from insert.prems have Fx: "f x \<ge> 0" and Fp: "\<forall> a \<in> F. f a \<ge> 0" by simp_all  | 
|
852  | 
from insert.hyps Fp setsum_nonneg[OF Fp]  | 
|
853  | 
have h: "setsum f F = 0 \<longleftrightarrow> (\<forall>a \<in>F. f a = 0)" by metis  | 
|
854  | 
from add_nonneg_eq_0_iff[OF Fx setsum_nonneg[OF Fp]] insert.hyps(1,2)  | 
|
855  | 
show ?case by (simp add: h)  | 
|
856  | 
qed  | 
|
857  | 
||
| 
35043
 
07dbdf60d5ad
dropped accidental duplication of "lin" prefix from cs. 108662d50512
 
haftmann 
parents: 
35028 
diff
changeset
 | 
858  | 
lemma dot_eq_0: "x \<bullet> x = 0 \<longleftrightarrow> (x::'a::{linordered_ring_strict,ring_no_zero_divisors} ^ 'n) = 0"
 | 
| 33175 | 859  | 
by (simp add: dot_def setsum_squares_eq_0_iff Cart_eq)  | 
860  | 
||
| 
35043
 
07dbdf60d5ad
dropped accidental duplication of "lin" prefix from cs. 108662d50512
 
haftmann 
parents: 
35028 
diff
changeset
 | 
861  | 
lemma dot_pos_lt[simp]: "(0 < x \<bullet> x) \<longleftrightarrow> (x::'a::{linordered_ring_strict,ring_no_zero_divisors} ^ 'n) \<noteq> 0" using dot_eq_0[of x] dot_pos_le[of x]
 | 
| 33175 | 862  | 
by (auto simp add: le_less)  | 
863  | 
||
864  | 
subsection{* The collapse of the general concepts to dimension one. *}
 | 
|
865  | 
||
866  | 
lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))"  | 
|
867  | 
by (simp add: Cart_eq forall_1)  | 
|
868  | 
||
869  | 
lemma forall_one: "(\<forall>(x::'a ^1). P x) \<longleftrightarrow> (\<forall>x. P(\<chi> i. x))"  | 
|
870  | 
apply auto  | 
|
871  | 
apply (erule_tac x= "x$1" in allE)  | 
|
872  | 
apply (simp only: vector_one[symmetric])  | 
|
873  | 
done  | 
|
874  | 
||
875  | 
lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)"  | 
|
876  | 
by (simp add: norm_vector_def UNIV_1)  | 
|
877  | 
||
878  | 
lemma norm_real: "norm(x::real ^ 1) = abs(x$1)"  | 
|
879  | 
by (simp add: norm_vector_1)  | 
|
880  | 
||
881  | 
lemma dist_real: "dist(x::real ^ 1) y = abs((x$1) - (y$1))"  | 
|
882  | 
by (auto simp add: norm_real dist_norm)  | 
|
883  | 
||
884  | 
subsection {* A connectedness or intermediate value lemma with several applications. *}
 | 
|
885  | 
||
886  | 
lemma connected_real_lemma:  | 
|
887  | 
fixes f :: "real \<Rightarrow> 'a::metric_space"  | 
|
888  | 
assumes ab: "a \<le> b" and fa: "f a \<in> e1" and fb: "f b \<in> e2"  | 
|
889  | 
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"  | 
|
890  | 
and e1: "\<forall>y \<in> e1. \<exists>e > 0. \<forall>y'. dist y' y < e \<longrightarrow> y' \<in> e1"  | 
|
891  | 
and e2: "\<forall>y \<in> e2. \<exists>e > 0. \<forall>y'. dist y' y < e \<longrightarrow> y' \<in> e2"  | 
|
892  | 
and e12: "~(\<exists>x \<ge> a. x <= b \<and> f x \<in> e1 \<and> f x \<in> e2)"  | 
|
893  | 
shows "\<exists>x \<ge> a. x <= b \<and> f x \<notin> e1 \<and> f x \<notin> e2" (is "\<exists> x. ?P x")  | 
|
894  | 
proof-  | 
|
895  | 
  let ?S = "{c. \<forall>x \<ge> a. x <= c \<longrightarrow> f x \<in> e1}"
 | 
|
896  | 
have Se: " \<exists>x. x \<in> ?S" apply (rule exI[where x=a]) by (auto simp add: fa)  | 
|
897  | 
have Sub: "\<exists>y. isUb UNIV ?S y"  | 
|
898  | 
apply (rule exI[where x= b])  | 
|
899  | 
using ab fb e12 by (auto simp add: isUb_def setle_def)  | 
|
900  | 
from reals_complete[OF Se Sub] obtain l where  | 
|
901  | 
l: "isLub UNIV ?S l"by blast  | 
|
902  | 
have alb: "a \<le> l" "l \<le> b" using l ab fa fb e12  | 
|
903  | 
apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def)  | 
|
904  | 
by (metis linorder_linear)  | 
|
905  | 
have ale1: "\<forall>z \<ge> a. z < l \<longrightarrow> f z \<in> e1" using l  | 
|
906  | 
apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def)  | 
|
907  | 
by (metis linorder_linear not_le)  | 
|
908  | 
have th1: "\<And>z x e d :: real. z <= x + e \<Longrightarrow> e < d ==> z < x \<or> abs(z - x) < d" by arith  | 
|
909  | 
have th2: "\<And>e x:: real. 0 < e ==> ~(x + e <= x)" by arith  | 
|
910  | 
have th3: "\<And>d::real. d > 0 \<Longrightarrow> \<exists>e > 0. e < d" by dlo  | 
|
911  | 
    {assume le2: "f l \<in> e2"
 | 
|
912  | 
from le2 fa fb e12 alb have la: "l \<noteq> a" by metis  | 
|
913  | 
hence lap: "l - a > 0" using alb by arith  | 
|
914  | 
from e2[rule_format, OF le2] obtain e where  | 
|
915  | 
e: "e > 0" "\<forall>y. dist y (f l) < e \<longrightarrow> y \<in> e2" by metis  | 
|
916  | 
from dst[OF alb e(1)] obtain d where  | 
|
917  | 
d: "d > 0" "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> dist (f y) (f l) < e" by metis  | 
|
918  | 
have "\<exists>d'. d' < d \<and> d' >0 \<and> l - d' > a" using lap d(1)  | 
|
919  | 
apply ferrack by arith  | 
|
920  | 
then obtain d' where d': "d' > 0" "d' < d" "l - d' > a" by metis  | 
|
921  | 
from d e have th0: "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> f y \<in> e2" by metis  | 
|
922  | 
from th0[rule_format, of "l - d'"] d' have "f (l - d') \<in> e2" by auto  | 
|
923  | 
moreover  | 
|
924  | 
have "f (l - d') \<in> e1" using ale1[rule_format, of "l -d'"] d' by auto  | 
|
925  | 
ultimately have False using e12 alb d' by auto}  | 
|
926  | 
moreover  | 
|
927  | 
    {assume le1: "f l \<in> e1"
 | 
|
928  | 
from le1 fa fb e12 alb have lb: "l \<noteq> b" by metis  | 
|
929  | 
hence blp: "b - l > 0" using alb by arith  | 
|
930  | 
from e1[rule_format, OF le1] obtain e where  | 
|
931  | 
e: "e > 0" "\<forall>y. dist y (f l) < e \<longrightarrow> y \<in> e1" by metis  | 
|
932  | 
from dst[OF alb e(1)] obtain d where  | 
|
933  | 
d: "d > 0" "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> dist (f y) (f l) < e" by metis  | 
|
934  | 
have "\<exists>d'. d' < d \<and> d' >0" using d(1) by dlo  | 
|
935  | 
then obtain d' where d': "d' > 0" "d' < d" by metis  | 
|
936  | 
from d e have th0: "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> f y \<in> e1" by auto  | 
|
937  | 
hence "\<forall>y. l \<le> y \<and> y \<le> l + d' \<longrightarrow> f y \<in> e1" using d' by auto  | 
|
938  | 
with ale1 have "\<forall>y. a \<le> y \<and> y \<le> l + d' \<longrightarrow> f y \<in> e1" by auto  | 
|
939  | 
with l d' have False  | 
|
940  | 
by (auto simp add: isLub_def isUb_def setle_def setge_def leastP_def) }  | 
|
941  | 
ultimately show ?thesis using alb by metis  | 
|
942  | 
qed  | 
|
943  | 
||
944  | 
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 @{typ "real^1"} *}
 | 
|
945  | 
||
946  | 
lemma square_bound_lemma: "(x::real) < (1 + x) * (1 + x)"  | 
|
947  | 
proof-  | 
|
948  | 
have "(x + 1/2)^2 + 3/4 > 0" using zero_le_power2[of "x+1/2"] by arith  | 
|
949  | 
thus ?thesis by (simp add: ring_simps power2_eq_square)  | 
|
950  | 
qed  | 
|
951  | 
||
952  | 
lemma square_continuous: "0 < (e::real) ==> \<exists>d. 0 < d \<and> (\<forall>y. abs(y - x) < d \<longrightarrow> abs(y * y - x * x) < e)"  | 
|
953  | 
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)  | 
|
954  | 
apply (rule_tac x="s" in exI)  | 
|
955  | 
apply auto  | 
|
956  | 
apply (erule_tac x=y in allE)  | 
|
957  | 
apply auto  | 
|
958  | 
done  | 
|
959  | 
||
960  | 
lemma real_le_lsqrt: "0 <= x \<Longrightarrow> 0 <= y \<Longrightarrow> x <= y^2 ==> sqrt x <= y"  | 
|
961  | 
using real_sqrt_le_iff[of x "y^2"] by simp  | 
|
962  | 
||
963  | 
lemma real_le_rsqrt: "x^2 \<le> y \<Longrightarrow> x \<le> sqrt y"  | 
|
964  | 
using real_sqrt_le_mono[of "x^2" y] by simp  | 
|
965  | 
||
966  | 
lemma real_less_rsqrt: "x^2 < y \<Longrightarrow> x < sqrt y"  | 
|
967  | 
using real_sqrt_less_mono[of "x^2" y] by simp  | 
|
968  | 
||
969  | 
lemma sqrt_even_pow2: assumes n: "even n"  | 
|
970  | 
shows "sqrt(2 ^ n) = 2 ^ (n div 2)"  | 
|
971  | 
proof-  | 
|
972  | 
from n obtain m where m: "n = 2*m" unfolding even_nat_equiv_def2  | 
|
973  | 
by (auto simp add: nat_number)  | 
|
974  | 
from m have "sqrt(2 ^ n) = sqrt ((2 ^ m) ^ 2)"  | 
|
975  | 
by (simp only: power_mult[symmetric] mult_commute)  | 
|
976  | 
then show ?thesis using m by simp  | 
|
977  | 
qed  | 
|
978  | 
||
979  | 
lemma real_div_sqrt: "0 <= x ==> x / sqrt(x) = sqrt(x)"  | 
|
980  | 
apply (cases "x = 0", simp_all)  | 
|
981  | 
using sqrt_divide_self_eq[of x]  | 
|
982  | 
apply (simp add: inverse_eq_divide real_sqrt_ge_0_iff field_simps)  | 
|
983  | 
done  | 
|
984  | 
||
985  | 
text{* Hence derive more interesting properties of the norm. *}
 | 
|
986  | 
||
987  | 
text {*
 | 
|
988  | 
This type-specific version is only here  | 
|
989  | 
  to make @{text normarith.ML} happy.
 | 
|
990  | 
*}  | 
|
991  | 
lemma norm_0: "norm (0::real ^ _) = 0"  | 
|
992  | 
by (rule norm_zero)  | 
|
993  | 
||
994  | 
lemma norm_mul[simp]: "norm(a *s x) = abs(a) * norm x"  | 
|
995  | 
by (simp add: norm_vector_def vector_component setL2_right_distrib  | 
|
996  | 
abs_mult cong: strong_setL2_cong)  | 
|
997  | 
lemma norm_eq_0_dot: "(norm x = 0) \<longleftrightarrow> (x \<bullet> x = (0::real))"  | 
|
998  | 
by (simp add: norm_vector_def dot_def setL2_def power2_eq_square)  | 
|
999  | 
lemma real_vector_norm_def: "norm x = sqrt (x \<bullet> x)"  | 
|
1000  | 
by (simp add: norm_vector_def setL2_def dot_def power2_eq_square)  | 
|
1001  | 
lemma norm_pow_2: "norm x ^ 2 = x \<bullet> x"  | 
|
1002  | 
by (simp add: real_vector_norm_def)  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1003  | 
lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n)" by (metis norm_eq_zero)  | 
| 33175 | 1004  | 
lemma vector_mul_eq_0[simp]: "(a *s x = 0) \<longleftrightarrow> a = (0::'a::idom) \<or> x = 0"  | 
1005  | 
by vector  | 
|
1006  | 
lemma vector_mul_lcancel[simp]: "a *s x = a *s y \<longleftrightarrow> a = (0::real) \<or> x = y"  | 
|
1007  | 
by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_ssub_ldistrib)  | 
|
1008  | 
lemma vector_mul_rcancel[simp]: "a *s x = b *s x \<longleftrightarrow> (a::real) = b \<or> x = 0"  | 
|
1009  | 
by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_sub_rdistrib)  | 
|
1010  | 
lemma vector_mul_lcancel_imp: "a \<noteq> (0::real) ==> a *s x = a *s y ==> (x = y)"  | 
|
1011  | 
by (metis vector_mul_lcancel)  | 
|
1012  | 
lemma vector_mul_rcancel_imp: "x \<noteq> 0 \<Longrightarrow> (a::real) *s x = b *s x ==> a = b"  | 
|
1013  | 
by (metis vector_mul_rcancel)  | 
|
1014  | 
lemma norm_cauchy_schwarz:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1015  | 
fixes x y :: "real ^ 'n"  | 
| 33175 | 1016  | 
shows "x \<bullet> y <= norm x * norm y"  | 
1017  | 
proof-  | 
|
1018  | 
  {assume "norm x = 0"
 | 
|
1019  | 
hence ?thesis by (simp add: dot_lzero dot_rzero)}  | 
|
1020  | 
moreover  | 
|
1021  | 
  {assume "norm y = 0"
 | 
|
1022  | 
hence ?thesis by (simp add: dot_lzero dot_rzero)}  | 
|
1023  | 
moreover  | 
|
1024  | 
  {assume h: "norm x \<noteq> 0" "norm y \<noteq> 0"
 | 
|
1025  | 
let ?z = "norm y *s x - norm x *s y"  | 
|
1026  | 
from h have p: "norm x * norm y > 0" by (metis norm_ge_zero le_less zero_compare_simps)  | 
|
1027  | 
from dot_pos_le[of ?z]  | 
|
1028  | 
have "(norm x * norm y) * (x \<bullet> y) \<le> norm x ^2 * norm y ^2"  | 
|
1029  | 
apply (simp add: dot_rsub dot_lsub dot_lmult dot_rmult ring_simps)  | 
|
1030  | 
by (simp add: norm_pow_2[symmetric] power2_eq_square dot_sym)  | 
|
1031  | 
hence "x\<bullet>y \<le> (norm x ^2 * norm y ^2) / (norm x * norm y)" using p  | 
|
1032  | 
by (simp add: field_simps)  | 
|
1033  | 
hence ?thesis using h by (simp add: power2_eq_square)}  | 
|
1034  | 
ultimately show ?thesis by metis  | 
|
1035  | 
qed  | 
|
1036  | 
||
1037  | 
lemma norm_cauchy_schwarz_abs:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1038  | 
fixes x y :: "real ^ 'n"  | 
| 33175 | 1039  | 
shows "\<bar>x \<bullet> y\<bar> \<le> norm x * norm y"  | 
1040  | 
using norm_cauchy_schwarz[of x y] norm_cauchy_schwarz[of x "-y"]  | 
|
1041  | 
by (simp add: real_abs_def dot_rneg)  | 
|
1042  | 
||
1043  | 
lemma norm_triangle_sub:  | 
|
1044  | 
fixes x y :: "'a::real_normed_vector"  | 
|
1045  | 
shows "norm x \<le> norm y + norm (x - y)"  | 
|
1046  | 
using norm_triangle_ineq[of "y" "x - y"] by (simp add: ring_simps)  | 
|
1047  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1048  | 
lemma component_le_norm: "\<bar>x$i\<bar> <= norm x"  | 
| 33175 | 1049  | 
apply (simp add: norm_vector_def)  | 
1050  | 
apply (rule member_le_setL2, simp_all)  | 
|
1051  | 
done  | 
|
1052  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1053  | 
lemma norm_bound_component_le: "norm x <= e ==> \<bar>x$i\<bar> <= e"  | 
| 33175 | 1054  | 
by (metis component_le_norm order_trans)  | 
1055  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1056  | 
lemma norm_bound_component_lt: "norm x < e ==> \<bar>x$i\<bar> < e"  | 
| 33175 | 1057  | 
by (metis component_le_norm basic_trans_rules(21))  | 
1058  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1059  | 
lemma norm_le_l1: "norm x <= setsum(\<lambda>i. \<bar>x$i\<bar>) UNIV"  | 
| 33175 | 1060  | 
by (simp add: norm_vector_def setL2_le_setsum)  | 
1061  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1062  | 
lemma real_abs_norm: "\<bar>norm x\<bar> = norm x"  | 
| 33175 | 1063  | 
by (rule abs_norm_cancel)  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1064  | 
lemma real_abs_sub_norm: "\<bar>norm (x::real ^ 'n) - norm y\<bar> <= norm(x - y)"  | 
| 33175 | 1065  | 
by (rule norm_triangle_ineq3)  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1066  | 
lemma norm_le: "norm(x::real ^ 'n) <= norm(y) \<longleftrightarrow> x \<bullet> x <= y \<bullet> y"  | 
| 33175 | 1067  | 
by (simp add: real_vector_norm_def)  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1068  | 
lemma norm_lt: "norm(x::real ^ 'n) < norm(y) \<longleftrightarrow> x \<bullet> x < y \<bullet> y"  | 
| 33175 | 1069  | 
by (simp add: real_vector_norm_def)  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1070  | 
lemma norm_eq: "norm(x::real ^ 'n) = norm y \<longleftrightarrow> x \<bullet> x = y \<bullet> y"  | 
| 33175 | 1071  | 
by (simp add: order_eq_iff norm_le)  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1072  | 
lemma norm_eq_1: "norm(x::real ^ 'n) = 1 \<longleftrightarrow> x \<bullet> x = 1"  | 
| 33175 | 1073  | 
by (simp add: real_vector_norm_def)  | 
1074  | 
||
1075  | 
text{* Squaring equations and inequalities involving norms.  *}
 | 
|
1076  | 
||
1077  | 
lemma dot_square_norm: "x \<bullet> x = norm(x)^2"  | 
|
1078  | 
by (simp add: real_vector_norm_def)  | 
|
1079  | 
||
1080  | 
lemma norm_eq_square: "norm(x) = a \<longleftrightarrow> 0 <= a \<and> x \<bullet> x = a^2"  | 
|
1081  | 
by (auto simp add: real_vector_norm_def)  | 
|
1082  | 
||
1083  | 
lemma real_abs_le_square_iff: "\<bar>x\<bar> \<le> \<bar>y\<bar> \<longleftrightarrow> (x::real)^2 \<le> y^2"  | 
|
1084  | 
proof-  | 
|
1085  | 
have "x^2 \<le> y^2 \<longleftrightarrow> (x -y) * (y + x) \<le> 0" by (simp add: ring_simps power2_eq_square)  | 
|
1086  | 
also have "\<dots> \<longleftrightarrow> \<bar>x\<bar> \<le> \<bar>y\<bar>" apply (simp add: zero_compare_simps real_abs_def not_less) by arith  | 
|
1087  | 
finally show ?thesis ..  | 
|
1088  | 
qed  | 
|
1089  | 
||
1090  | 
lemma norm_le_square: "norm(x) <= a \<longleftrightarrow> 0 <= a \<and> x \<bullet> x <= a^2"  | 
|
1091  | 
apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric])  | 
|
1092  | 
using norm_ge_zero[of x]  | 
|
1093  | 
apply arith  | 
|
1094  | 
done  | 
|
1095  | 
||
1096  | 
lemma norm_ge_square: "norm(x) >= a \<longleftrightarrow> a <= 0 \<or> x \<bullet> x >= a ^ 2"  | 
|
1097  | 
apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric])  | 
|
1098  | 
using norm_ge_zero[of x]  | 
|
1099  | 
apply arith  | 
|
1100  | 
done  | 
|
1101  | 
||
1102  | 
lemma norm_lt_square: "norm(x) < a \<longleftrightarrow> 0 < a \<and> x \<bullet> x < a^2"  | 
|
1103  | 
by (metis not_le norm_ge_square)  | 
|
1104  | 
lemma norm_gt_square: "norm(x) > a \<longleftrightarrow> a < 0 \<or> x \<bullet> x > a^2"  | 
|
1105  | 
by (metis norm_le_square not_less)  | 
|
1106  | 
||
1107  | 
text{* Dot product in terms of the norm rather than conversely. *}
 | 
|
1108  | 
||
1109  | 
lemma dot_norm: "x \<bullet> y = (norm(x + y) ^2 - norm x ^ 2 - norm y ^ 2) / 2"  | 
|
1110  | 
by (simp add: norm_pow_2 dot_ladd dot_radd dot_sym)  | 
|
1111  | 
||
1112  | 
lemma dot_norm_neg: "x \<bullet> y = ((norm x ^ 2 + norm y ^ 2) - norm(x - y) ^ 2) / 2"  | 
|
1113  | 
by (simp add: norm_pow_2 dot_ladd dot_radd dot_lsub dot_rsub dot_sym)  | 
|
1114  | 
||
1115  | 
||
1116  | 
text{* Equality of vectors in terms of @{term "op \<bullet>"} products.    *}
 | 
|
1117  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1118  | 
lemma vector_eq: "(x:: real ^ 'n) = y \<longleftrightarrow> x \<bullet> x = x \<bullet> y\<and> y \<bullet> y = x \<bullet> x" (is "?lhs \<longleftrightarrow> ?rhs")  | 
| 33175 | 1119  | 
proof  | 
1120  | 
assume "?lhs" then show ?rhs by simp  | 
|
1121  | 
next  | 
|
1122  | 
assume ?rhs  | 
|
1123  | 
then have "x \<bullet> x - x \<bullet> y = 0 \<and> x \<bullet> y - y\<bullet> y = 0" by simp  | 
|
1124  | 
hence "x \<bullet> (x - y) = 0 \<and> y \<bullet> (x - y) = 0"  | 
|
1125  | 
by (simp add: dot_rsub dot_lsub dot_sym)  | 
|
1126  | 
then have "(x - y) \<bullet> (x - y) = 0" by (simp add: ring_simps dot_lsub dot_rsub)  | 
|
1127  | 
then show "x = y" by (simp add: dot_eq_0)  | 
|
1128  | 
qed  | 
|
1129  | 
||
1130  | 
||
1131  | 
subsection{* General linear decision procedure for normed spaces. *}
 | 
|
1132  | 
||
1133  | 
lemma norm_cmul_rule_thm:  | 
|
1134  | 
fixes x :: "'a::real_normed_vector"  | 
|
1135  | 
shows "b >= norm(x) ==> \<bar>c\<bar> * b >= norm(scaleR c x)"  | 
|
1136  | 
unfolding norm_scaleR  | 
|
1137  | 
apply (erule mult_mono1)  | 
|
1138  | 
apply simp  | 
|
1139  | 
done  | 
|
1140  | 
||
1141  | 
(* FIXME: Move all these theorems into the ML code using lemma antiquotation *)  | 
|
1142  | 
lemma norm_add_rule_thm:  | 
|
1143  | 
fixes x1 x2 :: "'a::real_normed_vector"  | 
|
1144  | 
shows "norm x1 \<le> b1 \<Longrightarrow> norm x2 \<le> b2 \<Longrightarrow> norm (x1 + x2) \<le> b1 + b2"  | 
|
1145  | 
by (rule order_trans [OF norm_triangle_ineq add_mono])  | 
|
1146  | 
||
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34964 
diff
changeset
 | 
1147  | 
lemma ge_iff_diff_ge_0: "(a::'a::linordered_ring) \<ge> b == a - b \<ge> 0"  | 
| 33175 | 1148  | 
by (simp add: ring_simps)  | 
1149  | 
||
1150  | 
lemma pth_1:  | 
|
1151  | 
fixes x :: "'a::real_normed_vector"  | 
|
1152  | 
shows "x == scaleR 1 x" by simp  | 
|
1153  | 
||
1154  | 
lemma pth_2:  | 
|
1155  | 
fixes x :: "'a::real_normed_vector"  | 
|
1156  | 
shows "x - y == x + -y" by (atomize (full)) simp  | 
|
1157  | 
||
1158  | 
lemma pth_3:  | 
|
1159  | 
fixes x :: "'a::real_normed_vector"  | 
|
1160  | 
shows "- x == scaleR (-1) x" by simp  | 
|
1161  | 
||
1162  | 
lemma pth_4:  | 
|
1163  | 
fixes x :: "'a::real_normed_vector"  | 
|
1164  | 
shows "scaleR 0 x == 0" and "scaleR c 0 = (0::'a)" by simp_all  | 
|
1165  | 
||
1166  | 
lemma pth_5:  | 
|
1167  | 
fixes x :: "'a::real_normed_vector"  | 
|
1168  | 
shows "scaleR c (scaleR d x) == scaleR (c * d) x" by simp  | 
|
1169  | 
||
1170  | 
lemma pth_6:  | 
|
1171  | 
fixes x :: "'a::real_normed_vector"  | 
|
1172  | 
shows "scaleR c (x + y) == scaleR c x + scaleR c y"  | 
|
1173  | 
by (simp add: scaleR_right_distrib)  | 
|
1174  | 
||
1175  | 
lemma pth_7:  | 
|
1176  | 
fixes x :: "'a::real_normed_vector"  | 
|
1177  | 
shows "0 + x == x" and "x + 0 == x" by simp_all  | 
|
1178  | 
||
1179  | 
lemma pth_8:  | 
|
1180  | 
fixes x :: "'a::real_normed_vector"  | 
|
1181  | 
shows "scaleR c x + scaleR d x == scaleR (c + d) x"  | 
|
1182  | 
by (simp add: scaleR_left_distrib)  | 
|
1183  | 
||
1184  | 
lemma pth_9:  | 
|
1185  | 
fixes x :: "'a::real_normed_vector" shows  | 
|
1186  | 
"(scaleR c x + z) + scaleR d x == scaleR (c + d) x + z"  | 
|
1187  | 
"scaleR c x + (scaleR d x + z) == scaleR (c + d) x + z"  | 
|
1188  | 
"(scaleR c x + w) + (scaleR d x + z) == scaleR (c + d) x + (w + z)"  | 
|
1189  | 
by (simp_all add: algebra_simps)  | 
|
1190  | 
||
1191  | 
lemma pth_a:  | 
|
1192  | 
fixes x :: "'a::real_normed_vector"  | 
|
1193  | 
shows "scaleR 0 x + y == y" by simp  | 
|
1194  | 
||
1195  | 
lemma pth_b:  | 
|
1196  | 
fixes x :: "'a::real_normed_vector" shows  | 
|
1197  | 
"scaleR c x + scaleR d y == scaleR c x + scaleR d y"  | 
|
1198  | 
"(scaleR c x + z) + scaleR d y == scaleR c x + (z + scaleR d y)"  | 
|
1199  | 
"scaleR c x + (scaleR d y + z) == scaleR c x + (scaleR d y + z)"  | 
|
1200  | 
"(scaleR c x + w) + (scaleR d y + z) == scaleR c x + (w + (scaleR d y + z))"  | 
|
1201  | 
by (simp_all add: algebra_simps)  | 
|
1202  | 
||
1203  | 
lemma pth_c:  | 
|
1204  | 
fixes x :: "'a::real_normed_vector" shows  | 
|
1205  | 
"scaleR c x + scaleR d y == scaleR d y + scaleR c x"  | 
|
1206  | 
"(scaleR c x + z) + scaleR d y == scaleR d y + (scaleR c x + z)"  | 
|
1207  | 
"scaleR c x + (scaleR d y + z) == scaleR d y + (scaleR c x + z)"  | 
|
1208  | 
"(scaleR c x + w) + (scaleR d y + z) == scaleR d y + ((scaleR c x + w) + z)"  | 
|
1209  | 
by (simp_all add: algebra_simps)  | 
|
1210  | 
||
1211  | 
lemma pth_d:  | 
|
1212  | 
fixes x :: "'a::real_normed_vector"  | 
|
1213  | 
shows "x + 0 == x" by simp  | 
|
1214  | 
||
1215  | 
lemma norm_imp_pos_and_ge:  | 
|
1216  | 
fixes x :: "'a::real_normed_vector"  | 
|
1217  | 
shows "norm x == n \<Longrightarrow> norm x \<ge> 0 \<and> n \<ge> norm x"  | 
|
1218  | 
by atomize auto  | 
|
1219  | 
||
1220  | 
lemma real_eq_0_iff_le_ge_0: "(x::real) = 0 == x \<ge> 0 \<and> -x \<ge> 0" by arith  | 
|
1221  | 
||
1222  | 
lemma norm_pths:  | 
|
1223  | 
fixes x :: "'a::real_normed_vector" shows  | 
|
1224  | 
"x = y \<longleftrightarrow> norm (x - y) \<le> 0"  | 
|
1225  | 
"x \<noteq> y \<longleftrightarrow> \<not> (norm (x - y) \<le> 0)"  | 
|
1226  | 
using norm_ge_zero[of "x - y"] by auto  | 
|
1227  | 
||
1228  | 
lemma vector_dist_norm:  | 
|
1229  | 
fixes x :: "'a::real_normed_vector"  | 
|
1230  | 
shows "dist x y = norm (x - y)"  | 
|
1231  | 
by (rule dist_norm)  | 
|
1232  | 
||
1233  | 
use "normarith.ML"  | 
|
1234  | 
||
1235  | 
method_setup norm = {* Scan.succeed (SIMPLE_METHOD' o NormArith.norm_arith_tac)
 | 
|
1236  | 
*} "Proves simple linear statements about vector norms"  | 
|
1237  | 
||
1238  | 
||
1239  | 
text{* Hence more metric properties. *}
 | 
|
1240  | 
||
1241  | 
lemma dist_triangle_alt:  | 
|
1242  | 
fixes x y z :: "'a::metric_space"  | 
|
1243  | 
shows "dist y z <= dist x y + dist x z"  | 
|
1244  | 
using dist_triangle [of y z x] by (simp add: dist_commute)  | 
|
1245  | 
||
1246  | 
lemma dist_pos_lt:  | 
|
1247  | 
fixes x y :: "'a::metric_space"  | 
|
1248  | 
shows "x \<noteq> y ==> 0 < dist x y"  | 
|
1249  | 
by (simp add: zero_less_dist_iff)  | 
|
1250  | 
||
1251  | 
lemma dist_nz:  | 
|
1252  | 
fixes x y :: "'a::metric_space"  | 
|
1253  | 
shows "x \<noteq> y \<longleftrightarrow> 0 < dist x y"  | 
|
1254  | 
by (simp add: zero_less_dist_iff)  | 
|
1255  | 
||
1256  | 
lemma dist_triangle_le:  | 
|
1257  | 
fixes x y z :: "'a::metric_space"  | 
|
1258  | 
shows "dist x z + dist y z <= e \<Longrightarrow> dist x y <= e"  | 
|
1259  | 
by (rule order_trans [OF dist_triangle2])  | 
|
1260  | 
||
1261  | 
lemma dist_triangle_lt:  | 
|
1262  | 
fixes x y z :: "'a::metric_space"  | 
|
1263  | 
shows "dist x z + dist y z < e ==> dist x y < e"  | 
|
1264  | 
by (rule le_less_trans [OF dist_triangle2])  | 
|
1265  | 
||
1266  | 
lemma dist_triangle_half_l:  | 
|
1267  | 
fixes x1 x2 y :: "'a::metric_space"  | 
|
1268  | 
shows "dist x1 y < e / 2 \<Longrightarrow> dist x2 y < e / 2 \<Longrightarrow> dist x1 x2 < e"  | 
|
1269  | 
by (rule dist_triangle_lt [where z=y], simp)  | 
|
1270  | 
||
1271  | 
lemma dist_triangle_half_r:  | 
|
1272  | 
fixes x1 x2 y :: "'a::metric_space"  | 
|
1273  | 
shows "dist y x1 < e / 2 \<Longrightarrow> dist y x2 < e / 2 \<Longrightarrow> dist x1 x2 < e"  | 
|
1274  | 
by (rule dist_triangle_half_l, simp_all add: dist_commute)  | 
|
1275  | 
||
| 
35172
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35150 
diff
changeset
 | 
1276  | 
|
| 
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35150 
diff
changeset
 | 
1277  | 
lemma norm_triangle_half_r:  | 
| 
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35150 
diff
changeset
 | 
1278  | 
shows "norm (y - x1) < e / 2 \<Longrightarrow> norm (y - x2) < e / 2 \<Longrightarrow> norm (x1 - x2) < e"  | 
| 
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35150 
diff
changeset
 | 
1279  | 
using dist_triangle_half_r unfolding vector_dist_norm[THEN sym] by auto  | 
| 
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35150 
diff
changeset
 | 
1280  | 
|
| 
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35150 
diff
changeset
 | 
1281  | 
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
 | 
1282  | 
shows "norm (x - x') < e"  | 
| 
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35150 
diff
changeset
 | 
1283  | 
using dist_triangle_half_l[OF assms[unfolded vector_dist_norm[THEN sym]]]  | 
| 
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35150 
diff
changeset
 | 
1284  | 
unfolding vector_dist_norm[THEN sym] .  | 
| 
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35150 
diff
changeset
 | 
1285  | 
|
| 
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35150 
diff
changeset
 | 
1286  | 
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
 | 
1287  | 
by (metis order_trans norm_triangle_ineq)  | 
| 
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35150 
diff
changeset
 | 
1288  | 
|
| 
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35150 
diff
changeset
 | 
1289  | 
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
 | 
1290  | 
by (metis basic_trans_rules(21) norm_triangle_ineq)  | 
| 
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35150 
diff
changeset
 | 
1291  | 
|
| 33175 | 1292  | 
lemma dist_triangle_add:  | 
1293  | 
fixes x y x' y' :: "'a::real_normed_vector"  | 
|
1294  | 
shows "dist (x + y) (x' + y') <= dist x x' + dist y y'"  | 
|
1295  | 
by norm  | 
|
1296  | 
||
1297  | 
lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \<bar>c\<bar> * dist x y"  | 
|
1298  | 
unfolding dist_norm vector_ssub_ldistrib[symmetric] norm_mul ..  | 
|
1299  | 
||
1300  | 
lemma dist_triangle_add_half:  | 
|
1301  | 
fixes x x' y y' :: "'a::real_normed_vector"  | 
|
1302  | 
shows "dist x x' < e / 2 \<Longrightarrow> dist y y' < e / 2 \<Longrightarrow> dist(x + y) (x' + y') < e"  | 
|
1303  | 
by norm  | 
|
1304  | 
||
1305  | 
lemma setsum_component [simp]:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1306  | 
  fixes f:: " 'a \<Rightarrow> ('b::comm_monoid_add) ^'n"
 | 
| 33175 | 1307  | 
shows "(setsum f S)$i = setsum (\<lambda>x. (f x)$i) S"  | 
1308  | 
by (cases "finite S", induct S set: finite, simp_all)  | 
|
1309  | 
||
1310  | 
lemma setsum_eq: "setsum f S = (\<chi> i. setsum (\<lambda>x. (f x)$i ) S)"  | 
|
1311  | 
by (simp add: Cart_eq)  | 
|
1312  | 
||
1313  | 
lemma setsum_clauses:  | 
|
1314  | 
  shows "setsum f {} = 0"
 | 
|
1315  | 
and "finite S \<Longrightarrow> setsum f (insert x S) =  | 
|
1316  | 
(if x \<in> S then setsum f S else f x + setsum f S)"  | 
|
1317  | 
by (auto simp add: insert_absorb)  | 
|
1318  | 
||
1319  | 
lemma setsum_cmul:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1320  | 
  fixes f:: "'c \<Rightarrow> ('a::semiring_1)^'n"
 | 
| 33175 | 1321  | 
shows "setsum (\<lambda>x. c *s f x) S = c *s setsum f S"  | 
1322  | 
by (simp add: Cart_eq setsum_right_distrib)  | 
|
1323  | 
||
1324  | 
lemma setsum_norm:  | 
|
1325  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"  | 
|
1326  | 
assumes fS: "finite S"  | 
|
1327  | 
shows "norm (setsum f S) <= setsum (\<lambda>x. norm(f x)) S"  | 
|
1328  | 
proof(induct rule: finite_induct[OF fS])  | 
|
1329  | 
case 1 thus ?case by simp  | 
|
1330  | 
next  | 
|
1331  | 
case (2 x S)  | 
|
1332  | 
from "2.hyps" have "norm (setsum f (insert x S)) \<le> norm (f x) + norm (setsum f S)" by (simp add: norm_triangle_ineq)  | 
|
1333  | 
also have "\<dots> \<le> norm (f x) + setsum (\<lambda>x. norm(f x)) S"  | 
|
1334  | 
using "2.hyps" by simp  | 
|
1335  | 
finally show ?case using "2.hyps" by simp  | 
|
1336  | 
qed  | 
|
1337  | 
||
1338  | 
lemma real_setsum_norm:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1339  | 
fixes f :: "'a \<Rightarrow> real ^'n"  | 
| 33175 | 1340  | 
assumes fS: "finite S"  | 
1341  | 
shows "norm (setsum f S) <= setsum (\<lambda>x. norm(f x)) S"  | 
|
1342  | 
proof(induct rule: finite_induct[OF fS])  | 
|
1343  | 
case 1 thus ?case by simp  | 
|
1344  | 
next  | 
|
1345  | 
case (2 x S)  | 
|
1346  | 
from "2.hyps" have "norm (setsum f (insert x S)) \<le> norm (f x) + norm (setsum f S)" by (simp add: norm_triangle_ineq)  | 
|
1347  | 
also have "\<dots> \<le> norm (f x) + setsum (\<lambda>x. norm(f x)) S"  | 
|
1348  | 
using "2.hyps" by simp  | 
|
1349  | 
finally show ?case using "2.hyps" by simp  | 
|
1350  | 
qed  | 
|
1351  | 
||
1352  | 
lemma setsum_norm_le:  | 
|
1353  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"  | 
|
1354  | 
assumes fS: "finite S"  | 
|
1355  | 
and fg: "\<forall>x \<in> S. norm (f x) \<le> g x"  | 
|
1356  | 
shows "norm (setsum f S) \<le> setsum g S"  | 
|
1357  | 
proof-  | 
|
1358  | 
from fg have "setsum (\<lambda>x. norm(f x)) S <= setsum g S"  | 
|
1359  | 
by - (rule setsum_mono, simp)  | 
|
1360  | 
then show ?thesis using setsum_norm[OF fS, of f] fg  | 
|
1361  | 
by arith  | 
|
1362  | 
qed  | 
|
1363  | 
||
1364  | 
lemma real_setsum_norm_le:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1365  | 
fixes f :: "'a \<Rightarrow> real ^ 'n"  | 
| 33175 | 1366  | 
assumes fS: "finite S"  | 
1367  | 
and fg: "\<forall>x \<in> S. norm (f x) \<le> g x"  | 
|
1368  | 
shows "norm (setsum f S) \<le> setsum g S"  | 
|
1369  | 
proof-  | 
|
1370  | 
from fg have "setsum (\<lambda>x. norm(f x)) S <= setsum g S"  | 
|
1371  | 
by - (rule setsum_mono, simp)  | 
|
1372  | 
then show ?thesis using real_setsum_norm[OF fS, of f] fg  | 
|
1373  | 
by arith  | 
|
1374  | 
qed  | 
|
1375  | 
||
1376  | 
lemma setsum_norm_bound:  | 
|
1377  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"  | 
|
1378  | 
assumes fS: "finite S"  | 
|
1379  | 
and K: "\<forall>x \<in> S. norm (f x) \<le> K"  | 
|
1380  | 
shows "norm (setsum f S) \<le> of_nat (card S) * K"  | 
|
1381  | 
using setsum_norm_le[OF fS K] setsum_constant[symmetric]  | 
|
1382  | 
by simp  | 
|
1383  | 
||
1384  | 
lemma real_setsum_norm_bound:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1385  | 
fixes f :: "'a \<Rightarrow> real ^ 'n"  | 
| 33175 | 1386  | 
assumes fS: "finite S"  | 
1387  | 
and K: "\<forall>x \<in> S. norm (f x) \<le> K"  | 
|
1388  | 
shows "norm (setsum f S) \<le> of_nat (card S) * K"  | 
|
1389  | 
using real_setsum_norm_le[OF fS K] setsum_constant[symmetric]  | 
|
1390  | 
by simp  | 
|
1391  | 
||
1392  | 
lemma setsum_vmul:  | 
|
1393  | 
  fixes f :: "'a \<Rightarrow> 'b::{real_normed_vector,semiring, mult_zero}"
 | 
|
1394  | 
assumes fS: "finite S"  | 
|
1395  | 
shows "setsum f S *s v = setsum (\<lambda>x. f x *s v) S"  | 
|
1396  | 
proof(induct rule: finite_induct[OF fS])  | 
|
1397  | 
case 1 then show ?case by (simp add: vector_smult_lzero)  | 
|
1398  | 
next  | 
|
1399  | 
case (2 x F)  | 
|
1400  | 
from "2.hyps" have "setsum f (insert x F) *s v = (f x + setsum f F) *s v"  | 
|
1401  | 
by simp  | 
|
1402  | 
also have "\<dots> = f x *s v + setsum f F *s v"  | 
|
1403  | 
by (simp add: vector_sadd_rdistrib)  | 
|
1404  | 
also have "\<dots> = setsum (\<lambda>x. f x *s v) (insert x F)" using "2.hyps" by simp  | 
|
1405  | 
finally show ?case .  | 
|
1406  | 
qed  | 
|
1407  | 
||
1408  | 
(* FIXME : Problem thm setsum_vmul[of _ "f:: 'a \<Rightarrow> real ^'n"] ---  | 
|
1409  | 
Get rid of *s and use real_vector instead! Also prove that ^ creates a real_vector !! *)  | 
|
1410  | 
||
1411  | 
(* FIXME: Here too need stupid finiteness assumption on T!!! *)  | 
|
1412  | 
lemma setsum_group:  | 
|
1413  | 
assumes fS: "finite S" and fT: "finite T" and fST: "f ` S \<subseteq> T"  | 
|
1414  | 
  shows "setsum (\<lambda>y. setsum g {x. x\<in> S \<and> f x = y}) T = setsum g S"
 | 
|
1415  | 
||
1416  | 
apply (subst setsum_image_gen[OF fS, of g f])  | 
|
1417  | 
apply (rule setsum_mono_zero_right[OF fT fST])  | 
|
1418  | 
by (auto intro: setsum_0')  | 
|
1419  | 
||
1420  | 
lemma vsum_norm_allsubsets_bound:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1421  | 
fixes f:: "'a \<Rightarrow> real ^'n"  | 
| 33175 | 1422  | 
assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (setsum f Q) \<le> e"  | 
1423  | 
  shows "setsum (\<lambda>x. norm (f x)) P \<le> 2 * real CARD('n) *  e"
 | 
|
1424  | 
proof-  | 
|
1425  | 
  let ?d = "real CARD('n)"
 | 
|
1426  | 
let ?nf = "\<lambda>x. norm (f x)"  | 
|
1427  | 
let ?U = "UNIV :: 'n set"  | 
|
1428  | 
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"  | 
|
1429  | 
by (rule setsum_commute)  | 
|
1430  | 
have th1: "2 * ?d * e = of_nat (card ?U) * (2 * e)" by (simp add: real_of_nat_def)  | 
|
1431  | 
have "setsum ?nf P \<le> setsum (\<lambda>x. setsum (\<lambda>i. \<bar>f x $ i\<bar>) ?U) P"  | 
|
1432  | 
apply (rule setsum_mono)  | 
|
1433  | 
by (rule norm_le_l1)  | 
|
1434  | 
also have "\<dots> \<le> 2 * ?d * e"  | 
|
1435  | 
unfolding th0 th1  | 
|
1436  | 
proof(rule setsum_bounded)  | 
|
1437  | 
fix i assume i: "i \<in> ?U"  | 
|
1438  | 
    let ?Pp = "{x. x\<in> P \<and> f x $ i \<ge> 0}"
 | 
|
1439  | 
    let ?Pn = "{x. x \<in> P \<and> f x $ i < 0}"
 | 
|
1440  | 
have thp: "P = ?Pp \<union> ?Pn" by auto  | 
|
1441  | 
    have thp0: "?Pp \<inter> ?Pn ={}" by auto
 | 
|
1442  | 
have PpP: "?Pp \<subseteq> P" and PnP: "?Pn \<subseteq> P" by blast+  | 
|
1443  | 
have Ppe:"setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pp \<le> e"  | 
|
1444  | 
using component_le_norm[of "setsum (\<lambda>x. f x) ?Pp" i] fPs[OF PpP]  | 
|
1445  | 
by (auto intro: abs_le_D1)  | 
|
1446  | 
have Pne: "setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pn \<le> e"  | 
|
1447  | 
using component_le_norm[of "setsum (\<lambda>x. - f x) ?Pn" i] fPs[OF PnP]  | 
|
1448  | 
by (auto simp add: setsum_negf intro: abs_le_D1)  | 
|
1449  | 
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"  | 
|
1450  | 
apply (subst thp)  | 
|
1451  | 
apply (rule setsum_Un_zero)  | 
|
1452  | 
using fP thp0 by auto  | 
|
1453  | 
also have "\<dots> \<le> 2*e" using Pne Ppe by arith  | 
|
1454  | 
finally show "setsum (\<lambda>x. \<bar>f x $ i\<bar>) P \<le> 2*e" .  | 
|
1455  | 
qed  | 
|
1456  | 
finally show ?thesis .  | 
|
1457  | 
qed  | 
|
1458  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1459  | 
lemma dot_lsum: "finite S \<Longrightarrow> setsum f S \<bullet> (y::'a::{comm_ring}^'n) = setsum (\<lambda>x. f x \<bullet> y) S "
 | 
| 33175 | 1460  | 
by (induct rule: finite_induct, auto simp add: dot_lzero dot_ladd dot_radd)  | 
1461  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1462  | 
lemma dot_rsum: "finite S \<Longrightarrow> (y::'a::{comm_ring}^'n) \<bullet> setsum f S = setsum (\<lambda>x. y \<bullet> f x) S "
 | 
| 33175 | 1463  | 
by (induct rule: finite_induct, auto simp add: dot_rzero dot_radd)  | 
1464  | 
||
1465  | 
subsection{* Basis vectors in coordinate directions. *}
 | 
|
1466  | 
||
1467  | 
||
1468  | 
definition "basis k = (\<chi> i. if i = k then 1 else 0)"  | 
|
1469  | 
||
1470  | 
lemma basis_component [simp]: "basis k $ i = (if k=i then 1 else 0)"  | 
|
1471  | 
unfolding basis_def by simp  | 
|
1472  | 
||
1473  | 
lemma delta_mult_idempotent:  | 
|
1474  | 
"(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)" by (cases "k=a", auto)  | 
|
1475  | 
||
1476  | 
lemma norm_basis:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1477  | 
shows "norm (basis k :: real ^'n) = 1"  | 
| 33175 | 1478  | 
apply (simp add: basis_def real_vector_norm_def dot_def)  | 
1479  | 
apply (vector delta_mult_idempotent)  | 
|
1480  | 
using setsum_delta[of "UNIV :: 'n set" "k" "\<lambda>k. 1::real"]  | 
|
1481  | 
apply auto  | 
|
1482  | 
done  | 
|
1483  | 
||
1484  | 
lemma norm_basis_1: "norm(basis 1 :: real ^'n::{finite,one}) = 1"
 | 
|
1485  | 
by (rule norm_basis)  | 
|
1486  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1487  | 
lemma vector_choose_size: "0 <= c ==> \<exists>(x::real^'n). norm x = c"  | 
| 33175 | 1488  | 
apply (rule exI[where x="c *s basis arbitrary"])  | 
1489  | 
by (simp only: norm_mul norm_basis)  | 
|
1490  | 
||
1491  | 
lemma vector_choose_dist: assumes e: "0 <= e"  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1492  | 
shows "\<exists>(y::real^'n). dist x y = e"  | 
| 33175 | 1493  | 
proof-  | 
1494  | 
from vector_choose_size[OF e] obtain c:: "real ^'n" where "norm c = e"  | 
|
1495  | 
by blast  | 
|
1496  | 
then have "dist x (x - c) = e" by (simp add: dist_norm)  | 
|
1497  | 
then show ?thesis by blast  | 
|
1498  | 
qed  | 
|
1499  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1500  | 
lemma basis_inj: "inj (basis :: 'n \<Rightarrow> real ^'n)"  | 
| 33175 | 1501  | 
by (simp add: inj_on_def Cart_eq)  | 
1502  | 
||
1503  | 
lemma cond_value_iff: "f (if b then x else y) = (if b then f x else f y)"  | 
|
1504  | 
by auto  | 
|
1505  | 
||
1506  | 
lemma basis_expansion:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1507  | 
  "setsum (\<lambda>i. (x$i) *s basis i) UNIV = (x::('a::ring_1) ^'n)" (is "?lhs = ?rhs" is "setsum ?f ?S = _")
 | 
| 33175 | 1508  | 
by (auto simp add: Cart_eq cond_value_iff setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong)  | 
1509  | 
||
1510  | 
lemma basis_expansion_unique:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1511  | 
  "setsum (\<lambda>i. f i *s basis i) UNIV = (x::('a::comm_ring_1) ^'n) \<longleftrightarrow> (\<forall>i. f i = x$i)"
 | 
| 33175 | 1512  | 
by (simp add: Cart_eq setsum_delta cond_value_iff cong del: if_weak_cong)  | 
1513  | 
||
1514  | 
lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)"  | 
|
1515  | 
by auto  | 
|
1516  | 
||
1517  | 
lemma dot_basis:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1518  | 
shows "basis i \<bullet> x = x$i" "x \<bullet> (basis i :: 'a^'n) = (x$i :: 'a::semiring_1)"  | 
| 33175 | 1519  | 
by (auto simp add: dot_def basis_def cond_application_beta cond_value_iff setsum_delta cong del: if_weak_cong)  | 
1520  | 
||
1521  | 
lemma inner_basis:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1522  | 
  fixes x :: "'a::{real_inner, real_algebra_1} ^ 'n"
 | 
| 33175 | 1523  | 
shows "inner (basis i) x = inner 1 (x $ i)"  | 
1524  | 
and "inner x (basis i) = inner (x $ i) 1"  | 
|
1525  | 
unfolding inner_vector_def basis_def  | 
|
1526  | 
by (auto simp add: cond_application_beta cond_value_iff setsum_delta cong del: if_weak_cong)  | 
|
1527  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1528  | 
lemma basis_eq_0: "basis i = (0::'a::semiring_1^'n) \<longleftrightarrow> False"  | 
| 33175 | 1529  | 
by (auto simp add: Cart_eq)  | 
1530  | 
||
1531  | 
lemma basis_nonzero:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1532  | 
shows "basis k \<noteq> (0:: 'a::semiring_1 ^'n)"  | 
| 33175 | 1533  | 
by (simp add: basis_eq_0)  | 
1534  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1535  | 
lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = (z::'a::semiring_1^'n)"  | 
| 33175 | 1536  | 
apply (auto simp add: Cart_eq dot_basis)  | 
1537  | 
apply (erule_tac x="basis i" in allE)  | 
|
1538  | 
apply (simp add: dot_basis)  | 
|
1539  | 
apply (subgoal_tac "y = z")  | 
|
1540  | 
apply simp  | 
|
1541  | 
apply (simp add: Cart_eq)  | 
|
1542  | 
done  | 
|
1543  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1544  | 
lemma vector_eq_rdot: "(\<forall>z. x \<bullet> z = y \<bullet> z) \<longleftrightarrow> x = (y::'a::semiring_1^'n)"  | 
| 33175 | 1545  | 
apply (auto simp add: Cart_eq dot_basis)  | 
1546  | 
apply (erule_tac x="basis i" in allE)  | 
|
1547  | 
apply (simp add: dot_basis)  | 
|
1548  | 
apply (subgoal_tac "x = y")  | 
|
1549  | 
apply simp  | 
|
1550  | 
apply (simp add: Cart_eq)  | 
|
1551  | 
done  | 
|
1552  | 
||
1553  | 
subsection{* Orthogonality. *}
 | 
|
1554  | 
||
1555  | 
definition "orthogonal x y \<longleftrightarrow> (x \<bullet> y = 0)"  | 
|
1556  | 
||
1557  | 
lemma orthogonal_basis:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1558  | 
shows "orthogonal (basis i :: 'a^'n) x \<longleftrightarrow> x$i = (0::'a::ring_1)"  | 
| 33175 | 1559  | 
by (auto simp add: orthogonal_def dot_def basis_def cond_value_iff cond_application_beta setsum_delta cong del: if_weak_cong)  | 
1560  | 
||
1561  | 
lemma orthogonal_basis_basis:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1562  | 
shows "orthogonal (basis i :: 'a::ring_1^'n) (basis j) \<longleftrightarrow> i \<noteq> j"  | 
| 33175 | 1563  | 
unfolding orthogonal_basis[of i] basis_component[of j] by simp  | 
1564  | 
||
1565  | 
(* FIXME : Maybe some of these require less than comm_ring, but not all*)  | 
|
1566  | 
lemma orthogonal_clauses:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1567  | 
"orthogonal a (0::'a::comm_ring ^'n)"  | 
| 33175 | 1568  | 
"orthogonal a x ==> orthogonal a (c *s x)"  | 
1569  | 
"orthogonal a x ==> orthogonal a (-x)"  | 
|
1570  | 
"orthogonal a x \<Longrightarrow> orthogonal a y ==> orthogonal a (x + y)"  | 
|
1571  | 
"orthogonal a x \<Longrightarrow> orthogonal a y ==> orthogonal a (x - y)"  | 
|
1572  | 
"orthogonal 0 a"  | 
|
1573  | 
"orthogonal x a ==> orthogonal (c *s x) a"  | 
|
1574  | 
"orthogonal x a ==> orthogonal (-x) a"  | 
|
1575  | 
"orthogonal x a \<Longrightarrow> orthogonal y a ==> orthogonal (x + y) a"  | 
|
1576  | 
"orthogonal x a \<Longrightarrow> orthogonal y a ==> orthogonal (x - y) a"  | 
|
1577  | 
unfolding orthogonal_def dot_rneg dot_rmult dot_radd dot_rsub  | 
|
1578  | 
dot_lzero dot_rzero dot_lneg dot_lmult dot_ladd dot_lsub  | 
|
1579  | 
by simp_all  | 
|
1580  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1581  | 
lemma orthogonal_commute: "orthogonal (x::'a::{ab_semigroup_mult,comm_monoid_add} ^'n)y \<longleftrightarrow> orthogonal y x"
 | 
| 33175 | 1582  | 
by (simp add: orthogonal_def dot_sym)  | 
1583  | 
||
1584  | 
subsection{* Explicit vector construction from lists. *}
 | 
|
1585  | 
||
1586  | 
primrec from_nat :: "nat \<Rightarrow> 'a::{monoid_add,one}"
 | 
|
1587  | 
where "from_nat 0 = 0" | "from_nat (Suc n) = 1 + from_nat n"  | 
|
1588  | 
||
1589  | 
lemma from_nat [simp]: "from_nat = of_nat"  | 
|
1590  | 
by (rule ext, induct_tac x, simp_all)  | 
|
1591  | 
||
1592  | 
primrec  | 
|
1593  | 
list_fun :: "nat \<Rightarrow> _ list \<Rightarrow> _ \<Rightarrow> _"  | 
|
1594  | 
where  | 
|
1595  | 
"list_fun n [] = (\<lambda>x. 0)"  | 
|
1596  | 
| "list_fun n (x # xs) = fun_upd (list_fun (Suc n) xs) (from_nat n) x"  | 
|
1597  | 
||
1598  | 
definition "vector l = (\<chi> i. list_fun 1 l i)"  | 
|
1599  | 
(*definition "vector l = (\<chi> i. if i <= length l then l ! (i - 1) else 0)"*)  | 
|
1600  | 
||
1601  | 
lemma vector_1: "(vector[x]) $1 = x"  | 
|
1602  | 
unfolding vector_def by simp  | 
|
1603  | 
||
1604  | 
lemma vector_2:  | 
|
1605  | 
"(vector[x,y]) $1 = x"  | 
|
1606  | 
"(vector[x,y] :: 'a^2)$2 = (y::'a::zero)"  | 
|
1607  | 
unfolding vector_def by simp_all  | 
|
1608  | 
||
1609  | 
lemma vector_3:  | 
|
1610  | 
 "(vector [x,y,z] ::('a::zero)^3)$1 = x"
 | 
|
1611  | 
 "(vector [x,y,z] ::('a::zero)^3)$2 = y"
 | 
|
1612  | 
 "(vector [x,y,z] ::('a::zero)^3)$3 = z"
 | 
|
1613  | 
unfolding vector_def by simp_all  | 
|
1614  | 
||
1615  | 
lemma forall_vector_1: "(\<forall>v::'a::zero^1. P v) \<longleftrightarrow> (\<forall>x. P(vector[x]))"  | 
|
1616  | 
apply auto  | 
|
1617  | 
apply (erule_tac x="v$1" in allE)  | 
|
1618  | 
apply (subgoal_tac "vector [v$1] = v")  | 
|
1619  | 
apply simp  | 
|
1620  | 
apply (vector vector_def)  | 
|
1621  | 
apply (simp add: forall_1)  | 
|
1622  | 
done  | 
|
1623  | 
||
1624  | 
lemma forall_vector_2: "(\<forall>v::'a::zero^2. P v) \<longleftrightarrow> (\<forall>x y. P(vector[x, y]))"  | 
|
1625  | 
apply auto  | 
|
1626  | 
apply (erule_tac x="v$1" in allE)  | 
|
1627  | 
apply (erule_tac x="v$2" in allE)  | 
|
1628  | 
apply (subgoal_tac "vector [v$1, v$2] = v")  | 
|
1629  | 
apply simp  | 
|
1630  | 
apply (vector vector_def)  | 
|
1631  | 
apply (simp add: forall_2)  | 
|
1632  | 
done  | 
|
1633  | 
||
1634  | 
lemma forall_vector_3: "(\<forall>v::'a::zero^3. P v) \<longleftrightarrow> (\<forall>x y z. P(vector[x, y, z]))"  | 
|
1635  | 
apply auto  | 
|
1636  | 
apply (erule_tac x="v$1" in allE)  | 
|
1637  | 
apply (erule_tac x="v$2" in allE)  | 
|
1638  | 
apply (erule_tac x="v$3" in allE)  | 
|
1639  | 
apply (subgoal_tac "vector [v$1, v$2, v$3] = v")  | 
|
1640  | 
apply simp  | 
|
1641  | 
apply (vector vector_def)  | 
|
1642  | 
apply (simp add: forall_3)  | 
|
1643  | 
done  | 
|
1644  | 
||
1645  | 
subsection{* Linear functions. *}
 | 
|
1646  | 
||
1647  | 
definition "linear f \<longleftrightarrow> (\<forall>x y. f(x + y) = f x + f y) \<and> (\<forall>c x. f(c *s x) = c *s f x)"  | 
|
1648  | 
||
| 
33714
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
1649  | 
lemma linearI: assumes "\<And>x y. f (x + y) = f x + f y" "\<And>c x. f (c *s x) = c *s f x"  | 
| 
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
1650  | 
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
 | 
1651  | 
|
| 33175 | 1652  | 
lemma linear_compose_cmul: "linear f ==> linear (\<lambda>x. (c::'a::comm_semiring) *s f x)"  | 
1653  | 
by (vector linear_def Cart_eq ring_simps)  | 
|
1654  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1655  | 
lemma linear_compose_neg: "linear (f :: 'a ^'n \<Rightarrow> 'a::comm_ring ^'m) ==> linear (\<lambda>x. -(f(x)))" by (vector linear_def Cart_eq)  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1656  | 
|
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1657  | 
lemma linear_compose_add: "linear (f :: 'a ^'n \<Rightarrow> 'a::semiring_1 ^'m) \<Longrightarrow> linear g ==> linear (\<lambda>x. f(x) + g(x))"  | 
| 33175 | 1658  | 
by (vector linear_def Cart_eq ring_simps)  | 
1659  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1660  | 
lemma linear_compose_sub: "linear (f :: 'a ^'n \<Rightarrow> 'a::ring_1 ^'m) \<Longrightarrow> linear g ==> linear (\<lambda>x. f x - g x)"  | 
| 33175 | 1661  | 
by (vector linear_def Cart_eq ring_simps)  | 
1662  | 
||
1663  | 
lemma linear_compose: "linear f \<Longrightarrow> linear g ==> linear (g o f)"  | 
|
1664  | 
by (simp add: linear_def)  | 
|
1665  | 
||
1666  | 
lemma linear_id: "linear id" by (simp add: linear_def id_def)  | 
|
1667  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1668  | 
lemma linear_zero: "linear (\<lambda>x. 0::'a::semiring_1 ^ 'n)" by (simp add: linear_def)  | 
| 33175 | 1669  | 
|
1670  | 
lemma linear_compose_setsum:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1671  | 
assumes fS: "finite S" and lS: "\<forall>a \<in> S. linear (f a :: 'a::semiring_1 ^ 'n \<Rightarrow> 'a ^'m)"  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1672  | 
shows "linear(\<lambda>x. setsum (\<lambda>a. f a x :: 'a::semiring_1 ^'m) S)"  | 
| 33175 | 1673  | 
using lS  | 
1674  | 
apply (induct rule: finite_induct[OF fS])  | 
|
1675  | 
by (auto simp add: linear_zero intro: linear_compose_add)  | 
|
1676  | 
||
1677  | 
lemma linear_vmul_component:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1678  | 
fixes f:: "'a::semiring_1^'m \<Rightarrow> 'a^'n"  | 
| 33175 | 1679  | 
assumes lf: "linear f"  | 
1680  | 
shows "linear (\<lambda>x. f x $ k *s v)"  | 
|
1681  | 
using lf  | 
|
1682  | 
apply (auto simp add: linear_def )  | 
|
1683  | 
by (vector ring_simps)+  | 
|
1684  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1685  | 
lemma linear_0: "linear f ==> f 0 = (0::'a::semiring_1 ^'n)"  | 
| 33175 | 1686  | 
unfolding linear_def  | 
1687  | 
apply clarsimp  | 
|
1688  | 
apply (erule allE[where x="0::'a"])  | 
|
1689  | 
apply simp  | 
|
1690  | 
done  | 
|
1691  | 
||
1692  | 
lemma linear_cmul: "linear f ==> f(c*s x) = c *s f x" by (simp add: linear_def)  | 
|
1693  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1694  | 
lemma linear_neg: "linear (f :: 'a::ring_1 ^'n \<Rightarrow> _) ==> f (-x) = - f x"  | 
| 33175 | 1695  | 
unfolding vector_sneg_minus1  | 
1696  | 
using linear_cmul[of f] by auto  | 
|
1697  | 
||
1698  | 
lemma linear_add: "linear f ==> f(x + y) = f x + f y" by (metis linear_def)  | 
|
1699  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1700  | 
lemma linear_sub: "linear (f::'a::ring_1 ^'n \<Rightarrow> _) ==> f(x - y) = f x - f y"  | 
| 33175 | 1701  | 
by (simp add: diff_def linear_add linear_neg)  | 
1702  | 
||
1703  | 
lemma linear_setsum:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1704  | 
fixes f:: "'a::semiring_1^'n \<Rightarrow> _"  | 
| 33175 | 1705  | 
assumes lf: "linear f" and fS: "finite S"  | 
1706  | 
shows "f (setsum g S) = setsum (f o g) S"  | 
|
1707  | 
proof (induct rule: finite_induct[OF fS])  | 
|
1708  | 
case 1 thus ?case by (simp add: linear_0[OF lf])  | 
|
1709  | 
next  | 
|
1710  | 
case (2 x F)  | 
|
1711  | 
have "f (setsum g (insert x F)) = f (g x + setsum g F)" using "2.hyps"  | 
|
1712  | 
by simp  | 
|
1713  | 
also have "\<dots> = f (g x) + f (setsum g F)" using linear_add[OF lf] by simp  | 
|
1714  | 
also have "\<dots> = setsum (f o g) (insert x F)" using "2.hyps" by simp  | 
|
1715  | 
finally show ?case .  | 
|
1716  | 
qed  | 
|
1717  | 
||
1718  | 
lemma linear_setsum_mul:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1719  | 
fixes f:: "'a ^'n \<Rightarrow> 'a::semiring_1^'m"  | 
| 33175 | 1720  | 
assumes lf: "linear f" and fS: "finite S"  | 
1721  | 
shows "f (setsum (\<lambda>i. c i *s v i) S) = setsum (\<lambda>i. c i *s f (v i)) S"  | 
|
1722  | 
using linear_setsum[OF lf fS, of "\<lambda>i. c i *s v i" , unfolded o_def]  | 
|
1723  | 
linear_cmul[OF lf] by simp  | 
|
1724  | 
||
1725  | 
lemma linear_injective_0:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1726  | 
assumes lf: "linear (f:: 'a::ring_1 ^ 'n \<Rightarrow> _)"  | 
| 33175 | 1727  | 
shows "inj f \<longleftrightarrow> (\<forall>x. f x = 0 \<longrightarrow> x = 0)"  | 
1728  | 
proof-  | 
|
1729  | 
have "inj f \<longleftrightarrow> (\<forall> x y. f x = f y \<longrightarrow> x = y)" by (simp add: inj_on_def)  | 
|
1730  | 
also have "\<dots> \<longleftrightarrow> (\<forall> x y. f x - f y = 0 \<longrightarrow> x - y = 0)" by simp  | 
|
1731  | 
also have "\<dots> \<longleftrightarrow> (\<forall> x y. f (x - y) = 0 \<longrightarrow> x - y = 0)"  | 
|
1732  | 
by (simp add: linear_sub[OF lf])  | 
|
1733  | 
also have "\<dots> \<longleftrightarrow> (\<forall> x. f x = 0 \<longrightarrow> x = 0)" by auto  | 
|
1734  | 
finally show ?thesis .  | 
|
1735  | 
qed  | 
|
1736  | 
||
1737  | 
lemma linear_bounded:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1738  | 
fixes f:: "real ^'m \<Rightarrow> real ^'n"  | 
| 33175 | 1739  | 
assumes lf: "linear f"  | 
1740  | 
shows "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x"  | 
|
1741  | 
proof-  | 
|
1742  | 
let ?S = "UNIV:: 'm set"  | 
|
1743  | 
let ?B = "setsum (\<lambda>i. norm(f(basis i))) ?S"  | 
|
1744  | 
have fS: "finite ?S" by simp  | 
|
1745  | 
  {fix x:: "real ^ 'm"
 | 
|
1746  | 
let ?g = "(\<lambda>i. (x$i) *s (basis i) :: real ^ 'm)"  | 
|
1747  | 
have "norm (f x) = norm (f (setsum (\<lambda>i. (x$i) *s (basis i)) ?S))"  | 
|
1748  | 
by (simp only: basis_expansion)  | 
|
1749  | 
also have "\<dots> = norm (setsum (\<lambda>i. (x$i) *s f (basis i))?S)"  | 
|
1750  | 
using linear_setsum[OF lf fS, of ?g, unfolded o_def] linear_cmul[OF lf]  | 
|
1751  | 
by auto  | 
|
1752  | 
finally have th0: "norm (f x) = norm (setsum (\<lambda>i. (x$i) *s f (basis i))?S)" .  | 
|
1753  | 
    {fix i assume i: "i \<in> ?S"
 | 
|
1754  | 
from component_le_norm[of x i]  | 
|
1755  | 
have "norm ((x$i) *s f (basis i :: real ^'m)) \<le> norm (f (basis i)) * norm x"  | 
|
1756  | 
unfolding norm_mul  | 
|
1757  | 
apply (simp only: mult_commute)  | 
|
1758  | 
apply (rule mult_mono)  | 
|
1759  | 
by (auto simp add: ring_simps norm_ge_zero) }  | 
|
1760  | 
then have th: "\<forall>i\<in> ?S. norm ((x$i) *s f (basis i :: real ^'m)) \<le> norm (f (basis i)) * norm x" by metis  | 
|
1761  | 
from real_setsum_norm_le[OF fS, of "\<lambda>i. (x$i) *s (f (basis i))", OF th]  | 
|
1762  | 
have "norm (f x) \<le> ?B * norm x" unfolding th0 setsum_left_distrib by metis}  | 
|
1763  | 
then show ?thesis by blast  | 
|
1764  | 
qed  | 
|
1765  | 
||
1766  | 
lemma linear_bounded_pos:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1767  | 
fixes f:: "real ^'n \<Rightarrow> real ^'m"  | 
| 33175 | 1768  | 
assumes lf: "linear f"  | 
1769  | 
shows "\<exists>B > 0. \<forall>x. norm (f x) \<le> B * norm x"  | 
|
1770  | 
proof-  | 
|
1771  | 
from linear_bounded[OF lf] obtain B where  | 
|
1772  | 
B: "\<forall>x. norm (f x) \<le> B * norm x" by blast  | 
|
1773  | 
let ?K = "\<bar>B\<bar> + 1"  | 
|
1774  | 
have Kp: "?K > 0" by arith  | 
|
1775  | 
    {assume C: "B < 0"
 | 
|
1776  | 
have "norm (1::real ^ 'n) > 0" by (simp add: zero_less_norm_iff)  | 
|
1777  | 
with C have "B * norm (1:: real ^ 'n) < 0"  | 
|
1778  | 
by (simp add: zero_compare_simps)  | 
|
1779  | 
with B[rule_format, of 1] norm_ge_zero[of "f 1"] have False by simp  | 
|
1780  | 
}  | 
|
1781  | 
then have Bp: "B \<ge> 0" by ferrack  | 
|
1782  | 
    {fix x::"real ^ 'n"
 | 
|
1783  | 
have "norm (f x) \<le> ?K * norm x"  | 
|
1784  | 
using B[rule_format, of x] norm_ge_zero[of x] norm_ge_zero[of "f x"] Bp  | 
|
1785  | 
apply (auto simp add: ring_simps split add: abs_split)  | 
|
1786  | 
apply (erule order_trans, simp)  | 
|
1787  | 
done  | 
|
1788  | 
}  | 
|
1789  | 
then show ?thesis using Kp by blast  | 
|
1790  | 
qed  | 
|
1791  | 
||
1792  | 
lemma smult_conv_scaleR: "c *s x = scaleR c x"  | 
|
1793  | 
unfolding vector_scalar_mult_def vector_scaleR_def by simp  | 
|
1794  | 
||
1795  | 
lemma linear_conv_bounded_linear:  | 
|
1796  | 
fixes f :: "real ^ _ \<Rightarrow> real ^ _"  | 
|
1797  | 
shows "linear f \<longleftrightarrow> bounded_linear f"  | 
|
1798  | 
proof  | 
|
1799  | 
assume "linear f"  | 
|
1800  | 
show "bounded_linear f"  | 
|
1801  | 
proof  | 
|
1802  | 
fix x y show "f (x + y) = f x + f y"  | 
|
1803  | 
using `linear f` unfolding linear_def by simp  | 
|
1804  | 
next  | 
|
1805  | 
fix r x show "f (scaleR r x) = scaleR r (f x)"  | 
|
1806  | 
using `linear f` unfolding linear_def  | 
|
1807  | 
by (simp add: smult_conv_scaleR)  | 
|
1808  | 
next  | 
|
1809  | 
have "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x"  | 
|
1810  | 
using `linear f` by (rule linear_bounded)  | 
|
1811  | 
thus "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"  | 
|
1812  | 
by (simp add: mult_commute)  | 
|
1813  | 
qed  | 
|
1814  | 
next  | 
|
1815  | 
assume "bounded_linear f"  | 
|
1816  | 
then interpret f: bounded_linear f .  | 
|
1817  | 
show "linear f"  | 
|
1818  | 
unfolding linear_def smult_conv_scaleR  | 
|
1819  | 
by (simp add: f.add f.scaleR)  | 
|
1820  | 
qed  | 
|
1821  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1822  | 
lemma bounded_linearI': fixes f::"real^'n \<Rightarrow> real^'m"  | 
| 
33714
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
1823  | 
assumes "\<And>x y. f (x + y) = f x + f y" "\<And>c x. f (c *s x) = c *s f x"  | 
| 
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
1824  | 
shows "bounded_linear f" unfolding linear_conv_bounded_linear[THEN sym]  | 
| 
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
1825  | 
by(rule linearI[OF assms])  | 
| 
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
1826  | 
|
| 33175 | 1827  | 
subsection{* Bilinear functions. *}
 | 
1828  | 
||
1829  | 
definition "bilinear f \<longleftrightarrow> (\<forall>x. linear(\<lambda>y. f x y)) \<and> (\<forall>y. linear(\<lambda>x. f x y))"  | 
|
1830  | 
||
1831  | 
lemma bilinear_ladd: "bilinear h ==> h (x + y) z = (h x z) + (h y z)"  | 
|
1832  | 
by (simp add: bilinear_def linear_def)  | 
|
1833  | 
lemma bilinear_radd: "bilinear h ==> h x (y + z) = (h x y) + (h x z)"  | 
|
1834  | 
by (simp add: bilinear_def linear_def)  | 
|
1835  | 
||
1836  | 
lemma bilinear_lmul: "bilinear h ==> h (c *s x) y = c *s (h x y)"  | 
|
1837  | 
by (simp add: bilinear_def linear_def)  | 
|
1838  | 
||
1839  | 
lemma bilinear_rmul: "bilinear h ==> h x (c *s y) = c *s (h x y)"  | 
|
1840  | 
by (simp add: bilinear_def linear_def)  | 
|
1841  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1842  | 
lemma bilinear_lneg: "bilinear h ==> h (- (x:: 'a::ring_1 ^ 'n)) y = -(h x y)"  | 
| 33175 | 1843  | 
by (simp only: vector_sneg_minus1 bilinear_lmul)  | 
1844  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1845  | 
lemma bilinear_rneg: "bilinear h ==> h x (- (y:: 'a::ring_1 ^ 'n)) = - h x y"  | 
| 33175 | 1846  | 
by (simp only: vector_sneg_minus1 bilinear_rmul)  | 
1847  | 
||
1848  | 
lemma (in ab_group_add) eq_add_iff: "x = x + y \<longleftrightarrow> y = 0"  | 
|
1849  | 
using add_imp_eq[of x y 0] by auto  | 
|
1850  | 
||
1851  | 
lemma bilinear_lzero:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1852  | 
fixes h :: "'a::ring^'n \<Rightarrow> _" assumes bh: "bilinear h" shows "h 0 x = 0"  | 
| 33175 | 1853  | 
using bilinear_ladd[OF bh, of 0 0 x]  | 
1854  | 
by (simp add: eq_add_iff ring_simps)  | 
|
1855  | 
||
1856  | 
lemma bilinear_rzero:  | 
|
| 34289 | 1857  | 
fixes h :: "'a::ring^_ \<Rightarrow> _" assumes bh: "bilinear h" shows "h x 0 = 0"  | 
| 33175 | 1858  | 
using bilinear_radd[OF bh, of x 0 0 ]  | 
1859  | 
by (simp add: eq_add_iff ring_simps)  | 
|
1860  | 
||
| 34289 | 1861  | 
lemma bilinear_lsub: "bilinear h ==> h (x - (y:: 'a::ring_1 ^ _)) z = h x z - h y z"  | 
| 33175 | 1862  | 
by (simp add: diff_def bilinear_ladd bilinear_lneg)  | 
1863  | 
||
| 34289 | 1864  | 
lemma bilinear_rsub: "bilinear h ==> h z (x - (y:: 'a::ring_1 ^ _)) = h z x - h z y"  | 
| 33175 | 1865  | 
by (simp add: diff_def bilinear_radd bilinear_rneg)  | 
1866  | 
||
1867  | 
lemma bilinear_setsum:  | 
|
| 34289 | 1868  | 
fixes h:: "'a ^_ \<Rightarrow> 'a::semiring_1^_\<Rightarrow> 'a ^ _"  | 
| 33175 | 1869  | 
assumes bh: "bilinear h" and fS: "finite S" and fT: "finite T"  | 
1870  | 
shows "h (setsum f S) (setsum g T) = setsum (\<lambda>(i,j). h (f i) (g j)) (S \<times> T) "  | 
|
1871  | 
proof-  | 
|
1872  | 
have "h (setsum f S) (setsum g T) = setsum (\<lambda>x. h (f x) (setsum g T)) S"  | 
|
1873  | 
apply (rule linear_setsum[unfolded o_def])  | 
|
1874  | 
using bh fS by (auto simp add: bilinear_def)  | 
|
1875  | 
also have "\<dots> = setsum (\<lambda>x. setsum (\<lambda>y. h (f x) (g y)) T) S"  | 
|
1876  | 
apply (rule setsum_cong, simp)  | 
|
1877  | 
apply (rule linear_setsum[unfolded o_def])  | 
|
1878  | 
using bh fT by (auto simp add: bilinear_def)  | 
|
1879  | 
finally show ?thesis unfolding setsum_cartesian_product .  | 
|
1880  | 
qed  | 
|
1881  | 
||
1882  | 
lemma bilinear_bounded:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1883  | 
fixes h:: "real ^'m \<Rightarrow> real^'n \<Rightarrow> real ^'k"  | 
| 33175 | 1884  | 
assumes bh: "bilinear h"  | 
1885  | 
shows "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"  | 
|
1886  | 
proof-  | 
|
1887  | 
let ?M = "UNIV :: 'm set"  | 
|
1888  | 
let ?N = "UNIV :: 'n set"  | 
|
1889  | 
let ?B = "setsum (\<lambda>(i,j). norm (h (basis i) (basis j))) (?M \<times> ?N)"  | 
|
1890  | 
have fM: "finite ?M" and fN: "finite ?N" by simp_all  | 
|
1891  | 
  {fix x:: "real ^ 'm" and  y :: "real^'n"
 | 
|
1892  | 
have "norm (h x y) = norm (h (setsum (\<lambda>i. (x$i) *s basis i) ?M) (setsum (\<lambda>i. (y$i) *s basis i) ?N))" unfolding basis_expansion ..  | 
|
1893  | 
also have "\<dots> = norm (setsum (\<lambda> (i,j). h ((x$i) *s basis i) ((y$j) *s basis j)) (?M \<times> ?N))" unfolding bilinear_setsum[OF bh fM fN] ..  | 
|
1894  | 
finally have th: "norm (h x y) = \<dots>" .  | 
|
1895  | 
have "norm (h x y) \<le> ?B * norm x * norm y"  | 
|
1896  | 
apply (simp add: setsum_left_distrib th)  | 
|
1897  | 
apply (rule real_setsum_norm_le)  | 
|
1898  | 
using fN fM  | 
|
1899  | 
apply simp  | 
|
1900  | 
apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] norm_mul ring_simps)  | 
|
1901  | 
apply (rule mult_mono)  | 
|
1902  | 
apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm)  | 
|
1903  | 
apply (rule mult_mono)  | 
|
1904  | 
apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm)  | 
|
1905  | 
done}  | 
|
1906  | 
then show ?thesis by metis  | 
|
1907  | 
qed  | 
|
1908  | 
||
1909  | 
lemma bilinear_bounded_pos:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1910  | 
fixes h:: "real ^'m \<Rightarrow> real^'n \<Rightarrow> real ^'k"  | 
| 33175 | 1911  | 
assumes bh: "bilinear h"  | 
1912  | 
shows "\<exists>B > 0. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"  | 
|
1913  | 
proof-  | 
|
1914  | 
from bilinear_bounded[OF bh] obtain B where  | 
|
1915  | 
B: "\<forall>x y. norm (h x y) \<le> B * norm x * norm y" by blast  | 
|
1916  | 
let ?K = "\<bar>B\<bar> + 1"  | 
|
1917  | 
have Kp: "?K > 0" by arith  | 
|
1918  | 
have KB: "B < ?K" by arith  | 
|
1919  | 
  {fix x::"real ^'m" and y :: "real ^'n"
 | 
|
1920  | 
from KB Kp  | 
|
1921  | 
have "B * norm x * norm y \<le> ?K * norm x * norm y"  | 
|
1922  | 
apply -  | 
|
1923  | 
apply (rule mult_right_mono, rule mult_right_mono)  | 
|
1924  | 
by (auto simp add: norm_ge_zero)  | 
|
1925  | 
then have "norm (h x y) \<le> ?K * norm x * norm y"  | 
|
1926  | 
using B[rule_format, of x y] by simp}  | 
|
1927  | 
with Kp show ?thesis by blast  | 
|
1928  | 
qed  | 
|
1929  | 
||
1930  | 
lemma bilinear_conv_bounded_bilinear:  | 
|
1931  | 
fixes h :: "real ^ _ \<Rightarrow> real ^ _ \<Rightarrow> real ^ _"  | 
|
1932  | 
shows "bilinear h \<longleftrightarrow> bounded_bilinear h"  | 
|
1933  | 
proof  | 
|
1934  | 
assume "bilinear h"  | 
|
1935  | 
show "bounded_bilinear h"  | 
|
1936  | 
proof  | 
|
1937  | 
fix x y z show "h (x + y) z = h x z + h y z"  | 
|
1938  | 
using `bilinear h` unfolding bilinear_def linear_def by simp  | 
|
1939  | 
next  | 
|
1940  | 
fix x y z show "h x (y + z) = h x y + h x z"  | 
|
1941  | 
using `bilinear h` unfolding bilinear_def linear_def by simp  | 
|
1942  | 
next  | 
|
1943  | 
fix r x y show "h (scaleR r x) y = scaleR r (h x y)"  | 
|
1944  | 
using `bilinear h` unfolding bilinear_def linear_def  | 
|
1945  | 
by (simp add: smult_conv_scaleR)  | 
|
1946  | 
next  | 
|
1947  | 
fix r x y show "h x (scaleR r y) = scaleR r (h x y)"  | 
|
1948  | 
using `bilinear h` unfolding bilinear_def linear_def  | 
|
1949  | 
by (simp add: smult_conv_scaleR)  | 
|
1950  | 
next  | 
|
1951  | 
have "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"  | 
|
1952  | 
using `bilinear h` by (rule bilinear_bounded)  | 
|
1953  | 
thus "\<exists>K. \<forall>x y. norm (h x y) \<le> norm x * norm y * K"  | 
|
1954  | 
by (simp add: mult_ac)  | 
|
1955  | 
qed  | 
|
1956  | 
next  | 
|
1957  | 
assume "bounded_bilinear h"  | 
|
1958  | 
then interpret h: bounded_bilinear h .  | 
|
1959  | 
show "bilinear h"  | 
|
1960  | 
unfolding bilinear_def linear_conv_bounded_linear  | 
|
1961  | 
using h.bounded_linear_left h.bounded_linear_right  | 
|
1962  | 
by simp  | 
|
1963  | 
qed  | 
|
1964  | 
||
1965  | 
subsection{* Adjoints. *}
 | 
|
1966  | 
||
1967  | 
definition "adjoint f = (SOME f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y)"  | 
|
1968  | 
||
1969  | 
lemma choice_iff: "(\<forall>x. \<exists>y. P x y) \<longleftrightarrow> (\<exists>f. \<forall>x. P x (f x))" by metis  | 
|
1970  | 
||
1971  | 
lemma adjoint_works_lemma:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1972  | 
fixes f:: "'a::ring_1 ^'n \<Rightarrow> 'a ^'m"  | 
| 33175 | 1973  | 
assumes lf: "linear f"  | 
1974  | 
shows "\<forall>x y. f x \<bullet> y = x \<bullet> adjoint f y"  | 
|
1975  | 
proof-  | 
|
1976  | 
let ?N = "UNIV :: 'n set"  | 
|
1977  | 
let ?M = "UNIV :: 'm set"  | 
|
1978  | 
have fN: "finite ?N" by simp  | 
|
1979  | 
have fM: "finite ?M" by simp  | 
|
1980  | 
  {fix y:: "'a ^ 'm"
 | 
|
1981  | 
let ?w = "(\<chi> i. (f (basis i) \<bullet> y)) :: 'a ^ 'n"  | 
|
1982  | 
    {fix x
 | 
|
1983  | 
have "f x \<bullet> y = f (setsum (\<lambda>i. (x$i) *s basis i) ?N) \<bullet> y"  | 
|
1984  | 
by (simp only: basis_expansion)  | 
|
1985  | 
also have "\<dots> = (setsum (\<lambda>i. (x$i) *s f (basis i)) ?N) \<bullet> y"  | 
|
1986  | 
unfolding linear_setsum[OF lf fN]  | 
|
1987  | 
by (simp add: linear_cmul[OF lf])  | 
|
1988  | 
finally have "f x \<bullet> y = x \<bullet> ?w"  | 
|
1989  | 
apply (simp only: )  | 
|
1990  | 
apply (simp add: dot_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] ring_simps)  | 
|
1991  | 
done}  | 
|
1992  | 
}  | 
|
1993  | 
then show ?thesis unfolding adjoint_def  | 
|
1994  | 
some_eq_ex[of "\<lambda>f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y"]  | 
|
1995  | 
using choice_iff[of "\<lambda>a b. \<forall>x. f x \<bullet> a = x \<bullet> b "]  | 
|
1996  | 
by metis  | 
|
1997  | 
qed  | 
|
1998  | 
||
1999  | 
lemma adjoint_works:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2000  | 
fixes f:: "'a::ring_1 ^'n \<Rightarrow> 'a ^'m"  | 
| 33175 | 2001  | 
assumes lf: "linear f"  | 
2002  | 
shows "x \<bullet> adjoint f y = f x \<bullet> y"  | 
|
2003  | 
using adjoint_works_lemma[OF lf] by metis  | 
|
2004  | 
||
2005  | 
||
2006  | 
lemma adjoint_linear:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2007  | 
fixes f :: "'a::comm_ring_1 ^'n \<Rightarrow> 'a ^'m"  | 
| 33175 | 2008  | 
assumes lf: "linear f"  | 
2009  | 
shows "linear (adjoint f)"  | 
|
2010  | 
by (simp add: linear_def vector_eq_ldot[symmetric] dot_radd dot_rmult adjoint_works[OF lf])  | 
|
2011  | 
||
2012  | 
lemma adjoint_clauses:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2013  | 
fixes f:: "'a::comm_ring_1 ^'n \<Rightarrow> 'a ^'m"  | 
| 33175 | 2014  | 
assumes lf: "linear f"  | 
2015  | 
shows "x \<bullet> adjoint f y = f x \<bullet> y"  | 
|
2016  | 
and "adjoint f y \<bullet> x = y \<bullet> f x"  | 
|
2017  | 
by (simp_all add: adjoint_works[OF lf] dot_sym )  | 
|
2018  | 
||
2019  | 
lemma adjoint_adjoint:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2020  | 
fixes f:: "'a::comm_ring_1 ^ 'n \<Rightarrow> 'a ^'m"  | 
| 33175 | 2021  | 
assumes lf: "linear f"  | 
2022  | 
shows "adjoint (adjoint f) = f"  | 
|
2023  | 
apply (rule ext)  | 
|
2024  | 
by (simp add: vector_eq_ldot[symmetric] adjoint_clauses[OF adjoint_linear[OF lf]] adjoint_clauses[OF lf])  | 
|
2025  | 
||
2026  | 
lemma adjoint_unique:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2027  | 
fixes f:: "'a::comm_ring_1 ^ 'n \<Rightarrow> 'a ^'m"  | 
| 33175 | 2028  | 
assumes lf: "linear f" and u: "\<forall>x y. f' x \<bullet> y = x \<bullet> f y"  | 
2029  | 
shows "f' = adjoint f"  | 
|
2030  | 
apply (rule ext)  | 
|
2031  | 
using u  | 
|
2032  | 
by (simp add: vector_eq_rdot[symmetric] adjoint_clauses[OF lf])  | 
|
2033  | 
||
2034  | 
text{* Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"} *}
 | 
|
2035  | 
||
| 
34292
 
14fd037ccc47
remove overloaded star operator, use specific vector / matrix operators
 
hoelzl 
parents: 
34291 
diff
changeset
 | 
2036  | 
definition matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"  (infixl "**" 70)
 | 
| 
 
14fd037ccc47
remove overloaded star operator, use specific vector / matrix operators
 
hoelzl 
parents: 
34291 
diff
changeset
 | 
2037  | 
where "m ** m' == (\<chi> i j. setsum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m"  | 
| 
 
14fd037ccc47
remove overloaded star operator, use specific vector / matrix operators
 
hoelzl 
parents: 
34291 
diff
changeset
 | 
2038  | 
|
| 
 
14fd037ccc47
remove overloaded star operator, use specific vector / matrix operators
 
hoelzl 
parents: 
34291 
diff
changeset
 | 
2039  | 
definition matrix_vector_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"  (infixl "*v" 70)
 | 
| 
 
14fd037ccc47
remove overloaded star operator, use specific vector / matrix operators
 
hoelzl 
parents: 
34291 
diff
changeset
 | 
2040  | 
where "m *v x \<equiv> (\<chi> i. setsum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m"  | 
| 
 
14fd037ccc47
remove overloaded star operator, use specific vector / matrix operators
 
hoelzl 
parents: 
34291 
diff
changeset
 | 
2041  | 
|
| 
 
14fd037ccc47
remove overloaded star operator, use specific vector / matrix operators
 
hoelzl 
parents: 
34291 
diff
changeset
 | 
2042  | 
definition vector_matrix_mult :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "  (infixl "v*" 70)
 | 
| 
 
14fd037ccc47
remove overloaded star operator, use specific vector / matrix operators
 
hoelzl 
parents: 
34291 
diff
changeset
 | 
2043  | 
where "v v* m == (\<chi> j. setsum (\<lambda>i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n"  | 
| 33175 | 2044  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2045  | 
definition "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"  | 
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2046  | 
definition transpose where  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2047  | 
"(transpose::'a^'n^'m \<Rightarrow> 'a^'m^'n) A = (\<chi> i j. ((A$j)$i))"  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2048  | 
definition "(row::'m => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))"  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2049  | 
definition "(column::'n =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))"  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2050  | 
definition "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}"
 | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2051  | 
definition "columns(A::'a^'n^'m) = { column i A | i. i \<in> (UNIV :: 'n set)}"
 | 
| 33175 | 2052  | 
|
2053  | 
lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def)  | 
|
| 
34292
 
14fd037ccc47
remove overloaded star operator, use specific vector / matrix operators
 
hoelzl 
parents: 
34291 
diff
changeset
 | 
2054  | 
lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)"  | 
| 33175 | 2055  | 
by (vector matrix_matrix_mult_def setsum_addf[symmetric] ring_simps)  | 
2056  | 
||
2057  | 
lemma matrix_mul_lid:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2058  | 
fixes A :: "'a::semiring_1 ^ 'm ^ 'n"  | 
| 33175 | 2059  | 
shows "mat 1 ** A = A"  | 
2060  | 
apply (simp add: matrix_matrix_mult_def mat_def)  | 
|
2061  | 
apply vector  | 
|
2062  | 
by (auto simp only: cond_value_iff cond_application_beta setsum_delta'[OF finite] mult_1_left mult_zero_left if_True UNIV_I)  | 
|
2063  | 
||
2064  | 
||
2065  | 
lemma matrix_mul_rid:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2066  | 
fixes A :: "'a::semiring_1 ^ 'm ^ 'n"  | 
| 33175 | 2067  | 
shows "A ** mat 1 = A"  | 
2068  | 
apply (simp add: matrix_matrix_mult_def mat_def)  | 
|
2069  | 
apply vector  | 
|
2070  | 
by (auto simp only: cond_value_iff cond_application_beta setsum_delta[OF finite] mult_1_right mult_zero_right if_True UNIV_I cong: if_cong)  | 
|
2071  | 
||
2072  | 
lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C"  | 
|
2073  | 
apply (vector matrix_matrix_mult_def setsum_right_distrib setsum_left_distrib mult_assoc)  | 
|
2074  | 
apply (subst setsum_commute)  | 
|
2075  | 
apply simp  | 
|
2076  | 
done  | 
|
2077  | 
||
2078  | 
lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x"  | 
|
2079  | 
apply (vector matrix_matrix_mult_def matrix_vector_mult_def setsum_right_distrib setsum_left_distrib mult_assoc)  | 
|
2080  | 
apply (subst setsum_commute)  | 
|
2081  | 
apply simp  | 
|
2082  | 
done  | 
|
2083  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2084  | 
lemma matrix_vector_mul_lid: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)"  | 
| 33175 | 2085  | 
apply (vector matrix_vector_mult_def mat_def)  | 
2086  | 
by (simp add: cond_value_iff cond_application_beta  | 
|
2087  | 
setsum_delta' cong del: if_weak_cong)  | 
|
2088  | 
||
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2089  | 
lemma matrix_transpose_mul: "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)"  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2090  | 
by (simp add: matrix_matrix_mult_def transpose_def Cart_eq mult_commute)  | 
| 33175 | 2091  | 
|
2092  | 
lemma matrix_eq:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2093  | 
fixes A B :: "'a::semiring_1 ^ 'n ^ 'm"  | 
| 33175 | 2094  | 
shows "A = B \<longleftrightarrow> (\<forall>x. A *v x = B *v x)" (is "?lhs \<longleftrightarrow> ?rhs")  | 
2095  | 
apply auto  | 
|
2096  | 
apply (subst Cart_eq)  | 
|
2097  | 
apply clarify  | 
|
2098  | 
apply (clarsimp simp add: matrix_vector_mult_def basis_def cond_value_iff cond_application_beta Cart_eq cong del: if_weak_cong)  | 
|
2099  | 
apply (erule_tac x="basis ia" in allE)  | 
|
2100  | 
apply (erule_tac x="i" in allE)  | 
|
2101  | 
by (auto simp add: basis_def cond_value_iff cond_application_beta setsum_delta[OF finite] cong del: if_weak_cong)  | 
|
2102  | 
||
2103  | 
lemma matrix_vector_mul_component:  | 
|
| 34289 | 2104  | 
shows "((A::'a::semiring_1^_^_) *v x)$k = (A$k) \<bullet> x"  | 
| 33175 | 2105  | 
by (simp add: matrix_vector_mult_def dot_def)  | 
2106  | 
||
| 34289 | 2107  | 
lemma dot_lmul_matrix: "((x::'a::comm_semiring_1 ^_) v* A) \<bullet> y = x \<bullet> (A *v y)"  | 
| 33175 | 2108  | 
apply (simp add: dot_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib mult_ac)  | 
2109  | 
apply (subst setsum_commute)  | 
|
2110  | 
by simp  | 
|
2111  | 
||
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2112  | 
lemma transpose_mat: "transpose (mat n) = mat n"  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2113  | 
by (vector transpose_def mat_def)  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2114  | 
|
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2115  | 
lemma transpose_transpose: "transpose(transpose A) = A"  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2116  | 
by (vector transpose_def)  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2117  | 
|
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2118  | 
lemma row_transpose:  | 
| 34289 | 2119  | 
fixes A:: "'a::semiring_1^_^_"  | 
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2120  | 
shows "row i (transpose A) = column i A"  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2121  | 
by (simp add: row_def column_def transpose_def Cart_eq)  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2122  | 
|
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2123  | 
lemma column_transpose:  | 
| 34289 | 2124  | 
fixes A:: "'a::semiring_1^_^_"  | 
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2125  | 
shows "column i (transpose A) = row i A"  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2126  | 
by (simp add: row_def column_def transpose_def Cart_eq)  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2127  | 
|
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2128  | 
lemma rows_transpose: "rows(transpose (A::'a::semiring_1^_^_)) = columns A"  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2129  | 
by (auto simp add: rows_def columns_def row_transpose intro: set_ext)  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2130  | 
|
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2131  | 
lemma columns_transpose: "columns(transpose (A::'a::semiring_1^_^_)) = rows A" by (metis transpose_transpose rows_transpose)  | 
| 33175 | 2132  | 
|
2133  | 
text{* Two sometimes fruitful ways of looking at matrix-vector multiplication. *}
 | 
|
2134  | 
||
2135  | 
lemma matrix_mult_dot: "A *v x = (\<chi> i. A$i \<bullet> x)"  | 
|
2136  | 
by (simp add: matrix_vector_mult_def dot_def)  | 
|
2137  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2138  | 
lemma matrix_mult_vsum: "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\<lambda>i. (x$i) *s column i A) (UNIV:: 'n set)"  | 
| 33175 | 2139  | 
by (simp add: matrix_vector_mult_def Cart_eq column_def mult_commute)  | 
2140  | 
||
2141  | 
lemma vector_componentwise:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2142  | 
"(x::'a::ring_1^'n) = (\<chi> j. setsum (\<lambda>i. (x$i) * (basis i :: 'a^'n)$j) (UNIV :: 'n set))"  | 
| 33175 | 2143  | 
apply (subst basis_expansion[symmetric])  | 
2144  | 
by (vector Cart_eq setsum_component)  | 
|
2145  | 
||
2146  | 
lemma linear_componentwise:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2147  | 
fixes f:: "'a::ring_1 ^'m \<Rightarrow> 'a ^ _"  | 
| 33175 | 2148  | 
assumes lf: "linear f"  | 
2149  | 
shows "(f x)$j = setsum (\<lambda>i. (x$i) * (f (basis i)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")  | 
|
2150  | 
proof-  | 
|
2151  | 
let ?M = "(UNIV :: 'm set)"  | 
|
2152  | 
let ?N = "(UNIV :: 'n set)"  | 
|
2153  | 
have fM: "finite ?M" by simp  | 
|
2154  | 
have "?rhs = (setsum (\<lambda>i.(x$i) *s f (basis i) ) ?M)$j"  | 
|
2155  | 
unfolding vector_smult_component[symmetric]  | 
|
2156  | 
unfolding setsum_component[of "(\<lambda>i.(x$i) *s f (basis i :: 'a^'m))" ?M]  | 
|
2157  | 
..  | 
|
2158  | 
then show ?thesis unfolding linear_setsum_mul[OF lf fM, symmetric] basis_expansion ..  | 
|
2159  | 
qed  | 
|
2160  | 
||
2161  | 
text{* Inverse matrices  (not necessarily square) *}
 | 
|
2162  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2163  | 
definition "invertible(A::'a::semiring_1^'n^'m) \<longleftrightarrow> (\<exists>A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2164  | 
|
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2165  | 
definition "matrix_inv(A:: 'a::semiring_1^'n^'m) =  | 
| 33175 | 2166  | 
(SOME A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"  | 
2167  | 
||
2168  | 
text{* Correspondence between matrices and linear operators. *}
 | 
|
2169  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2170  | 
definition matrix:: "('a::{plus,times, one, zero}^'m \<Rightarrow> 'a ^ 'n) \<Rightarrow> 'a^'m^'n"
 | 
| 33175 | 2171  | 
where "matrix f = (\<chi> i j. (f(basis j))$i)"  | 
2172  | 
||
| 34289 | 2173  | 
lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::'a::comm_semiring_1 ^ _))"  | 
| 33175 | 2174  | 
by (simp add: linear_def matrix_vector_mult_def Cart_eq ring_simps setsum_right_distrib setsum_addf)  | 
2175  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2176  | 
lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::'a::comm_ring_1 ^ 'n)"  | 
| 33175 | 2177  | 
apply (simp add: matrix_def matrix_vector_mult_def Cart_eq mult_commute)  | 
2178  | 
apply clarify  | 
|
2179  | 
apply (rule linear_componentwise[OF lf, symmetric])  | 
|
2180  | 
done  | 
|
2181  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2182  | 
lemma matrix_vector_mul: "linear f ==> f = (\<lambda>x. matrix f *v (x::'a::comm_ring_1 ^ 'n))" by (simp add: ext matrix_works)  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2183  | 
|
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2184  | 
lemma matrix_of_matrix_vector_mul: "matrix(\<lambda>x. A *v (x :: 'a:: comm_ring_1 ^ 'n)) = A"  | 
| 33175 | 2185  | 
by (simp add: matrix_eq matrix_vector_mul_linear matrix_works)  | 
2186  | 
||
2187  | 
lemma matrix_compose:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2188  | 
assumes lf: "linear (f::'a::comm_ring_1^'n \<Rightarrow> 'a^'m)"  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2189  | 
and lg: "linear (g::'a::comm_ring_1^'m \<Rightarrow> 'a^_)"  | 
| 33175 | 2190  | 
shows "matrix (g o f) = matrix g ** matrix f"  | 
2191  | 
using lf lg linear_compose[OF lf lg] matrix_works[OF linear_compose[OF lf lg]]  | 
|
2192  | 
by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def)  | 
|
2193  | 
||
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2194  | 
lemma matrix_vector_column:"(A::'a::comm_semiring_1^'n^_) *v x = setsum (\<lambda>i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)"  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2195  | 
by (simp add: matrix_vector_mult_def transpose_def Cart_eq mult_commute)  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2196  | 
|
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2197  | 
lemma adjoint_matrix: "adjoint(\<lambda>x. (A::'a::comm_ring_1^'n^'m) *v x) = (\<lambda>x. transpose A *v x)"  | 
| 33175 | 2198  | 
apply (rule adjoint_unique[symmetric])  | 
2199  | 
apply (rule matrix_vector_mul_linear)  | 
|
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2200  | 
apply (simp add: transpose_def dot_def matrix_vector_mult_def setsum_left_distrib setsum_right_distrib)  | 
| 33175 | 2201  | 
apply (subst setsum_commute)  | 
2202  | 
apply (auto simp add: mult_ac)  | 
|
2203  | 
done  | 
|
2204  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2205  | 
lemma matrix_adjoint: assumes lf: "linear (f :: 'a::comm_ring_1^'n \<Rightarrow> 'a ^'m)"  | 
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
2206  | 
shows "matrix(adjoint f) = transpose(matrix f)"  | 
| 33175 | 2207  | 
apply (subst matrix_vector_mul[OF lf])  | 
2208  | 
unfolding adjoint_matrix matrix_of_matrix_vector_mul ..  | 
|
2209  | 
||
2210  | 
subsection{* Interlude: Some properties of real sets *}
 | 
|
2211  | 
||
2212  | 
lemma seq_mono_lemma: assumes "\<forall>(n::nat) \<ge> m. (d n :: real) < e n" and "\<forall>n \<ge> m. e n <= e m"  | 
|
2213  | 
shows "\<forall>n \<ge> m. d n < e m"  | 
|
2214  | 
using prems apply auto  | 
|
2215  | 
apply (erule_tac x="n" in allE)  | 
|
2216  | 
apply (erule_tac x="n" in allE)  | 
|
2217  | 
apply auto  | 
|
2218  | 
done  | 
|
2219  | 
||
2220  | 
||
2221  | 
lemma real_convex_bound_lt:  | 
|
2222  | 
assumes xa: "(x::real) < a" and ya: "y < a" and u: "0 <= u" and v: "0 <= v"  | 
|
2223  | 
and uv: "u + v = 1"  | 
|
2224  | 
shows "u * x + v * y < a"  | 
|
2225  | 
proof-  | 
|
2226  | 
have uv': "u = 0 \<longrightarrow> v \<noteq> 0" using u v uv by arith  | 
|
2227  | 
have "a = a * (u + v)" unfolding uv by simp  | 
|
2228  | 
hence th: "u * a + v * a = a" by (simp add: ring_simps)  | 
|
2229  | 
from xa u have "u \<noteq> 0 \<Longrightarrow> u*x < u*a" by (simp add: mult_compare_simps)  | 
|
2230  | 
from ya v have "v \<noteq> 0 \<Longrightarrow> v * y < v * a" by (simp add: mult_compare_simps)  | 
|
2231  | 
from xa ya u v have "u * x + v * y < u * a + v * a"  | 
|
2232  | 
apply (cases "u = 0", simp_all add: uv')  | 
|
2233  | 
apply(rule mult_strict_left_mono)  | 
|
2234  | 
using uv' apply simp_all  | 
|
2235  | 
||
2236  | 
apply (rule add_less_le_mono)  | 
|
2237  | 
apply(rule mult_strict_left_mono)  | 
|
2238  | 
apply simp_all  | 
|
2239  | 
apply (rule mult_left_mono)  | 
|
2240  | 
apply simp_all  | 
|
2241  | 
done  | 
|
2242  | 
thus ?thesis unfolding th .  | 
|
2243  | 
qed  | 
|
2244  | 
||
2245  | 
lemma real_convex_bound_le:  | 
|
2246  | 
assumes xa: "(x::real) \<le> a" and ya: "y \<le> a" and u: "0 <= u" and v: "0 <= v"  | 
|
2247  | 
and uv: "u + v = 1"  | 
|
2248  | 
shows "u * x + v * y \<le> a"  | 
|
2249  | 
proof-  | 
|
2250  | 
from xa ya u v have "u * x + v * y \<le> u * a + v * a" by (simp add: add_mono mult_left_mono)  | 
|
2251  | 
also have "\<dots> \<le> (u + v) * a" by (simp add: ring_simps)  | 
|
2252  | 
finally show ?thesis unfolding uv by simp  | 
|
2253  | 
qed  | 
|
2254  | 
||
2255  | 
lemma infinite_enumerate: assumes fS: "infinite S"  | 
|
2256  | 
shows "\<exists>r. subseq r \<and> (\<forall>n. r n \<in> S)"  | 
|
2257  | 
unfolding subseq_def  | 
|
2258  | 
using enumerate_in_set[OF fS] enumerate_mono[of _ _ S] fS by auto  | 
|
2259  | 
||
2260  | 
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)"  | 
|
2261  | 
apply auto  | 
|
2262  | 
apply (rule_tac x="d/2" in exI)  | 
|
2263  | 
apply auto  | 
|
2264  | 
done  | 
|
2265  | 
||
2266  | 
||
2267  | 
lemma triangle_lemma:  | 
|
2268  | 
assumes x: "0 <= (x::real)" and y:"0 <= y" and z: "0 <= z" and xy: "x^2 <= y^2 + z^2"  | 
|
2269  | 
shows "x <= y + z"  | 
|
2270  | 
proof-  | 
|
2271  | 
have "y^2 + z^2 \<le> y^2 + 2*y*z + z^2" using z y by (simp add: zero_compare_simps)  | 
|
2272  | 
with xy have th: "x ^2 \<le> (y+z)^2" by (simp add: power2_eq_square ring_simps)  | 
|
2273  | 
from y z have yz: "y + z \<ge> 0" by arith  | 
|
2274  | 
from power2_le_imp_le[OF th yz] show ?thesis .  | 
|
2275  | 
qed  | 
|
2276  | 
||
2277  | 
||
2278  | 
lemma lambda_skolem: "(\<forall>i. \<exists>x. P i x) \<longleftrightarrow>  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2279  | 
(\<exists>x::'a ^ 'n. \<forall>i. P i (x$i))" (is "?lhs \<longleftrightarrow> ?rhs")  | 
| 33175 | 2280  | 
proof-  | 
2281  | 
let ?S = "(UNIV :: 'n set)"  | 
|
2282  | 
  {assume H: "?rhs"
 | 
|
2283  | 
then have ?lhs by auto}  | 
|
2284  | 
moreover  | 
|
2285  | 
  {assume H: "?lhs"
 | 
|
2286  | 
then obtain f where f:"\<forall>i. P i (f i)" unfolding choice_iff by metis  | 
|
2287  | 
let ?x = "(\<chi> i. (f i)) :: 'a ^ 'n"  | 
|
2288  | 
    {fix i
 | 
|
2289  | 
from f have "P i (f i)" by metis  | 
|
2290  | 
then have "P i (?x$i)" by auto  | 
|
2291  | 
}  | 
|
2292  | 
hence "\<forall>i. P i (?x$i)" by metis  | 
|
2293  | 
hence ?rhs by metis }  | 
|
2294  | 
ultimately show ?thesis by metis  | 
|
2295  | 
qed  | 
|
2296  | 
||
2297  | 
subsection{* Operator norm. *}
 | 
|
2298  | 
||
| 33270 | 2299  | 
definition "onorm f = Sup {norm (f x)| x. norm x = 1}"
 | 
| 33175 | 2300  | 
|
2301  | 
lemma norm_bound_generalize:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2302  | 
fixes f:: "real ^'n \<Rightarrow> real^'m"  | 
| 33175 | 2303  | 
assumes lf: "linear f"  | 
2304  | 
shows "(\<forall>x. norm x = 1 \<longrightarrow> norm (f x) \<le> b) \<longleftrightarrow> (\<forall>x. norm (f x) \<le> b * norm x)" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
2305  | 
proof-  | 
|
2306  | 
  {assume H: ?rhs
 | 
|
2307  | 
    {fix x :: "real^'n" assume x: "norm x = 1"
 | 
|
2308  | 
from H[rule_format, of x] x have "norm (f x) \<le> b" by simp}  | 
|
2309  | 
then have ?lhs by blast }  | 
|
2310  | 
||
2311  | 
moreover  | 
|
2312  | 
  {assume H: ?lhs
 | 
|
2313  | 
from H[rule_format, of "basis arbitrary"]  | 
|
2314  | 
have bp: "b \<ge> 0" using norm_ge_zero[of "f (basis arbitrary)"]  | 
|
2315  | 
by (auto simp add: norm_basis elim: order_trans [OF norm_ge_zero])  | 
|
2316  | 
    {fix x :: "real ^'n"
 | 
|
2317  | 
      {assume "x = 0"
 | 
|
2318  | 
then have "norm (f x) \<le> b * norm x" by (simp add: linear_0[OF lf] bp)}  | 
|
2319  | 
moreover  | 
|
2320  | 
      {assume x0: "x \<noteq> 0"
 | 
|
2321  | 
hence n0: "norm x \<noteq> 0" by (metis norm_eq_zero)  | 
|
2322  | 
let ?c = "1/ norm x"  | 
|
2323  | 
have "norm (?c*s x) = 1" using x0 by (simp add: n0 norm_mul)  | 
|
2324  | 
with H have "norm (f(?c*s x)) \<le> b" by blast  | 
|
2325  | 
hence "?c * norm (f x) \<le> b"  | 
|
2326  | 
by (simp add: linear_cmul[OF lf] norm_mul)  | 
|
2327  | 
hence "norm (f x) \<le> b * norm x"  | 
|
2328  | 
using n0 norm_ge_zero[of x] by (auto simp add: field_simps)}  | 
|
2329  | 
ultimately have "norm (f x) \<le> b * norm x" by blast}  | 
|
2330  | 
then have ?rhs by blast}  | 
|
2331  | 
ultimately show ?thesis by blast  | 
|
2332  | 
qed  | 
|
2333  | 
||
2334  | 
lemma onorm:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2335  | 
fixes f:: "real ^'n \<Rightarrow> real ^'m"  | 
| 33175 | 2336  | 
assumes lf: "linear f"  | 
2337  | 
shows "norm (f x) <= onorm f * norm x"  | 
|
2338  | 
and "\<forall>x. norm (f x) <= b * norm x \<Longrightarrow> onorm f <= b"  | 
|
2339  | 
proof-  | 
|
2340  | 
  {
 | 
|
2341  | 
    let ?S = "{norm (f x) |x. norm x = 1}"
 | 
|
2342  | 
    have Se: "?S \<noteq> {}" using  norm_basis by auto
 | 
|
2343  | 
from linear_bounded[OF lf] have b: "\<exists> b. ?S *<= b"  | 
|
2344  | 
unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def)  | 
|
| 33270 | 2345  | 
    {from Sup[OF Se b, unfolded onorm_def[symmetric]]
 | 
| 33175 | 2346  | 
show "norm (f x) <= onorm f * norm x"  | 
2347  | 
apply -  | 
|
2348  | 
apply (rule spec[where x = x])  | 
|
2349  | 
unfolding norm_bound_generalize[OF lf, symmetric]  | 
|
2350  | 
by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)}  | 
|
2351  | 
    {
 | 
|
2352  | 
show "\<forall>x. norm (f x) <= b * norm x \<Longrightarrow> onorm f <= b"  | 
|
| 33270 | 2353  | 
using Sup[OF Se b, unfolded onorm_def[symmetric]]  | 
| 33175 | 2354  | 
unfolding norm_bound_generalize[OF lf, symmetric]  | 
2355  | 
by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)}  | 
|
2356  | 
}  | 
|
2357  | 
qed  | 
|
2358  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2359  | 
lemma onorm_pos_le: assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" shows "0 <= onorm f"  | 
| 33175 | 2360  | 
using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "basis arbitrary"], unfolded norm_basis] by simp  | 
2361  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2362  | 
lemma onorm_eq_0: assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)"  | 
| 33175 | 2363  | 
shows "onorm f = 0 \<longleftrightarrow> (\<forall>x. f x = 0)"  | 
2364  | 
using onorm[OF lf]  | 
|
2365  | 
apply (auto simp add: onorm_pos_le)  | 
|
2366  | 
apply atomize  | 
|
2367  | 
apply (erule allE[where x="0::real"])  | 
|
2368  | 
using onorm_pos_le[OF lf]  | 
|
2369  | 
apply arith  | 
|
2370  | 
done  | 
|
2371  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2372  | 
lemma onorm_const: "onorm(\<lambda>x::real^'n. (y::real ^'m)) = norm y"  | 
| 33175 | 2373  | 
proof-  | 
2374  | 
let ?f = "\<lambda>x::real^'n. (y::real ^ 'm)"  | 
|
2375  | 
  have th: "{norm (?f x)| x. norm x = 1} = {norm y}"
 | 
|
2376  | 
by(auto intro: vector_choose_size set_ext)  | 
|
2377  | 
show ?thesis  | 
|
2378  | 
unfolding onorm_def th  | 
|
| 33270 | 2379  | 
apply (rule Sup_unique) by (simp_all add: setle_def)  | 
| 33175 | 2380  | 
qed  | 
2381  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2382  | 
lemma onorm_pos_lt: assumes lf: "linear (f::real ^ 'n \<Rightarrow> real ^'m)"  | 
| 33175 | 2383  | 
shows "0 < onorm f \<longleftrightarrow> ~(\<forall>x. f x = 0)"  | 
2384  | 
unfolding onorm_eq_0[OF lf, symmetric]  | 
|
2385  | 
using onorm_pos_le[OF lf] by arith  | 
|
2386  | 
||
2387  | 
lemma onorm_compose:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2388  | 
assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)"  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2389  | 
and lg: "linear (g::real^'k \<Rightarrow> real^'n)"  | 
| 33175 | 2390  | 
shows "onorm (f o g) <= onorm f * onorm g"  | 
2391  | 
apply (rule onorm(2)[OF linear_compose[OF lg lf], rule_format])  | 
|
2392  | 
unfolding o_def  | 
|
2393  | 
apply (subst mult_assoc)  | 
|
2394  | 
apply (rule order_trans)  | 
|
2395  | 
apply (rule onorm(1)[OF lf])  | 
|
2396  | 
apply (rule mult_mono1)  | 
|
2397  | 
apply (rule onorm(1)[OF lg])  | 
|
2398  | 
apply (rule onorm_pos_le[OF lf])  | 
|
2399  | 
done  | 
|
2400  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2401  | 
lemma onorm_neg_lemma: assumes lf: "linear (f::real ^'n \<Rightarrow> real^'m)"  | 
| 33175 | 2402  | 
shows "onorm (\<lambda>x. - f x) \<le> onorm f"  | 
2403  | 
using onorm[OF linear_compose_neg[OF lf]] onorm[OF lf]  | 
|
2404  | 
unfolding norm_minus_cancel by metis  | 
|
2405  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2406  | 
lemma onorm_neg: assumes lf: "linear (f::real ^'n \<Rightarrow> real^'m)"  | 
| 33175 | 2407  | 
shows "onorm (\<lambda>x. - f x) = onorm f"  | 
2408  | 
using onorm_neg_lemma[OF lf] onorm_neg_lemma[OF linear_compose_neg[OF lf]]  | 
|
2409  | 
by simp  | 
|
2410  | 
||
2411  | 
lemma onorm_triangle:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2412  | 
assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" and lg: "linear g"  | 
| 33175 | 2413  | 
shows "onorm (\<lambda>x. f x + g x) <= onorm f + onorm g"  | 
2414  | 
apply(rule onorm(2)[OF linear_compose_add[OF lf lg], rule_format])  | 
|
2415  | 
apply (rule order_trans)  | 
|
2416  | 
apply (rule norm_triangle_ineq)  | 
|
2417  | 
apply (simp add: distrib)  | 
|
2418  | 
apply (rule add_mono)  | 
|
2419  | 
apply (rule onorm(1)[OF lf])  | 
|
2420  | 
apply (rule onorm(1)[OF lg])  | 
|
2421  | 
done  | 
|
2422  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2423  | 
lemma onorm_triangle_le: "linear (f::real ^'n \<Rightarrow> real ^'m) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) <= e  | 
| 33175 | 2424  | 
\<Longrightarrow> onorm(\<lambda>x. f x + g x) <= e"  | 
2425  | 
apply (rule order_trans)  | 
|
2426  | 
apply (rule onorm_triangle)  | 
|
2427  | 
apply assumption+  | 
|
2428  | 
done  | 
|
2429  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2430  | 
lemma onorm_triangle_lt: "linear (f::real ^'n \<Rightarrow> real ^'m) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) < e  | 
| 33175 | 2431  | 
==> onorm(\<lambda>x. f x + g x) < e"  | 
2432  | 
apply (rule order_le_less_trans)  | 
|
2433  | 
apply (rule onorm_triangle)  | 
|
2434  | 
by assumption+  | 
|
2435  | 
||
2436  | 
(* "lift" from 'a to 'a^1 and "drop" from 'a^1 to 'a -- FIXME: potential use of transfer *)  | 
|
2437  | 
||
| 34964 | 2438  | 
abbreviation vec1:: "'a \<Rightarrow> 'a ^ 1" where "vec1 x \<equiv> vec x"  | 
2439  | 
||
2440  | 
abbreviation dest_vec1:: "'a ^1 \<Rightarrow> 'a"  | 
|
2441  | 
where "dest_vec1 x \<equiv> (x$1)"  | 
|
| 33175 | 2442  | 
|
2443  | 
lemma vec1_component[simp]: "(vec1 x)$1 = x"  | 
|
| 34964 | 2444  | 
by (simp add: )  | 
| 33175 | 2445  | 
|
2446  | 
lemma vec1_dest_vec1[simp]: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y"  | 
|
| 34964 | 2447  | 
by (simp_all add: Cart_eq forall_1)  | 
| 33175 | 2448  | 
|
2449  | 
lemma forall_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (vec1 x))" by (metis vec1_dest_vec1)  | 
|
2450  | 
||
2451  | 
lemma exists_vec1: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P(vec1 x))" by (metis vec1_dest_vec1)  | 
|
2452  | 
||
2453  | 
lemma vec1_eq[simp]: "vec1 x = vec1 y \<longleftrightarrow> x = y" by (metis vec1_dest_vec1)  | 
|
2454  | 
||
2455  | 
lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \<longleftrightarrow> x = y" by (metis vec1_dest_vec1)  | 
|
2456  | 
||
| 34964 | 2457  | 
lemma vec_in_image_vec: "vec x \<in> (vec ` S) \<longleftrightarrow> x \<in> S" by auto  | 
2458  | 
||
2459  | 
lemma vec_add: "vec(x + y) = vec x + vec y" by (vector vec_def)  | 
|
2460  | 
lemma vec_sub: "vec(x - y) = vec x - vec y" by (vector vec_def)  | 
|
2461  | 
lemma vec_cmul: "vec(c* x) = c *s vec x " by (vector vec_def)  | 
|
2462  | 
lemma vec_neg: "vec(- x) = - vec x " by (vector vec_def)  | 
|
| 
33714
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
2463  | 
|
| 
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
2464  | 
lemma range_vec1[simp]:"range vec1 = UNIV" apply(rule set_ext,rule) unfolding image_iff defer  | 
| 
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
2465  | 
apply(rule_tac x="dest_vec1 x" in bexI) by auto  | 
| 
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
2466  | 
|
| 34964 | 2467  | 
lemma vec_setsum: assumes fS: "finite S"  | 
2468  | 
shows "vec(setsum f S) = setsum (vec o f) S"  | 
|
| 33175 | 2469  | 
apply (induct rule: finite_induct[OF fS])  | 
| 34964 | 2470  | 
apply (simp)  | 
2471  | 
apply (auto simp add: vec_add)  | 
|
| 33175 | 2472  | 
done  | 
2473  | 
||
2474  | 
lemma dest_vec1_lambda: "dest_vec1(\<chi> i. x i) = x 1"  | 
|
| 34964 | 2475  | 
by (simp)  | 
| 33175 | 2476  | 
|
2477  | 
lemma dest_vec1_vec: "dest_vec1(vec x) = x"  | 
|
| 34964 | 2478  | 
by (simp)  | 
| 33175 | 2479  | 
|
2480  | 
lemma dest_vec1_sum: assumes fS: "finite S"  | 
|
2481  | 
shows "dest_vec1(setsum f S) = setsum (dest_vec1 o f) S"  | 
|
2482  | 
apply (induct rule: finite_induct[OF fS])  | 
|
2483  | 
apply (simp add: dest_vec1_vec)  | 
|
| 34964 | 2484  | 
apply (auto simp add:vector_minus_component)  | 
| 33175 | 2485  | 
done  | 
2486  | 
||
2487  | 
lemma norm_vec1: "norm(vec1 x) = abs(x)"  | 
|
| 34964 | 2488  | 
by (simp add: vec_def norm_real)  | 
| 33175 | 2489  | 
|
2490  | 
lemma dist_vec1: "dist(vec1 x) (vec1 y) = abs(x - y)"  | 
|
2491  | 
by (simp only: dist_real vec1_component)  | 
|
2492  | 
lemma abs_dest_vec1: "norm x = \<bar>dest_vec1 x\<bar>"  | 
|
2493  | 
by (metis vec1_dest_vec1 norm_vec1)  | 
|
2494  | 
||
| 34964 | 2495  | 
lemmas vec1_dest_vec1_simps = forall_vec1 vec_add[THEN sym] dist_vec1 vec_sub[THEN sym] vec1_dest_vec1 norm_vec1 vector_smult_component  | 
2496  | 
vec1_eq vec_cmul[THEN sym] smult_conv_scaleR[THEN sym] o_def dist_real_def norm_vec1 real_norm_def  | 
|
| 
33714
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
2497  | 
|
| 
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
2498  | 
lemma bounded_linear_vec1:"bounded_linear (vec1::real\<Rightarrow>real^1)"  | 
| 
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
2499  | 
unfolding bounded_linear_def additive_def bounded_linear_axioms_def  | 
| 
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
2500  | 
unfolding smult_conv_scaleR[THEN sym] unfolding vec1_dest_vec1_simps  | 
| 
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
2501  | 
apply(rule conjI) defer apply(rule conjI) defer apply(rule_tac x=1 in exI) by auto  | 
| 
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33270 
diff
changeset
 | 
2502  | 
|
| 33175 | 2503  | 
lemma linear_vmul_dest_vec1:  | 
| 34289 | 2504  | 
fixes f:: "'a::semiring_1^_ \<Rightarrow> 'a^1"  | 
| 33175 | 2505  | 
shows "linear f \<Longrightarrow> linear (\<lambda>x. dest_vec1(f x) *s v)"  | 
2506  | 
apply (rule linear_vmul_component)  | 
|
2507  | 
by auto  | 
|
2508  | 
||
2509  | 
lemma linear_from_scalars:  | 
|
| 34289 | 2510  | 
assumes lf: "linear (f::'a::comm_ring_1 ^1 \<Rightarrow> 'a^_)"  | 
| 33175 | 2511  | 
shows "f = (\<lambda>x. dest_vec1 x *s column 1 (matrix f))"  | 
2512  | 
apply (rule ext)  | 
|
2513  | 
apply (subst matrix_works[OF lf, symmetric])  | 
|
| 34964 | 2514  | 
apply (auto simp add: Cart_eq matrix_vector_mult_def column_def mult_commute UNIV_1)  | 
| 33175 | 2515  | 
done  | 
2516  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2517  | 
lemma linear_to_scalars: assumes lf: "linear (f::'a::comm_ring_1 ^'n \<Rightarrow> 'a^1)"  | 
| 33175 | 2518  | 
shows "f = (\<lambda>x. vec1(row 1 (matrix f) \<bullet> x))"  | 
2519  | 
apply (rule ext)  | 
|
2520  | 
apply (subst matrix_works[OF lf, symmetric])  | 
|
| 34964 | 2521  | 
apply (simp add: Cart_eq matrix_vector_mult_def row_def dot_def mult_commute forall_1)  | 
| 33175 | 2522  | 
done  | 
2523  | 
||
2524  | 
lemma dest_vec1_eq_0: "dest_vec1 x = 0 \<longleftrightarrow> x = 0"  | 
|
2525  | 
by (simp add: dest_vec1_eq[symmetric])  | 
|
2526  | 
||
2527  | 
lemma setsum_scalars: assumes fS: "finite S"  | 
|
2528  | 
shows "setsum f S = vec1 (setsum (dest_vec1 o f) S)"  | 
|
| 34964 | 2529  | 
unfolding vec_setsum[OF fS] by simp  | 
| 33175 | 2530  | 
|
2531  | 
lemma dest_vec1_wlog_le: "(\<And>(x::'a::linorder ^ 1) y. P x y \<longleftrightarrow> P y x) \<Longrightarrow> (\<And>x y. dest_vec1 x <= dest_vec1 y ==> P x y) \<Longrightarrow> P x y"  | 
|
2532  | 
apply (cases "dest_vec1 x \<le> dest_vec1 y")  | 
|
2533  | 
apply simp  | 
|
2534  | 
apply (subgoal_tac "dest_vec1 y \<le> dest_vec1 x")  | 
|
2535  | 
apply (auto)  | 
|
2536  | 
done  | 
|
2537  | 
||
2538  | 
text{* Pasting vectors. *}
 | 
|
2539  | 
||
| 34964 | 2540  | 
lemma linear_fstcart[intro]: "linear fstcart"  | 
| 33175 | 2541  | 
by (auto simp add: linear_def Cart_eq)  | 
2542  | 
||
| 34964 | 2543  | 
lemma linear_sndcart[intro]: "linear sndcart"  | 
| 33175 | 2544  | 
by (auto simp add: linear_def Cart_eq)  | 
2545  | 
||
2546  | 
lemma fstcart_vec[simp]: "fstcart(vec x) = vec x"  | 
|
2547  | 
by (simp add: Cart_eq)  | 
|
2548  | 
||
| 34289 | 2549  | 
lemma fstcart_add[simp]:"fstcart(x + y) = fstcart (x::'a::{plus,times}^('b::finite + 'c::finite)) + fstcart y"
 | 
| 33175 | 2550  | 
by (simp add: Cart_eq)  | 
2551  | 
||
| 34289 | 2552  | 
lemma fstcart_cmul[simp]:"fstcart(c*s x) = c*s fstcart (x::'a::{plus,times}^('b::finite + 'c::finite))"
 | 
| 33175 | 2553  | 
by (simp add: Cart_eq)  | 
2554  | 
||
| 34289 | 2555  | 
lemma fstcart_neg[simp]:"fstcart(- x) = - fstcart (x::'a::ring_1^(_ + _))"  | 
| 33175 | 2556  | 
by (simp add: Cart_eq)  | 
2557  | 
||
| 34289 | 2558  | 
lemma fstcart_sub[simp]:"fstcart(x - y) = fstcart (x::'a::ring_1^(_ + _)) - fstcart y"  | 
| 33175 | 2559  | 
by (simp add: Cart_eq)  | 
2560  | 
||
2561  | 
lemma fstcart_setsum:  | 
|
2562  | 
fixes f:: "'d \<Rightarrow> 'a::semiring_1^_"  | 
|
2563  | 
assumes fS: "finite S"  | 
|
2564  | 
shows "fstcart (setsum f S) = setsum (\<lambda>i. fstcart (f i)) S"  | 
|
2565  | 
by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0)  | 
|
2566  | 
||
2567  | 
lemma sndcart_vec[simp]: "sndcart(vec x) = vec x"  | 
|
2568  | 
by (simp add: Cart_eq)  | 
|
2569  | 
||
| 34289 | 2570  | 
lemma sndcart_add[simp]:"sndcart(x + y) = sndcart (x::'a::{plus,times}^(_ + _)) + sndcart y"
 | 
| 33175 | 2571  | 
by (simp add: Cart_eq)  | 
2572  | 
||
| 34289 | 2573  | 
lemma sndcart_cmul[simp]:"sndcart(c*s x) = c*s sndcart (x::'a::{plus,times}^(_ + _))"
 | 
| 33175 | 2574  | 
by (simp add: Cart_eq)  | 
2575  | 
||
| 34289 | 2576  | 
lemma sndcart_neg[simp]:"sndcart(- x) = - sndcart (x::'a::ring_1^(_ + _))"  | 
| 33175 | 2577  | 
by (simp add: Cart_eq)  | 
2578  | 
||
| 34289 | 2579  | 
lemma sndcart_sub[simp]:"sndcart(x - y) = sndcart (x::'a::ring_1^(_ + _)) - sndcart y"  | 
| 33175 | 2580  | 
by (simp add: Cart_eq)  | 
2581  | 
||
2582  | 
lemma sndcart_setsum:  | 
|
2583  | 
fixes f:: "'d \<Rightarrow> 'a::semiring_1^_"  | 
|
2584  | 
assumes fS: "finite S"  | 
|
2585  | 
shows "sndcart (setsum f S) = setsum (\<lambda>i. sndcart (f i)) S"  | 
|
2586  | 
by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0)  | 
|
2587  | 
||
2588  | 
lemma pastecart_vec[simp]: "pastecart (vec x) (vec x) = vec x"  | 
|
2589  | 
by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)  | 
|
2590  | 
||
2591  | 
lemma pastecart_add[simp]:"pastecart (x1::'a::{plus,times}^_) y1 + pastecart x2 y2 = pastecart (x1 + x2) (y1 + y2)"
 | 
|
2592  | 
by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)  | 
|
2593  | 
||
2594  | 
lemma pastecart_cmul[simp]: "pastecart (c *s (x1::'a::{plus,times}^_)) (c *s y1) = c *s pastecart x1 y1"
 | 
|
2595  | 
by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)  | 
|
2596  | 
||
2597  | 
lemma pastecart_neg[simp]: "pastecart (- (x::'a::ring_1^_)) (- y) = - pastecart x y"  | 
|
2598  | 
unfolding vector_sneg_minus1 pastecart_cmul ..  | 
|
2599  | 
||
2600  | 
lemma pastecart_sub: "pastecart (x1::'a::ring_1^_) y1 - pastecart x2 y2 = pastecart (x1 - x2) (y1 - y2)"  | 
|
2601  | 
by (simp add: diff_def pastecart_neg[symmetric] del: pastecart_neg)  | 
|
2602  | 
||
2603  | 
lemma pastecart_setsum:  | 
|
2604  | 
fixes f:: "'d \<Rightarrow> 'a::semiring_1^_"  | 
|
2605  | 
assumes fS: "finite S"  | 
|
2606  | 
shows "pastecart (setsum f S) (setsum g S) = setsum (\<lambda>i. pastecart (f i) (g i)) S"  | 
|
2607  | 
by (simp add: pastecart_eq fstcart_setsum[OF fS] sndcart_setsum[OF fS] fstcart_pastecart sndcart_pastecart)  | 
|
2608  | 
||
2609  | 
lemma setsum_Plus:  | 
|
2610  | 
"\<lbrakk>finite A; finite B\<rbrakk> \<Longrightarrow>  | 
|
2611  | 
(\<Sum>x\<in>A <+> B. g x) = (\<Sum>x\<in>A. g (Inl x)) + (\<Sum>x\<in>B. g (Inr x))"  | 
|
2612  | 
unfolding Plus_def  | 
|
2613  | 
by (subst setsum_Un_disjoint, auto simp add: setsum_reindex)  | 
|
2614  | 
||
2615  | 
lemma setsum_UNIV_sum:  | 
|
2616  | 
fixes g :: "'a::finite + 'b::finite \<Rightarrow> _"  | 
|
2617  | 
shows "(\<Sum>x\<in>UNIV. g x) = (\<Sum>x\<in>UNIV. g (Inl x)) + (\<Sum>x\<in>UNIV. g (Inr x))"  | 
|
2618  | 
apply (subst UNIV_Plus_UNIV [symmetric])  | 
|
2619  | 
apply (rule setsum_Plus [OF finite finite])  | 
|
2620  | 
done  | 
|
2621  | 
||
2622  | 
lemma norm_fstcart: "norm(fstcart x) <= norm (x::real ^('n::finite + 'm::finite))"
 | 
|
2623  | 
proof-  | 
|
2624  | 
have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))"  | 
|
2625  | 
by (simp add: pastecart_fst_snd)  | 
|
2626  | 
have th1: "fstcart x \<bullet> fstcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)"  | 
|
2627  | 
by (simp add: dot_def setsum_UNIV_sum pastecart_def setsum_nonneg)  | 
|
2628  | 
then show ?thesis  | 
|
2629  | 
unfolding th0  | 
|
2630  | 
unfolding real_vector_norm_def real_sqrt_le_iff id_def  | 
|
2631  | 
by (simp add: dot_def)  | 
|
2632  | 
qed  | 
|
2633  | 
||
2634  | 
lemma dist_fstcart: "dist(fstcart (x::real^_)) (fstcart y) <= dist x y"  | 
|
2635  | 
unfolding dist_norm by (metis fstcart_sub[symmetric] norm_fstcart)  | 
|
2636  | 
||
2637  | 
lemma norm_sndcart: "norm(sndcart x) <= norm (x::real ^('n::finite + 'm::finite))"
 | 
|
2638  | 
proof-  | 
|
2639  | 
have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))"  | 
|
2640  | 
by (simp add: pastecart_fst_snd)  | 
|
2641  | 
have th1: "sndcart x \<bullet> sndcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)"  | 
|
2642  | 
by (simp add: dot_def setsum_UNIV_sum pastecart_def setsum_nonneg)  | 
|
2643  | 
then show ?thesis  | 
|
2644  | 
unfolding th0  | 
|
2645  | 
unfolding real_vector_norm_def real_sqrt_le_iff id_def  | 
|
2646  | 
by (simp add: dot_def)  | 
|
2647  | 
qed  | 
|
2648  | 
||
2649  | 
lemma dist_sndcart: "dist(sndcart (x::real^_)) (sndcart y) <= dist x y"  | 
|
2650  | 
unfolding dist_norm by (metis sndcart_sub[symmetric] norm_sndcart)  | 
|
2651  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2652  | 
lemma dot_pastecart: "(pastecart (x1::'a::{times,comm_monoid_add}^'n) (x2::'a::{times,comm_monoid_add}^'m)) \<bullet> (pastecart y1 y2) =  x1 \<bullet> y1 + x2 \<bullet> y2"
 | 
| 33175 | 2653  | 
by (simp add: dot_def setsum_UNIV_sum pastecart_def)  | 
2654  | 
||
2655  | 
text {* TODO: move to NthRoot *}
 | 
|
2656  | 
lemma sqrt_add_le_add_sqrt:  | 
|
2657  | 
assumes x: "0 \<le> x" and y: "0 \<le> y"  | 
|
2658  | 
shows "sqrt (x + y) \<le> sqrt x + sqrt y"  | 
|
2659  | 
apply (rule power2_le_imp_le)  | 
|
2660  | 
apply (simp add: real_sum_squared_expand add_nonneg_nonneg x y)  | 
|
2661  | 
apply (simp add: mult_nonneg_nonneg x y)  | 
|
2662  | 
apply (simp add: add_nonneg_nonneg x y)  | 
|
2663  | 
done  | 
|
2664  | 
||
2665  | 
lemma norm_pastecart: "norm (pastecart x y) <= norm x + norm y"  | 
|
2666  | 
unfolding norm_vector_def setL2_def setsum_UNIV_sum  | 
|
2667  | 
by (simp add: sqrt_add_le_add_sqrt setsum_nonneg)  | 
|
2668  | 
||
2669  | 
subsection {* A generic notion of "hull" (convex, affine, conic hull and closure). *}
 | 
|
2670  | 
||
2671  | 
definition hull :: "'a set set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "hull" 75) where  | 
|
2672  | 
  "S hull s = Inter {t. t \<in> S \<and> s \<subseteq> t}"
 | 
|
2673  | 
||
2674  | 
lemma hull_same: "s \<in> S \<Longrightarrow> S hull s = s"  | 
|
2675  | 
unfolding hull_def by auto  | 
|
2676  | 
||
2677  | 
lemma hull_in: "(\<And>T. T \<subseteq> S ==> Inter T \<in> S) ==> (S hull s) \<in> S"  | 
|
2678  | 
unfolding hull_def subset_iff by auto  | 
|
2679  | 
||
2680  | 
lemma hull_eq: "(\<And>T. T \<subseteq> S ==> Inter T \<in> S) ==> (S hull s) = s \<longleftrightarrow> s \<in> S"  | 
|
2681  | 
using hull_same[of s S] hull_in[of S s] by metis  | 
|
2682  | 
||
2683  | 
||
2684  | 
lemma hull_hull: "S hull (S hull s) = S hull s"  | 
|
2685  | 
unfolding hull_def by blast  | 
|
2686  | 
||
| 34964 | 2687  | 
lemma hull_subset[intro]: "s \<subseteq> (S hull s)"  | 
| 33175 | 2688  | 
unfolding hull_def by blast  | 
2689  | 
||
2690  | 
lemma hull_mono: " s \<subseteq> t ==> (S hull s) \<subseteq> (S hull t)"  | 
|
2691  | 
unfolding hull_def by blast  | 
|
2692  | 
||
2693  | 
lemma hull_antimono: "S \<subseteq> T ==> (T hull s) \<subseteq> (S hull s)"  | 
|
2694  | 
unfolding hull_def by blast  | 
|
2695  | 
||
2696  | 
lemma hull_minimal: "s \<subseteq> t \<Longrightarrow> t \<in> S ==> (S hull s) \<subseteq> t"  | 
|
2697  | 
unfolding hull_def by blast  | 
|
2698  | 
||
2699  | 
lemma subset_hull: "t \<in> S ==> S hull s \<subseteq> t \<longleftrightarrow> s \<subseteq> t"  | 
|
2700  | 
unfolding hull_def by blast  | 
|
2701  | 
||
2702  | 
lemma hull_unique: "s \<subseteq> t \<Longrightarrow> t \<in> S \<Longrightarrow> (\<And>t'. s \<subseteq> t' \<Longrightarrow> t' \<in> S ==> t \<subseteq> t')  | 
|
2703  | 
==> (S hull s = t)"  | 
|
2704  | 
unfolding hull_def by auto  | 
|
2705  | 
||
2706  | 
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"
 | 
|
2707  | 
  using hull_minimal[of S "{x. P x}" Q]
 | 
|
2708  | 
by (auto simp add: subset_eq Collect_def mem_def)  | 
|
2709  | 
||
2710  | 
lemma hull_inc: "x \<in> S \<Longrightarrow> x \<in> P hull S" by (metis hull_subset subset_eq)  | 
|
2711  | 
||
2712  | 
lemma hull_union_subset: "(S hull s) \<union> (S hull t) \<subseteq> (S hull (s \<union> t))"  | 
|
2713  | 
unfolding Un_subset_iff by (metis hull_mono Un_upper1 Un_upper2)  | 
|
2714  | 
||
2715  | 
lemma hull_union: assumes T: "\<And>T. T \<subseteq> S ==> Inter T \<in> S"  | 
|
2716  | 
shows "S hull (s \<union> t) = S hull (S hull s \<union> S hull t)"  | 
|
2717  | 
apply rule  | 
|
2718  | 
apply (rule hull_mono)  | 
|
2719  | 
unfolding Un_subset_iff  | 
|
2720  | 
apply (metis hull_subset Un_upper1 Un_upper2 subset_trans)  | 
|
2721  | 
apply (rule hull_minimal)  | 
|
2722  | 
apply (metis hull_union_subset)  | 
|
2723  | 
apply (metis hull_in T)  | 
|
2724  | 
done  | 
|
2725  | 
||
2726  | 
lemma hull_redundant_eq: "a \<in> (S hull s) \<longleftrightarrow> (S hull (insert a s) = S hull s)"  | 
|
2727  | 
unfolding hull_def by blast  | 
|
2728  | 
||
2729  | 
lemma hull_redundant: "a \<in> (S hull s) ==> (S hull (insert a s) = S hull s)"  | 
|
2730  | 
by (metis hull_redundant_eq)  | 
|
2731  | 
||
2732  | 
text{* Archimedian properties and useful consequences. *}
 | 
|
2733  | 
||
2734  | 
lemma real_arch_simple: "\<exists>n. x <= real (n::nat)"  | 
|
2735  | 
using reals_Archimedean2[of x] apply auto by (rule_tac x="Suc n" in exI, auto)  | 
|
2736  | 
lemmas real_arch_lt = reals_Archimedean2  | 
|
2737  | 
||
2738  | 
lemmas real_arch = reals_Archimedean3  | 
|
2739  | 
||
2740  | 
lemma real_arch_inv: "0 < e \<longleftrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> 0 < inverse (real n) \<and> inverse (real n) < e)"  | 
|
2741  | 
using reals_Archimedean  | 
|
2742  | 
apply (auto simp add: field_simps inverse_positive_iff_positive)  | 
|
2743  | 
apply (subgoal_tac "inverse (real n) > 0")  | 
|
2744  | 
apply arith  | 
|
2745  | 
apply simp  | 
|
2746  | 
done  | 
|
2747  | 
||
2748  | 
lemma real_pow_lbound: "0 <= x ==> 1 + real n * x <= (1 + x) ^ n"  | 
|
2749  | 
proof(induct n)  | 
|
2750  | 
case 0 thus ?case by simp  | 
|
2751  | 
next  | 
|
2752  | 
case (Suc n)  | 
|
2753  | 
hence h: "1 + real n * x \<le> (1 + x) ^ n" by simp  | 
|
2754  | 
from h have p: "1 \<le> (1 + x) ^ n" using Suc.prems by simp  | 
|
2755  | 
from h have "1 + real n * x + x \<le> (1 + x) ^ n + x" by simp  | 
|
2756  | 
also have "\<dots> \<le> (1 + x) ^ Suc n" apply (subst diff_le_0_iff_le[symmetric])  | 
|
2757  | 
apply (simp add: ring_simps)  | 
|
2758  | 
using mult_left_mono[OF p Suc.prems] by simp  | 
|
2759  | 
finally show ?case by (simp add: real_of_nat_Suc ring_simps)  | 
|
2760  | 
qed  | 
|
2761  | 
||
2762  | 
lemma real_arch_pow: assumes x: "1 < (x::real)" shows "\<exists>n. y < x^n"  | 
|
2763  | 
proof-  | 
|
2764  | 
from x have x0: "x - 1 > 0" by arith  | 
|
2765  | 
from real_arch[OF x0, rule_format, of y]  | 
|
2766  | 
obtain n::nat where n:"y < real n * (x - 1)" by metis  | 
|
2767  | 
from x0 have x00: "x- 1 \<ge> 0" by arith  | 
|
2768  | 
from real_pow_lbound[OF x00, of n] n  | 
|
2769  | 
have "y < x^n" by auto  | 
|
2770  | 
then show ?thesis by metis  | 
|
2771  | 
qed  | 
|
2772  | 
||
2773  | 
lemma real_arch_pow2: "\<exists>n. (x::real) < 2^ n"  | 
|
2774  | 
using real_arch_pow[of 2 x] by simp  | 
|
2775  | 
||
2776  | 
lemma real_arch_pow_inv: assumes y: "(y::real) > 0" and x1: "x < 1"  | 
|
2777  | 
shows "\<exists>n. x^n < y"  | 
|
2778  | 
proof-  | 
|
2779  | 
  {assume x0: "x > 0"
 | 
|
2780  | 
from x0 x1 have ix: "1 < 1/x" by (simp add: field_simps)  | 
|
2781  | 
from real_arch_pow[OF ix, of "1/y"]  | 
|
2782  | 
obtain n where n: "1/y < (1/x)^n" by blast  | 
|
2783  | 
then  | 
|
2784  | 
have ?thesis using y x0 by (auto simp add: field_simps power_divide) }  | 
|
2785  | 
moreover  | 
|
2786  | 
  {assume "\<not> x > 0" with y x1 have ?thesis apply auto by (rule exI[where x=1], auto)}
 | 
|
2787  | 
ultimately show ?thesis by metis  | 
|
2788  | 
qed  | 
|
2789  | 
||
2790  | 
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)"  | 
|
2791  | 
by (metis real_arch_inv)  | 
|
2792  | 
||
2793  | 
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"  | 
|
2794  | 
apply (rule forall_pos_mono)  | 
|
2795  | 
apply auto  | 
|
2796  | 
apply (atomize)  | 
|
2797  | 
apply (erule_tac x="n - 1" in allE)  | 
|
2798  | 
apply auto  | 
|
2799  | 
done  | 
|
2800  | 
||
2801  | 
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"  | 
|
2802  | 
shows "x = 0"  | 
|
2803  | 
proof-  | 
|
2804  | 
  {assume "x \<noteq> 0" with x0 have xp: "x > 0" by arith
 | 
|
2805  | 
from real_arch[OF xp, rule_format, of c] obtain n::nat where n: "c < real n * x" by blast  | 
|
2806  | 
with xc[rule_format, of n] have "n = 0" by arith  | 
|
2807  | 
with n c have False by simp}  | 
|
2808  | 
then show ?thesis by blast  | 
|
2809  | 
qed  | 
|
2810  | 
||
2811  | 
(* ------------------------------------------------------------------------- *)  | 
|
2812  | 
(* Geometric progression. *)  | 
|
2813  | 
(* ------------------------------------------------------------------------- *)  | 
|
2814  | 
||
2815  | 
lemma sum_gp_basic: "((1::'a::{field}) - x) * setsum (\<lambda>i. x^i) {0 .. n} = (1 - x^(Suc n))"
 | 
|
2816  | 
(is "?lhs = ?rhs")  | 
|
2817  | 
proof-  | 
|
2818  | 
  {assume x1: "x = 1" hence ?thesis by simp}
 | 
|
2819  | 
moreover  | 
|
2820  | 
  {assume x1: "x\<noteq>1"
 | 
|
2821  | 
hence x1': "x - 1 \<noteq> 0" "1 - x \<noteq> 0" "x - 1 = - (1 - x)" "- (1 - x) \<noteq> 0" by auto  | 
|
2822  | 
from geometric_sum[OF x1, of "Suc n", unfolded x1']  | 
|
2823  | 
    have "(- (1 - x)) * setsum (\<lambda>i. x^i) {0 .. n} = - (1 - x^(Suc n))"
 | 
|
2824  | 
unfolding atLeastLessThanSuc_atLeastAtMost  | 
|
2825  | 
using x1' apply (auto simp only: field_simps)  | 
|
2826  | 
apply (simp add: ring_simps)  | 
|
2827  | 
done  | 
|
2828  | 
then have ?thesis by (simp add: ring_simps) }  | 
|
2829  | 
ultimately show ?thesis by metis  | 
|
2830  | 
qed  | 
|
2831  | 
||
2832  | 
lemma sum_gp_multiplied: assumes mn: "m <= n"  | 
|
2833  | 
  shows "((1::'a::{field}) - x) * setsum (op ^ x) {m..n} = x^m - x^ Suc n"
 | 
|
2834  | 
(is "?lhs = ?rhs")  | 
|
2835  | 
proof-  | 
|
2836  | 
  let ?S = "{0..(n - m)}"
 | 
|
2837  | 
from mn have mn': "n - m \<ge> 0" by arith  | 
|
2838  | 
let ?f = "op + m"  | 
|
2839  | 
have i: "inj_on ?f ?S" unfolding inj_on_def by auto  | 
|
2840  | 
  have f: "?f ` ?S = {m..n}"
 | 
|
2841  | 
using mn apply (auto simp add: image_iff Bex_def) by arith  | 
|
2842  | 
have th: "op ^ x o op + m = (\<lambda>i. x^m * x^i)"  | 
|
2843  | 
by (rule ext, simp add: power_add power_mult)  | 
|
2844  | 
from setsum_reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]]  | 
|
2845  | 
  have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})" by simp
 | 
|
2846  | 
then show ?thesis unfolding sum_gp_basic using mn  | 
|
2847  | 
by (simp add: ring_simps power_add[symmetric])  | 
|
2848  | 
qed  | 
|
2849  | 
||
2850  | 
lemma sum_gp: "setsum (op ^ (x::'a::{field})) {m .. n} =
 | 
|
2851  | 
(if n < m then 0 else if x = 1 then of_nat ((n + 1) - m)  | 
|
2852  | 
else (x^ m - x^ (Suc n)) / (1 - x))"  | 
|
2853  | 
proof-  | 
|
2854  | 
  {assume nm: "n < m" hence ?thesis by simp}
 | 
|
2855  | 
moreover  | 
|
2856  | 
  {assume "\<not> n < m" hence nm: "m \<le> n" by arith
 | 
|
2857  | 
    {assume x: "x = 1"  hence ?thesis by simp}
 | 
|
2858  | 
moreover  | 
|
2859  | 
    {assume x: "x \<noteq> 1" hence nz: "1 - x \<noteq> 0" by simp
 | 
|
2860  | 
from sum_gp_multiplied[OF nm, of x] nz have ?thesis by (simp add: field_simps)}  | 
|
2861  | 
ultimately have ?thesis by metis  | 
|
2862  | 
}  | 
|
2863  | 
ultimately show ?thesis by metis  | 
|
2864  | 
qed  | 
|
2865  | 
||
2866  | 
lemma sum_gp_offset: "setsum (op ^ (x::'a::{field})) {m .. m+n} =
 | 
|
2867  | 
(if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"  | 
|
2868  | 
unfolding sum_gp[of x m "m + n"] power_Suc  | 
|
2869  | 
by (simp add: ring_simps power_add)  | 
|
2870  | 
||
2871  | 
||
2872  | 
subsection{* A bit of linear algebra. *}
 | 
|
2873  | 
||
2874  | 
definition "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 *s x \<in>S )"  | 
|
2875  | 
definition "span S = (subspace hull S)"  | 
|
2876  | 
definition "dependent S \<longleftrightarrow> (\<exists>a \<in> S. a \<in> span(S - {a}))"
 | 
|
2877  | 
abbreviation "independent s == ~(dependent s)"  | 
|
2878  | 
||
2879  | 
(* Closure properties of subspaces. *)  | 
|
2880  | 
||
2881  | 
lemma subspace_UNIV[simp]: "subspace(UNIV)" by (simp add: subspace_def)  | 
|
2882  | 
||
2883  | 
lemma subspace_0: "subspace S ==> 0 \<in> S" by (metis subspace_def)  | 
|
2884  | 
||
2885  | 
lemma subspace_add: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S ==> x + y \<in> S"  | 
|
2886  | 
by (metis subspace_def)  | 
|
2887  | 
||
2888  | 
lemma subspace_mul: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> c *s x \<in> S"  | 
|
2889  | 
by (metis subspace_def)  | 
|
2890  | 
||
| 34289 | 2891  | 
lemma subspace_neg: "subspace S \<Longrightarrow> (x::'a::ring_1^_) \<in> S \<Longrightarrow> - x \<in> S"  | 
| 33175 | 2892  | 
by (metis vector_sneg_minus1 subspace_mul)  | 
2893  | 
||
| 34289 | 2894  | 
lemma subspace_sub: "subspace S \<Longrightarrow> (x::'a::ring_1^_) \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x - y \<in> S"  | 
| 33175 | 2895  | 
by (metis diff_def subspace_add subspace_neg)  | 
2896  | 
||
2897  | 
lemma subspace_setsum:  | 
|
2898  | 
assumes sA: "subspace A" and fB: "finite B"  | 
|
2899  | 
and f: "\<forall>x\<in> B. f x \<in> A"  | 
|
2900  | 
shows "setsum f B \<in> A"  | 
|
2901  | 
using fB f sA  | 
|
2902  | 
apply(induct rule: finite_induct[OF fB])  | 
|
2903  | 
by (simp add: subspace_def sA, auto simp add: sA subspace_add)  | 
|
2904  | 
||
2905  | 
lemma subspace_linear_image:  | 
|
| 34289 | 2906  | 
assumes lf: "linear (f::'a::semiring_1^_ \<Rightarrow> _)" and sS: "subspace S"  | 
| 33175 | 2907  | 
shows "subspace(f ` S)"  | 
2908  | 
using lf sS linear_0[OF lf]  | 
|
2909  | 
unfolding linear_def subspace_def  | 
|
2910  | 
apply (auto simp add: image_iff)  | 
|
2911  | 
apply (rule_tac x="x + y" in bexI, auto)  | 
|
2912  | 
apply (rule_tac x="c*s x" in bexI, auto)  | 
|
2913  | 
done  | 
|
2914  | 
||
| 34289 | 2915  | 
lemma subspace_linear_preimage: "linear (f::'a::semiring_1^_ \<Rightarrow> _) ==> subspace S ==> subspace {x. f x \<in> S}"
 | 
| 33175 | 2916  | 
by (auto simp add: subspace_def linear_def linear_0[of f])  | 
2917  | 
||
2918  | 
lemma subspace_trivial: "subspace {0::'a::semiring_1 ^_}"
 | 
|
2919  | 
by (simp add: subspace_def)  | 
|
2920  | 
||
2921  | 
lemma subspace_inter: "subspace A \<Longrightarrow> subspace B ==> subspace (A \<inter> B)"  | 
|
2922  | 
by (simp add: subspace_def)  | 
|
2923  | 
||
2924  | 
||
2925  | 
lemma span_mono: "A \<subseteq> B ==> span A \<subseteq> span B"  | 
|
2926  | 
by (metis span_def hull_mono)  | 
|
2927  | 
||
2928  | 
lemma subspace_span: "subspace(span S)"  | 
|
2929  | 
unfolding span_def  | 
|
2930  | 
apply (rule hull_in[unfolded mem_def])  | 
|
2931  | 
apply (simp only: subspace_def Inter_iff Int_iff subset_eq)  | 
|
2932  | 
apply auto  | 
|
2933  | 
apply (erule_tac x="X" in ballE)  | 
|
2934  | 
apply (simp add: mem_def)  | 
|
2935  | 
apply blast  | 
|
2936  | 
apply (erule_tac x="X" in ballE)  | 
|
2937  | 
apply (erule_tac x="X" in ballE)  | 
|
2938  | 
apply (erule_tac x="X" in ballE)  | 
|
2939  | 
apply (clarsimp simp add: mem_def)  | 
|
2940  | 
apply simp  | 
|
2941  | 
apply simp  | 
|
2942  | 
apply simp  | 
|
2943  | 
apply (erule_tac x="X" in ballE)  | 
|
2944  | 
apply (erule_tac x="X" in ballE)  | 
|
2945  | 
apply (simp add: mem_def)  | 
|
2946  | 
apply simp  | 
|
2947  | 
apply simp  | 
|
2948  | 
done  | 
|
2949  | 
||
2950  | 
lemma span_clauses:  | 
|
2951  | 
"a \<in> S ==> a \<in> span S"  | 
|
2952  | 
"0 \<in> span S"  | 
|
2953  | 
"x\<in> span S \<Longrightarrow> y \<in> span S ==> x + y \<in> span S"  | 
|
2954  | 
"x \<in> span S \<Longrightarrow> c *s x \<in> span S"  | 
|
2955  | 
by (metis span_def hull_subset subset_eq subspace_span subspace_def)+  | 
|
2956  | 
||
2957  | 
lemma span_induct: assumes SP: "\<And>x. x \<in> S ==> P x"  | 
|
2958  | 
and P: "subspace P" and x: "x \<in> span S" shows "P x"  | 
|
2959  | 
proof-  | 
|
2960  | 
from SP have SP': "S \<subseteq> P" by (simp add: mem_def subset_eq)  | 
|
2961  | 
from P have P': "P \<in> subspace" by (simp add: mem_def)  | 
|
2962  | 
from x hull_minimal[OF SP' P', unfolded span_def[symmetric]]  | 
|
2963  | 
show "P x" by (metis mem_def subset_eq)  | 
|
2964  | 
qed  | 
|
2965  | 
||
| 34289 | 2966  | 
lemma span_empty: "span {} = {(0::'a::semiring_0 ^ _)}"
 | 
| 33175 | 2967  | 
apply (simp add: span_def)  | 
2968  | 
apply (rule hull_unique)  | 
|
2969  | 
apply (auto simp add: mem_def subspace_def)  | 
|
| 34289 | 2970  | 
unfolding mem_def[of "0::'a^_", symmetric]  | 
| 33175 | 2971  | 
apply simp  | 
2972  | 
done  | 
|
2973  | 
||
2974  | 
lemma independent_empty: "independent {}"
 | 
|
2975  | 
by (simp add: dependent_def)  | 
|
2976  | 
||
2977  | 
lemma independent_mono: "independent A \<Longrightarrow> B \<subseteq> A ==> independent B"  | 
|
2978  | 
apply (clarsimp simp add: dependent_def span_mono)  | 
|
2979  | 
  apply (subgoal_tac "span (B - {a}) \<le> span (A - {a})")
 | 
|
2980  | 
apply force  | 
|
2981  | 
apply (rule span_mono)  | 
|
2982  | 
apply auto  | 
|
2983  | 
done  | 
|
2984  | 
||
2985  | 
lemma span_subspace: "A \<subseteq> B \<Longrightarrow> B \<le> span A \<Longrightarrow> subspace B \<Longrightarrow> span A = B"  | 
|
2986  | 
by (metis order_antisym span_def hull_minimal mem_def)  | 
|
2987  | 
||
2988  | 
lemma span_induct': assumes SP: "\<forall>x \<in> S. P x"  | 
|
2989  | 
and P: "subspace P" shows "\<forall>x \<in> span S. P x"  | 
|
2990  | 
using span_induct SP P by blast  | 
|
2991  | 
||
| 34289 | 2992  | 
inductive span_induct_alt_help for S:: "'a::semiring_1^_ \<Rightarrow> bool"  | 
| 33175 | 2993  | 
where  | 
2994  | 
span_induct_alt_help_0: "span_induct_alt_help S 0"  | 
|
2995  | 
| span_induct_alt_help_S: "x \<in> S \<Longrightarrow> span_induct_alt_help S z \<Longrightarrow> span_induct_alt_help S (c *s x + z)"  | 
|
2996  | 
||
2997  | 
lemma span_induct_alt':  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
2998  | 
assumes h0: "h (0::'a::semiring_1^'n)" and hS: "\<And>c x y. x \<in> S \<Longrightarrow> h y \<Longrightarrow> h (c*s x + y)" shows "\<forall>x \<in> span S. h x"  | 
| 33175 | 2999  | 
proof-  | 
3000  | 
  {fix x:: "'a^'n" assume x: "span_induct_alt_help S x"
 | 
|
3001  | 
have "h x"  | 
|
3002  | 
apply (rule span_induct_alt_help.induct[OF x])  | 
|
3003  | 
apply (rule h0)  | 
|
3004  | 
apply (rule hS, assumption, assumption)  | 
|
3005  | 
done}  | 
|
3006  | 
note th0 = this  | 
|
3007  | 
  {fix x assume x: "x \<in> span S"
 | 
|
3008  | 
||
3009  | 
have "span_induct_alt_help S x"  | 
|
3010  | 
proof(rule span_induct[where x=x and S=S])  | 
|
3011  | 
show "x \<in> span S" using x .  | 
|
3012  | 
next  | 
|
3013  | 
fix x assume xS : "x \<in> S"  | 
|
3014  | 
from span_induct_alt_help_S[OF xS span_induct_alt_help_0, of 1]  | 
|
3015  | 
show "span_induct_alt_help S x" by simp  | 
|
3016  | 
next  | 
|
3017  | 
have "span_induct_alt_help S 0" by (rule span_induct_alt_help_0)  | 
|
3018  | 
moreover  | 
|
3019  | 
        {fix x y assume h: "span_induct_alt_help S x" "span_induct_alt_help S y"
 | 
|
3020  | 
from h  | 
|
3021  | 
have "span_induct_alt_help S (x + y)"  | 
|
3022  | 
apply (induct rule: span_induct_alt_help.induct)  | 
|
3023  | 
apply simp  | 
|
3024  | 
unfolding add_assoc  | 
|
3025  | 
apply (rule span_induct_alt_help_S)  | 
|
3026  | 
apply assumption  | 
|
3027  | 
apply simp  | 
|
3028  | 
done}  | 
|
3029  | 
moreover  | 
|
3030  | 
        {fix c x assume xt: "span_induct_alt_help S x"
 | 
|
3031  | 
then have "span_induct_alt_help S (c*s x)"  | 
|
3032  | 
apply (induct rule: span_induct_alt_help.induct)  | 
|
3033  | 
apply (simp add: span_induct_alt_help_0)  | 
|
3034  | 
apply (simp add: vector_smult_assoc vector_add_ldistrib)  | 
|
3035  | 
apply (rule span_induct_alt_help_S)  | 
|
3036  | 
apply assumption  | 
|
3037  | 
apply simp  | 
|
3038  | 
done  | 
|
3039  | 
}  | 
|
3040  | 
ultimately show "subspace (span_induct_alt_help S)"  | 
|
3041  | 
unfolding subspace_def mem_def Ball_def by blast  | 
|
3042  | 
qed}  | 
|
3043  | 
with th0 show ?thesis by blast  | 
|
3044  | 
qed  | 
|
3045  | 
||
3046  | 
lemma span_induct_alt:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3047  | 
assumes h0: "h (0::'a::semiring_1^'n)" and hS: "\<And>c x y. x \<in> S \<Longrightarrow> h y \<Longrightarrow> h (c*s x + y)" and x: "x \<in> span S"  | 
| 33175 | 3048  | 
shows "h x"  | 
3049  | 
using span_induct_alt'[of h S] h0 hS x by blast  | 
|
3050  | 
||
3051  | 
(* Individual closure properties. *)  | 
|
3052  | 
||
3053  | 
lemma span_superset: "x \<in> S ==> x \<in> span S" by (metis span_clauses)  | 
|
3054  | 
||
3055  | 
lemma span_0: "0 \<in> span S" by (metis subspace_span subspace_0)  | 
|
3056  | 
||
3057  | 
lemma span_add: "x \<in> span S \<Longrightarrow> y \<in> span S ==> x + y \<in> span S"  | 
|
3058  | 
by (metis subspace_add subspace_span)  | 
|
3059  | 
||
3060  | 
lemma span_mul: "x \<in> span S ==> (c *s x) \<in> span S"  | 
|
3061  | 
by (metis subspace_span subspace_mul)  | 
|
3062  | 
||
| 34289 | 3063  | 
lemma span_neg: "x \<in> span S ==> - (x::'a::ring_1^_) \<in> span S"  | 
| 33175 | 3064  | 
by (metis subspace_neg subspace_span)  | 
3065  | 
||
| 34289 | 3066  | 
lemma span_sub: "(x::'a::ring_1^_) \<in> span S \<Longrightarrow> y \<in> span S ==> x - y \<in> span S"  | 
| 33175 | 3067  | 
by (metis subspace_span subspace_sub)  | 
3068  | 
||
3069  | 
lemma span_setsum: "finite A \<Longrightarrow> \<forall>x \<in> A. f x \<in> span S ==> setsum f A \<in> span S"  | 
|
3070  | 
apply (rule subspace_setsum)  | 
|
3071  | 
by (metis subspace_span subspace_setsum)+  | 
|
3072  | 
||
| 34289 | 3073  | 
lemma span_add_eq: "(x::'a::ring_1^_) \<in> span S \<Longrightarrow> x + y \<in> span S \<longleftrightarrow> y \<in> span S"  | 
| 33175 | 3074  | 
apply (auto simp only: span_add span_sub)  | 
3075  | 
apply (subgoal_tac "(x + y) - x \<in> span S", simp)  | 
|
3076  | 
by (simp only: span_add span_sub)  | 
|
3077  | 
||
3078  | 
(* Mapping under linear image. *)  | 
|
3079  | 
||
| 34289 | 3080  | 
lemma span_linear_image: assumes lf: "linear (f::'a::semiring_1 ^ _ => _)"  | 
| 33175 | 3081  | 
shows "span (f ` S) = f ` (span S)"  | 
3082  | 
proof-  | 
|
3083  | 
  {fix x
 | 
|
3084  | 
assume x: "x \<in> span (f ` S)"  | 
|
3085  | 
have "x \<in> f ` span S"  | 
|
3086  | 
apply (rule span_induct[where x=x and S = "f ` S"])  | 
|
3087  | 
apply (clarsimp simp add: image_iff)  | 
|
3088  | 
apply (frule span_superset)  | 
|
3089  | 
apply blast  | 
|
3090  | 
apply (simp only: mem_def)  | 
|
3091  | 
apply (rule subspace_linear_image[OF lf])  | 
|
3092  | 
apply (rule subspace_span)  | 
|
3093  | 
apply (rule x)  | 
|
3094  | 
done}  | 
|
3095  | 
moreover  | 
|
3096  | 
  {fix x assume x: "x \<in> span S"
 | 
|
3097  | 
    have th0:"(\<lambda>a. f a \<in> span (f ` S)) = {x. f x \<in> span (f ` S)}" apply (rule set_ext)
 | 
|
3098  | 
unfolding mem_def Collect_def ..  | 
|
3099  | 
have "f x \<in> span (f ` S)"  | 
|
3100  | 
apply (rule span_induct[where S=S])  | 
|
3101  | 
apply (rule span_superset)  | 
|
3102  | 
apply simp  | 
|
3103  | 
apply (subst th0)  | 
|
3104  | 
apply (rule subspace_linear_preimage[OF lf subspace_span, of "f ` S"])  | 
|
3105  | 
apply (rule x)  | 
|
3106  | 
done}  | 
|
3107  | 
ultimately show ?thesis by blast  | 
|
3108  | 
qed  | 
|
3109  | 
||
3110  | 
(* The key breakdown property. *)  | 
|
3111  | 
||
3112  | 
lemma span_breakdown:  | 
|
| 34289 | 3113  | 
assumes bS: "(b::'a::ring_1 ^ _) \<in> S" and aS: "a \<in> span S"  | 
| 33175 | 3114  | 
  shows "\<exists>k. a - k*s b \<in> span (S - {b})" (is "?P a")
 | 
3115  | 
proof-  | 
|
3116  | 
  {fix x assume xS: "x \<in> S"
 | 
|
3117  | 
    {assume ab: "x = b"
 | 
|
3118  | 
then have "?P x"  | 
|
3119  | 
apply simp  | 
|
3120  | 
apply (rule exI[where x="1"], simp)  | 
|
3121  | 
by (rule span_0)}  | 
|
3122  | 
moreover  | 
|
3123  | 
    {assume ab: "x \<noteq> b"
 | 
|
3124  | 
then have "?P x" using xS  | 
|
3125  | 
apply -  | 
|
3126  | 
apply (rule exI[where x=0])  | 
|
3127  | 
apply (rule span_superset)  | 
|
3128  | 
by simp}  | 
|
3129  | 
ultimately have "?P x" by blast}  | 
|
3130  | 
moreover have "subspace ?P"  | 
|
3131  | 
unfolding subspace_def  | 
|
3132  | 
apply auto  | 
|
3133  | 
apply (simp add: mem_def)  | 
|
3134  | 
apply (rule exI[where x=0])  | 
|
3135  | 
    using span_0[of "S - {b}"]
 | 
|
3136  | 
apply (simp add: mem_def)  | 
|
3137  | 
apply (clarsimp simp add: mem_def)  | 
|
3138  | 
apply (rule_tac x="k + ka" in exI)  | 
|
3139  | 
apply (subgoal_tac "x + y - (k + ka) *s b = (x - k*s b) + (y - ka *s b)")  | 
|
3140  | 
apply (simp only: )  | 
|
3141  | 
apply (rule span_add[unfolded mem_def])  | 
|
3142  | 
apply assumption+  | 
|
3143  | 
apply (vector ring_simps)  | 
|
3144  | 
apply (clarsimp simp add: mem_def)  | 
|
3145  | 
apply (rule_tac x= "c*k" in exI)  | 
|
3146  | 
apply (subgoal_tac "c *s x - (c * k) *s b = c*s (x - k*s b)")  | 
|
3147  | 
apply (simp only: )  | 
|
3148  | 
apply (rule span_mul[unfolded mem_def])  | 
|
3149  | 
apply assumption  | 
|
3150  | 
by (vector ring_simps)  | 
|
3151  | 
ultimately show "?P a" using aS span_induct[where S=S and P= "?P"] by metis  | 
|
3152  | 
qed  | 
|
3153  | 
||
3154  | 
lemma span_breakdown_eq:  | 
|
| 34289 | 3155  | 
"(x::'a::ring_1^_) \<in> span (insert a S) \<longleftrightarrow> (\<exists>k. (x - k *s a) \<in> span S)" (is "?lhs \<longleftrightarrow> ?rhs")  | 
| 33175 | 3156  | 
proof-  | 
3157  | 
  {assume x: "x \<in> span (insert a S)"
 | 
|
3158  | 
from x span_breakdown[of "a" "insert a S" "x"]  | 
|
3159  | 
have ?rhs apply clarsimp  | 
|
3160  | 
apply (rule_tac x= "k" in exI)  | 
|
3161  | 
      apply (rule set_rev_mp[of _ "span (S - {a})" _])
 | 
|
3162  | 
apply assumption  | 
|
3163  | 
apply (rule span_mono)  | 
|
3164  | 
apply blast  | 
|
3165  | 
done}  | 
|
3166  | 
moreover  | 
|
3167  | 
  { fix k assume k: "x - k *s a \<in> span S"
 | 
|
3168  | 
have eq: "x = (x - k *s a) + k *s a" by vector  | 
|
3169  | 
have "(x - k *s a) + k *s a \<in> span (insert a S)"  | 
|
3170  | 
apply (rule span_add)  | 
|
3171  | 
apply (rule set_rev_mp[of _ "span S" _])  | 
|
3172  | 
apply (rule k)  | 
|
3173  | 
apply (rule span_mono)  | 
|
3174  | 
apply blast  | 
|
3175  | 
apply (rule span_mul)  | 
|
3176  | 
apply (rule span_superset)  | 
|
3177  | 
apply blast  | 
|
3178  | 
done  | 
|
3179  | 
then have ?lhs using eq by metis}  | 
|
3180  | 
ultimately show ?thesis by blast  | 
|
3181  | 
qed  | 
|
3182  | 
||
3183  | 
(* Hence some "reversal" results.*)  | 
|
3184  | 
||
3185  | 
lemma in_span_insert:  | 
|
| 34289 | 3186  | 
assumes a: "(a::'a::field^_) \<in> span (insert b S)" and na: "a \<notin> span S"  | 
| 33175 | 3187  | 
shows "b \<in> span (insert a S)"  | 
3188  | 
proof-  | 
|
3189  | 
from span_breakdown[of b "insert b S" a, OF insertI1 a]  | 
|
3190  | 
  obtain k where k: "a - k*s b \<in> span (S - {b})" by auto
 | 
|
3191  | 
  {assume k0: "k = 0"
 | 
|
3192  | 
with k have "a \<in> span S"  | 
|
3193  | 
apply (simp)  | 
|
3194  | 
apply (rule set_rev_mp)  | 
|
3195  | 
apply assumption  | 
|
3196  | 
apply (rule span_mono)  | 
|
3197  | 
apply blast  | 
|
3198  | 
done  | 
|
3199  | 
with na have ?thesis by blast}  | 
|
3200  | 
moreover  | 
|
3201  | 
  {assume k0: "k \<noteq> 0"
 | 
|
3202  | 
have eq: "b = (1/k) *s a - ((1/k) *s a - b)" by vector  | 
|
3203  | 
from k0 have eq': "(1/k) *s (a - k*s b) = (1/k) *s a - b"  | 
|
3204  | 
by (vector field_simps)  | 
|
3205  | 
    from k have "(1/k) *s (a - k*s b) \<in> span (S - {b})"
 | 
|
3206  | 
by (rule span_mul)  | 
|
3207  | 
    hence th: "(1/k) *s a - b \<in> span (S - {b})"
 | 
|
3208  | 
unfolding eq' .  | 
|
3209  | 
||
3210  | 
from k  | 
|
3211  | 
have ?thesis  | 
|
3212  | 
apply (subst eq)  | 
|
3213  | 
apply (rule span_sub)  | 
|
3214  | 
apply (rule span_mul)  | 
|
3215  | 
apply (rule span_superset)  | 
|
3216  | 
apply blast  | 
|
3217  | 
apply (rule set_rev_mp)  | 
|
3218  | 
apply (rule th)  | 
|
3219  | 
apply (rule span_mono)  | 
|
3220  | 
using na by blast}  | 
|
3221  | 
ultimately show ?thesis by blast  | 
|
3222  | 
qed  | 
|
3223  | 
||
3224  | 
lemma in_span_delete:  | 
|
| 34289 | 3225  | 
assumes a: "(a::'a::field^_) \<in> span S"  | 
| 33175 | 3226  | 
  and na: "a \<notin> span (S-{b})"
 | 
3227  | 
  shows "b \<in> span (insert a (S - {b}))"
 | 
|
3228  | 
apply (rule in_span_insert)  | 
|
3229  | 
apply (rule set_rev_mp)  | 
|
3230  | 
apply (rule a)  | 
|
3231  | 
apply (rule span_mono)  | 
|
3232  | 
apply blast  | 
|
3233  | 
apply (rule na)  | 
|
3234  | 
done  | 
|
3235  | 
||
3236  | 
(* Transitivity property. *)  | 
|
3237  | 
||
3238  | 
lemma span_trans:  | 
|
| 34289 | 3239  | 
assumes x: "(x::'a::ring_1^_) \<in> span S" and y: "y \<in> span (insert x S)"  | 
| 33175 | 3240  | 
shows "y \<in> span S"  | 
3241  | 
proof-  | 
|
3242  | 
from span_breakdown[of x "insert x S" y, OF insertI1 y]  | 
|
3243  | 
  obtain k where k: "y -k*s x \<in> span (S - {x})" by auto
 | 
|
3244  | 
have eq: "y = (y - k *s x) + k *s x" by vector  | 
|
3245  | 
show ?thesis  | 
|
3246  | 
apply (subst eq)  | 
|
3247  | 
apply (rule span_add)  | 
|
3248  | 
apply (rule set_rev_mp)  | 
|
3249  | 
apply (rule k)  | 
|
3250  | 
apply (rule span_mono)  | 
|
3251  | 
apply blast  | 
|
3252  | 
apply (rule span_mul)  | 
|
3253  | 
by (rule x)  | 
|
3254  | 
qed  | 
|
3255  | 
||
3256  | 
(* ------------------------------------------------------------------------- *)  | 
|
3257  | 
(* An explicit expansion is sometimes needed. *)  | 
|
3258  | 
(* ------------------------------------------------------------------------- *)  | 
|
3259  | 
||
3260  | 
lemma span_explicit:  | 
|
| 34289 | 3261  | 
  "span P = {y::'a::semiring_1^_. \<exists>S u. finite S \<and> S \<subseteq> P \<and> setsum (\<lambda>v. u v *s v) S = y}"
 | 
| 33175 | 3262  | 
  (is "_ = ?E" is "_ = {y. ?h y}" is "_ = {y. \<exists>S u. ?Q S u y}")
 | 
3263  | 
proof-  | 
|
3264  | 
  {fix x assume x: "x \<in> ?E"
 | 
|
3265  | 
then obtain S u where fS: "finite S" and SP: "S\<subseteq>P" and u: "setsum (\<lambda>v. u v *s v) S = x"  | 
|
3266  | 
by blast  | 
|
3267  | 
have "x \<in> span P"  | 
|
3268  | 
unfolding u[symmetric]  | 
|
3269  | 
apply (rule span_setsum[OF fS])  | 
|
3270  | 
using span_mono[OF SP]  | 
|
3271  | 
by (auto intro: span_superset span_mul)}  | 
|
3272  | 
moreover  | 
|
3273  | 
have "\<forall>x \<in> span P. x \<in> ?E"  | 
|
3274  | 
unfolding mem_def Collect_def  | 
|
3275  | 
proof(rule span_induct_alt')  | 
|
3276  | 
show "?h 0"  | 
|
3277  | 
      apply (rule exI[where x="{}"]) by simp
 | 
|
3278  | 
next  | 
|
3279  | 
fix c x y  | 
|
3280  | 
assume x: "x \<in> P" and hy: "?h y"  | 
|
3281  | 
from hy obtain S u where fS: "finite S" and SP: "S\<subseteq>P"  | 
|
3282  | 
and u: "setsum (\<lambda>v. u v *s v) S = y" by blast  | 
|
3283  | 
let ?S = "insert x S"  | 
|
3284  | 
let ?u = "\<lambda>y. if y = x then (if x \<in> S then u y + c else c)  | 
|
3285  | 
else u y"  | 
|
3286  | 
from fS SP x have th0: "finite (insert x S)" "insert x S \<subseteq> P" by blast+  | 
|
3287  | 
    {assume xS: "x \<in> S"
 | 
|
3288  | 
      have S1: "S = (S - {x}) \<union> {x}"
 | 
|
3289  | 
        and Sss:"finite (S - {x})" "finite {x}" "(S -{x}) \<inter> {x} = {}" using xS fS by auto
 | 
|
3290  | 
      have "setsum (\<lambda>v. ?u v *s v) ?S =(\<Sum>v\<in>S - {x}. u v *s v) + (u x + c) *s x"
 | 
|
3291  | 
using xS  | 
|
3292  | 
by (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]]  | 
|
3293  | 
setsum_clauses(2)[OF fS] cong del: if_weak_cong)  | 
|
3294  | 
also have "\<dots> = (\<Sum>v\<in>S. u v *s v) + c *s x"  | 
|
3295  | 
apply (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]])  | 
|
3296  | 
by (vector ring_simps)  | 
|
3297  | 
also have "\<dots> = c*s x + y"  | 
|
3298  | 
by (simp add: add_commute u)  | 
|
3299  | 
finally have "setsum (\<lambda>v. ?u v *s v) ?S = c*s x + y" .  | 
|
3300  | 
then have "?Q ?S ?u (c*s x + y)" using th0 by blast}  | 
|
3301  | 
moreover  | 
|
3302  | 
  {assume xS: "x \<notin> S"
 | 
|
3303  | 
have th00: "(\<Sum>v\<in>S. (if v = x then c else u v) *s v) = y"  | 
|
3304  | 
unfolding u[symmetric]  | 
|
3305  | 
apply (rule setsum_cong2)  | 
|
3306  | 
using xS by auto  | 
|
3307  | 
have "?Q ?S ?u (c*s x + y)" using fS xS th0  | 
|
3308  | 
by (simp add: th00 setsum_clauses add_commute cong del: if_weak_cong)}  | 
|
3309  | 
ultimately have "?Q ?S ?u (c*s x + y)"  | 
|
3310  | 
by (cases "x \<in> S", simp, simp)  | 
|
3311  | 
then show "?h (c*s x + y)"  | 
|
3312  | 
apply -  | 
|
3313  | 
apply (rule exI[where x="?S"])  | 
|
3314  | 
apply (rule exI[where x="?u"]) by metis  | 
|
3315  | 
qed  | 
|
3316  | 
ultimately show ?thesis by blast  | 
|
3317  | 
qed  | 
|
3318  | 
||
3319  | 
lemma dependent_explicit:  | 
|
| 34289 | 3320  | 
  "dependent P \<longleftrightarrow> (\<exists>S u. finite S \<and> S \<subseteq> P \<and> (\<exists>(v::'a::{idom,field}^_) \<in>S. u v \<noteq> 0 \<and> setsum (\<lambda>v. u v *s v) S = 0))" (is "?lhs = ?rhs")
 | 
| 33175 | 3321  | 
proof-  | 
3322  | 
  {assume dP: "dependent P"
 | 
|
3323  | 
then obtain a S u where aP: "a \<in> P" and fS: "finite S"  | 
|
3324  | 
      and SP: "S \<subseteq> P - {a}" and ua: "setsum (\<lambda>v. u v *s v) S = a"
 | 
|
3325  | 
unfolding dependent_def span_explicit by blast  | 
|
3326  | 
let ?S = "insert a S"  | 
|
3327  | 
let ?u = "\<lambda>y. if y = a then - 1 else u y"  | 
|
3328  | 
let ?v = a  | 
|
3329  | 
from aP SP have aS: "a \<notin> S" by blast  | 
|
3330  | 
from fS SP aP have th0: "finite ?S" "?S \<subseteq> P" "?v \<in> ?S" "?u ?v \<noteq> 0" by auto  | 
|
3331  | 
have s0: "setsum (\<lambda>v. ?u v *s v) ?S = 0"  | 
|
3332  | 
using fS aS  | 
|
3333  | 
apply (simp add: vector_smult_lneg vector_smult_lid setsum_clauses ring_simps )  | 
|
3334  | 
apply (subst (2) ua[symmetric])  | 
|
3335  | 
apply (rule setsum_cong2)  | 
|
3336  | 
by auto  | 
|
3337  | 
with th0 have ?rhs  | 
|
3338  | 
apply -  | 
|
3339  | 
apply (rule exI[where x= "?S"])  | 
|
3340  | 
apply (rule exI[where x= "?u"])  | 
|
3341  | 
by clarsimp}  | 
|
3342  | 
moreover  | 
|
3343  | 
  {fix S u v assume fS: "finite S"
 | 
|
3344  | 
and SP: "S \<subseteq> P" and vS: "v \<in> S" and uv: "u v \<noteq> 0"  | 
|
3345  | 
and u: "setsum (\<lambda>v. u v *s v) S = 0"  | 
|
3346  | 
let ?a = v  | 
|
3347  | 
    let ?S = "S - {v}"
 | 
|
3348  | 
let ?u = "\<lambda>i. (- u i) / u v"  | 
|
3349  | 
have th0: "?a \<in> P" "finite ?S" "?S \<subseteq> P" using fS SP vS by auto  | 
|
3350  | 
have "setsum (\<lambda>v. ?u v *s v) ?S = setsum (\<lambda>v. (- (inverse (u ?a))) *s (u v *s v)) S - ?u v *s v"  | 
|
3351  | 
using fS vS uv  | 
|
3352  | 
by (simp add: setsum_diff1 vector_smult_lneg divide_inverse  | 
|
3353  | 
vector_smult_assoc field_simps)  | 
|
3354  | 
also have "\<dots> = ?a"  | 
|
3355  | 
unfolding setsum_cmul u  | 
|
3356  | 
using uv by (simp add: vector_smult_lneg)  | 
|
3357  | 
finally have "setsum (\<lambda>v. ?u v *s v) ?S = ?a" .  | 
|
3358  | 
with th0 have ?lhs  | 
|
3359  | 
unfolding dependent_def span_explicit  | 
|
3360  | 
apply -  | 
|
3361  | 
apply (rule bexI[where x= "?a"])  | 
|
3362  | 
apply simp_all  | 
|
3363  | 
apply (rule exI[where x= "?S"])  | 
|
3364  | 
by auto}  | 
|
3365  | 
ultimately show ?thesis by blast  | 
|
3366  | 
qed  | 
|
3367  | 
||
3368  | 
||
3369  | 
lemma span_finite:  | 
|
3370  | 
assumes fS: "finite S"  | 
|
| 34289 | 3371  | 
  shows "span S = {(y::'a::semiring_1^_). \<exists>u. setsum (\<lambda>v. u v *s v) S = y}"
 | 
| 33175 | 3372  | 
(is "_ = ?rhs")  | 
3373  | 
proof-  | 
|
3374  | 
  {fix y assume y: "y \<in> span S"
 | 
|
3375  | 
from y obtain S' u where fS': "finite S'" and SS': "S' \<subseteq> S" and  | 
|
3376  | 
u: "setsum (\<lambda>v. u v *s v) S' = y" unfolding span_explicit by blast  | 
|
3377  | 
let ?u = "\<lambda>x. if x \<in> S' then u x else 0"  | 
|
3378  | 
from setsum_restrict_set[OF fS, of "\<lambda>v. u v *s v" S', symmetric] SS'  | 
|
3379  | 
have "setsum (\<lambda>v. ?u v *s v) S = setsum (\<lambda>v. u v *s v) S'"  | 
|
3380  | 
unfolding cond_value_iff cond_application_beta  | 
|
3381  | 
by (simp add: cond_value_iff inf_absorb2 cong del: if_weak_cong)  | 
|
3382  | 
hence "setsum (\<lambda>v. ?u v *s v) S = y" by (metis u)  | 
|
3383  | 
hence "y \<in> ?rhs" by auto}  | 
|
3384  | 
moreover  | 
|
3385  | 
  {fix y u assume u: "setsum (\<lambda>v. u v *s v) S = y"
 | 
|
3386  | 
then have "y \<in> span S" using fS unfolding span_explicit by auto}  | 
|
3387  | 
ultimately show ?thesis by blast  | 
|
3388  | 
qed  | 
|
3389  | 
||
3390  | 
||
3391  | 
(* Standard bases are a spanning set, and obviously finite. *)  | 
|
3392  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3393  | 
lemma span_stdbasis:"span {basis i :: 'a::ring_1^'n | i. i \<in> (UNIV :: 'n set)} = UNIV"
 | 
| 33175 | 3394  | 
apply (rule set_ext)  | 
3395  | 
apply auto  | 
|
3396  | 
apply (subst basis_expansion[symmetric])  | 
|
3397  | 
apply (rule span_setsum)  | 
|
3398  | 
apply simp  | 
|
3399  | 
apply auto  | 
|
3400  | 
apply (rule span_mul)  | 
|
3401  | 
apply (rule span_superset)  | 
|
3402  | 
apply (auto simp add: Collect_def mem_def)  | 
|
3403  | 
done  | 
|
3404  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3405  | 
lemma finite_stdbasis: "finite {basis i ::real^'n |i. i\<in> (UNIV:: 'n set)}" (is "finite ?S")
 | 
| 33175 | 3406  | 
proof-  | 
3407  | 
have eq: "?S = basis ` UNIV" by blast  | 
|
| 33715 | 3408  | 
show ?thesis unfolding eq by auto  | 
| 33175 | 3409  | 
qed  | 
3410  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3411  | 
lemma card_stdbasis: "card {basis i ::real^'n |i. i\<in> (UNIV :: 'n set)} = CARD('n)" (is "card ?S = _")
 | 
| 33715 | 3412  | 
proof-  | 
3413  | 
have eq: "?S = basis ` UNIV" by blast  | 
|
3414  | 
show ?thesis unfolding eq using card_image[OF basis_inj] by simp  | 
|
3415  | 
qed  | 
|
3416  | 
||
| 33175 | 3417  | 
|
3418  | 
lemma independent_stdbasis_lemma:  | 
|
| 34289 | 3419  | 
assumes x: "(x::'a::semiring_1 ^ _) \<in> span (basis ` S)"  | 
| 33175 | 3420  | 
and iS: "i \<notin> S"  | 
3421  | 
shows "(x$i) = 0"  | 
|
3422  | 
proof-  | 
|
3423  | 
let ?U = "UNIV :: 'n set"  | 
|
3424  | 
let ?B = "basis ` S"  | 
|
| 34289 | 3425  | 
let ?P = "\<lambda>(x::'a^_). \<forall>i\<in> ?U. i \<notin> S \<longrightarrow> x$i =0"  | 
3426  | 
 {fix x::"'a^_" assume xS: "x\<in> ?B"
 | 
|
| 33175 | 3427  | 
from xS have "?P x" by auto}  | 
3428  | 
moreover  | 
|
3429  | 
have "subspace ?P"  | 
|
3430  | 
by (auto simp add: subspace_def Collect_def mem_def)  | 
|
3431  | 
ultimately show ?thesis  | 
|
3432  | 
using x span_induct[of ?B ?P x] iS by blast  | 
|
3433  | 
qed  | 
|
3434  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3435  | 
lemma independent_stdbasis: "independent {basis i ::real^'n |i. i\<in> (UNIV :: 'n set)}"
 | 
| 33175 | 3436  | 
proof-  | 
3437  | 
let ?I = "UNIV :: 'n set"  | 
|
3438  | 
let ?b = "basis :: _ \<Rightarrow> real ^'n"  | 
|
3439  | 
let ?B = "?b ` ?I"  | 
|
3440  | 
  have eq: "{?b i|i. i \<in> ?I} = ?B"
 | 
|
3441  | 
by auto  | 
|
3442  | 
  {assume d: "dependent ?B"
 | 
|
3443  | 
    then obtain k where k: "k \<in> ?I" "?b k \<in> span (?B - {?b k})"
 | 
|
3444  | 
unfolding dependent_def by auto  | 
|
3445  | 
    have eq1: "?B - {?b k} = ?B - ?b ` {k}"  by simp
 | 
|
3446  | 
    have eq2: "?B - {?b k} = ?b ` (?I - {k})"
 | 
|
3447  | 
unfolding eq1  | 
|
3448  | 
apply (rule inj_on_image_set_diff[symmetric])  | 
|
3449  | 
apply (rule basis_inj) using k(1) by auto  | 
|
3450  | 
    from k(2) have th0: "?b k \<in> span (?b ` (?I - {k}))" unfolding eq2 .
 | 
|
3451  | 
from independent_stdbasis_lemma[OF th0, of k, simplified]  | 
|
3452  | 
have False by simp}  | 
|
3453  | 
then show ?thesis unfolding eq dependent_def ..  | 
|
3454  | 
qed  | 
|
3455  | 
||
3456  | 
(* This is useful for building a basis step-by-step. *)  | 
|
3457  | 
||
3458  | 
lemma independent_insert:  | 
|
| 34289 | 3459  | 
"independent(insert (a::'a::field ^_) S) \<longleftrightarrow>  | 
| 33175 | 3460  | 
(if a \<in> S then independent S  | 
3461  | 
else independent S \<and> a \<notin> span S)" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
3462  | 
proof-  | 
|
3463  | 
  {assume aS: "a \<in> S"
 | 
|
3464  | 
hence ?thesis using insert_absorb[OF aS] by simp}  | 
|
3465  | 
moreover  | 
|
3466  | 
  {assume aS: "a \<notin> S"
 | 
|
3467  | 
    {assume i: ?lhs
 | 
|
3468  | 
then have ?rhs using aS  | 
|
3469  | 
apply simp  | 
|
3470  | 
apply (rule conjI)  | 
|
3471  | 
apply (rule independent_mono)  | 
|
3472  | 
apply assumption  | 
|
3473  | 
apply blast  | 
|
3474  | 
by (simp add: dependent_def)}  | 
|
3475  | 
moreover  | 
|
3476  | 
    {assume i: ?rhs
 | 
|
3477  | 
have ?lhs using i aS  | 
|
3478  | 
apply simp  | 
|
3479  | 
apply (auto simp add: dependent_def)  | 
|
3480  | 
apply (case_tac "aa = a", auto)  | 
|
3481  | 
        apply (subgoal_tac "insert a S - {aa} = insert a (S - {aa})")
 | 
|
3482  | 
apply simp  | 
|
3483  | 
        apply (subgoal_tac "a \<in> span (insert aa (S - {aa}))")
 | 
|
3484  | 
        apply (subgoal_tac "insert aa (S - {aa}) = S")
 | 
|
3485  | 
apply simp  | 
|
3486  | 
apply blast  | 
|
3487  | 
apply (rule in_span_insert)  | 
|
3488  | 
apply assumption  | 
|
3489  | 
apply blast  | 
|
3490  | 
apply blast  | 
|
3491  | 
done}  | 
|
3492  | 
ultimately have ?thesis by blast}  | 
|
3493  | 
ultimately show ?thesis by blast  | 
|
3494  | 
qed  | 
|
3495  | 
||
3496  | 
(* The degenerate case of the Exchange Lemma. *)  | 
|
3497  | 
||
3498  | 
lemma mem_delete: "x \<in> (A - {a}) \<longleftrightarrow> x \<noteq> a \<and> x \<in> A"
 | 
|
3499  | 
by blast  | 
|
3500  | 
||
3501  | 
lemma span_span: "span (span A) = span A"  | 
|
3502  | 
unfolding span_def hull_hull ..  | 
|
3503  | 
||
3504  | 
lemma span_inc: "S \<subseteq> span S"  | 
|
3505  | 
by (metis subset_eq span_superset)  | 
|
3506  | 
||
3507  | 
lemma spanning_subset_independent:  | 
|
| 34289 | 3508  | 
  assumes BA: "B \<subseteq> A" and iA: "independent (A::('a::field ^_) set)"
 | 
| 33175 | 3509  | 
and AsB: "A \<subseteq> span B"  | 
3510  | 
shows "A = B"  | 
|
3511  | 
proof  | 
|
3512  | 
from BA show "B \<subseteq> A" .  | 
|
3513  | 
next  | 
|
3514  | 
from span_mono[OF BA] span_mono[OF AsB]  | 
|
3515  | 
have sAB: "span A = span B" unfolding span_span by blast  | 
|
3516  | 
||
3517  | 
  {fix x assume x: "x \<in> A"
 | 
|
3518  | 
    from iA have th0: "x \<notin> span (A - {x})"
 | 
|
3519  | 
unfolding dependent_def using x by blast  | 
|
3520  | 
from x have xsA: "x \<in> span A" by (blast intro: span_superset)  | 
|
3521  | 
    have "A - {x} \<subseteq> A" by blast
 | 
|
3522  | 
    hence th1:"span (A - {x}) \<subseteq> span A" by (metis span_mono)
 | 
|
3523  | 
    {assume xB: "x \<notin> B"
 | 
|
3524  | 
      from xB BA have "B \<subseteq> A -{x}" by blast
 | 
|
3525  | 
      hence "span B \<subseteq> span (A - {x})" by (metis span_mono)
 | 
|
3526  | 
with th1 th0 sAB have "x \<notin> span A" by blast  | 
|
3527  | 
with x have False by (metis span_superset)}  | 
|
3528  | 
then have "x \<in> B" by blast}  | 
|
3529  | 
then show "A \<subseteq> B" by blast  | 
|
3530  | 
qed  | 
|
3531  | 
||
3532  | 
(* The general case of the Exchange Lemma, the key to what follows. *)  | 
|
3533  | 
||
3534  | 
lemma exchange_lemma:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3535  | 
  assumes f:"finite (t:: ('a::field^'n) set)" and i: "independent s"
 | 
| 33175 | 3536  | 
and sp:"s \<subseteq> span t"  | 
| 33715 | 3537  | 
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 | 3538  | 
using f i sp  | 
| 34915 | 3539  | 
proof(induct "card (t - s)" arbitrary: s t rule: less_induct)  | 
3540  | 
case less  | 
|
3541  | 
note ft = `finite t` and s = `independent s` and sp = `s \<subseteq> span t`  | 
|
| 33715 | 3542  | 
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 | 3543  | 
let ?ths = "\<exists>t'. ?P t'"  | 
3544  | 
  {assume st: "s \<subseteq> t"
 | 
|
3545  | 
from st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t])  | 
|
| 33715 | 3546  | 
by (auto intro: span_superset)}  | 
| 33175 | 3547  | 
moreover  | 
3548  | 
  {assume st: "t \<subseteq> s"
 | 
|
3549  | 
||
3550  | 
from spanning_subset_independent[OF st s sp]  | 
|
3551  | 
st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t])  | 
|
| 33715 | 3552  | 
by (auto intro: span_superset)}  | 
| 33175 | 3553  | 
moreover  | 
3554  | 
  {assume st: "\<not> s \<subseteq> t" "\<not> t \<subseteq> s"
 | 
|
3555  | 
from st(2) obtain b where b: "b \<in> t" "b \<notin> s" by blast  | 
|
3556  | 
      from b have "t - {b} - s \<subset> t - s" by blast
 | 
|
| 34915 | 3557  | 
      then have cardlt: "card (t - {b} - s) < card (t - s)" using ft
 | 
| 33175 | 3558  | 
by (auto intro: psubset_card_mono)  | 
3559  | 
from b ft have ct0: "card t \<noteq> 0" by auto  | 
|
3560  | 
    {assume stb: "s \<subseteq> span(t -{b})"
 | 
|
3561  | 
      from ft have ftb: "finite (t -{b})" by auto
 | 
|
| 34915 | 3562  | 
from less(1)[OF cardlt ftb s stb]  | 
| 33715 | 3563  | 
      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 | 3564  | 
let ?w = "insert b u"  | 
3565  | 
have th0: "s \<subseteq> insert b u" using u by blast  | 
|
3566  | 
from u(3) b have "u \<subseteq> s \<union> t" by blast  | 
|
3567  | 
then have th1: "insert b u \<subseteq> s \<union> t" using u b by blast  | 
|
3568  | 
have bu: "b \<notin> u" using b u by blast  | 
|
| 33715 | 3569  | 
from u(1) ft b have "card u = (card t - 1)" by auto  | 
| 33175 | 3570  | 
then  | 
| 33715 | 3571  | 
have th2: "card (insert b u) = card t"  | 
3572  | 
using card_insert_disjoint[OF fu bu] ct0 by auto  | 
|
| 33175 | 3573  | 
from u(4) have "s \<subseteq> span u" .  | 
3574  | 
also have "\<dots> \<subseteq> span (insert b u)" apply (rule span_mono) by blast  | 
|
| 33715 | 3575  | 
finally have th3: "s \<subseteq> span (insert b u)" .  | 
3576  | 
from th0 th1 th2 th3 fu have th: "?P ?w" by blast  | 
|
| 33175 | 3577  | 
from th have ?ths by blast}  | 
3578  | 
moreover  | 
|
3579  | 
    {assume stb: "\<not> s \<subseteq> span(t -{b})"
 | 
|
3580  | 
      from stb obtain a where a: "a \<in> s" "a \<notin> span (t - {b})" by blast
 | 
|
3581  | 
have ab: "a \<noteq> b" using a b by blast  | 
|
3582  | 
      have at: "a \<notin> t" using a ab span_superset[of a "t- {b}"] by auto
 | 
|
| 34915 | 3583  | 
      have mlt: "card ((insert a (t - {b})) - s) < card (t - s)"
 | 
3584  | 
using cardlt ft a b by auto  | 
|
| 33175 | 3585  | 
      have ft': "finite (insert a (t - {b}))" using ft by auto
 | 
3586  | 
      {fix x assume xs: "x \<in> s"
 | 
|
3587  | 
        have t: "t \<subseteq> (insert b (insert a (t -{b})))" using b by auto
 | 
|
3588  | 
from b(1) have "b \<in> span t" by (simp add: span_superset)  | 
|
3589  | 
        have bs: "b \<in> span (insert a (t - {b}))"
 | 
|
3590  | 
by (metis in_span_delete a sp mem_def subset_eq)  | 
|
3591  | 
from xs sp have "x \<in> span t" by blast  | 
|
3592  | 
with span_mono[OF t]  | 
|
3593  | 
        have x: "x \<in> span (insert b (insert a (t - {b})))" ..
 | 
|
3594  | 
        from span_trans[OF bs x] have "x \<in> span (insert a (t - {b}))"  .}
 | 
|
3595  | 
      then have sp': "s \<subseteq> span (insert a (t - {b}))" by blast
 | 
|
3596  | 
||
| 34915 | 3597  | 
from less(1)[OF mlt ft' s sp'] obtain u where  | 
| 33715 | 3598  | 
        u: "card u = card (insert a (t -{b}))" "finite u" "s \<subseteq> u" "u \<subseteq> s \<union> insert a (t -{b})"
 | 
| 33175 | 3599  | 
"s \<subseteq> span u" by blast  | 
| 33715 | 3600  | 
from u a b ft at ct0 have "?P u" by auto  | 
| 33175 | 3601  | 
then have ?ths by blast }  | 
3602  | 
ultimately have ?ths by blast  | 
|
3603  | 
}  | 
|
3604  | 
ultimately  | 
|
3605  | 
show ?ths by blast  | 
|
3606  | 
qed  | 
|
3607  | 
||
3608  | 
(* This implies corresponding size bounds. *)  | 
|
3609  | 
||
3610  | 
lemma independent_span_bound:  | 
|
| 34289 | 3611  | 
  assumes f: "finite t" and i: "independent (s::('a::field^_) set)" and sp:"s \<subseteq> span t"
 | 
| 33175 | 3612  | 
shows "finite s \<and> card s \<le> card t"  | 
| 33715 | 3613  | 
by (metis exchange_lemma[OF f i sp] finite_subset card_mono)  | 
| 33175 | 3614  | 
|
3615  | 
||
3616  | 
lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x\<in> (UNIV::'a::finite set)}"
 | 
|
3617  | 
proof-  | 
|
3618  | 
  have eq: "{f x |x. x\<in> UNIV} = f ` UNIV" by auto
 | 
|
3619  | 
show ?thesis unfolding eq  | 
|
3620  | 
apply (rule finite_imageI)  | 
|
3621  | 
apply (rule finite)  | 
|
3622  | 
done  | 
|
3623  | 
qed  | 
|
3624  | 
||
3625  | 
||
3626  | 
lemma independent_bound:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3627  | 
fixes S:: "(real^'n) set"  | 
| 33175 | 3628  | 
  shows "independent S \<Longrightarrow> finite S \<and> card S <= CARD('n)"
 | 
3629  | 
apply (subst card_stdbasis[symmetric])  | 
|
3630  | 
apply (rule independent_span_bound)  | 
|
3631  | 
apply (rule finite_Atleast_Atmost_nat)  | 
|
3632  | 
apply assumption  | 
|
3633  | 
unfolding span_stdbasis  | 
|
3634  | 
apply (rule subset_UNIV)  | 
|
3635  | 
done  | 
|
3636  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3637  | 
lemma dependent_biggerset: "(finite (S::(real ^'n) set) ==> card S > CARD('n)) ==> dependent S"
 | 
| 33175 | 3638  | 
by (metis independent_bound not_less)  | 
3639  | 
||
3640  | 
(* Hence we can create a maximal independent subset. *)  | 
|
3641  | 
||
3642  | 
lemma maximal_independent_subset_extend:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3643  | 
assumes sv: "(S::(real^'n) set) \<subseteq> V" and iS: "independent S"  | 
| 33175 | 3644  | 
shows "\<exists>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B"  | 
3645  | 
using sv iS  | 
|
| 34915 | 3646  | 
proof(induct "CARD('n) - card S" arbitrary: S rule: less_induct)
 | 
3647  | 
case less  | 
|
3648  | 
note sv = `S \<subseteq> V` and i = `independent S`  | 
|
| 33175 | 3649  | 
let ?P = "\<lambda>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B"  | 
3650  | 
let ?ths = "\<exists>x. ?P x"  | 
|
3651  | 
  let ?d = "CARD('n)"
 | 
|
3652  | 
  {assume "V \<subseteq> span S"
 | 
|
3653  | 
then have ?ths using sv i by blast }  | 
|
3654  | 
moreover  | 
|
3655  | 
  {assume VS: "\<not> V \<subseteq> span S"
 | 
|
3656  | 
from VS obtain a where a: "a \<in> V" "a \<notin> span S" by blast  | 
|
3657  | 
from a have aS: "a \<notin> S" by (auto simp add: span_superset)  | 
|
3658  | 
have th0: "insert a S \<subseteq> V" using a sv by blast  | 
|
3659  | 
from independent_insert[of a S] i a  | 
|
3660  | 
have th1: "independent (insert a S)" by auto  | 
|
| 34915 | 3661  | 
have mlt: "?d - card (insert a S) < ?d - card S"  | 
3662  | 
using aS a independent_bound[OF th1]  | 
|
| 33175 | 3663  | 
by auto  | 
3664  | 
||
| 34915 | 3665  | 
from less(1)[OF mlt th0 th1]  | 
| 33175 | 3666  | 
obtain B where B: "insert a S \<subseteq> B" "B \<subseteq> V" "independent B" " V \<subseteq> span B"  | 
3667  | 
by blast  | 
|
3668  | 
from B have "?P B" by auto  | 
|
3669  | 
then have ?ths by blast}  | 
|
3670  | 
ultimately show ?ths by blast  | 
|
3671  | 
qed  | 
|
3672  | 
||
3673  | 
lemma maximal_independent_subset:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3674  | 
"\<exists>(B:: (real ^'n) set). B\<subseteq> V \<and> independent B \<and> V \<subseteq> span B"  | 
| 33175 | 3675  | 
  by (metis maximal_independent_subset_extend[of "{}:: (real ^'n) set"] empty_subsetI independent_empty)
 | 
3676  | 
||
3677  | 
(* Notion of dimension. *)  | 
|
3678  | 
||
| 33715 | 3679  | 
definition "dim V = (SOME n. \<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (card B = n))"  | 
3680  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3681  | 
lemma basis_exists: "\<exists>B. (B :: (real ^'n) set) \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (card B = dim V)"  | 
| 33715 | 3682  | 
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 | 3683  | 
using maximal_independent_subset[of V] independent_bound  | 
3684  | 
by auto  | 
|
3685  | 
||
3686  | 
(* Consequences of independence or spanning for cardinality. *)  | 
|
3687  | 
||
| 33715 | 3688  | 
lemma independent_card_le_dim:  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3689  | 
assumes "(B::(real ^'n) set) \<subseteq> V" and "independent B" shows "card B \<le> dim V"  | 
| 33715 | 3690  | 
proof -  | 
3691  | 
from basis_exists[of V] `B \<subseteq> V`  | 
|
3692  | 
obtain B' where "independent B'" and "B \<subseteq> span B'" and "card B' = dim V" by blast  | 
|
3693  | 
with independent_span_bound[OF _ `independent B` `B \<subseteq> span B'`] independent_bound[of B']  | 
|
3694  | 
show ?thesis by auto  | 
|
3695  | 
qed  | 
|
| 33175 | 3696  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3697  | 
lemma span_card_ge_dim: "(B::(real ^'n) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> finite B \<Longrightarrow> dim V \<le> card B"  | 
| 33715 | 3698  | 
by (metis basis_exists[of V] independent_span_bound subset_trans)  | 
| 33175 | 3699  | 
|
3700  | 
lemma basis_card_eq_dim:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3701  | 
"B \<subseteq> (V:: (real ^'n) set) \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> finite B \<and> card B = dim V"  | 
| 33715 | 3702  | 
by (metis order_eq_iff independent_card_le_dim span_card_ge_dim independent_mono independent_bound)  | 
3703  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3704  | 
lemma dim_unique: "(B::(real ^'n) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> card B = n \<Longrightarrow> dim V = n"  | 
| 33715 | 3705  | 
by (metis basis_card_eq_dim)  | 
| 33175 | 3706  | 
|
3707  | 
(* More lemmas about dimension. *)  | 
|
3708  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3709  | 
lemma dim_univ: "dim (UNIV :: (real^'n) set) = CARD('n)"
 | 
| 33175 | 3710  | 
  apply (rule dim_unique[of "{basis i |i. i\<in> (UNIV :: 'n set)}"])
 | 
| 33715 | 3711  | 
by (auto simp only: span_stdbasis card_stdbasis finite_stdbasis independent_stdbasis)  | 
| 33175 | 3712  | 
|
3713  | 
lemma dim_subset:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3714  | 
"(S:: (real ^'n) set) \<subseteq> T \<Longrightarrow> dim S \<le> dim T"  | 
| 33175 | 3715  | 
using basis_exists[of T] basis_exists[of S]  | 
| 33715 | 3716  | 
by (metis independent_card_le_dim subset_trans)  | 
| 33175 | 3717  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3718  | 
lemma dim_subset_univ: "dim (S:: (real^'n) set) \<le> CARD('n)"
 | 
| 33175 | 3719  | 
by (metis dim_subset subset_UNIV dim_univ)  | 
3720  | 
||
3721  | 
(* Converses to those. *)  | 
|
3722  | 
||
3723  | 
lemma card_ge_dim_independent:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3724  | 
assumes BV:"(B::(real ^'n) set) \<subseteq> V" and iB:"independent B" and dVB:"dim V \<le> card B"  | 
| 33175 | 3725  | 
shows "V \<subseteq> span B"  | 
3726  | 
proof-  | 
|
3727  | 
  {fix a assume aV: "a \<in> V"
 | 
|
3728  | 
    {assume aB: "a \<notin> span B"
 | 
|
3729  | 
then have iaB: "independent (insert a B)" using iB aV BV by (simp add: independent_insert)  | 
|
3730  | 
from aV BV have th0: "insert a B \<subseteq> V" by blast  | 
|
3731  | 
from aB have "a \<notin>B" by (auto simp add: span_superset)  | 
|
| 33715 | 3732  | 
with independent_card_le_dim[OF th0 iaB] dVB independent_bound[OF iB] have False by auto }  | 
| 33175 | 3733  | 
then have "a \<in> span B" by blast}  | 
3734  | 
then show ?thesis by blast  | 
|
3735  | 
qed  | 
|
3736  | 
||
3737  | 
lemma card_le_dim_spanning:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3738  | 
assumes BV: "(B:: (real ^'n) set) \<subseteq> V" and VB: "V \<subseteq> span B"  | 
| 33175 | 3739  | 
and fB: "finite B" and dVB: "dim V \<ge> card B"  | 
3740  | 
shows "independent B"  | 
|
3741  | 
proof-  | 
|
3742  | 
  {fix a assume a: "a \<in> B" "a \<in> span (B -{a})"
 | 
|
3743  | 
from a fB have c0: "card B \<noteq> 0" by auto  | 
|
3744  | 
    from a fB have cb: "card (B -{a}) = card B - 1" by auto
 | 
|
3745  | 
    from BV a have th0: "B -{a} \<subseteq> V" by blast
 | 
|
3746  | 
    {fix x assume x: "x \<in> V"
 | 
|
3747  | 
      from a have eq: "insert a (B -{a}) = B" by blast
 | 
|
3748  | 
from x VB have x': "x \<in> span B" by blast  | 
|
3749  | 
from span_trans[OF a(2), unfolded eq, OF x']  | 
|
3750  | 
      have "x \<in> span (B -{a})" . }
 | 
|
3751  | 
    then have th1: "V \<subseteq> span (B -{a})" by blast
 | 
|
3752  | 
    have th2: "finite (B -{a})" using fB by auto
 | 
|
3753  | 
from span_card_ge_dim[OF th0 th1 th2]  | 
|
3754  | 
    have c: "dim V \<le> card (B -{a})" .
 | 
|
3755  | 
from c c0 dVB cb have False by simp}  | 
|
3756  | 
then show ?thesis unfolding dependent_def by blast  | 
|
3757  | 
qed  | 
|
3758  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3759  | 
lemma card_eq_dim: "(B:: (real ^'n) set) \<subseteq> V \<Longrightarrow> card B = dim V \<Longrightarrow> finite B \<Longrightarrow> independent B \<longleftrightarrow> V \<subseteq> span B"  | 
| 33715 | 3760  | 
by (metis order_eq_iff card_le_dim_spanning  | 
| 33175 | 3761  | 
card_ge_dim_independent)  | 
3762  | 
||
3763  | 
(* ------------------------------------------------------------------------- *)  | 
|
3764  | 
(* More general size bound lemmas. *)  | 
|
3765  | 
(* ------------------------------------------------------------------------- *)  | 
|
3766  | 
||
3767  | 
lemma independent_bound_general:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3768  | 
"independent (S:: (real^'n) set) \<Longrightarrow> finite S \<and> card S \<le> dim S"  | 
| 33175 | 3769  | 
by (metis independent_card_le_dim independent_bound subset_refl)  | 
3770  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3771  | 
lemma dependent_biggerset_general: "(finite (S:: (real^'n) set) \<Longrightarrow> card S > dim S) \<Longrightarrow> dependent S"  | 
| 33175 | 3772  | 
using independent_bound_general[of S] by (metis linorder_not_le)  | 
3773  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3774  | 
lemma dim_span: "dim (span (S:: (real ^'n) set)) = dim S"  | 
| 33175 | 3775  | 
proof-  | 
3776  | 
have th0: "dim S \<le> dim (span S)"  | 
|
3777  | 
by (auto simp add: subset_eq intro: dim_subset span_superset)  | 
|
3778  | 
from basis_exists[of S]  | 
|
| 33715 | 3779  | 
obtain B where B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" by blast  | 
3780  | 
from B have fB: "finite B" "card B = dim S" using independent_bound by blast+  | 
|
| 33175 | 3781  | 
have bSS: "B \<subseteq> span S" using B(1) by (metis subset_eq span_inc)  | 
3782  | 
have sssB: "span S \<subseteq> span B" using span_mono[OF B(3)] by (simp add: span_span)  | 
|
3783  | 
from span_card_ge_dim[OF bSS sssB fB(1)] th0 show ?thesis  | 
|
3784  | 
using fB(2) by arith  | 
|
3785  | 
qed  | 
|
3786  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3787  | 
lemma subset_le_dim: "(S:: (real ^'n) set) \<subseteq> span T \<Longrightarrow> dim S \<le> dim T"  | 
| 33175 | 3788  | 
by (metis dim_span dim_subset)  | 
3789  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3790  | 
lemma span_eq_dim: "span (S:: (real ^'n) set) = span T ==> dim S = dim T"  | 
| 33175 | 3791  | 
by (metis dim_span)  | 
3792  | 
||
3793  | 
lemma spans_image:  | 
|
| 34289 | 3794  | 
assumes lf: "linear (f::'a::semiring_1^_ \<Rightarrow> _)" and VB: "V \<subseteq> span B"  | 
| 33175 | 3795  | 
shows "f ` V \<subseteq> span (f ` B)"  | 
3796  | 
unfolding span_linear_image[OF lf]  | 
|
3797  | 
by (metis VB image_mono)  | 
|
3798  | 
||
3799  | 
lemma dim_image_le:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3800  | 
fixes f :: "real^'n \<Rightarrow> real^'m"  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3801  | 
assumes lf: "linear f" shows "dim (f ` S) \<le> dim (S:: (real ^'n) set)"  | 
| 33175 | 3802  | 
proof-  | 
3803  | 
from basis_exists[of S] obtain B where  | 
|
| 33715 | 3804  | 
B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" by blast  | 
3805  | 
from B have fB: "finite B" "card B = dim S" using independent_bound by blast+  | 
|
| 33175 | 3806  | 
have "dim (f ` S) \<le> card (f ` B)"  | 
3807  | 
apply (rule span_card_ge_dim)  | 
|
3808  | 
using lf B fB by (auto simp add: span_linear_image spans_image subset_image_iff)  | 
|
3809  | 
also have "\<dots> \<le> dim S" using card_image_le[OF fB(1)] fB by simp  | 
|
3810  | 
finally show ?thesis .  | 
|
3811  | 
qed  | 
|
3812  | 
||
3813  | 
(* Relation between bases and injectivity/surjectivity of map. *)  | 
|
3814  | 
||
3815  | 
lemma spanning_surjective_image:  | 
|
| 34289 | 3816  | 
  assumes us: "UNIV \<subseteq> span (S:: ('a::semiring_1 ^_) set)"
 | 
| 33175 | 3817  | 
and lf: "linear f" and sf: "surj f"  | 
3818  | 
shows "UNIV \<subseteq> span (f ` S)"  | 
|
3819  | 
proof-  | 
|
3820  | 
have "UNIV \<subseteq> f ` UNIV" using sf by (auto simp add: surj_def)  | 
|
3821  | 
also have " \<dots> \<subseteq> span (f ` S)" using spans_image[OF lf us] .  | 
|
3822  | 
finally show ?thesis .  | 
|
3823  | 
qed  | 
|
3824  | 
||
3825  | 
lemma independent_injective_image:  | 
|
| 34289 | 3826  | 
  assumes iS: "independent (S::('a::semiring_1^_) set)" and lf: "linear f" and fi: "inj f"
 | 
| 33175 | 3827  | 
shows "independent (f ` S)"  | 
3828  | 
proof-  | 
|
3829  | 
  {fix a assume a: "a \<in> S" "f a \<in> span (f ` S - {f a})"
 | 
|
3830  | 
    have eq: "f ` S - {f a} = f ` (S - {a})" using fi
 | 
|
3831  | 
by (auto simp add: inj_on_def)  | 
|
3832  | 
    from a have "f a \<in> f ` span (S -{a})"
 | 
|
3833  | 
      unfolding eq span_linear_image[OF lf, of "S - {a}"]  by blast
 | 
|
3834  | 
    hence "a \<in> span (S -{a})" using fi by (auto simp add: inj_on_def)
 | 
|
3835  | 
with a(1) iS have False by (simp add: dependent_def) }  | 
|
3836  | 
then show ?thesis unfolding dependent_def by blast  | 
|
3837  | 
qed  | 
|
3838  | 
||
3839  | 
(* ------------------------------------------------------------------------- *)  | 
|
3840  | 
(* Picking an orthogonal replacement for a spanning set. *)  | 
|
3841  | 
(* ------------------------------------------------------------------------- *)  | 
|
3842  | 
(* FIXME : Move to some general theory ?*)  | 
|
3843  | 
definition "pairwise R S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y\<in> S. x\<noteq>y \<longrightarrow> R x y)"  | 
|
3844  | 
||
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34964 
diff
changeset
 | 
3845  | 
lemma vector_sub_project_orthogonal: "(b::'a::linordered_field^'n) \<bullet> (x - ((b \<bullet> x) / (b\<bullet>b)) *s b) = 0"  | 
| 33175 | 3846  | 
apply (cases "b = 0", simp)  | 
3847  | 
apply (simp add: dot_rsub dot_rmult)  | 
|
3848  | 
unfolding times_divide_eq_right[symmetric]  | 
|
3849  | 
by (simp add: field_simps dot_eq_0)  | 
|
3850  | 
||
3851  | 
lemma basis_orthogonal:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3852  | 
fixes B :: "(real ^'n) set"  | 
| 33175 | 3853  | 
assumes fB: "finite B"  | 
3854  | 
shows "\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C"  | 
|
3855  | 
(is " \<exists>C. ?P B C")  | 
|
3856  | 
proof(induct rule: finite_induct[OF fB])  | 
|
3857  | 
  case 1 thus ?case apply (rule exI[where x="{}"]) by (auto simp add: pairwise_def)
 | 
|
3858  | 
next  | 
|
3859  | 
case (2 a B)  | 
|
3860  | 
note fB = `finite B` and aB = `a \<notin> B`  | 
|
3861  | 
from `\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C`  | 
|
3862  | 
obtain C where C: "finite C" "card C \<le> card B"  | 
|
3863  | 
"span C = span B" "pairwise orthogonal C" by blast  | 
|
3864  | 
let ?a = "a - setsum (\<lambda>x. (x\<bullet>a / (x\<bullet>x)) *s x) C"  | 
|
3865  | 
let ?C = "insert ?a C"  | 
|
3866  | 
from C(1) have fC: "finite ?C" by simp  | 
|
3867  | 
from fB aB C(1,2) have cC: "card ?C \<le> card (insert a B)" by (simp add: card_insert_if)  | 
|
3868  | 
  {fix x k
 | 
|
3869  | 
have th0: "\<And>(a::'b::comm_ring) b c. a - (b - c) = c + (a - b)" by (simp add: ring_simps)  | 
|
3870  | 
have "x - k *s (a - (\<Sum>x\<in>C. (x \<bullet> a / (x \<bullet> x)) *s x)) \<in> span C \<longleftrightarrow> x - k *s a \<in> span C"  | 
|
3871  | 
apply (simp only: vector_ssub_ldistrib th0)  | 
|
3872  | 
apply (rule span_add_eq)  | 
|
3873  | 
apply (rule span_mul)  | 
|
3874  | 
apply (rule span_setsum[OF C(1)])  | 
|
3875  | 
apply clarify  | 
|
3876  | 
apply (rule span_mul)  | 
|
3877  | 
by (rule span_superset)}  | 
|
3878  | 
then have SC: "span ?C = span (insert a B)"  | 
|
3879  | 
unfolding expand_set_eq span_breakdown_eq C(3)[symmetric] by auto  | 
|
3880  | 
thm pairwise_def  | 
|
3881  | 
  {fix x y assume xC: "x \<in> ?C" and yC: "y \<in> ?C" and xy: "x \<noteq> y"
 | 
|
3882  | 
    {assume xa: "x = ?a" and ya: "y = ?a"
 | 
|
3883  | 
have "orthogonal x y" using xa ya xy by blast}  | 
|
3884  | 
moreover  | 
|
3885  | 
    {assume xa: "x = ?a" and ya: "y \<noteq> ?a" "y \<in> C"
 | 
|
3886  | 
      from ya have Cy: "C = insert y (C - {y})" by blast
 | 
|
3887  | 
      have fth: "finite (C - {y})" using C by simp
 | 
|
3888  | 
have "orthogonal x y"  | 
|
3889  | 
using xa ya  | 
|
3890  | 
unfolding orthogonal_def xa dot_lsub dot_rsub diff_eq_0_iff_eq  | 
|
3891  | 
apply simp  | 
|
3892  | 
apply (subst Cy)  | 
|
3893  | 
using C(1) fth  | 
|
3894  | 
apply (simp only: setsum_clauses)  | 
|
3895  | 
thm dot_ladd  | 
|
3896  | 
apply (auto simp add: dot_ladd dot_radd dot_lmult dot_rmult dot_eq_0 dot_sym[of y a] dot_lsum[OF fth])  | 
|
3897  | 
apply (rule setsum_0')  | 
|
3898  | 
apply clarsimp  | 
|
3899  | 
apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])  | 
|
3900  | 
by auto}  | 
|
3901  | 
moreover  | 
|
3902  | 
    {assume xa: "x \<noteq> ?a" "x \<in> C" and ya: "y = ?a"
 | 
|
3903  | 
      from xa have Cx: "C = insert x (C - {x})" by blast
 | 
|
3904  | 
      have fth: "finite (C - {x})" using C by simp
 | 
|
3905  | 
have "orthogonal x y"  | 
|
3906  | 
using xa ya  | 
|
3907  | 
unfolding orthogonal_def ya dot_rsub dot_lsub diff_eq_0_iff_eq  | 
|
3908  | 
apply simp  | 
|
3909  | 
apply (subst Cx)  | 
|
3910  | 
using C(1) fth  | 
|
3911  | 
apply (simp only: setsum_clauses)  | 
|
3912  | 
apply (subst dot_sym[of x])  | 
|
3913  | 
apply (auto simp add: dot_radd dot_rmult dot_eq_0 dot_sym[of x a] dot_rsum[OF fth])  | 
|
3914  | 
apply (rule setsum_0')  | 
|
3915  | 
apply clarsimp  | 
|
3916  | 
apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])  | 
|
3917  | 
by auto}  | 
|
3918  | 
moreover  | 
|
3919  | 
    {assume xa: "x \<in> C" and ya: "y \<in> C"
 | 
|
3920  | 
have "orthogonal x y" using xa ya xy C(4) unfolding pairwise_def by blast}  | 
|
3921  | 
ultimately have "orthogonal x y" using xC yC by blast}  | 
|
3922  | 
then have CPO: "pairwise orthogonal ?C" unfolding pairwise_def by blast  | 
|
3923  | 
from fC cC SC CPO have "?P (insert a B) ?C" by blast  | 
|
3924  | 
then show ?case by blast  | 
|
3925  | 
qed  | 
|
3926  | 
||
3927  | 
lemma orthogonal_basis_exists:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3928  | 
fixes V :: "(real ^'n) set"  | 
| 33715 | 3929  | 
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 | 3930  | 
proof-  | 
| 33715 | 3931  | 
from basis_exists[of V] obtain B where B: "B \<subseteq> V" "independent B" "V \<subseteq> span B" "card B = dim V" by blast  | 
3932  | 
from B have fB: "finite B" "card B = dim V" using independent_bound by auto  | 
|
| 33175 | 3933  | 
from basis_orthogonal[OF fB(1)] obtain C where  | 
3934  | 
C: "finite C" "card C \<le> card B" "span C = span B" "pairwise orthogonal C" by blast  | 
|
3935  | 
from C B  | 
|
3936  | 
have CSV: "C \<subseteq> span V" by (metis span_inc span_mono subset_trans)  | 
|
3937  | 
from span_mono[OF B(3)] C have SVC: "span V \<subseteq> span C" by (simp add: span_span)  | 
|
3938  | 
from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB  | 
|
3939  | 
have iC: "independent C" by (simp add: dim_span)  | 
|
3940  | 
from C fB have "card C \<le> dim V" by simp  | 
|
3941  | 
moreover have "dim V \<le> card C" using span_card_ge_dim[OF CSV SVC C(1)]  | 
|
3942  | 
by (simp add: dim_span)  | 
|
| 33715 | 3943  | 
ultimately have CdV: "card C = dim V" using C(1) by simp  | 
| 33175 | 3944  | 
from C B CSV CdV iC show ?thesis by auto  | 
3945  | 
qed  | 
|
3946  | 
||
3947  | 
lemma span_eq: "span S = span T \<longleftrightarrow> S \<subseteq> span T \<and> T \<subseteq> span S"  | 
|
3948  | 
by (metis set_eq_subset span_mono span_span span_inc) (* FIXME: slow *)  | 
|
3949  | 
||
3950  | 
(* ------------------------------------------------------------------------- *)  | 
|
3951  | 
(* Low-dimensional subset is in a hyperplane (weak orthogonal complement). *)  | 
|
3952  | 
(* ------------------------------------------------------------------------- *)  | 
|
3953  | 
||
3954  | 
lemma span_not_univ_orthogonal:  | 
|
3955  | 
assumes sU: "span S \<noteq> UNIV"  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3956  | 
shows "\<exists>(a:: real ^'n). a \<noteq>0 \<and> (\<forall>x \<in> span S. a \<bullet> x = 0)"  | 
| 33175 | 3957  | 
proof-  | 
3958  | 
from sU obtain a where a: "a \<notin> span S" by blast  | 
|
3959  | 
from orthogonal_basis_exists obtain B where  | 
|
| 33715 | 3960  | 
B: "independent B" "B \<subseteq> span S" "S \<subseteq> span B" "card B = dim S" "pairwise orthogonal B"  | 
| 33175 | 3961  | 
by blast  | 
| 33715 | 3962  | 
from B have fB: "finite B" "card B = dim S" using independent_bound by auto  | 
| 33175 | 3963  | 
from span_mono[OF B(2)] span_mono[OF B(3)]  | 
3964  | 
have sSB: "span S = span B" by (simp add: span_span)  | 
|
3965  | 
let ?a = "a - setsum (\<lambda>b. (a\<bullet>b / (b\<bullet>b)) *s b) B"  | 
|
3966  | 
have "setsum (\<lambda>b. (a\<bullet>b / (b\<bullet>b)) *s b) B \<in> span S"  | 
|
3967  | 
unfolding sSB  | 
|
3968  | 
apply (rule span_setsum[OF fB(1)])  | 
|
3969  | 
apply clarsimp  | 
|
3970  | 
apply (rule span_mul)  | 
|
3971  | 
by (rule span_superset)  | 
|
3972  | 
with a have a0:"?a \<noteq> 0" by auto  | 
|
3973  | 
have "\<forall>x\<in>span B. ?a \<bullet> x = 0"  | 
|
3974  | 
proof(rule span_induct')  | 
|
3975  | 
show "subspace (\<lambda>x. ?a \<bullet> x = 0)"  | 
|
3976  | 
by (auto simp add: subspace_def mem_def dot_radd dot_rmult)  | 
|
3977  | 
next  | 
|
3978  | 
    {fix x assume x: "x \<in> B"
 | 
|
3979  | 
      from x have B': "B = insert x (B - {x})" by blast
 | 
|
3980  | 
      have fth: "finite (B - {x})" using fB by simp
 | 
|
3981  | 
have "?a \<bullet> x = 0"  | 
|
3982  | 
apply (subst B') using fB fth  | 
|
3983  | 
unfolding setsum_clauses(2)[OF fth]  | 
|
3984  | 
apply simp  | 
|
3985  | 
apply (clarsimp simp add: dot_lsub dot_ladd dot_lmult dot_lsum dot_eq_0)  | 
|
3986  | 
apply (rule setsum_0', rule ballI)  | 
|
3987  | 
unfolding dot_sym  | 
|
3988  | 
by (auto simp add: x field_simps dot_eq_0 intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])}  | 
|
3989  | 
then show "\<forall>x \<in> B. ?a \<bullet> x = 0" by blast  | 
|
3990  | 
qed  | 
|
3991  | 
with a0 show ?thesis unfolding sSB by (auto intro: exI[where x="?a"])  | 
|
3992  | 
qed  | 
|
3993  | 
||
3994  | 
lemma span_not_univ_subset_hyperplane:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
3995  | 
assumes SU: "span S \<noteq> (UNIV ::(real^'n) set)"  | 
| 33175 | 3996  | 
  shows "\<exists> a. a \<noteq>0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
 | 
3997  | 
using span_not_univ_orthogonal[OF SU] by auto  | 
|
3998  | 
||
3999  | 
lemma lowdim_subset_hyperplane:  | 
|
4000  | 
  assumes d: "dim S < CARD('n::finite)"
 | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4001  | 
  shows "\<exists>(a::real ^'n). a  \<noteq> 0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
 | 
| 33175 | 4002  | 
proof-  | 
4003  | 
  {assume "span S = UNIV"
 | 
|
4004  | 
hence "dim (span S) = dim (UNIV :: (real ^'n) set)" by simp  | 
|
4005  | 
    hence "dim S = CARD('n)" by (simp add: dim_span dim_univ)
 | 
|
4006  | 
with d have False by arith}  | 
|
4007  | 
hence th: "span S \<noteq> UNIV" by blast  | 
|
4008  | 
from span_not_univ_subset_hyperplane[OF th] show ?thesis .  | 
|
4009  | 
qed  | 
|
4010  | 
||
4011  | 
(* We can extend a linear basis-basis injection to the whole set. *)  | 
|
4012  | 
||
4013  | 
lemma linear_indep_image_lemma:  | 
|
4014  | 
assumes lf: "linear f" and fB: "finite B"  | 
|
4015  | 
and ifB: "independent (f ` B)"  | 
|
4016  | 
and fi: "inj_on f B" and xsB: "x \<in> span B"  | 
|
| 34289 | 4017  | 
and fx: "f (x::'a::field^_) = 0"  | 
| 33175 | 4018  | 
shows "x = 0"  | 
4019  | 
using fB ifB fi xsB fx  | 
|
4020  | 
proof(induct arbitrary: x rule: finite_induct[OF fB])  | 
|
4021  | 
case 1 thus ?case by (auto simp add: span_empty)  | 
|
4022  | 
next  | 
|
4023  | 
case (2 a b x)  | 
|
4024  | 
have fb: "finite b" using "2.prems" by simp  | 
|
4025  | 
have th0: "f ` b \<subseteq> f ` (insert a b)"  | 
|
4026  | 
apply (rule image_mono) by blast  | 
|
4027  | 
from independent_mono[ OF "2.prems"(2) th0]  | 
|
4028  | 
have ifb: "independent (f ` b)" .  | 
|
4029  | 
have fib: "inj_on f b"  | 
|
4030  | 
apply (rule subset_inj_on [OF "2.prems"(3)])  | 
|
4031  | 
by blast  | 
|
4032  | 
from span_breakdown[of a "insert a b", simplified, OF "2.prems"(4)]  | 
|
4033  | 
  obtain k where k: "x - k*s a \<in> span (b -{a})" by blast
 | 
|
4034  | 
have "f (x - k*s a) \<in> span (f ` b)"  | 
|
4035  | 
unfolding span_linear_image[OF lf]  | 
|
4036  | 
apply (rule imageI)  | 
|
4037  | 
    using k span_mono[of "b-{a}" b] by blast
 | 
|
4038  | 
hence "f x - k*s f a \<in> span (f ` b)"  | 
|
4039  | 
by (simp add: linear_sub[OF lf] linear_cmul[OF lf])  | 
|
4040  | 
hence th: "-k *s f a \<in> span (f ` b)"  | 
|
4041  | 
using "2.prems"(5) by (simp add: vector_smult_lneg)  | 
|
4042  | 
  {assume k0: "k = 0"
 | 
|
4043  | 
    from k0 k have "x \<in> span (b -{a})" by simp
 | 
|
4044  | 
    then have "x \<in> span b" using span_mono[of "b-{a}" b]
 | 
|
4045  | 
by blast}  | 
|
4046  | 
moreover  | 
|
4047  | 
  {assume k0: "k \<noteq> 0"
 | 
|
4048  | 
from span_mul[OF th, of "- 1/ k"] k0  | 
|
4049  | 
have th1: "f a \<in> span (f ` b)"  | 
|
4050  | 
by (auto simp add: vector_smult_assoc)  | 
|
4051  | 
    from inj_on_image_set_diff[OF "2.prems"(3), of "insert a b " "{a}", symmetric]
 | 
|
4052  | 
    have tha: "f ` insert a b - f ` {a} = f ` (insert a b - {a})" by blast
 | 
|
4053  | 
from "2.prems"(2)[unfolded dependent_def bex_simps(10), rule_format, of "f a"]  | 
|
4054  | 
have "f a \<notin> span (f ` b)" using tha  | 
|
4055  | 
using "2.hyps"(2)  | 
|
4056  | 
"2.prems"(3) by auto  | 
|
4057  | 
with th1 have False by blast  | 
|
4058  | 
then have "x \<in> span b" by blast}  | 
|
4059  | 
ultimately have xsb: "x \<in> span b" by blast  | 
|
4060  | 
from "2.hyps"(3)[OF fb ifb fib xsb "2.prems"(5)]  | 
|
4061  | 
show "x = 0" .  | 
|
4062  | 
qed  | 
|
4063  | 
||
4064  | 
(* We can extend a linear mapping from basis. *)  | 
|
4065  | 
||
4066  | 
lemma linear_independent_extend_lemma:  | 
|
4067  | 
assumes fi: "finite B" and ib: "independent B"  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4068  | 
shows "\<exists>g. (\<forall>x\<in> span B. \<forall>y\<in> span B. g ((x::'a::field^'n) + y) = g x + g y)  | 
| 33175 | 4069  | 
\<and> (\<forall>x\<in> span B. \<forall>c. g (c*s x) = c *s g x)  | 
4070  | 
\<and> (\<forall>x\<in> B. g x = f x)"  | 
|
4071  | 
using ib fi  | 
|
4072  | 
proof(induct rule: finite_induct[OF fi])  | 
|
4073  | 
case 1 thus ?case by (auto simp add: span_empty)  | 
|
4074  | 
next  | 
|
4075  | 
case (2 a b)  | 
|
4076  | 
from "2.prems" "2.hyps" have ibf: "independent b" "finite b"  | 
|
4077  | 
by (simp_all add: independent_insert)  | 
|
4078  | 
from "2.hyps"(3)[OF ibf] obtain g where  | 
|
4079  | 
g: "\<forall>x\<in>span b. \<forall>y\<in>span b. g (x + y) = g x + g y"  | 
|
4080  | 
"\<forall>x\<in>span b. \<forall>c. g (c *s x) = c *s g x" "\<forall>x\<in>b. g x = f x" by blast  | 
|
4081  | 
let ?h = "\<lambda>z. SOME k. (z - k *s a) \<in> span b"  | 
|
4082  | 
  {fix z assume z: "z \<in> span (insert a b)"
 | 
|
4083  | 
have th0: "z - ?h z *s a \<in> span b"  | 
|
4084  | 
apply (rule someI_ex)  | 
|
4085  | 
unfolding span_breakdown_eq[symmetric]  | 
|
4086  | 
using z .  | 
|
4087  | 
    {fix k assume k: "z - k *s a \<in> span b"
 | 
|
4088  | 
have eq: "z - ?h z *s a - (z - k*s a) = (k - ?h z) *s a"  | 
|
4089  | 
by (simp add: ring_simps vector_sadd_rdistrib[symmetric])  | 
|
4090  | 
from span_sub[OF th0 k]  | 
|
4091  | 
have khz: "(k - ?h z) *s a \<in> span b" by (simp add: eq)  | 
|
4092  | 
      {assume "k \<noteq> ?h z" hence k0: "k - ?h z \<noteq> 0" by simp
 | 
|
4093  | 
from k0 span_mul[OF khz, of "1 /(k - ?h z)"]  | 
|
4094  | 
have "a \<in> span b" by (simp add: vector_smult_assoc)  | 
|
4095  | 
with "2.prems"(1) "2.hyps"(2) have False  | 
|
4096  | 
by (auto simp add: dependent_def)}  | 
|
4097  | 
then have "k = ?h z" by blast}  | 
|
4098  | 
with th0 have "z - ?h z *s a \<in> span b \<and> (\<forall>k. z - k *s a \<in> span b \<longrightarrow> k = ?h z)" by blast}  | 
|
4099  | 
note h = this  | 
|
4100  | 
let ?g = "\<lambda>z. ?h z *s f a + g (z - ?h z *s a)"  | 
|
4101  | 
  {fix x y assume x: "x \<in> span (insert a b)" and y: "y \<in> span (insert a b)"
 | 
|
4102  | 
have tha: "\<And>(x::'a^'n) y a k l. (x + y) - (k + l) *s a = (x - k *s a) + (y - l *s a)"  | 
|
4103  | 
by (vector ring_simps)  | 
|
4104  | 
have addh: "?h (x + y) = ?h x + ?h y"  | 
|
4105  | 
apply (rule conjunct2[OF h, rule_format, symmetric])  | 
|
4106  | 
apply (rule span_add[OF x y])  | 
|
4107  | 
unfolding tha  | 
|
4108  | 
by (metis span_add x y conjunct1[OF h, rule_format])  | 
|
4109  | 
have "?g (x + y) = ?g x + ?g y"  | 
|
4110  | 
unfolding addh tha  | 
|
4111  | 
g(1)[rule_format,OF conjunct1[OF h, OF x] conjunct1[OF h, OF y]]  | 
|
4112  | 
by (simp add: vector_sadd_rdistrib)}  | 
|
4113  | 
moreover  | 
|
4114  | 
  {fix x:: "'a^'n" and c:: 'a  assume x: "x \<in> span (insert a b)"
 | 
|
4115  | 
have tha: "\<And>(x::'a^'n) c k a. c *s x - (c * k) *s a = c *s (x - k *s a)"  | 
|
4116  | 
by (vector ring_simps)  | 
|
4117  | 
have hc: "?h (c *s x) = c * ?h x"  | 
|
4118  | 
apply (rule conjunct2[OF h, rule_format, symmetric])  | 
|
4119  | 
apply (metis span_mul x)  | 
|
4120  | 
by (metis tha span_mul x conjunct1[OF h])  | 
|
4121  | 
have "?g (c *s x) = c*s ?g x"  | 
|
4122  | 
unfolding hc tha g(2)[rule_format, OF conjunct1[OF h, OF x]]  | 
|
4123  | 
by (vector ring_simps)}  | 
|
4124  | 
moreover  | 
|
4125  | 
  {fix x assume x: "x \<in> (insert a b)"
 | 
|
4126  | 
    {assume xa: "x = a"
 | 
|
4127  | 
have ha1: "1 = ?h a"  | 
|
4128  | 
apply (rule conjunct2[OF h, rule_format])  | 
|
4129  | 
apply (metis span_superset insertI1)  | 
|
4130  | 
using conjunct1[OF h, OF span_superset, OF insertI1]  | 
|
4131  | 
by (auto simp add: span_0)  | 
|
4132  | 
||
4133  | 
from xa ha1[symmetric] have "?g x = f x"  | 
|
4134  | 
apply simp  | 
|
4135  | 
using g(2)[rule_format, OF span_0, of 0]  | 
|
4136  | 
by simp}  | 
|
4137  | 
moreover  | 
|
4138  | 
    {assume xb: "x \<in> b"
 | 
|
4139  | 
have h0: "0 = ?h x"  | 
|
4140  | 
apply (rule conjunct2[OF h, rule_format])  | 
|
4141  | 
apply (metis span_superset insertI1 xb x)  | 
|
4142  | 
apply simp  | 
|
4143  | 
apply (metis span_superset xb)  | 
|
4144  | 
done  | 
|
4145  | 
have "?g x = f x"  | 
|
4146  | 
by (simp add: h0[symmetric] g(3)[rule_format, OF xb])}  | 
|
4147  | 
ultimately have "?g x = f x" using x by blast }  | 
|
4148  | 
ultimately show ?case apply - apply (rule exI[where x="?g"]) by blast  | 
|
4149  | 
qed  | 
|
4150  | 
||
4151  | 
lemma linear_independent_extend:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4152  | 
assumes iB: "independent (B:: (real ^'n) set)"  | 
| 33175 | 4153  | 
shows "\<exists>g. linear g \<and> (\<forall>x\<in>B. g x = f x)"  | 
4154  | 
proof-  | 
|
4155  | 
from maximal_independent_subset_extend[of B UNIV] iB  | 
|
4156  | 
obtain C where C: "B \<subseteq> C" "independent C" "\<And>x. x \<in> span C" by auto  | 
|
4157  | 
||
4158  | 
from C(2) independent_bound[of C] linear_independent_extend_lemma[of C f]  | 
|
4159  | 
obtain g where g: "(\<forall>x\<in> span C. \<forall>y\<in> span C. g (x + y) = g x + g y)  | 
|
4160  | 
\<and> (\<forall>x\<in> span C. \<forall>c. g (c*s x) = c *s g x)  | 
|
4161  | 
\<and> (\<forall>x\<in> C. g x = f x)" by blast  | 
|
4162  | 
from g show ?thesis unfolding linear_def using C  | 
|
4163  | 
apply clarsimp by blast  | 
|
4164  | 
qed  | 
|
4165  | 
||
4166  | 
(* Can construct an isomorphism between spaces of same dimension. *)  | 
|
4167  | 
||
4168  | 
lemma card_le_inj: assumes fA: "finite A" and fB: "finite B"  | 
|
4169  | 
and c: "card A \<le> card B" shows "(\<exists>f. f ` A \<subseteq> B \<and> inj_on f A)"  | 
|
4170  | 
using fB c  | 
|
4171  | 
proof(induct arbitrary: B rule: finite_induct[OF fA])  | 
|
4172  | 
case 1 thus ?case by simp  | 
|
4173  | 
next  | 
|
4174  | 
case (2 x s t)  | 
|
4175  | 
thus ?case  | 
|
4176  | 
proof(induct rule: finite_induct[OF "2.prems"(1)])  | 
|
4177  | 
case 1 then show ?case by simp  | 
|
4178  | 
next  | 
|
4179  | 
case (2 y t)  | 
|
4180  | 
from "2.prems"(1,2,5) "2.hyps"(1,2) have cst:"card s \<le> card t" by simp  | 
|
4181  | 
from "2.prems"(3) [OF "2.hyps"(1) cst] obtain f where  | 
|
4182  | 
f: "f ` s \<subseteq> t \<and> inj_on f s" by blast  | 
|
4183  | 
from f "2.prems"(2) "2.hyps"(2) show ?case  | 
|
4184  | 
apply -  | 
|
4185  | 
apply (rule exI[where x = "\<lambda>z. if z = x then y else f z"])  | 
|
4186  | 
by (auto simp add: inj_on_def)  | 
|
4187  | 
qed  | 
|
4188  | 
qed  | 
|
4189  | 
||
4190  | 
lemma card_subset_eq: assumes fB: "finite B" and AB: "A \<subseteq> B" and  | 
|
4191  | 
c: "card A = card B"  | 
|
4192  | 
shows "A = B"  | 
|
4193  | 
proof-  | 
|
4194  | 
from fB AB have fA: "finite A" by (auto intro: finite_subset)  | 
|
4195  | 
from fA fB have fBA: "finite (B - A)" by auto  | 
|
4196  | 
  have e: "A \<inter> (B - A) = {}" by blast
 | 
|
4197  | 
have eq: "A \<union> (B - A) = B" using AB by blast  | 
|
4198  | 
from card_Un_disjoint[OF fA fBA e, unfolded eq c]  | 
|
4199  | 
have "card (B - A) = 0" by arith  | 
|
4200  | 
  hence "B - A = {}" unfolding card_eq_0_iff using fA fB by simp
 | 
|
4201  | 
with AB show "A = B" by blast  | 
|
4202  | 
qed  | 
|
4203  | 
||
4204  | 
lemma subspace_isomorphism:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4205  | 
assumes s: "subspace (S:: (real ^'n) set)"  | 
| 
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4206  | 
and t: "subspace (T :: (real ^'m) set)"  | 
| 33175 | 4207  | 
and d: "dim S = dim T"  | 
4208  | 
shows "\<exists>f. linear f \<and> f ` S = T \<and> inj_on f S"  | 
|
4209  | 
proof-  | 
|
| 33715 | 4210  | 
from basis_exists[of S] independent_bound obtain B where  | 
4211  | 
B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" and fB: "finite B" by blast  | 
|
4212  | 
from basis_exists[of T] independent_bound obtain C where  | 
|
4213  | 
C: "C \<subseteq> T" "independent C" "T \<subseteq> span C" "card C = dim T" and fC: "finite C" by blast  | 
|
| 33175 | 4214  | 
from B(4) C(4) card_le_inj[of B C] d obtain f where  | 
| 33715 | 4215  | 
f: "f ` B \<subseteq> C" "inj_on f B" using `finite B` `finite C` by auto  | 
| 33175 | 4216  | 
from linear_independent_extend[OF B(2)] obtain g where  | 
4217  | 
g: "linear g" "\<forall>x\<in> B. g x = f x" by blast  | 
|
4218  | 
from inj_on_iff_eq_card[OF fB, of f] f(2)  | 
|
4219  | 
have "card (f ` B) = card B" by simp  | 
|
4220  | 
with B(4) C(4) have ceq: "card (f ` B) = card C" using d  | 
|
| 33715 | 4221  | 
by simp  | 
| 33175 | 4222  | 
have "g ` B = f ` B" using g(2)  | 
4223  | 
by (auto simp add: image_iff)  | 
|
4224  | 
also have "\<dots> = C" using card_subset_eq[OF fC f(1) ceq] .  | 
|
4225  | 
finally have gBC: "g ` B = C" .  | 
|
4226  | 
have gi: "inj_on g B" using f(2) g(2)  | 
|
4227  | 
by (auto simp add: inj_on_def)  | 
|
4228  | 
note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi]  | 
|
4229  | 
  {fix x y assume x: "x \<in> S" and y: "y \<in> S" and gxy:"g x = g y"
 | 
|
4230  | 
from B(3) x y have x': "x \<in> span B" and y': "y \<in> span B" by blast+  | 
|
4231  | 
from gxy have th0: "g (x - y) = 0" by (simp add: linear_sub[OF g(1)])  | 
|
4232  | 
have th1: "x - y \<in> span B" using x' y' by (metis span_sub)  | 
|
4233  | 
have "x=y" using g0[OF th1 th0] by simp }  | 
|
4234  | 
then have giS: "inj_on g S"  | 
|
4235  | 
unfolding inj_on_def by blast  | 
|
4236  | 
from span_subspace[OF B(1,3) s]  | 
|
4237  | 
have "g ` S = span (g ` B)" by (simp add: span_linear_image[OF g(1)])  | 
|
4238  | 
also have "\<dots> = span C" unfolding gBC ..  | 
|
4239  | 
also have "\<dots> = T" using span_subspace[OF C(1,3) t] .  | 
|
4240  | 
finally have gS: "g ` S = T" .  | 
|
4241  | 
from g(1) gS giS show ?thesis by blast  | 
|
4242  | 
qed  | 
|
4243  | 
||
4244  | 
(* linear functions are equal on a subspace if they are on a spanning set. *)  | 
|
4245  | 
||
4246  | 
lemma subspace_kernel:  | 
|
| 34289 | 4247  | 
assumes lf: "linear (f::'a::semiring_1 ^_ \<Rightarrow> _)"  | 
| 33175 | 4248  | 
  shows "subspace {x. f x = 0}"
 | 
4249  | 
apply (simp add: subspace_def)  | 
|
4250  | 
by (simp add: linear_add[OF lf] linear_cmul[OF lf] linear_0[OF lf])  | 
|
4251  | 
||
4252  | 
lemma linear_eq_0_span:  | 
|
4253  | 
assumes lf: "linear f" and f0: "\<forall>x\<in>B. f x = 0"  | 
|
| 34289 | 4254  | 
shows "\<forall>x \<in> span B. f x = (0::'a::semiring_1 ^_)"  | 
| 33175 | 4255  | 
proof  | 
4256  | 
fix x assume x: "x \<in> span B"  | 
|
4257  | 
let ?P = "\<lambda>x. f x = 0"  | 
|
4258  | 
from subspace_kernel[OF lf] have "subspace ?P" unfolding Collect_def .  | 
|
4259  | 
with x f0 span_induct[of B "?P" x] show "f x = 0" by blast  | 
|
4260  | 
qed  | 
|
4261  | 
||
4262  | 
lemma linear_eq_0:  | 
|
4263  | 
assumes lf: "linear f" and SB: "S \<subseteq> span B" and f0: "\<forall>x\<in>B. f x = 0"  | 
|
| 34289 | 4264  | 
shows "\<forall>x \<in> S. f x = (0::'a::semiring_1^_)"  | 
| 33175 | 4265  | 
by (metis linear_eq_0_span[OF lf] subset_eq SB f0)  | 
4266  | 
||
4267  | 
lemma linear_eq:  | 
|
| 34289 | 4268  | 
assumes lf: "linear (f::'a::ring_1^_ \<Rightarrow> _)" and lg: "linear g" and S: "S \<subseteq> span B"  | 
| 33175 | 4269  | 
and fg: "\<forall> x\<in> B. f x = g x"  | 
4270  | 
shows "\<forall>x\<in> S. f x = g x"  | 
|
4271  | 
proof-  | 
|
4272  | 
let ?h = "\<lambda>x. f x - g x"  | 
|
4273  | 
from fg have fg': "\<forall>x\<in> B. ?h x = 0" by simp  | 
|
4274  | 
from linear_eq_0[OF linear_compose_sub[OF lf lg] S fg']  | 
|
4275  | 
show ?thesis by simp  | 
|
4276  | 
qed  | 
|
4277  | 
||
4278  | 
lemma linear_eq_stdbasis:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4279  | 
assumes lf: "linear (f::'a::ring_1^'m \<Rightarrow> 'a^'n)" and lg: "linear g"  | 
| 33175 | 4280  | 
and fg: "\<forall>i. f (basis i) = g(basis i)"  | 
4281  | 
shows "f = g"  | 
|
4282  | 
proof-  | 
|
4283  | 
let ?U = "UNIV :: 'm set"  | 
|
4284  | 
  let ?I = "{basis i:: 'a^'m|i. i \<in> ?U}"
 | 
|
4285  | 
  {fix x assume x: "x \<in> (UNIV :: ('a^'m) set)"
 | 
|
4286  | 
from equalityD2[OF span_stdbasis]  | 
|
4287  | 
    have IU: " (UNIV :: ('a^'m) set) \<subseteq> span ?I" by blast
 | 
|
4288  | 
from linear_eq[OF lf lg IU] fg x  | 
|
4289  | 
have "f x = g x" unfolding Collect_def Ball_def mem_def by metis}  | 
|
4290  | 
then show ?thesis by (auto intro: ext)  | 
|
4291  | 
qed  | 
|
4292  | 
||
4293  | 
(* Similar results for bilinear functions. *)  | 
|
4294  | 
||
4295  | 
lemma bilinear_eq:  | 
|
| 34289 | 4296  | 
assumes bf: "bilinear (f:: 'a::ring^_ \<Rightarrow> 'a^_ \<Rightarrow> 'a^_)"  | 
| 33175 | 4297  | 
and bg: "bilinear g"  | 
4298  | 
and SB: "S \<subseteq> span B" and TC: "T \<subseteq> span C"  | 
|
4299  | 
and fg: "\<forall>x\<in> B. \<forall>y\<in> C. f x y = g x y"  | 
|
4300  | 
shows "\<forall>x\<in>S. \<forall>y\<in>T. f x y = g x y "  | 
|
4301  | 
proof-  | 
|
4302  | 
let ?P = "\<lambda>x. \<forall>y\<in> span C. f x y = g x y"  | 
|
4303  | 
from bf bg have sp: "subspace ?P"  | 
|
4304  | 
unfolding bilinear_def linear_def subspace_def bf bg  | 
|
4305  | 
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])  | 
|
4306  | 
||
4307  | 
have "\<forall>x \<in> span B. \<forall>y\<in> span C. f x y = g x y"  | 
|
4308  | 
apply -  | 
|
4309  | 
apply (rule ballI)  | 
|
4310  | 
apply (rule span_induct[of B ?P])  | 
|
4311  | 
defer  | 
|
4312  | 
apply (rule sp)  | 
|
4313  | 
apply assumption  | 
|
4314  | 
apply (clarsimp simp add: Ball_def)  | 
|
4315  | 
apply (rule_tac P="\<lambda>y. f xa y = g xa y" and S=C in span_induct)  | 
|
4316  | 
using fg  | 
|
4317  | 
apply (auto simp add: subspace_def)  | 
|
4318  | 
using bf bg unfolding bilinear_def linear_def  | 
|
4319  | 
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])  | 
|
4320  | 
then show ?thesis using SB TC by (auto intro: ext)  | 
|
4321  | 
qed  | 
|
4322  | 
||
4323  | 
lemma bilinear_eq_stdbasis:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4324  | 
assumes bf: "bilinear (f:: 'a::ring_1^'m \<Rightarrow> 'a^'n \<Rightarrow> 'a^_)"  | 
| 33175 | 4325  | 
and bg: "bilinear g"  | 
4326  | 
and fg: "\<forall>i j. f (basis i) (basis j) = g (basis i) (basis j)"  | 
|
4327  | 
shows "f = g"  | 
|
4328  | 
proof-  | 
|
4329  | 
  from fg have th: "\<forall>x \<in> {basis i| i. i\<in> (UNIV :: 'm set)}. \<forall>y\<in>  {basis j |j. j \<in> (UNIV :: 'n set)}. f x y = g x y" by blast
 | 
|
4330  | 
from bilinear_eq[OF bf bg equalityD2[OF span_stdbasis] equalityD2[OF span_stdbasis] th] show ?thesis by (blast intro: ext)  | 
|
4331  | 
qed  | 
|
4332  | 
||
4333  | 
(* Detailed theorems about left and right invertibility in general case. *)  | 
|
4334  | 
||
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4335  | 
lemma left_invertible_transpose:  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4336  | 
"(\<exists>(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). A ** B = mat 1)"  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4337  | 
by (metis matrix_transpose_mul transpose_mat transpose_transpose)  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4338  | 
|
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4339  | 
lemma right_invertible_transpose:  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4340  | 
"(\<exists>(B). transpose (A) ** B = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). B ** A = mat 1)"  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4341  | 
by (metis matrix_transpose_mul transpose_mat transpose_transpose)  | 
| 33175 | 4342  | 
|
4343  | 
lemma linear_injective_left_inverse:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4344  | 
assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" and fi: "inj f"  | 
| 33175 | 4345  | 
shows "\<exists>g. linear g \<and> g o f = id"  | 
4346  | 
proof-  | 
|
4347  | 
from linear_independent_extend[OF independent_injective_image, OF independent_stdbasis, OF lf fi]  | 
|
4348  | 
  obtain h:: "real ^'m \<Rightarrow> real ^'n" where h: "linear h" " \<forall>x \<in> f ` {basis i|i. i \<in> (UNIV::'n set)}. h x = inv f x" by blast
 | 
|
4349  | 
from h(2)  | 
|
4350  | 
have th: "\<forall>i. (h \<circ> f) (basis i) = id (basis i)"  | 
|
4351  | 
using inv_o_cancel[OF fi, unfolded stupid_ext[symmetric] id_def o_def]  | 
|
4352  | 
by auto  | 
|
4353  | 
||
4354  | 
from linear_eq_stdbasis[OF linear_compose[OF lf h(1)] linear_id th]  | 
|
4355  | 
have "h o f = id" .  | 
|
4356  | 
then show ?thesis using h(1) by blast  | 
|
4357  | 
qed  | 
|
4358  | 
||
4359  | 
lemma linear_surjective_right_inverse:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4360  | 
assumes lf: "linear (f:: real ^'m \<Rightarrow> real ^'n)" and sf: "surj f"  | 
| 33175 | 4361  | 
shows "\<exists>g. linear g \<and> f o g = id"  | 
4362  | 
proof-  | 
|
4363  | 
from linear_independent_extend[OF independent_stdbasis]  | 
|
4364  | 
obtain h:: "real ^'n \<Rightarrow> real ^'m" where  | 
|
4365  | 
    h: "linear h" "\<forall> x\<in> {basis i| i. i\<in> (UNIV :: 'n set)}. h x = inv f x" by blast
 | 
|
4366  | 
from h(2)  | 
|
4367  | 
have th: "\<forall>i. (f o h) (basis i) = id (basis i)"  | 
|
4368  | 
using sf  | 
|
4369  | 
apply (auto simp add: surj_iff o_def stupid_ext[symmetric])  | 
|
4370  | 
apply (erule_tac x="basis i" in allE)  | 
|
4371  | 
by auto  | 
|
4372  | 
||
4373  | 
from linear_eq_stdbasis[OF linear_compose[OF h(1) lf] linear_id th]  | 
|
4374  | 
have "f o h = id" .  | 
|
4375  | 
then show ?thesis using h(1) by blast  | 
|
4376  | 
qed  | 
|
4377  | 
||
4378  | 
lemma matrix_left_invertible_injective:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4379  | 
"(\<exists>B. (B::real^'m^'n) ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x y. A *v x = A *v y \<longrightarrow> x = y)"  | 
| 33175 | 4380  | 
proof-  | 
4381  | 
  {fix B:: "real^'m^'n" and x y assume B: "B ** A = mat 1" and xy: "A *v x = A*v y"
 | 
|
4382  | 
from xy have "B*v (A *v x) = B *v (A*v y)" by simp  | 
|
4383  | 
hence "x = y"  | 
|
4384  | 
unfolding matrix_vector_mul_assoc B matrix_vector_mul_lid .}  | 
|
4385  | 
moreover  | 
|
4386  | 
  {assume A: "\<forall>x y. A *v x = A *v y \<longrightarrow> x = y"
 | 
|
4387  | 
hence i: "inj (op *v A)" unfolding inj_on_def by auto  | 
|
4388  | 
from linear_injective_left_inverse[OF matrix_vector_mul_linear i]  | 
|
4389  | 
obtain g where g: "linear g" "g o op *v A = id" by blast  | 
|
4390  | 
have "matrix g ** A = mat 1"  | 
|
4391  | 
unfolding matrix_eq matrix_vector_mul_lid matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]  | 
|
4392  | 
using g(2) by (simp add: o_def id_def stupid_ext)  | 
|
4393  | 
then have "\<exists>B. (B::real ^'m^'n) ** A = mat 1" by blast}  | 
|
4394  | 
ultimately show ?thesis by blast  | 
|
4395  | 
qed  | 
|
4396  | 
||
4397  | 
lemma matrix_left_invertible_ker:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4398  | 
"(\<exists>B. (B::real ^'m^'n) ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x. A *v x = 0 \<longrightarrow> x = 0)"  | 
| 33175 | 4399  | 
unfolding matrix_left_invertible_injective  | 
4400  | 
using linear_injective_0[OF matrix_vector_mul_linear, of A]  | 
|
4401  | 
by (simp add: inj_on_def)  | 
|
4402  | 
||
4403  | 
lemma matrix_right_invertible_surjective:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4404  | 
"(\<exists>B. (A::real^'n^'m) ** (B::real^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)"  | 
| 33175 | 4405  | 
proof-  | 
4406  | 
  {fix B :: "real ^'m^'n"  assume AB: "A ** B = mat 1"
 | 
|
4407  | 
    {fix x :: "real ^ 'm"
 | 
|
4408  | 
have "A *v (B *v x) = x"  | 
|
4409  | 
by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB)}  | 
|
4410  | 
hence "surj (op *v A)" unfolding surj_def by metis }  | 
|
4411  | 
moreover  | 
|
4412  | 
  {assume sf: "surj (op *v A)"
 | 
|
4413  | 
from linear_surjective_right_inverse[OF matrix_vector_mul_linear sf]  | 
|
4414  | 
obtain g:: "real ^'m \<Rightarrow> real ^'n" where g: "linear g" "op *v A o g = id"  | 
|
4415  | 
by blast  | 
|
4416  | 
||
4417  | 
have "A ** (matrix g) = mat 1"  | 
|
4418  | 
unfolding matrix_eq matrix_vector_mul_lid  | 
|
4419  | 
matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]  | 
|
4420  | 
using g(2) unfolding o_def stupid_ext[symmetric] id_def  | 
|
4421  | 
.  | 
|
4422  | 
hence "\<exists>B. A ** (B::real^'m^'n) = mat 1" by blast  | 
|
4423  | 
}  | 
|
4424  | 
ultimately show ?thesis unfolding surj_def by blast  | 
|
4425  | 
qed  | 
|
4426  | 
||
4427  | 
lemma matrix_left_invertible_independent_columns:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4428  | 
fixes A :: "real^'n^'m"  | 
| 33175 | 4429  | 
shows "(\<exists>(B::real ^'m^'n). B ** A = mat 1) \<longleftrightarrow> (\<forall>c. setsum (\<lambda>i. c i *s column i A) (UNIV :: 'n set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"  | 
4430  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
4431  | 
proof-  | 
|
4432  | 
let ?U = "UNIV :: 'n set"  | 
|
4433  | 
  {assume k: "\<forall>x. A *v x = 0 \<longrightarrow> x = 0"
 | 
|
4434  | 
    {fix c i assume c: "setsum (\<lambda>i. c i *s column i A) ?U = 0"
 | 
|
4435  | 
and i: "i \<in> ?U"  | 
|
4436  | 
let ?x = "\<chi> i. c i"  | 
|
4437  | 
have th0:"A *v ?x = 0"  | 
|
4438  | 
using c  | 
|
4439  | 
unfolding matrix_mult_vsum Cart_eq  | 
|
4440  | 
by auto  | 
|
4441  | 
from k[rule_format, OF th0] i  | 
|
4442  | 
have "c i = 0" by (vector Cart_eq)}  | 
|
4443  | 
hence ?rhs by blast}  | 
|
4444  | 
moreover  | 
|
4445  | 
  {assume H: ?rhs
 | 
|
4446  | 
    {fix x assume x: "A *v x = 0"
 | 
|
4447  | 
let ?c = "\<lambda>i. ((x$i ):: real)"  | 
|
4448  | 
from H[rule_format, of ?c, unfolded matrix_mult_vsum[symmetric], OF x]  | 
|
4449  | 
have "x = 0" by vector}}  | 
|
4450  | 
ultimately show ?thesis unfolding matrix_left_invertible_ker by blast  | 
|
4451  | 
qed  | 
|
4452  | 
||
4453  | 
lemma matrix_right_invertible_independent_rows:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4454  | 
fixes A :: "real^'n^'m"  | 
| 33175 | 4455  | 
shows "(\<exists>(B::real^'m^'n). A ** B = mat 1) \<longleftrightarrow> (\<forall>c. setsum (\<lambda>i. c i *s row i A) (UNIV :: 'm set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"  | 
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4456  | 
unfolding left_invertible_transpose[symmetric]  | 
| 33175 | 4457  | 
matrix_left_invertible_independent_columns  | 
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4458  | 
by (simp add: column_transpose)  | 
| 33175 | 4459  | 
|
4460  | 
lemma matrix_right_invertible_span_columns:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4461  | 
"(\<exists>(B::real ^'n^'m). (A::real ^'m^'n) ** B = mat 1) \<longleftrightarrow> span (columns A) = UNIV" (is "?lhs = ?rhs")  | 
| 33175 | 4462  | 
proof-  | 
4463  | 
let ?U = "UNIV :: 'm set"  | 
|
4464  | 
have fU: "finite ?U" by simp  | 
|
4465  | 
have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y)"  | 
|
4466  | 
unfolding matrix_right_invertible_surjective matrix_mult_vsum surj_def  | 
|
4467  | 
apply (subst eq_commute) ..  | 
|
4468  | 
have rhseq: "?rhs \<longleftrightarrow> (\<forall>x. x \<in> span (columns A))" by blast  | 
|
4469  | 
  {assume h: ?lhs
 | 
|
4470  | 
    {fix x:: "real ^'n"
 | 
|
4471  | 
from h[unfolded lhseq, rule_format, of x] obtain y:: "real ^'m"  | 
|
4472  | 
where y: "setsum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast  | 
|
4473  | 
have "x \<in> span (columns A)"  | 
|
4474  | 
unfolding y[symmetric]  | 
|
4475  | 
apply (rule span_setsum[OF fU])  | 
|
4476  | 
apply clarify  | 
|
4477  | 
apply (rule span_mul)  | 
|
4478  | 
apply (rule span_superset)  | 
|
4479  | 
unfolding columns_def  | 
|
4480  | 
by blast}  | 
|
4481  | 
then have ?rhs unfolding rhseq by blast}  | 
|
4482  | 
moreover  | 
|
4483  | 
  {assume h:?rhs
 | 
|
4484  | 
let ?P = "\<lambda>(y::real ^'n). \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y"  | 
|
4485  | 
    {fix y have "?P y"
 | 
|
4486  | 
proof(rule span_induct_alt[of ?P "columns A"])  | 
|
4487  | 
show "\<exists>x\<Colon>real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0"  | 
|
4488  | 
apply (rule exI[where x=0])  | 
|
4489  | 
by (simp add: zero_index vector_smult_lzero)  | 
|
4490  | 
next  | 
|
4491  | 
fix c y1 y2 assume y1: "y1 \<in> columns A" and y2: "?P y2"  | 
|
4492  | 
from y1 obtain i where i: "i \<in> ?U" "y1 = column i A"  | 
|
4493  | 
unfolding columns_def by blast  | 
|
4494  | 
from y2 obtain x:: "real ^'m" where  | 
|
4495  | 
x: "setsum (\<lambda>i. (x$i) *s column i A) ?U = y2" by blast  | 
|
4496  | 
let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::real^'m"  | 
|
4497  | 
show "?P (c*s y1 + y2)"  | 
|
4498  | 
proof(rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] cond_value_iff right_distrib cond_application_beta cong del: if_weak_cong)  | 
|
4499  | 
fix j  | 
|
4500  | 
have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)  | 
|
4501  | 
else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" using i(1)  | 
|
4502  | 
by (simp add: ring_simps)  | 
|
4503  | 
have "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)  | 
|
4504  | 
else (x$xa) * ((column xa A$j))) ?U = setsum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"  | 
|
4505  | 
apply (rule setsum_cong[OF refl])  | 
|
4506  | 
using th by blast  | 
|
4507  | 
also have "\<dots> = setsum (\<lambda>xa. if xa = i then c * ((column i A)$j) else 0) ?U + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"  | 
|
4508  | 
by (simp add: setsum_addf)  | 
|
4509  | 
also have "\<dots> = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"  | 
|
4510  | 
unfolding setsum_delta[OF fU]  | 
|
4511  | 
using i(1) by simp  | 
|
4512  | 
finally show "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)  | 
|
4513  | 
else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" .  | 
|
4514  | 
qed  | 
|
4515  | 
next  | 
|
4516  | 
show "y \<in> span (columns A)" unfolding h by blast  | 
|
4517  | 
qed}  | 
|
4518  | 
then have ?lhs unfolding lhseq ..}  | 
|
4519  | 
ultimately show ?thesis by blast  | 
|
4520  | 
qed  | 
|
4521  | 
||
4522  | 
lemma matrix_left_invertible_span_rows:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4523  | 
"(\<exists>(B::real^'m^'n). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV"  | 
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4524  | 
unfolding right_invertible_transpose[symmetric]  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4525  | 
unfolding columns_transpose[symmetric]  | 
| 33175 | 4526  | 
unfolding matrix_right_invertible_span_columns  | 
4527  | 
..  | 
|
4528  | 
||
4529  | 
(* An injective map real^'n->real^'n is also surjective. *)  | 
|
4530  | 
||
4531  | 
lemma linear_injective_imp_surjective:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4532  | 
assumes lf: "linear (f:: real ^'n \<Rightarrow> real ^'n)" and fi: "inj f"  | 
| 33175 | 4533  | 
shows "surj f"  | 
4534  | 
proof-  | 
|
4535  | 
let ?U = "UNIV :: (real ^'n) set"  | 
|
4536  | 
from basis_exists[of ?U] obtain B  | 
|
| 33715 | 4537  | 
where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" "card B = dim ?U"  | 
| 33175 | 4538  | 
by blast  | 
| 33715 | 4539  | 
from B(4) have d: "dim ?U = card B" by simp  | 
| 33175 | 4540  | 
have th: "?U \<subseteq> span (f ` B)"  | 
4541  | 
apply (rule card_ge_dim_independent)  | 
|
4542  | 
apply blast  | 
|
4543  | 
apply (rule independent_injective_image[OF B(2) lf fi])  | 
|
4544  | 
apply (rule order_eq_refl)  | 
|
4545  | 
apply (rule sym)  | 
|
4546  | 
unfolding d  | 
|
4547  | 
apply (rule card_image)  | 
|
4548  | 
apply (rule subset_inj_on[OF fi])  | 
|
4549  | 
by blast  | 
|
4550  | 
from th show ?thesis  | 
|
4551  | 
unfolding span_linear_image[OF lf] surj_def  | 
|
4552  | 
using B(3) by blast  | 
|
4553  | 
qed  | 
|
4554  | 
||
4555  | 
(* And vice versa. *)  | 
|
4556  | 
||
4557  | 
lemma surjective_iff_injective_gen:  | 
|
4558  | 
assumes fS: "finite S" and fT: "finite T" and c: "card S = card T"  | 
|
4559  | 
and ST: "f ` S \<subseteq> T"  | 
|
4560  | 
shows "(\<forall>y \<in> T. \<exists>x \<in> S. f x = y) \<longleftrightarrow> inj_on f S" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
4561  | 
proof-  | 
|
4562  | 
  {assume h: "?lhs"
 | 
|
4563  | 
    {fix x y assume x: "x \<in> S" and y: "y \<in> S" and f: "f x = f y"
 | 
|
4564  | 
from x fS have S0: "card S \<noteq> 0" by auto  | 
|
4565  | 
      {assume xy: "x \<noteq> y"
 | 
|
4566  | 
        have th: "card S \<le> card (f ` (S - {y}))"
 | 
|
4567  | 
unfolding c  | 
|
4568  | 
apply (rule card_mono)  | 
|
4569  | 
apply (rule finite_imageI)  | 
|
4570  | 
using fS apply simp  | 
|
4571  | 
using h xy x y f unfolding subset_eq image_iff  | 
|
4572  | 
apply auto  | 
|
4573  | 
apply (case_tac "xa = f x")  | 
|
4574  | 
apply (rule bexI[where x=x])  | 
|
4575  | 
apply auto  | 
|
4576  | 
done  | 
|
4577  | 
        also have " \<dots> \<le> card (S -{y})"
 | 
|
4578  | 
apply (rule card_image_le)  | 
|
4579  | 
using fS by simp  | 
|
4580  | 
also have "\<dots> \<le> card S - 1" using y fS by simp  | 
|
4581  | 
finally have False using S0 by arith }  | 
|
4582  | 
then have "x = y" by blast}  | 
|
4583  | 
then have ?rhs unfolding inj_on_def by blast}  | 
|
4584  | 
moreover  | 
|
4585  | 
  {assume h: ?rhs
 | 
|
4586  | 
have "f ` S = T"  | 
|
4587  | 
apply (rule card_subset_eq[OF fT ST])  | 
|
4588  | 
unfolding card_image[OF h] using c .  | 
|
4589  | 
then have ?lhs by blast}  | 
|
4590  | 
ultimately show ?thesis by blast  | 
|
4591  | 
qed  | 
|
4592  | 
||
4593  | 
lemma linear_surjective_imp_injective:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4594  | 
assumes lf: "linear (f::real ^'n => real ^'n)" and sf: "surj f"  | 
| 33175 | 4595  | 
shows "inj f"  | 
4596  | 
proof-  | 
|
4597  | 
let ?U = "UNIV :: (real ^'n) set"  | 
|
4598  | 
from basis_exists[of ?U] obtain B  | 
|
| 33715 | 4599  | 
where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" and d: "card B = dim ?U"  | 
| 33175 | 4600  | 
by blast  | 
4601  | 
  {fix x assume x: "x \<in> span B" and fx: "f x = 0"
 | 
|
| 33715 | 4602  | 
from B(2) have fB: "finite B" using independent_bound by auto  | 
| 33175 | 4603  | 
have fBi: "independent (f ` B)"  | 
4604  | 
apply (rule card_le_dim_spanning[of "f ` B" ?U])  | 
|
4605  | 
apply blast  | 
|
4606  | 
using sf B(3)  | 
|
4607  | 
unfolding span_linear_image[OF lf] surj_def subset_eq image_iff  | 
|
4608  | 
apply blast  | 
|
4609  | 
using fB apply (blast intro: finite_imageI)  | 
|
| 33715 | 4610  | 
unfolding d[symmetric]  | 
| 33175 | 4611  | 
apply (rule card_image_le)  | 
4612  | 
apply (rule fB)  | 
|
4613  | 
done  | 
|
4614  | 
have th0: "dim ?U \<le> card (f ` B)"  | 
|
4615  | 
apply (rule span_card_ge_dim)  | 
|
4616  | 
apply blast  | 
|
4617  | 
unfolding span_linear_image[OF lf]  | 
|
4618  | 
apply (rule subset_trans[where B = "f ` UNIV"])  | 
|
4619  | 
using sf unfolding surj_def apply blast  | 
|
4620  | 
apply (rule image_mono)  | 
|
4621  | 
apply (rule B(3))  | 
|
4622  | 
apply (metis finite_imageI fB)  | 
|
4623  | 
done  | 
|
4624  | 
||
4625  | 
moreover have "card (f ` B) \<le> card B"  | 
|
4626  | 
by (rule card_image_le, rule fB)  | 
|
4627  | 
ultimately have th1: "card B = card (f ` B)" unfolding d by arith  | 
|
4628  | 
have fiB: "inj_on f B"  | 
|
4629  | 
unfolding surjective_iff_injective_gen[OF fB finite_imageI[OF fB] th1 subset_refl, symmetric] by blast  | 
|
4630  | 
from linear_indep_image_lemma[OF lf fB fBi fiB x] fx  | 
|
4631  | 
have "x = 0" by blast}  | 
|
4632  | 
note th = this  | 
|
4633  | 
from th show ?thesis unfolding linear_injective_0[OF lf]  | 
|
4634  | 
using B(3) by blast  | 
|
4635  | 
qed  | 
|
4636  | 
||
4637  | 
(* Hence either is enough for isomorphism. *)  | 
|
4638  | 
||
4639  | 
lemma left_right_inverse_eq:  | 
|
4640  | 
assumes fg: "f o g = id" and gh: "g o h = id"  | 
|
4641  | 
shows "f = h"  | 
|
4642  | 
proof-  | 
|
4643  | 
have "f = f o (g o h)" unfolding gh by simp  | 
|
4644  | 
also have "\<dots> = (f o g) o h" by (simp add: o_assoc)  | 
|
4645  | 
finally show "f = h" unfolding fg by simp  | 
|
4646  | 
qed  | 
|
4647  | 
||
4648  | 
lemma isomorphism_expand:  | 
|
4649  | 
"f o g = id \<and> g o f = id \<longleftrightarrow> (\<forall>x. f(g x) = x) \<and> (\<forall>x. g(f x) = x)"  | 
|
4650  | 
by (simp add: expand_fun_eq o_def id_def)  | 
|
4651  | 
||
4652  | 
lemma linear_injective_isomorphism:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4653  | 
assumes lf: "linear (f :: real^'n \<Rightarrow> real ^'n)" and fi: "inj f"  | 
| 33175 | 4654  | 
shows "\<exists>f'. linear f' \<and> (\<forall>x. f' (f x) = x) \<and> (\<forall>x. f (f' x) = x)"  | 
4655  | 
unfolding isomorphism_expand[symmetric]  | 
|
4656  | 
using linear_surjective_right_inverse[OF lf linear_injective_imp_surjective[OF lf fi]] linear_injective_left_inverse[OF lf fi]  | 
|
4657  | 
by (metis left_right_inverse_eq)  | 
|
4658  | 
||
4659  | 
lemma linear_surjective_isomorphism:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4660  | 
assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'n)" and sf: "surj f"  | 
| 33175 | 4661  | 
shows "\<exists>f'. linear f' \<and> (\<forall>x. f' (f x) = x) \<and> (\<forall>x. f (f' x) = x)"  | 
4662  | 
unfolding isomorphism_expand[symmetric]  | 
|
4663  | 
using linear_surjective_right_inverse[OF lf sf] linear_injective_left_inverse[OF lf linear_surjective_imp_injective[OF lf sf]]  | 
|
4664  | 
by (metis left_right_inverse_eq)  | 
|
4665  | 
||
4666  | 
(* Left and right inverses are the same for R^N->R^N. *)  | 
|
4667  | 
||
4668  | 
lemma linear_inverse_left:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4669  | 
assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'n)" and lf': "linear f'"  | 
| 33175 | 4670  | 
shows "f o f' = id \<longleftrightarrow> f' o f = id"  | 
4671  | 
proof-  | 
|
4672  | 
  {fix f f':: "real ^'n \<Rightarrow> real ^'n"
 | 
|
4673  | 
assume lf: "linear f" "linear f'" and f: "f o f' = id"  | 
|
4674  | 
from f have sf: "surj f"  | 
|
4675  | 
||
4676  | 
apply (auto simp add: o_def stupid_ext[symmetric] id_def surj_def)  | 
|
4677  | 
by metis  | 
|
4678  | 
from linear_surjective_isomorphism[OF lf(1) sf] lf f  | 
|
4679  | 
have "f' o f = id" unfolding stupid_ext[symmetric] o_def id_def  | 
|
4680  | 
by metis}  | 
|
4681  | 
then show ?thesis using lf lf' by metis  | 
|
4682  | 
qed  | 
|
4683  | 
||
4684  | 
(* Moreover, a one-sided inverse is automatically linear. *)  | 
|
4685  | 
||
4686  | 
lemma left_inverse_linear:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4687  | 
assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'n)" and gf: "g o f = id"  | 
| 33175 | 4688  | 
shows "linear g"  | 
4689  | 
proof-  | 
|
4690  | 
from gf have fi: "inj f" apply (auto simp add: inj_on_def o_def id_def stupid_ext[symmetric])  | 
|
4691  | 
by metis  | 
|
4692  | 
from linear_injective_isomorphism[OF lf fi]  | 
|
4693  | 
obtain h:: "real ^'n \<Rightarrow> real ^'n" where  | 
|
4694  | 
h: "linear h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x" by blast  | 
|
4695  | 
have "h = g" apply (rule ext) using gf h(2,3)  | 
|
4696  | 
apply (simp add: o_def id_def stupid_ext[symmetric])  | 
|
4697  | 
by metis  | 
|
4698  | 
with h(1) show ?thesis by blast  | 
|
4699  | 
qed  | 
|
4700  | 
||
4701  | 
lemma right_inverse_linear:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4702  | 
assumes lf: "linear (f:: real ^'n \<Rightarrow> real ^'n)" and gf: "f o g = id"  | 
| 33175 | 4703  | 
shows "linear g"  | 
4704  | 
proof-  | 
|
4705  | 
from gf have fi: "surj f" apply (auto simp add: surj_def o_def id_def stupid_ext[symmetric])  | 
|
4706  | 
by metis  | 
|
4707  | 
from linear_surjective_isomorphism[OF lf fi]  | 
|
4708  | 
obtain h:: "real ^'n \<Rightarrow> real ^'n" where  | 
|
4709  | 
h: "linear h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x" by blast  | 
|
4710  | 
have "h = g" apply (rule ext) using gf h(2,3)  | 
|
4711  | 
apply (simp add: o_def id_def stupid_ext[symmetric])  | 
|
4712  | 
by metis  | 
|
4713  | 
with h(1) show ?thesis by blast  | 
|
4714  | 
qed  | 
|
4715  | 
||
4716  | 
(* The same result in terms of square matrices. *)  | 
|
4717  | 
||
4718  | 
lemma matrix_left_right_inverse:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4719  | 
fixes A A' :: "real ^'n^'n"  | 
| 33175 | 4720  | 
shows "A ** A' = mat 1 \<longleftrightarrow> A' ** A = mat 1"  | 
4721  | 
proof-  | 
|
4722  | 
  {fix A A' :: "real ^'n^'n" assume AA': "A ** A' = mat 1"
 | 
|
4723  | 
have sA: "surj (op *v A)"  | 
|
4724  | 
unfolding surj_def  | 
|
4725  | 
apply clarify  | 
|
4726  | 
apply (rule_tac x="(A' *v y)" in exI)  | 
|
4727  | 
by (simp add: matrix_vector_mul_assoc AA' matrix_vector_mul_lid)  | 
|
4728  | 
from linear_surjective_isomorphism[OF matrix_vector_mul_linear sA]  | 
|
4729  | 
obtain f' :: "real ^'n \<Rightarrow> real ^'n"  | 
|
4730  | 
where f': "linear f'" "\<forall>x. f' (A *v x) = x" "\<forall>x. A *v f' x = x" by blast  | 
|
4731  | 
have th: "matrix f' ** A = mat 1"  | 
|
4732  | 
by (simp add: matrix_eq matrix_works[OF f'(1)] matrix_vector_mul_assoc[symmetric] matrix_vector_mul_lid f'(2)[rule_format])  | 
|
4733  | 
hence "(matrix f' ** A) ** A' = mat 1 ** A'" by simp  | 
|
4734  | 
hence "matrix f' = A'" by (simp add: matrix_mul_assoc[symmetric] AA' matrix_mul_rid matrix_mul_lid)  | 
|
4735  | 
hence "matrix f' ** A = A' ** A" by simp  | 
|
4736  | 
hence "A' ** A = mat 1" by (simp add: th)}  | 
|
4737  | 
then show ?thesis by blast  | 
|
4738  | 
qed  | 
|
4739  | 
||
4740  | 
(* Considering an n-element vector as an n-by-1 or 1-by-n matrix. *)  | 
|
4741  | 
||
4742  | 
definition "rowvector v = (\<chi> i j. (v$j))"  | 
|
4743  | 
||
4744  | 
definition "columnvector v = (\<chi> i j. (v$i))"  | 
|
4745  | 
||
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4746  | 
lemma transpose_columnvector:  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4747  | 
"transpose(columnvector v) = rowvector v"  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4748  | 
by (simp add: transpose_def rowvector_def columnvector_def Cart_eq)  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4749  | 
|
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4750  | 
lemma transpose_rowvector: "transpose(rowvector v) = columnvector v"  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4751  | 
by (simp add: transpose_def columnvector_def rowvector_def Cart_eq)  | 
| 33175 | 4752  | 
|
4753  | 
lemma dot_rowvector_columnvector:  | 
|
4754  | 
"columnvector (A *v v) = A ** columnvector v"  | 
|
4755  | 
by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def)  | 
|
4756  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4757  | 
lemma dot_matrix_product: "(x::'a::semiring_1^'n) \<bullet> y = (((rowvector x ::'a^'n^1) ** (columnvector y :: 'a^1^'n))$1)$1"  | 
| 33175 | 4758  | 
by (vector matrix_matrix_mult_def rowvector_def columnvector_def dot_def)  | 
4759  | 
||
4760  | 
lemma dot_matrix_vector_mul:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4761  | 
fixes A B :: "real ^'n ^'n" and x y :: "real ^'n"  | 
| 33175 | 4762  | 
shows "(A *v x) \<bullet> (B *v y) =  | 
| 
35150
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4763  | 
(((rowvector x :: real^'n^1) ** ((transpose A ** B) ** (columnvector y :: real ^1^'n)))$1)$1"  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4764  | 
unfolding dot_matrix_product transpose_columnvector[symmetric]  | 
| 
 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 
hoelzl 
parents: 
35043 
diff
changeset
 | 
4765  | 
dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc ..  | 
| 33175 | 4766  | 
|
4767  | 
(* Infinity norm. *)  | 
|
4768  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4769  | 
definition "infnorm (x::real^'n) = Sup {abs(x$i) |i. i\<in> (UNIV :: 'n set)}"
 | 
| 33175 | 4770  | 
|
4771  | 
lemma numseg_dimindex_nonempty: "\<exists>i. i \<in> (UNIV :: 'n set)"  | 
|
4772  | 
by auto  | 
|
4773  | 
||
4774  | 
lemma infnorm_set_image:  | 
|
| 34289 | 4775  | 
  "{abs(x$i) |i. i\<in> (UNIV :: _ set)} =
 | 
4776  | 
(\<lambda>i. abs(x$i)) ` (UNIV)" by blast  | 
|
| 33175 | 4777  | 
|
4778  | 
lemma infnorm_set_lemma:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4779  | 
  shows "finite {abs((x::'a::abs ^'n)$i) |i. i\<in> (UNIV :: 'n set)}"
 | 
| 33175 | 4780  | 
  and "{abs(x$i) |i. i\<in> (UNIV :: 'n::finite set)} \<noteq> {}"
 | 
4781  | 
unfolding infnorm_set_image  | 
|
4782  | 
by (auto intro: finite_imageI)  | 
|
4783  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4784  | 
lemma infnorm_pos_le: "0 \<le> infnorm (x::real^'n)"  | 
| 33175 | 4785  | 
unfolding infnorm_def  | 
| 33270 | 4786  | 
unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma]  | 
| 33175 | 4787  | 
unfolding infnorm_set_image  | 
4788  | 
by auto  | 
|
4789  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4790  | 
lemma infnorm_triangle: "infnorm ((x::real^'n) + y) \<le> infnorm x + infnorm y"  | 
| 33175 | 4791  | 
proof-  | 
4792  | 
have th: "\<And>x y (z::real). x - y <= z \<longleftrightarrow> x - z <= y" by arith  | 
|
4793  | 
  have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
 | 
|
4794  | 
have th2: "\<And>x (y::real). abs(x + y) - abs(x) <= abs(y)" by arith  | 
|
4795  | 
show ?thesis  | 
|
4796  | 
unfolding infnorm_def  | 
|
| 33270 | 4797  | 
unfolding Sup_finite_le_iff[ OF infnorm_set_lemma]  | 
| 33175 | 4798  | 
apply (subst diff_le_eq[symmetric])  | 
| 33270 | 4799  | 
unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma]  | 
| 33175 | 4800  | 
unfolding infnorm_set_image bex_simps  | 
4801  | 
apply (subst th)  | 
|
4802  | 
unfolding th1  | 
|
| 33270 | 4803  | 
unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma]  | 
| 33175 | 4804  | 
|
4805  | 
unfolding infnorm_set_image ball_simps bex_simps  | 
|
4806  | 
apply simp  | 
|
4807  | 
apply (metis th2)  | 
|
4808  | 
done  | 
|
4809  | 
qed  | 
|
4810  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4811  | 
lemma infnorm_eq_0: "infnorm x = 0 \<longleftrightarrow> (x::real ^'n) = 0"  | 
| 33175 | 4812  | 
proof-  | 
4813  | 
have "infnorm x <= 0 \<longleftrightarrow> x = 0"  | 
|
4814  | 
unfolding infnorm_def  | 
|
| 33270 | 4815  | 
unfolding Sup_finite_le_iff[OF infnorm_set_lemma]  | 
| 33175 | 4816  | 
unfolding infnorm_set_image ball_simps  | 
4817  | 
by vector  | 
|
4818  | 
then show ?thesis using infnorm_pos_le[of x] by simp  | 
|
4819  | 
qed  | 
|
4820  | 
||
4821  | 
lemma infnorm_0: "infnorm 0 = 0"  | 
|
4822  | 
by (simp add: infnorm_eq_0)  | 
|
4823  | 
||
4824  | 
lemma infnorm_neg: "infnorm (- x) = infnorm x"  | 
|
4825  | 
unfolding infnorm_def  | 
|
| 33270 | 4826  | 
apply (rule cong[of "Sup" "Sup"])  | 
| 33175 | 4827  | 
apply blast  | 
4828  | 
apply (rule set_ext)  | 
|
4829  | 
apply auto  | 
|
4830  | 
done  | 
|
4831  | 
||
4832  | 
lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)"  | 
|
4833  | 
proof-  | 
|
4834  | 
have "y - x = - (x - y)" by simp  | 
|
4835  | 
then show ?thesis by (metis infnorm_neg)  | 
|
4836  | 
qed  | 
|
4837  | 
||
4838  | 
lemma real_abs_sub_infnorm: "\<bar> infnorm x - infnorm y\<bar> \<le> infnorm (x - y)"  | 
|
4839  | 
proof-  | 
|
4840  | 
have th: "\<And>(nx::real) n ny. nx <= n + ny \<Longrightarrow> ny <= n + nx ==> \<bar>nx - ny\<bar> <= n"  | 
|
4841  | 
by arith  | 
|
4842  | 
from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"]  | 
|
4843  | 
have ths: "infnorm x \<le> infnorm (x - y) + infnorm y"  | 
|
4844  | 
"infnorm y \<le> infnorm (x - y) + infnorm x"  | 
|
4845  | 
by (simp_all add: ring_simps infnorm_neg diff_def[symmetric])  | 
|
4846  | 
from th[OF ths] show ?thesis .  | 
|
4847  | 
qed  | 
|
4848  | 
||
4849  | 
lemma real_abs_infnorm: " \<bar>infnorm x\<bar> = infnorm x"  | 
|
4850  | 
using infnorm_pos_le[of x] by arith  | 
|
4851  | 
||
4852  | 
lemma component_le_infnorm:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4853  | 
shows "\<bar>x$i\<bar> \<le> infnorm (x::real^'n)"  | 
| 33175 | 4854  | 
proof-  | 
4855  | 
let ?U = "UNIV :: 'n set"  | 
|
4856  | 
  let ?S = "{\<bar>x$i\<bar> |i. i\<in> ?U}"
 | 
|
4857  | 
have fS: "finite ?S" unfolding image_Collect[symmetric]  | 
|
4858  | 
apply (rule finite_imageI) unfolding Collect_def mem_def by simp  | 
|
4859  | 
  have S0: "?S \<noteq> {}" by blast
 | 
|
4860  | 
  have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
 | 
|
| 33270 | 4861  | 
from Sup_finite_in[OF fS S0]  | 
4862  | 
show ?thesis unfolding infnorm_def infnorm_set_image  | 
|
4863  | 
by (metis Sup_finite_ge_iff finite finite_imageI UNIV_not_empty image_is_empty  | 
|
4864  | 
rangeI real_le_refl)  | 
|
| 33175 | 4865  | 
qed  | 
4866  | 
||
4867  | 
lemma infnorm_mul_lemma: "infnorm(a *s x) <= \<bar>a\<bar> * infnorm x"  | 
|
4868  | 
apply (subst infnorm_def)  | 
|
| 33270 | 4869  | 
unfolding Sup_finite_le_iff[OF infnorm_set_lemma]  | 
| 33175 | 4870  | 
unfolding infnorm_set_image ball_simps  | 
4871  | 
apply (simp add: abs_mult)  | 
|
4872  | 
apply (rule allI)  | 
|
4873  | 
apply (cut_tac component_le_infnorm[of x])  | 
|
4874  | 
apply (rule mult_mono)  | 
|
4875  | 
apply auto  | 
|
4876  | 
done  | 
|
4877  | 
||
4878  | 
lemma infnorm_mul: "infnorm(a *s x) = abs a * infnorm x"  | 
|
4879  | 
proof-  | 
|
4880  | 
  {assume a0: "a = 0" hence ?thesis by (simp add: infnorm_0) }
 | 
|
4881  | 
moreover  | 
|
4882  | 
  {assume a0: "a \<noteq> 0"
 | 
|
4883  | 
from a0 have th: "(1/a) *s (a *s x) = x"  | 
|
4884  | 
by (simp add: vector_smult_assoc)  | 
|
4885  | 
from a0 have ap: "\<bar>a\<bar> > 0" by arith  | 
|
4886  | 
from infnorm_mul_lemma[of "1/a" "a *s x"]  | 
|
4887  | 
have "infnorm x \<le> 1/\<bar>a\<bar> * infnorm (a*s x)"  | 
|
4888  | 
unfolding th by simp  | 
|
4889  | 
with ap have "\<bar>a\<bar> * infnorm x \<le> \<bar>a\<bar> * (1/\<bar>a\<bar> * infnorm (a *s x))" by (simp add: field_simps)  | 
|
4890  | 
then have "\<bar>a\<bar> * infnorm x \<le> infnorm (a*s x)"  | 
|
4891  | 
using ap by (simp add: field_simps)  | 
|
4892  | 
with infnorm_mul_lemma[of a x] have ?thesis by arith }  | 
|
4893  | 
ultimately show ?thesis by blast  | 
|
4894  | 
qed  | 
|
4895  | 
||
4896  | 
lemma infnorm_pos_lt: "infnorm x > 0 \<longleftrightarrow> x \<noteq> 0"  | 
|
4897  | 
using infnorm_pos_le[of x] infnorm_eq_0[of x] by arith  | 
|
4898  | 
||
4899  | 
(* Prove that it differs only up to a bound from Euclidean norm. *)  | 
|
4900  | 
||
4901  | 
lemma infnorm_le_norm: "infnorm x \<le> norm x"  | 
|
| 33270 | 4902  | 
unfolding infnorm_def Sup_finite_le_iff[OF infnorm_set_lemma]  | 
| 33175 | 4903  | 
unfolding infnorm_set_image ball_simps  | 
4904  | 
by (metis component_le_norm)  | 
|
4905  | 
lemma card_enum: "card {1 .. n} = n" by auto
 | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4906  | 
lemma norm_le_infnorm: "norm(x) <= sqrt(real CARD('n)) * infnorm(x::real ^'n)"
 | 
| 33175 | 4907  | 
proof-  | 
4908  | 
  let ?d = "CARD('n)"
 | 
|
4909  | 
have "real ?d \<ge> 0" by simp  | 
|
4910  | 
hence d2: "(sqrt (real ?d))^2 = real ?d"  | 
|
4911  | 
by (auto intro: real_sqrt_pow2)  | 
|
4912  | 
have th: "sqrt (real ?d) * infnorm x \<ge> 0"  | 
|
4913  | 
by (simp add: zero_le_mult_iff real_sqrt_ge_0_iff infnorm_pos_le)  | 
|
4914  | 
have th1: "x\<bullet>x \<le> (sqrt (real ?d) * infnorm x)^2"  | 
|
4915  | 
unfolding power_mult_distrib d2  | 
|
4916  | 
apply (subst power2_abs[symmetric])  | 
|
4917  | 
unfolding real_of_nat_def dot_def power2_eq_square[symmetric]  | 
|
4918  | 
apply (subst power2_abs[symmetric])  | 
|
4919  | 
apply (rule setsum_bounded)  | 
|
4920  | 
apply (rule power_mono)  | 
|
4921  | 
unfolding abs_of_nonneg[OF infnorm_pos_le]  | 
|
| 33270 | 4922  | 
unfolding infnorm_def Sup_finite_ge_iff[OF infnorm_set_lemma]  | 
| 33175 | 4923  | 
unfolding infnorm_set_image bex_simps  | 
4924  | 
apply blast  | 
|
4925  | 
by (rule abs_ge_zero)  | 
|
4926  | 
from real_le_lsqrt[OF dot_pos_le th th1]  | 
|
4927  | 
show ?thesis unfolding real_vector_norm_def id_def .  | 
|
4928  | 
qed  | 
|
4929  | 
||
4930  | 
(* Equality in Cauchy-Schwarz and triangle inequalities. *)  | 
|
4931  | 
||
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4932  | 
lemma norm_cauchy_schwarz_eq: "(x::real ^'n) \<bullet> y = norm x * norm y \<longleftrightarrow> norm x *s y = norm y *s x" (is "?lhs \<longleftrightarrow> ?rhs")  | 
| 33175 | 4933  | 
proof-  | 
4934  | 
  {assume h: "x = 0"
 | 
|
4935  | 
hence ?thesis by simp}  | 
|
4936  | 
moreover  | 
|
4937  | 
  {assume h: "y = 0"
 | 
|
4938  | 
hence ?thesis by simp}  | 
|
4939  | 
moreover  | 
|
4940  | 
  {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
 | 
|
4941  | 
from dot_eq_0[of "norm y *s x - norm x *s y"]  | 
|
4942  | 
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)"  | 
|
4943  | 
using x y  | 
|
4944  | 
unfolding dot_rsub dot_lsub dot_lmult dot_rmult  | 
|
4945  | 
unfolding norm_pow_2[symmetric] power2_eq_square diff_eq_0_iff_eq apply (simp add: dot_sym)  | 
|
4946  | 
apply (simp add: ring_simps)  | 
|
4947  | 
apply metis  | 
|
4948  | 
done  | 
|
4949  | 
also have "\<dots> \<longleftrightarrow> (2 * norm x * norm y * (norm x * norm y - x \<bullet> y) = 0)" using x y  | 
|
4950  | 
by (simp add: ring_simps dot_sym)  | 
|
4951  | 
also have "\<dots> \<longleftrightarrow> ?lhs" using x y  | 
|
4952  | 
apply simp  | 
|
4953  | 
by metis  | 
|
4954  | 
finally have ?thesis by blast}  | 
|
4955  | 
ultimately show ?thesis by blast  | 
|
4956  | 
qed  | 
|
4957  | 
||
4958  | 
lemma norm_cauchy_schwarz_abs_eq:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4959  | 
fixes x y :: "real ^ 'n"  | 
| 33175 | 4960  | 
shows "abs(x \<bullet> y) = norm x * norm y \<longleftrightarrow>  | 
4961  | 
norm x *s y = norm y *s x \<or> norm(x) *s y = - norm y *s x" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
4962  | 
proof-  | 
|
4963  | 
have th: "\<And>(x::real) a. a \<ge> 0 \<Longrightarrow> abs x = a \<longleftrightarrow> x = a \<or> x = - a" by arith  | 
|
4964  | 
have "?rhs \<longleftrightarrow> norm x *s y = norm y *s x \<or> norm (- x) *s y = norm y *s (- x)"  | 
|
4965  | 
apply simp by vector  | 
|
4966  | 
also have "\<dots> \<longleftrightarrow>(x \<bullet> y = norm x * norm y \<or>  | 
|
4967  | 
(-x) \<bullet> y = norm x * norm y)"  | 
|
4968  | 
unfolding norm_cauchy_schwarz_eq[symmetric]  | 
|
4969  | 
unfolding norm_minus_cancel  | 
|
4970  | 
norm_mul by blast  | 
|
4971  | 
also have "\<dots> \<longleftrightarrow> ?lhs"  | 
|
4972  | 
unfolding th[OF mult_nonneg_nonneg, OF norm_ge_zero[of x] norm_ge_zero[of y]] dot_lneg  | 
|
4973  | 
by arith  | 
|
4974  | 
finally show ?thesis ..  | 
|
4975  | 
qed  | 
|
4976  | 
||
4977  | 
lemma norm_triangle_eq:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
4978  | 
fixes x y :: "real ^ 'n"  | 
| 33175 | 4979  | 
shows "norm(x + y) = norm x + norm y \<longleftrightarrow> norm x *s y = norm y *s x"  | 
4980  | 
proof-  | 
|
4981  | 
  {assume x: "x =0 \<or> y =0"
 | 
|
4982  | 
hence ?thesis by (cases "x=0", simp_all)}  | 
|
4983  | 
moreover  | 
|
4984  | 
  {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
 | 
|
4985  | 
hence "norm x \<noteq> 0" "norm y \<noteq> 0"  | 
|
4986  | 
by simp_all  | 
|
4987  | 
hence n: "norm x > 0" "norm y > 0"  | 
|
4988  | 
using norm_ge_zero[of x] norm_ge_zero[of y]  | 
|
4989  | 
by arith+  | 
|
4990  | 
have th: "\<And>(a::real) b c. a + b + c \<noteq> 0 ==> (a = b + c \<longleftrightarrow> a^2 = (b + c)^2)" by algebra  | 
|
4991  | 
have "norm(x + y) = norm x + norm y \<longleftrightarrow> norm(x + y)^ 2 = (norm x + norm y) ^2"  | 
|
4992  | 
apply (rule th) using n norm_ge_zero[of "x + y"]  | 
|
4993  | 
by arith  | 
|
4994  | 
also have "\<dots> \<longleftrightarrow> norm x *s y = norm y *s x"  | 
|
4995  | 
unfolding norm_cauchy_schwarz_eq[symmetric]  | 
|
4996  | 
unfolding norm_pow_2 dot_ladd dot_radd  | 
|
4997  | 
by (simp add: norm_pow_2[symmetric] power2_eq_square dot_sym ring_simps)  | 
|
4998  | 
finally have ?thesis .}  | 
|
4999  | 
ultimately show ?thesis by blast  | 
|
5000  | 
qed  | 
|
5001  | 
||
5002  | 
(* Collinearity.*)  | 
|
5003  | 
||
5004  | 
definition "collinear S \<longleftrightarrow> (\<exists>u. \<forall>x \<in> S. \<forall> y \<in> S. \<exists>c. x - y = c *s u)"  | 
|
5005  | 
||
5006  | 
lemma collinear_empty:  "collinear {}" by (simp add: collinear_def)
 | 
|
5007  | 
||
| 34289 | 5008  | 
lemma collinear_sing: "collinear {(x::'a::ring_1^_)}"
 | 
| 33175 | 5009  | 
apply (simp add: collinear_def)  | 
5010  | 
apply (rule exI[where x=0])  | 
|
5011  | 
by simp  | 
|
5012  | 
||
| 34289 | 5013  | 
lemma collinear_2: "collinear {(x::'a::ring_1^_),y}"
 | 
| 33175 | 5014  | 
apply (simp add: collinear_def)  | 
5015  | 
apply (rule exI[where x="x - y"])  | 
|
5016  | 
apply auto  | 
|
5017  | 
apply (rule exI[where x=0], simp)  | 
|
5018  | 
apply (rule exI[where x=1], simp)  | 
|
5019  | 
apply (rule exI[where x="- 1"], simp add: vector_sneg_minus1[symmetric])  | 
|
5020  | 
apply (rule exI[where x=0], simp)  | 
|
5021  | 
done  | 
|
5022  | 
||
| 34289 | 5023  | 
lemma collinear_lemma: "collinear {(0::real^_),x,y} \<longleftrightarrow> x = 0 \<or> y = 0 \<or> (\<exists>c. y = c *s x)" (is "?lhs \<longleftrightarrow> ?rhs")
 | 
| 33175 | 5024  | 
proof-  | 
5025  | 
  {assume "x=0 \<or> y = 0" hence ?thesis
 | 
|
5026  | 
by (cases "x = 0", simp_all add: collinear_2 insert_commute)}  | 
|
5027  | 
moreover  | 
|
5028  | 
  {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
 | 
|
5029  | 
    {assume h: "?lhs"
 | 
|
5030  | 
      then obtain u where u: "\<forall> x\<in> {0,x,y}. \<forall>y\<in> {0,x,y}. \<exists>c. x - y = c *s u" unfolding collinear_def by blast
 | 
|
5031  | 
from u[rule_format, of x 0] u[rule_format, of y 0]  | 
|
5032  | 
obtain cx and cy where  | 
|
5033  | 
cx: "x = cx*s u" and cy: "y = cy*s u"  | 
|
5034  | 
by auto  | 
|
5035  | 
from cx x have cx0: "cx \<noteq> 0" by auto  | 
|
5036  | 
from cy y have cy0: "cy \<noteq> 0" by auto  | 
|
5037  | 
let ?d = "cy / cx"  | 
|
5038  | 
from cx cy cx0 have "y = ?d *s x"  | 
|
5039  | 
by (simp add: vector_smult_assoc)  | 
|
5040  | 
hence ?rhs using x y by blast}  | 
|
5041  | 
moreover  | 
|
5042  | 
    {assume h: "?rhs"
 | 
|
5043  | 
then obtain c where c: "y = c*s x" using x y by blast  | 
|
5044  | 
have ?lhs unfolding collinear_def c  | 
|
5045  | 
apply (rule exI[where x=x])  | 
|
5046  | 
apply auto  | 
|
5047  | 
apply (rule exI[where x="- 1"], simp only: vector_smult_lneg vector_smult_lid)  | 
|
5048  | 
apply (rule exI[where x= "-c"], simp only: vector_smult_lneg)  | 
|
5049  | 
apply (rule exI[where x=1], simp)  | 
|
5050  | 
apply (rule exI[where x="1 - c"], simp add: vector_smult_lneg vector_sub_rdistrib)  | 
|
5051  | 
apply (rule exI[where x="c - 1"], simp add: vector_smult_lneg vector_sub_rdistrib)  | 
|
5052  | 
done}  | 
|
5053  | 
ultimately have ?thesis by blast}  | 
|
5054  | 
ultimately show ?thesis by blast  | 
|
5055  | 
qed  | 
|
5056  | 
||
5057  | 
lemma norm_cauchy_schwarz_equal:  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
5058  | 
fixes x y :: "real ^ 'n"  | 
| 33175 | 5059  | 
  shows "abs(x \<bullet> y) = norm x * norm y \<longleftrightarrow> collinear {(0::real^'n),x,y}"
 | 
5060  | 
unfolding norm_cauchy_schwarz_abs_eq  | 
|
5061  | 
apply (cases "x=0", simp_all add: collinear_2)  | 
|
5062  | 
apply (cases "y=0", simp_all add: collinear_2 insert_commute)  | 
|
5063  | 
unfolding collinear_lemma  | 
|
5064  | 
apply simp  | 
|
5065  | 
apply (subgoal_tac "norm x \<noteq> 0")  | 
|
5066  | 
apply (subgoal_tac "norm y \<noteq> 0")  | 
|
5067  | 
apply (rule iffI)  | 
|
5068  | 
apply (cases "norm x *s y = norm y *s x")  | 
|
5069  | 
apply (rule exI[where x="(1/norm x) * norm y"])  | 
|
5070  | 
apply (drule sym)  | 
|
5071  | 
unfolding vector_smult_assoc[symmetric]  | 
|
5072  | 
apply (simp add: vector_smult_assoc field_simps)  | 
|
5073  | 
apply (rule exI[where x="(1/norm x) * - norm y"])  | 
|
5074  | 
apply clarify  | 
|
5075  | 
apply (drule sym)  | 
|
5076  | 
unfolding vector_smult_assoc[symmetric]  | 
|
5077  | 
apply (simp add: vector_smult_assoc field_simps)  | 
|
5078  | 
apply (erule exE)  | 
|
5079  | 
apply (erule ssubst)  | 
|
5080  | 
unfolding vector_smult_assoc  | 
|
5081  | 
unfolding norm_mul  | 
|
5082  | 
apply (subgoal_tac "norm x * c = \<bar>c\<bar> * norm x \<or> norm x * c = - \<bar>c\<bar> * norm x")  | 
|
5083  | 
apply (case_tac "c <= 0", simp add: ring_simps)  | 
|
5084  | 
apply (simp add: ring_simps)  | 
|
5085  | 
apply (case_tac "c <= 0", simp add: ring_simps)  | 
|
5086  | 
apply (simp add: ring_simps)  | 
|
5087  | 
apply simp  | 
|
5088  | 
apply simp  | 
|
5089  | 
done  | 
|
5090  | 
||
5091  | 
end  |