| author | nipkow | 
| Thu, 11 Aug 2022 13:23:00 +0200 | |
| changeset 75805 | 3581dcee70db | 
| parent 75455 | 91c16c5ad3e9 | 
| child 76836 | 30182f9e1818 | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/Linear_Algebra.thy | 
| 44133 | 2 | Author: Amine Chaieb, University of Cambridge | 
| 3 | *) | |
| 4 | ||
| 69517 | 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 | |
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 33 | lemma substdbasis_expansion_unique: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 34 | includes inner_syntax | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 35 | assumes d: "d \<subseteq> Basis" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 36 | shows "(\<Sum>i\<in>d. f i *\<^sub>R i) = (x::'a::euclidean_space) \<longleftrightarrow> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 37 | (\<forall>i\<in>Basis. (i \<in> d \<longrightarrow> f i = x \<bullet> i) \<and> (i \<notin> d \<longrightarrow> x \<bullet> i = 0))" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 38 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 39 | have *: "\<And>x a b P. x * (if P then a else b) = (if P then x * a else x * b)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 40 | by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 41 | have **: "finite d" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 42 | by (auto intro: finite_subset[OF assms]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 43 | have ***: "\<And>i. i \<in> Basis \<Longrightarrow> (\<Sum>i\<in>d. f i *\<^sub>R i) \<bullet> i = (\<Sum>x\<in>d. if x = i then f x else 0)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 44 | using d | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 45 | by (auto intro!: sum.cong simp: inner_Basis inner_sum_left) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 46 | show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 47 | unfolding euclidean_eq_iff[where 'a='a] by (auto simp: sum.delta[OF **] ***) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 48 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 49 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 50 | lemma independent_substdbasis: "d \<subseteq> Basis \<Longrightarrow> independent d" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 51 | by (rule independent_mono[OF independent_Basis]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 52 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 53 | lemma subset_translation_eq [simp]: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 54 | fixes a :: "'a::real_vector" shows "(+) a ` s \<subseteq> (+) a ` t \<longleftrightarrow> s \<subseteq> t" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 55 | by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 56 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 57 | lemma translate_inj_on: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 58 | fixes A :: "'a::ab_group_add set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 59 | shows "inj_on (\<lambda>x. a + x) A" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 60 | unfolding inj_on_def by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 61 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 62 | lemma translation_assoc: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 63 | fixes a b :: "'a::ab_group_add" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 64 | shows "(\<lambda>x. b + x) ` ((\<lambda>x. a + x) ` S) = (\<lambda>x. (a + b) + x) ` S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 65 | by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 66 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 67 | lemma translation_invert: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 68 | fixes a :: "'a::ab_group_add" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 69 | assumes "(\<lambda>x. a + x) ` A = (\<lambda>x. a + x) ` B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 70 | shows "A = B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 71 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 72 | have "(\<lambda>x. -a + x) ` ((\<lambda>x. a + x) ` A) = (\<lambda>x. - a + x) ` ((\<lambda>x. a + x) ` B)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 73 | using assms by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 74 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 75 | using translation_assoc[of "-a" a A] translation_assoc[of "-a" a B] by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 76 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 77 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 78 | lemma translation_galois: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 79 | fixes a :: "'a::ab_group_add" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 80 | shows "T = ((\<lambda>x. a + x) ` S) \<longleftrightarrow> S = ((\<lambda>x. (- a) + x) ` T)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 81 | using translation_assoc[of "-a" a S] | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 82 | apply auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 83 | using translation_assoc[of a "-a" T] | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 84 | apply auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 85 | done | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 86 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 87 | lemma translation_inverse_subset: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 88 | assumes "((\<lambda>x. - a + x) ` V) \<le> (S :: 'n::ab_group_add set)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 89 | shows "V \<le> ((\<lambda>x. a + x) ` S)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 90 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 91 |   {
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 92 | fix x | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 93 | assume "x \<in> V" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 94 | then have "x-a \<in> S" using assms by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 95 |     then have "x \<in> {a + v |v. v \<in> S}"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 96 | apply auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 97 | apply (rule exI[of _ "x-a"], simp) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 98 | done | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 99 | then have "x \<in> ((\<lambda>x. a+x) ` S)" by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 100 | } | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 101 | then show ?thesis by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 102 | qed | 
| 53406 | 103 | |
| 70136 | 104 | subsection\<^marker>\<open>tag unimportant\<close> \<open>More interesting properties of the norm\<close> | 
| 63050 | 105 | |
| 69674 | 106 | unbundle inner_syntax | 
| 63050 | 107 | |
| 69597 | 108 | text\<open>Equality of vectors in terms of \<^term>\<open>(\<bullet>)\<close> products.\<close> | 
| 63050 | 109 | |
| 110 | lemma linear_componentwise: | |
| 111 | fixes f:: "'a::euclidean_space \<Rightarrow> 'b::real_inner" | |
| 112 | assumes lf: "linear f" | |
| 113 | shows "(f x) \<bullet> j = (\<Sum>i\<in>Basis. (x\<bullet>i) * (f i\<bullet>j))" (is "?lhs = ?rhs") | |
| 114 | proof - | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 115 | interpret linear f by fact | 
| 63050 | 116 | have "?rhs = (\<Sum>i\<in>Basis. (x\<bullet>i) *\<^sub>R (f i))\<bullet>j" | 
| 64267 | 117 | by (simp add: inner_sum_left) | 
| 63050 | 118 | then show ?thesis | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 119 | by (simp add: euclidean_representation sum[symmetric] scale[symmetric]) | 
| 63050 | 120 | qed | 
| 121 | ||
| 122 | lemma vector_eq: "x = y \<longleftrightarrow> x \<bullet> x = x \<bullet> y \<and> y \<bullet> y = x \<bullet> x" | |
| 123 | (is "?lhs \<longleftrightarrow> ?rhs") | |
| 124 | proof | |
| 125 | assume ?lhs | |
| 126 | then show ?rhs by simp | |
| 127 | next | |
| 128 | assume ?rhs | |
| 129 | then have "x \<bullet> x - x \<bullet> y = 0 \<and> x \<bullet> y - y \<bullet> y = 0" | |
| 130 | by simp | |
| 131 | then have "x \<bullet> (x - y) = 0 \<and> y \<bullet> (x - y) = 0" | |
| 132 | by (simp add: inner_diff inner_commute) | |
| 133 | then have "(x - y) \<bullet> (x - y) = 0" | |
| 134 | by (simp add: field_simps inner_diff inner_commute) | |
| 135 | then show "x = y" by simp | |
| 136 | qed | |
| 137 | ||
| 138 | lemma norm_triangle_half_r: | |
| 139 | "norm (y - x1) < e / 2 \<Longrightarrow> norm (y - x2) < e / 2 \<Longrightarrow> norm (x1 - x2) < e" | |
| 140 | using dist_triangle_half_r unfolding dist_norm[symmetric] by auto | |
| 141 | ||
| 142 | lemma norm_triangle_half_l: | |
| 143 | assumes "norm (x - y) < e / 2" | |
| 144 | and "norm (x' - y) < e / 2" | |
| 145 | shows "norm (x - x') < e" | |
| 146 | using dist_triangle_half_l[OF assms[unfolded dist_norm[symmetric]]] | |
| 147 | unfolding dist_norm[symmetric] . | |
| 148 | ||
| 66420 | 149 | lemma abs_triangle_half_r: | 
| 150 | fixes y :: "'a::linordered_field" | |
| 151 | shows "abs (y - x1) < e / 2 \<Longrightarrow> abs (y - x2) < e / 2 \<Longrightarrow> abs (x1 - x2) < e" | |
| 152 | by linarith | |
| 153 | ||
| 154 | lemma abs_triangle_half_l: | |
| 155 | fixes y :: "'a::linordered_field" | |
| 156 | assumes "abs (x - y) < e / 2" | |
| 157 | and "abs (x' - y) < e / 2" | |
| 158 | shows "abs (x - x') < e" | |
| 159 | using assms by linarith | |
| 160 | ||
| 64267 | 161 | lemma sum_clauses: | 
| 162 |   shows "sum f {} = 0"
 | |
| 163 | and "finite S \<Longrightarrow> sum f (insert x S) = (if x \<in> S then sum f S else f x + sum f S)" | |
| 63050 | 164 | by (auto simp add: insert_absorb) | 
| 165 | ||
| 166 | lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = z" | |
| 167 | proof | |
| 168 | assume "\<forall>x. x \<bullet> y = x \<bullet> z" | |
| 169 | then have "\<forall>x. x \<bullet> (y - z) = 0" | |
| 170 | by (simp add: inner_diff) | |
| 171 | then have "(y - z) \<bullet> (y - z) = 0" .. | |
| 172 | then show "y = z" by simp | |
| 173 | qed simp | |
| 174 | ||
| 175 | lemma vector_eq_rdot: "(\<forall>z. x \<bullet> z = y \<bullet> z) \<longleftrightarrow> x = y" | |
| 176 | proof | |
| 177 | assume "\<forall>z. x \<bullet> z = y \<bullet> z" | |
| 178 | then have "\<forall>z. (x - y) \<bullet> z = 0" | |
| 179 | by (simp add: inner_diff) | |
| 180 | then have "(x - y) \<bullet> (x - y) = 0" .. | |
| 181 | then show "x = y" by simp | |
| 182 | qed simp | |
| 183 | ||
| 69619 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 184 | subsection \<open>Substandard Basis\<close> | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 185 | |
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 186 | lemma ex_card: | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 187 | assumes "n \<le> card A" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 188 | shows "\<exists>S\<subseteq>A. card S = n" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 189 | proof (cases "finite A") | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 190 | case True | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 191 |   from ex_bij_betw_nat_finite[OF this] obtain f where f: "bij_betw f {0..<card A} A" ..
 | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 192 |   moreover from f \<open>n \<le> card A\<close> have "{..< n} \<subseteq> {..< card A}" "inj_on f {..< n}"
 | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 193 | by (auto simp: bij_betw_def intro: subset_inj_on) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 194 |   ultimately have "f ` {..< n} \<subseteq> A" "card (f ` {..< n}) = n"
 | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 195 | by (auto simp: bij_betw_def card_image) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 196 | then show ?thesis by blast | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 197 | next | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 198 | case False | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 199 | with \<open>n \<le> card A\<close> show ?thesis by force | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 200 | qed | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 201 | |
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 202 | lemma subspace_substandard: "subspace {x::'a::euclidean_space. (\<forall>i\<in>Basis. P i \<longrightarrow> x\<bullet>i = 0)}"
 | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 203 | by (auto simp: subspace_def inner_add_left) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 204 | |
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 205 | lemma dim_substandard: | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 206 | assumes d: "d \<subseteq> Basis" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 207 |   shows "dim {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0} = card d" (is "dim ?A = _")
 | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 208 | proof (rule dim_unique) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 209 | from d show "d \<subseteq> ?A" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 210 | by (auto simp: inner_Basis) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 211 | from d show "independent d" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 212 | by (rule independent_mono [OF independent_Basis]) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 213 | have "x \<in> span d" if "\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0" for x | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 214 | proof - | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 215 | have "finite d" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 216 | by (rule finite_subset [OF d finite_Basis]) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 217 | then have "(\<Sum>i\<in>d. (x \<bullet> i) *\<^sub>R i) \<in> span d" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 218 | by (simp add: span_sum span_clauses) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 219 | also have "(\<Sum>i\<in>d. (x \<bullet> i) *\<^sub>R i) = (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i)" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 220 | by (rule sum.mono_neutral_cong_left [OF finite_Basis d]) (auto simp: that) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 221 | finally show "x \<in> span d" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 222 | by (simp only: euclidean_representation) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 223 | qed | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 224 | then show "?A \<subseteq> span d" by auto | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 225 | qed simp | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 226 | |
