src/HOL/Hahn_Banach/Function_Norm.thy
author wenzelm
Mon, 02 Nov 2015 11:43:02 +0100
changeset 61540 f92bf6674699
parent 61539 a29295dac1ca
child 61879 e4f9d8f094fe
permissions -rw-r--r--
tuned document;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31795
be3e1cc5005c standard naming conventions for session and theories;
wenzelm
parents: 29291
diff changeset
     1
(*  Title:      HOL/Hahn_Banach/Function_Norm.thy
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     2
    Author:     Gertrud Bauer, TU Munich
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     3
*)
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     4
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 58744
diff changeset
     5
section \<open>The norm of a function\<close>
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
     6
31795
be3e1cc5005c standard naming conventions for session and theories;
wenzelm
parents: 29291
diff changeset
     7
theory Function_Norm
be3e1cc5005c standard naming conventions for session and theories;
wenzelm
parents: 29291
diff changeset
     8
imports Normed_Space Function_Order
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
     9
begin
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    10
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
    11
subsection \<open>Continuous linear forms\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    12
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
    13
text \<open>
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    14
  A linear form \<open>f\<close> on a normed vector space \<open>(V, \<parallel>\<cdot>\<parallel>)\<close> is \<^emph>\<open>continuous\<close>,
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    15
  iff it is bounded, i.e.
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    16
  \begin{center}
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
    17
  \<open>\<exists>c \<in> R. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>\<close>
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    18
  \end{center}
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    19
  In our application no other functions than linear forms are considered, so
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    20
  we can define continuous linear forms as bounded linear forms:
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
    21
\<close>
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    22
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
    23
locale continuous = linearform +
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
    24
  fixes norm :: "_ \<Rightarrow> real"    ("\<parallel>_\<parallel>")
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    25
  assumes bounded: "\<exists>c. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    26
14254
342634f38451 Isar/Locales: <loc>.intro and <loc>.axioms no longer intro? and elim? by
ballarin
parents: 14214
diff changeset
    27
declare continuous.intro [intro?] continuous_axioms.intro [intro?]
342634f38451 Isar/Locales: <loc>.intro and <loc>.axioms no longer intro? and elim? by
ballarin
parents: 14214
diff changeset
    28
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    29
lemma continuousI [intro]:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
    30
  fixes norm :: "_ \<Rightarrow> real"  ("\<parallel>_\<parallel>")
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
    31
  assumes "linearform V f"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    32
  assumes r: "\<And>x. x \<in> V \<Longrightarrow> \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
    33
  shows "continuous V f norm"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    34
proof
23378
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 22931
diff changeset
    35
  show "linearform V f" by fact
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    36
  from r have "\<exists>c. \<forall>x\<in>V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" by blast
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
    37
  then show "continuous_axioms V f norm" ..
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    38
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    39
11472
wenzelm
parents: 10752
diff changeset
    40
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
    41
subsection \<open>The norm of a linear form\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    42
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
    43
text \<open>
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
    44
  The least real number \<open>c\<close> for which holds
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    45
  \begin{center}
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
    46
  \<open>\<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>\<close>
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    47
  \end{center}
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
    48
  is called the \<^emph>\<open>norm\<close> of \<open>f\<close>.
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    49
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    50
  For non-trivial vector spaces \<open>V \<noteq> {0}\<close> the norm can be defined as
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    51
  \begin{center}
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
    52
  \<open>\<parallel>f\<parallel> = \<sup>x \<noteq> 0. \<bar>f x\<bar> / \<parallel>x\<parallel>\<close>
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    53
  \end{center}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    54
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    55
  For the case \<open>V = {0}\<close> the supremum would be taken from an empty set.
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    56
  Since \<open>\<real>\<close> is unbounded, there would be no supremum. To avoid this
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    57
  situation it must be guaranteed that there is an element in this set. This
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    58
  element must be \<open>{} \<ge> 0\<close> so that \<open>fn_norm\<close> has the norm properties.
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    59
  Furthermore it does not have to change the norm in all other cases, so it
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    60
  must be \<open>0\<close>, as all other elements are \<open>{} \<ge> 0\<close>.
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    61
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    62
  Thus we define the set \<open>B\<close> where the supremum is taken from as follows:
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    63
  \begin{center}
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
    64
  \<open>{0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel>. x \<noteq> 0 \<and> x \<in> F}\<close>
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    65
  \end{center}
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    66
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    67
  \<open>fn_norm\<close> is equal to the supremum of \<open>B\<close>, if the supremum exists
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    68
  (otherwise it is undefined).
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
    69
