| author | wenzelm | 
| Wed, 26 Dec 2018 20:57:23 +0100 | |
| changeset 69506 | 7d59af98af29 | 
| parent 68901 | 4824cc40f42e | 
| child 69510 | 0f31dd2e540d | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/Linear_Algebra.thy | 
| 44133 | 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 | |
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66447diff
changeset | 10 | "HOL-Library.Infinite_Set" | 
| 44133 | 11 | begin | 
| 12 | ||
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63881diff
changeset | 13 | lemma linear_simps: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63881diff
changeset | 14 | assumes "bounded_linear f" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63881diff
changeset | 15 | shows | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63881diff
changeset | 16 | "f (a + b) = f a + f b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63881diff
changeset | 17 | "f (a - b) = f a - f b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63881diff
changeset | 18 | "f 0 = 0" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63881diff
changeset | 19 | "f (- a) = - f a" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63881diff
changeset | 20 | "f (s *\<^sub>R v) = s *\<^sub>R (f v)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63881diff
changeset | 21 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63881diff
changeset | 22 | interpret f: bounded_linear f by fact | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63881diff
changeset | 23 | show "f (a + b) = f a + f b" by (rule f.add) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63881diff
changeset | 24 | show "f (a - b) = f a - f b" by (rule f.diff) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63881diff
changeset | 25 | show "f 0 = 0" by (rule f.zero) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 26 | show "f (- a) = - f a" by (rule f.neg) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 27 | show "f (s *\<^sub>R v) = s *\<^sub>R (f v)" by (rule f.scale) | 
| 44133 | 28 | qed | 
| 29 | ||
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 30 | lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x \<in> (UNIV::'a::finite set)}"
 | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 31 | using finite finite_image_set by blast | 
| 44133 | 32 | |
| 53406 | 33 | |
| 68901 | 34 | subsection%unimportant \<open>More interesting properties of the norm\<close> | 
| 63050 | 35 | |
| 36 | notation inner (infix "\<bullet>" 70) | |
| 37 | ||
| 67399 | 38 | text\<open>Equality of vectors in terms of @{term "(\<bullet>)"} products.\<close>
 | 
| 63050 | 39 | |
| 40 | lemma linear_componentwise: | |
| 41 | fixes f:: "'a::euclidean_space \<Rightarrow> 'b::real_inner" | |
| 42 | assumes lf: "linear f" | |
| 43 | shows "(f x) \<bullet> j = (\<Sum>i\<in>Basis. (x\<bullet>i) * (f i\<bullet>j))" (is "?lhs = ?rhs") | |
| 44 | proof - | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 45 | interpret linear f by fact | 
| 63050 | 46 | have "?rhs = (\<Sum>i\<in>Basis. (x\<bullet>i) *\<^sub>R (f i))\<bullet>j" | 
| 64267 | 47 | by (simp add: inner_sum_left) | 
| 63050 | 48 | then show ?thesis | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 49 | by (simp add: euclidean_representation sum[symmetric] scale[symmetric]) | 
| 63050 | 50 | qed | 
| 51 | ||
| 52 | lemma vector_eq: "x = y \<longleftrightarrow> x \<bullet> x = x \<bullet> y \<and> y \<bullet> y = x \<bullet> x" | |
| 53 | (is "?lhs \<longleftrightarrow> ?rhs") | |
| 54 | proof | |
| 55 | assume ?lhs | |
| 56 | then show ?rhs by simp | |
| 57 | next | |
| 58 | assume ?rhs | |
| 59 | then have "x \<bullet> x - x \<bullet> y = 0 \<and> x \<bullet> y - y \<bullet> y = 0" | |
| 60 | by simp | |
| 61 | then have "x \<bullet> (x - y) = 0 \<and> y \<bullet> (x - y) = 0" | |
| 62 | by (simp add: inner_diff inner_commute) | |
| 63 | then have "(x - y) \<bullet> (x - y) = 0" | |
| 64 | by (simp add: field_simps inner_diff inner_commute) | |
| 65 | then show "x = y" by simp | |
| 66 | qed | |
| 67 | ||
| 68 | lemma norm_triangle_half_r: | |
| 69 | "norm (y - x1) < e / 2 \<Longrightarrow> norm (y - x2) < e / 2 \<Longrightarrow> norm (x1 - x2) < e" | |
| 70 | using dist_triangle_half_r unfolding dist_norm[symmetric] by auto | |
| 71 | ||
| 72 | lemma norm_triangle_half_l: | |
| 73 | assumes "norm (x - y) < e / 2" | |
| 74 | and "norm (x' - y) < e / 2" | |
| 75 | shows "norm (x - x') < e" | |
| 76 | using dist_triangle_half_l[OF assms[unfolded dist_norm[symmetric]]] | |
| 77 | unfolding dist_norm[symmetric] . | |
| 78 | ||
| 79 | lemma norm_triangle_le: "norm x + norm y \<le> e \<Longrightarrow> norm (x + y) \<le> e" | |
| 80 | by (rule norm_triangle_ineq [THEN order_trans]) | |
| 81 | ||
| 82 | lemma norm_triangle_lt: "norm x + norm y < e \<Longrightarrow> norm (x + y) < e" | |
| 83 | by (rule norm_triangle_ineq [THEN le_less_trans]) | |
| 84 | ||
| 66420 | 85 | lemma abs_triangle_half_r: | 
| 86 | fixes y :: "'a::linordered_field" | |
| 87 | shows "abs (y - x1) < e / 2 \<Longrightarrow> abs (y - x2) < e / 2 \<Longrightarrow> abs (x1 - x2) < e" | |
| 88 | by linarith | |
| 89 | ||
| 90 | lemma abs_triangle_half_l: | |
| 91 | fixes y :: "'a::linordered_field" | |
| 92 | assumes "abs (x - y) < e / 2" | |
| 93 | and "abs (x' - y) < e / 2" | |
| 94 | shows "abs (x - x') < e" | |
| 95 | using assms by linarith | |
| 96 | ||
| 64267 | 97 | lemma sum_clauses: | 
| 98 |   shows "sum f {} = 0"
 | |
| 99 | and "finite S \<Longrightarrow> sum f (insert x S) = (if x \<in> S then sum f S else f x + sum f S)" | |
| 63050 | 100 | by (auto simp add: insert_absorb) | 
| 101 | ||
| 64267 | 102 | lemma sum_norm_bound: | 
| 63050 | 103 | fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64267diff
changeset | 104 | assumes K: "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> K" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64267diff
changeset | 105 | shows "norm (sum f S) \<le> of_nat (card S)*K" | 
| 64267 | 106 | using sum_norm_le[OF K] sum_constant[symmetric] | 
| 63050 | 107 | by simp | 
| 108 | ||
| 64267 | 109 | lemma sum_group: | 
| 63050 | 110 | assumes fS: "finite S" and fT: "finite T" and fST: "f ` S \<subseteq> T" | 
| 64267 | 111 |   shows "sum (\<lambda>y. sum g {x. x \<in> S \<and> f x = y}) T = sum g S"
 | 
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 112 | unfolding sum_image_gen[OF fS, of g f] | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 113 | by (auto intro: sum.neutral sum.mono_neutral_right[OF fT fST]) | 
| 63050 | 114 | |
| 115 | lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = z" | |
| 116 | proof | |
| 117 | assume "\<forall>x. x \<bullet> y = x \<bullet> z" | |
| 118 | then have "\<forall>x. x \<bullet> (y - z) = 0" | |
| 119 | by (simp add: inner_diff) | |
| 120 | then have "(y - z) \<bullet> (y - z) = 0" .. | |
| 121 | then show "y = z" by simp | |
| 122 | qed simp | |
| 123 | ||
| 124 | lemma vector_eq_rdot: "(\<forall>z. x \<bullet> z = y \<bullet> z) \<longleftrightarrow> x = y" | |
| 125 | proof | |
| 126 | assume "\<forall>z. x \<bullet> z = y \<bullet> z" | |
| 127 | then have "\<forall>z. (x - y) \<bullet> z = 0" | |
| 128 | by (simp add: inner_diff) | |
| 129 | then have "(x - y) \<bullet> (x - y) = 0" .. | |
| 130 | then show "x = y" by simp | |
| 131 | qed simp | |
| 132 | ||
| 133 | ||
| 68901 | 134 | subsection \<open>Orthogonality\<close> | 
| 63050 | 135 | |
| 67962 | 136 | definition%important (in real_inner) "orthogonal x y \<longleftrightarrow> x \<bullet> y = 0" | 
| 137 | ||
| 63050 | 138 | context real_inner | 
| 139 | begin | |
| 140 | ||
| 63072 | 141 | lemma orthogonal_self: "orthogonal x x \<longleftrightarrow> x = 0" | 
| 142 | by (simp add: orthogonal_def) | |
| 143 | ||
| 63050 | 144 | lemma orthogonal_clauses: | 
| 145 | "orthogonal a 0" | |
| 146 | "orthogonal a x \<Longrightarrow> orthogonal a (c *\<^sub>R x)" | |
| 147 | "orthogonal a x \<Longrightarrow> orthogonal a (- x)" | |
| 148 | "orthogonal a x \<Longrightarrow> orthogonal a y \<Longrightarrow> orthogonal a (x + y)" | |
| 149 | "orthogonal a x \<Longrightarrow> orthogonal a y \<Longrightarrow> orthogonal a (x - y)" | |
| 150 | "orthogonal 0 a" | |
| 151 | "orthogonal x a \<Longrightarrow> orthogonal (c *\<^sub>R x) a" | |
| 152 | "orthogonal x a \<Longrightarrow> orthogonal (- x) a" | |
| 153 | "orthogonal x a \<Longrightarrow> orthogonal y a \<Longrightarrow> orthogonal (x + y) a" | |
| 154 | "orthogonal x a \<Longrightarrow> orthogonal y a \<Longrightarrow> orthogonal (x - y) a" | |
| 155 | unfolding orthogonal_def inner_add inner_diff by auto | |
| 156 | ||
| 157 | end | |
| 158 | ||
| 159 | lemma orthogonal_commute: "orthogonal x y \<longleftrightarrow> orthogonal y x" | |
| 160 | by (simp add: orthogonal_def inner_commute) | |
| 161 | ||
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 162 | lemma orthogonal_scaleR [simp]: "c \<noteq> 0 \<Longrightarrow> orthogonal (c *\<^sub>R x) = orthogonal x" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 163 | by (rule ext) (simp add: orthogonal_def) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 164 | |
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 165 | lemma pairwise_ortho_scaleR: | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 166 | "pairwise (\<lambda>i j. orthogonal (f i) (g j)) B | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 167 | \<Longrightarrow> pairwise (\<lambda>i j. orthogonal (a i *\<^sub>R f i) (a j *\<^sub>R g j)) B" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 168 | by (auto simp: pairwise_def orthogonal_clauses) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 169 | |
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 170 | lemma orthogonal_rvsum: | 
| 64267 | 171 | "\<lbrakk>finite s; \<And>y. y \<in> s \<Longrightarrow> orthogonal x (f y)\<rbrakk> \<Longrightarrow> orthogonal x (sum f s)" | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 172 | by (induction s rule: finite_induct) (auto simp: orthogonal_clauses) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 173 | |
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 174 | lemma orthogonal_lvsum: | 
| 64267 | 175 | "\<lbrakk>finite s; \<And>x. x \<in> s \<Longrightarrow> orthogonal (f x) y\<rbrakk> \<Longrightarrow> orthogonal (sum f s) y" | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 176 | by (induction s rule: finite_induct) (auto simp: orthogonal_clauses) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 177 | |
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 178 | lemma norm_add_Pythagorean: | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 179 | assumes "orthogonal a b" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 180 | shows "norm(a + b) ^ 2 = norm a ^ 2 + norm b ^ 2" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 181 | proof - | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 182 | from assms have "(a - (0 - b)) \<bullet> (a - (0 - b)) = a \<bullet> a - (0 - b \<bullet> b)" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 183 | by (simp add: algebra_simps orthogonal_def inner_commute) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 184 | then show ?thesis | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 185 | by (simp add: power2_norm_eq_inner) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 186 | qed | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 187 | |
| 64267 | 188 | lemma norm_sum_Pythagorean: | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 189 | assumes "finite I" "pairwise (\<lambda>i j. orthogonal (f i) (f j)) I" | 
| 64267 | 190 | shows "(norm (sum f I))\<^sup>2 = (\<Sum>i\<in>I. (norm (f i))\<^sup>2)" | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 191 | using assms | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 192 | proof (induction I rule: finite_induct) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 193 | case empty then show ?case by simp | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 194 | next | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 195 | case (insert x I) | 
| 64267 | 196 | then have "orthogonal (f x) (sum f I)" | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 197 | by (metis pairwise_insert orthogonal_rvsum) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 198 | with insert show ?case | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 199 | by (simp add: pairwise_insert norm_add_Pythagorean) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 200 | qed | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 201 | |
| 63050 | 202 | |
| 68901 | 203 | subsection \<open>Bilinear functions\<close> | 
| 63050 | 204 | |
| 67962 | 205 | definition%important "bilinear f \<longleftrightarrow> (\<forall>x. linear (\<lambda>y. f x y)) \<and> (\<forall>y. linear (\<lambda>x. f x y))" | 
| 63050 | 206 | |
| 207 | lemma bilinear_ladd: "bilinear h \<Longrightarrow> h (x + y) z = h x z + h y z" | |
| 208 | by (simp add: bilinear_def linear_iff) | |
| 209 | ||
| 210 | lemma bilinear_radd: "bilinear h \<Longrightarrow> h x (y + z) = h x y + h x z" | |
| 211 | by (simp add: bilinear_def linear_iff) | |
| 212 | ||
| 213 | lemma bilinear_lmul: "bilinear h \<Longrightarrow> h (c *\<^sub>R x) y = c *\<^sub>R h x y" | |
| 214 | by (simp add: bilinear_def linear_iff) | |
| 215 | ||
| 216 | lemma bilinear_rmul: "bilinear h \<Longrightarrow> h x (c *\<^sub>R y) = c *\<^sub>R h x y" | |
| 217 | by (simp add: bilinear_def linear_iff) | |
| 218 | ||
| 219 | lemma bilinear_lneg: "bilinear h \<Longrightarrow> h (- x) y = - h x y" | |
| 220 | by (drule bilinear_lmul [of _ "- 1"]) simp | |
| 221 | ||
| 222 | lemma bilinear_rneg: "bilinear h \<Longrightarrow> h x (- y) = - h x y" | |
| 223 | by (drule bilinear_rmul [of _ _ "- 1"]) simp | |
| 224 | ||
| 225 | lemma (in ab_group_add) eq_add_iff: "x = x + y \<longleftrightarrow> y = 0" | |
| 226 | using add_left_imp_eq[of x y 0] by auto | |
| 227 | ||
| 228 | lemma bilinear_lzero: | |
| 229 | assumes "bilinear h" | |
| 230 | shows "h 0 x = 0" | |
| 231 | using bilinear_ladd [OF assms, of 0 0 x] by (simp add: eq_add_iff field_simps) | |
| 232 | ||
| 233 | lemma bilinear_rzero: | |
| 234 | assumes "bilinear h" | |
| 235 | shows "h x 0 = 0" | |
| 236 | using bilinear_radd [OF assms, of x 0 0 ] by (simp add: eq_add_iff field_simps) | |
| 237 | ||
| 238 | lemma bilinear_lsub: "bilinear h \<Longrightarrow> h (x - y) z = h x z - h y z" | |
| 239 | using bilinear_ladd [of h x "- y"] by (simp add: bilinear_lneg) | |
| 240 | ||
| 241 | lemma bilinear_rsub: "bilinear h \<Longrightarrow> h z (x - y) = h z x - h z y" | |
| 242 | using bilinear_radd [of h _ x "- y"] by (simp add: bilinear_rneg) | |
| 243 | ||
| 64267 | 244 | lemma bilinear_sum: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 245 | assumes "bilinear h" | 
| 64267 | 246 | shows "h (sum f S) (sum g T) = sum (\<lambda>(i,j). h (f i) (g j)) (S \<times> T) " | 
| 63050 | 247 | proof - | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 248 | interpret l: linear "\<lambda>x. h x y" for y using assms by (simp add: bilinear_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 249 | interpret r: linear "\<lambda>y. h x y" for x using assms by (simp add: bilinear_def) | 
| 64267 | 250 | have "h (sum f S) (sum g T) = sum (\<lambda>x. h (f x) (sum g T)) S" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 251 | by (simp add: l.sum) | 
| 64267 | 252 | also have "\<dots> = sum (\<lambda>x. sum (\<lambda>y. h (f x) (g y)) T) S" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 253 | by (rule sum.cong) (simp_all add: r.sum) | 
| 63050 | 254 | finally show ?thesis | 
| 64267 | 255 | unfolding sum.cartesian_product . | 
| 63050 | 256 | qed | 
| 257 | ||
| 258 | ||
| 68901 | 259 | subsection \<open>Adjoints\<close> | 
| 63050 | 260 | |
| 67962 | 261 | definition%important "adjoint f = (SOME f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y)" | 
| 63050 | 262 | |
| 263 | lemma adjoint_unique: | |
| 264 | assumes "\<forall>x y. inner (f x) y = inner x (g y)" | |
| 265 | shows "adjoint f = g" | |
| 266 | unfolding adjoint_def | |
| 267 | proof (rule some_equality) | |
| 268 | show "\<forall>x y. inner (f x) y = inner x (g y)" | |
| 269 | by (rule assms) | |
| 270 | next | |
| 271 | fix h | |
| 272 | assume "\<forall>x y. inner (f x) y = inner x (h y)" | |
| 273 | then have "\<forall>x y. inner x (g y) = inner x (h y)" | |
| 274 | using assms by simp | |
| 275 | then have "\<forall>x y. inner x (g y - h y) = 0" | |
| 276 | by (simp add: inner_diff_right) | |
| 277 | then have "\<forall>y. inner (g y - h y) (g y - h y) = 0" | |
| 278 | by simp | |
| 279 | then have "\<forall>y. h y = g y" | |
| 280 | by simp | |
| 281 | then show "h = g" by (simp add: ext) | |
| 282 | qed | |
| 283 | ||
| 284 | text \<open>TODO: The following lemmas about adjoints should hold for any | |
| 63680 | 285 | Hilbert space (i.e. complete inner product space). | 
| 68224 | 286 | (see \<^url>\<open>https://en.wikipedia.org/wiki/Hermitian_adjoint\<close>) | 
| 63050 | 287 | \<close> | 
| 288 | ||
| 289 | lemma adjoint_works: | |
| 290 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | |
| 291 | assumes lf: "linear f" | |
| 292 | shows "x \<bullet> adjoint f y = f x \<bullet> y" | |
| 293 | proof - | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 294 | interpret linear f by fact | 
| 63050 | 295 | have "\<forall>y. \<exists>w. \<forall>x. f x \<bullet> y = x \<bullet> w" | 
| 296 | proof (intro allI exI) | |
| 297 | fix y :: "'m" and x | |
| 298 | let ?w = "(\<Sum>i\<in>Basis. (f i \<bullet> y) *\<^sub>R i) :: 'n" | |
| 299 | have "f x \<bullet> y = f (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i) \<bullet> y" | |
| 300 | by (simp add: euclidean_representation) | |
| 301 | also have "\<dots> = (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R f i) \<bullet> y" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 302 | by (simp add: sum scale) | 
| 63050 | 303 | finally show "f x \<bullet> y = x \<bullet> ?w" | 
| 64267 | 304 | by (simp add: inner_sum_left inner_sum_right mult.commute) | 
| 63050 | 305 | qed | 
| 306 | then show ?thesis | |
| 307 | unfolding adjoint_def choice_iff | |
| 308 | by (intro someI2_ex[where Q="\<lambda>f'. x \<bullet> f' y = f x \<bullet> y"]) auto | |
| 309 | qed | |
| 310 | ||
| 311 | lemma adjoint_clauses: | |
| 312 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | |
| 313 | assumes lf: "linear f" | |
| 314 | shows "x \<bullet> adjoint f y = f x \<bullet> y" | |
| 315 | and "adjoint f y \<bullet> x = y \<bullet> f x" | |
| 316 | by (simp_all add: adjoint_works[OF lf] inner_commute) | |
| 317 | ||
| 318 | lemma adjoint_linear: | |
| 319 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | |
| 320 | assumes lf: "linear f" | |
| 321 | shows "linear (adjoint f)" | |
| 322 | by (simp add: lf linear_iff euclidean_eq_iff[where 'a='n] euclidean_eq_iff[where 'a='m] | |
| 323 | adjoint_clauses[OF lf] inner_distrib) | |
| 324 | ||
| 325 | lemma adjoint_adjoint: | |
| 326 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | |
| 327 | assumes lf: "linear f" | |
| 328 | shows "adjoint (adjoint f) = f" | |
| 329 | by (rule adjoint_unique, simp add: adjoint_clauses [OF lf]) | |
| 330 | ||
| 331 | ||
| 67962 | 332 | subsection%unimportant \<open>Interlude: Some properties of real sets\<close> | 
| 63050 | 333 | |
| 334 | lemma seq_mono_lemma: | |
| 335 | assumes "\<forall>(n::nat) \<ge> m. (d n :: real) < e n" | |
| 336 | and "\<forall>n \<ge> m. e n \<le> e m" | |
| 337 | shows "\<forall>n \<ge> m. d n < e m" | |
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 338 | using assms by force | 
| 63050 | 339 | |
| 340 | lemma infinite_enumerate: | |
| 341 | assumes fS: "infinite S" | |
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
66420diff
changeset | 342 | shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (\<forall>n. r n \<in> S)" | 
| 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
66420diff
changeset | 343 | unfolding strict_mono_def | 
| 63050 | 344 | using enumerate_in_set[OF fS] enumerate_mono[of _ _ S] fS by auto | 
| 345 | ||
| 346 | 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)" | |
| 347 | apply auto | |
| 348 | apply (rule_tac x="d/2" in exI) | |
| 349 | apply auto | |
| 350 | done | |
| 351 | ||
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67399diff
changeset | 352 | lemma approachable_lt_le2: \<comment> \<open>like the above, but pushes aside an extra formula\<close> | 
| 63050 | 353 | "(\<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)" | 
| 354 | apply auto | |
| 355 | apply (rule_tac x="d/2" in exI, auto) | |
| 356 | done | |
| 357 | ||
| 358 | lemma triangle_lemma: | |
| 359 | fixes x y z :: real | |
| 360 | assumes x: "0 \<le> x" | |
| 361 | and y: "0 \<le> y" | |
| 362 | and z: "0 \<le> z" | |
| 363 | and xy: "x\<^sup>2 \<le> y\<^sup>2 + z\<^sup>2" | |
| 364 | shows "x \<le> y + z" | |
| 365 | proof - | |
| 366 | have "y\<^sup>2 + z\<^sup>2 \<le> y\<^sup>2 + 2 * y * z + z\<^sup>2" | |
| 367 | using z y by simp | |
| 368 | with xy have th: "x\<^sup>2 \<le> (y + z)\<^sup>2" | |
| 369 | by (simp add: power2_eq_square field_simps) | |
| 370 | from y z have yz: "y + z \<ge> 0" | |
| 371 | by arith | |
| 372 | from power2_le_imp_le[OF th yz] show ?thesis . | |
| 373 | qed | |
| 374 | ||
| 375 | ||
| 376 | ||
| 377 | subsection \<open>Archimedean properties and useful consequences\<close> | |
| 378 | ||
| 379 | text\<open>Bernoulli's inequality\<close> | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68224diff
changeset | 380 | proposition Bernoulli_inequality: | 
| 63050 | 381 | fixes x :: real | 
| 382 | assumes "-1 \<le> x" | |
| 383 | shows "1 + n * x \<le> (1 + x) ^ n" | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68224diff
changeset | 384 | proof (induct n) | 
| 63050 | 385 | case 0 | 
| 386 | then show ?case by simp | |
| 387 | next | |
| 388 | case (Suc n) | |
| 389 | have "1 + Suc n * x \<le> 1 + (Suc n)*x + n * x^2" | |
| 390 | by (simp add: algebra_simps) | |
| 391 | also have "... = (1 + x) * (1 + n*x)" | |
| 392 | by (auto simp: power2_eq_square algebra_simps of_nat_Suc) | |
| 393 | also have "... \<le> (1 + x) ^ Suc n" | |
| 394 | using Suc.hyps assms mult_left_mono by fastforce | |
| 395 | finally show ?case . | |
| 396 | qed | |
| 397 | ||
| 398 | corollary Bernoulli_inequality_even: | |
| 399 | fixes x :: real | |
| 400 | assumes "even n" | |
| 401 | shows "1 + n * x \<le> (1 + x) ^ n" | |
| 402 | proof (cases "-1 \<le> x \<or> n=0") | |
| 403 | case True | |
| 404 | then show ?thesis | |
| 405 | by (auto simp: Bernoulli_inequality) | |
| 406 | next | |
| 407 | case False | |
| 408 | then have "real n \<ge> 1" | |
| 409 | by simp | |
| 410 | with False have "n * x \<le> -1" | |
| 411 | by (metis linear minus_zero mult.commute mult.left_neutral mult_left_mono_neg neg_le_iff_le order_trans zero_le_one) | |
| 412 | then have "1 + n * x \<le> 0" | |
| 413 | by auto | |
| 414 | also have "... \<le> (1 + x) ^ n" | |
| 415 | using assms | |
| 416 | using zero_le_even_power by blast | |
| 417 | finally show ?thesis . | |
| 418 | qed | |
| 419 | ||
| 420 | corollary real_arch_pow: | |
| 421 | fixes x :: real | |
| 422 | assumes x: "1 < x" | |
| 423 | shows "\<exists>n. y < x^n" | |
| 424 | proof - | |
| 425 | from x have x0: "x - 1 > 0" | |
| 426 | by arith | |
| 427 | from reals_Archimedean3[OF x0, rule_format, of y] | |
| 428 | obtain n :: nat where n: "y < real n * (x - 1)" by metis | |
| 429 | from x0 have x00: "x- 1 \<ge> -1" by arith | |
| 430 | from Bernoulli_inequality[OF x00, of n] n | |
| 431 | have "y < x^n" by auto | |
| 432 | then show ?thesis by metis | |
| 433 | qed | |
| 434 | ||
| 435 | corollary real_arch_pow_inv: | |
| 436 | fixes x y :: real | |
| 437 | assumes y: "y > 0" | |
| 438 | and x1: "x < 1" | |
| 439 | shows "\<exists>n. x^n < y" | |
| 440 | proof (cases "x > 0") | |
| 441 | case True | |
| 442 | with x1 have ix: "1 < 1/x" by (simp add: field_simps) | |
| 443 | from real_arch_pow[OF ix, of "1/y"] | |
| 444 | obtain n where n: "1/y < (1/x)^n" by blast | |
| 445 | then show ?thesis using y \<open>x > 0\<close> | |
| 446 | by (auto simp add: field_simps) | |
| 447 | next | |
| 448 | case False | |
| 449 | with y x1 show ?thesis | |
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 450 | by (metis less_le_trans not_less power_one_right) | 
| 63050 | 451 | qed | 
| 452 | ||
| 453 | lemma forall_pos_mono: | |
| 454 | "(\<And>d e::real. d < e \<Longrightarrow> P d \<Longrightarrow> P e) \<Longrightarrow> | |
| 455 | (\<And>n::nat. n \<noteq> 0 \<Longrightarrow> P (inverse (real n))) \<Longrightarrow> (\<And>e. 0 < e \<Longrightarrow> P e)" | |
| 456 | by (metis real_arch_inverse) | |
| 457 | ||
| 458 | lemma forall_pos_mono_1: | |
| 459 | "(\<And>d e::real. d < e \<Longrightarrow> P d \<Longrightarrow> P e) \<Longrightarrow> | |
| 460 | (\<And>n. P (inverse (real (Suc n)))) \<Longrightarrow> 0 < e \<Longrightarrow> P e" | |
| 461 | apply (rule forall_pos_mono) | |
| 462 | apply auto | |
| 463 | apply (metis Suc_pred of_nat_Suc) | |
| 464 | done | |
| 465 | ||
| 466 | ||
| 67962 | 467 | subsection%unimportant \<open>Euclidean Spaces as Typeclass\<close> | 
| 44133 | 468 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50105diff
changeset | 469 | lemma independent_Basis: "independent Basis" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 470 | by (rule independent_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 | 471 | |
| 53939 | 472 | lemma span_Basis [simp]: "span Basis = UNIV" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 473 | by (rule span_Basis) | 
| 44133 | 474 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50105diff
changeset | 475 | 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 | 476 | 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 | 477 | |
| 53406 | 478 | |
| 67962 | 479 | subsection%unimportant \<open>Linearity and Bilinearity continued\<close> | 
| 44133 | 480 | |
| 481 | lemma linear_bounded: | |
| 56444 | 482 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 44133 | 483 | assumes lf: "linear f" | 
| 484 | shows "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x" | |
| 53939 | 485 | proof | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 486 | interpret linear f by fact | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50105diff
changeset | 487 | let ?B = "\<Sum>b\<in>Basis. norm (f b)" | 
| 53939 | 488 | show "\<forall>x. norm (f x) \<le> ?B * norm x" | 
| 489 | proof | |
| 53406 | 490 | 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 | 491 | 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 | 492 | 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 | 493 | unfolding euclidean_representation .. | 
| 64267 | 494 | also have "\<dots> = norm (sum ?g Basis)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 495 | by (simp add: sum scale) | 
| 64267 | 496 | finally have th0: "norm (f x) = norm (sum ?g Basis)" . | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64267diff
changeset | 497 | have th: "norm (?g i) \<le> norm (f i) * norm x" if "i \<in> Basis" for i | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64267diff
changeset | 498 | proof - | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64267diff
changeset | 499 | from Basis_le_norm[OF that, of x] | 
| 53939 | 500 | show "norm (?g i) \<le> norm (f i) * norm x" | 
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 501 | unfolding norm_scaleR by (metis mult.commute mult_left_mono norm_ge_zero) | 
| 53939 | 502 | qed | 
| 64267 | 503 | from sum_norm_le[of _ ?g, OF th] | 
| 53939 | 504 | show "norm (f x) \<le> ?B * norm x" | 
| 64267 | 505 | unfolding th0 sum_distrib_right by metis | 
| 53939 | 506 | qed | 
| 44133 | 507 | qed | 
| 508 | ||
| 509 | lemma linear_conv_bounded_linear: | |
| 510 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | |
| 511 | shows "linear f \<longleftrightarrow> bounded_linear f" | |
| 512 | proof | |
| 513 | assume "linear f" | |
| 53939 | 514 | then interpret f: linear f . | 
| 44133 | 515 | show "bounded_linear f" | 
| 516 | proof | |
| 517 | have "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x" | |
| 60420 | 518 | using \<open>linear f\<close> by (rule linear_bounded) | 
| 49522 | 519 | 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 | 520 | by (simp add: mult.commute) | 
| 44133 | 521 | qed | 
| 522 | next | |
| 523 | assume "bounded_linear f" | |
| 524 | then interpret f: bounded_linear f . | |
| 53939 | 525 | show "linear f" .. | 
| 526 | qed | |
| 527 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61306diff
changeset | 528 | lemmas linear_linear = linear_conv_bounded_linear[symmetric] | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61306diff
changeset | 529 | |
| 53939 | 530 | lemma linear_bounded_pos: | 
| 56444 | 531 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 53939 | 532 | assumes lf: "linear f" | 
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 533 | obtains B where "B > 0" "\<And>x. norm (f x) \<le> B * norm x" | 
| 53939 | 534 | proof - | 
| 535 | have "\<exists>B > 0. \<forall>x. norm (f x) \<le> norm x * B" | |
| 536 | using lf unfolding linear_conv_bounded_linear | |
| 537 | by (rule bounded_linear.pos_bounded) | |
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 538 | with that show ?thesis | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 539 | by (auto simp: mult.commute) | 
| 44133 | 540 | qed | 
| 541 | ||
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 542 | lemma linear_invertible_bounded_below_pos: | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 543 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 544 | assumes "linear f" "linear g" "g \<circ> f = id" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 545 | obtains B where "B > 0" "\<And>x. B * norm x \<le> norm(f x)" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 546 | proof - | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 547 | obtain B where "B > 0" and B: "\<And>x. norm (g x) \<le> B * norm x" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 548 | using linear_bounded_pos [OF \<open>linear g\<close>] by blast | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 549 | show thesis | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 550 | proof | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 551 | show "0 < 1/B" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 552 | by (simp add: \<open>B > 0\<close>) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 553 | show "1/B * norm x \<le> norm (f x)" for x | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 554 | proof - | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 555 | have "1/B * norm x = 1/B * norm (g (f x))" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 556 | using assms by (simp add: pointfree_idE) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 557 | also have "\<dots> \<le> norm (f x)" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 558 | using B [of "f x"] by (simp add: \<open>B > 0\<close> mult.commute pos_divide_le_eq) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 559 | finally show ?thesis . | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 560 | qed | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 561 | qed | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 562 | qed | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 563 | |
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 564 | lemma linear_inj_bounded_below_pos: | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 565 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 566 | assumes "linear f" "inj f" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 567 | obtains B where "B > 0" "\<And>x. B * norm x \<le> norm(f x)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 568 | using linear_injective_left_inverse [OF assms] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 569 | linear_invertible_bounded_below_pos assms by blast | 
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 570 | |
| 49522 | 571 | lemma bounded_linearI': | 
| 56444 | 572 | fixes f ::"'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 53406 | 573 | assumes "\<And>x y. f (x + y) = f x + f y" | 
| 574 | and "\<And>c x. f (c *\<^sub>R x) = c *\<^sub>R f x" | |
| 49522 | 575 | shows "bounded_linear f" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 576 | using assms linearI linear_conv_bounded_linear by blast | 
| 44133 | 577 | |
| 578 | lemma bilinear_bounded: | |
| 56444 | 579 | fixes h :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'k::real_normed_vector" | 
| 44133 | 580 | assumes bh: "bilinear h" | 
| 581 | 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 | 582 | proof (clarify intro!: exI[of _ "\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. norm (h i j)"]) | 
| 53406 | 583 | fix x :: 'm | 
| 584 | fix y :: 'n | |
| 64267 | 585 | have "norm (h x y) = norm (h (sum (\<lambda>i. (x \<bullet> i) *\<^sub>R i) Basis) (sum (\<lambda>i. (y \<bullet> i) *\<^sub>R i) Basis))" | 
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 586 | by (simp add: euclidean_representation) | 
| 64267 | 587 | also have "\<dots> = norm (sum (\<lambda> (i,j). h ((x \<bullet> i) *\<^sub>R i) ((y \<bullet> j) *\<^sub>R j)) (Basis \<times> Basis))" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 588 | unfolding bilinear_sum[OF bh] .. | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50105diff
changeset | 589 | finally have th: "norm (h x y) = \<dots>" . | 
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 590 | have "\<And>i j. \<lbrakk>i \<in> Basis; j \<in> Basis\<rbrakk> | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 591 | \<Longrightarrow> \<bar>x \<bullet> i\<bar> * (\<bar>y \<bullet> j\<bar> * norm (h i j)) \<le> norm x * (norm y * norm (h i j))" | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 592 | by (auto simp add: zero_le_mult_iff Basis_le_norm mult_mono) | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 593 | then show "norm (h x y) \<le> (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. norm (h i j)) * norm x * norm y" | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 594 | unfolding sum_distrib_right th sum.cartesian_product | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 595 | by (clarsimp simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 596 | field_simps simp del: scaleR_scaleR intro!: sum_norm_le) | 
| 44133 | 597 | qed | 
| 598 | ||
| 599 | lemma bilinear_conv_bounded_bilinear: | |
| 600 | fixes h :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::real_normed_vector" | |
| 601 | shows "bilinear h \<longleftrightarrow> bounded_bilinear h" | |
| 602 | proof | |
| 603 | assume "bilinear h" | |
| 604 | show "bounded_bilinear h" | |
| 605 | proof | |
| 53406 | 606 | fix x y z | 
| 607 | show "h (x + y) z = h x z + h y z" | |
| 60420 | 608 | using \<open>bilinear h\<close> unfolding bilinear_def linear_iff by simp | 
| 44133 | 609 | next | 
| 53406 | 610 | fix x y z | 
| 611 | show "h x (y + z) = h x y + h x z" | |
| 60420 | 612 | using \<open>bilinear h\<close> unfolding bilinear_def linear_iff by simp | 
| 44133 | 613 | next | 
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 614 | show "h (scaleR r x) y = scaleR r (h x y)" "h x (scaleR r y) = scaleR r (h x y)" for r x y | 
| 60420 | 615 | using \<open>bilinear h\<close> unfolding bilinear_def linear_iff | 
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 616 | by simp_all | 
| 44133 | 617 | next | 
| 618 | have "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y" | |
| 60420 | 619 | using \<open>bilinear h\<close> by (rule bilinear_bounded) | 
| 49522 | 620 | 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 | 621 | by (simp add: ac_simps) | 
| 44133 | 622 | qed | 
| 623 | next | |
| 624 | assume "bounded_bilinear h" | |
| 625 | then interpret h: bounded_bilinear h . | |
| 626 | show "bilinear h" | |
| 627 | unfolding bilinear_def linear_conv_bounded_linear | |
| 49522 | 628 | using h.bounded_linear_left h.bounded_linear_right by simp | 
| 44133 | 629 | qed | 
| 630 | ||
| 53939 | 631 | lemma bilinear_bounded_pos: | 
| 56444 | 632 | fixes h :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::real_normed_vector" | 
| 53939 | 633 | assumes bh: "bilinear h" | 
| 634 | shows "\<exists>B > 0. \<forall>x y. norm (h x y) \<le> B * norm x * norm y" | |
| 635 | proof - | |
| 636 | have "\<exists>B > 0. \<forall>x y. norm (h x y) \<le> norm x * norm y * B" | |
| 637 | using bh [unfolded bilinear_conv_bounded_bilinear] | |
| 638 | by (rule bounded_bilinear.pos_bounded) | |
| 639 | then show ?thesis | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 640 | by (simp only: ac_simps) | 
| 53939 | 641 | qed | 
| 642 | ||
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 643 | lemma bounded_linear_imp_has_derivative: "bounded_linear f \<Longrightarrow> (f has_derivative f) net" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 644 | by (auto simp add: has_derivative_def linear_diff linear_linear linear_def | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 645 | dest: bounded_linear.linear) | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 646 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 647 | lemma linear_imp_has_derivative: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 648 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 649 | shows "linear f \<Longrightarrow> (f has_derivative f) net" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 650 | by (simp add: bounded_linear_imp_has_derivative linear_conv_bounded_linear) | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 651 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 652 | lemma bounded_linear_imp_differentiable: "bounded_linear f \<Longrightarrow> f differentiable net" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 653 | using bounded_linear_imp_has_derivative differentiable_def by blast | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 654 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 655 | lemma linear_imp_differentiable: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 656 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 657 | shows "linear f \<Longrightarrow> f differentiable net" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 658 | by (metis linear_imp_has_derivative differentiable_def) | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 659 | |
| 49522 | 660 | |
| 68901 | 661 | subsection%unimportant \<open>We continue\<close> | 
| 44133 | 662 | |
| 663 | lemma independent_bound: | |
| 53716 | 664 | fixes S :: "'a::euclidean_space set" | 
| 665 |   shows "independent S \<Longrightarrow> finite S \<and> card S \<le> DIM('a)"
 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 666 | by (metis dim_subset_UNIV finiteI_independent dim_span_eq_card_independent) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 667 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 668 | lemmas independent_imp_finite = finiteI_independent | 
| 44133 | 669 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 670 | corollary | 
| 60303 | 671 | fixes S :: "'a::euclidean_space set" | 
| 672 | assumes "independent S" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 673 |   shows independent_card_le:"card S \<le> DIM('a)"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 674 | using assms independent_bound by auto | 
| 63075 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 675 | |
| 49663 | 676 | lemma dependent_biggerset: | 
| 56444 | 677 | fixes S :: "'a::euclidean_space set" | 
| 678 |   shows "(finite S \<Longrightarrow> card S > DIM('a)) \<Longrightarrow> dependent S"
 | |
| 44133 | 679 | by (metis independent_bound not_less) | 
| 680 | ||
| 60420 | 681 | text \<open>Picking an orthogonal replacement for a spanning set.\<close> | 
| 44133 | 682 | |
| 53406 | 683 | lemma vector_sub_project_orthogonal: | 
| 684 | fixes b x :: "'a::euclidean_space" | |
| 685 | shows "b \<bullet> (x - ((b \<bullet> x) / (b \<bullet> b)) *\<^sub>R b) = 0" | |
| 44133 | 686 | unfolding inner_simps by auto | 
| 687 | ||
| 44528 | 688 | lemma pairwise_orthogonal_insert: | 
| 689 | assumes "pairwise orthogonal S" | |
| 49522 | 690 | and "\<And>y. y \<in> S \<Longrightarrow> orthogonal x y" | 
| 44528 | 691 | shows "pairwise orthogonal (insert x S)" | 
| 692 | using assms unfolding pairwise_def | |
| 693 | by (auto simp add: orthogonal_commute) | |
| 694 | ||
| 44133 | 695 | lemma basis_orthogonal: | 
| 53406 | 696 | fixes B :: "'a::real_inner set" | 
| 44133 | 697 | assumes fB: "finite B" | 
| 698 | shows "\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C" | |
| 699 | (is " \<exists>C. ?P B C") | |
| 49522 | 700 | using fB | 
| 701 | proof (induct rule: finite_induct) | |
| 702 | case empty | |
| 53406 | 703 | then show ?case | 
| 704 |     apply (rule exI[where x="{}"])
 | |
| 705 | apply (auto simp add: pairwise_def) | |
| 706 | done | |
| 44133 | 707 | next | 
| 49522 | 708 | case (insert a B) | 
| 60420 | 709 | note fB = \<open>finite B\<close> and aB = \<open>a \<notin> B\<close> | 
| 710 | from \<open>\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C\<close> | |
| 44133 | 711 | obtain C where C: "finite C" "card C \<le> card B" | 
| 712 | "span C = span B" "pairwise orthogonal C" by blast | |
| 64267 | 713 | let ?a = "a - sum (\<lambda>x. (x \<bullet> a / (x \<bullet> x)) *\<^sub>R x) C" | 
| 44133 | 714 | let ?C = "insert ?a C" | 
| 53406 | 715 | from C(1) have fC: "finite ?C" | 
| 716 | by simp | |
| 49522 | 717 | from fB aB C(1,2) have cC: "card ?C \<le> card (insert a B)" | 
| 718 | by (simp add: card_insert_if) | |
| 53406 | 719 |   {
 | 
| 720 | fix x k | |
| 49522 | 721 | have th0: "\<And>(a::'a) b c. a - (b - c) = c + (a - b)" | 
| 722 | by (simp add: field_simps) | |
| 44133 | 723 | 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" | 
| 724 | apply (simp only: scaleR_right_diff_distrib th0) | |
| 725 | apply (rule span_add_eq) | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 726 | apply (rule span_scale) | 
| 64267 | 727 | apply (rule span_sum) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 728 | apply (rule span_scale) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 729 | apply (rule span_base) | 
| 49522 | 730 | apply assumption | 
| 53406 | 731 | done | 
| 732 | } | |
| 44133 | 733 | then have SC: "span ?C = span (insert a B)" | 
| 734 | unfolding set_eq_iff span_breakdown_eq C(3)[symmetric] by auto | |
| 53406 | 735 |   {
 | 
| 736 | fix y | |
| 737 | assume yC: "y \<in> C" | |
| 738 |     then have Cy: "C = insert y (C - {y})"
 | |
| 739 | by blast | |
| 740 |     have fth: "finite (C - {y})"
 | |
| 741 | using C by simp | |
| 44528 | 742 | have "orthogonal ?a y" | 
| 743 | unfolding orthogonal_def | |
| 64267 | 744 | unfolding inner_diff inner_sum_left right_minus_eq | 
| 745 | unfolding sum.remove [OF \<open>finite C\<close> \<open>y \<in> C\<close>] | |
| 44528 | 746 | apply (clarsimp simp add: inner_commute[of y a]) | 
| 64267 | 747 | apply (rule sum.neutral) | 
| 44528 | 748 | apply clarsimp | 
| 749 | apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format]) | |
| 60420 | 750 | using \<open>y \<in> C\<close> by auto | 
| 53406 | 751 | } | 
| 60420 | 752 | with \<open>pairwise orthogonal C\<close> have CPO: "pairwise orthogonal ?C" | 
| 44528 | 753 | by (rule pairwise_orthogonal_insert) | 
| 53406 | 754 | from fC cC SC CPO have "?P (insert a B) ?C" | 
| 755 | by blast | |
| 44133 | 756 | then show ?case by blast | 
| 757 | qed | |
| 758 | ||
| 759 | lemma orthogonal_basis_exists: | |
| 760 |   fixes V :: "('a::euclidean_space) set"
 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 761 | shows "\<exists>B. independent B \<and> B \<subseteq> span V \<and> V \<subseteq> span B \<and> | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 762 | (card B = dim V) \<and> pairwise orthogonal B" | 
| 49663 | 763 | proof - | 
| 49522 | 764 | from basis_exists[of V] obtain B where | 
| 53406 | 765 | B: "B \<subseteq> V" "independent B" "V \<subseteq> span B" "card B = dim V" | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immlerdiff
changeset | 766 | by force | 
| 53406 | 767 | from B have fB: "finite B" "card B = dim V" | 
| 768 | using independent_bound by auto | |
| 44133 | 769 | from basis_orthogonal[OF fB(1)] obtain C where | 
| 53406 | 770 | C: "finite C" "card C \<le> card B" "span C = span B" "pairwise orthogonal C" | 
| 771 | by blast | |
| 772 | from C B have CSV: "C \<subseteq> span V" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 773 | by (metis span_superset span_mono subset_trans) | 
| 53406 | 774 | from span_mono[OF B(3)] C have SVC: "span V \<subseteq> span C" | 
| 775 | by (simp add: span_span) | |
| 44133 | 776 | from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB | 
| 53406 | 777 | have iC: "independent C" | 
| 44133 | 778 | by (simp add: dim_span) | 
| 53406 | 779 | from C fB have "card C \<le> dim V" | 
| 780 | by simp | |
| 781 | moreover have "dim V \<le> card C" | |
| 782 | using span_card_ge_dim[OF CSV SVC C(1)] | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 783 | by simp | 
| 53406 | 784 | ultimately have CdV: "card C = dim V" | 
| 785 | using C(1) by simp | |
| 786 | from C B CSV CdV iC show ?thesis | |
| 787 | by auto | |
| 44133 | 788 | qed | 
| 789 | ||
| 60420 | 790 | text \<open>Low-dimensional subset is in a hyperplane (weak orthogonal complement).\<close> | 
| 44133 | 791 | |
| 49522 | 792 | lemma span_not_univ_orthogonal: | 
| 53406 | 793 | fixes S :: "'a::euclidean_space set" | 
| 44133 | 794 | assumes sU: "span S \<noteq> UNIV" | 
| 56444 | 795 | shows "\<exists>a::'a. a \<noteq> 0 \<and> (\<forall>x \<in> span S. a \<bullet> x = 0)" | 
| 49522 | 796 | proof - | 
| 53406 | 797 | from sU obtain a where a: "a \<notin> span S" | 
| 798 | by blast | |
| 44133 | 799 | from orthogonal_basis_exists obtain B where | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 800 | B: "independent B" "B \<subseteq> span S" "S \<subseteq> span B" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 801 | "card B = dim S" "pairwise orthogonal B" | 
| 44133 | 802 | by blast | 
| 53406 | 803 | from B have fB: "finite B" "card B = dim S" | 
| 804 | using independent_bound by auto | |
| 44133 | 805 | from span_mono[OF B(2)] span_mono[OF B(3)] | 
| 53406 | 806 | have sSB: "span S = span B" | 
| 807 | by (simp add: span_span) | |
| 64267 | 808 | let ?a = "a - sum (\<lambda>b. (a \<bullet> b / (b \<bullet> b)) *\<^sub>R b) B" | 
| 809 | have "sum (\<lambda>b. (a \<bullet> b / (b \<bullet> b)) *\<^sub>R b) B \<in> span S" | |
| 44133 | 810 | unfolding sSB | 
| 64267 | 811 | apply (rule span_sum) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 812 | apply (rule span_scale) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 813 | apply (rule span_base) | 
| 49522 | 814 | apply assumption | 
| 815 | done | |
| 53406 | 816 | with a have a0:"?a \<noteq> 0" | 
| 817 | by auto | |
| 68058 | 818 | have "?a \<bullet> x = 0" if "x\<in>span B" for x | 
| 819 | proof (rule span_induct [OF that]) | |
| 49522 | 820 |     show "subspace {x. ?a \<bullet> x = 0}"
 | 
| 821 | by (auto simp add: subspace_def inner_add) | |
| 822 | next | |
| 53406 | 823 |     {
 | 
| 824 | fix x | |
| 825 | assume x: "x \<in> B" | |
| 826 |       from x have B': "B = insert x (B - {x})"
 | |
| 827 | by blast | |
| 828 |       have fth: "finite (B - {x})"
 | |
| 829 | using fB by simp | |
| 44133 | 830 | have "?a \<bullet> x = 0" | 
| 53406 | 831 | apply (subst B') | 
| 832 | using fB fth | |
| 64267 | 833 | unfolding sum_clauses(2)[OF fth] | 
| 44133 | 834 | apply simp unfolding inner_simps | 
| 64267 | 835 | apply (clarsimp simp add: inner_add inner_sum_left) | 
| 836 | apply (rule sum.neutral, rule ballI) | |
| 63170 | 837 | apply (simp only: inner_commute) | 
| 49711 | 838 | apply (auto simp add: x field_simps | 
| 839 | intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format]) | |
| 53406 | 840 | done | 
| 841 | } | |
| 68058 | 842 | then show "?a \<bullet> x = 0" if "x \<in> B" for x | 
| 843 | using that by blast | |
| 844 | qed | |
| 53406 | 845 | with a0 show ?thesis | 
| 846 | unfolding sSB by (auto intro: exI[where x="?a"]) | |
| 44133 | 847 | qed | 
| 848 | ||
| 849 | lemma span_not_univ_subset_hyperplane: | |
| 53406 | 850 | fixes S :: "'a::euclidean_space set" | 
| 851 | assumes SU: "span S \<noteq> UNIV" | |
| 44133 | 852 |   shows "\<exists> a. a \<noteq>0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
 | 
| 853 | using span_not_univ_orthogonal[OF SU] by auto | |
| 854 | ||
| 49663 | 855 | lemma lowdim_subset_hyperplane: | 
| 53406 | 856 | fixes S :: "'a::euclidean_space set" | 
| 44133 | 857 |   assumes d: "dim S < DIM('a)"
 | 
| 56444 | 858 |   shows "\<exists>a::'a. a \<noteq> 0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
 | 
| 49522 | 859 | proof - | 
| 53406 | 860 |   {
 | 
| 861 | assume "span S = UNIV" | |
| 862 |     then have "dim (span S) = dim (UNIV :: ('a) set)"
 | |
| 863 | by simp | |
| 864 |     then have "dim S = DIM('a)"
 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 865 | by (metis Euclidean_Space.dim_UNIV dim_span) | 
| 53406 | 866 | with d have False by arith | 
| 867 | } | |
| 868 | then have th: "span S \<noteq> UNIV" | |
| 869 | by blast | |
| 44133 | 870 | from span_not_univ_subset_hyperplane[OF th] show ?thesis . | 
| 871 | qed | |
| 872 | ||
| 873 | lemma linear_eq_stdbasis: | |
| 56444 | 874 | fixes f :: "'a::euclidean_space \<Rightarrow> _" | 
| 875 | assumes lf: "linear f" | |
| 49663 | 876 | and lg: "linear g" | 
| 68058 | 877 | and fg: "\<And>b. b \<in> Basis \<Longrightarrow> f b = g b" | 
| 44133 | 878 | shows "f = g" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 879 | using linear_eq_on_span[OF lf lg, of Basis] fg | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 880 | by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 881 | |
| 44133 | 882 | |
| 60420 | 883 | text \<open>Similar results for bilinear functions.\<close> | 
| 44133 | 884 | |
| 885 | lemma bilinear_eq: | |
| 886 | assumes bf: "bilinear f" | |
| 49522 | 887 | and bg: "bilinear g" | 
| 53406 | 888 | and SB: "S \<subseteq> span B" | 
| 889 | and TC: "T \<subseteq> span C" | |
| 68058 | 890 | and "x\<in>S" "y\<in>T" | 
| 891 | and fg: "\<And>x y. \<lbrakk>x \<in> B; y\<in> C\<rbrakk> \<Longrightarrow> f x y = g x y" | |
| 892 | shows "f x y = g x y" | |
| 49663 | 893 | proof - | 
| 44170 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 huffman parents: 
44166diff
changeset | 894 |   let ?P = "{x. \<forall>y\<in> span C. f x y = g x y}"
 | 
| 44133 | 895 | from bf bg have sp: "subspace ?P" | 
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53596diff
changeset | 896 | unfolding bilinear_def linear_iff subspace_def bf bg | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 897 | by (auto simp add: span_zero bilinear_lzero[OF bf] bilinear_lzero[OF bg] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 898 | span_add Ball_def | 
| 49663 | 899 | intro: bilinear_ladd[OF bf]) | 
| 68058 | 900 |   have sfg: "\<And>x. x \<in> B \<Longrightarrow> subspace {a. f x a = g x a}"
 | 
| 44133 | 901 | apply (auto simp add: subspace_def) | 
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53596diff
changeset | 902 | using bf bg unfolding bilinear_def linear_iff | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 903 | apply (auto simp add: span_zero bilinear_rzero[OF bf] bilinear_rzero[OF bg] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 904 | span_add Ball_def | 
| 49663 | 905 | intro: bilinear_ladd[OF bf]) | 
| 49522 | 906 | done | 
| 68058 | 907 | have "\<forall>y\<in> span C. f x y = g x y" if "x \<in> span B" for x | 
| 908 | apply (rule span_induct [OF that sp]) | |
| 68062 | 909 | using fg sfg span_induct by blast | 
| 53406 | 910 | then show ?thesis | 
| 68058 | 911 | using SB TC assms by auto | 
| 44133 | 912 | qed | 
| 913 | ||
| 49522 | 914 | lemma bilinear_eq_stdbasis: | 
| 53406 | 915 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> _" | 
| 44133 | 916 | assumes bf: "bilinear f" | 
| 49522 | 917 | and bg: "bilinear g" | 
| 68058 | 918 | and fg: "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> f i j = g i j" | 
| 44133 | 919 | shows "f = g" | 
| 68074 | 920 | using bilinear_eq[OF bf bg equalityD2[OF span_Basis] equalityD2[OF span_Basis]] fg by blast | 
| 49522 | 921 | |
| 60420 | 922 | subsection \<open>Infinity norm\<close> | 
| 44133 | 923 | |
| 67962 | 924 | definition%important "infnorm (x::'a::euclidean_space) = Sup {\<bar>x \<bullet> b\<bar> |b. b \<in> Basis}"
 | 
| 44133 | 925 | |
| 926 | lemma infnorm_set_image: | |
| 53716 | 927 | fixes x :: "'a::euclidean_space" | 
| 56444 | 928 |   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 | 929 | by blast | 
| 44133 | 930 | |
| 53716 | 931 | lemma infnorm_Max: | 
| 932 | fixes x :: "'a::euclidean_space" | |
| 56444 | 933 | shows "infnorm x = Max ((\<lambda>i. \<bar>x \<bullet> i\<bar>) ` Basis)" | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61973diff
changeset | 934 | by (simp add: infnorm_def infnorm_set_image cSup_eq_Max) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 935 | |
| 44133 | 936 | lemma infnorm_set_lemma: | 
| 53716 | 937 | fixes x :: "'a::euclidean_space" | 
| 56444 | 938 |   shows "finite {\<bar>x \<bullet> i\<bar> |i. i \<in> Basis}"
 | 
| 939 |     and "{\<bar>x \<bullet> i\<bar> |i. i \<in> Basis} \<noteq> {}"
 | |
| 44133 | 940 | unfolding infnorm_set_image | 
| 941 | by auto | |
| 942 | ||
| 53406 | 943 | lemma infnorm_pos_le: | 
| 944 | fixes x :: "'a::euclidean_space" | |
| 945 | 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 | 946 | by (simp add: infnorm_Max Max_ge_iff ex_in_conv) | 
| 44133 | 947 | |
| 53406 | 948 | lemma infnorm_triangle: | 
| 949 | fixes x :: "'a::euclidean_space" | |
| 950 | shows "infnorm (x + y) \<le> infnorm x + infnorm y" | |
| 49522 | 951 | proof - | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 952 | 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 | 953 | by simp | 
| 44133 | 954 | show ?thesis | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 955 | by (auto simp: infnorm_Max inner_add_left intro!: *) | 
| 44133 | 956 | qed | 
| 957 | ||
| 53406 | 958 | lemma infnorm_eq_0: | 
| 959 | fixes x :: "'a::euclidean_space" | |
| 960 | shows "infnorm x = 0 \<longleftrightarrow> x = 0" | |
| 49522 | 961 | proof - | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 962 | 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 | 963 | 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 | 964 | then show ?thesis | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 965 | using infnorm_pos_le[of x] by simp | 
| 44133 | 966 | qed | 
| 967 | ||
| 968 | lemma infnorm_0: "infnorm 0 = 0" | |
| 969 | by (simp add: infnorm_eq_0) | |
| 970 | ||
| 971 | lemma infnorm_neg: "infnorm (- x) = infnorm x" | |
| 68062 | 972 | unfolding infnorm_def by simp | 
| 44133 | 973 | |
| 974 | lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)" | |
| 68062 | 975 | by (metis infnorm_neg minus_diff_eq) | 
| 976 | ||
| 977 | lemma absdiff_infnorm: "\<bar>infnorm x - infnorm y\<bar> \<le> infnorm (x - y)" | |
| 49522 | 978 | proof - | 
| 68062 | 979 | have *: "\<And>(nx::real) n ny. nx \<le> n + ny \<Longrightarrow> ny \<le> n + nx \<Longrightarrow> \<bar>nx - ny\<bar> \<le> n" | 
| 44133 | 980 | by arith | 
| 68062 | 981 | show ?thesis | 
| 982 | proof (rule *) | |
| 983 | from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"] | |
| 984 | show "infnorm x \<le> infnorm (x - y) + infnorm y" "infnorm y \<le> infnorm (x - y) + infnorm x" | |
| 985 | by (simp_all add: field_simps infnorm_neg) | |
| 986 | qed | |
| 44133 | 987 | qed | 
| 988 | ||
| 53406 | 989 | lemma real_abs_infnorm: "\<bar>infnorm x\<bar> = infnorm x" | 
| 44133 | 990 | using infnorm_pos_le[of x] by arith | 
| 991 | ||
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50105diff
changeset | 992 | lemma Basis_le_infnorm: | 
| 53406 | 993 | fixes x :: "'a::euclidean_space" | 
| 994 | 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 | 995 | by (simp add: infnorm_Max) | 
| 44133 | 996 | |
| 56444 | 997 | 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 | 998 | unfolding infnorm_Max | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 999 | 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 | 1000 | let ?B = "(\<lambda>i. \<bar>x \<bullet> i\<bar>) ` Basis" | 
| 68062 | 1001 |   { fix b :: 'a
 | 
| 53406 | 1002 | assume "b \<in> Basis" | 
| 1003 | then show "\<bar>a *\<^sub>R x \<bullet> b\<bar> \<le> \<bar>a\<bar> * Max ?B" | |
| 1004 | by (simp add: abs_mult mult_left_mono) | |
| 1005 | next | |
| 1006 | from Max_in[of ?B] obtain b where "b \<in> Basis" "Max ?B = \<bar>x \<bullet> b\<bar>" | |
| 1007 | by (auto simp del: Max_in) | |
| 1008 | 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" | |
| 1009 | by (intro image_eqI[where x=b]) (auto simp: abs_mult) | |
| 1010 | } | |
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 1011 | qed simp | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 1012 | |
| 53406 | 1013 | 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 | 1014 | unfolding infnorm_mul .. | 
| 44133 | 1015 | |
| 1016 | lemma infnorm_pos_lt: "infnorm x > 0 \<longleftrightarrow> x \<noteq> 0" | |
| 1017 | using infnorm_pos_le[of x] infnorm_eq_0[of x] by arith | |
| 1018 | ||
| 60420 | 1019 | text \<open>Prove that it differs only up to a bound from Euclidean norm.\<close> | 
| 44133 | 1020 | |
| 1021 | 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 | 1022 | 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 | 1023 | |
| 53716 | 1024 | lemma norm_le_infnorm: | 
| 1025 | fixes x :: "'a::euclidean_space" | |
| 1026 |   shows "norm x \<le> sqrt DIM('a) * infnorm x"
 | |
