| author | hoelzl | 
| Tue, 05 Nov 2013 09:44:59 +0100 | |
| changeset 54260 | 6a967667fd45 | 
| parent 53688 | 63892cfef47f | 
| child 54263 | c4159fe6fa46 | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Multivariate_Analysis/Operator_Norm.thy | 
| 36581 | 2 | Author: Amine Chaieb, University of Cambridge | 
| 3 | *) | |
| 4 | ||
| 5 | header {* Operator Norm *}
 | |
| 6 | ||
| 7 | theory Operator_Norm | |
| 44133 
691c52e900ca
split Linear_Algebra.thy from Euclidean_Space.thy
 huffman parents: 
41959diff
changeset | 8 | imports Linear_Algebra | 
| 36581 | 9 | begin | 
| 10 | ||
| 11 | definition "onorm f = Sup {norm (f x)| x. norm x = 1}"
 | |
| 12 | ||
| 13 | lemma norm_bound_generalize: | |
| 53253 | 14 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 36581 | 15 | assumes lf: "linear f" | 
| 53253 | 16 | shows "(\<forall>x. norm x = 1 \<longrightarrow> norm (f x) \<le> b) \<longleftrightarrow> (\<forall>x. norm (f x) \<le> b * norm x)" | 
| 17 | (is "?lhs \<longleftrightarrow> ?rhs") | |
| 18 | proof | |
| 19 | assume H: ?rhs | |
| 20 |   {
 | |
| 21 | fix x :: "'a" | |
| 22 | assume x: "norm x = 1" | |
| 53688 | 23 | from H[rule_format, of x] x have "norm (f x) \<le> b" | 
| 24 | by simp | |
| 53253 | 25 | } | 
| 26 | then show ?lhs by blast | |
| 27 | next | |
| 28 | assume H: ?lhs | |
| 29 | have bp: "b \<ge> 0" | |
| 30 | apply - | |
| 31 | apply (rule order_trans [OF norm_ge_zero]) | |
| 32 | apply (rule H[rule_format, of "SOME x::'a. x \<in> Basis"]) | |
| 33 | apply (auto intro: SOME_Basis norm_Basis) | |
| 34 | done | |
| 35 |   {
 | |
| 36 | fix x :: "'a" | |
| 37 |     {
 | |
| 38 | assume "x = 0" | |
| 39 | then have "norm (f x) \<le> b * norm x" | |
| 40 | by (simp add: linear_0[OF lf] bp) | |
| 41 | } | |
| 42 | moreover | |
| 43 |     {
 | |
| 44 | assume x0: "x \<noteq> 0" | |
| 53688 | 45 | then have n0: "norm x \<noteq> 0" | 
| 46 | by (metis norm_eq_zero) | |
| 53253 | 47 | let ?c = "1/ norm x" | 
| 53688 | 48 | have "norm (?c *\<^sub>R x) = 1" | 
| 49 | using x0 by (simp add: n0) | |
| 50 | with H have "norm (f (?c *\<^sub>R x)) \<le> b" | |
| 51 | by blast | |
| 53253 | 52 | then have "?c * norm (f x) \<le> b" | 
| 53 | by (simp add: linear_cmul[OF lf]) | |
| 54 | then have "norm (f x) \<le> b * norm x" | |
| 53688 | 55 | using n0 norm_ge_zero[of x] | 
| 56 | by (auto simp add: field_simps) | |
| 53253 | 57 | } | 
| 53688 | 58 | ultimately have "norm (f x) \<le> b * norm x" | 
| 59 | by blast | |
| 53253 | 60 | } | 
| 61 | then show ?rhs by blast | |
| 62 | qed | |
| 36581 | 63 | |
| 64 | lemma onorm: | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36593diff
changeset | 65 | fixes f:: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 36581 | 66 | assumes lf: "linear f" | 
| 53253 | 67 | shows "norm (f x) \<le> onorm f * norm x" | 
| 68 | and "\<forall>x. norm (f x) \<le> b * norm x \<Longrightarrow> onorm f \<le> b" | |
| 69 | proof - | |
| 70 |   let ?S = "{norm (f x) |x. norm x = 1}"
 | |
| 71 | have "norm (f (SOME i. i \<in> Basis)) \<in> ?S" | |
| 72 | by (auto intro!: exI[of _ "SOME i. i \<in> Basis"] norm_Basis SOME_Basis) | |
| 53688 | 73 |   then have Se: "?S \<noteq> {}"
 | 