\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    70
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
    71
locale fn_norm =
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
    72
  fixes norm :: "_ \<Rightarrow> real"    ("\<parallel>_\<parallel>")
13547
wenzelm
parents: 13524
diff changeset
    73
  fixes B defines "B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
wenzelm
parents: 13524
diff changeset
    74
  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    75
  defines "\<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    76
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
    77
locale normed_vectorspace_with_fn_norm = normed_vectorspace + fn_norm
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
    78
13547
wenzelm
parents: 13524
diff changeset
    79
lemma (in fn_norm) B_not_empty [intro]: "0 \<in> B V f"
wenzelm
parents: 13524
diff changeset
    80
  by (simp add: B_def)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    81
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
    82
text \<open>
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    83
  The following lemma states that every continuous linear form on a normed
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    84
  space \<open>(V, \<parallel>\<cdot>\<parallel>)\<close> has a function norm.
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
    85
\<close>
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    86
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
    87
lemma (in normed_vectorspace_with_fn_norm) fn_norm_works:
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
    88
  assumes "continuous V f norm"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    89
  shows "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    90
proof -
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
    91
  interpret continuous V f norm by fact
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
    92
  txt \<open>The existence of the supremum is shown using the
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    93
    completeness of the reals. Completeness means, that every
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
    94
    non-empty bounded set of reals has a supremum.\<close>
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    95
  have "\<exists>a. lub (B V f) a"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    96
  proof (rule real_complete)
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
    97
    txt \<open>First we have to show that \<open>B\<close> is non-empty:\<close>
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    98
    have "0 \<in> B V f" ..
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    99
    then show "\<exists>x. x \<in> B V f" ..
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   100
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   101
    txt \<open>Then we have to show that \<open>B\<close> is bounded:\<close>
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   102
    show "\<exists>c. \<forall>y \<in> B V f. y \<le> c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   103
    proof -
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   104
      txt \<open>We know that \<open>f\<close> is bounded by some value \<open>c\<close>.\<close>
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   105
      from bounded obtain c where c: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   106
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   107
      txt \<open>To prove the thesis, we have to show that there is some \<open>b\<close>, such
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   108
        that \<open>y \<le> b\<close> for all \<open>y \<in> B\<close>. Due to the definition of \<open>B\<close> there are
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   109
        two cases.\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   110
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   111
      def b \<equiv> "max c 0"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   112
      have "\<forall>y \<in> B V f. y \<le> b"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   113
      proof
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   114
        fix y assume y: "y \<in> B V f"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   115
        show "y \<le> b"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   116
        proof cases
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   117
          assume "y = 0"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   118
          then show ?thesis unfolding b_def by arith
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   119
        next
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   120
          txt \<open>The second case is \<open>y = \<bar>f x\<bar> / \<parallel>x\<parallel>\<close> for some
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   121
            \<open>x \<in> V\<close> with \<open>x \<noteq> 0\<close>.\<close>
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   122
          assume "y \<noteq> 0"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   123
          with y obtain x where y_rep: "y = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   124
              and x: "x \<in> V" and neq: "x \<noteq> 0"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 32960
diff changeset
   125
            by (auto simp add: B_def divide_inverse)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   126
          from x neq have gt: "0 < \<parallel>x\<parallel>" ..
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   127
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
   128
          txt \<open>The thesis follows by a short calculation using the
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   129
            fact that \<open>f\<close> is bounded.\<close>
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   130
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   131
          note y_rep
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   132
          also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14254
diff changeset
   133
          proof (rule mult_right_mono)