| 68062 | 1027 | unfolding norm_eq_sqrt_inner id_def | 
| 1028 | proof (rule real_le_lsqrt[OF inner_ge_zero]) | |
| 1029 |   show "sqrt DIM('a) * infnorm x \<ge> 0"
 | |
| 44133 | 1030 | by (simp add: zero_le_mult_iff infnorm_pos_le) | 
| 68062 | 1031 | have "x \<bullet> x \<le> (\<Sum>b\<in>Basis. x \<bullet> b * (x \<bullet> b))" | 
| 1032 | by (metis euclidean_inner order_refl) | |
| 1033 |   also have "... \<le> DIM('a) * \<bar>infnorm x\<bar>\<^sup>2"
 | |
| 1034 | by (rule sum_bounded_above) (metis Basis_le_infnorm abs_le_square_iff power2_eq_square real_abs_infnorm) | |
| 1035 |   also have "... \<le> (sqrt DIM('a) * infnorm x)\<^sup>2"
 | |
| 1036 | by (simp add: power_mult_distrib) | |
| 1037 |   finally show "x \<bullet> x \<le> (sqrt DIM('a) * infnorm x)\<^sup>2" .
 | |
| 44133 | 1038 | qed | 
| 1039 | ||
| 44646 | 1040 | lemma tendsto_infnorm [tendsto_intros]: | 
| 61973 | 1041 | assumes "(f \<longlongrightarrow> a) F" | 
| 1042 | shows "((\<lambda>x. infnorm (f x)) \<longlongrightarrow> infnorm a) F" | |
| 44646 | 1043 | proof (rule tendsto_compose [OF LIM_I assms]) | 
| 53406 | 1044 | fix r :: real | 
| 1045 | assume "r > 0" | |
| 49522 | 1046 | then show "\<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (infnorm x - infnorm a) < r" | 
| 68062 | 1047 | by (metis real_norm_def le_less_trans absdiff_infnorm infnorm_le_norm) | 
| 44646 | 1048 | qed | 
| 1049 | ||
| 60420 | 1050 | text \<open>Equality in Cauchy-Schwarz and triangle inequalities.\<close> | 
| 44133 | 1051 | |
| 53406 | 1052 | lemma norm_cauchy_schwarz_eq: "x \<bullet> y = norm x * norm y \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x" | 
| 1053 | (is "?lhs \<longleftrightarrow> ?rhs") | |
| 68062 | 1054 | proof (cases "x=0") | 
| 1055 | case True | |
| 1056 | then show ?thesis | |
| 1057 | by auto | |
| 1058 | next | |
| 1059 | case False | |
| 1060 | from inner_eq_zero_iff[of "norm y *\<^sub>R x - norm x *\<^sub>R y"] | |
| 1061 | have "?rhs \<longleftrightarrow> | |
| 49522 | 1062 | (norm y * (norm y * norm x * norm x - norm x * (x \<bullet> y)) - | 
| 1063 | norm x * (norm y * (y \<bullet> x) - norm x * norm y * norm y) = 0)" | |
| 68062 | 1064 | using False unfolding inner_simps | 
| 1065 | by (auto simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps) | |
| 1066 | also have "\<dots> \<longleftrightarrow> (2 * norm x * norm y * (norm x * norm y - x \<bullet> y) = 0)" | |
| 1067 | using False by (simp add: field_simps inner_commute) | |
| 1068 | also have "\<dots> \<longleftrightarrow> ?lhs" | |
| 1069 | using False by auto | |
| 1070 | finally show ?thesis by metis | |
| 44133 | 1071 | qed | 
| 1072 | ||
| 1073 | lemma norm_cauchy_schwarz_abs_eq: | |
| 56444 | 1074 | "\<bar>x \<bullet> y\<bar> = norm x * norm y \<longleftrightarrow> | 
| 53716 | 1075 | norm x *\<^sub>R y = norm y *\<^sub>R x \<or> norm x *\<^sub>R y = - norm y *\<^sub>R x" | 
| 53406 | 1076 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 49522 | 1077 | proof - | 
| 56444 | 1078 | have th: "\<And>(x::real) a. a \<ge> 0 \<Longrightarrow> \<bar>x\<bar> = a \<longleftrightarrow> x = a \<or> x = - a" | 
| 53406 | 1079 | by arith | 
| 44133 | 1080 | have "?rhs \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x \<or> norm (- x) *\<^sub>R y = norm y *\<^sub>R (- x)" | 
| 1081 | by simp | |
| 68062 | 1082 | also have "\<dots> \<longleftrightarrow> (x \<bullet> y = norm x * norm y \<or> (- x) \<bullet> y = norm x * norm y)" | 
| 44133 | 1083 | unfolding norm_cauchy_schwarz_eq[symmetric] | 
| 1084 | unfolding norm_minus_cancel norm_scaleR .. | |
| 1085 | also have "\<dots> \<longleftrightarrow> ?lhs" | |
| 53406 | 1086 | unfolding th[OF mult_nonneg_nonneg, OF norm_ge_zero[of x] norm_ge_zero[of y]] inner_simps | 
| 1087 | by auto | |
| 44133 | 1088 | finally show ?thesis .. | 
| 1089 | qed | |
| 1090 | ||
| 1091 | lemma norm_triangle_eq: | |
| 1092 | fixes x y :: "'a::real_inner" | |
| 53406 | 1093 | shows "norm (x + y) = norm x + norm y \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x" | 
| 68062 | 1094 | proof (cases "x = 0 \<or> y = 0") | 
| 1095 | case True | |
| 1096 | then show ?thesis | |
| 1097 | by force | |
| 1098 | next | |
| 1099 | case False | |
| 1100 | then have n: "norm x > 0" "norm y > 0" | |
| 1101 | by auto | |
| 1102 | have "norm (x + y) = norm x + norm y \<longleftrightarrow> (norm (x + y))\<^sup>2 = (norm x + norm y)\<^sup>2" | |
| 1103 | by simp | |
| 1104 | also have "\<dots> \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x" | |
| 1105 | unfolding norm_cauchy_schwarz_eq[symmetric] | |
| 1106 | unfolding power2_norm_eq_inner inner_simps | |
| 1107 | by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps) | |
| 1108 | finally show ?thesis . | |
| 44133 | 1109 | qed | 
| 1110 | ||
| 49522 | 1111 | |
| 60420 | 1112 | subsection \<open>Collinearity\<close> | 
| 44133 | 1113 | |
| 67962 | 1114 | definition%important collinear :: "'a::real_vector set \<Rightarrow> bool" | 
| 49522 | 1115 | where "collinear S \<longleftrightarrow> (\<exists>u. \<forall>x \<in> S. \<forall> y \<in> S. \<exists>c. x - y = c *\<^sub>R u)" | 
| 44133 | 1116 | |
| 66287 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1117 | lemma collinear_alt: | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1118 | "collinear S \<longleftrightarrow> (\<exists>u v. \<forall>x \<in> S. \<exists>c. x = u + c *\<^sub>R v)" (is "?lhs = ?rhs") | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1119 | proof | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1120 | assume ?lhs | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1121 | then show ?rhs | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1122 | unfolding collinear_def by (metis Groups.add_ac(2) diff_add_cancel) | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1123 | next | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1124 | assume ?rhs | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1125 | then obtain u v where *: "\<And>x. x \<in> S \<Longrightarrow> \<exists>c. x = u + c *\<^sub>R v" | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1126 | by (auto simp: ) | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1127 | have "\<exists>c. x - y = c *\<^sub>R v" if "x \<in> S" "y \<in> S" for x y | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1128 | by (metis *[OF \<open>x \<in> S\<close>] *[OF \<open>y \<in> S\<close>] scaleR_left.diff add_diff_cancel_left) | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1129 | then show ?lhs | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1130 | using collinear_def by blast | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1131 | qed | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1132 | |
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1133 | lemma collinear: | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1134 |   fixes S :: "'a::{perfect_space,real_vector} set"
 | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1135 | shows "collinear S \<longleftrightarrow> (\<exists>u. u \<noteq> 0 \<and> (\<forall>x \<in> S. \<forall> y \<in> S. \<exists>c. x - y = c *\<^sub>R u))" | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1136 | proof - | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1137 | have "\<exists>v. v \<noteq> 0 \<and> (\<forall>x\<in>S. \<forall>y\<in>S. \<exists>c. x - y = c *\<^sub>R v)" | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1138 | if "\<forall>x\<in>S. \<forall>y\<in>S. \<exists>c. x - y = c *\<^sub>R u" "u=0" for u | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1139 | proof - | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1140 | have "\<forall>x\<in>S. \<forall>y\<in>S. x = y" | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1141 | using that by auto | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1142 | moreover | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1143 | obtain v::'a where "v \<noteq> 0" | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1144 | using UNIV_not_singleton [of 0] by auto | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1145 | ultimately have "\<forall>x\<in>S. \<forall>y\<in>S. \<exists>c. x - y = c *\<^sub>R v" | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1146 | by auto | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1147 | then show ?thesis | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1148 | using \<open>v \<noteq> 0\<close> by blast | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1149 | qed | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1150 | then show ?thesis | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1151 | apply (clarsimp simp: collinear_def) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 1152 | by (metis scaleR_zero_right vector_fraction_eq_iff) | 
