src/HOL/Real/HahnBanach/HahnBanach.thy
author wenzelm
Thu, 03 Aug 2000 00:34:22 +0200
changeset 9503 3324cbbecef8
parent 9475 b24516d96847
child 9560 b87a6afe5881
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;
9503
wenzelm
parents: 9475
diff changeset
    67
  is_linearform F f; \<forall>x \<in> F. f x <= p x |] 
wenzelm
parents: 9475
diff changeset
    68
  ==> \<exists>h. is_linearform E h \<and> (\<forall>x \<in> F. h x = f x)
wenzelm
parents: 9475
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" 
9503
wenzelm
parents: 9475
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
  {
9503
wenzelm
parents: 9475
diff changeset
    80
    fix c assume "c \<in> chain M" "\<exists>x. x \<in> c"
wenzelm
parents: 9475
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)
9503
wenzelm
parents: 9475
diff changeset
    85
      show "\<exists>H h. graph H h = \<Union>c
wenzelm
parents: 9475
diff changeset
    86
              \<and> is_linearform H h 
wenzelm
parents: 9475
diff changeset
    87
              \<and> is_subspace H E 
wenzelm
parents: 9475
diff changeset
    88
              \<and> is_subspace F H 
wenzelm
parents: 9475
diff changeset
    89
              \<and> graph F f \<subseteq> graph H h
wenzelm
parents: 9475
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)
9503
wenzelm
parents: 9475
diff changeset
    92
        let ?H = "domain (\<Union>c)"
wenzelm
parents: 9475
diff changeset
    93
        let ?h = "funct (\<Union>c)"
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    94
9503
wenzelm
parents: 9475
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)
9503
wenzelm
parents: 9475
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!)+
9503
wenzelm
parents: 9475
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!)+
9503
wenzelm
parents: 9475
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
  }
9503
wenzelm
parents: 9475
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: *}
9503
wenzelm
parents: 9475
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!)+ 
9503
wenzelm
parents: 9475
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
9503
wenzelm
parents: 9475
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 *}
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   129
    obtain H h where "graph H h = g" "is_linearform H h" 
9503
wenzelm
parents: 9475
diff changeset
   130
      "is_subspace H E" "is_subspace F H" "graph F f \<subseteq> graph H h" 
