src/HOL/HahnBanach/HahnBanach.thy
changeset 31795 be3e1cc5005c
parent 31794 71af1fd6a5e4
child 31797 294b955d0e80
child 31799 117300d72398
     1.1 --- a/src/HOL/HahnBanach/HahnBanach.thy	Wed Jun 24 21:28:02 2009 +0200
     1.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
     1.3 @@ -1,509 +0,0 @@
     1.4 -(*  Title:      HOL/Real/HahnBanach/HahnBanach.thy
     1.5 -    Author:     Gertrud Bauer, TU Munich
     1.6 -*)
     1.7 -
     1.8 -header {* The Hahn-Banach Theorem *}
     1.9 -
    1.10 -theory HahnBanach
    1.11 -imports HahnBanachLemmas
    1.12 -begin
    1.13 -
    1.14 -text {*
    1.15 -  We present the proof of two different versions of the Hahn-Banach
    1.16 -  Theorem, closely following \cite[\S36]{Heuser:1986}.
    1.17 -*}
    1.18 -
    1.19 -subsection {* The Hahn-Banach Theorem for vector spaces *}
    1.20 -
    1.21 -text {*
    1.22 -  \textbf{Hahn-Banach Theorem.} Let @{text F} be a subspace of a real
    1.23 -  vector space @{text E}, let @{text p} be a semi-norm on @{text E},
    1.24 -  and @{text f} be a linear form defined on @{text F} such that @{text
    1.25 -  f} is bounded by @{text p}, i.e.  @{text "\<forall>x \<in> F. f x \<le> p x"}.  Then
    1.26 -  @{text f} can be extended to a linear form @{text h} on @{text E}
    1.27 -  such that @{text h} is norm-preserving, i.e. @{text h} is also
    1.28 -  bounded by @{text p}.
    1.29 -
    1.30 -  \bigskip
    1.31 -  \textbf{Proof Sketch.}
    1.32 -  \begin{enumerate}
    1.33 -
    1.34 -  \item Define @{text M} as the set of norm-preserving extensions of
    1.35 -  @{text f} to subspaces of @{text E}. The linear forms in @{text M}
    1.36 -  are ordered by domain extension.
    1.37 -
    1.38 -  \item We show that every non-empty chain in @{text M} has an upper
    1.39 -  bound in @{text M}.
    1.40 -
    1.41 -  \item With Zorn's Lemma we conclude that there is a maximal function
    1.42 -  @{text g} in @{text M}.
    1.43 -
    1.44 -  \item The domain @{text H} of @{text g} is the whole space @{text
    1.45 -  E}, as shown by classical contradiction:
    1.46 -
    1.47 -  \begin{itemize}
    1.48 -
    1.49 -  \item Assuming @{text g} is not defined on whole @{text E}, it can
    1.50 -  still be extended in a norm-preserving way to a super-space @{text
    1.51 -  H'} of @{text H}.
    1.52 -
    1.53 -  \item Thus @{text g} can not be maximal. Contradiction!
    1.54 -
    1.55 -  \end{itemize}
    1.56 -  \end{enumerate}
    1.57 -*}
    1.58 -
    1.59 -theorem HahnBanach:
    1.60 -  assumes E: "vectorspace E" and "subspace F E"
    1.61 -    and "seminorm E p" and "linearform F f"
    1.62 -  assumes fp: "\<forall>x \<in> F. f x \<le> p x"
    1.63 -  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)"
    1.64 -    -- {* Let @{text E} be a vector space, @{text F} a subspace of @{text E}, @{text p} a seminorm on @{text E}, *}
    1.65 -    -- {* and @{text f} a linear form on @{text F} such that @{text f} is bounded by @{text p}, *}
    1.66 -    -- {* then @{text f} can be extended to a linear form @{text h} on @{text E} in a norm-preserving way. \skp *}
    1.67 -proof -
    1.68 -  interpret vectorspace E by fact
    1.69 -  interpret subspace F E by fact
    1.70 -  interpret seminorm E p by fact
    1.71 -  interpret linearform F f by fact
    1.72 -  def M \<equiv> "norm_pres_extensions E p F f"
    1.73 -  then have M: "M = \<dots>" by (simp only:)
    1.74 -  from E have F: "vectorspace F" ..
    1.75 -  note FE = `F \<unlhd> E`
    1.76 -  {
    1.77 -    fix c assume cM: "c \<in> chain M" and ex: "\<exists>x. x \<in> c"
    1.78 -    have "\<Union>c \<in> M"
    1.79 -      -- {* Show that every non-empty chain @{text c} of @{text M} has an upper bound in @{text M}: *}
    1.80 -      -- {* @{text "\<Union>c"} is greater than any element of the chain @{text c}, so it suffices to show @{text "\<Union>c \<in> M"}. *}
    1.81 -      unfolding M_def
    1.82 -    proof (rule norm_pres_extensionI)
    1.83 -      let ?H = "domain (\<Union>c)"
    1.84 -      let ?h = "funct (\<Union>c)"
    1.85 -
    1.86 -      have a: "graph ?H ?h = \<Union>c"
    1.87 -      proof (rule graph_domain_funct)
    1.88 -        fix x y z assume "(x, y) \<in> \<Union>c" and "(x, z) \<in> \<Union>c"
    1.89 -        with M_def cM show "z = y" by (rule sup_definite)
    1.90 -      qed
    1.91 -      moreover from M cM a have "linearform ?H ?h"
    1.92 -        by (rule sup_lf)
    1.93 -      moreover from a M cM ex FE E have "?H \<unlhd> E"
    1.94 -        by (rule sup_subE)
    1.95 -      moreover from a M cM ex FE have "F \<unlhd> ?H"
    1.96 -        by (rule sup_supF)
    1.97 -      moreover from a M cM ex have "graph F f \<subseteq> graph ?H ?h"
    1.98 -        by (rule sup_ext)
    1.99 -      moreover from a M cM have "\<forall>x \<in> ?H. ?h x \<le> p x"
   1.100 -        by (rule sup_norm_pres)
   1.101 -      ultimately show "\<exists>H h. \<Union>c = graph H h
   1.102 -          \<and> linearform H h
   1.103 -          \<and> H \<unlhd> E
   1.104 -          \<and> F \<unlhd> H
   1.105 -          \<and> graph F f \<subseteq> graph H h
   1.106 -          \<and> (\<forall>x \<in> H. h x \<le> p x)" by blast
   1.107 -    qed
   1.108 -  }
   1.109 -  then have "\<exists>g \<in> M. \<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
   1.110 -  -- {* With Zorn's Lemma we can conclude that there is a maximal element in @{text M}. \skp *}
   1.111 -  proof (rule Zorn's_Lemma)
   1.112 -      -- {* We show that @{text M} is non-empty: *}
   1.113 -    show "graph F f \<in> M"
   1.114 -      unfolding M_def
   1.115 -    proof (rule norm_pres_extensionI2)
   1.116 -      show "linearform F f" by fact
   1.117 -      show "F \<unlhd> E" by fact
   1.118 -      from F show "F \<unlhd> F" by (rule vectorspace.subspace_refl)
   1.119 -      show "graph F f \<subseteq> graph F f" ..
   1.120 -      show "\<forall>x\<in>F. f x \<le> p x" by fact
   1.121 -    qed
   1.122 -  qed
   1.123 -  then obtain g where gM: "g \<in> M" and gx: "\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
   1.124 -    by blast
   1.125 -  from gM obtain H h where
   1.126 -      g_rep: "g = graph H h"
   1.127 -    and linearform: "linearform H h"
   1.128 -    and HE: "H \<unlhd> E" and FH: "F \<unlhd> H"
   1.129 -    and graphs: "graph F f \<subseteq> graph H h"
   1.130 -    and hp: "\<forall>x \<in> H. h x \<le> p x" unfolding M_def ..
   1.131 -      -- {* @{text g} is a norm-preserving extension of @{text f}, in other words: *}
   1.132 -      -- {* @{text g} is the graph of some linear form @{text h} defined on a subspace @{text H} of @{text E}, *}
   1.133 -      -- {* and @{text h} is an extension of @{text f} that is again bounded by @{text p}. \skp *}
   1.134 -  from HE E have H: "vectorspace H"
   1.135 -    by (rule subspace.vectorspace)
   1.136 -
   1.137 -  have HE_eq: "H = E"
   1.138 -    -- {* We show that @{text h} is defined on whole @{text E} by classical contradiction. \skp *}
   1.139 -  proof (rule classical)
   1.140 -    assume neq: "H \<noteq> E"
   1.141 -      -- {* Assume @{text h} is not defined on whole @{text E}. Then show that @{text h} can be extended *}
   1.142 -      -- {* in a norm-preserving way to a function @{text h'} with the graph @{text g'}. \skp *}
   1.143 -    have "\<exists>g' \<in> M. g \<subseteq> g' \<and> g \<noteq> g'"
   1.144 -    proof -
   1.145 -      from HE have "H \<subseteq> E" ..
   1.146 -      with neq obtain x' where x'E: "x' \<in> E" and "x' \<notin> H" by blast
   1.147 -      obtain x': "x' \<noteq> 0"
   1.148 -      proof
   1.149 -        show "x' \<noteq> 0"
   1.150 -        proof
   1.151 -          assume "x' = 0"
   1.152 -          with H have "x' \<in> H" by (simp only: vectorspace.zero)
   1.153 -          with `x' \<notin> H` show False by contradiction
   1.154 -        qed
   1.155 -      qed
   1.156 -
   1.157 -      def H' \<equiv> "H + lin x'"
   1.158 -        -- {* Define @{text H'} as the direct sum of @{text H} and the linear closure of @{text x'}. \skp *}
   1.159 -      have HH': "H \<unlhd> H'"
   1.160 -      proof (unfold H'_def)
   1.161 -        from x'E have "vectorspace (lin x')" ..
   1.162 -        with H show "H \<unlhd> H + lin x'" ..
   1.163 -      qed
   1.164 -
   1.165 -      obtain xi where
   1.166 -        xi: "\<forall>y \<in> H. - p (y + x') - h y \<le> xi
   1.167 -          \<and> xi \<le> p (y + x') - h y"
   1.168 -        -- {* Pick a real number @{text \<xi>} that fulfills certain inequations; this will *}
   1.169 -        -- {* be used to establish that @{text h'} is a norm-preserving extension of @{text h}.
   1.170 -           \label{ex-xi-use}\skp *}
   1.171 -      proof -
   1.172 -        from H have "\<exists>xi. \<forall>y \<in> H. - p (y + x') - h y \<le> xi
   1.173 -            \<and> xi \<le> p (y + x') - h y"
   1.174 -        proof (rule ex_xi)
   1.175 -          fix u v assume u: "u \<in> H" and v: "v \<in> H"
   1.176 -          with HE have uE: "u \<in> E" and vE: "v \<in> E" by auto
   1.177 -          from H u v linearform have "h v - h u = h (v - u)"
   1.178 -            by (simp add: linearform.diff)
   1.179 -          also from hp and H u v have "\<dots> \<le> p (v - u)"
   1.180 -            by (simp only: vectorspace.diff_closed)
   1.181 -          also from x'E uE vE have "v - u = x' + - x' + v + - u"
   1.182 -            by (simp add: diff_eq1)
   1.183 -          also from x'E uE vE have "\<dots> = v + x' + - (u + x')"
   1.184 -            by (simp add: add_ac)
   1.185 -          also from x'E uE vE have "\<dots> = (v + x') - (u + x')"
   1.186 -            by (simp add: diff_eq1)
   1.187 -          also from x'E uE vE E have "p \<dots> \<le> p (v + x') + p (u + x')"
   1.188 -            by (simp add: diff_subadditive)
   1.189 -          finally have "h v - h u \<le> p (v + x') + p (u + x')" .
   1.190 -          then show "- p (u + x') - h u \<le> p (v + x') - h v" by simp
   1.191 -        qed
   1.192 -        then show thesis by (blast intro: that)
   1.193 -      qed
   1.194 -
   1.195 -      def h' \<equiv> "\<lambda>x. let (y, a) =
   1.196 -          SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H in h y + a * xi"
   1.197 -        -- {* Define the extension @{text h'} of @{text h} to @{text H'} using @{text \<xi>}. \skp *}
   1.198 -
   1.199 -      have "g \<subseteq> graph H' h' \<and> g \<noteq> graph H' h'"
   1.200 -        -- {* @{text h'} is an extension of @{text h} \dots \skp *}
   1.201 -      proof
   1.202 -        show "g \<subseteq> graph H' h'"
   1.203 -        proof -
   1.204 -          have  "graph H h \<subseteq> graph H' h'"
   1.205 -          proof (rule graph_extI)
   1.206 -            fix t assume t: "t \<in> H"
   1.207 -            from E HE t have "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
   1.208 -	      using `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` by (rule decomp_H'_H)
   1.209 -            with h'_def show "h t = h' t" by (simp add: Let_def)
   1.210 -          next
   1.211 -            from HH' show "H \<subseteq> H'" ..
   1.212 -          qed
   1.213 -          with g_rep show ?thesis by (simp only:)
   1.214 -        qed
   1.215 -
   1.216 -        show "g \<noteq> graph H' h'"
   1.217 -        proof -
   1.218 -          have "graph H h \<noteq> graph H' h'"
   1.219 -          proof
   1.220 -            assume eq: "graph H h = graph H' h'"
   1.221 -            have "x' \<in> H'"
   1.222 -	      unfolding H'_def
   1.223 -            proof
   1.224 -              from H show "0 \<in> H" by (rule vectorspace.zero)
   1.225 -              from x'E show "x' \<in> lin x'" by (rule x_lin_x)
   1.226 -              from x'E show "x' = 0 + x'" by simp
   1.227 -            qed
   1.228 -            then have "(x', h' x') \<in> graph H' h'" ..
   1.229 -            with eq have "(x', h' x') \<in> graph H h" by (simp only:)
   1.230 -            then have "x' \<in> H" ..
   1.231 -            with `x' \<notin> H` show False by contradiction
   1.232 -          qed
   1.233 -          with g_rep show ?thesis by simp
   1.234 -        qed
   1.235 -      qed
   1.236 -      moreover have "graph H' h' \<in> M"
   1.237 -        -- {* and @{text h'} is norm-preserving. \skp *}
   1.238 -      proof (unfold M_def)
   1.239 -        show "graph H' h' \<in> norm_pres_extensions E p F f"
   1.240 -        proof (rule norm_pres_extensionI2)
   1.241 -          show "linearform H' h'"
   1.242 -	    using h'_def H'_def HE linearform `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E
   1.243 -	    by (rule h'_lf)
   1.244 -          show "H' \<unlhd> E"
   1.245 -	  unfolding H'_def
   1.246 -          proof
   1.247 -            show "H \<unlhd> E" by fact
   1.248 -            show "vectorspace E" by fact
   1.249 -            from x'E show "lin x' \<unlhd> E" ..
   1.250 -          qed
   1.251 -          from H `F \<unlhd> H` HH' show FH': "F \<unlhd> H'"
   1.252 -            by (rule vectorspace.subspace_trans)
   1.253 -          show "graph F f \<subseteq> graph H' h'"
   1.254 -          proof (rule graph_extI)
   1.255 -            fix x assume x: "x \<in> F"
   1.256 -            with graphs have "f x = h x" ..
   1.257 -            also have "\<dots> = h x + 0 * xi" by simp
   1.258 -            also have "\<dots> = (let (y, a) = (x, 0) in h y + a * xi)"
   1.259 -              by (simp add: Let_def)
   1.260 -            also have "(x, 0) =
   1.261 -                (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)"
   1.262 -	      using E HE
   1.263 -            proof (rule decomp_H'_H [symmetric])
   1.264 -              from FH x show "x \<in> H" ..
   1.265 -              from x' show "x' \<noteq> 0" .
   1.266 -	      show "x' \<notin> H" by fact
   1.267 -	      show "x' \<in> E" by fact
   1.268 -            qed
   1.269 -            also have
   1.270 -              "(let (y, a) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)
   1.271 -              in h y + a * xi) = h' x" by (simp only: h'_def)
   1.272 -            finally show "f x = h' x" .
   1.273 -          next
   1.274 -            from FH' show "F \<subseteq> H'" ..
   1.275 -          qed
   1.276 -          show "\<forall>x \<in> H'. h' x \<le> p x"
   1.277 -	    using h'_def H'_def `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E HE
   1.278 -	      `seminorm E p` linearform and hp xi
   1.279 -	    by (rule h'_norm_pres)
   1.280 -        qed
   1.281 -      qed
   1.282 -      ultimately show ?thesis ..
   1.283 -    qed
   1.284 -    then have "\<not> (\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x)" by simp
   1.285 -      -- {* So the graph @{text g} of @{text h} cannot be maximal. Contradiction! \skp *}
   1.286 -    with gx show "H = E" by contradiction
   1.287 -  qed
   1.288 -
   1.289 -  from HE_eq and linearform have "linearform E h"
   1.290 -    by (simp only:)
   1.291 -  moreover have "\<forall>x \<in> F. h x = f x"
   1.292 -  proof
   1.293 -    fix x assume "x \<in> F"
   1.294 -    with graphs have "f x = h x" ..
   1.295 -    then show "h x = f x" ..
   1.296 -  qed
   1.297 -  moreover from HE_eq and hp have "\<forall>x \<in> E. h x \<le> p x"
   1.298 -    by (simp only:)
   1.299 -  ultimately show ?thesis by blast
   1.300 -qed
   1.301 -
   1.302 -
   1.303 -subsection  {* Alternative formulation *}
   1.304 -
   1.305 -text {*
   1.306 -  The following alternative formulation of the Hahn-Banach
   1.307 -  Theorem\label{abs-HahnBanach} uses the fact that for a real linear
   1.308 -  form @{text f} and a seminorm @{text p} the following inequations
   1.309 -  are equivalent:\footnote{This was shown in lemma @{thm [source]
   1.310 -  abs_ineq_iff} (see page \pageref{abs-ineq-iff}).}
   1.311 -  \begin{center}
   1.312 -  \begin{tabular}{lll}
   1.313 -  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
   1.314 -  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
   1.315 -  \end{tabular}
   1.316 -  \end{center}
   1.317 -*}
   1.318 -
   1.319 -theorem abs_HahnBanach:
   1.320 -  assumes E: "vectorspace E" and FE: "subspace F E"
   1.321 -    and lf: "linearform F f" and sn: "seminorm E p"
   1.322 -  assumes fp: "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
   1.323 -  shows "\<exists>g. linearform E g
   1.324 -    \<and> (\<forall>x \<in> F. g x = f x)
   1.325 -    \<and> (\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x)"
   1.326 -proof -
   1.327 -  interpret vectorspace E by fact
   1.328 -  interpret subspace F E by fact
   1.329 -  interpret linearform F f by fact
   1.330 -  interpret seminorm E p by fact
   1.331 -  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)"
   1.332 -    using E FE sn lf
   1.333 -  proof (rule HahnBanach)
   1.334 -    show "\<forall>x \<in> F. f x \<le> p x"
   1.335 -      using FE E sn lf and fp by (rule abs_ineq_iff [THEN iffD1])
   1.336 -  qed
   1.337 -  then obtain g where lg: "linearform E g" and *: "\<forall>x \<in> F. g x = f x"
   1.338 -      and **: "\<forall>x \<in> E. g x \<le> p x" by blast
   1.339 -  have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
   1.340 -    using _ E sn lg **
   1.341 -  proof (rule abs_ineq_iff [THEN iffD2])
   1.342 -    show "E \<unlhd> E" ..
   1.343 -  qed
   1.344 -  with lg * show ?thesis by blast
   1.345 -qed
   1.346 -
   1.347 -
   1.348 -subsection {* The Hahn-Banach Theorem for normed spaces *}
   1.349 -
   1.350 -text {*
   1.351 -  Every continuous linear form @{text f} on a subspace @{text F} of a
   1.352 -  norm space @{text E}, can be extended to a continuous linear form
   1.353 -  @{text g} on @{text E} such that @{text "\<parallel>f\<parallel> = \<parallel>g\<parallel>"}.
   1.354 -*}
   1.355 -
   1.356 -theorem norm_HahnBanach:
   1.357 -  fixes V and norm ("\<parallel>_\<parallel>")
   1.358 -  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}"
   1.359 -  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
   1.360 -  defines "\<And>V f. \<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
   1.361 -  assumes E_norm: "normed_vectorspace E norm" and FE: "subspace F E"
   1.362 -    and linearform: "linearform F f" and "continuous F norm f"
   1.363 -  shows "\<exists>g. linearform E g
   1.364 -     \<and> continuous E norm g
   1.365 -     \<and> (\<forall>x \<in> F. g x = f x)
   1.366 -     \<and> \<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
   1.367 -proof -
   1.368 -  interpret normed_vectorspace E norm by fact
   1.369 -  interpret normed_vectorspace_with_fn_norm E norm B fn_norm
   1.370 -    by (auto simp: B_def fn_norm_def) intro_locales
   1.371 -  interpret subspace F E by fact
   1.372 -  interpret linearform F f by fact
   1.373 -  interpret continuous F norm f by fact
   1.374 -  have E: "vectorspace E" by intro_locales
   1.375 -  have F: "vectorspace F" by rule intro_locales
   1.376 -  have F_norm: "normed_vectorspace F norm"
   1.377 -    using FE E_norm by (rule subspace_normed_vs)
   1.378 -  have ge_zero: "0 \<le> \<parallel>f\<parallel>\<hyphen>F"
   1.379 -    by (rule normed_vectorspace_with_fn_norm.fn_norm_ge_zero
   1.380 -      [OF normed_vectorspace_with_fn_norm.intro,
   1.381 -       OF F_norm `continuous F norm f` , folded B_def fn_norm_def])
   1.382 -  txt {* We define a function @{text p} on @{text E} as follows:
   1.383 -    @{text "p x = \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"} *}
   1.384 -  def p \<equiv> "\<lambda>x. \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
   1.385 -
   1.386 -  txt {* @{text p} is a seminorm on @{text E}: *}
   1.387 -  have q: "seminorm E p"
   1.388 -  proof
   1.389 -    fix x y a assume x: "x \<in> E" and y: "y \<in> E"
   1.390 -    
   1.391 -    txt {* @{text p} is positive definite: *}
   1.392 -    have "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
   1.393 -    moreover from x have "0 \<le> \<parallel>x\<parallel>" ..
   1.394 -    ultimately show "0 \<le> p x"  
   1.395 -      by (simp add: p_def zero_le_mult_iff)
   1.396 -
   1.397 -    txt {* @{text p} is absolutely homogenous: *}
   1.398 -
   1.399 -    show "p (a \<cdot> x) = \<bar>a\<bar> * p x"
   1.400 -    proof -
   1.401 -      have "p (a \<cdot> x) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>a \<cdot> x\<parallel>" by (simp only: p_def)
   1.402 -      also from x have "\<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>" by (rule abs_homogenous)
   1.403 -      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
   1.404 -      also have "\<dots> = \<bar>a\<bar> * p x" by (simp only: p_def)
   1.405 -      finally show ?thesis .
   1.406 -    qed
   1.407 -
   1.408 -    txt {* Furthermore, @{text p} is subadditive: *}
   1.409 -
   1.410 -    show "p (x + y) \<le> p x + p y"
   1.411 -    proof -
   1.412 -      have "p (x + y) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel>" by (simp only: p_def)
   1.413 -      also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
   1.414 -      from x y have "\<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>" ..
   1.415 -      with a have " \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>F * (\<parallel>x\<parallel> + \<parallel>y\<parallel>)"
   1.416 -        by (simp add: mult_left_mono)
   1.417 -      also have "\<dots> = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel> + \<parallel>f\<parallel>\<hyphen>F * \<parallel>y\<parallel>" by (simp only: right_distrib)
   1.418 -      also have "\<dots> = p x + p y" by (simp only: p_def)
   1.419 -      finally show ?thesis .
   1.420 -    qed
   1.421 -  qed
   1.422 -
   1.423 -  txt {* @{text f} is bounded by @{text p}. *}
   1.424 -
   1.425 -  have "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
   1.426 -  proof
   1.427 -    fix x assume "x \<in> F"
   1.428 -    with `continuous F norm f` and linearform
   1.429 -    show "\<bar>f x\<bar> \<le> p x"
   1.430 -      unfolding p_def by (rule normed_vectorspace_with_fn_norm.fn_norm_le_cong
   1.431 -        [OF normed_vectorspace_with_fn_norm.intro,
   1.432 -         OF F_norm, folded B_def fn_norm_def])
   1.433 -  qed
   1.434 -
   1.435 -  txt {* Using the fact that @{text p} is a seminorm and @{text f} is bounded
   1.436 -    by @{text p} we can apply the Hahn-Banach Theorem for real vector
   1.437 -    spaces. So @{text f} can be extended in a norm-preserving way to
   1.438 -    some function @{text g} on the whole vector space @{text E}. *}
   1.439 -
   1.440 -  with E FE linearform q obtain g where
   1.441 -      linearformE: "linearform E g"
   1.442 -    and a: "\<forall>x \<in> F. g x = f x"
   1.443 -    and b: "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
   1.444 -    by (rule abs_HahnBanach [elim_format]) iprover
   1.445 -
   1.446 -  txt {* We furthermore have to show that @{text g} is also continuous: *}
   1.447 -
   1.448 -  have g_cont: "continuous E norm g" using linearformE
   1.449 -  proof
   1.450 -    fix x assume "x \<in> E"
   1.451 -    with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
   1.452 -      by (simp only: p_def)
   1.453 -  qed
   1.454 -
   1.455 -  txt {* To complete the proof, we show that @{text "\<parallel>g\<parallel> = \<parallel>f\<parallel>"}. *}
   1.456 -
   1.457 -  have "\<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
   1.458 -  proof (rule order_antisym)
   1.459 -    txt {*
   1.460 -      First we show @{text "\<parallel>g\<parallel> \<le> \<parallel>f\<parallel>"}.  The function norm @{text
   1.461 -      "\<parallel>g\<parallel>"} is defined as the smallest @{text "c \<in> \<real>"} such that
   1.462 -      \begin{center}
   1.463 -      \begin{tabular}{l}
   1.464 -      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
   1.465 -      \end{tabular}
   1.466 -      \end{center}
   1.467 -      \noindent Furthermore holds
   1.468 -      \begin{center}
   1.469 -      \begin{tabular}{l}
   1.470 -      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
   1.471 -      \end{tabular}
   1.472 -      \end{center}
   1.473 -    *}
   1.474 -
   1.475 -    have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
   1.476 -    proof
   1.477 -      fix x assume "x \<in> E"
   1.478 -      with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
   1.479 -        by (simp only: p_def)
   1.480 -    qed
   1.481 -    from g_cont this ge_zero
   1.482 -    show "\<parallel>g\<parallel>\<hyphen>E \<le> \<parallel>f\<parallel>\<hyphen>F"
   1.483 -      by (rule fn_norm_least [of g, folded B_def fn_norm_def])
   1.484 -
   1.485 -    txt {* The other direction is achieved by a similar argument. *}
   1.486 -
   1.487 -    show "\<parallel>f\<parallel>\<hyphen>F \<le> \<parallel>g\<parallel>\<hyphen>E"
   1.488 -    proof (rule normed_vectorspace_with_fn_norm.fn_norm_least
   1.489 -	[OF normed_vectorspace_with_fn_norm.intro,
   1.490 -	 OF F_norm, folded B_def fn_norm_def])
   1.491 -      show "\<forall>x \<in> F. \<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
   1.492 -      proof
   1.493 -	fix x assume x: "x \<in> F"
   1.494 -	from a x have "g x = f x" ..
   1.495 -	then have "\<bar>f x\<bar> = \<bar>g x\<bar>" by (simp only:)
   1.496 -	also from g_cont
   1.497 -	have "\<dots> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
   1.498 -	proof (rule fn_norm_le_cong [OF _ linearformE, folded B_def fn_norm_def])
   1.499 -	  from FE x show "x \<in> E" ..
   1.500 -	qed
   1.501 -	finally show "\<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>" .
   1.502 -      qed
   1.503 -      show "0 \<le> \<parallel>g\<parallel>\<hyphen>E"
   1.504 -	using g_cont
   1.505 -	by (rule fn_norm_ge_zero [of g, folded B_def fn_norm_def])
   1.506 -      show "continuous F norm f" by fact
   1.507 -    qed
   1.508 -  qed
   1.509 -  with linearformE a g_cont show ?thesis by blast
   1.510 -qed
   1.511 -
   1.512 -end