| 66287 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1153 | qed | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1154 | |
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63680diff
changeset | 1155 | lemma collinear_subset: "\<lbrakk>collinear T; S \<subseteq> T\<rbrakk> \<Longrightarrow> collinear S" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63680diff
changeset | 1156 | by (meson collinear_def subsetCE) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63680diff
changeset | 1157 | |
| 60762 | 1158 | lemma collinear_empty [iff]: "collinear {}"
 | 
| 53406 | 1159 | by (simp add: collinear_def) | 
| 44133 | 1160 | |
| 60762 | 1161 | lemma collinear_sing [iff]: "collinear {x}"
 | 
| 44133 | 1162 | by (simp add: collinear_def) | 
| 1163 | ||
| 60762 | 1164 | lemma collinear_2 [iff]: "collinear {x, y}"
 | 
| 44133 | 1165 | apply (simp add: collinear_def) | 
| 1166 | apply (rule exI[where x="x - y"]) | |
| 68062 | 1167 | by (metis minus_diff_eq scaleR_left.minus scaleR_one) | 
| 44133 | 1168 | |
| 56444 | 1169 | lemma collinear_lemma: "collinear {0, x, y} \<longleftrightarrow> x = 0 \<or> y = 0 \<or> (\<exists>c. y = c *\<^sub>R x)"
 | 
