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