| 63050 | 227 | |
| 68901 | 228 | subsection \<open>Orthogonality\<close> | 
| 63050 | 229 | |
| 70136 | 230 | definition\<^marker>\<open>tag important\<close> (in real_inner) "orthogonal x y \<longleftrightarrow> x \<bullet> y = 0" | 
| 67962 | 231 | |
| 63050 | 232 | context real_inner | 
| 233 | begin | |
| 234 | ||
| 63072 | 235 | lemma orthogonal_self: "orthogonal x x \<longleftrightarrow> x = 0" | 
| 236 | by (simp add: orthogonal_def) | |
| 237 | ||
| 63050 | 238 | lemma orthogonal_clauses: | 
| 239 | "orthogonal a 0" | |
| 240 | "orthogonal a x \<Longrightarrow> orthogonal a (c *\<^sub>R x)" | |
| 241 | "orthogonal a x \<Longrightarrow> orthogonal a (- x)" | |
| 242 | "orthogonal a x \<Longrightarrow> orthogonal a y \<Longrightarrow> orthogonal a (x + y)" | |
| 243 | "orthogonal a x \<Longrightarrow> orthogonal a y \<Longrightarrow> orthogonal a (x - y)" | |
| 244 | "orthogonal 0 a" | |
| 245 | "orthogonal x a \<Longrightarrow> orthogonal (c *\<^sub>R x) a" | |
| 246 | "orthogonal x a \<Longrightarrow> orthogonal (- x) a" | |
| 247 | "orthogonal x a \<Longrightarrow> orthogonal y a \<Longrightarrow> orthogonal (x + y) a" | |
| 248 | "orthogonal x a \<Longrightarrow> orthogonal y a \<Longrightarrow> orthogonal (x - y) a" | |
| 249 | unfolding orthogonal_def inner_add inner_diff by auto | |
| 250 | ||
| 251 | end | |
| 252 | ||
| 253 | lemma orthogonal_commute: "orthogonal x y \<longleftrightarrow> orthogonal y x" | |
| 254 | by (simp add: orthogonal_def inner_commute) | |
| 255 | ||
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 256 | 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 | 257 | by (rule ext) (simp add: orthogonal_def) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 258 | |
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 259 | lemma pairwise_ortho_scaleR: | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 260 | "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 | 261 | \<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 | 262 | by (auto simp: pairwise_def orthogonal_clauses) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 263 | |
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 264 | lemma orthogonal_rvsum: | 
| 64267 | 265 | "\<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 | 266 | 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 | 267 | |
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 268 | lemma orthogonal_lvsum: | 
| 64267 | 269 | "\<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 | 270 | 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 | 271 | |
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 272 | lemma norm_add_Pythagorean: | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 273 | assumes "orthogonal a b" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 274 | 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 | 275 | proof - | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 276 | 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 | 277 | 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 | 278 | then show ?thesis | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 279 | by (simp add: power2_norm_eq_inner) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 280 | qed | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 281 | |
| 64267 | 282 | lemma norm_sum_Pythagorean: | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 283 | assumes "finite I" "pairwise (\<lambda>i j. orthogonal (f i) (f j)) I" | 
| 64267 | 284 | 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 | 285 | using assms | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 286 | proof (induction I rule: finite_induct) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 287 | case empty then show ?case by simp | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 288 | next | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 289 | case (insert x I) | 
| 64267 | 290 | 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 | 291 | by (metis pairwise_insert orthogonal_rvsum) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 292 | with insert show ?case | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 293 | by (simp add: pairwise_insert norm_add_Pythagorean) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 294 | qed | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63075diff
changeset | 295 | |
| 63050 | 296 | |
| 69683 | 297 | subsection \<open>Orthogonality of a transformation\<close> | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 298 | |
| 70136 | 299 | definition\<^marker>\<open>tag important\<close> "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v w. f v \<bullet> f w = v \<bullet> w)" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 300 | |
| 70136 | 301 | lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 302 | "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v. norm (f v) = norm v)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 303 | unfolding orthogonal_transformation_def | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 304 | apply auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 305 | apply (erule_tac x=v in allE)+ | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 306 | apply (simp add: norm_eq_sqrt_inner) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 307 | apply (simp add: dot_norm linear_add[symmetric]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 308 | done | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 309 | |
| 70136 | 310 | lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_id [simp]: "orthogonal_transformation (\<lambda>x. x)" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 311 | by (simp add: linear_iff orthogonal_transformation_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 312 | |
| 70136 | 313 | lemma\<^marker>\<open>tag unimportant\<close> orthogonal_orthogonal_transformation: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 314 | "orthogonal_transformation f \<Longrightarrow> orthogonal (f x) (f y) \<longleftrightarrow> orthogonal x y" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 315 | by (simp add: orthogonal_def orthogonal_transformation_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 316 | |
| 70136 | 317 | lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_compose: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 318 | "\<lbrakk>orthogonal_transformation f; orthogonal_transformation g\<rbrakk> \<Longrightarrow> orthogonal_transformation(f \<circ> g)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 319 | by (auto simp: orthogonal_transformation_def linear_compose) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 320 | |
| 70136 | 321 | lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_neg: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 322 | "orthogonal_transformation(\<lambda>x. -(f x)) \<longleftrightarrow> orthogonal_transformation f" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 323 | by (auto simp: orthogonal_transformation_def dest: linear_compose_neg) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 324 | |
| 70136 | 325 | lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_scaleR: "orthogonal_transformation f \<Longrightarrow> f (c *\<^sub>R v) = c *\<^sub>R f v" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 326 | by (simp add: linear_iff orthogonal_transformation_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 327 | |
| 70136 | 328 | lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_linear: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 329 | "orthogonal_transformation f \<Longrightarrow> linear f" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 330 | by (simp add: orthogonal_transformation_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 331 | |
| 70136 | 332 | lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_inj: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 333 | "orthogonal_transformation f \<Longrightarrow> inj f" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 334 | unfolding orthogonal_transformation_def inj_on_def | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 335 | by (metis vector_eq) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 336 | |
| 70136 | 337 | lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_surj: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 338 | "orthogonal_transformation f \<Longrightarrow> surj f" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 339 | for f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 340 | by (simp add: linear_injective_imp_surjective orthogonal_transformation_inj orthogonal_transformation_linear) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 341 | |
| 70136 | 342 | lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_bij: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 343 | "orthogonal_transformation f \<Longrightarrow> bij f" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 344 | for f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 345 | by (simp add: bij_def orthogonal_transformation_inj orthogonal_transformation_surj) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 346 | |
| 70136 | 347 | lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_inv: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 348 | "orthogonal_transformation f \<Longrightarrow> orthogonal_transformation (inv f)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 349 | for f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space" | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73795diff
changeset | 350 | by (metis (no_types, opaque_lifting) bijection.inv_right bijection_def inj_linear_imp_inv_linear orthogonal_transformation orthogonal_transformation_bij orthogonal_transformation_inj) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 351 | |
| 70136 | 352 | lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_norm: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 353 | "orthogonal_transformation f \<Longrightarrow> norm (f x) = norm x" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 354 | by (metis orthogonal_transformation) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 355 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 356 | |
| 68901 | 357 | subsection \<open>Bilinear functions\<close> | 
| 63050 | 358 | |
| 70136 | 359 | definition\<^marker>\<open>tag important\<close> | 
| 69600 | 360 | bilinear :: "('a::real_vector \<Rightarrow> 'b::real_vector \<Rightarrow> 'c::real_vector) \<Rightarrow> bool" where
 | 
| 361 | "bilinear f \<longleftrightarrow> (\<forall>x. linear (\<lambda>y. f x y)) \<and> (\<forall>y. linear (\<lambda>x. f x y))" | |
| 63050 | 362 | |
| 363 | lemma bilinear_ladd: "bilinear h \<Longrightarrow> h (x + y) z = h x z + h y z" | |
| 364 | by (simp add: bilinear_def linear_iff) | |
| 365 | ||
| 366 | lemma bilinear_radd: "bilinear h \<Longrightarrow> h x (y + z) = h x y + h x z" | |
| 367 | by (simp add: bilinear_def linear_iff) | |
| 368 | ||
| 70707 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 369 | lemma bilinear_times: | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 370 | fixes c::"'a::real_algebra" shows "bilinear (\<lambda>x y::'a. x*y)" | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 371 | by (auto simp: bilinear_def distrib_left distrib_right intro!: linearI) | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70688diff
changeset | 372 | |
| 63050 | 373 | lemma bilinear_lmul: "bilinear h \<Longrightarrow> h (c *\<^sub>R x) y = c *\<^sub>R h x y" | 
| 374 | by (simp add: bilinear_def linear_iff) | |
| 375 | ||
| 376 | lemma bilinear_rmul: "bilinear h \<Longrightarrow> h x (c *\<^sub>R y) = c *\<^sub>R h x y" | |
| 377 | by (simp add: bilinear_def linear_iff) | |
| 378 | ||
| 379 | lemma bilinear_lneg: "bilinear h \<Longrightarrow> h (- x) y = - h x y" | |
| 380 | by (drule bilinear_lmul [of _ "- 1"]) simp | |
| 381 | ||
| 382 | lemma bilinear_rneg: "bilinear h \<Longrightarrow> h x (- y) = - h x y" | |
| 383 | by (drule bilinear_rmul [of _ _ "- 1"]) simp | |
| 384 | ||
| 385 | lemma (in ab_group_add) eq_add_iff: "x = x + y \<longleftrightarrow> y = 0" | |
| 386 | using add_left_imp_eq[of x y 0] by auto | |
| 387 | ||
| 388 | lemma bilinear_lzero: | |
| 389 | assumes "bilinear h" | |
| 390 | shows "h 0 x = 0" | |
| 391 | using bilinear_ladd [OF assms, of 0 0 x] by (simp add: eq_add_iff field_simps) | |
| 392 | ||
| 393 | lemma bilinear_rzero: | |
| 394 | assumes "bilinear h" | |
| 395 | shows "h x 0 = 0" | |
| 396 | using bilinear_radd [OF assms, of x 0 0 ] by (simp add: eq_add_iff field_simps) | |
| 397 | ||
| 398 | lemma bilinear_lsub: "bilinear h \<Longrightarrow> h (x - y) z = h x z - h y z" | |
| 399 | using bilinear_ladd [of h x "- y"] by (simp add: bilinear_lneg) | |
| 400 | ||
| 401 | lemma bilinear_rsub: "bilinear h \<Longrightarrow> h z (x - y) = h z x - h z y" | |
| 402 | using bilinear_radd [of h _ x "- y"] by (simp add: bilinear_rneg) | |
| 403 | ||
| 64267 | 404 | lemma bilinear_sum: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 405 | assumes "bilinear h" | 
| 64267 | 406 | shows "h (sum f S) (sum g T) = sum (\<lambda>(i,j). h (f i) (g j)) (S \<times> T) " | 
| 63050 | 407 | proof - | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 408 | 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 | 409 | interpret r: linear "\<lambda>y. h x y" for x using assms by (simp add: bilinear_def) | 
| 64267 | 410 | 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 | 411 | by (simp add: l.sum) | 
| 64267 | 412 | 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 | 413 | by (rule sum.cong) (simp_all add: r.sum) | 
| 63050 | 414 | finally show ?thesis | 
| 64267 | 415 | unfolding sum.cartesian_product . | 
| 63050 | 416 | qed | 
| 417 | ||
| 418 | ||
| 68901 | 419 | subsection \<open>Adjoints\<close> | 
| 63050 | 420 | |
| 70136 | 421 | definition\<^marker>\<open>tag important\<close> adjoint :: "(('a::real_inner) \<Rightarrow> ('b::real_inner)) \<Rightarrow> 'b \<Rightarrow> 'a" where
 | 
| 69600 | 422 | "adjoint f = (SOME f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y)" | 
| 63050 | 423 | |
| 424 | lemma adjoint_unique: | |
| 425 | assumes "\<forall>x y. inner (f x) y = inner x (g y)" | |
| 426 | shows "adjoint f = g" | |
| 427 | unfolding adjoint_def | |
| 428 | proof (rule some_equality) | |
| 429 | show "\<forall>x y. inner (f x) y = inner x (g y)" | |
| 430 | by (rule assms) | |
| 431 | next | |
| 432 | fix h | |
| 433 | assume "\<forall>x y. inner (f x) y = inner x (h y)" | |
| 434 | then have "\<forall>x y. inner x (g y) = inner x (h y)" | |
| 435 | using assms by simp | |
| 436 | then have "\<forall>x y. inner x (g y - h y) = 0" | |
| 437 | by (simp add: inner_diff_right) | |
| 438 | then have "\<forall>y. inner (g y - h y) (g y - h y) = 0" | |
| 439 | by simp | |
| 440 | then have "\<forall>y. h y = g y" | |
| 441 | by simp | |
| 442 | then show "h = g" by (simp add: ext) | |
| 443 | qed | |
| 444 | ||
| 445 | text \<open>TODO: The following lemmas about adjoints should hold for any | |
| 63680 | 446 | Hilbert space (i.e. complete inner product space). | 
| 68224 | 447 | (see \<^url>\<open>https://en.wikipedia.org/wiki/Hermitian_adjoint\<close>) | 
| 63050 | 448 | \<close> | 
| 449 | ||
| 450 | lemma adjoint_works: | |
| 451 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | |
| 452 | assumes lf: "linear f" | |
| 453 | shows "x \<bullet> adjoint f y = f x \<bullet> y" | |
| 454 | proof - | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 455 | interpret linear f by fact | 
| 63050 | 456 | have "\<forall>y. \<exists>w. \<forall>x. f x \<bullet> y = x \<bullet> w" | 
| 457 | proof (intro allI exI) | |
| 458 | fix y :: "'m" and x | |
| 459 | let ?w = "(\<Sum>i\<in>Basis. (f i \<bullet> y) *\<^sub>R i) :: 'n" | |
| 460 | have "f x \<bullet> y = f (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i) \<bullet> y" | |
| 461 | by (simp add: euclidean_representation) | |
| 462 | 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 | 463 | by (simp add: sum scale) | 
| 63050 | 464 | finally show "f x \<bullet> y = x \<bullet> ?w" | 
| 64267 | 465 | by (simp add: inner_sum_left inner_sum_right mult.commute) | 
| 63050 | 466 | qed | 
| 467 | then show ?thesis | |
| 468 | unfolding adjoint_def choice_iff | |
| 469 | by (intro someI2_ex[where Q="\<lambda>f'. x \<bullet> f' y = f x \<bullet> y"]) auto | |
| 470 | qed | |
| 471 | ||
| 472 | lemma adjoint_clauses: | |
| 473 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | |
| 474 | assumes lf: "linear f" | |
| 475 | shows "x \<bullet> adjoint f y = f x \<bullet> y" | |
| 476 | and "adjoint f y \<bullet> x = y \<bullet> f x" | |
| 477 | by (simp_all add: adjoint_works[OF lf] inner_commute) | |
| 478 | ||
| 479 | lemma adjoint_linear: | |
| 480 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | |
| 481 | assumes lf: "linear f" | |
| 482 | shows "linear (adjoint f)" | |
| 483 | by (simp add: lf linear_iff euclidean_eq_iff[where 'a='n] euclidean_eq_iff[where 'a='m] | |
| 484 | adjoint_clauses[OF lf] inner_distrib) | |
| 485 | ||
| 486 | lemma adjoint_adjoint: | |
| 487 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | |
| 488 | assumes lf: "linear f" | |
| 489 | shows "adjoint (adjoint f) = f" | |
| 490 | by (rule adjoint_unique, simp add: adjoint_clauses [OF lf]) | |
| 491 | ||
| 492 | ||
| 70136 | 493 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Euclidean Spaces as Typeclass\<close> | 
| 44133 | 494 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50105diff
changeset | 495 | 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 | 496 | 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 | 497 | |
| 53939 | 498 | 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 | 499 | by (rule span_Basis) | 
| 44133 | 500 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50105diff
changeset | 501 | 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 | 502 | 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 | 503 | |
| 53406 | 504 | |
| 70136 | 505 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Linearity and Bilinearity continued\<close> | 
| 44133 | 506 | |
| 507 | lemma linear_bounded: | |
| 56444 | 508 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 44133 | 509 | assumes lf: "linear f" | 
| 510 | shows "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x" | |
| 53939 | 511 | proof | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 512 | 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 | 513 | let ?B = "\<Sum>b\<in>Basis. norm (f b)" | 
| 53939 | 514 | show "\<forall>x. norm (f x) \<le> ?B * norm x" | 
| 515 | proof | |
| 53406 | 516 | 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 | 517 | 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 | 518 | 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 | 519 | unfolding euclidean_representation .. | 
| 64267 | 520 | 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 | 521 | by (simp add: sum scale) | 
| 64267 | 522 | 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 | 523 | 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 | 524 | proof - | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64267diff
changeset | 525 | from Basis_le_norm[OF that, of x] | 
| 53939 | 526 | 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 | 527 | unfolding norm_scaleR by (metis mult.commute mult_left_mono norm_ge_zero) | 
| 53939 | 528 | qed | 
| 64267 | 529 | from sum_norm_le[of _ ?g, OF th] | 
| 53939 | 530 | show "norm (f x) \<le> ?B * norm x" | 
| 64267 | 531 | unfolding th0 sum_distrib_right by metis | 
| 53939 | 532 | qed | 
| 44133 | 533 | qed | 
| 534 | ||
| 535 | lemma linear_conv_bounded_linear: | |
| 536 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | |
| 537 | shows "linear f \<longleftrightarrow> bounded_linear f" | |
| 538 | proof | |
| 539 | assume "linear f" | |
| 53939 | 540 | then interpret f: linear f . | 
| 44133 | 541 | show "bounded_linear f" | 
| 542 | proof | |
| 543 | have "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x" | |
| 60420 | 544 | using \<open>linear f\<close> by (rule linear_bounded) | 
| 49522 | 545 | 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 | 546 | by (simp add: mult.commute) | 
| 44133 | 547 | qed | 
| 548 | next | |
| 549 | assume "bounded_linear f" | |
| 550 | then interpret f: bounded_linear f . | |
| 53939 | 551 | show "linear f" .. | 
| 552 | qed | |
| 553 | ||
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61306diff
changeset | 554 | lemmas linear_linear = linear_conv_bounded_linear[symmetric] | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61306diff
changeset | 555 | |
| 70999 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 556 | lemma inj_linear_imp_inv_bounded_linear: | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 557 | fixes f::"'a::euclidean_space \<Rightarrow> 'a" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 558 | shows "\<lbrakk>bounded_linear f; inj f\<rbrakk> \<Longrightarrow> bounded_linear (inv f)" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 559 | by (simp add: inj_linear_imp_inv_linear linear_linear) | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 560 | |
| 53939 | 561 | lemma linear_bounded_pos: | 
| 56444 | 562 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 53939 | 563 | 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 | 564 | obtains B where "B > 0" "\<And>x. norm (f x) \<le> B * norm x" | 
| 53939 | 565 | proof - | 
| 566 | have "\<exists>B > 0. \<forall>x. norm (f x) \<le> norm x * B" | |
| 567 | using lf unfolding linear_conv_bounded_linear | |
| 568 | 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 | 569 | with that show ?thesis | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 570 | by (auto simp: mult.commute) | 
| 44133 | 571 | qed | 
| 572 | ||
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 573 | 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 | 574 | 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 | 575 | 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 | 576 | 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 | 577 | proof - | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 578 | 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 | 579 | 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 | 580 | show thesis | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 581 | proof | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 582 | show "0 < 1/B" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 583 | 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 | 584 | 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 | 585 | proof - | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 586 | 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 | 587 | 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 | 588 | 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 | 589 | 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 | 590 | finally show ?thesis . | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 591 | qed | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 592 | qed | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 593 | qed | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 594 | |
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 595 | 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 | 596 | 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 | 597 | 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 | 598 | 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 | 599 | 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 | 600 | 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 | 601 | |
| 49522 | 602 | lemma bounded_linearI': | 
| 56444 | 603 | fixes f ::"'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 53406 | 604 | assumes "\<And>x y. f (x + y) = f x + f y" | 
| 605 | and "\<And>c x. f (c *\<^sub>R x) = c *\<^sub>R f x" | |
| 49522 | 606 | 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 | 607 | using assms linearI linear_conv_bounded_linear by blast | 
| 44133 | 608 | |
| 609 | lemma bilinear_bounded: | |
| 56444 | 610 | fixes h :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'k::real_normed_vector" | 
| 44133 | 611 | assumes bh: "bilinear h" | 
| 612 | 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 | 613 | proof (clarify intro!: exI[of _ "\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. norm (h i j)"]) | 
| 53406 | 614 | fix x :: 'm | 
| 615 | fix y :: 'n | |
| 64267 | 616 | 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 | 617 | by (simp add: euclidean_representation) | 
| 64267 | 618 | 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 | 619 | 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 | 620 | 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 | 621 | 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 | 622 | \<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 | 623 | 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 | 624 | 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 | 625 | unfolding sum_distrib_right th sum.cartesian_product | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 626 | 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 | 627 | field_simps simp del: scaleR_scaleR intro!: sum_norm_le) | 
| 44133 | 628 | qed | 
| 629 | ||
| 630 | lemma bilinear_conv_bounded_bilinear: | |
| 631 | fixes h :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::real_normed_vector" | |
| 632 | shows "bilinear h \<longleftrightarrow> bounded_bilinear h" | |
| 633 | proof | |
| 634 | assume "bilinear h" | |
| 635 | show "bounded_bilinear h" | |
| 636 | proof | |
| 53406 | 637 | fix x y z | 
| 638 | show "h (x + y) z = h x z + h y z" | |
| 60420 | 639 | using \<open>bilinear h\<close> unfolding bilinear_def linear_iff by simp | 
| 44133 | 640 | next | 
| 53406 | 641 | fix x y z | 
| 642 | show "h x (y + z) = h x y + h x z" | |
| 60420 | 643 | using \<open>bilinear h\<close> unfolding bilinear_def linear_iff by simp | 
| 44133 | 644 | next | 
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68062diff
changeset | 645 | 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 | 646 | 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 | 647 | by simp_all | 
| 44133 | 648 | next | 
| 649 | have "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y" | |
| 60420 | 650 | using \<open>bilinear h\<close> by (rule bilinear_bounded) | 
| 49522 | 651 | 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 | 652 | by (simp add: ac_simps) | 
| 44133 | 653 | qed | 
| 654 | next | |
| 655 | assume "bounded_bilinear h" | |
| 656 | then interpret h: bounded_bilinear h . | |
| 657 | show "bilinear h" | |
| 658 | unfolding bilinear_def linear_conv_bounded_linear | |
| 49522 | 659 | using h.bounded_linear_left h.bounded_linear_right by simp | 
| 44133 | 660 | qed | 
| 661 | ||
| 53939 | 662 | lemma bilinear_bounded_pos: | 
| 56444 | 663 | fixes h :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::real_normed_vector" | 
| 53939 | 664 | assumes bh: "bilinear h" | 
| 665 | shows "\<exists>B > 0. \<forall>x y. norm (h x y) \<le> B * norm x * norm y" | |
| 666 | proof - | |
| 667 | have "\<exists>B > 0. \<forall>x y. norm (h x y) \<le> norm x * norm y * B" | |
| 668 | using bh [unfolded bilinear_conv_bounded_bilinear] | |
| 669 | by (rule bounded_bilinear.pos_bounded) | |
| 670 | then show ?thesis | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 671 | by (simp only: ac_simps) | 
| 53939 | 672 | qed | 
| 673 | ||
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 674 | 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 | 675 | 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 | 676 | dest: bounded_linear.linear) | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 677 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 678 | lemma linear_imp_has_derivative: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 679 | 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 | 680 | 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 | 681 | 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 | 682 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 683 | 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 | 684 | 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 | 685 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 686 | lemma linear_imp_differentiable: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63170diff
changeset | 687 | 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 | 688 | 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 | 689 | 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 | 690 | |
| 73795 | 691 | lemma of_real_differentiable [simp,derivative_intros]: "of_real differentiable F" | 
| 692 | by (simp add: bounded_linear_imp_differentiable bounded_linear_of_real) | |
| 693 | ||
| 49522 | 694 | |
| 70136 | 695 | subsection\<^marker>\<open>tag unimportant\<close> \<open>We continue\<close> | 
| 44133 | 696 | |
| 697 | lemma independent_bound: | |
| 53716 | 698 | fixes S :: "'a::euclidean_space set" | 
| 699 |   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 | 700 | 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 | 701 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 702 | lemmas independent_imp_finite = finiteI_independent | 
| 44133 | 703 | |
| 71120 | 704 | corollary\<^marker>\<open>tag unimportant\<close> independent_card_le: | 
| 60303 | 705 | fixes S :: "'a::euclidean_space set" | 
| 706 | assumes "independent S" | |
| 71120 | 707 |   shows "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 | 708 | using assms independent_bound by auto | 
| 63075 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63072diff
changeset | 709 | |
| 49663 | 710 | lemma dependent_biggerset: | 
| 56444 | 711 | fixes S :: "'a::euclidean_space set" | 
| 712 |   shows "(finite S \<Longrightarrow> card S > DIM('a)) \<Longrightarrow> dependent S"
 | |
| 44133 | 713 | by (metis independent_bound not_less) | 
| 714 | ||
| 60420 | 715 | text \<open>Picking an orthogonal replacement for a spanning set.\<close> | 
| 44133 | 716 | |
| 53406 | 717 | lemma vector_sub_project_orthogonal: | 
| 718 | fixes b x :: "'a::euclidean_space" | |
| 719 | shows "b \<bullet> (x - ((b \<bullet> x) / (b \<bullet> b)) *\<^sub>R b) = 0" | |
| 44133 | 720 | unfolding inner_simps by auto | 
| 721 | ||
| 44528 | 722 | lemma pairwise_orthogonal_insert: | 
| 723 | assumes "pairwise orthogonal S" | |
| 49522 | 724 | and "\<And>y. y \<in> S \<Longrightarrow> orthogonal x y" | 
| 44528 | 725 | shows "pairwise orthogonal (insert x S)" | 
| 726 | using assms unfolding pairwise_def | |
| 727 | by (auto simp add: orthogonal_commute) | |
| 728 | ||
| 44133 | 729 | lemma basis_orthogonal: | 
| 53406 | 730 | fixes B :: "'a::real_inner set" | 
| 44133 | 731 | assumes fB: "finite B" | 
| 732 | shows "\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C" | |
| 733 | (is " \<exists>C. ?P B C") | |
| 49522 | 734 | using fB | 
| 735 | proof (induct rule: finite_induct) | |
| 736 | case empty | |
| 53406 | 737 | then show ?case | 
| 738 |     apply (rule exI[where x="{}"])
 | |
| 739 | apply (auto simp add: pairwise_def) | |
| 740 | done | |
| 44133 | 741 | next | 
| 49522 | 742 | case (insert a B) | 
| 60420 | 743 | note fB = \<open>finite B\<close> and aB = \<open>a \<notin> B\<close> | 
| 744 | from \<open>\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C\<close> | |
| 44133 | 745 | obtain C where C: "finite C" "card C \<le> card B" | 
| 746 | "span C = span B" "pairwise orthogonal C" by blast | |
| 64267 | 747 | let ?a = "a - sum (\<lambda>x. (x \<bullet> a / (x \<bullet> x)) *\<^sub>R x) C" | 
| 44133 | 748 | let ?C = "insert ?a C" | 
| 53406 | 749 | from C(1) have fC: "finite ?C" | 
| 750 | by simp | |
| 49522 | 751 | from fB aB C(1,2) have cC: "card ?C \<le> card (insert a B)" | 
| 752 | by (simp add: card_insert_if) | |
| 53406 | 753 |   {
 | 
| 754 | fix x k | |
| 49522 | 755 | have th0: "\<And>(a::'a) b c. a - (b - c) = c + (a - b)" | 
| 756 | by (simp add: field_simps) | |
| 44133 | 757 | 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" | 
| 758 | apply (simp only: scaleR_right_diff_distrib th0) | |
| 759 | 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 | 760 | apply (rule span_scale) | 
| 64267 | 761 | 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 | 762 | apply (rule span_scale) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 763 | apply (rule span_base) | 
| 49522 | 764 | apply assumption | 
| 53406 | 765 | done | 
| 766 | } | |
| 44133 | 767 | then have SC: "span ?C = span (insert a B)" | 
| 768 | unfolding set_eq_iff span_breakdown_eq C(3)[symmetric] by auto | |
| 53406 | 769 |   {
 | 
| 770 | fix y | |
| 771 | assume yC: "y \<in> C" | |
| 772 |     then have Cy: "C = insert y (C - {y})"
 | |
| 773 | by blast | |
| 774 |     have fth: "finite (C - {y})"
 | |
| 775 | using C by simp | |
| 44528 | 776 | have "orthogonal ?a y" | 
| 777 | unfolding orthogonal_def | |
| 64267 | 778 | unfolding inner_diff inner_sum_left right_minus_eq | 
| 779 | unfolding sum.remove [OF \<open>finite C\<close> \<open>y \<in> C\<close>] | |
| 44528 | 780 | apply (clarsimp simp add: inner_commute[of y a]) | 
| 64267 | 781 | apply (rule sum.neutral) | 
| 44528 | 782 | apply clarsimp | 
| 783 | apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format]) | |
| 60420 | 784 | using \<open>y \<in> C\<close> by auto | 
| 53406 | 785 | } | 
| 60420 | 786 | with \<open>pairwise orthogonal C\<close> have CPO: "pairwise orthogonal ?C" | 
| 44528 | 787 | by (rule pairwise_orthogonal_insert) | 
| 53406 | 788 | from fC cC SC CPO have "?P (insert a B) ?C" | 
| 789 | by blast | |
| 44133 | 790 | then show ?case by blast | 
| 791 | qed | |
| 792 | ||
| 793 | lemma orthogonal_basis_exists: | |
| 794 |   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 | 795 | 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 | 796 | (card B = dim V) \<and> pairwise orthogonal B" | 
| 49663 | 797 | proof - | 
| 49522 | 798 | from basis_exists[of V] obtain B where | 
| 53406 | 799 | 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 | 800 | by force | 
| 53406 | 801 | from B have fB: "finite B" "card B = dim V" | 
| 802 | using independent_bound by auto | |
| 44133 | 803 | from basis_orthogonal[OF fB(1)] obtain C where | 
| 53406 | 804 | C: "finite C" "card C \<le> card B" "span C = span B" "pairwise orthogonal C" | 
| 805 | by blast | |
| 806 | 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 | 807 | by (metis span_superset span_mono subset_trans) | 
| 53406 | 808 | from span_mono[OF B(3)] C have SVC: "span V \<subseteq> span C" | 
| 809 | by (simp add: span_span) | |
| 44133 | 810 | from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB | 
| 53406 | 811 | have iC: "independent C" | 
| 71044 | 812 | by (simp) | 
| 53406 | 813 | from C fB have "card C \<le> dim V" | 
| 814 | by simp | |
| 815 | moreover have "dim V \<le> card C" | |
| 816 | 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 | 817 | by simp | 
| 53406 | 818 | ultimately have CdV: "card C = dim V" | 
| 819 | using C(1) by simp | |
| 820 | from C B CSV CdV iC show ?thesis | |
| 821 | by auto | |
| 44133 | 822 | qed | 
| 823 | ||
| 60420 | 824 | text \<open>Low-dimensional subset is in a hyperplane (weak orthogonal complement).\<close> | 
| 44133 | 825 | |
| 49522 | 826 | lemma span_not_univ_orthogonal: | 
| 53406 | 827 | fixes S :: "'a::euclidean_space set" | 
| 44133 | 828 | assumes sU: "span S \<noteq> UNIV" | 
| 56444 | 829 | shows "\<exists>a::'a. a \<noteq> 0 \<and> (\<forall>x \<in> span S. a \<bullet> x = 0)" | 
| 49522 | 830 | proof - | 
| 53406 | 831 | from sU obtain a where a: "a \<notin> span S" | 
| 832 | by blast | |
| 44133 | 833 | 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 | 834 | 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 | 835 | "card B = dim S" "pairwise orthogonal B" | 
| 44133 | 836 | by blast | 
| 53406 | 837 | from B have fB: "finite B" "card B = dim S" | 
| 838 | using independent_bound by auto | |
| 44133 | 839 | from span_mono[OF B(2)] span_mono[OF B(3)] | 
| 53406 | 840 | have sSB: "span S = span B" | 
| 841 | by (simp add: span_span) | |
| 64267 | 842 | let ?a = "a - sum (\<lambda>b. (a \<bullet> b / (b \<bullet> b)) *\<^sub>R b) B" | 
| 843 | have "sum (\<lambda>b. (a \<bullet> b / (b \<bullet> b)) *\<^sub>R b) B \<in> span S" | |
| 44133 | 844 | unfolding sSB | 
| 64267 | 845 | 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 | 846 | apply (rule span_scale) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 847 | apply (rule span_base) | 
| 49522 | 848 | apply assumption | 
| 849 | done | |
| 53406 | 850 | with a have a0:"?a \<noteq> 0" | 
| 851 | by auto | |
| 68058 | 852 | have "?a \<bullet> x = 0" if "x\<in>span B" for x | 
| 853 | proof (rule span_induct [OF that]) | |
| 49522 | 854 |     show "subspace {x. ?a \<bullet> x = 0}"
 | 
| 855 | by (auto simp add: subspace_def inner_add) | |
| 856 | next | |
| 53406 | 857 |     {
 | 
| 858 | fix x | |
| 859 | assume x: "x \<in> B" | |
| 860 |       from x have B': "B = insert x (B - {x})"
 | |
| 861 | by blast | |
| 862 |       have fth: "finite (B - {x})"
 | |
| 863 | using fB by simp | |
| 44133 | 864 | have "?a \<bullet> x = 0" | 
| 53406 | 865 | apply (subst B') | 
| 866 | using fB fth | |
| 64267 | 867 | unfolding sum_clauses(2)[OF fth] | 
| 44133 | 868 | apply simp unfolding inner_simps | 
| 64267 | 869 | apply (clarsimp simp add: inner_add inner_sum_left) | 
| 870 | apply (rule sum.neutral, rule ballI) | |
| 63170 | 871 | apply (simp only: inner_commute) | 
| 49711 | 872 | apply (auto simp add: x field_simps | 
| 873 | intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format]) | |
| 53406 | 874 | done | 
| 875 | } | |
| 68058 | 876 | then show "?a \<bullet> x = 0" if "x \<in> B" for x | 
| 877 | using that by blast | |
| 878 | qed | |
| 53406 | 879 | with a0 show ?thesis | 
| 880 | unfolding sSB by (auto intro: exI[where x="?a"]) | |
| 44133 | 881 | qed | 
| 882 | ||
| 883 | lemma span_not_univ_subset_hyperplane: | |
| 53406 | 884 | fixes S :: "'a::euclidean_space set" | 
| 885 | assumes SU: "span S \<noteq> UNIV" | |
| 44133 | 886 |   shows "\<exists> a. a \<noteq>0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
 | 
| 887 | using span_not_univ_orthogonal[OF SU] by auto | |
| 888 | ||
| 49663 | 889 | lemma lowdim_subset_hyperplane: | 
| 53406 | 890 | fixes S :: "'a::euclidean_space set" | 
| 44133 | 891 |   assumes d: "dim S < DIM('a)"
 | 
| 56444 | 892 |   shows "\<exists>a::'a. a \<noteq> 0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
 | 
| 49522 | 893 | proof - | 
| 53406 | 894 |   {
 | 
| 895 | assume "span S = UNIV" | |
| 896 |     then have "dim (span S) = dim (UNIV :: ('a) set)"
 | |
| 897 | by simp | |
| 898 |     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 | 899 | by (metis Euclidean_Space.dim_UNIV dim_span) | 
| 53406 | 900 | with d have False by arith | 
| 901 | } | |
| 902 | then have th: "span S \<noteq> UNIV" | |
| 903 | by blast | |
| 44133 | 904 | from span_not_univ_subset_hyperplane[OF th] show ?thesis . | 
| 905 | qed | |
| 906 | ||
| 907 | lemma linear_eq_stdbasis: | |
| 56444 | 908 | fixes f :: "'a::euclidean_space \<Rightarrow> _" | 
| 909 | assumes lf: "linear f" | |
| 49663 | 910 | and lg: "linear g" | 
| 68058 | 911 | and fg: "\<And>b. b \<in> Basis \<Longrightarrow> f b = g b" | 
| 44133 | 912 | shows "f = g" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 913 | 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 | 914 | by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67982diff
changeset | 915 | |
| 44133 | 916 | |
| 60420 | 917 | text \<open>Similar results for bilinear functions.\<close> | 
| 44133 | 918 | |
| 919 | lemma bilinear_eq: | |
| 920 | assumes bf: "bilinear f" | |
| 49522 | 921 | and bg: "bilinear g" | 
| 53406 | 922 | and SB: "S \<subseteq> span B" | 
| 923 | and TC: "T \<subseteq> span C" | |
| 68058 | 924 | and "x\<in>S" "y\<in>T" | 
| 925 | and fg: "\<And>x y. \<lbrakk>x \<in> B; y\<in> C\<rbrakk> \<Longrightarrow> f x y = g x y" | |
| 926 | shows "f x y = g x y" | |
| 49663 | 927 | proof - | 
| 44170 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 huffman parents: 
44166diff
changeset | 928 |   let ?P = "{x. \<forall>y\<in> span C. f x y = g x y}"
 | 
| 44133 | 929 | from bf bg have sp: "subspace ?P" | 
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53596diff
changeset | 930 | 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 | 931 | 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 | 932 | span_add Ball_def | 
| 49663 | 933 | intro: bilinear_ladd[OF bf]) | 
| 68058 | 934 |   have sfg: "\<And>x. x \<in> B \<Longrightarrow> subspace {a. f x a = g x a}"
 | 
