author | eberlm <eberlm@in.tum.de> |
Thu, 17 Aug 2017 18:19:16 +0200 | |
changeset 66448 | 97ad7a583457 |
parent 63627 | 6ddb43c6b711 |
child 67685 | bdff8bf0a75b |
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:
54263
diff
changeset
|
3 |
Author: Brian Huffman |
36581 | 4 |
*) |
5 |
||
60420 | 6 |
section \<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:
54263
diff
changeset
|
13 |
|
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
14 |
definition onorm :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> real" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
16 |
|
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
17 |
lemma onorm_bound: |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
19 |
shows "onorm f \<le> b" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
20 |
unfolding onorm_def |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
21 |
proof (rule cSUP_least) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
22 |
fix x |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
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:
54263
diff
changeset
|
25 |
qed simp |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
26 |
|
60420 | 27 |
text \<open>In non-trivial vector spaces, the first assumption is redundant.\<close> |
36581 | 28 |
|
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
29 |
lemma onorm_le: |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
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:
54263
diff
changeset
|
32 |
shows "onorm f \<le> b" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
33 |
proof (rule onorm_bound [OF _ assms]) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
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:
54263
diff
changeset
|
36 |
have "0 \<le> b * norm a" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
39 |
by (simp add: zero_le_mult_iff) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
40 |
qed |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
41 |
|
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
42 |
lemma le_onorm: |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
43 |
assumes "bounded_linear f" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
44 |
shows "norm (f x) / norm x \<le> onorm f" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
45 |
proof - |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
46 |
interpret f: bounded_linear f by fact |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
48 |
using f.nonneg_bounded by auto |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
50 |
by (clarify, case_tac "x = 0", |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
56319
diff
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:
54263
diff
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:
54263
diff
changeset
|
53 |
unfolding bdd_above_def by fast |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
54 |
with UNIV_I show ?thesis |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
55 |
unfolding onorm_def by (rule cSUP_upper) |
53253 | 56 |
qed |
36581 | 57 |
|
58 |
lemma onorm: |
|
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
59 |
assumes "bounded_linear f" |
53253 | 60 |
shows "norm (f x) \<le> onorm f * norm x" |
61 |
proof - |
|
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
62 |
interpret f: bounded_linear f by fact |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
63 |
show ?thesis |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
64 |
proof (cases) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
65 |
assume "x = 0" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
67 |
next |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
68 |
assume "x \<noteq> 0" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
70 |
by (rule le_onorm [OF assms]) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
73 |
qed |
36581 | 74 |
qed |
75 |
||
53253 | 76 |
lemma onorm_pos_le: |
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
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:
54263
diff
changeset
|
80 |
|
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
81 |
lemma onorm_zero: "onorm (\<lambda>x. 0) = 0" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
82 |
proof (rule order_antisym) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
83 |
show "onorm (\<lambda>x. 0) \<le> 0" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
84 |
by (simp add: onorm_bound) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
85 |
show "0 \<le> onorm (\<lambda>x. 0)" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
87 |
qed |
36581 | 88 |
|
53253 | 89 |
lemma onorm_eq_0: |
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
92 |
using onorm [OF f] by (auto simp: fun_eq_iff [symmetric] onorm_zero) |
36581 | 93 |
|
53253 | 94 |
lemma onorm_pos_lt: |
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
97 |
by (simp add: less_le onorm_pos_le [OF f] onorm_eq_0 [OF f]) |
36581 | 98 |
|
61915
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
99 |
lemma onorm_id_le: "onorm (\<lambda>x. x) \<le> 1" |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
100 |
by (rule onorm_bound) simp_all |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
101 |
|
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
102 |
lemma onorm_id: "onorm (\<lambda>x. x::'a::{real_normed_vector, perfect_space}) = 1" |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
103 |
proof (rule antisym[OF onorm_id_le]) |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
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:
61808
diff
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:
61808
diff
changeset
|
106 |
hence "1 \<le> norm x / norm x" |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
107 |
by simp |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
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:
61808
diff
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:
61808
diff
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:
61808
diff
changeset
|
111 |
qed |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
112 |
|
36581 | 113 |
lemma onorm_compose: |
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
114 |
assumes f: "bounded_linear f" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
115 |
assumes g: "bounded_linear g" |
53688 | 116 |
shows "onorm (f \<circ> g) \<le> onorm f * onorm g" |
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
117 |
proof (rule onorm_bound) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
118 |
show "0 \<le> onorm f * onorm g" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
120 |
next |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
121 |
fix x |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
123 |
by (rule onorm [OF f]) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
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:
54263
diff
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:
56319
diff
changeset
|
127 |
by (simp add: mult.assoc) |
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
128 |
qed |
36581 | 129 |
|
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
130 |
lemma onorm_scaleR_lemma: |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
131 |
assumes f: "bounded_linear f" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
133 |
proof (rule onorm_bound) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
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:
54263
diff
changeset
|
136 |
next |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
137 |
fix x |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
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:
54263
diff
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:
56319
diff
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:
54263
diff
changeset
|
142 |
qed |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
143 |
|
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
144 |
lemma onorm_scaleR: |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
145 |
assumes f: "bounded_linear f" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
146 |
shows "onorm (\<lambda>x. r *\<^sub>R f x) = \<bar>r\<bar> * onorm f" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
147 |
proof (cases "r = 0") |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
148 |
assume "r \<noteq> 0" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
149 |
show ?thesis |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
150 |
proof (rule order_antisym) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
152 |
using f by (rule onorm_scaleR_lemma) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
153 |
next |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
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:
54263
diff
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:
54263
diff
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:
56319
diff
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:
54263
diff
changeset
|
160 |
qed |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
161 |
qed (simp add: onorm_zero) |
36581 | 162 |
|
61915
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
163 |
lemma onorm_scaleR_left_lemma: |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
164 |
assumes r: "bounded_linear r" |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
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:
61808
diff
changeset
|
166 |
proof (rule onorm_bound) |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
167 |
fix x |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
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:
61808
diff
changeset
|
169 |
by simp |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
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:
61808
diff
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:
61808
diff
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:
61808
diff
changeset
|
173 |
by (simp add: ac_simps) |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
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:
61808
diff
changeset
|
175 |
|
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
176 |
lemma onorm_scaleR_left: |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
177 |
assumes f: "bounded_linear r" |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
178 |
shows "onorm (\<lambda>x. r x *\<^sub>R f) = onorm r * norm f" |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
179 |
proof (cases "f = 0") |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
180 |
assume "f \<noteq> 0" |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
181 |
show ?thesis |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
182 |
proof (rule order_antisym) |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
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:
61808
diff
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:
61808
diff
changeset
|
185 |
next |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
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:
61808
diff
changeset
|
187 |
by (metis bounded_linear_scaleR_const f) |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
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:
61808
diff
changeset
|
189 |
by (metis bounded_linear_mult_const f) |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
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:
61808
diff
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:
61808
diff
changeset
|
193 |
by (simp add: inverse_eq_divide) |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
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:
61808
diff
changeset
|
195 |
by (rule onorm_bound) |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
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:
61808
diff
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:
61808
diff
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:
61808
diff
changeset
|
200 |
qed |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
201 |
qed (simp add: onorm_zero) |
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61808
diff
changeset
|
202 |
|
53253 | 203 |
lemma 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:
54263
diff
changeset
|
205 |
unfolding onorm_def by simp |
36581 | 206 |
|
207 |
lemma onorm_triangle: |
|
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
208 |
assumes f: "bounded_linear f" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
209 |
assumes g: "bounded_linear g" |
53253 | 210 |
shows "onorm (\<lambda>x. f x + g x) \<le> onorm f + onorm g" |
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
211 |
proof (rule onorm_bound) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
212 |
show "0 \<le> onorm f + onorm g" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
214 |
next |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
215 |
fix x |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
217 |
by (rule norm_triangle_ineq) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
219 |
by (intro add_mono onorm f g) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
221 |
by (simp only: distrib_right) |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
222 |
qed |
36581 | 223 |
|
53253 | 224 |
lemma onorm_triangle_le: |
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
225 |
assumes "bounded_linear f" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
226 |
assumes "bounded_linear g" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
227 |
assumes "onorm f + onorm g \<le> e" |
53688 | 228 |
shows "onorm (\<lambda>x. f x + g x) \<le> e" |
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
229 |
using assms by (rule onorm_triangle [THEN order_trans]) |
36581 | 230 |
|
53253 | 231 |
lemma onorm_triangle_lt: |
56223
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
232 |
assumes "bounded_linear f" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
changeset
|
233 |
assumes "bounded_linear g" |
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
huffman
parents:
54263
diff
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:
54263
diff
changeset
|
236 |
using assms by (rule onorm_triangle [THEN order_le_less_trans]) |
36581 | 237 |
|
238 |
end |