23378
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 22931
diff changeset
   134
            from c x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14254
diff changeset
   135
            from gt have "0 < inverse \<parallel>x\<parallel>" 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14254
diff changeset
   136
              by (rule positive_imp_inverse_positive)
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   137
            then show "0 \<le> inverse \<parallel>x\<parallel>" by (rule order_less_imp_le)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   138
          qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   139
          also have "\<dots> = c * (\<parallel>x\<parallel> * inverse \<parallel>x\<parallel>)"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 50918
diff changeset
   140
            by (rule Groups.mult.assoc)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   141
          also
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   142
          from gt have "\<parallel>x\<parallel> \<noteq> 0" by simp
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   143
          then have "\<parallel>x\<parallel> * inverse \<parallel>x\<parallel> = 1" by simp 
44887
7ca82df6e951 misc tuning and clarification;
wenzelm
parents: 36778
diff changeset
   144
          also have "c * 1 \<le> b" by (simp add: b_def)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   145
          finally show "y \<le> b" .
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   146
        qed
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   147
      qed
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   148
      then show ?thesis ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   149
    qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   150
  qed
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   151
  then show ?thesis unfolding fn_norm_def by (rule the_lubI_ex)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   152
qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   153
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   154
lemma (in normed_vectorspace_with_fn_norm) fn_norm_ub [iff?]:
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
   155
  assumes "continuous V f norm"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   156
  assumes b: "b \<in> B V f"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   157
  shows "b \<le> \<parallel>f\<parallel>\<hyphen>V"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   158
proof -
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
   159
  interpret continuous V f norm by fact
13547
wenzelm
parents: 13524
diff changeset
   160
  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
   161
    using \<open>continuous V f norm\<close> by (rule fn_norm_works)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   162
  from this and b show ?thesis ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   163
qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   164
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   165
lemma (in normed_vectorspace_with_fn_norm) fn_norm_leastB:
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
   166
  assumes "continuous V f norm"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   167
  assumes b: "\<And>b. b \<in> B V f \<Longrightarrow> b \<le> y"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   168
  shows "\<parallel>f\<parallel>\<hyphen>V \<le> y"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   169
proof -
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
   170
  interpret continuous V f norm by fact
13547
wenzelm
parents: 13524
diff changeset
   171
  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
   172
    using \<open>continuous V f norm\<close> by (rule fn_norm_works)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   173
  from this and b show ?thesis ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   174
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   175
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   176
text \<open>The norm of a continuous function is always \<open>\<ge> 0\<close>.\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   177
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   178
lemma (in normed_vectorspace_with_fn_norm) fn_norm_ge_zero [iff]:
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
   179
  assumes "continuous V f norm"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   180
  shows "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   181
proof -
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
   182
  interpret continuous V f norm by fact
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   183
  txt \<open>The function norm is defined as the supremum of \<open>B\<close>.
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   184
    So it is \<open>\<ge> 0\<close> if all elements in \<open>B\<close> are \<open>\<ge>
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   185
    0\<close>, provided the supremum exists and \<open>B\<close> is not empty.\<close>
13547
wenzelm
parents: 13524
diff changeset
   186
  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
   187
    using \<open>continuous V f norm\<close> by (rule fn_norm_works)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   188
  moreover have "0 \<in> B V f" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   189
  ultimately show ?thesis ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   190
qed
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   191
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
   192
text \<open>
61486
3590367b0ce9 tuned document;
wenzelm
parents: 58889
diff changeset
   193
  \<^medskip>
3590367b0ce9 tuned document;
wenzelm
parents: 58889
diff changeset
   194
  The fundamental property of function norms is:
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   195
  \begin{center}
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   196
  \<open>\<bar>f x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>\<close>
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   197
  \end{center}
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
   198
\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   199
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   200
lemma (in normed_vectorspace_with_fn_norm) fn_norm_le_cong:
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
   201
  assumes "continuous V f norm" "linearform V f"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   202
  assumes x: "x \<in> V"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   203
  shows "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   204
proof -
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
   205
  interpret continuous V f norm by fact