| 44133 | 935 | apply (auto simp add: subspace_def) | 
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53596diff
changeset | 936 | 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 | 937 | 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 | 938 | span_add Ball_def | 
| 49663 | 939 | intro: bilinear_ladd[OF bf]) | 
| 49522 | 940 | done | 
| 68058 | 941 | have "\<forall>y\<in> span C. f x y = g x y" if "x \<in> span B" for x | 
| 942 | apply (rule span_induct [OF that sp]) | |
| 68062 | 943 | using fg sfg span_induct by blast | 
| 53406 | 944 | then show ?thesis | 
| 68058 | 945 | using SB TC assms by auto | 
| 44133 | 946 | qed | 
| 947 | ||
| 49522 | 948 | lemma bilinear_eq_stdbasis: | 
| 53406 | 949 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> _" | 
| 44133 | 950 | assumes bf: "bilinear f" | 
| 49522 | 951 | and bg: "bilinear g" | 
| 68058 | 952 | and fg: "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> f i j = g i j" | 
| 44133 | 953 | shows "f = g" | 
| 68074 | 954 | using bilinear_eq[OF bf bg equalityD2[OF span_Basis] equalityD2[OF span_Basis]] fg by blast | 
| 49522 | 955 | |
| 69619 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69600diff
changeset | 956 | |
| 60420 | 957 | subsection \<open>Infinity norm\<close> | 
| 44133 | 958 | |
| 70136 | 959 | definition\<^marker>\<open>tag important\<close> "infnorm (x::'a::euclidean_space) = Sup {\<bar>x \<bullet> b\<bar> |b. b \<in> Basis}"
 | 
