src/HOL/Analysis/Operator_Norm.thy
author wenzelm
Fri, 15 Oct 2021 22:00:28 +0200
changeset 74530 823ccd84b879
parent 70136 f03a01a18c6e
permissions -rw-r--r--
revert bbfed17243af, breaks HOL-Proofs extraction;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 61975
diff changeset
     1
(*  Title:      HOL/Analysis/Operator_Norm.thy
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
     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
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
     4
*)
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
     5
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
     6
section \<open>Operator Norm\<close>
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
     7
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
     8
theory Operator_Norm
56319
3bc95e0fa651 minimized imports
huffman
parents: 56223
diff changeset
     9
imports Complex_Main
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
    10
begin
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
    11
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60420
diff changeset
    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
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69607
diff changeset
    14
text\<^marker>\<open>tag important\<close> \<open>%whitespace\<close>
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69607
diff changeset
    15
definition\<^marker>\<open>tag important\<close>
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    16
onorm :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> real" where
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    17
"onorm f = (SUP x. norm (f x) / norm x)"
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    18
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    19
proposition onorm_bound:
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    20
  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
    21
  shows "onorm f \<le> b"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    22
  unfolding onorm_def
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    23
proof (rule cSUP_least)
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    24
  fix x
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    25
  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
    26
    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
    27
qed simp
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    28
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 58877
diff changeset
    29
text \<open>In non-trivial vector spaces, the first assumption is redundant.\<close>
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
    30
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    31
lemma onorm_le:
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    32
  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
    33
  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
    34
  shows "onorm f \<le> b"
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    35
proof (rule onorm_bound [OF _ assms])
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    36
  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
    37
  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
    38
  have "0 \<le> b * norm a"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    39
    by (rule order_trans [OF norm_ge_zero assms])
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 58877
diff changeset
    40
  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
    41
    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
    42
