| author | wenzelm | 
| Wed, 26 Dec 2018 20:57:23 +0100 | |
| changeset 69506 | 7d59af98af29 | 
| parent 68833 | fde093888c16 | 
| child 69518 | bf88364c9e94 | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/Operator_Norm.thy | 
| 36581 | 2 | Author: Amine Chaieb, University of Cambridge | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 3 | Author: Brian Huffman | 
| 36581 | 4 | *) | 
| 5 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 6 | section%important \<open>Operator Norm\<close> | 
| 36581 | 7 | |
| 8 | theory Operator_Norm | |
| 56319 | 9 | imports Complex_Main | 
| 36581 | 10 | begin | 
| 11 | ||
| 61808 | 12 | text \<open>This formulation yields zero if \<open>'a\<close> is the trivial vector space.\<close> | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 13 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 14 | definition%important onorm :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> real"
 | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 15 | where "onorm f = (SUP x. norm (f x) / norm x)" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 16 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 17 | lemma%important onorm_bound: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 18 | assumes "0 \<le> b" and "\<And>x. norm (f x) \<le> b * norm x" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 19 | shows "onorm f \<le> b" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 20 | unfolding onorm_def | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 21 | proof%unimportant (rule cSUP_least) | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 22 | fix x | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 23 | show "norm (f x) / norm x \<le> b" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 24 | using assms by (cases "x = 0") (simp_all add: pos_divide_le_eq) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 25 | qed simp | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 26 | |
| 60420 | 27 | text \<open>In non-trivial vector spaces, the first assumption is redundant.\<close> | 
| 36581 | 28 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 29 | lemma%important onorm_le: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 30 |   fixes f :: "'a::{real_normed_vector, perfect_space} \<Rightarrow> 'b::real_normed_vector"
 | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 31 | assumes "\<And>x. norm (f x) \<le> b * norm x" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 32 | shows "onorm f \<le> b" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 33 | proof%unimportant (rule onorm_bound [OF _ assms]) | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 34 |   have "{0::'a} \<noteq> UNIV" by (metis not_open_singleton open_UNIV)
 | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 35 | then obtain a :: 'a where "a \<noteq> 0" by fast | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 36 | have "0 \<le> b * norm a" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 37 | by (rule order_trans [OF norm_ge_zero assms]) | 
