src/HOL/Real/HahnBanach/HahnBanach.thy
author wenzelm
Thu, 20 Jul 2000 12:04:47 +0200
changeset 9394 1ff8a6234c6a
parent 9379 21cfeae6659d
child 9475 b24516d96847
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
     1
(*  Title:      HOL/Real/HahnBanach/HahnBanach.thy
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
     2
    ID:         $Id$
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
     3
    Author:     Gertrud Bauer, TU Munich
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
     4
*)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
     5
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
     6
header {* The Hahn-Banach Theorem *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
     7
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
     8
theory HahnBanach = HahnBanachLemmas: 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
     9
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    10
text {*
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    11
  We present the proof of two different versions of the Hahn-Banach 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    12
  Theorem, closely following \cite[\S36]{Heuser:1986}.
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    13
*}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    14
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    15
subsection {* The Hahn-Banach Theorem for vector spaces *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    16
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    17
text {*
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    18
{\bf Hahn-Banach Theorem.}\quad
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    19
  Let $F$ be a subspace of a real vector space $E$, let $p$ be a semi-norm on
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    20
  $E$, and $f$ be a linear form defined on $F$ such that $f$ is bounded by
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    21
  $p$, i.e. $\All {x\in F} f\ap x \leq p\ap x$.  Then $f$ can be extended to
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    22
  a linear form $h$ on $E$ such that $h$ is norm-preserving, i.e. $h$ is also
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    23
  bounded by $p$.
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    24
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    25
\bigskip
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    26
{\bf Proof Sketch.}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    27
\begin{enumerate}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    28
\item Define $M$ as the set of norm-preserving extensions of $f$ to subspaces
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    29
  of $E$. The linear forms in $M$ are ordered by domain extension.
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    30
\item We show that every non-empty chain in $M$ has an upper bound in $M$.
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    31
\item With Zorn's Lemma we conclude that there is a maximal function $g$ in
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    32
  $M$.
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    33
\item The domain $H$ of $g$ is the whole space $E$, as shown by classical
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    34
  contradiction:
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    35
\begin{itemize}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    36
\item Assuming $g$ is not defined on whole $E$, it can still be extended in a
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    37
  norm-preserving way to a super-space $H'$ of $H$.
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    38
\item Thus $g$ can not be maximal. Contradiction!
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    39
\end{itemize}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    40
\end{enumerate}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    41
\bigskip
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    42
*}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    43
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    44
(*
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    45
text {* {\bf Theorem.} Let $f$ be a linear form defined on a subspace 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    46
 $F$ of a real vector space $E$, such that $f$ is bounded by a seminorm 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    47
 $p$. 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    48
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    49
 Then $f$ can be extended  to a linear form $h$  on $E$ that is again
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    50
 bounded by $p$.
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    51
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    52
 \bigskip{\bf Proof Outline.}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    53
 First we define the set $M$ of all norm-preserving extensions of $f$.
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    54
 We show that every chain in $M$ has an upper bound in $M$.
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    55
 With Zorn's lemma we can conclude that $M$ has a maximal element $g$.
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    56
 We further show by contradiction that the domain $H$ of $g$ is the whole
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    57
 vector space $E$. 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    58
 If $H \neq E$, then $g$ can be extended in 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    59
 a norm-preserving way to a greater vector space $H_0$. 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    60
 So $g$ cannot be maximal in $M$.
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    61
 \bigskip
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    62
*}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    63
*)
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    64
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    65
theorem HahnBanach:
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    66
  "[| is_vectorspace E; is_subspace F E; is_seminorm E p;
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
    67
  is_linearform F f; \\<forall>x \\<in> F. f x <= p x |] 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
    68
  ==> \\<exists>h. is_linearform E h \\<and> (\\<forall>x \\<in> F. h x = f x)
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
    69
        \\<and> (\\<forall>x \\<in> E. h x <= p x)"   
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    70
    -- {* Let $E$ be a vector space, $F$ a subspace of $E$, $p$ a seminorm on $E$, *}
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    71
    -- {* and $f$ a linear form on $F$ such that $f$ is bounded by $p$, *}
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    72
    -- {* then $f$ can be extended to a linear form $h$ on $E$ in a norm-preserving way. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    73
proof -
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    74
  assume "is_vectorspace E" "is_subspace F E" "is_seminorm E p" 
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
    75
   and "is_linearform F f" "\\<forall>x \\<in> F. f x <= p x"
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    76
  -- {* Assume the context of the theorem. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    77
  def M == "norm_pres_extensions E p F f"
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    78
  -- {* Define $M$ as the set of all norm-preserving extensions of $F$. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    79
  {
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
    80
    fix c assume "c \\<in> chain M" "\\<exists>x. x \\<in> c"
9394
wenzelm
parents: 9379
diff changeset
    81
    have "\\<Union>c \\<in> M"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    82
    -- {* Show that every non-empty chain $c$ of $M$ has an upper bound in $M$: *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
    83
    -- {* $\Union c$ is greater than any element of the chain $c$, so it suffices to show $\Union c \in M$. *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    84
    proof (unfold M_def, rule norm_pres_extensionI)
9394
wenzelm
parents: 9379
diff changeset
    85
      show "\\<exists>H h. graph H h = \\<Union>c
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
    86
              \\<and> is_linearform H h 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
    87
              \\<and> is_subspace H E 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
    88
              \\<and> is_subspace F H 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
    89
              \\<and> graph F f \\<subseteq> graph H h
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
    90
              \\<and> (\\<forall>x \\<in> H. h x <= p x)"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    91
      proof (intro exI conjI)
9394
wenzelm
parents: 9379
diff changeset
    92
        let ?H = "domain (\\<Union>c)"
wenzelm
parents: 9379
diff changeset
    93
        let ?h = "funct (\\<Union>c)"
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    94
9394
wenzelm
parents: 9379
diff changeset
    95
        show a: "graph ?H ?h = \\<Union>c" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    96
        proof (rule graph_domain_funct)
9394
wenzelm
parents: 9379
diff changeset
    97
          fix x y z assume "(x, y) \\<in> \\<Union>c" "(x, z) \\<in> \\<Union>c"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    98
          show "z = y" by (rule sup_definite)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    99
        qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   100
        show "is_linearform ?H ?h" 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   101
          by (simp! add: sup_lf a)
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   102
        show "is_subspace ?H E" 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   103
          by (rule sup_subE, rule a) (simp!)+
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   104
        show "is_subspace F ?H" 
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   105
          by (rule sup_supF, rule a) (simp!)+
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   106
        show "graph F f \\<subseteq> graph ?H ?h" 
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   107
          by (rule sup_ext, rule a) (simp!)+
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   108
        show "\\<forall>x \\<in> ?H. ?h x <= p x" 
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   109
          by (rule sup_norm_pres, rule a) (simp!)+
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   110
      qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   111
    qed
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   112
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   113
  }
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   114
  hence "\\<exists>g \\<in> M. \\<forall>x \\<in> M. g \\<subseteq> x --> g = x" 
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   115
  -- {* With Zorn's Lemma we can conclude that there is a maximal element in $M$.\skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   116
  proof (rule Zorn's_Lemma)
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   117
    -- {* We show that $M$ is non-empty: *}
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   118
    have "graph F f \\<in> norm_pres_extensions E p F f"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   119
    proof (rule norm_pres_extensionI2)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   120
      have "is_vectorspace F" ..
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   121
      thus "is_subspace F F" ..
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   122
    qed (blast!)+ 
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   123
    thus "graph F f \\<in> M" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   124
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   125
  thus ?thesis
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   126
  proof
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   127
    fix g assume "g \\<in> M" "\\<forall>x \\<in> M. g \\<subseteq> x --> g = x"
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   128
    -- {* We consider such a maximal element $g \in M$. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   129
    show ?thesis
8109
aca11f954993 obtain: renamed 'in' to 'where';
wenzelm
parents: 8108
diff changeset
   130
      obtain H h where "graph H h = g" "is_linearform H h" 
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   131
        "is_subspace H E" "is_subspace F H" "graph F f \\<subseteq> graph H h" 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   132
        "\\<forall>x \\<in> H. h x <= p x" 
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   133
        -- {* $g$ is a norm-preserving extension of $f$, in other words: *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   134
        -- {* $g$ is the graph of some linear form $h$ defined on a subspace $H$ of $E$, *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   135
        -- {* and $h$ is an extension of $f$ that is again bounded by $p$. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   136
      proof -
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   137
        have "\\<exists>H h. graph H h = g \\<and> is_linearform H h 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   138
          \\<and> is_subspace H E \\<and> is_subspace F H
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   139
          \\<and> graph F f \\<subseteq> graph H h
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   140
          \\<and> (\\<forall>x \\<in> H. h x <= p x)" 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   141
          by (simp! add: norm_pres_extension_D)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   142
        thus ?thesis by (elim exE conjE) rule
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   143
      qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   144
      have h: "is_vectorspace H" ..
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   145
      have "H = E"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   146
      -- {* We show that $h$ is defined on whole $E$ by classical contradiction. \skp *} 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   147
      proof (rule classical)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   148
        assume "H \\<noteq> E"
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   149
        -- {* Assume $h$ is not defined on whole $E$. Then show that $h$ can be extended *}
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   150
        -- {* in a norm-preserving way to a function $h'$ with the graph $g'$. \skp *}
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   151
        have "\\<exists>g' \\<in> M. g \\<subseteq> g' \\<and> g \\<noteq> g'" 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   152
          obtain x' where "x' \\<in> E" "x' \\<notin> H" 
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   153
          -- {* Pick $x' \in E \setminus H$. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   154
          proof -
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   155
            have "\\<exists>x' \\<in> E. x' \\<notin> H"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   156
            proof (rule set_less_imp_diff_not_empty)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   157
              have "H \\<subseteq> E" ..
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   158
              thus "H \\<subset> E" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   159
            qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   160
            thus ?thesis by blast
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   161
          qed
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   162
          have x': "x' \\<noteq> 0"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   163
          proof (rule classical)
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   164
            presume "x' = 0"
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   165
            with h have "x' \\<in> H" by simp
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   166
            thus ?thesis by contradiction
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   167
          qed blast
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   168
          def H' == "H + lin x'"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   169
          -- {* Define $H'$ as the direct sum of $H$ and the linear closure of $x'$. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   170
          show ?thesis
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   171
            obtain xi where "\\<forall>y \\<in> H. - p (y + x') - h y <= xi 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   172
                              \\<and> xi <= p (y + x') - h y" 
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   173
            -- {* Pick a real number $\xi$ that fulfills certain inequations; this will *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   174
            -- {* be used to establish that $h'$ is a norm-preserving extension of $h$. 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   175
               \label{ex-xi-use}\skp *}
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   176
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   177
            proof -
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   178
	      from h have "\\<exists>xi. \\<forall>y \\<in> H. - p (y + x') - h y <= xi 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   179
                              \\<and> xi <= p (y + x') - h y" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   180
              proof (rule ex_xi)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   181
                fix u v assume "u \\<in> H" "v \\<in> H"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   182
                from h have "h v - h u = h (v - u)"
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   183
                  by (simp! add: linearform_diff)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   184
                also have "... <= p (v - u)"
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   185
                  by (simp!)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   186
                also have "v - u = x' + - x' + v + - u"
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   187
                  by (simp! add: diff_eq1)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   188
                also have "... = v + x' + - (u + x')"
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   189
                  by (simp!)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   190
                also have "... = (v + x') - (u + x')"
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   191
                  by (simp! add: diff_eq1)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   192
                also have "p ... <= p (v + x') + p (u + x')"
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   193
                  by (rule seminorm_diff_subadditive) (simp!)+
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   194
                finally have "h v - h u <= p (v + x') + p (u + x')" .
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   195
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   196
                thus "- p (u + x') - h u <= p (v + x') - h v" 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   197
                  by (rule real_diff_ineq_swap)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   198
              qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   199
              thus ?thesis by rule rule
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   200
            qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   201
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   202
            def h' == "\\<lambda>x. let (y,a) = SOME (y,a). x = y + a \\<cdot> x' \\<and> y \\<in> H
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   203
                           in (h y) + a * xi"
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   204
            -- {* Define the extension $h'$ of $h$ to $H'$ using $\xi$. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   205
            show ?thesis
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   206
            proof
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   207
              show "g \\<subseteq> graph H' h' \\<and> g \\<noteq> graph H' h'" 
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   208
              -- {* Show that $h'$ is an extension of $h$ \dots \skp *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   209
              proof
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   210
		show "g \\<subseteq> graph H' h'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   211
		proof -
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   212
		  have  "graph H h \\<subseteq> graph H' h'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   213
                  proof (rule graph_extI)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   214
		    fix t assume "t \\<in> H" 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   215
		    have "(SOME (y, a). t = y + a \\<cdot> x' \\<and> y \\<in> H)
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   216
                         = (t, #0)"
9394
wenzelm
parents: 9379
diff changeset
   217
		      by (rule decomp_H'_H) (assumption+, rule x')
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   218
		    thus "h t = h' t" by (simp! add: Let_def)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   219
		  next
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   220
		    show "H \\<subseteq> H'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   221
		    proof (rule subspace_subset)
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   222
		      show "is_subspace H H'"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   223
		      proof (unfold H'_def, rule subspace_vs_sum1)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   224
			show "is_vectorspace H" ..
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   225
			show "is_vectorspace (lin x')" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   226
		      qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   227
		    qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   228
		  qed 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   229
		  thus ?thesis by (simp!)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   230
		qed
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   231
                show "g \\<noteq> graph H' h'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   232
		proof -
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   233
		  have "graph H h \\<noteq> graph H' h'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   234
		  proof
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   235
		    assume e: "graph H h = graph H' h'"
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   236
		    have "x' \\<in> H'" 
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   237
		    proof (unfold H'_def, rule vs_sumI)
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   238
		      show "x' = 0 + x'" by (simp!)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   239
		      from h show "0 \\<in> H" ..
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   240
		      show "x' \\<in> lin x'" by (rule x_lin_x)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   241
		    qed
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   242
		    hence "(x', h' x') \\<in> graph H' h'" ..
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   243
		    with e have "(x', h' x') \\<in> graph H h" by simp
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   244
		    hence "x' \\<in> H" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   245
		    thus False by contradiction
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   246
		  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   247
		  thus ?thesis by (simp!)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   248
		qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   249
              qed
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   250
	      show "graph H' h' \\<in> M" 
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   251
              -- {* and $h'$ is norm-preserving. \skp *} 
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   252
              proof -
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   253
		have "graph H' h' \\<in> norm_pres_extensions E p F f"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   254
		proof (rule norm_pres_extensionI2)
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   255
		  show "is_linearform H' h'"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   256
		    by (rule h'_lf) (simp! add: x')+
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   257
                  show "is_subspace H' E" 
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   258
		    by (unfold H'_def) 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   259
                      (rule vs_sum_subspace [OF _ lin_subspace])
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   260
		  have "is_subspace F H" .
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   261
		  also from h lin_vs 
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   262
		  have [fold H'_def]: "is_subspace H (H + lin x')" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   263
		  finally (subspace_trans [OF _ h]) 
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   264
		  show f_h': "is_subspace F H'" .
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
   265
		
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   266
		  show "graph F f \\<subseteq> graph H' h'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   267
		  proof (rule graph_extI)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   268
		    fix x assume "x \\<in> F"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   269
		    have "f x = h x" ..
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   270
		    also have " ... = h x + #0 * xi" by simp
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   271
		    also 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   272
                    have "... = (let (y,a) = (x, #0) in h y + a * xi)"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   273
		      by (simp add: Let_def)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   274
		    also have 
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   275
		      "(x, #0) = (SOME (y, a). x = y + a \\<cdot> x' \\<and> y \\<in> H)"
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   276
		      by (rule decomp_H'_H [RS sym]) (simp! add: x')+
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   277
		    also have 
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   278
		      "(let (y,a) = (SOME (y,a). x = y + a \\<cdot> x' \\<and> y \\<in> H)
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   279
                        in h y + a * xi) = h' x" by (simp!)
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   280
		    finally show "f x = h' x" .
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   281
		  next
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   282
		    from f_h' show "F \\<subseteq> H'" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   283
		  qed
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
   284
		
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   285
		  show "\\<forall>x \\<in> H'. h' x <= p x"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   286
		    by (rule h'_norm_pres) (assumption+, rule x')
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   287
		qed
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   288
		thus "graph H' h' \\<in> M" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   289
	      qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   290
            qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   291
          qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   292
        qed
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   293
        hence "\\<not> (\\<forall>x \\<in> M. g \\<subseteq> x --> g = x)" by simp
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   294
        -- {* So the graph $g$ of $h$ cannot be maximal. Contradiction! \skp *}
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   295
        thus "H = E" by contradiction
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   296
      qed
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   297
      thus "\\<exists>h. is_linearform E h \\<and> (\\<forall>x \\<in> F. h x = f x) 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   298
          \\<and> (\\<forall>x \\<in> E. h x <= p x)" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   299
      proof (intro exI conjI)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   300
        assume eq: "H = E"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   301
	from eq show "is_linearform E h" by (simp!)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   302
	show "\\<forall>x \\<in> F. h x = f x" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   303
	proof (intro ballI, rule sym)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   304
	  fix x assume "x \\<in> F" show "f x = h x " ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   305
	qed
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   306
	from eq show "\\<forall>x \\<in> E. h x <= p x" by (force!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   307
      qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   308
    qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   309
  qed
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   310
qed 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   311
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   312
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   313
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   314
subsection  {* Alternative formulation *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   315
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   316
text {* The following alternative formulation of the Hahn-Banach
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   317
Theorem\label{abs-HahnBanach} uses the fact that for a real linear form
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   318
$f$ and a seminorm $p$ the
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   319
following inequations are equivalent:\footnote{This was shown in lemma
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   320
$\idt{abs{\dsh}ineq{\dsh}iff}$ (see page \pageref{abs-ineq-iff}).}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   321
\begin{matharray}{ll}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   322
\forall x\in H.\ap |h\ap x|\leq p\ap x& {\rm and}\\
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   323
\forall x\in H.\ap h\ap x\leq p\ap x\\
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   324
\end{matharray}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   325
*}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   326
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   327
theorem abs_HahnBanach:
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   328
  "[| is_vectorspace E; is_subspace F E; is_linearform F f; 
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   329
  is_seminorm E p; \\<forall>x \\<in> F. |f x| <= p x |]
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   330
  ==> \\<exists>g. is_linearform E g \\<and> (\\<forall>x \\<in> F. g x = f x)
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   331
   \\<and> (\\<forall>x \\<in> E. |g x| <= p x)"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   332
proof -
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   333
  assume "is_vectorspace E" "is_subspace F E" "is_seminorm E p" 
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   334
    "is_linearform F f"  "\\<forall>x \\<in> F. |f x| <= p x"
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   335
  have "\\<forall>x \\<in> F. f x <= p x"  by (rule abs_ineq_iff [RS iffD1])
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   336
  hence "\\<exists>g. is_linearform E g \\<and> (\\<forall>x \\<in> F. g x = f x) 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   337
              \\<and> (\\<forall>x \\<in> E. g x <= p x)"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   338
    by (simp! only: HahnBanach)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   339
  thus ?thesis 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   340
  proof (elim exE conjE)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   341
    fix g assume "is_linearform E g" "\\<forall>x \\<in> F. g x = f x" 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   342
                  "\\<forall>x \\<in> E. g x <= p x"
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   343
    hence "\\<forall>x \\<in> E. |g x| <= p x"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   344
      by (simp! add: abs_ineq_iff [OF subspace_refl])
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   345
    thus ?thesis by (intro exI conjI)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   346
  qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   347
qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   348
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   349
subsection {* The Hahn-Banach Theorem for normed spaces *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   350
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   351
text {* Every continuous linear form $f$ on a subspace $F$ of a
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   352
norm space $E$, can be extended to a continuous linear form $g$ on
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   353
$E$ such that $\fnorm{f} = \fnorm {g}$. *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   354
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   355
theorem norm_HahnBanach:
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   356
  "[| is_normed_vectorspace E norm; is_subspace F E; 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   357
  is_linearform F f; is_continuous F norm f |] 
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   358
  ==> \\<exists>g. is_linearform E g
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   359
         \\<and> is_continuous E norm g 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   360
         \\<and> (\\<forall>x \\<in> F. g x = f x) 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   361
         \\<and> \\<parallel>g\\<parallel>E,norm = \\<parallel>f\\<parallel>F,norm"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   362
proof -
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   363
  assume e_norm: "is_normed_vectorspace E norm"
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   364
  assume f: "is_subspace F E" "is_linearform F f"
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   365
  assume f_cont: "is_continuous F norm f"
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   366
  have e: "is_vectorspace E" ..
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   367
  hence f_norm: "is_normed_vectorspace F norm" ..
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   368
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   369
  txt{* We define a function $p$ on $E$ as follows:
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   370
  \begin{matharray}{l}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   371
  p \: x = \fnorm f \cdot \norm x\\
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   372
  \end{matharray}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   373
  *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   374
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   375
  def p == "\\<lambda>x. \\<parallel>f\\<parallel>F,norm * norm x"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   376
  
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   377
  txt{* $p$ is a seminorm on $E$: *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   378
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   379
  have q: "is_seminorm E p"
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   380
  proof
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   381
    fix x y a assume "x \\<in> E" "y \\<in> E"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   382
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   383
    txt{* $p$ is positive definite: *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   384
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   385
    show "#0 <= p x"
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   386
    proof (unfold p_def, rule real_le_mult_order1a)
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   387
      from f_cont f_norm show "#0 <= \\<parallel>f\\<parallel>F,norm" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   388
      show "#0 <= norm x" ..
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   389
    qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   390
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   391
    txt{* $p$ is absolutely homogenous: *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   392
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   393
    show "p (a \\<cdot> x) = |a| * p x"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   394
    proof - 
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   395
      have "p (a \\<cdot> x) = \\<parallel>f\\<parallel>F,norm * norm (a \\<cdot> x)"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   396
        by (simp!)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   397
      also have "norm (a \\<cdot> x) = |a| * norm x" 
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   398
        by (rule normed_vs_norm_abs_homogenous)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   399
      also have "\\<parallel>f\\<parallel>F,norm * ( |a| * norm x ) 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   400
        = |a| * (\\<parallel>f\\<parallel>F,norm * norm x)"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   401
        by (simp! only: real_mult_left_commute)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   402
      also have "... = |a| * p x" by (simp!)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   403
      finally show ?thesis .
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   404
    qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   405
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   406
    txt{* Furthermore, $p$ is subadditive: *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   407
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   408
    show "p (x + y) <= p x + p y"
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   409
    proof -
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   410
      have "p (x + y) = \\<parallel>f\\<parallel>F,norm * norm (x + y)"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   411
        by (simp!)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   412
      also 
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   413
      have "... <= \\<parallel>f\\<parallel>F,norm * (norm x + norm y)"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   414
      proof (rule real_mult_le_le_mono1a)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   415
        from f_cont f_norm show "#0 <= \\<parallel>f\\<parallel>F,norm" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   416
        show "norm (x + y) <= norm x + norm y" ..
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   417
      qed
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   418
      also have "... = \\<parallel>f\\<parallel>F,norm * norm x 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   419
                        + \\<parallel>f\\<parallel>F,norm * norm y"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   420
        by (simp! only: real_add_mult_distrib2)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   421
      finally show ?thesis by (simp!)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   422
    qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   423
  qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   424
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   425
  txt{* $f$ is bounded by $p$. *} 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   426
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   427
  have "\\<forall>x \\<in> F. |f x| <= p x"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   428
  proof
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   429
    fix x assume "x \\<in> F"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   430
     from f_norm show "|f x| <= p x" 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   431
       by (simp! add: norm_fx_le_norm_f_norm_x)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   432
  qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   433
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   434
  txt{* Using the fact that $p$ is a seminorm and 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   435
  $f$ is bounded by $p$ we can apply the Hahn-Banach Theorem 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   436
  for real vector spaces. 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   437
  So $f$ can be extended in a norm-preserving way to some function
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   438
  $g$ on the whole vector space $E$. *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   439
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   440
  with e f q 
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   441
  have "\\<exists>g. is_linearform E g \\<and> (\\<forall>x \\<in> F. g x = f x) 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   442
            \\<and> (\\<forall>x \\<in> E. |g x| <= p x)"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   443
    by (simp! add: abs_HahnBanach)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   444
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   445
  thus ?thesis
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   446
  proof (elim exE conjE) 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   447
    fix g
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   448
    assume "is_linearform E g" and a: "\\<forall>x \\<in> F. g x = f x" 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   449
       and b: "\\<forall>x \\<in> E. |g x| <= p x"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   450
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   451
    show "\\<exists>g. is_linearform E g 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   452
            \\<and> is_continuous E norm g 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   453
            \\<and> (\\<forall>x \\<in> F. g x = f x) 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   454
            \\<and> \\<parallel>g\\<parallel>E,norm = \\<parallel>f\\<parallel>F,norm"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   455
    proof (intro exI conjI)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   456
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   457
    txt{* We furthermore have to show that 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   458
    $g$ is also continuous: *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   459
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   460
      show g_cont: "is_continuous E norm g"
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   461
      proof
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   462
        fix x assume "x \\<in> E"
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   463
        with b show "|g x| <= \\<parallel>f\\<parallel>F,norm * norm x"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   464
          by (simp add: p_def) 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   465
      qed 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   466
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   467
      txt {* To complete the proof, we show that 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   468
      $\fnorm g = \fnorm f$. \label{order_antisym} *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   469
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   470
      show "\\<parallel>g\\<parallel>E,norm = \\<parallel>f\\<parallel>F,norm"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   471
        (is "?L = ?R")
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   472
      proof (rule order_antisym)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   473
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   474
        txt{* First we show $\fnorm g \leq \fnorm f$.  The function norm
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   475
        $\fnorm g$ is defined as the smallest $c\in\bbbR$ such that
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   476
        \begin{matharray}{l}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   477
        \All {x\in E} {|g\ap x| \leq c \cdot \norm x}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   478
        \end{matharray}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   479
        Furthermore holds
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   480
        \begin{matharray}{l}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   481
        \All {x\in E} {|g\ap x| \leq \fnorm f \cdot \norm x}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   482
        \end{matharray}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   483
        *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   484
 
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   485
        have "\\<forall>x \\<in> E. |g x| <= \\<parallel>f\\<parallel>F,norm * norm x"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   486
        proof
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   487
          fix x assume "x \\<in> E" 
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   488
          show "|g x| <= \\<parallel>f\\<parallel>F,norm * norm x"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   489
            by (simp!)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   490
        qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   491
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   492
        with g_cont e_norm show "?L <= ?R"
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   493
        proof (rule fnorm_le_ub)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   494
          from f_cont f_norm show "#0 <= \\<parallel>f\\<parallel>F,norm" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   495
        qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   496
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   497
        txt{* The other direction is achieved by a similar 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   498
        argument. *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   499
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   500
        have "\\<forall>x \\<in> F. |f x| <= \\<parallel>g\\<parallel>E,norm * norm x"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   501
        proof
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   502
          fix x assume "x \\<in> F" 
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   503
          from a have "g x = f x" ..
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   504
          hence "|f x| = |g x|" by simp
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   505
          also from g_cont
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   506
          have "... <= \\<parallel>g\\<parallel>E,norm * norm x"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   507
          proof (rule norm_fx_le_norm_f_norm_x)
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   508
            show "x \\<in> E" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   509
          qed
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   510
          finally show "|f x| <= \\<parallel>g\\<parallel>E,norm * norm x" .
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   511
        qed 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   512
        thus "?R <= ?L" 
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   513
        proof (rule fnorm_le_ub [OF f_cont f_norm])
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   514
          from g_cont show "#0 <= \\<parallel>g\\<parallel>E,norm" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   515
        qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   516
      qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   517
    qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   518
  qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   519
qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   520
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   521
end