| 44133 | 960 | |
| 961 | lemma infnorm_set_image: | |
| 53716 | 962 | fixes x :: "'a::euclidean_space" | 
| 56444 | 963 |   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 | 964 | by blast | 
| 44133 | 965 | |
| 53716 | 966 | lemma infnorm_Max: | 
| 967 | fixes x :: "'a::euclidean_space" | |
| 56444 | 968 | 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 | 969 | 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 | 970 | |
| 44133 | 971 | lemma infnorm_set_lemma: | 
| 53716 | 972 | fixes x :: "'a::euclidean_space" | 
| 56444 | 973 |   shows "finite {\<bar>x \<bullet> i\<bar> |i. i \<in> Basis}"
 | 
| 974 |     and "{\<bar>x \<bullet> i\<bar> |i. i \<in> Basis} \<noteq> {}"
 | |
| 44133 | 975 | unfolding infnorm_set_image | 
| 976 | by auto | |
| 977 | ||
| 53406 | 978 | lemma infnorm_pos_le: | 
| 979 | fixes x :: "'a::euclidean_space" | |
| 980 | 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 | 981 | by (simp add: infnorm_Max Max_ge_iff ex_in_conv) | 
| 44133 | 982 | |
| 53406 | 983 | lemma infnorm_triangle: | 
| 984 | fixes x :: "'a::euclidean_space" | |
| 985 | shows "infnorm (x + y) \<le> infnorm x + infnorm y" | |
| 49522 | 986 | proof - | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 987 | 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 | 988 | by simp | 
| 44133 | 989 | show ?thesis | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 990 | by (auto simp: infnorm_Max inner_add_left intro!: *) | 
| 44133 | 991 | qed | 
| 992 | ||
| 53406 | 993 | lemma infnorm_eq_0: | 
| 994 | fixes x :: "'a::euclidean_space" | |
| 995 | shows "infnorm x = 0 \<longleftrightarrow> x = 0" | |
| 49522 | 996 | proof - | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 997 | 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 | 998 | 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 | 999 | then show ?thesis | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 1000 | using infnorm_pos_le[of x] by simp | 
| 44133 | 1001 | qed | 
| 1002 | ||
| 1003 | lemma infnorm_0: "infnorm 0 = 0" | |
| 1004 | by (simp add: infnorm_eq_0) | |
| 1005 | ||
| 1006 | lemma infnorm_neg: "infnorm (- x) = infnorm x" | |
| 68062 | 1007 | unfolding infnorm_def by simp | 
| 44133 | 1008 | |
| 1009 | lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)" | |
| 68062 | 1010 | by (metis infnorm_neg minus_diff_eq) | 
| 1011 | ||
| 1012 | lemma absdiff_infnorm: "\<bar>infnorm x - infnorm y\<bar> \<le> infnorm (x - y)" | |
| 49522 | 1013 | proof - | 
| 68062 | 1014 | have *: "\<And>(nx::real) n ny. nx \<le> n + ny \<Longrightarrow> ny \<le> n + nx \<Longrightarrow> \<bar>nx - ny\<bar> \<le> n" | 
| 44133 | 1015 | by arith | 
| 68062 | 1016 | show ?thesis | 
| 1017 | proof (rule *) | |
| 1018 | from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"] | |
| 1019 | show "infnorm x \<le> infnorm (x - y) + infnorm y" "infnorm y \<le> infnorm (x - y) + infnorm x" | |
| 1020 | by (simp_all add: field_simps infnorm_neg) | |
| 1021 | qed | |
| 44133 | 1022 | qed | 
| 1023 | ||
| 53406 | 1024 | lemma real_abs_infnorm: "\<bar>infnorm x\<bar> = infnorm x" | 
| 44133 | 1025 | using infnorm_pos_le[of x] by arith | 
| 1026 | ||
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50105diff
changeset | 1027 | lemma Basis_le_infnorm: | 
| 53406 | 1028 | fixes x :: "'a::euclidean_space" | 
| 1029 | 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 | 1030 | by (simp add: infnorm_Max) | 
| 44133 | 1031 | |
| 56444 | 1032 | 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 | 1033 | unfolding infnorm_Max | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 1034 | 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 | 1035 | let ?B = "(\<lambda>i. \<bar>x \<bullet> i\<bar>) ` Basis" | 
| 68062 | 1036 |   { fix b :: 'a
 | 
| 53406 | 1037 | assume "b \<in> Basis" | 
| 1038 | then show "\<bar>a *\<^sub>R x \<bullet> b\<bar> \<le> \<bar>a\<bar> * Max ?B" | |
| 1039 | by (simp add: abs_mult mult_left_mono) | |
| 1040 | next | |
| 1041 | from Max_in[of ?B] obtain b where "b \<in> Basis" "Max ?B = \<bar>x \<bullet> b\<bar>" | |
| 1042 | by (auto simp del: Max_in) | |
| 1043 | 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" | |
| 1044 | by (intro image_eqI[where x=b]) (auto simp: abs_mult) | |
| 1045 | } | |
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 1046 | qed simp | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
50526diff
changeset | 1047 | |
| 53406 | 1048 | 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 | 1049 | unfolding infnorm_mul .. | 
| 44133 | 1050 | |
| 1051 | lemma infnorm_pos_lt: "infnorm x > 0 \<longleftrightarrow> x \<noteq> 0" | |
| 1052 | using infnorm_pos_le[of x] infnorm_eq_0[of x] by arith | |
| 1053 | ||
| 60420 | 1054 | text \<open>Prove that it differs only up to a bound from Euclidean norm.\<close> | 
| 44133 | 1055 | |
| 1056 | 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 | 1057 | 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 | 1058 | |
| 53716 | 1059 | lemma norm_le_infnorm: | 
| 1060 | fixes x :: "'a::euclidean_space" | |
| 1061 |   shows "norm x \<le> sqrt DIM('a) * infnorm x"
 | |