| 60420 | 38 | with \<open>a \<noteq> 0\<close> show "0 \<le> b" | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 39 | by (simp add: zero_le_mult_iff) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 40 | qed | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 41 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 42 | lemma%important le_onorm: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 43 | assumes "bounded_linear f" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 44 | shows "norm (f x) / norm x \<le> onorm f" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 45 | proof%unimportant - | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 46 | interpret f: bounded_linear f by fact | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 47 | obtain b where "0 \<le> b" and "\<forall>x. norm (f x) \<le> norm x * b" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 48 | using f.nonneg_bounded by auto | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 49 | then have "\<forall>x. norm (f x) / norm x \<le> b" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 50 | by (clarify, case_tac "x = 0", | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56319diff
changeset | 51 | simp_all add: f.zero pos_divide_le_eq mult.commute) | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 52 | then have "bdd_above (range (\<lambda>x. norm (f x) / norm x))" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 53 | unfolding bdd_above_def by fast | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 54 | with UNIV_I show ?thesis | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 55 | unfolding onorm_def by (rule cSUP_upper) | 
| 53253 | 56 | qed | 
| 36581 | 57 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 58 | lemma%important onorm: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 59 | assumes "bounded_linear f" | 
| 53253 | 60 | shows "norm (f x) \<le> onorm f * norm x" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 61 | proof%unimportant - | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 62 | interpret f: bounded_linear f by fact | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 63 | show ?thesis | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 64 | proof (cases) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 65 | assume "x = 0" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 66 | then show ?thesis by (simp add: f.zero) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 67 | next | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 68 | assume "x \<noteq> 0" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 69 | have "norm (f x) / norm x \<le> onorm f" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 70 | by (rule le_onorm [OF assms]) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 71 | then show "norm (f x) \<le> onorm f * norm x" | 
| 60420 | 72 | by (simp add: pos_divide_le_eq \<open>x \<noteq> 0\<close>) | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 73 | qed | 
| 36581 | 74 | qed | 
| 75 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 76 | lemma%unimportant onorm_pos_le: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 77 | assumes f: "bounded_linear f" | 
| 53253 | 78 | shows "0 \<le> onorm f" | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 79 | using le_onorm [OF f, where x=0] by simp | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 80 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 81 | lemma%unimportant onorm_zero: "onorm (\<lambda>x. 0) = 0" | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 82 | proof (rule order_antisym) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 83 | show "onorm (\<lambda>x. 0) \<le> 0" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 84 | by (simp add: onorm_bound) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 85 | show "0 \<le> onorm (\<lambda>x. 0)" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 86 | using bounded_linear_zero by (rule onorm_pos_le) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 87 | qed | 
| 36581 | 88 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 89 | lemma%unimportant onorm_eq_0: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 90 | assumes f: "bounded_linear f" | 
| 36581 | 91 | shows "onorm f = 0 \<longleftrightarrow> (\<forall>x. f x = 0)" | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 92 | using onorm [OF f] by (auto simp: fun_eq_iff [symmetric] onorm_zero) | 
| 36581 | 93 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 94 | lemma%unimportant onorm_pos_lt: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 95 | assumes f: "bounded_linear f" | 
| 53688 | 96 | shows "0 < onorm f \<longleftrightarrow> \<not> (\<forall>x. f x = 0)" | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 97 | by (simp add: less_le onorm_pos_le [OF f] onorm_eq_0 [OF f]) | 
| 36581 | 98 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 99 | lemma%unimportant onorm_id_le: "onorm (\<lambda>x. x) \<le> 1" | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 100 | by (rule onorm_bound) simp_all | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 101 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 102 | lemma%unimportant onorm_id: "onorm (\<lambda>x. x::'a::{real_normed_vector, perfect_space}) = 1"
 | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 103 | proof (rule antisym[OF onorm_id_le]) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 104 |   have "{0::'a} \<noteq> UNIV" by (metis not_open_singleton open_UNIV)
 | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 105 | then obtain x :: 'a where "x \<noteq> 0" by fast | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 106 | hence "1 \<le> norm x / norm x" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 107 | by simp | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 108 | also have "\<dots> \<le> onorm (\<lambda>x::'a. x)" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 109 | by (rule le_onorm) (rule bounded_linear_ident) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 110 | finally show "1 \<le> onorm (\<lambda>x::'a. x)" . | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 111 | qed | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 112 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 113 | lemma%important onorm_compose: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 114 | assumes f: "bounded_linear f" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 115 | assumes g: "bounded_linear g" | 
| 53688 | 116 | shows "onorm (f \<circ> g) \<le> onorm f * onorm g" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 117 | proof%unimportant (rule onorm_bound) | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 118 | show "0 \<le> onorm f * onorm g" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 119 | by (intro mult_nonneg_nonneg onorm_pos_le f g) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 120 | next | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 121 | fix x | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 122 | have "norm (f (g x)) \<le> onorm f * norm (g x)" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 123 | by (rule onorm [OF f]) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 124 | also have "onorm f * norm (g x) \<le> onorm f * (onorm g * norm x)" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 125 | by (rule mult_left_mono [OF onorm [OF g] onorm_pos_le [OF f]]) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 126 | finally show "norm ((f \<circ> g) x) \<le> onorm f * onorm g * norm x" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56319diff
changeset | 127 | by (simp add: mult.assoc) | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 128 | qed | 
| 36581 | 129 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 130 | lemma%unimportant onorm_scaleR_lemma: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 131 | assumes f: "bounded_linear f" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 132 | shows "onorm (\<lambda>x. r *\<^sub>R f x) \<le> \<bar>r\<bar> * onorm f" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 133 | proof (rule onorm_bound) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 134 | show "0 \<le> \<bar>r\<bar> * onorm f" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 135 | by (intro mult_nonneg_nonneg onorm_pos_le abs_ge_zero f) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 136 | next | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 137 | fix x | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 138 | have "\<bar>r\<bar> * norm (f x) \<le> \<bar>r\<bar> * (onorm f * norm x)" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 139 | by (intro mult_left_mono onorm abs_ge_zero f) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 140 | then show "norm (r *\<^sub>R f x) \<le> \<bar>r\<bar> * onorm f * norm x" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56319diff
changeset | 141 | by (simp only: norm_scaleR mult.assoc) | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 142 | qed | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 143 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 144 | lemma%important onorm_scaleR: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 145 | assumes f: "bounded_linear f" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 146 | shows "onorm (\<lambda>x. r *\<^sub>R f x) = \<bar>r\<bar> * onorm f" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 147 | proof%unimportant (cases "r = 0") | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 148 | assume "r \<noteq> 0" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 149 | show ?thesis | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 150 | proof (rule order_antisym) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 151 | show "onorm (\<lambda>x. r *\<^sub>R f x) \<le> \<bar>r\<bar> * onorm f" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 152 | using f by (rule onorm_scaleR_lemma) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 153 | next | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 154 | have "bounded_linear (\<lambda>x. r *\<^sub>R f x)" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 155 | using bounded_linear_scaleR_right f by (rule bounded_linear_compose) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 156 | then have "onorm (\<lambda>x. inverse r *\<^sub>R r *\<^sub>R f x) \<le> \<bar>inverse r\<bar> * onorm (\<lambda>x. r *\<^sub>R f x)" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 157 | by (rule onorm_scaleR_lemma) | 
| 60420 | 158 | with \<open>r \<noteq> 0\<close> show "\<bar>r\<bar> * onorm f \<le> onorm (\<lambda>x. r *\<^sub>R f x)" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56319diff
changeset | 159 | by (simp add: inverse_eq_divide pos_le_divide_eq mult.commute) | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 160 | qed | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 161 | qed (simp add: onorm_zero) | 
| 36581 | 162 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 163 | lemma%unimportant onorm_scaleR_left_lemma: | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 164 | assumes r: "bounded_linear r" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 165 | shows "onorm (\<lambda>x. r x *\<^sub>R f) \<le> onorm r * norm f" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 166 | proof (rule onorm_bound) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 167 | fix x | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 168 | have "norm (r x *\<^sub>R f) = norm (r x) * norm f" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 169 | by simp | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 170 | also have "\<dots> \<le> onorm r * norm x * norm f" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 171 | by (intro mult_right_mono onorm r norm_ge_zero) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 172 | finally show "norm (r x *\<^sub>R f) \<le> onorm r * norm f * norm x" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 173 | by (simp add: ac_simps) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 174 | qed (intro mult_nonneg_nonneg norm_ge_zero onorm_pos_le r) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 175 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 176 | lemma%important onorm_scaleR_left: | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 177 | assumes f: "bounded_linear r" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 178 | shows "onorm (\<lambda>x. r x *\<^sub>R f) = onorm r * norm f" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 179 | proof%unimportant (cases "f = 0") | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 180 | assume "f \<noteq> 0" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 181 | show ?thesis | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 182 | proof (rule order_antisym) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 183 | show "onorm (\<lambda>x. r x *\<^sub>R f) \<le> onorm r * norm f" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 184 | using f by (rule onorm_scaleR_left_lemma) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 185 | next | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 186 | have bl1: "bounded_linear (\<lambda>x. r x *\<^sub>R f)" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 187 | by (metis bounded_linear_scaleR_const f) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 188 | have "bounded_linear (\<lambda>x. r x * norm f)" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 189 | by (metis bounded_linear_mult_const f) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 190 | from onorm_scaleR_left_lemma[OF this, of "inverse (norm f)"] | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 191 | have "onorm r \<le> onorm (\<lambda>x. r x * norm f) * inverse (norm f)" | 
| 61975 | 192 | using \<open>f \<noteq> 0\<close> | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 193 | by (simp add: inverse_eq_divide) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 194 | also have "onorm (\<lambda>x. r x * norm f) \<le> onorm (\<lambda>x. r x *\<^sub>R f)" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 195 | by (rule onorm_bound) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 196 | (auto simp: abs_mult bl1 onorm_pos_le intro!: order_trans[OF _ onorm]) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 197 | finally show "onorm r * norm f \<le> onorm (\<lambda>x. r x *\<^sub>R f)" | 
| 61975 | 198 | using \<open>f \<noteq> 0\<close> | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 199 | by (simp add: inverse_eq_divide pos_le_divide_eq mult.commute) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 200 | qed | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 201 | qed (simp add: onorm_zero) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61808diff
changeset | 202 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 203 | lemma%unimportant onorm_neg: | 
| 36581 | 204 | shows "onorm (\<lambda>x. - f x) = onorm f" | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 205 | unfolding onorm_def by simp | 
| 36581 | 206 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 207 | lemma%important onorm_triangle: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 208 | assumes f: "bounded_linear f" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 209 | assumes g: "bounded_linear g" | 
| 53253 | 210 | shows "onorm (\<lambda>x. f x + g x) \<le> onorm f + onorm g" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 211 | proof%unimportant (rule onorm_bound) | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 212 | show "0 \<le> onorm f + onorm g" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 213 | by (intro add_nonneg_nonneg onorm_pos_le f g) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 214 | next | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 215 | fix x | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 216 | have "norm (f x + g x) \<le> norm (f x) + norm (g x)" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 217 | by (rule norm_triangle_ineq) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 218 | also have "norm (f x) + norm (g x) \<le> onorm f * norm x + onorm g * norm x" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 219 | by (intro add_mono onorm f g) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 220 | finally show "norm (f x + g x) \<le> (onorm f + onorm g) * norm x" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 221 | by (simp only: distrib_right) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 222 | qed | 
| 36581 | 223 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 224 | lemma%important onorm_triangle_le: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 225 | assumes "bounded_linear f" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 226 | assumes "bounded_linear g" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 227 | assumes "onorm f + onorm g \<le> e" | 
| 53688 | 228 | shows "onorm (\<lambda>x. f x + g x) \<le> e" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 229 | using%unimportant assms by%unimportant (rule onorm_triangle [THEN order_trans]) | 
| 36581 | 230 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 231 | lemma%unimportant onorm_triangle_lt: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 232 | assumes "bounded_linear f" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 233 | assumes "bounded_linear g" | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 234 | assumes "onorm f + onorm g < e" | 
| 53688 | 235 | shows "onorm (\<lambda>x. f x + g x) < e" | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
54263diff
changeset | 236 | using assms by (rule onorm_triangle [THEN order_le_less_trans]) | 
| 36581 | 237 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 238 | lemma%important onorm_sum: | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
63627diff
changeset | 239 | assumes "finite S" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
63627diff
changeset | 240 | assumes "\<And>s. s \<in> S \<Longrightarrow> bounded_linear (f s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
63627diff
changeset | 241 | shows "onorm (\<lambda>x. sum (\<lambda>s. f s x) S) \<le> sum (\<lambda>s. onorm (f s)) S" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 242 | using%unimportant assms | 
| 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
67685diff
changeset | 243 | by%unimportant (induction) (auto simp: onorm_zero intro!: onorm_triangle_le bounded_linear_sum) | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
63627diff
changeset | 244 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
63627diff
changeset | 245 | lemmas onorm_sum_le = onorm_sum[THEN order_trans] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
63627diff
changeset | 246 | |
| 36581 | 247 | end |