| 53406 | 1170 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 68062 | 1171 | proof (cases "x = 0 \<or> y = 0") | 
| 1172 | case True | |
| 1173 | then show ?thesis | |
| 1174 | by (auto simp: insert_commute) | |
| 1175 | next | |
| 1176 | case False | |
| 1177 | show ?thesis | |
| 1178 | proof | |
| 1179 | assume h: "?lhs" | |
| 1180 |     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"
 | |
| 1181 | unfolding collinear_def by blast | |
| 1182 | from u[rule_format, of x 0] u[rule_format, of y 0] | |
| 1183 | obtain cx and cy where | |
| 1184 | cx: "x = cx *\<^sub>R u" and cy: "y = cy *\<^sub>R u" | |
| 1185 | by auto | |
| 1186 | from cx cy False have cx0: "cx \<noteq> 0" and cy0: "cy \<noteq> 0" by auto | |
| 1187 | let ?d = "cy / cx" | |
| 1188 | from cx cy cx0 have "y = ?d *\<^sub>R x" | |
| 1189 | by simp | |
| 1190 | then show ?rhs using False by blast | |
| 1191 | next | |
| 1192 | assume h: "?rhs" | |
| 1193 | then obtain c where c: "y = c *\<^sub>R x" | |
| 1194 | using False by blast | |
| 1195 | show ?lhs | |
| 1196 | unfolding collinear_def c | |
| 1197 | apply (rule exI[where x=x]) | |
| 1198 | apply auto | |
| 1199 | apply (rule exI[where x="- 1"], simp) | |
| 1200 | apply (rule exI[where x= "-c"], simp) | |
| 44133 | 1201 | apply (rule exI[where x=1], simp) | 
| 68062 | 1202 | apply (rule exI[where x="1 - c"], simp add: scaleR_left_diff_distrib) | 
| 1203 | apply (rule exI[where x="c - 1"], simp add: scaleR_left_diff_distrib) | |
| 1204 | done | |
| 1205 | qed | |
| 44133 | 1206 | qed | 
| 1207 | ||
| 56444 | 1208 | lemma norm_cauchy_schwarz_equal: "\<bar>x \<bullet> y\<bar> = norm x * norm y \<longleftrightarrow> collinear {0, x, y}"
 | 
| 68062 | 1209 | proof (cases "x=0") | 
| 1210 | case True | |
| 1211 | then show ?thesis | |
| 1212 | by (auto simp: insert_commute) | |
| 1213 | next | |
| 1214 | case False | |
| 1215 | then have nnz: "norm x \<noteq> 0" | |
| 1216 | by auto | |
| 1217 | show ?thesis | |
| 1218 | proof | |
| 1219 | assume "\<bar>x \<bullet> y\<bar> = norm x * norm y" | |
| 1220 |     then show "collinear {0, x, y}"
 | |
| 1221 | unfolding norm_cauchy_schwarz_abs_eq collinear_lemma | |
| 1222 | by (meson eq_vector_fraction_iff nnz) | |
| 1223 | next | |
| 1224 |     assume "collinear {0, x, y}"
 | |
| 1225 | with False show "\<bar>x \<bullet> y\<bar> = norm x * norm y" | |
| 1226 | unfolding norm_cauchy_schwarz_abs_eq collinear_lemma by (auto simp: abs_if) | |
| 1227 | qed | |
| 1228 | qed | |
| 49522 | 1229 | |
| 54776 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
54703diff
changeset | 1230 | end |