| 73795 | 1062 | unfolding norm_eq_sqrt_inner id_def | 
| 68062 | 1063 | proof (rule real_le_lsqrt[OF inner_ge_zero]) | 
| 1064 |   show "sqrt DIM('a) * infnorm x \<ge> 0"
 | |
| 44133 | 1065 | by (simp add: zero_le_mult_iff infnorm_pos_le) | 
| 68062 | 1066 | have "x \<bullet> x \<le> (\<Sum>b\<in>Basis. x \<bullet> b * (x \<bullet> b))" | 
| 1067 | by (metis euclidean_inner order_refl) | |
| 1068 |   also have "... \<le> DIM('a) * \<bar>infnorm x\<bar>\<^sup>2"
 | |
| 1069 | by (rule sum_bounded_above) (metis Basis_le_infnorm abs_le_square_iff power2_eq_square real_abs_infnorm) | |
| 1070 |   also have "... \<le> (sqrt DIM('a) * infnorm x)\<^sup>2"
 | |
| 1071 | by (simp add: power_mult_distrib) | |
| 1072 |   finally show "x \<bullet> x \<le> (sqrt DIM('a) * infnorm x)\<^sup>2" .
 | |
| 44133 | 1073 | qed | 
| 1074 | ||
| 44646 | 1075 | lemma tendsto_infnorm [tendsto_intros]: | 
| 61973 | 1076 | assumes "(f \<longlongrightarrow> a) F" | 
| 1077 | shows "((\<lambda>x. infnorm (f x)) \<longlongrightarrow> infnorm a) F" | |
| 44646 | 1078 | proof (rule tendsto_compose [OF LIM_I assms]) | 
| 53406 | 1079 | fix r :: real | 
| 1080 | assume "r > 0" | |
| 49522 | 1081 | then show "\<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (infnorm x - infnorm a) < r" | 
| 68062 | 1082 | by (metis real_norm_def le_less_trans absdiff_infnorm infnorm_le_norm) | 
| 44646 | 1083 | qed | 
| 1084 | ||
| 60420 | 1085 | text \<open>Equality in Cauchy-Schwarz and triangle inequalities.\<close> | 
| 44133 | 1086 | |
| 53406 | 1087 | lemma norm_cauchy_schwarz_eq: "x \<bullet> y = norm x * norm y \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x" | 
| 1088 | (is "?lhs \<longleftrightarrow> ?rhs") | |
| 68062 | 1089 | proof (cases "x=0") | 
| 1090 | case True | |
| 73795 | 1091 | then show ?thesis | 
| 68062 | 1092 | by auto | 
| 1093 | next | |
| 1094 | case False | |
| 1095 | from inner_eq_zero_iff[of "norm y *\<^sub>R x - norm x *\<^sub>R y"] | |
| 1096 | have "?rhs \<longleftrightarrow> | |
| 49522 | 1097 | (norm y * (norm y * norm x * norm x - norm x * (x \<bullet> y)) - | 
| 1098 | norm x * (norm y * (y \<bullet> x) - norm x * norm y * norm y) = 0)" | |
| 68062 | 1099 | using False unfolding inner_simps | 
| 1100 | by (auto simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps) | |
| 73795 | 1101 | also have "\<dots> \<longleftrightarrow> (2 * norm x * norm y * (norm x * norm y - x \<bullet> y) = 0)" | 
| 68062 | 1102 | using False by (simp add: field_simps inner_commute) | 
| 73795 | 1103 | also have "\<dots> \<longleftrightarrow> ?lhs" | 
| 68062 | 1104 | using False by auto | 
| 1105 | finally show ?thesis by metis | |
| 44133 | 1106 | qed | 
| 1107 | ||
| 1108 | lemma norm_cauchy_schwarz_abs_eq: | |
| 56444 | 1109 | "\<bar>x \<bullet> y\<bar> = norm x * norm y \<longleftrightarrow> | 
| 53716 | 1110 | norm x *\<^sub>R y = norm y *\<^sub>R x \<or> norm x *\<^sub>R y = - norm y *\<^sub>R x" | 
| 53406 | 1111 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 49522 | 1112 | proof - | 
| 56444 | 1113 | have th: "\<And>(x::real) a. a \<ge> 0 \<Longrightarrow> \<bar>x\<bar> = a \<longleftrightarrow> x = a \<or> x = - a" | 
| 53406 | 1114 | by arith | 
| 44133 | 1115 | have "?rhs \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x \<or> norm (- x) *\<^sub>R y = norm y *\<^sub>R (- x)" | 
| 1116 | by simp | |
| 68062 | 1117 | also have "\<dots> \<longleftrightarrow> (x \<bullet> y = norm x * norm y \<or> (- x) \<bullet> y = norm x * norm y)" | 
| 44133 | 1118 | unfolding norm_cauchy_schwarz_eq[symmetric] | 
| 1119 | unfolding norm_minus_cancel norm_scaleR .. | |
| 1120 | also have "\<dots> \<longleftrightarrow> ?lhs" | |
| 53406 | 1121 | unfolding th[OF mult_nonneg_nonneg, OF norm_ge_zero[of x] norm_ge_zero[of y]] inner_simps | 
| 1122 | by auto | |
| 44133 | 1123 | finally show ?thesis .. | 
| 1124 | qed | |
| 1125 | ||
| 1126 | lemma norm_triangle_eq: | |
| 1127 | fixes x y :: "'a::real_inner" | |
| 53406 | 1128 | shows "norm (x + y) = norm x + norm y \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x" | 
| 68062 | 1129 | proof (cases "x = 0 \<or> y = 0") | 
| 1130 | case True | |
| 73795 | 1131 | then show ?thesis | 
| 68062 | 1132 | by force | 
| 1133 | next | |
| 1134 | case False | |
| 1135 | then have n: "norm x > 0" "norm y > 0" | |
| 1136 | by auto | |
| 1137 | have "norm (x + y) = norm x + norm y \<longleftrightarrow> (norm (x + y))\<^sup>2 = (norm x + norm y)\<^sup>2" | |
| 1138 | by simp | |
| 1139 | also have "\<dots> \<longleftrightarrow> norm x *\<^sub>R y = norm y *\<^sub>R x" | |
| 1140 | unfolding norm_cauchy_schwarz_eq[symmetric] | |
| 1141 | unfolding power2_norm_eq_inner inner_simps | |
| 1142 | by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps) | |
| 1143 | finally show ?thesis . | |
| 44133 | 1144 | qed | 
| 1145 | ||
| 74729 
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 1146 | lemma dist_triangle_eq: | 
| 
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 1147 | fixes x y z :: "'a::real_inner" | 
| 
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 1148 | shows "dist x z = dist x y + dist y z \<longleftrightarrow> | 
| 
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 1149 | norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)" | 
| 
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 1150 | proof - | 
| 
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 1151 | have *: "x - y + (y - z) = x - z" by auto | 
| 
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 1152 | show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *] | 
| 
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 1153 | by (auto simp:norm_minus_commute) | 
| 
64b3d8d9bd10
new lemmas about convex, concave functions, + tidying
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 1154 | qed | 
| 49522 | 1155 | |
| 60420 | 1156 | subsection \<open>Collinearity\<close> | 
| 44133 | 1157 | |
| 70136 | 1158 | definition\<^marker>\<open>tag important\<close> collinear :: "'a::real_vector set \<Rightarrow> bool" | 
| 49522 | 1159 | where "collinear S \<longleftrightarrow> (\<exists>u. \<forall>x \<in> S. \<forall> y \<in> S. \<exists>c. x - y = c *\<^sub>R u)" | 
| 44133 | 1160 | |
| 66287 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1161 | lemma collinear_alt: | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1162 | "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 | 1163 | proof | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1164 | assume ?lhs | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1165 | then show ?rhs | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1166 | 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 | 1167 | next | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1168 | assume ?rhs | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1169 | then obtain u v where *: "\<And>x. x \<in> S \<Longrightarrow> \<exists>c. x = u + c *\<^sub>R v" | 
| 75455 
91c16c5ad3e9
tidied auto / simp with null arguments
 paulson <lp15@cam.ac.uk> parents: 
74729diff
changeset | 1170 | by auto | 
| 66287 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1171 | 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 | 1172 | 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 | 1173 | then show ?lhs | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1174 | using collinear_def by blast | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1175 | qed | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1176 | |
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1177 | lemma collinear: | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1178 |   fixes S :: "'a::{perfect_space,real_vector} set"
 | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1179 | 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 | 1180 | proof - | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1181 | 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 | 1182 | 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 | 1183 | proof - | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1184 | 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 | 1185 | using that by auto | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1186 | moreover | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1187 | obtain v::'a where "v \<noteq> 0" | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1188 | using UNIV_not_singleton [of 0] by auto | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1189 | 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 | 1190 | by auto | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1191 | then show ?thesis | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1192 | using \<open>v \<noteq> 0\<close> by blast | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1193 | qed | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1194 | then show ?thesis | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1195 | 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 | 1196 | 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 | 1197 | qed | 
| 
005a30862ed0
new material: Colinearity, convex sets, polytopes
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 1198 | |
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63680diff
changeset | 1199 | 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 | 1200 | by (meson collinear_def subsetCE) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63680diff
changeset | 1201 | |
| 60762 | 1202 | lemma collinear_empty [iff]: "collinear {}"
 | 
| 53406 | 1203 | by (simp add: collinear_def) | 
| 44133 | 1204 | |
| 60762 | 1205 | lemma collinear_sing [iff]: "collinear {x}"
 | 
| 44133 | 1206 | by (simp add: collinear_def) | 
| 1207 | ||
| 60762 | 1208 | lemma collinear_2 [iff]: "collinear {x, y}"
 | 
| 44133 | 1209 | apply (simp add: collinear_def) | 
| 1210 | apply (rule exI[where x="x - y"]) | |
| 68062 | 1211 | by (metis minus_diff_eq scaleR_left.minus scaleR_one) | 
| 44133 | 1212 | |
| 56444 | 1213 | lemma collinear_lemma: "collinear {0, x, y} \<longleftrightarrow> x = 0 \<or> y = 0 \<or> (\<exists>c. y = c *\<^sub>R x)"
 | 