| 74 | by auto | |
| 53253 | 75 | from linear_bounded[OF lf] have b: "\<exists> b. ?S *<= b" | 
| 53688 | 76 | unfolding norm_bound_generalize[OF lf, symmetric] | 
| 77 | by (auto simp add: setle_def) | |
| 53253 | 78 | from isLub_cSup[OF Se b, unfolded onorm_def[symmetric]] | 
| 53688 | 79 | show "norm (f x) \<le> onorm f * norm x" | 
| 53253 | 80 | apply - | 
| 81 | apply (rule spec[where x = x]) | |
| 82 | unfolding norm_bound_generalize[OF lf, symmetric] | |
| 83 | apply (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def) | |
| 84 | done | |
| 53688 | 85 | show "\<forall>x. norm (f x) \<le> b * norm x \<Longrightarrow> onorm f \<le> b" | 
| 53253 | 86 | using isLub_cSup[OF Se b, unfolded onorm_def[symmetric]] | 
| 87 | unfolding norm_bound_generalize[OF lf, symmetric] | |
| 88 | by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def) | |
| 36581 | 89 | qed | 
| 90 | ||
| 53253 | 91 | lemma onorm_pos_le: | 
| 53688 | 92 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 93 | assumes lf: "linear f" | |
| 53253 | 94 | shows "0 \<le> onorm f" | 
| 95 | using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "SOME i. i \<in> Basis"]] | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
44133diff
changeset | 96 | by (simp add: SOME_Basis) | 
| 36581 | 97 | |
| 53253 | 98 | lemma onorm_eq_0: | 
| 53688 | 99 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 100 | assumes lf: "linear f" | |
| 36581 | 101 | shows "onorm f = 0 \<longleftrightarrow> (\<forall>x. f x = 0)" | 
| 102 | using onorm[OF lf] | |
| 103 | apply (auto simp add: onorm_pos_le) | |
| 104 | apply atomize | |
| 105 | apply (erule allE[where x="0::real"]) | |
| 106 | using onorm_pos_le[OF lf] | |
| 107 | apply arith | |
| 108 | done | |
| 109 | ||
| 53688 | 110 | lemma onorm_const: | 
| 111 | "onorm (\<lambda>x::'a::euclidean_space. y::'b::euclidean_space) = norm y" | |
| 53253 | 112 | proof - | 
| 53688 | 113 | let ?f = "\<lambda>x::'a. y::'b" | 
| 36581 | 114 |   have th: "{norm (?f x)| x. norm x = 1} = {norm y}"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
44133diff
changeset | 115 | by (auto simp: SOME_Basis intro!: exI[of _ "SOME i. i \<in> Basis"]) | 
| 36581 | 116 | show ?thesis | 
| 117 | unfolding onorm_def th | |
| 53253 | 118 | apply (rule cSup_unique) | 
| 119 | apply (simp_all add: setle_def) | |
| 120 | done | |
| 36581 | 121 | qed | 
| 122 | ||
| 53253 | 123 | lemma onorm_pos_lt: | 
| 53688 | 124 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 125 | assumes lf: "linear f" | |
| 126 | shows "0 < onorm f \<longleftrightarrow> \<not> (\<forall>x. f x = 0)" | |
| 36581 | 127 | unfolding onorm_eq_0[OF lf, symmetric] | 
| 128 | using onorm_pos_le[OF lf] by arith | |
| 129 | ||
| 130 | lemma onorm_compose: | |
| 53688 | 131 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 132 | and g :: "'k::euclidean_space \<Rightarrow> 'n::euclidean_space" | |
| 133 | assumes lf: "linear f" | |
| 134 | and lg: "linear g" | |
| 135 | shows "onorm (f \<circ> g) \<le> onorm f * onorm g" | |
| 53253 | 136 | apply (rule onorm(2)[OF linear_compose[OF lg lf], rule_format]) | 
| 137 | unfolding o_def | |
| 138 | apply (subst mult_assoc) | |
| 139 | apply (rule order_trans) | |
| 140 | apply (rule onorm(1)[OF lf]) | |
| 141 | apply (rule mult_left_mono) | |
| 142 | apply (rule onorm(1)[OF lg]) | |
| 143 | apply (rule onorm_pos_le[OF lf]) | |
| 144 | done | |
| 36581 | 145 | |
| 53253 | 146 | lemma onorm_neg_lemma: | 
| 53688 | 147 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 148 | assumes lf: "linear f" | |
| 36581 | 149 | shows "onorm (\<lambda>x. - f x) \<le> onorm f" | 
| 150 | using onorm[OF linear_compose_neg[OF lf]] onorm[OF lf] | |
| 151 | unfolding norm_minus_cancel by metis | |
| 152 | ||
| 53253 | 153 | lemma onorm_neg: | 
| 53688 | 154 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 155 | assumes lf: "linear f" | |
| 36581 | 156 | shows "onorm (\<lambda>x. - f x) = onorm f" | 
| 157 | using onorm_neg_lemma[OF lf] onorm_neg_lemma[OF linear_compose_neg[OF lf]] | |
| 158 | by simp | |
| 159 | ||
| 160 | lemma onorm_triangle: | |
| 53688 | 161 | fixes f g :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 162 | assumes lf: "linear f" | |
| 53253 | 163 | and lg: "linear g" | 
| 164 | shows "onorm (\<lambda>x. f x + g x) \<le> onorm f + onorm g" | |
| 36581 | 165 | apply(rule onorm(2)[OF linear_compose_add[OF lf lg], rule_format]) | 
| 166 | apply (rule order_trans) | |
| 167 | apply (rule norm_triangle_ineq) | |
| 168 | apply (simp add: distrib) | |
| 169 | apply (rule add_mono) | |
| 170 | apply (rule onorm(1)[OF lf]) | |
| 171 | apply (rule onorm(1)[OF lg]) | |
| 172 | done | |
| 173 | ||
| 53253 | 174 | lemma onorm_triangle_le: | 
| 53688 | 175 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 176 | assumes "linear f" | |
| 177 | and "linear g" | |
| 178 | and "onorm f + onorm g \<le> e" | |
| 179 | shows "onorm (\<lambda>x. f x + g x) \<le> e" | |
| 36581 | 180 | apply (rule order_trans) | 
| 181 | apply (rule onorm_triangle) | |
| 53688 | 182 | apply (rule assms)+ | 
| 36581 | 183 | done | 
| 184 | ||
| 53253 | 185 | lemma onorm_triangle_lt: | 
| 53688 | 186 | fixes f g :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 187 | assumes "linear f" | |
| 188 | and "linear g" | |
| 189 | and "onorm f + onorm g < e" | |
| 190 | shows "onorm (\<lambda>x. f x + g x) < e" | |
| 36581 | 191 | apply (rule order_le_less_trans) | 
| 192 | apply (rule onorm_triangle) | |
| 53688 | 193 | apply (rule assms)+ | 
| 53253 | 194 | done | 
| 36581 | 195 | |
| 196 | end |