29291
d3cc5398bad5 avoid implicit use of prems;
wenzelm
parents: 29252
diff changeset
   206
  interpret linearform V f by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   207
  show ?thesis
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   208
  proof cases
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   209
    assume "x = 0"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   210
    then have "\<bar>f x\<bar> = \<bar>f 0\<bar>" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   211
    also have "f 0 = 0" by rule unfold_locales
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   212
    also have "\<bar>\<dots>\<bar> = 0" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   213
    also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
   214
      using \<open>continuous V f norm\<close> by (rule fn_norm_ge_zero)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   215
    from x have "0 \<le> norm x" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   216
    with a have "0 \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" by (simp add: zero_le_mult_iff)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   217
    finally show "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" .
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   218
  next
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   219
    assume "x \<noteq> 0"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   220
    with x have neq: "\<parallel>x\<parallel> \<noteq> 0" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   221
    then have "\<bar>f x\<bar> = (\<bar>f x\<bar> * inverse \<parallel>x\<parallel>) * \<parallel>x\<parallel>" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   222
    also have "\<dots> \<le>  \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   223
    proof (rule mult_right_mono)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   224
      from x show "0 \<le> \<parallel>x\<parallel>" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   225
      from x and neq have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<in> B V f"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 32960
diff changeset
   226
        by (auto simp add: B_def divide_inverse)
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
   227
      with \<open>continuous V f norm\<close> show "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>V"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   228
        by (rule fn_norm_ub)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   229
    qed
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   230
    finally show ?thesis .
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   231
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   232
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   233
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
   234
text \<open>
61486
3590367b0ce9 tuned document;
wenzelm
parents: 58889
diff changeset
   235
  \<^medskip>
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   236
  The function norm is the least positive real number for which the
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   237
  following inequality holds:
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   238
  \begin{center}
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   239
    \<open>\<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>\<close>
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   240
  \end{center}
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
   241
\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   242
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   243
lemma (in normed_vectorspace_with_fn_norm) fn_norm_least [intro?]:
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
   244
  assumes "continuous V f norm"
50918
3b6417e9f73e tuned proofs;
wenzelm
parents: 46867
diff changeset
   245
  assumes ineq: "\<And>x. x \<in> V \<Longrightarrow> \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" and ge: "0 \<le> c"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   246
  shows "\<parallel>f\<parallel>\<hyphen>V \<le> c"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   247
proof -
46867
0883804b67bb modernized locale expression, with some fluctuation of arguments;
wenzelm
parents: 44887
diff changeset
   248
  interpret continuous V f norm by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   249
  show ?thesis
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   250
  proof (rule fn_norm_leastB [folded B_def fn_norm_def])
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   251
    fix b assume b: "b \<in> B V f"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   252
    show "b \<le> c"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   253
    proof cases
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   254
      assume "b = 0"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   255
      with ge show ?thesis by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   256
    next
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   257
      assume "b \<noteq> 0"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   258
      with b obtain x where b_rep: "b = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   259
        and x_neq: "x \<noteq> 0" and x: "x \<in> V"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 32960
diff changeset
   260
        by (auto simp add: B_def divide_inverse)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   261
      note b_rep
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   262
      also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   263
      proof (rule mult_right_mono)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   264
        have "0 < \<parallel>x\<parallel>" using x x_neq ..
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   265
        then show "0 \<le> inverse \<parallel>x\<parallel>" by simp
50918
3b6417e9f73e tuned proofs;
wenzelm
parents: 46867
diff changeset
   266
        from x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" by (rule ineq)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   267
      qed
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   268
      also have "\<dots> = c"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   269
      proof -
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   270
        from x_neq and x have "\<parallel>x\<parallel> \<noteq> 0" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   271
        then show ?thesis by simp
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   272
      qed
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   273
      finally show ?thesis .
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   274
    qed
58744
c434e37f290e update_cartouches;
wenzelm
parents: 57512
diff changeset
   275
  qed (insert \<open>continuous V f norm\<close>, simp_all add: continuous_def)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   276
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   277
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   278
end