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