| author | haftmann | 
| Thu, 20 May 2010 16:35:52 +0200 | |
| changeset 37020 | 6c699a8e6927 | 
| parent 36593 | fb69c8cd27bd | 
| child 37489 | 44e42d392c6e | 
| permissions | -rw-r--r-- | 
| 36581 | 1 | (* Title: Library/Operator_Norm.thy | 
| 2 | Author: Amine Chaieb, University of Cambridge | |
| 3 | *) | |
| 4 | ||
| 5 | header {* Operator Norm *}
 | |
| 6 | ||
| 7 | theory Operator_Norm | |
| 8 | imports Euclidean_Space | |
| 9 | begin | |
| 10 | ||
| 11 | definition "onorm f = Sup {norm (f x)| x. norm x = 1}"
 | |
| 12 | ||
| 13 | lemma norm_bound_generalize: | |
| 14 | fixes f:: "real ^'n \<Rightarrow> real^'m" | |
| 15 | assumes lf: "linear f" | |
| 16 | shows "(\<forall>x. norm x = 1 \<longrightarrow> norm (f x) \<le> b) \<longleftrightarrow> (\<forall>x. norm (f x) \<le> b * norm x)" (is "?lhs \<longleftrightarrow> ?rhs") | |
| 17 | proof- | |
| 18 |   {assume H: ?rhs
 | |
| 19 |     {fix x :: "real^'n" assume x: "norm x = 1"
 | |
| 20 | from H[rule_format, of x] x have "norm (f x) \<le> b" by simp} | |
| 21 | then have ?lhs by blast } | |
| 22 | ||
| 23 | moreover | |
| 24 |   {assume H: ?lhs
 | |
| 25 | from H[rule_format, of "basis arbitrary"] | |
| 26 | have bp: "b \<ge> 0" using norm_ge_zero[of "f (basis arbitrary)"] | |
| 27 | by (auto simp add: norm_basis elim: order_trans [OF norm_ge_zero]) | |
| 28 |     {fix x :: "real ^'n"
 | |
| 29 |       {assume "x = 0"
 | |
| 30 | then have "norm (f x) \<le> b * norm x" by (simp add: linear_0[OF lf] bp)} | |
| 31 | moreover | |
| 32 |       {assume x0: "x \<noteq> 0"
 | |
| 33 | hence n0: "norm x \<noteq> 0" by (metis norm_eq_zero) | |
| 34 | let ?c = "1/ norm x" | |
| 36593 
fb69c8cd27bd
define linear algebra concepts using scaleR instead of (op *s); generalized many lemmas, though a few theorems that used to work on type int^'n are a bit less general
 huffman parents: 
36581diff
changeset | 35 | have "norm (?c *\<^sub>R x) = 1" using x0 by (simp add: n0) | 
| 
fb69c8cd27bd
define linear algebra concepts using scaleR instead of (op *s); generalized many lemmas, though a few theorems that used to work on type int^'n are a bit less general
 huffman parents: 
36581diff
changeset | 36 | with H have "norm (f (?c *\<^sub>R x)) \<le> b" by blast | 
| 36581 | 37 | hence "?c * norm (f x) \<le> b" | 
| 38 | by (simp add: linear_cmul[OF lf]) | |
| 39 | hence "norm (f x) \<le> b * norm x" | |
| 40 | using n0 norm_ge_zero[of x] by (auto simp add: field_simps)} | |
| 41 | ultimately have "norm (f x) \<le> b * norm x" by blast} | |
| 42 | then have ?rhs by blast} | |
| 43 | ultimately show ?thesis by blast | |
| 44 | qed | |
| 45 | ||
| 46 | lemma onorm: | |
| 47 | fixes f:: "real ^'n \<Rightarrow> real ^'m" | |
| 48 | assumes lf: "linear f" | |
| 49 | shows "norm (f x) <= onorm f * norm x" | |
| 50 | and "\<forall>x. norm (f x) <= b * norm x \<Longrightarrow> onorm f <= b" | |
| 51 | proof- | |
| 52 |   {
 | |
| 53 |     let ?S = "{norm (f x) |x. norm x = 1}"
 | |
| 54 |     have Se: "?S \<noteq> {}" using  norm_basis by auto
 | |
| 55 | from linear_bounded[OF lf] have b: "\<exists> b. ?S *<= b" | |
| 56 | unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def) | |
| 57 |     {from Sup[OF Se b, unfolded onorm_def[symmetric]]
 | |
| 58 | show "norm (f x) <= onorm f * norm x" | |
| 59 | apply - | |
| 60 | apply (rule spec[where x = x]) | |
| 61 | unfolding norm_bound_generalize[OF lf, symmetric] | |
| 62 | by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)} | |
| 63 |     {
 | |
| 64 | show "\<forall>x. norm (f x) <= b * norm x \<Longrightarrow> onorm f <= b" | |
| 65 | using Sup[OF Se b, unfolded onorm_def[symmetric]] | |
| 66 | unfolding norm_bound_generalize[OF lf, symmetric] | |
| 67 | by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)} | |
| 68 | } | |
| 69 | qed | |
| 70 | ||
| 71 | lemma onorm_pos_le: assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" shows "0 <= onorm f" | |
| 72 | using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "basis arbitrary"], unfolded norm_basis] by simp | |
| 73 | ||
| 74 | lemma onorm_eq_0: assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" | |
| 75 | shows "onorm f = 0 \<longleftrightarrow> (\<forall>x. f x = 0)" | |
| 76 | using onorm[OF lf] | |
| 77 | apply (auto simp add: onorm_pos_le) | |
| 78 | apply atomize | |
| 79 | apply (erule allE[where x="0::real"]) | |
| 80 | using onorm_pos_le[OF lf] | |
| 81 | apply arith | |
| 82 | done | |
| 83 | ||
| 84 | lemma onorm_const: "onorm(\<lambda>x::real^'n. (y::real ^'m)) = norm y" | |
| 85 | proof- | |
| 86 | let ?f = "\<lambda>x::real^'n. (y::real ^ 'm)" | |
| 87 |   have th: "{norm (?f x)| x. norm x = 1} = {norm y}"
 | |