| 53406 | 1214 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 68062 | 1215 | proof (cases "x = 0 \<or> y = 0") | 
| 1216 | case True | |
| 1217 | then show ?thesis | |
| 1218 | by (auto simp: insert_commute) | |
| 1219 | next | |
| 1220 | case False | |
| 73795 | 1221 | show ?thesis | 
| 68062 | 1222 | proof | 
| 1223 | assume h: "?lhs" | |
| 1224 |     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"
 | |
| 1225 | unfolding collinear_def by blast | |
| 1226 | from u[rule_format, of x 0] u[rule_format, of y 0] | |
| 1227 | obtain cx and cy where | |
| 1228 | cx: "x = cx *\<^sub>R u" and cy: "y = cy *\<^sub>R u" | |
| 1229 | by auto | |
| 1230 | from cx cy False have cx0: "cx \<noteq> 0" and cy0: "cy \<noteq> 0" by auto | |
| 1231 | let ?d = "cy / cx" | |
| 1232 | from cx cy cx0 have "y = ?d *\<^sub>R x" | |
| 1233 | by simp | |
| 1234 | then show ?rhs using False by blast | |
| 1235 | next | |
| 1236 | assume h: "?rhs" | |
| 1237 | then obtain c where c: "y = c *\<^sub>R x" | |
| 1238 | using False by blast | |
| 1239 | show ?lhs | |
| 1240 | unfolding collinear_def c | |
| 1241 | apply (rule exI[where x=x]) | |
| 1242 | apply auto | |
| 1243 | apply (rule exI[where x="- 1"], simp) | |
| 1244 | apply (rule exI[where x= "-c"], simp) | |
| 44133 | 1245 | apply (rule exI[where x=1], simp) | 
| 68062 | 1246 | apply (rule exI[where x="1 - c"], simp add: scaleR_left_diff_distrib) | 
| 1247 | apply (rule exI[where x="c - 1"], simp add: scaleR_left_diff_distrib) | |
| 1248 | done | |
| 1249 | qed | |
| 44133 | 1250 | qed | 
| 1251 | ||
| 73885 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1252 | lemma collinear_iff_Reals: "collinear {0::complex,w,z} \<longleftrightarrow> z/w \<in> \<real>"
 | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1253 | proof | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1254 |   show "z/w \<in> \<real> \<Longrightarrow> collinear {0,w,z}"
 | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1255 | by (metis Reals_cases collinear_lemma nonzero_divide_eq_eq scaleR_conv_of_real) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1256 | qed (auto simp: collinear_lemma scaleR_conv_of_real) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1257 | |
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1258 | lemma collinear_scaleR_iff: "collinear {0, \<alpha> *\<^sub>R w, \<beta> *\<^sub>R z} \<longleftrightarrow> collinear {0,w,z} \<or> \<alpha>=0 \<or> \<beta>=0"
 | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1259 | (is "?lhs = ?rhs") | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1260 | proof (cases "\<alpha>=0 \<or> \<beta>=0") | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1261 | case False | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1262 | then have "(\<exists>c. \<beta> *\<^sub>R z = (c * \<alpha>) *\<^sub>R w) = (\<exists>c. z = c *\<^sub>R w)" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1263 | by (metis mult.commute scaleR_scaleR vector_fraction_eq_iff) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1264 | then show ?thesis | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1265 | by (auto simp add: collinear_lemma) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1266 | qed (auto simp: collinear_lemma) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1267 | |
| 56444 | 1268 | lemma norm_cauchy_schwarz_equal: "\<bar>x \<bullet> y\<bar> = norm x * norm y \<longleftrightarrow> collinear {0, x, y}"
 | 
| 68062 | 1269 | proof (cases "x=0") | 
| 1270 | case True | |
| 1271 | then show ?thesis | |
| 1272 | by (auto simp: insert_commute) | |
| 1273 | next | |
| 1274 | case False | |
| 1275 | then have nnz: "norm x \<noteq> 0" | |
| 1276 | by auto | |
| 1277 | show ?thesis | |
| 1278 | proof | |
| 1279 | assume "\<bar>x \<bullet> y\<bar> = norm x * norm y" | |
| 1280 |     then show "collinear {0, x, y}"
 | |
| 73795 | 1281 | unfolding norm_cauchy_schwarz_abs_eq collinear_lemma | 
| 68062 | 1282 | by (meson eq_vector_fraction_iff nnz) | 
| 1283 | next | |
| 1284 |     assume "collinear {0, x, y}"
 | |