qed
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    43
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    44
lemma le_onorm:
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    45
  assumes "bounded_linear f"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    46
  shows "norm (f x) / norm x \<le> onorm f"
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    47
proof -
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    48
  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
    49
  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
    50
    using f.nonneg_bounded by auto
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    51
  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
    52
    by (clarify, case_tac "x = 0",
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56319
diff changeset
    53
      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
    54
  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
    55
    unfolding bdd_above_def by fast
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    56
  with UNIV_I show ?thesis
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    57
    unfolding onorm_def by (rule cSUP_upper)
53253
220f306f5c4e tuned proofs;
wenzelm
parents: 51475
diff changeset
    58
qed
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
    59
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    60
lemma onorm:
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    61
  assumes "bounded_linear f"
53253
220f306f5c4e tuned proofs;
wenzelm
parents: 51475
diff changeset
    62
  shows "norm (f x) \<le> onorm f * norm x"
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    63
proof -
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    64
  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
    65
  show ?thesis
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    66
  proof (cases)
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    67
    assume "x = 0"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    68
    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
    69
  next
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    70
    assume "x \<noteq> 0"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    71
    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
    72
      by (rule le_onorm [OF assms])
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    73
    then show "norm (f x) \<le> onorm f * norm x"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 58877
diff changeset
    74
      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
    75
  qed
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
    76
qed
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
    77
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    78
lemma onorm_pos_le:
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    79
  assumes f: "bounded_linear f"
53253
220f306f5c4e tuned proofs;
wenzelm
parents: 51475
diff changeset
    80
  shows "0 \<le> onorm f"
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    81
  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
    82
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    83
lemma onorm_zero: "onorm (\<lambda>x. 0) = 0"
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    84
proof (rule order_antisym)
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    85
  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
    86
    by (simp add: onorm_bound)
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    87
  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
    88
    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
    89
qed
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
    90
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    91
lemma onorm_eq_0:
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    92
  assumes f: "bounded_linear f"
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
    93
  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
    94
  using onorm [OF f] by (auto simp: fun_eq_iff [symmetric] onorm_zero)
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
    95
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
    96
lemma onorm_pos_lt:
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
    97
  assumes f: "bounded_linear f"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53253
diff changeset
    98
  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
    99
  by (simp add: less_le onorm_pos_le [OF f] onorm_eq_0 [OF f])
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
   100
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   101
lemma 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: 61808
diff changeset
   102
  by (rule onorm_bound) simp_all
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   103
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   104
lemma 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: 61808
diff changeset
   105
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
   106
  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
   107
  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
   108
  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
   109
    by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   110
  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
   111
    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
   112
  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
   113
qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   114
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   115
lemma onorm_compose:
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   116
  assumes f: "bounded_linear f"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   117
  assumes g: "bounded_linear g"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53253
diff changeset
   118
  shows "onorm (f \<circ> g) \<le> onorm f * onorm g"
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   119
proof (rule onorm_bound)
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   120
  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
   121
    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
   122
next
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   123
  fix x
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   124
  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
   125
    by (rule onorm [OF f])
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   126
  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
   127
    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
   128
  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
   129
    by (simp add: mult.assoc)
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   130
qed
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
   131
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   132
lemma onorm_scaleR_lemma:
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   133
  assumes f: "bounded_linear f"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   134
  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
   135
proof (rule onorm_bound)
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   136
  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
   137
    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
   138
next
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   139
  fix x
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   140
  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
   141
    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
   142
  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
   143
    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
   144
qed
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   145
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   146
lemma onorm_scaleR:
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   147
  assumes f: "bounded_linear f"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   148
  shows "onorm (\<lambda>x. r *\<^sub>R f x) = \<bar>r\<bar> * onorm f"
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   149
proof (cases "r = 0")
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   150
  assume "r \<noteq> 0"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   151
  show ?thesis
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   152
  proof (rule order_antisym)
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   153
    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
   154
      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
   155
  next
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   156
    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
   157
      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
   158
    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
   159
      by (rule onorm_scaleR_lemma)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 58877
diff changeset
   160
    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
   161
      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
   162
  qed
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   163
qed (simp add: onorm_zero)
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
   164
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   165
lemma onorm_scaleR_left_lemma:
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   166
  assumes r: "bounded_linear r"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   167
  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
   168
proof (rule onorm_bound)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   169
  fix x
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   170
  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
   171
    by simp
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   172
  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
   173
    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
   174
  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
   175
    by (simp add: ac_simps)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   176
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
   177
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   178
lemma onorm_scaleR_left:
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   179
  assumes f: "bounded_linear r"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   180
  shows "onorm (\<lambda>x. r x *\<^sub>R f) = onorm r * norm f"
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   181
proof (cases "f = 0")
61915
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   182
  assume "f \<noteq> 0"
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   183
  show ?thesis
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   184
  proof (rule order_antisym)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   185
    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
   186
      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
   187
  next
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   188
    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
   189
      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
   190
    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
   191
      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
   192
    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
   193
    have "onorm r \<le> onorm (\<lambda>x. r x * norm f) * inverse (norm f)"
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61915
diff changeset
   194
      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
   195
      by (simp add: inverse_eq_divide)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   196
    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
   197
      by (rule onorm_bound)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   198
        (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
   199
    finally show "onorm r * norm f \<le> onorm (\<lambda>x. r x *\<^sub>R f)"
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61915
diff changeset
   200
      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
   201
      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
   202
  qed
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   203
qed (simp add: onorm_zero)
e9812a95d108 theory for type of bounded linear functions; differentiation under the integral sign
immler
parents: 61808
diff changeset
   204
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   205
lemma onorm_neg:
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
   206
  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
   207
  unfolding onorm_def by simp
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
   208
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   209
lemma onorm_triangle:
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   210
  assumes f: "bounded_linear f"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   211
  assumes g: "bounded_linear g"
53253
220f306f5c4e tuned proofs;
wenzelm
parents: 51475
diff changeset
   212
  shows "onorm (\<lambda>x. f x + g x) \<le> onorm f + onorm g"
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   213
proof (rule onorm_bound)
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   214
  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
   215
    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
   216
next
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   217
  fix x
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   218
  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
   219
    by (rule norm_triangle_ineq)
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   220
  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
   221
    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
   222
  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
   223
    by (simp only: distrib_right)
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   224
qed
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
   225
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   226
lemma onorm_triangle_le:
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   227
  assumes "bounded_linear f"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   228
  assumes "bounded_linear g"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   229
  assumes "onorm f + onorm g \<le> e"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53253
diff changeset
   230
  shows "onorm (\<lambda>x. f x + g x) \<le> e"
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   231
  using assms by (rule onorm_triangle [THEN order_trans])
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
   232
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   233
lemma onorm_triangle_lt:
56223
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   234
  assumes "bounded_linear f"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   235
  assumes "bounded_linear g"
7696903b9e61 generalize theory of operator norms to work with class real_normed_vector
huffman
parents: 54263
diff changeset
   236
  assumes "onorm f + onorm g < e"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53253
diff changeset
   237
  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
   238
  using assms by (rule onorm_triangle [THEN order_le_less_trans])
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
   239
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   240
lemma onorm_sum:
67685
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 63627
diff changeset
   241
  assumes "finite S"
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 63627
diff changeset
   242
  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: 63627
diff changeset
   243
  shows "onorm (\<lambda>x. sum (\<lambda>s. f s x) S) \<le> sum (\<lambda>s. onorm (f s)) S"
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   244
  using assms
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 68833
diff changeset
   245
  by (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: 63627
diff changeset
   246
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 63627
diff changeset
   247
lemmas onorm_sum_le = onorm_sum[THEN order_trans]
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 63627
diff changeset
   248
36581
bbea7f52e8e1 move operator norm stuff to new theory file
huffman
parents:
diff changeset
   249
end