| 88 | by(auto intro: vector_choose_size set_ext) | |
| 89 | show ?thesis | |
| 90 | unfolding onorm_def th | |
| 91 | apply (rule Sup_unique) by (simp_all add: setle_def) | |
| 92 | qed | |
| 93 | ||
| 94 | lemma onorm_pos_lt: assumes lf: "linear (f::real ^ 'n \<Rightarrow> real ^'m)" | |
| 95 | shows "0 < onorm f \<longleftrightarrow> ~(\<forall>x. f x = 0)" | |
| 96 | unfolding onorm_eq_0[OF lf, symmetric] | |
| 97 | using onorm_pos_le[OF lf] by arith | |
| 98 | ||
| 99 | lemma onorm_compose: | |
| 100 | assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" | |
| 101 | and lg: "linear (g::real^'k \<Rightarrow> real^'n)" | |
| 102 | shows "onorm (f o g) <= onorm f * onorm g" | |
| 103 | apply (rule onorm(2)[OF linear_compose[OF lg lf], rule_format]) | |
| 104 | unfolding o_def | |
| 105 | apply (subst mult_assoc) | |
| 106 | apply (rule order_trans) | |
| 107 | apply (rule onorm(1)[OF lf]) | |
| 108 | apply (rule mult_mono1) | |
| 109 | apply (rule onorm(1)[OF lg]) | |
| 110 | apply (rule onorm_pos_le[OF lf]) | |
| 111 | done | |
| 112 | ||
| 113 | lemma onorm_neg_lemma: assumes lf: "linear (f::real ^'n \<Rightarrow> real^'m)" | |
| 114 | shows "onorm (\<lambda>x. - f x) \<le> onorm f" | |
| 115 | using onorm[OF linear_compose_neg[OF lf]] onorm[OF lf] | |
| 116 | unfolding norm_minus_cancel by metis | |
| 117 | ||
| 118 | lemma onorm_neg: assumes lf: "linear (f::real ^'n \<Rightarrow> real^'m)" | |
| 119 | shows "onorm (\<lambda>x. - f x) = onorm f" | |
| 120 | using onorm_neg_lemma[OF lf] onorm_neg_lemma[OF linear_compose_neg[OF lf]] | |
| 121 | by simp | |
| 122 | ||
| 123 | lemma onorm_triangle: | |
| 124 | assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" and lg: "linear g" | |
| 125 | shows "onorm (\<lambda>x. f x + g x) <= onorm f + onorm g" | |
| 126 | apply(rule onorm(2)[OF linear_compose_add[OF lf lg], rule_format]) | |
| 127 | apply (rule order_trans) | |
| 128 | apply (rule norm_triangle_ineq) | |
| 129 | apply (simp add: distrib) | |
| 130 | apply (rule add_mono) | |
| 131 | apply (rule onorm(1)[OF lf]) | |
| 132 | apply (rule onorm(1)[OF lg]) | |
| 133 | done | |
| 134 | ||
| 135 | lemma onorm_triangle_le: "linear (f::real ^'n \<Rightarrow> real ^'m) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) <= e | |
| 136 | \<Longrightarrow> onorm(\<lambda>x. f x + g x) <= e" | |
| 137 | apply (rule order_trans) | |
| 138 | apply (rule onorm_triangle) | |
| 139 | apply assumption+ | |
| 140 | done | |
| 141 | ||
| 142 | lemma onorm_triangle_lt: "linear (f::real ^'n \<Rightarrow> real ^'m) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) < e | |
| 143 | ==> onorm(\<lambda>x. f x + g x) < e" | |
| 144 | apply (rule order_le_less_trans) | |
| 145 | apply (rule onorm_triangle) | |
| 146 | by assumption+ | |
| 147 | ||
| 148 | end |