src/HOL/Real/HahnBanach/HahnBanach.thy
changeset 29197 6d4cb27ed19c
parent 29189 ee8572f3bb57
child 29198 418ed6411847
equal deleted inserted replaced
29189:ee8572f3bb57 29197:6d4cb27ed19c
     1 (*  Title:      HOL/Real/HahnBanach/HahnBanach.thy
       
     2     ID:         $Id$
       
     3     Author:     Gertrud Bauer, TU Munich
       
     4 *)
       
     5 
       
     6 header {* The Hahn-Banach Theorem *}
       
     7 
       
     8 theory HahnBanach
       
     9 imports HahnBanachLemmas
       
    10 begin
       
    11 
       
    12 text {*
       
    13   We present the proof of two different versions of the Hahn-Banach
       
    14   Theorem, closely following \cite[\S36]{Heuser:1986}.
       
    15 *}
       
    16 
       
    17 subsection {* The Hahn-Banach Theorem for vector spaces *}
       
    18 
       
    19 text {*
       
    20   \textbf{Hahn-Banach Theorem.} Let @{text F} be a subspace of a real
       
    21   vector space @{text E}, let @{text p} be a semi-norm on @{text E},
       
    22   and @{text f} be a linear form defined on @{text F} such that @{text
       
    23   f} is bounded by @{text p}, i.e.  @{text "\<forall>x \<in> F. f x \<le> p x"}.  Then
       
    24   @{text f} can be extended to a linear form @{text h} on @{text E}
       
    25   such that @{text h} is norm-preserving, i.e. @{text h} is also
       
    26   bounded by @{text p}.
       
    27 
       
    28   \bigskip
       
    29   \textbf{Proof Sketch.}
       
    30   \begin{enumerate}
       
    31 
       
    32   \item Define @{text M} as the set of norm-preserving extensions of
       
    33   @{text f} to subspaces of @{text E}. The linear forms in @{text M}
       
    34   are ordered by domain extension.
       
    35 
       
    36   \item We show that every non-empty chain in @{text M} has an upper
       
    37   bound in @{text M}.
       
    38 
       
    39   \item With Zorn's Lemma we conclude that there is a maximal function
       
    40   @{text g} in @{text M}.
       
    41 
       
    42   \item The domain @{text H} of @{text g} is the whole space @{text
       
    43   E}, as shown by classical contradiction:
       
    44 
       
    45   \begin{itemize}
       
    46 
       
    47   \item Assuming @{text g} is not defined on whole @{text E}, it can
       
    48   still be extended in a norm-preserving way to a super-space @{text
       
    49   H'} of @{text H}.
       
    50 
       
    51   \item Thus @{text g} can not be maximal. Contradiction!
       
    52 
       
    53   \end{itemize}
       
    54   \end{enumerate}
       
    55 *}
       
    56 
       
    57 theorem HahnBanach:
       
    58   assumes E: "vectorspace E" and "subspace F E"
       
    59     and "seminorm E p" and "linearform F f"
       
    60   assumes fp: "\<forall>x \<in> F. f x \<le> p x"
       
    61   shows "\<exists>h. linearform E h \<and> (\<forall>x \<in> F. h x = f x) \<and> (\<forall>x \<in> E. h x \<le> p x)"
       
    62     -- {* Let @{text E} be a vector space, @{text F} a subspace of @{text E}, @{text p} a seminorm on @{text E}, *}
       
    63     -- {* and @{text f} a linear form on @{text F} such that @{text f} is bounded by @{text p}, *}
       
    64     -- {* then @{text f} can be extended to a linear form @{text h} on @{text E} in a norm-preserving way. \skp *}
       
    65 proof -
       
    66   interpret vectorspace [E] by fact
       
    67   interpret subspace [F E] by fact
       
    68   interpret seminorm [E p] by fact
       
    69   interpret linearform [F f] by fact
       
    70   def M \<equiv> "norm_pres_extensions E p F f"
       
    71   then have M: "M = \<dots>" by (simp only:)
       
    72   from E have F: "vectorspace F" ..
       
    73   note FE = `F \<unlhd> E`
       
    74   {
       
    75     fix c assume cM: "c \<in> chain M" and ex: "\<exists>x. x \<in> c"
       
    76     have "\<Union>c \<in> M"
       
    77       -- {* Show that every non-empty chain @{text c} of @{text M} has an upper bound in @{text M}: *}
       
    78       -- {* @{text "\<Union>c"} is greater than any element of the chain @{text c}, so it suffices to show @{text "\<Union>c \<in> M"}. *}
       
    79       unfolding M_def
       
    80     proof (rule norm_pres_extensionI)
       
    81       let ?H = "domain (\<Union>c)"
       
    82       let ?h = "funct (\<Union>c)"
       
    83 
       
    84       have a: "graph ?H ?h = \<Union>c"
       
    85       proof (rule graph_domain_funct)
       
    86         fix x y z assume "(x, y) \<in> \<Union>c" and "(x, z) \<in> \<Union>c"
       
    87         with M_def cM show "z = y" by (rule sup_definite)
       
    88       qed
       
    89       moreover from M cM a have "linearform ?H ?h"
       
    90         by (rule sup_lf)
       
    91       moreover from a M cM ex FE E have "?H \<unlhd> E"
       
    92         by (rule sup_subE)
       
    93       moreover from a M cM ex FE have "F \<unlhd> ?H"
       
    94         by (rule sup_supF)
       
    95       moreover from a M cM ex have "graph F f \<subseteq> graph ?H ?h"
       
    96         by (rule sup_ext)
       
    97       moreover from a M cM have "\<forall>x \<in> ?H. ?h x \<le> p x"
       
    98         by (rule sup_norm_pres)
       
    99       ultimately show "\<exists>H h. \<Union>c = graph H h
       
   100           \<and> linearform H h
       
   101           \<and> H \<unlhd> E
       
   102           \<and> F \<unlhd> H
       
   103           \<and> graph F f \<subseteq> graph H h
       
   104           \<and> (\<forall>x \<in> H. h x \<le> p x)" by blast
       
   105     qed
       
   106   }
       
   107   then have "\<exists>g \<in> M. \<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
       
   108   -- {* With Zorn's Lemma we can conclude that there is a maximal element in @{text M}. \skp *}
       
   109   proof (rule Zorn's_Lemma)
       
   110       -- {* We show that @{text M} is non-empty: *}
       
   111     show "graph F f \<in> M"
       
   112       unfolding M_def
       
   113     proof (rule norm_pres_extensionI2)
       
   114       show "linearform F f" by fact
       
   115       show "F \<unlhd> E" by fact
       
   116       from F show "F \<unlhd> F" by (rule vectorspace.subspace_refl)
       
   117       show "graph F f \<subseteq> graph F f" ..
       
   118       show "\<forall>x\<in>F. f x \<le> p x" by fact
       
   119     qed
       
   120   qed
       
   121   then obtain g where gM: "g \<in> M" and gx: "\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
       
   122     by blast
       
   123   from gM obtain H h where
       
   124       g_rep: "g = graph H h"
       
   125     and linearform: "linearform H h"
       
   126     and HE: "H \<unlhd> E" and FH: "F \<unlhd> H"
       
   127     and graphs: "graph F f \<subseteq> graph H h"
       
   128     and hp: "\<forall>x \<in> H. h x \<le> p x" unfolding M_def ..
       
   129       -- {* @{text g} is a norm-preserving extension of @{text f}, in other words: *}
       
   130       -- {* @{text g} is the graph of some linear form @{text h} defined on a subspace @{text H} of @{text E}, *}
       
   131       -- {* and @{text h} is an extension of @{text f} that is again bounded by @{text p}. \skp *}
       
   132   from HE E have H: "vectorspace H"
       
   133     by (rule subspace.vectorspace)
       
   134 
       
   135   have HE_eq: "H = E"
       
   136     -- {* We show that @{text h} is defined on whole @{text E} by classical contradiction. \skp *}
       
   137   proof (rule classical)
       
   138     assume neq: "H \<noteq> E"
       
   139       -- {* Assume @{text h} is not defined on whole @{text E}. Then show that @{text h} can be extended *}
       
   140       -- {* in a norm-preserving way to a function @{text h'} with the graph @{text g'}. \skp *}
       
   141     have "\<exists>g' \<in> M. g \<subseteq> g' \<and> g \<noteq> g'"
       
   142     proof -
       
   143       from HE have "H \<subseteq> E" ..
       
   144       with neq obtain x' where x'E: "x' \<in> E" and "x' \<notin> H" by blast
       
   145       obtain x': "x' \<noteq> 0"
       
   146       proof
       
   147         show "x' \<noteq> 0"
       
   148         proof
       
   149           assume "x' = 0"
       
   150           with H have "x' \<in> H" by (simp only: vectorspace.zero)
       
   151           with `x' \<notin> H` show False by contradiction
       
   152         qed
       
   153       qed
       
   154 
       
   155       def H' \<equiv> "H + lin x'"
       
   156         -- {* Define @{text H'} as the direct sum of @{text H} and the linear closure of @{text x'}. \skp *}
       
   157       have HH': "H \<unlhd> H'"
       
   158       proof (unfold H'_def)
       
   159         from x'E have "vectorspace (lin x')" ..
       
   160         with H show "H \<unlhd> H + lin x'" ..
       
   161       qed
       
   162 
       
   163       obtain xi where
       
   164         xi: "\<forall>y \<in> H. - p (y + x') - h y \<le> xi
       
   165           \<and> xi \<le> p (y + x') - h y"
       
   166         -- {* Pick a real number @{text \<xi>} that fulfills certain inequations; this will *}
       
   167         -- {* be used to establish that @{text h'} is a norm-preserving extension of @{text h}.
       
   168            \label{ex-xi-use}\skp *}
       
   169       proof -
       
   170         from H have "\<exists>xi. \<forall>y \<in> H. - p (y + x') - h y \<le> xi
       
   171             \<and> xi \<le> p (y + x') - h y"
       
   172         proof (rule ex_xi)
       
   173           fix u v assume u: "u \<in> H" and v: "v \<in> H"
       
   174           with HE have uE: "u \<in> E" and vE: "v \<in> E" by auto
       
   175           from H u v linearform have "h v - h u = h (v - u)"
       
   176             by (simp add: linearform.diff)
       
   177           also from hp and H u v have "\<dots> \<le> p (v - u)"
       
   178             by (simp only: vectorspace.diff_closed)
       
   179           also from x'E uE vE have "v - u = x' + - x' + v + - u"
       
   180             by (simp add: diff_eq1)
       
   181           also from x'E uE vE have "\<dots> = v + x' + - (u + x')"
       
   182             by (simp add: add_ac)
       
   183           also from x'E uE vE have "\<dots> = (v + x') - (u + x')"
       
   184             by (simp add: diff_eq1)
       
   185           also from x'E uE vE E have "p \<dots> \<le> p (v + x') + p (u + x')"
       
   186             by (simp add: diff_subadditive)
       
   187           finally have "h v - h u \<le> p (v + x') + p (u + x')" .
       
   188           then show "- p (u + x') - h u \<le> p (v + x') - h v" by simp
       
   189         qed
       
   190         then show thesis by (blast intro: that)
       
   191       qed
       
   192 
       
   193       def h' \<equiv> "\<lambda>x. let (y, a) =
       
   194           SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H in h y + a * xi"
       
   195         -- {* Define the extension @{text h'} of @{text h} to @{text H'} using @{text \<xi>}. \skp *}
       
   196 
       
   197       have "g \<subseteq> graph H' h' \<and> g \<noteq> graph H' h'"
       
   198         -- {* @{text h'} is an extension of @{text h} \dots \skp *}
       
   199       proof
       
   200         show "g \<subseteq> graph H' h'"
       
   201         proof -
       
   202           have  "graph H h \<subseteq> graph H' h'"
       
   203           proof (rule graph_extI)
       
   204             fix t assume t: "t \<in> H"
       
   205             from E HE t have "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
       
   206 	      using `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` by (rule decomp_H'_H)
       
   207             with h'_def show "h t = h' t" by (simp add: Let_def)
       
   208           next
       
   209             from HH' show "H \<subseteq> H'" ..
       
   210           qed
       
   211           with g_rep show ?thesis by (simp only:)
       
   212         qed
       
   213 
       
   214         show "g \<noteq> graph H' h'"
       
   215         proof -
       
   216           have "graph H h \<noteq> graph H' h'"
       
   217           proof
       
   218             assume eq: "graph H h = graph H' h'"
       
   219             have "x' \<in> H'"
       
   220 	      unfolding H'_def
       
   221             proof
       
   222               from H show "0 \<in> H" by (rule vectorspace.zero)
       
   223               from x'E show "x' \<in> lin x'" by (rule x_lin_x)
       
   224               from x'E show "x' = 0 + x'" by simp
       
   225             qed
       
   226             then have "(x', h' x') \<in> graph H' h'" ..
       
   227             with eq have "(x', h' x') \<in> graph H h" by (simp only:)
       
   228             then have "x' \<in> H" ..
       
   229             with `x' \<notin> H` show False by contradiction
       
   230           qed
       
   231           with g_rep show ?thesis by simp
       
   232         qed
       
   233       qed
       
   234       moreover have "graph H' h' \<in> M"
       
   235         -- {* and @{text h'} is norm-preserving. \skp *}
       
   236       proof (unfold M_def)
       
   237         show "graph H' h' \<in> norm_pres_extensions E p F f"
       
   238         proof (rule norm_pres_extensionI2)
       
   239           show "linearform H' h'"
       
   240 	    using h'_def H'_def HE linearform `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E
       
   241 	    by (rule h'_lf)
       
   242           show "H' \<unlhd> E"
       
   243 	  unfolding H'_def
       
   244           proof
       
   245             show "H \<unlhd> E" by fact
       
   246             show "vectorspace E" by fact
       
   247             from x'E show "lin x' \<unlhd> E" ..
       
   248           qed
       
   249           from H `F \<unlhd> H` HH' show FH': "F \<unlhd> H'"
       
   250             by (rule vectorspace.subspace_trans)
       
   251           show "graph F f \<subseteq> graph H' h'"
       
   252           proof (rule graph_extI)
       
   253             fix x assume x: "x \<in> F"
       
   254             with graphs have "f x = h x" ..
       
   255             also have "\<dots> = h x + 0 * xi" by simp
       
   256             also have "\<dots> = (let (y, a) = (x, 0) in h y + a * xi)"
       
   257               by (simp add: Let_def)
       
   258             also have "(x, 0) =
       
   259                 (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)"
       
   260 	      using E HE
       
   261             proof (rule decomp_H'_H [symmetric])
       
   262               from FH x show "x \<in> H" ..
       
   263               from x' show "x' \<noteq> 0" .
       
   264 	      show "x' \<notin> H" by fact
       
   265 	      show "x' \<in> E" by fact
       
   266             qed
       
   267             also have
       
   268               "(let (y, a) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)
       
   269               in h y + a * xi) = h' x" by (simp only: h'_def)
       
   270             finally show "f x = h' x" .
       
   271           next
       
   272             from FH' show "F \<subseteq> H'" ..
       
   273           qed
       
   274           show "\<forall>x \<in> H'. h' x \<le> p x"
       
   275 	    using h'_def H'_def `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E HE
       
   276 	      `seminorm E p` linearform and hp xi
       
   277 	    by (rule h'_norm_pres)
       
   278         qed
       
   279       qed
       
   280       ultimately show ?thesis ..
       
   281     qed
       
   282     then have "\<not> (\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x)" by simp
       
   283       -- {* So the graph @{text g} of @{text h} cannot be maximal. Contradiction! \skp *}
       
   284     with gx show "H = E" by contradiction
       
   285   qed
       
   286 
       
   287   from HE_eq and linearform have "linearform E h"
       
   288     by (simp only:)
       
   289   moreover have "\<forall>x \<in> F. h x = f x"
       
   290   proof
       
   291     fix x assume "x \<in> F"
       
   292     with graphs have "f x = h x" ..
       
   293     then show "h x = f x" ..
       
   294   qed
       
   295   moreover from HE_eq and hp have "\<forall>x \<in> E. h x \<le> p x"
       
   296     by (simp only:)
       
   297   ultimately show ?thesis by blast
       
   298 qed
       
   299 
       
   300 
       
   301 subsection  {* Alternative formulation *}
       
   302 
       
   303 text {*
       
   304   The following alternative formulation of the Hahn-Banach
       
   305   Theorem\label{abs-HahnBanach} uses the fact that for a real linear
       
   306   form @{text f} and a seminorm @{text p} the following inequations
       
   307   are equivalent:\footnote{This was shown in lemma @{thm [source]
       
   308   abs_ineq_iff} (see page \pageref{abs-ineq-iff}).}
       
   309   \begin{center}
       
   310   \begin{tabular}{lll}
       
   311   @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
       
   312   @{text "\<forall>x \<in> H. h x \<le> p x"} \\
       
   313   \end{tabular}
       
   314   \end{center}
       
   315 *}
       
   316 
       
   317 theorem abs_HahnBanach:
       
   318   assumes E: "vectorspace E" and FE: "subspace F E"
       
   319     and lf: "linearform F f" and sn: "seminorm E p"
       
   320   assumes fp: "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
       
   321   shows "\<exists>g. linearform E g
       
   322     \<and> (\<forall>x \<in> F. g x = f x)
       
   323     \<and> (\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x)"
       
   324 proof -
       
   325   interpret vectorspace [E] by fact
       
   326   interpret subspace [F E] by fact
       
   327   interpret linearform [F f] by fact
       
   328   interpret seminorm [E p] by fact
       
   329   have "\<exists>g. linearform E g \<and> (\<forall>x \<in> F. g x = f x) \<and> (\<forall>x \<in> E. g x \<le> p x)"
       
   330     using E FE sn lf
       
   331   proof (rule HahnBanach)
       
   332     show "\<forall>x \<in> F. f x \<le> p x"
       
   333       using FE E sn lf and fp by (rule abs_ineq_iff [THEN iffD1])
       
   334   qed
       
   335   then obtain g where lg: "linearform E g" and *: "\<forall>x \<in> F. g x = f x"
       
   336       and **: "\<forall>x \<in> E. g x \<le> p x" by blast
       
   337   have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
       
   338     using _ E sn lg **
       
   339   proof (rule abs_ineq_iff [THEN iffD2])
       
   340     show "E \<unlhd> E" ..
       
   341   qed
       
   342   with lg * show ?thesis by blast
       
   343 qed
       
   344 
       
   345 
       
   346 subsection {* The Hahn-Banach Theorem for normed spaces *}
       
   347 
       
   348 text {*
       
   349   Every continuous linear form @{text f} on a subspace @{text F} of a
       
   350   norm space @{text E}, can be extended to a continuous linear form
       
   351   @{text g} on @{text E} such that @{text "\<parallel>f\<parallel> = \<parallel>g\<parallel>"}.
       
   352 *}
       
   353 
       
   354 theorem norm_HahnBanach:
       
   355   fixes V and norm ("\<parallel>_\<parallel>")
       
   356   fixes B defines "\<And>V f. B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
       
   357   fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
       
   358   defines "\<And>V f. \<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
       
   359   assumes E_norm: "normed_vectorspace E norm" and FE: "subspace F E"
       
   360     and linearform: "linearform F f" and "continuous F norm f"
       
   361   shows "\<exists>g. linearform E g
       
   362      \<and> continuous E norm g
       
   363      \<and> (\<forall>x \<in> F. g x = f x)
       
   364      \<and> \<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
       
   365 proof -
       
   366   interpret normed_vectorspace [E norm] by fact
       
   367   interpret normed_vectorspace_with_fn_norm [E norm B fn_norm]
       
   368     by (auto simp: B_def fn_norm_def) intro_locales
       
   369   interpret subspace [F E] by fact
       
   370   interpret linearform [F f] by fact
       
   371   interpret continuous [F norm f] by fact
       
   372   have E: "vectorspace E" by intro_locales
       
   373   have F: "vectorspace F" by rule intro_locales
       
   374   have F_norm: "normed_vectorspace F norm"
       
   375     using FE E_norm by (rule subspace_normed_vs)
       
   376   have ge_zero: "0 \<le> \<parallel>f\<parallel>\<hyphen>F"
       
   377     by (rule normed_vectorspace_with_fn_norm.fn_norm_ge_zero
       
   378       [OF normed_vectorspace_with_fn_norm.intro,
       
   379        OF F_norm `continuous F norm f` , folded B_def fn_norm_def])
       
   380   txt {* We define a function @{text p} on @{text E} as follows:
       
   381     @{text "p x = \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"} *}
       
   382   def p \<equiv> "\<lambda>x. \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
       
   383 
       
   384   txt {* @{text p} is a seminorm on @{text E}: *}
       
   385   have q: "seminorm E p"
       
   386   proof
       
   387     fix x y a assume x: "x \<in> E" and y: "y \<in> E"
       
   388     
       
   389     txt {* @{text p} is positive definite: *}
       
   390     have "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
       
   391     moreover from x have "0 \<le> \<parallel>x\<parallel>" ..
       
   392     ultimately show "0 \<le> p x"  
       
   393       by (simp add: p_def zero_le_mult_iff)
       
   394 
       
   395     txt {* @{text p} is absolutely homogenous: *}
       
   396 
       
   397     show "p (a \<cdot> x) = \<bar>a\<bar> * p x"
       
   398     proof -
       
   399       have "p (a \<cdot> x) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>a \<cdot> x\<parallel>" by (simp only: p_def)
       
   400       also from x have "\<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>" by (rule abs_homogenous)
       
   401       also have "\<parallel>f\<parallel>\<hyphen>F * (\<bar>a\<bar> * \<parallel>x\<parallel>) = \<bar>a\<bar> * (\<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>)" by simp
       
   402       also have "\<dots> = \<bar>a\<bar> * p x" by (simp only: p_def)
       
   403       finally show ?thesis .
       
   404     qed
       
   405 
       
   406     txt {* Furthermore, @{text p} is subadditive: *}
       
   407 
       
   408     show "p (x + y) \<le> p x + p y"
       
   409     proof -
       
   410       have "p (x + y) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel>" by (simp only: p_def)
       
   411       also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
       
   412       from x y have "\<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>" ..
       
   413       with a have " \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>F * (\<parallel>x\<parallel> + \<parallel>y\<parallel>)"
       
   414         by (simp add: mult_left_mono)
       
   415       also have "\<dots> = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel> + \<parallel>f\<parallel>\<hyphen>F * \<parallel>y\<parallel>" by (simp only: right_distrib)
       
   416       also have "\<dots> = p x + p y" by (simp only: p_def)
       
   417       finally show ?thesis .
       
   418     qed
       
   419   qed
       
   420 
       
   421   txt {* @{text f} is bounded by @{text p}. *}
       
   422 
       
   423   have "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
       
   424   proof
       
   425     fix x assume "x \<in> F"
       
   426     with `continuous F norm f` and linearform
       
   427     show "\<bar>f x\<bar> \<le> p x"
       
   428       unfolding p_def by (rule normed_vectorspace_with_fn_norm.fn_norm_le_cong
       
   429         [OF normed_vectorspace_with_fn_norm.intro,
       
   430          OF F_norm, folded B_def fn_norm_def])
       
   431   qed
       
   432 
       
   433   txt {* Using the fact that @{text p} is a seminorm and @{text f} is bounded
       
   434     by @{text p} we can apply the Hahn-Banach Theorem for real vector
       
   435     spaces. So @{text f} can be extended in a norm-preserving way to
       
   436     some function @{text g} on the whole vector space @{text E}. *}
       
   437 
       
   438   with E FE linearform q obtain g where
       
   439       linearformE: "linearform E g"
       
   440     and a: "\<forall>x \<in> F. g x = f x"
       
   441     and b: "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
       
   442     by (rule abs_HahnBanach [elim_format]) iprover
       
   443 
       
   444   txt {* We furthermore have to show that @{text g} is also continuous: *}
       
   445 
       
   446   have g_cont: "continuous E norm g" using linearformE
       
   447   proof
       
   448     fix x assume "x \<in> E"
       
   449     with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
       
   450       by (simp only: p_def)
       
   451   qed
       
   452 
       
   453   txt {* To complete the proof, we show that @{text "\<parallel>g\<parallel> = \<parallel>f\<parallel>"}. *}
       
   454 
       
   455   have "\<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
       
   456   proof (rule order_antisym)
       
   457     txt {*
       
   458       First we show @{text "\<parallel>g\<parallel> \<le> \<parallel>f\<parallel>"}.  The function norm @{text
       
   459       "\<parallel>g\<parallel>"} is defined as the smallest @{text "c \<in> \<real>"} such that
       
   460       \begin{center}
       
   461       \begin{tabular}{l}
       
   462       @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
       
   463       \end{tabular}
       
   464       \end{center}
       
   465       \noindent Furthermore holds
       
   466       \begin{center}
       
   467       \begin{tabular}{l}
       
   468       @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
       
   469       \end{tabular}
       
   470       \end{center}
       
   471     *}
       
   472 
       
   473     have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
       
   474     proof
       
   475       fix x assume "x \<in> E"
       
   476       with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
       
   477         by (simp only: p_def)
       
   478     qed
       
   479     from g_cont this ge_zero
       
   480     show "\<parallel>g\<parallel>\<hyphen>E \<le> \<parallel>f\<parallel>\<hyphen>F"
       
   481       by (rule fn_norm_least [of g, folded B_def fn_norm_def])
       
   482 
       
   483     txt {* The other direction is achieved by a similar argument. *}
       
   484 
       
   485     show "\<parallel>f\<parallel>\<hyphen>F \<le> \<parallel>g\<parallel>\<hyphen>E"
       
   486     proof (rule normed_vectorspace_with_fn_norm.fn_norm_least
       
   487 	[OF normed_vectorspace_with_fn_norm.intro,
       
   488 	 OF F_norm, folded B_def fn_norm_def])
       
   489       show "\<forall>x \<in> F. \<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
       
   490       proof
       
   491 	fix x assume x: "x \<in> F"
       
   492 	from a x have "g x = f x" ..
       
   493 	then have "\<bar>f x\<bar> = \<bar>g x\<bar>" by (simp only:)
       
   494 	also from g_cont
       
   495 	have "\<dots> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
       
   496 	proof (rule fn_norm_le_cong [of g, folded B_def fn_norm_def])
       
   497 	  from FE x show "x \<in> E" ..
       
   498 	qed
       
   499 	finally show "\<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>" .
       
   500       qed
       
   501       show "0 \<le> \<parallel>g\<parallel>\<hyphen>E"
       
   502 	using g_cont
       
   503 	by (rule fn_norm_ge_zero [of g, folded B_def fn_norm_def])
       
   504       show "continuous F norm f" by fact
       
   505     qed
       
   506   qed
       
   507   with linearformE a g_cont show ?thesis by blast
       
   508 qed
       
   509 
       
   510 end