wenzelm
parents: 9475
diff changeset
   131
      "\<forall>x \<in> H. h x <= p x" 
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   132
      -- {* $g$ is a norm-preserving extension of $f$, in other words: *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   133
      -- {* $g$ is the graph of some linear form $h$ defined on a subspace $H$ of $E$, *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   134
      -- {* and $h$ is an extension of $f$ that is again bounded by $p$. \skp *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   135
    proof -
9503
wenzelm
parents: 9475
diff changeset
   136
      have "\<exists>H h. graph H h = g \<and> is_linearform H h 
wenzelm
parents: 9475
diff changeset
   137
        \<and> is_subspace H E \<and> is_subspace F H
wenzelm
parents: 9475
diff changeset
   138
        \<and> graph F f \<subseteq> graph H h
wenzelm
parents: 9475
diff changeset
   139
        \<and> (\<forall>x \<in> H. h x <= p x)" 
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   140
        by (simp! add: norm_pres_extension_D)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   141
      thus ?thesis by (elim exE conjE) rule
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   142
    qed
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   143
    have h: "is_vectorspace H" ..
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   144
    have "H = E"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   145
    -- {* We show that $h$ is defined on whole $E$ by classical contradiction. \skp *} 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   146
    proof (rule classical)
9503
wenzelm
parents: 9475
diff changeset
   147
      assume "H \<noteq> E"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   148
      -- {* Assume $h$ is not defined on whole $E$. Then show that $h$ can be extended *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   149
      -- {* in a norm-preserving way to a function $h'$ with the graph $g'$. \skp *}
9503
wenzelm
parents: 9475
diff changeset
   150
      have "\<exists>g' \<in> M. g \<subseteq> g' \<and> g \<noteq> g'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   151
      proof -
9503
wenzelm
parents: 9475
diff changeset
   152
        obtain x' where "x' \<in> E" "x' \<notin> H" 
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   153
        -- {* Pick $x' \in E \setminus H$. \skp *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   154
        proof -
9503
wenzelm
parents: 9475
diff changeset
   155
          have "\<exists>x' \<in> E. x' \<notin> H"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   156
          proof (rule set_less_imp_diff_not_empty)
9503
wenzelm
parents: 9475
diff changeset
   157
            have "H \<subseteq> E" ..
wenzelm
parents: 9475
diff changeset
   158
            thus "H \<subset> E" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   159
          qed
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   160
          thus ?thesis by blast
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   161
        qed
9503
wenzelm
parents: 9475
diff changeset
   162
        have x': "x' \<noteq> 0"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   163
        proof (rule classical)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   164
          presume "x' = 0"
9503
wenzelm
parents: 9475
diff changeset
   165
          with h have "x' \<in> H" by simp
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   166
          thus ?thesis by contradiction
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   167
        qed blast
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   168
        def H' == "H + lin x'"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   169
        -- {* Define $H'$ as the direct sum of $H$ and the linear closure of $x'$. \skp *}
9503
wenzelm
parents: 9475
diff changeset
   170
        obtain xi where "\<forall>y \<in> H. - p (y + x') - h y <= xi 
wenzelm
parents: 9475
diff changeset
   171
                          \<and> xi <= p (y + x') - h y" 
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   172
        -- {* Pick a real number $\xi$ that fulfills certain inequations; this will *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   173
        -- {* be used to establish that $h'$ is a norm-preserving extension of $h$. 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   174
           \label{ex-xi-use}\skp *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   175
        proof -
9503
wenzelm
parents: 9475
diff changeset
   176
          from h have "\<exists>xi. \<forall>y \<in> H. - p (y + x') - h y <= xi 
wenzelm
parents: 9475
diff changeset
   177
                          \<and> xi <= p (y + x') - h y" 
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   178
          proof (rule ex_xi)
9503
wenzelm
parents: 9475
diff changeset
   179
            fix u v assume "u \<in> H" "v \<in> H"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   180
            from h have "h v - h u = h (v - u)"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   181
              by (simp! add: linearform_diff)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   182
            also have "... <= p (v - u)"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   183
              by (simp!)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   184
            also have "v - u = x' + - x' + v + - u"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   185
              by (simp! add: diff_eq1)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   186
            also have "... = v + x' + - (u + x')"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   187
              by (simp!)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   188
            also have "... = (v + x') - (u + x')"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   189
              by (simp! add: diff_eq1)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   190
            also have "p ... <= p (v + x') + p (u + x')"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   191
              by (rule seminorm_diff_subadditive) (simp_all!)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   192
            finally have "h v - h u <= p (v + x') + p (u + x')" .
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   193
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   194
            thus "- p (u + x') - h u <= p (v + x') - h v" 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   195
              by (rule real_diff_ineq_swap)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   196
          qed
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   197
          thus ?thesis by rule rule
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   198
        qed
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   199
9503
wenzelm
parents: 9475
diff changeset
   200
        def h' == "\<lambda>x. let (y,a) = SOME (y,a). x = y + a \<cdot> x' \<and> y \<in> H
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   201
                       in (h y) + a * xi"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   202
        -- {* Define the extension $h'$ of $h$ to $H'$ using $\xi$. \skp *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   203
        show ?thesis
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   204
        proof
9503
wenzelm
parents: 9475
diff changeset
   205
          show "g \<subseteq> graph H' h' \<and> g \<noteq> graph H' h'" 
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   206
          -- {* Show that $h'$ is an extension of $h$ \dots \skp *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   207
          proof
9503
wenzelm
parents: 9475
diff changeset
   208
            show "g \<subseteq> graph H' h'"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   209
            proof -
9503
wenzelm
parents: 9475
diff changeset
   210
              have  "graph H h \<subseteq> graph H' h'"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   211
              proof (rule graph_extI)
9503
wenzelm
parents: 9475
diff changeset
   212
                fix t assume "t \<in> H" 
wenzelm
parents: 9475
diff changeset
   213
                have "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H)
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   214
                     = (t, #0)"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   215
                  by (rule decomp_H'_H) (assumption+, rule x')
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   216
                thus "h t = h' t" by (simp! add: Let_def)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   217
              next
9503
wenzelm
parents: 9475
diff changeset
   218
                show "H \<subseteq> H'"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   219
                proof (rule subspace_subset)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   220
                  show "is_subspace H H'"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   221
                  proof (unfold H'_def, rule subspace_vs_sum1)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   222
                    show "is_vectorspace H" ..
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   223
                    show "is_vectorspace (lin x')" ..
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   224
                  qed
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   225
                qed
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   226
              qed 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   227
              thus ?thesis by (simp!)
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   228
            qed
9503
wenzelm
parents: 9475
diff changeset
   229
            show "g \<noteq> graph H' h'"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   230
            proof -
9503
wenzelm
parents: 9475
diff changeset
   231
              have "graph H h \<noteq> graph H' h'"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   232
              proof
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   233
                assume e: "graph H h = graph H' h'"
9503
wenzelm
parents: 9475
diff changeset
   234
                have "x' \<in> H'" 
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   235
                proof (unfold H'_def, rule vs_sumI)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   236
                  show "x' = 0 + x'" by (simp!)
9503
wenzelm
parents: 9475
diff changeset
   237
                  from h show "0 \<in> H" ..
wenzelm
parents: 9475
diff changeset
   238
                  show "x' \<in> lin x'" by (rule x_lin_x)
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   239
                qed
9503
wenzelm
parents: 9475
diff changeset
   240
                hence "(x', h' x') \<in> graph H' h'" ..
wenzelm
parents: 9475
diff changeset
   241
                with e have "(x', h' x') \<in> graph H h" by simp
wenzelm
parents: 9475
diff changeset
   242
                hence "x' \<in> H" ..
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   243
                thus False by contradiction
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   244
              qed
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   245
              thus ?thesis by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   246
            qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   247
          qed
9503
wenzelm
parents: 9475
diff changeset
   248
          show "graph H' h' \<in> M" 
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   249
          -- {* and $h'$ is norm-preserving. \skp *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   250
          proof -
9503
wenzelm
parents: 9475
diff changeset
   251
            have "graph H' h' \<in> norm_pres_extensions E p F f"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   252
            proof (rule norm_pres_extensionI2)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   253
              show "is_linearform H' h'"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   254
                by (rule h'_lf) (simp! add: x')+
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   255
              show "is_subspace H' E" 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   256
                by (unfold H'_def) 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   257
                  (rule vs_sum_subspace [OF _ lin_subspace])
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   258
              have "is_subspace F H" .
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   259
              also from h lin_vs 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   260
              have [fold H'_def]: "is_subspace H (H + lin x')" ..
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   261
              finally (subspace_trans [OF _ h]) 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   262
              show f_h': "is_subspace F H'" .
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   263
            
9503
wenzelm
parents: 9475
diff changeset
   264
              show "graph F f \<subseteq> graph H' h'"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   265
              proof (rule graph_extI)
9503
wenzelm
parents: 9475
diff changeset
   266
                fix x assume "x \<in> F"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   267
                have "f x = h x" ..
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   268
                also have " ... = h x + #0 * xi" by simp
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   269
                also 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   270
                have "... = (let (y,a) = (x, #0) in h y + a * xi)"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   271
                  by (simp add: Let_def)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   272
                also have 
9503
wenzelm
parents: 9475
diff changeset
   273
                  "(x, #0) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   274
                  by (rule decomp_H'_H [RS sym]) (simp! add: x')+
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   275
                also have 
9503
wenzelm
parents: 9475
diff changeset
   276
                  "(let (y,a) = (SOME (y,a). x = y + a \<cdot> x' \<and> y \<in> H)
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   277
                    in h y + a * xi) = h' x" by (simp!)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   278
                finally show "f x = h' x" .
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   279
              next
9503
wenzelm
parents: 9475
diff changeset
   280
                from f_h' show "F \<subseteq> H'" ..
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   281
              qed
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   282
            
9503
wenzelm
parents: 9475
diff changeset
   283
              show "\<forall>x \<in> H'. h' x <= p x"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   284
                by (rule h'_norm_pres) (assumption+, rule x')
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   285
            qed
9503
wenzelm
parents: 9475
diff changeset
   286
            thus "graph H' h' \<in> M" by (simp!)
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   287
          qed
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   288
        qed
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   289
      qed
9503
wenzelm
parents: 9475
diff changeset
   290
      hence "\<not> (\<forall>x \<in> M. g \<subseteq> x --> g = x)" by simp
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   291
	-- {* So the graph $g$ of $h$ cannot be maximal. Contradiction! \skp *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   292
      thus "H = E" by contradiction
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   293
    qed
9503
wenzelm
parents: 9475
diff changeset
   294
    thus "\<exists>h. is_linearform E h \<and> (\<forall>x \<in> F. h x = f x) 
wenzelm
parents: 9475
diff changeset
   295
      \<and> (\<forall>x \<in> E. h x <= p x)" 
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   296
    proof (intro exI conjI)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   297
      assume eq: "H = E"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   298
      from eq show "is_linearform E h" by (simp!)
9503
wenzelm
parents: 9475
diff changeset
   299
      show "\<forall>x \<in> F. h x = f x"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   300
      proof
9503
wenzelm
parents: 9475
diff changeset
   301
	fix x assume "x \<in> F" have "f x = h x " ..
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   302
	thus "h x = f x" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   303
      qed
9503
wenzelm
parents: 9475
diff changeset
   304
      from eq show "\<forall>x \<in> E. h x <= p x" by (force!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   305
    qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   306
  qed
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   307
qed
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   308
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   309
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   310
subsection  {* Alternative formulation *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   311
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   312
text {* The following alternative formulation of the Hahn-Banach
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   313
Theorem\label{abs-HahnBanach} uses the fact that for a real linear form
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   314
$f$ and a seminorm $p$ the
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   315
following inequations are equivalent:\footnote{This was shown in lemma
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   316
$\idt{abs{\dsh}ineq{\dsh}iff}$ (see page \pageref{abs-ineq-iff}).}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   317
\begin{matharray}{ll}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   318
\forall x\in H.\ap |h\ap x|\leq p\ap x& {\rm and}\\
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   319
\forall x\in H.\ap h\ap x\leq p\ap x\\
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   320
\end{matharray}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   321
*}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   322
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   323
theorem abs_HahnBanach:
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   324
"[| is_vectorspace E; is_subspace F E; is_linearform F f; 
9503
wenzelm
parents: 9475
diff changeset
   325
is_seminorm E p; \<forall>x \<in> F. |f x| <= p x |]
wenzelm
parents: 9475
diff changeset
   326
==> \<exists>g. is_linearform E g \<and> (\<forall>x \<in> F. g x = f x)
wenzelm
parents: 9475
diff changeset
   327
 \<and> (\<forall>x \<in> E. |g x| <= p x)"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   328
proof -
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   329
assume "is_vectorspace E" "is_subspace F E" "is_seminorm E p" 
9503
wenzelm
parents: 9475
diff changeset
   330
"is_linearform F f"  "\<forall>x \<in> F. |f x| <= p x"
wenzelm
parents: 9475
diff changeset
   331
have "\<forall>x \<in> F. f x <= p x"  by (rule abs_ineq_iff [RS iffD1])
wenzelm
parents: 9475
diff changeset
   332
hence "\<exists>g. is_linearform E g \<and> (\<forall>x \<in> F. g x = f x) 
wenzelm
parents: 9475
diff changeset
   333
          \<and> (\<forall>x \<in> E. g x <= p x)"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   334
by (simp! only: HahnBanach)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   335
thus ?thesis 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   336
proof (elim exE conjE)
9503
wenzelm
parents: 9475
diff changeset
   337
fix g assume "is_linearform E g" "\<forall>x \<in> F. g x = f x" 
wenzelm
parents: 9475
diff changeset
   338
              "\<forall>x \<in> E. g x <= p x"
wenzelm
parents: 9475
diff changeset
   339
hence "\<forall>x \<in> E. |g x| <= p x"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   340
  by (simp! add: abs_ineq_iff [OF subspace_refl])
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   341
thus ?thesis by (intro exI conjI)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   342
qed
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   343
qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   344
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   345
subsection {* The Hahn-Banach Theorem for normed spaces *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   346
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   347
text {* Every continuous linear form $f$ on a subspace $F$ of a
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   348
norm space $E$, can be extended to a continuous linear form $g$ on
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   349
$E$ such that $\fnorm{f} = \fnorm {g}$. *}
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   350
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   351
theorem norm_HahnBanach:
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   352
"[| is_normed_vectorspace E norm; is_subspace F E; 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   353
is_linearform F f; is_continuous F norm f |] 
9503
wenzelm
parents: 9475
diff changeset
   354
==> \<exists>g. is_linearform E g
wenzelm
parents: 9475
diff changeset
   355
     \<and> is_continuous E norm g 
wenzelm
parents: 9475
diff changeset
   356
     \<and> (\<forall>x \<in> F. g x = f x) 
wenzelm
parents: 9475
diff changeset
   357
     \<and> \<parallel>g\<parallel>E,norm = \<parallel>f\<parallel>F,norm"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   358
proof -
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   359
assume e_norm: "is_normed_vectorspace E norm"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   360
assume f: "is_subspace F E" "is_linearform F f"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   361
assume f_cont: "is_continuous F norm f"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   362
have e: "is_vectorspace E" ..
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   363
hence f_norm: "is_normed_vectorspace F norm" ..
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   364
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   365
txt{* We define a function $p$ on $E$ as follows:
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   366
\begin{matharray}{l}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   367
p \: x = \fnorm f \cdot \norm x\\
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   368
\end{matharray}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   369
*}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   370
9503
wenzelm
parents: 9475
diff changeset
   371
def p == "\<lambda>x. \<parallel>f\<parallel>F,norm * norm x"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   372
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   373
txt{* $p$ is a seminorm on $E$: *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   374
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   375
have q: "is_seminorm E p"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   376
proof
9503
wenzelm
parents: 9475
diff changeset
   377
fix x y a assume "x \<in> E" "y \<in> E"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   378
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   379
txt{* $p$ is positive definite: *}
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   380
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   381
show "#0 <= p x"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   382
proof (unfold p_def, rule real_le_mult_order1a)
9503
wenzelm
parents: 9475
diff changeset
   383
  from f_cont f_norm show "#0 <= \<parallel>f\<parallel>F,norm" ..
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   384
  show "#0 <= norm x" ..
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   385
qed
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   386
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   387
txt{* $p$ is absolutely homogenous: *}
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   388
9503
wenzelm
parents: 9475
diff changeset
   389
show "p (a \<cdot> x) = |a| * p x"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   390
proof - 
9503
wenzelm
parents: 9475
diff changeset
   391
  have "p (a \<cdot> x) = \<parallel>f\<parallel>F,norm * norm (a \<cdot> x)"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   392
    by (simp!)
9503
wenzelm
parents: 9475
diff changeset
   393
  also have "norm (a \<cdot> x) = |a| * norm x" 
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   394
    by (rule normed_vs_norm_abs_homogenous)
9503
wenzelm
parents: 9475
diff changeset
   395
  also have "\<parallel>f\<parallel>F,norm * ( |a| * norm x ) 
wenzelm
parents: 9475
diff changeset
   396
    = |a| * (\<parallel>f\<parallel>F,norm * norm x)"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   397
    by (simp! only: real_mult_left_commute)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   398
  also have "... = |a| * p x" by (simp!)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   399
  finally show ?thesis .
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   400
qed
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   401
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   402
txt{* Furthermore, $p$ is subadditive: *}
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   403
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   404
show "p (x + y) <= p x + p y"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   405
proof -
9503
wenzelm
parents: 9475
diff changeset
   406
  have "p (x + y) = \<parallel>f\<parallel>F,norm * norm (x + y)"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   407
    by (simp!)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   408
  also 
9503
wenzelm
parents: 9475
diff changeset
   409
  have "... <= \<parallel>f\<parallel>F,norm * (norm x + norm y)"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   410
  proof (rule real_mult_le_le_mono1a)
9503
wenzelm
parents: 9475
diff changeset
   411
    from f_cont f_norm show "#0 <= \<parallel>f\<parallel>F,norm" ..
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   412
    show "norm (x + y) <= norm x + norm y" ..
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   413
  qed
9503
wenzelm
parents: 9475
diff changeset
   414
  also have "... = \<parallel>f\<parallel>F,norm * norm x 
wenzelm
parents: 9475
diff changeset
   415
                    + \<parallel>f\<parallel>F,norm * norm y"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   416
    by (simp! only: real_add_mult_distrib2)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   417
  finally show ?thesis by (simp!)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   418
qed
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   419
qed
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   420
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   421
txt{* $f$ is bounded by $p$. *} 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   422
9503
wenzelm
parents: 9475
diff changeset
   423
have "\<forall>x \<in> F. |f x| <= p x"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   424
proof
9503
wenzelm
parents: 9475
diff changeset
   425
fix x assume "x \<in> F"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   426
 from f_norm show "|f x| <= p x" 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   427
   by (simp! add: norm_fx_le_norm_f_norm_x)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   428
qed
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   429
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   430
txt{* Using the fact that $p$ is a seminorm and 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   431
$f$ is bounded by $p$ we can apply the Hahn-Banach Theorem 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   432
for real vector spaces. 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   433
So $f$ can be extended in a norm-preserving way to some function
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   434
$g$ on the whole vector space $E$. *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   435
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   436
with e f q 
9503
wenzelm
parents: 9475
diff changeset
   437
have "\<exists>g. is_linearform E g \<and> (\<forall>x \<in> F. g x = f x) 
wenzelm
parents: 9475
diff changeset
   438
        \<and> (\<forall>x \<in> E. |g x| <= p x)"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   439
by (simp! add: abs_HahnBanach)
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   440
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   441
thus ?thesis
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   442
proof (elim exE conjE) 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   443
fix g
9503
wenzelm
parents: 9475
diff changeset
   444
assume "is_linearform E g" and a: "\<forall>x \<in> F. g x = f x" 
wenzelm
parents: 9475
diff changeset
   445
   and b: "\<forall>x \<in> E. |g x| <= p x"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   446
9503
wenzelm
parents: 9475
diff changeset
   447
show "\<exists>g. is_linearform E g 
wenzelm
parents: 9475
diff changeset
   448
        \<and> is_continuous E norm g 
wenzelm
parents: 9475
diff changeset
   449
        \<and> (\<forall>x \<in> F. g x = f x) 
wenzelm
parents: 9475
diff changeset
   450
        \<and> \<parallel>g\<parallel>E,norm = \<parallel>f\<parallel>F,norm"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   451
proof (intro exI conjI)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   452
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   453
txt{* We furthermore have to show that 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   454
$g$ is also continuous: *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   455
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   456
  show g_cont: "is_continuous E norm g"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   457
  proof
9503
wenzelm
parents: 9475
diff changeset
   458
    fix x assume "x \<in> E"
wenzelm
parents: 9475
diff changeset
   459
    with b show "|g x| <= \<parallel>f\<parallel>F,norm * norm x"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   460
      by (simp add: p_def) 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   461
  qed 
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   462
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   463
  txt {* To complete the proof, we show that 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   464
  $\fnorm g = \fnorm f$. \label{order_antisym} *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   465
9503
wenzelm
parents: 9475
diff changeset
   466
  show "\<parallel>g\<parallel>E,norm = \<parallel>f\<parallel>F,norm"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   467
    (is "?L = ?R")
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   468
  proof (rule order_antisym)
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   469
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   470
    txt{* First we show $\fnorm g \leq \fnorm f$.  The function norm
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   471
    $\fnorm g$ is defined as the smallest $c\in\bbbR$ such that
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   472
    \begin{matharray}{l}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   473
    \All {x\in E} {|g\ap x| \leq c \cdot \norm x}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   474
    \end{matharray}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   475
    Furthermore holds
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   476
    \begin{matharray}{l}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   477
    \All {x\in E} {|g\ap x| \leq \fnorm f \cdot \norm x}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   478
    \end{matharray}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   479
    *}
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   480
 
9503
wenzelm
parents: 9475
diff changeset
   481
    have "\<forall>x \<in> E. |g x| <= \<parallel>f\<parallel>F,norm * norm x"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   482
    proof
9503
wenzelm
parents: 9475
diff changeset
   483
      fix x assume "x \<in> E" 
wenzelm
parents: 9475
diff changeset
   484
      show "|g x| <= \<parallel>f\<parallel>F,norm * norm x"
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   485
        by (simp!)
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   486
    qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   487
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   488
    with g_cont e_norm show "?L <= ?R"
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   489
    proof (rule fnorm_le_ub)
9503
wenzelm
parents: 9475
diff changeset
   490
      from f_cont f_norm show "#0 <= \<parallel>f\<parallel>F,norm" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   491
    qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   492
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   493
    txt{* The other direction is achieved by a similar 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   494
    argument. *}
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   495
9503
wenzelm
parents: 9475
diff changeset
   496
    have "\<forall>x \<in> F. |f x| <= \<parallel>g\<parallel>E,norm * norm x"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   497
    proof
9503
wenzelm
parents: 9475
diff changeset
   498
      fix x assume "x \<in> F" 
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   499
      from a have "g x = f x" ..
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   500
      hence "|f x| = |g x|" by simp
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   501
      also from g_cont
9503
wenzelm
parents: 9475
diff changeset
   502
      have "... <= \<parallel>g\<parallel>E,norm * norm x"
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   503
      proof (rule norm_fx_le_norm_f_norm_x)
9503
wenzelm
parents: 9475
diff changeset
   504
        show "x \<in> E" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   505
      qed
9503
wenzelm
parents: 9475
diff changeset
   506
      finally show "|f x| <= \<parallel>g\<parallel>E,norm * norm x" .
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   507
    qed 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   508
    thus "?R <= ?L" 
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   509
    proof (rule fnorm_le_ub [OF f_cont f_norm])
9503
wenzelm
parents: 9475
diff changeset
   510
      from g_cont show "#0 <= \<parallel>g\<parallel>E,norm" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   511
    qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   512
  qed
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   513
qed
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   514
qed
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   515
qed
9374
153853af318b - xsymbols for
bauerg
parents: 9261
diff changeset
   516
9475
b24516d96847 adapted obtain;
wenzelm
parents: 9394
diff changeset
   517
end