| 1285 | with False show "\<bar>x \<bullet> y\<bar> = norm x * norm y" | |
| 1286 | unfolding norm_cauchy_schwarz_abs_eq collinear_lemma by (auto simp: abs_if) | |
| 1287 | qed | |
| 1288 | qed | |
| 49522 | 1289 | |
| 73885 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1290 | lemma norm_triangle_eq_imp_collinear: | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1291 | fixes x y :: "'a::real_inner" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1292 | assumes "norm (x + y) = norm x + norm y" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1293 |   shows "collinear{0,x,y}"
 | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1294 | proof (cases "x = 0 \<or> y = 0") | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1295 | case False | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1296 | with assms show ?thesis | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1297 | by (meson norm_cauchy_schwarz_abs_eq norm_cauchy_schwarz_equal norm_triangle_eq) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1298 | qed (use collinear_lemma in blast) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1299 | |
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1300 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1301 | subsection\<open>Properties of special hyperplanes\<close> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1302 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1303 | lemma subspace_hyperplane: "subspace {x. a \<bullet> x = 0}"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1304 | by (simp add: subspace_def inner_right_distrib) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1305 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1306 | lemma subspace_hyperplane2: "subspace {x. x \<bullet> a = 0}"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1307 | by (simp add: inner_commute inner_right_distrib subspace_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1308 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1309 | lemma special_hyperplane_span: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1310 | fixes S :: "'n::euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1311 | assumes "k \<in> Basis" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1312 |   shows "{x. k \<bullet> x = 0} = span (Basis - {k})"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1313 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1314 |   have *: "x \<in> span (Basis - {k})" if "k \<bullet> x = 0" for x
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1315 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1316 | have "x = (\<Sum>b\<in>Basis. (x \<bullet> b) *\<^sub>R b)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1317 | by (simp add: euclidean_representation) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1318 |     also have "... = (\<Sum>b \<in> Basis - {k}. (x \<bullet> b) *\<^sub>R b)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1319 | by (auto simp: sum.remove [of _ k] inner_commute assms that) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1320 |     finally have "x = (\<Sum>b\<in>Basis - {k}. (x \<bullet> b) *\<^sub>R b)" .
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1321 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1322 | by (simp add: span_finite) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1323 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1324 | show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1325 | apply (rule span_subspace [symmetric]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1326 | using assms | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1327 | apply (auto simp: inner_not_same_Basis intro: * subspace_hyperplane) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1328 | done | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1329 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1330 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1331 | lemma dim_special_hyperplane: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1332 | fixes k :: "'n::euclidean_space" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1333 |   shows "k \<in> Basis \<Longrightarrow> dim {x. k \<bullet> x = 0} = DIM('n) - 1"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1334 | apply (simp add: special_hyperplane_span) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1335 | apply (rule dim_unique [OF subset_refl]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1336 | apply (auto simp: independent_substdbasis) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1337 | apply (metis member_remove remove_def span_base) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1338 | done | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1339 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1340 | proposition dim_hyperplane: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1341 | fixes a :: "'a::euclidean_space" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1342 | assumes "a \<noteq> 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1343 |     shows "dim {x. a \<bullet> x = 0} = DIM('a) - 1"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1344 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1345 |   have span0: "span {x. a \<bullet> x = 0} = {x. a \<bullet> x = 0}"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1346 | by (rule span_unique) (auto simp: subspace_hyperplane) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1347 | then obtain B where "independent B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1348 |               and Bsub: "B \<subseteq> {x. a \<bullet> x = 0}"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1349 |               and subspB: "{x. a \<bullet> x = 0} \<subseteq> span B"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1350 |               and card0: "(card B = dim {x. a \<bullet> x = 0})"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1351 | and ortho: "pairwise orthogonal B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1352 | using orthogonal_basis_exists by metis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1353 | with assms have "a \<notin> span B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1354 | by (metis (mono_tags, lifting) span_eq inner_eq_zero_iff mem_Collect_eq span0) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1355 | then have ind: "independent (insert a B)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1356 | by (simp add: \<open>independent B\<close> independent_insert) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1357 | have "finite B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1358 | using \<open>independent B\<close> independent_bound by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1359 | have "UNIV \<subseteq> span (insert a B)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1360 | proof fix y::'a | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1361 | obtain r z where z: "y = r *\<^sub>R a + z" "a \<bullet> z = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1362 | apply (rule_tac r="(a \<bullet> y) / (a \<bullet> a)" and z = "y - ((a \<bullet> y) / (a \<bullet> a)) *\<^sub>R a" in that) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1363 | using assms | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1364 | by (auto simp: algebra_simps) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1365 | show "y \<in> span (insert a B)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1366 | by (metis (mono_tags, lifting) z Bsub span_eq_iff | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1367 | add_diff_cancel_left' mem_Collect_eq span0 span_breakdown_eq span_subspace subspB) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1368 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1369 |   then have dima: "DIM('a) = dim(insert a B)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1370 | by (metis independent_Basis span_Basis dim_eq_card top.extremum_uniqueI) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1371 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1372 | by (metis (mono_tags, lifting) Bsub Diff_insert_absorb \<open>a \<notin> span B\<close> ind card0 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1373 | card_Diff_singleton dim_span indep_card_eq_dim_span insertI1 subsetCE | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1374 | subspB) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1375 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1376 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1377 | lemma lowdim_eq_hyperplane: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1378 | fixes S :: "'a::euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1379 |   assumes "dim S = DIM('a) - 1"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1380 |   obtains a where "a \<noteq> 0" and "span S = {x. a \<bullet> x = 0}"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1381 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1382 |   have dimS: "dim S < DIM('a)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1383 | by (simp add: assms) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1384 |   then obtain b where b: "b \<noteq> 0" "span S \<subseteq> {a. b \<bullet> a = 0}"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1385 | using lowdim_subset_hyperplane [of S] by fastforce | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1386 | show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1387 | apply (rule that[OF b(1)]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1388 | apply (rule subspace_dim_equal) | 
| 71044 | 1389 | by (auto simp: assms b dim_hyperplane subspace_hyperplane) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1390 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1391 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1392 | lemma dim_eq_hyperplane: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1393 | fixes S :: "'n::euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1394 |   shows "dim S = DIM('n) - 1 \<longleftrightarrow> (\<exists>a. a \<noteq> 0 \<and> span S = {x. a \<bullet> x = 0})"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1395 | by (metis One_nat_def dim_hyperplane dim_span lowdim_eq_hyperplane) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1396 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1397 | |
| 71044 | 1398 | subsection\<open> Orthogonal bases and Gram-Schmidt process\<close> | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1399 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1400 | lemma pairwise_orthogonal_independent: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1401 | assumes "pairwise orthogonal S" and "0 \<notin> S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1402 | shows "independent S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1403 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1404 | have 0: "\<And>x y. \<lbrakk>x \<noteq> y; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1405 | using assms by (simp add: pairwise_def orthogonal_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1406 |   have "False" if "a \<in> S" and a: "a \<in> span (S - {a})" for a
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1407 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1408 |     obtain T U where "T \<subseteq> S - {a}" "a = (\<Sum>v\<in>T. U v *\<^sub>R v)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1409 | using a by (force simp: span_explicit) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1410 | then have "a \<bullet> a = a \<bullet> (\<Sum>v\<in>T. U v *\<^sub>R v)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1411 | by simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1412 | also have "... = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1413 | apply (simp add: inner_sum_right) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1414 | apply (rule comm_monoid_add_class.sum.neutral) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1415 |       by (metis "0" DiffE \<open>T \<subseteq> S - {a}\<close> mult_not_zero singletonI subsetCE \<open>a \<in> S\<close>)
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1416 | finally show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1417 | using \<open>0 \<notin> S\<close> \<open>a \<in> S\<close> by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1418 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1419 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1420 | by (force simp: dependent_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1421 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1422 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1423 | lemma pairwise_orthogonal_imp_finite: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1424 | fixes S :: "'a::euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1425 | assumes "pairwise orthogonal S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1426 | shows "finite S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1427 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1428 |   have "independent (S - {0})"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1429 | apply (rule pairwise_orthogonal_independent) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1430 | apply (metis Diff_iff assms pairwise_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1431 | by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1432 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1433 | by (meson independent_imp_finite infinite_remove) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1434 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1435 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1436 | lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1437 | by (simp add: subspace_def orthogonal_clauses) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1438 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1439 | lemma subspace_orthogonal_to_vectors: "subspace {y. \<forall>x \<in> S. orthogonal x y}"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1440 | by (simp add: subspace_def orthogonal_clauses) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1441 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1442 | lemma orthogonal_to_span: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1443 | assumes a: "a \<in> span S" and x: "\<And>y. y \<in> S \<Longrightarrow> orthogonal x y" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1444 | shows "orthogonal x a" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1445 | by (metis a orthogonal_clauses(1,2,4) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1446 | span_induct_alt x) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1447 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1448 | proposition Gram_Schmidt_step: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1449 | fixes S :: "'a::euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1450 | assumes S: "pairwise orthogonal S" and x: "x \<in> span S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1451 | shows "orthogonal x (a - (\<Sum>b\<in>S. (b \<bullet> a / (b \<bullet> b)) *\<^sub>R b))" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1452 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1453 | have "finite S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1454 | by (simp add: S pairwise_orthogonal_imp_finite) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1455 | have "orthogonal (a - (\<Sum>b\<in>S. (b \<bullet> a / (b \<bullet> b)) *\<^sub>R b)) x" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1456 | if "x \<in> S" for x | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1457 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1458 | have "a \<bullet> x = (\<Sum>y\<in>S. if y = x then y \<bullet> a else 0)" | 
| 71044 | 1459 | by (simp add: \<open>finite S\<close> inner_commute that) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1460 | also have "... = (\<Sum>b\<in>S. b \<bullet> a * (b \<bullet> x) / (b \<bullet> b))" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1461 | apply (rule sum.cong [OF refl], simp) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1462 | by (meson S orthogonal_def pairwise_def that) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1463 | finally show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1464 | by (simp add: orthogonal_def algebra_simps inner_sum_left) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1465 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1466 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1467 | using orthogonal_to_span orthogonal_commute x by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1468 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1469 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1470 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1471 | lemma orthogonal_extension_aux: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1472 | fixes S :: "'a::euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1473 | assumes "finite T" "finite S" "pairwise orthogonal S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1474 | shows "\<exists>U. pairwise orthogonal (S \<union> U) \<and> span (S \<union> U) = span (S \<union> T)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1475 | using assms | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1476 | proof (induction arbitrary: S) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1477 | case empty then show ?case | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1478 | by simp (metis sup_bot_right) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1479 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1480 | case (insert a T) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1481 | have 0: "\<And>x y. \<lbrakk>x \<noteq> y; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1482 | using insert by (simp add: pairwise_def orthogonal_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1483 | define a' where "a' = a - (\<Sum>b\<in>S. (b \<bullet> a / (b \<bullet> b)) *\<^sub>R b)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1484 | obtain U where orthU: "pairwise orthogonal (S \<union> insert a' U)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1485 | and spanU: "span (insert a' S \<union> U) = span (insert a' S \<union> T)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1486 | by (rule exE [OF insert.IH [of "insert a' S"]]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1487 | (auto simp: Gram_Schmidt_step a'_def insert.prems orthogonal_commute | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1488 | pairwise_orthogonal_insert span_clauses) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1489 | have orthS: "\<And>x. x \<in> S \<Longrightarrow> a' \<bullet> x = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1490 | apply (simp add: a'_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1491 | using Gram_Schmidt_step [OF \<open>pairwise orthogonal S\<close>] | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1492 | apply (force simp: orthogonal_def inner_commute span_superset [THEN subsetD]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1493 | done | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1494 | have "span (S \<union> insert a' U) = span (insert a' (S \<union> T))" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1495 | using spanU by simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1496 | also have "... = span (insert a (S \<union> T))" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1497 | apply (rule eq_span_insert_eq) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1498 | apply (simp add: a'_def span_neg span_sum span_base span_mul) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1499 | done | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1500 | also have "... = span (S \<union> insert a T)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1501 | by simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1502 | finally show ?case | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1503 | by (rule_tac x="insert a' U" in exI) (use orthU in auto) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1504 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1505 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1506 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1507 | proposition orthogonal_extension: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1508 | fixes S :: "'a::euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1509 | assumes S: "pairwise orthogonal S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1510 | obtains U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> T)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1511 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1512 | obtain B where "finite B" "span B = span T" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1513 | using basis_subspace_exists [of "span T"] subspace_span by metis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1514 | with orthogonal_extension_aux [of B S] | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1515 | obtain U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> B)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1516 | using assms pairwise_orthogonal_imp_finite by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1517 | with \<open>span B = span T\<close> show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1518 | by (rule_tac U=U in that) (auto simp: span_Un) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1519 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1520 | |
| 70136 | 1521 | corollary\<^marker>\<open>tag unimportant\<close> orthogonal_extension_strong: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1522 | fixes S :: "'a::euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1523 | assumes S: "pairwise orthogonal S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1524 |   obtains U where "U \<inter> (insert 0 S) = {}" "pairwise orthogonal (S \<union> U)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1525 | "span (S \<union> U) = span (S \<union> T)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1526 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1527 | obtain U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> T)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1528 | using orthogonal_extension assms by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1529 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1530 | apply (rule_tac U = "U - (insert 0 S)" in that) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1531 | apply blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1532 | apply (force simp: pairwise_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1533 | apply (metis Un_Diff_cancel Un_insert_left span_redundant span_zero) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1534 | done | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1535 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1536 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1537 | subsection\<open>Decomposing a vector into parts in orthogonal subspaces\<close> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1538 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1539 | text\<open>existence of orthonormal basis for a subspace.\<close> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1540 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1541 | lemma orthogonal_spanningset_subspace: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1542 | fixes S :: "'a :: euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1543 | assumes "subspace S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1544 | obtains B where "B \<subseteq> S" "pairwise orthogonal B" "span B = S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1545 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1546 | obtain B where "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1547 | using basis_exists by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1548 |   with orthogonal_extension [of "{}" B]
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1549 | show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1550 | by (metis Un_empty_left assms pairwise_empty span_superset span_subspace that) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1551 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1552 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1553 | lemma orthogonal_basis_subspace: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1554 | fixes S :: "'a :: euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1555 | assumes "subspace S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1556 | obtains B where "0 \<notin> B" "B \<subseteq> S" "pairwise orthogonal B" "independent B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1557 | "card B = dim S" "span B = S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1558 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1559 | obtain B where "B \<subseteq> S" "pairwise orthogonal B" "span B = S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1560 | using assms orthogonal_spanningset_subspace by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1561 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1562 |     apply (rule_tac B = "B - {0}" in that)
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1563 | apply (auto simp: indep_card_eq_dim_span pairwise_subset pairwise_orthogonal_independent elim: pairwise_subset) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1564 | done | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1565 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1566 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1567 | proposition orthonormal_basis_subspace: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1568 | fixes S :: "'a :: euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1569 | assumes "subspace S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1570 | obtains B where "B \<subseteq> S" "pairwise orthogonal B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1571 | and "\<And>x. x \<in> B \<Longrightarrow> norm x = 1" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1572 | and "independent B" "card B = dim S" "span B = S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1573 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1574 | obtain B where "0 \<notin> B" "B \<subseteq> S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1575 | and orth: "pairwise orthogonal B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1576 | and "independent B" "card B = dim S" "span B = S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1577 | by (blast intro: orthogonal_basis_subspace [OF assms]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1578 | have 1: "(\<lambda>x. x /\<^sub>R norm x) ` B \<subseteq> S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1579 | using \<open>span B = S\<close> span_superset span_mul by fastforce | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1580 | have 2: "pairwise orthogonal ((\<lambda>x. x /\<^sub>R norm x) ` B)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1581 | using orth by (force simp: pairwise_def orthogonal_clauses) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1582 | have 3: "\<And>x. x \<in> (\<lambda>x. x /\<^sub>R norm x) ` B \<Longrightarrow> norm x = 1" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1583 | by (metis (no_types, lifting) \<open>0 \<notin> B\<close> image_iff norm_sgn sgn_div_norm) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1584 | have 4: "independent ((\<lambda>x. x /\<^sub>R norm x) ` B)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1585 | by (metis "2" "3" norm_zero pairwise_orthogonal_independent zero_neq_one) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1586 | have "inj_on (\<lambda>x. x /\<^sub>R norm x) B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1587 | proof | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1588 | fix x y | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1589 | assume "x \<in> B" "y \<in> B" "x /\<^sub>R norm x = y /\<^sub>R norm y" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1590 | moreover have "\<And>i. i \<in> B \<Longrightarrow> norm (i /\<^sub>R norm i) = 1" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1591 | using 3 by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1592 | ultimately show "x = y" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1593 | by (metis norm_eq_1 orth orthogonal_clauses(7) orthogonal_commute orthogonal_def pairwise_def zero_neq_one) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1594 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1595 | then have 5: "card ((\<lambda>x. x /\<^sub>R norm x) ` B) = dim S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1596 | by (metis \<open>card B = dim S\<close> card_image) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1597 | have 6: "span ((\<lambda>x. x /\<^sub>R norm x) ` B) = S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1598 | by (metis "1" "4" "5" assms card_eq_dim independent_imp_finite span_subspace) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1599 | show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1600 | by (rule that [OF 1 2 3 4 5 6]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1601 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1602 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1603 | |
| 70136 | 1604 | proposition\<^marker>\<open>tag unimportant\<close> orthogonal_to_subspace_exists_gen: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1605 | fixes S :: "'a :: euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1606 | assumes "span S \<subset> span T" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1607 | obtains x where "x \<noteq> 0" "x \<in> span T" "\<And>y. y \<in> span S \<Longrightarrow> orthogonal x y" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1608 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1609 | obtain B where "B \<subseteq> span S" and orthB: "pairwise orthogonal B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1610 | and "\<And>x. x \<in> B \<Longrightarrow> norm x = 1" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1611 | and "independent B" "card B = dim S" "span B = span S" | 
| 71044 | 1612 | by (rule orthonormal_basis_subspace [of "span S", OF subspace_span]) (auto) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1613 | with assms obtain u where spanBT: "span B \<subseteq> span T" and "u \<notin> span B" "u \<in> span T" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1614 | by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1615 |   obtain C where orthBC: "pairwise orthogonal (B \<union> C)" and spanBC: "span (B \<union> C) = span (B \<union> {u})"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1616 | by (blast intro: orthogonal_extension [OF orthB]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1617 | show thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1618 | proof (cases "C \<subseteq> insert 0 B") | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1619 | case True | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1620 | then have "C \<subseteq> span B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1621 | using span_eq | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1622 | by (metis span_insert_0 subset_trans) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1623 | moreover have "u \<in> span (B \<union> C)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1624 |       using \<open>span (B \<union> C) = span (B \<union> {u})\<close> span_superset by force
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1625 | ultimately show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1626 | using True \<open>u \<notin> span B\<close> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1627 | by (metis Un_insert_left span_insert_0 sup.orderE) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1628 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1629 | case False | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1630 | then obtain x where "x \<in> C" "x \<noteq> 0" "x \<notin> B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1631 | by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1632 | then have "x \<in> span T" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1633 | by (metis (no_types, lifting) Un_insert_right Un_upper2 \<open>u \<in> span T\<close> spanBT spanBC | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1634 | \<open>u \<in> span T\<close> insert_subset span_superset span_mono | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1635 | span_span subsetCE subset_trans sup_bot.comm_neutral) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1636 | moreover have "orthogonal x y" if "y \<in> span B" for y | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1637 | using that | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1638 | proof (rule span_induct) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1639 |       show "subspace {a. orthogonal x a}"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1640 | by (simp add: subspace_orthogonal_to_vector) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1641 | show "\<And>b. b \<in> B \<Longrightarrow> orthogonal x b" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1642 | by (metis Un_iff \<open>x \<in> C\<close> \<open>x \<notin> B\<close> orthBC pairwise_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1643 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1644 | ultimately show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1645 | using \<open>x \<noteq> 0\<close> that \<open>span B = span S\<close> by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1646 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1647 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1648 | |
| 70136 | 1649 | corollary\<^marker>\<open>tag unimportant\<close> orthogonal_to_subspace_exists: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1650 | fixes S :: "'a :: euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1651 |   assumes "dim S < DIM('a)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1652 | obtains x where "x \<noteq> 0" "\<And>y. y \<in> span S \<Longrightarrow> orthogonal x y" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1653 | proof - | 
| 71044 | 1654 | have "span S \<subset> UNIV" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1655 | by (metis (mono_tags) UNIV_I assms inner_eq_zero_iff less_le lowdim_subset_hyperplane | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1656 | mem_Collect_eq top.extremum_strict top.not_eq_extremum) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1657 | with orthogonal_to_subspace_exists_gen [of S UNIV] that show ?thesis | 
| 71044 | 1658 | by (auto) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1659 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1660 | |
| 70136 | 1661 | corollary\<^marker>\<open>tag unimportant\<close> orthogonal_to_vector_exists: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1662 | fixes x :: "'a :: euclidean_space" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1663 |   assumes "2 \<le> DIM('a)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1664 | obtains y where "y \<noteq> 0" "orthogonal x y" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1665 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1666 |   have "dim {x} < DIM('a)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1667 | using assms by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1668 | then show thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1669 | by (rule orthogonal_to_subspace_exists) (simp add: orthogonal_commute span_base that) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1670 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1671 | |
| 70136 | 1672 | proposition\<^marker>\<open>tag unimportant\<close> orthogonal_subspace_decomp_exists: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1673 | fixes S :: "'a :: euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1674 | obtains y z | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1675 | where "y \<in> span S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1676 | and "\<And>w. w \<in> span S \<Longrightarrow> orthogonal z w" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1677 | and "x = y + z" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1678 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1679 | obtain T where "0 \<notin> T" "T \<subseteq> span S" "pairwise orthogonal T" "independent T" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1680 | "card T = dim (span S)" "span T = span S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1681 | using orthogonal_basis_subspace subspace_span by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1682 | let ?a = "\<Sum>b\<in>T. (b \<bullet> x / (b \<bullet> b)) *\<^sub>R b" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1683 | have orth: "orthogonal (x - ?a) w" if "w \<in> span S" for w | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1684 | by (simp add: Gram_Schmidt_step \<open>pairwise orthogonal T\<close> \<open>span T = span S\<close> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1685 | orthogonal_commute that) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1686 | show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1687 | apply (rule_tac y = "?a" and z = "x - ?a" in that) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1688 | apply (meson \<open>T \<subseteq> span S\<close> span_scale span_sum subsetCE) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1689 | apply (fact orth, simp) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1690 | done | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1691 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1692 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1693 | lemma orthogonal_subspace_decomp_unique: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1694 | fixes S :: "'a :: euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1695 | assumes "x + y = x' + y'" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1696 | and ST: "x \<in> span S" "x' \<in> span S" "y \<in> span T" "y' \<in> span T" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1697 | and orth: "\<And>a b. \<lbrakk>a \<in> S; b \<in> T\<rbrakk> \<Longrightarrow> orthogonal a b" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1698 | shows "x = x' \<and> y = y'" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1699 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1700 | have "x + y - y' = x'" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1701 | by (simp add: assms) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1702 | moreover have "\<And>a b. \<lbrakk>a \<in> span S; b \<in> span T\<rbrakk> \<Longrightarrow> orthogonal a b" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1703 | by (meson orth orthogonal_commute orthogonal_to_span) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1704 | ultimately have "0 = x' - x" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1705 | by (metis (full_types) add_diff_cancel_left' ST diff_right_commute orthogonal_clauses(10) orthogonal_clauses(5) orthogonal_self) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1706 | with assms show ?thesis by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1707 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1708 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1709 | lemma vector_in_orthogonal_spanningset: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1710 | fixes a :: "'a::euclidean_space" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1711 | obtains S where "a \<in> S" "pairwise orthogonal S" "span S = UNIV" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1712 | by (metis UNIV_I Un_iff empty_iff insert_subset orthogonal_extension pairwise_def | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1713 | pairwise_orthogonal_insert span_UNIV subsetI subset_antisym) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1714 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1715 | lemma vector_in_orthogonal_basis: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1716 | fixes a :: "'a::euclidean_space" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1717 | assumes "a \<noteq> 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1718 | obtains S where "a \<in> S" "0 \<notin> S" "pairwise orthogonal S" "independent S" "finite S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1719 |                   "span S = UNIV" "card S = DIM('a)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1720 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1721 | obtain S where S: "a \<in> S" "pairwise orthogonal S" "span S = UNIV" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1722 | using vector_in_orthogonal_spanningset . | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1723 | show thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1724 | proof | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1725 |     show "pairwise orthogonal (S - {0})"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1726 | using pairwise_mono S(2) by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1727 |     show "independent (S - {0})"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1728 |       by (simp add: \<open>pairwise orthogonal (S - {0})\<close> pairwise_orthogonal_independent)
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1729 |     show "finite (S - {0})"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1730 |       using \<open>independent (S - {0})\<close> independent_imp_finite by blast
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1731 |     show "card (S - {0}) = DIM('a)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1732 | using span_delete_0 [of S] S | 
| 71044 | 1733 |       by (simp add: \<open>independent (S - {0})\<close> indep_card_eq_dim_span)
 | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1734 | qed (use S \<open>a \<noteq> 0\<close> in auto) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1735 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1736 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1737 | lemma vector_in_orthonormal_basis: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1738 | fixes a :: "'a::euclidean_space" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1739 | assumes "norm a = 1" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1740 | obtains S where "a \<in> S" "pairwise orthogonal S" "\<And>x. x \<in> S \<Longrightarrow> norm x = 1" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1741 |     "independent S" "card S = DIM('a)" "span S = UNIV"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1742 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1743 | have "a \<noteq> 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1744 | using assms by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1745 | then obtain S where "a \<in> S" "0 \<notin> S" "finite S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1746 |           and S: "pairwise orthogonal S" "independent S" "span S = UNIV" "card S = DIM('a)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1747 | by (metis vector_in_orthogonal_basis) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1748 | let ?S = "(\<lambda>x. x /\<^sub>R norm x) ` S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1749 | show thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1750 | proof | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1751 | show "a \<in> ?S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1752 | using \<open>a \<in> S\<close> assms image_iff by fastforce | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1753 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1754 | show "pairwise orthogonal ?S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1755 | using \<open>pairwise orthogonal S\<close> by (auto simp: pairwise_def orthogonal_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1756 | show "\<And>x. x \<in> (\<lambda>x. x /\<^sub>R norm x) ` S \<Longrightarrow> norm x = 1" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70707diff
changeset | 1757 | using \<open>0 \<notin> S\<close> by (auto simp: field_split_simps) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1758 | then show "independent ?S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1759 | by (metis \<open>pairwise orthogonal ((\<lambda>x. x /\<^sub>R norm x) ` S)\<close> norm_zero pairwise_orthogonal_independent zero_neq_one) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1760 | have "inj_on (\<lambda>x. x /\<^sub>R norm x) S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1761 | unfolding inj_on_def | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1762 | by (metis (full_types) S(1) \<open>0 \<notin> S\<close> inverse_nonzero_iff_nonzero norm_eq_zero orthogonal_scaleR orthogonal_self pairwise_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1763 |     then show "card ?S = DIM('a)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1764 | by (simp add: card_image S) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1765 | show "span ?S = UNIV" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1766 | by (metis (no_types) \<open>0 \<notin> S\<close> \<open>finite S\<close> \<open>span S = UNIV\<close> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1767 | field_class.field_inverse_zero inverse_inverse_eq less_irrefl span_image_scale | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1768 | zero_less_norm_iff) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1769 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1770 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1771 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1772 | proposition dim_orthogonal_sum: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1773 | fixes A :: "'a::euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1774 | assumes "\<And>x y. \<lbrakk>x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1775 | shows "dim(A \<union> B) = dim A + dim B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1776 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1777 | have 1: "\<And>x y. \<lbrakk>x \<in> span A; y \<in> B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1778 | by (erule span_induct [OF _ subspace_hyperplane2]; simp add: assms) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1779 | have "\<And>x y. \<lbrakk>x \<in> span A; y \<in> span B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1780 | using 1 by (simp add: span_induct [OF _ subspace_hyperplane]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1781 | then have 0: "\<And>x y. \<lbrakk>x \<in> span A; y \<in> span B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1782 | by simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1783 | have "dim(A \<union> B) = dim (span (A \<union> B))" | 
| 71044 | 1784 | by (simp) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1785 | also have "span (A \<union> B) = ((\<lambda>(a, b). a + b) ` (span A \<times> span B))" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1786 | by (auto simp add: span_Un image_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1787 |   also have "dim \<dots> = dim {x + y |x y. x \<in> span A \<and> y \<in> span B}"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1788 | by (auto intro!: arg_cong [where f=dim]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1789 |   also have "... = dim {x + y |x y. x \<in> span A \<and> y \<in> span B} + dim(span A \<inter> span B)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1790 | by (auto simp: dest: 0) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1791 | also have "... = dim (span A) + dim (span B)" | 
| 71044 | 1792 | by (rule dim_sums_Int) (auto) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1793 | also have "... = dim A + dim B" | 
| 71044 | 1794 | by (simp) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1795 | finally show ?thesis . | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1796 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1797 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1798 | lemma dim_subspace_orthogonal_to_vectors: | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1799 | fixes A :: "'a::euclidean_space set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1800 | assumes "subspace A" "subspace B" "A \<subseteq> B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1801 |     shows "dim {y \<in> B. \<forall>x \<in> A. orthogonal x y} + dim A = dim B"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1802 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1803 |   have "dim (span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)) = dim (span B)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1804 | proof (rule arg_cong [where f=dim, OF subset_antisym]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1805 |     show "span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A) \<subseteq> span B"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1806 | by (simp add: \<open>A \<subseteq> B\<close> Collect_restrict span_mono) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1807 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1808 |     have *: "x \<in> span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1809 | if "x \<in> B" for x | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1810 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1811 | obtain y z where "x = y + z" "y \<in> span A" and orth: "\<And>w. w \<in> span A \<Longrightarrow> orthogonal z w" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1812 | using orthogonal_subspace_decomp_exists [of A x] that by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1813 | have "y \<in> span B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1814 | using \<open>y \<in> span A\<close> assms(3) span_mono by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1815 |       then have "z \<in> {a \<in> B. \<forall>x. x \<in> A \<longrightarrow> orthogonal x a}"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1816 | apply simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1817 | using \<open>x = y + z\<close> assms(1) assms(2) orth orthogonal_commute span_add_eq | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1818 | span_eq_iff that by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1819 |       then have z: "z \<in> span {y \<in> B. \<forall>x\<in>A. orthogonal x y}"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1820 | by (meson span_superset subset_iff) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1821 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1822 | apply (auto simp: span_Un image_def \<open>x = y + z\<close> \<open>y \<in> span A\<close>) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1823 | using \<open>y \<in> span A\<close> add.commute by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1824 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1825 |     show "span B \<subseteq> span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)"
 | 
| 71044 | 1826 | by (rule span_minimal) (auto intro: * span_minimal) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1827 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1828 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1829 | by (metis (no_types, lifting) dim_orthogonal_sum dim_span mem_Collect_eq | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1830 | orthogonal_commute orthogonal_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1831 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69674diff
changeset | 1832 | |
| 70688 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1833 | subsection\<open>Linear functions are (uniformly) continuous on any set\<close> | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1834 | |
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1835 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Topological properties of linear functions\<close> | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1836 | |
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1837 | lemma linear_lim_0: | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1838 | assumes "bounded_linear f" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1839 | shows "(f \<longlongrightarrow> 0) (at (0))" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1840 | proof - | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1841 | interpret f: bounded_linear f by fact | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1842 | have "(f \<longlongrightarrow> f 0) (at 0)" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1843 | using tendsto_ident_at by (rule f.tendsto) | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1844 | then show ?thesis unfolding f.zero . | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1845 | qed | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1846 | |
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1847 | lemma linear_continuous_at: | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1848 | assumes "bounded_linear f" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1849 | shows "continuous (at a) f" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1850 | unfolding continuous_at using assms | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1851 | apply (rule bounded_linear.tendsto) | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1852 | apply (rule tendsto_ident_at) | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1853 | done | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1854 | |
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1855 | lemma linear_continuous_within: | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1856 | "bounded_linear f \<Longrightarrow> continuous (at x within s) f" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1857 | using continuous_at_imp_continuous_at_within linear_continuous_at by blast | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1858 | |
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1859 | lemma linear_continuous_on: | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1860 | "bounded_linear f \<Longrightarrow> continuous_on s f" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1861 | using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1862 | |
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1863 | lemma Lim_linear: | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1864 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" and h :: "'b \<Rightarrow> 'c::real_normed_vector" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1865 | assumes "(f \<longlongrightarrow> l) F" "linear h" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1866 | shows "((\<lambda>x. h(f x)) \<longlongrightarrow> h l) F" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1867 | proof - | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1868 | obtain B where B: "B > 0" "\<And>x. norm (h x) \<le> B * norm x" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1869 | using linear_bounded_pos [OF \<open>linear h\<close>] by blast | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1870 | show ?thesis | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1871 | unfolding tendsto_iff | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1872 | proof (intro allI impI) | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1873 | show "\<forall>\<^sub>F x in F. dist (h (f x)) (h l) < e" if "e > 0" for e | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1874 | proof - | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1875 | have "\<forall>\<^sub>F x in F. dist (f x) l < e/B" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1876 | by (simp add: \<open>0 < B\<close> assms(1) tendstoD that) | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1877 | then show ?thesis | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1878 | unfolding dist_norm | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1879 | proof (rule eventually_mono) | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1880 | show "norm (h (f x) - h l) < e" if "norm (f x - l) < e / B" for x | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1881 | using that B | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70707diff
changeset | 1882 | apply (simp add: field_split_simps) | 
| 71044 | 1883 | by (metis \<open>linear h\<close> le_less_trans linear_diff) | 
| 70688 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1884 | qed | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1885 | qed | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1886 | qed | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1887 | qed | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1888 | |
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1889 | lemma linear_continuous_compose: | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1890 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" and g :: "'b \<Rightarrow> 'c::real_normed_vector" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1891 | assumes "continuous F f" "linear g" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1892 | shows "continuous F (\<lambda>x. g(f x))" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1893 | using assms unfolding continuous_def by (rule Lim_linear) | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1894 | |
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1895 | lemma linear_continuous_on_compose: | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1896 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" and g :: "'b \<Rightarrow> 'c::real_normed_vector" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1897 | assumes "continuous_on S f" "linear g" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1898 | shows "continuous_on S (\<lambda>x. g(f x))" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1899 | using assms by (simp add: continuous_on_eq_continuous_within linear_continuous_compose) | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1900 | |
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1901 | text\<open>Also bilinear functions, in composition form\<close> | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1902 | |
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1903 | lemma bilinear_continuous_compose: | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1904 | fixes h :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::real_normed_vector" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1905 | assumes "continuous F f" "continuous F g" "bilinear h" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1906 | shows "continuous F (\<lambda>x. h (f x) (g x))" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1907 | using assms bilinear_conv_bounded_bilinear bounded_bilinear.continuous by blast | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1908 | |
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1909 | lemma bilinear_continuous_on_compose: | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1910 | fixes h :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::real_normed_vector" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1911 | and f :: "'d::t2_space \<Rightarrow> 'a" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1912 | assumes "continuous_on S f" "continuous_on S g" "bilinear h" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1913 | shows "continuous_on S (\<lambda>x. h (f x) (g x))" | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1914 | using assms by (simp add: continuous_on_eq_continuous_within bilinear_continuous_compose) | 
| 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1915 | |
| 54776 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
54703diff
changeset | 1916 | end |