Theory Hahn_Banach

(*  Title:      HOL/Hahn_Banach/Hahn_Banach.thy
    Author:     Gertrud Bauer, TU Munich
*)

section ‹The Hahn-Banach Theorem›

theory Hahn_Banach
imports Hahn_Banach_Lemmas
begin

text ‹
  We present the proof of two different versions of the Hahn-Banach Theorem,
  closely following cite‹\S36› in "Heuser:1986".
›


subsection ‹The Hahn-Banach Theorem for vector spaces›

paragraph ‹Hahn-Banach Theorem.›
text ‹
  Let F› be a subspace of a real vector space E›, let p› be a semi-norm on
  E›, and f› be a linear form defined on F› such that f› is bounded by
  p›, i.e. ∀x ∈ F. f x ≤ p x›. Then f› can be extended to a linear form h›
  on E› such that h› is norm-preserving, i.e. h› is also bounded by p›.
›

paragraph ‹Proof Sketch.›
text  Define M› as the set of norm-preserving extensions of f› to subspaces of
  E›. The linear forms in M› are ordered by domain extension.

   We show that every non-empty chain in M› has an upper bound in M›.

   With Zorn's Lemma we conclude that there is a maximal function g› in M›.

   The domain H› of g› is the whole space E›, as shown by classical
  contradiction:

     Assuming g› is not defined on whole E›, it can still be extended in a
    norm-preserving way to a super-space H'› of H›.

     Thus g› can not be maximal. Contradiction!
›

theorem Hahn_Banach:
  assumes E: "vectorspace E" and "subspace F E"
    and "seminorm E p" and "linearform F f"
  assumes fp: "x  F. f x  p x"
  shows "h. linearform E h  (x  F. h x = f x)  (x  E. h x  p x)"
    ― ‹Let E› be a vector space, F› a subspace of E›, p› a seminorm on E›,›
    ― ‹and f› a linear form on F› such that f› is bounded by p›,›
    ― ‹then f› can be extended to a linear form h› on E› in a norm-preserving way. 
proof -
  interpret vectorspace E by fact
  interpret subspace F E by fact
  interpret seminorm E p by fact
  interpret linearform F f by fact
  define M where "M = norm_pres_extensions E p F f"
  then have M: "M = " by (simp only:)
  from E have F: "vectorspace F" ..
  note FE = F  E
  {
    fix c assume cM: "c  chains M" and ex: "x. x  c"
    have "c  M"
      ― ‹Show that every non-empty chain c› of M› has an upper bound in M›:›
      ― ‹⋃c› is greater than any element of the chain c›, so it suffices to show ⋃c ∈ M›.›
      unfolding M_def
    proof (rule norm_pres_extensionI)
      let ?H = "domain (c)"
      let ?h = "funct (c)"

      have a: "graph ?H ?h = c"
      proof (rule graph_domain_funct)
        fix x y z assume "(x, y)  c" and "(x, z)  c"
        with M_def cM show "z = y" by (rule sup_definite)
      qed
      moreover from M cM a have "linearform ?H ?h"
        by (rule sup_lf)
      moreover from a M cM ex FE E have "?H  E"
        by (rule sup_subE)
      moreover from a M cM ex FE have "F  ?H"
        by (rule sup_supF)
      moreover from a M cM ex have "graph F f  graph ?H ?h"
        by (rule sup_ext)
      moreover from a M cM have "x  ?H. ?h x  p x"
        by (rule sup_norm_pres)
      ultimately show "H h. c = graph H h
           linearform H h
           H  E
           F  H
           graph F f  graph H h
           (x  H. h x  p x)" by blast
    qed
  }
  then have "g  M. x  M. g  x  x = g"
  ― ‹With Zorn's Lemma we can conclude that there is a maximal element in M›. 
  proof (rule Zorn's_Lemma)
      ― ‹We show that M› is non-empty:›
    show "graph F f  M"
      unfolding M_def
    proof (rule norm_pres_extensionI2)
      show "linearform F f" by fact
      show "F  E" by fact
      from F show "F  F" by (rule vectorspace.subspace_refl)
      show "graph F f  graph F f" ..
      show "xF. f x  p x" by fact
    qed
  qed
  then obtain g where gM: "g  M" and gx: "x  M. g  x  g = x"
    by blast
  from gM obtain H h where
      g_rep: "g = graph H h"
    and linearform: "linearform H h"
    and HE: "H  E" and FH: "F  H"
    and graphs: "graph F f  graph H h"
    and hp: "x  H. h x  p x" unfolding M_def ..
      ― ‹g› is a norm-preserving extension of f›, in other words:›
      ― ‹g› is the graph of some linear form h› defined on a subspace H› of E›,›
      ― ‹and h› is an extension of f› that is again bounded by p›. 
  from HE E have H: "vectorspace H"
    by (rule subspace.vectorspace)

  have HE_eq: "H = E"
    ― ‹We show that h› is defined on whole E› by classical contradiction. 
  proof (rule classical)
    assume neq: "H  E"
      ― ‹Assume h› is not defined on whole E›. Then show that h› can be extended›
      ― ‹in a norm-preserving way to a function h'› with the graph g'›. 
    have "g'  M. g  g'  g  g'"
    proof -
      from HE have "H  E" ..
      with neq obtain x' where x'E: "x'  E" and "x'  H" by blast
      obtain x': "x'  0"
      proof
        show "x'  0"
        proof
          assume "x' = 0"
          with H have "x'  H" by (simp only: vectorspace.zero)
          with x'  H show False by contradiction
        qed
      qed

      define H' where "H' = H + lin x'"
        ― ‹Define H'› as the direct sum of H› and the linear closure of x'›. 
      have HH': "H  H'"
      proof (unfold H'_def)
        from x'E have "vectorspace (lin x')" ..
        with H show "H  H + lin x'" ..
      qed

      obtain xi where
        xi: "y  H. - p (y + x') - h y  xi
           xi  p (y + x') - h y"
        ― ‹Pick a real number ξ› that fulfills certain inequality; this will›
        ― ‹be used to establish that h'› is a norm-preserving extension of h›.
           \label{ex-xi-use}
      proof -
        from H have "xi. y  H. - p (y + x') - h y  xi
             xi  p (y + x') - h y"
        proof (rule ex_xi)
          fix u v assume u: "u  H" and v: "v  H"
          with HE have uE: "u  E" and vE: "v  E" by auto
          from H u v linearform have "h v - h u = h (v - u)"
            by (simp add: linearform.diff)
          also from hp and H u v have "  p (v - u)"
            by (simp only: vectorspace.diff_closed)
          also from x'E uE vE have "v - u = x' + - x' + v + - u"
            by (simp add: diff_eq1)
          also from x'E uE vE have " = v + x' + - (u + x')"
            by (simp add: add_ac)
          also from x'E uE vE have " = (v + x') - (u + x')"
            by (simp add: diff_eq1)
          also from x'E uE vE E have "p   p (v + x') + p (u + x')"
            by (simp add: diff_subadditive)
          finally have "h v - h u  p (v + x') + p (u + x')" .
          then show "- p (u + x') - h u  p (v + x') - h v" by simp
        qed
        then show thesis by (blast intro: that)
      qed

      define h' where "h' x = (let (y, a) =
          SOME (y, a). x = y + a  x'  y  H in h y + a * xi)" for x
        ― ‹Define the extension h'› of h› to H'› using ξ›. 

      have "g  graph H' h'  g  graph H' h'"
        ― ‹h'› is an extension of h› \dots 
      proof
        show "g  graph H' h'"
        proof -
          have "graph H h  graph H' h'"
          proof (rule graph_extI)
            fix t assume t: "t  H"
            from E HE t have "(SOME (y, a). t = y + a  x'  y  H) = (t, 0)"
              using x'  H x'  E x'  0 by (rule decomp_H'_H)
            with h'_def show "h t = h' t" by (simp add: Let_def)
          next
            from HH' show "H  H'" ..
          qed
          with g_rep show ?thesis by (simp only:)
        qed

        show "g  graph H' h'"
        proof -
          have "graph H h  graph H' h'"
          proof
            assume eq: "graph H h = graph H' h'"
            have "x'  H'"
              unfolding H'_def
            proof
              from H show "0  H" by (rule vectorspace.zero)
              from x'E show "x'  lin x'" by (rule x_lin_x)
              from x'E show "x' = 0 + x'" by simp
            qed
            then have "(x', h' x')  graph H' h'" ..
            with eq have "(x', h' x')  graph H h" by (simp only:)
            then have "x'  H" ..
            with x'  H show False by contradiction
          qed
          with g_rep show ?thesis by simp
        qed
      qed
      moreover have "graph H' h'  M"
        ― ‹and h'› is norm-preserving. 
      proof (unfold M_def)
        show "graph H' h'  norm_pres_extensions E p F f"
        proof (rule norm_pres_extensionI2)
          show "linearform H' h'"
            using h'_def H'_def HE linearform x'  H x'  E x'  0 E
            by (rule h'_lf)
          show "H'  E"
          unfolding H'_def
          proof
            show "H  E" by fact
            show "vectorspace E" by fact
            from x'E show "lin x'  E" ..
          qed
          from H F  H HH' show FH': "F  H'"
            by (rule vectorspace.subspace_trans)
          show "graph F f  graph H' h'"
          proof (rule graph_extI)
            fix x assume x: "x  F"
            with graphs have "f x = h x" ..
            also have " = h x + 0 * xi" by simp
            also have " = (let (y, a) = (x, 0) in h y + a * xi)"
              by (simp add: Let_def)
            also have "(x, 0) =
                (SOME (y, a). x = y + a  x'  y  H)"
              using E HE
            proof (rule decomp_H'_H [symmetric])
              from FH x show "x  H" ..
              from x' show "x'  0" .
              show "x'  H" by fact
              show "x'  E" by fact
            qed
            also have
              "(let (y, a) = (SOME (y, a). x = y + a  x'  y  H)
              in h y + a * xi) = h' x" by (simp only: h'_def)
            finally show "f x = h' x" .
          next
            from FH' show "F  H'" ..
          qed
          show "x  H'. h' x  p x"
            using h'_def H'_def x'  H x'  E x'  0 E HE
              seminorm E p linearform and hp xi
            by (rule h'_norm_pres)
        qed
      qed
      ultimately show ?thesis ..
    qed
    then have "¬ (x  M. g  x  g = x)" by simp
      ― ‹So the graph g› of h› cannot be maximal. Contradiction! 
    with gx show "H = E" by contradiction
  qed

  from HE_eq and linearform have "linearform E h"
    by (simp only:)
  moreover have "x  F. h x = f x"
  proof
    fix x assume "x  F"
    with graphs have "f x = h x" ..
    then show "h x = f x" ..
  qed
  moreover from HE_eq and hp have "x  E. h x  p x"
    by (simp only:)
  ultimately show ?thesis by blast
qed


subsection ‹Alternative formulation›

text ‹
  The following alternative formulation of the Hahn-Banach
  Theorem\label{abs-Hahn-Banach} uses the fact that for a real linear form f›
  and a seminorm p› the following inequality are equivalent:\footnote{This
  was shown in lemma @{thm [source] abs_ineq_iff} (see page
  \pageref{abs-ineq-iff}).}
  \begin{center}
  \begin{tabular}{lll}
  ∀x ∈ H. ¦h x¦ ≤ p x› & and & ∀x ∈ H. h x ≤ p x› \\
  \end{tabular}
  \end{center}
›

theorem abs_Hahn_Banach:
  assumes E: "vectorspace E" and FE: "subspace F E"
    and lf: "linearform F f" and sn: "seminorm E p"
  assumes fp: "x  F. ¦f x¦  p x"
  shows "g. linearform E g
     (x  F. g x = f x)
     (x  E. ¦g x¦  p x)"
proof -
  interpret vectorspace E by fact
  interpret subspace F E by fact
  interpret linearform F f by fact
  interpret seminorm E p by fact
  have "g. linearform E g  (x  F. g x = f x)  (x  E. g x  p x)"
    using E FE sn lf
  proof (rule Hahn_Banach)
    show "x  F. f x  p x"
      using FE E sn lf and fp by (rule abs_ineq_iff [THEN iffD1])
  qed
  then obtain g where lg: "linearform E g" and *: "x  F. g x = f x"
      and **: "x  E. g x  p x" by blast
  have "x  E. ¦g x¦  p x"
    using _ E sn lg **
  proof (rule abs_ineq_iff [THEN iffD2])
    show "E  E" ..
  qed
  with lg * show ?thesis by blast
qed


subsection ‹The Hahn-Banach Theorem for normed spaces›

text ‹
  Every continuous linear form f› on a subspace F› of a norm space E›, can
  be extended to a continuous linear form g› on E› such that ∥f∥ = ∥g∥›.
›

theorem norm_Hahn_Banach:
  fixes V and norm ("_")
  fixes B defines "V f. B V f  {0}  {¦f x¦ / x | x. x  0  x  V}"
  fixes fn_norm ("_∥‐_" [0, 1000] 999)
  defines "V f. f∥‐V  (B V f)"
  assumes E_norm: "normed_vectorspace E norm" and FE: "subspace F E"
    and linearform: "linearform F f" and "continuous F f norm"
  shows "g. linearform E g
      continuous E g norm
      (x  F. g x = f x)
      g∥‐E = f∥‐F"
proof -
  interpret normed_vectorspace E norm by fact
  interpret normed_vectorspace_with_fn_norm E norm B fn_norm
    by (auto simp: B_def fn_norm_def) intro_locales
  interpret subspace F E by fact
  interpret linearform F f by fact
  interpret continuous F f norm by fact
  have E: "vectorspace E" by intro_locales
  have F: "vectorspace F" by rule intro_locales
  have F_norm: "normed_vectorspace F norm"
    using FE E_norm by (rule subspace_normed_vs)
  have ge_zero: "0  f∥‐F"
    by (rule normed_vectorspace_with_fn_norm.fn_norm_ge_zero
      [OF normed_vectorspace_with_fn_norm.intro,
       OF F_norm continuous F f norm , folded B_def fn_norm_def])
  txt ‹We define a function p› on E› as follows:
    p x = ∥f∥ ⋅ ∥x∥›
  define p where "p x = f∥‐F * x" for x

  txt p› is a seminorm on E›:›
  have q: "seminorm E p"
  proof
    fix x y a assume x: "x  E" and y: "y  E"
    
    txt p› is positive definite:›
    have "0  f∥‐F" by (rule ge_zero)
    moreover from x have "0  x" ..
    ultimately show "0  p x"  
      by (simp add: p_def zero_le_mult_iff)

    txt p› is absolutely homogeneous:›

    show "p (a  x) = ¦a¦ * p x"
    proof -
      have "p (a  x) = f∥‐F * a  x" by (simp only: p_def)
      also from x have "a  x = ¦a¦ * x" by (rule abs_homogenous)
      also have "f∥‐F * (¦a¦ * x) = ¦a¦ * (f∥‐F * x)" by simp
      also have " = ¦a¦ * p x" by (simp only: p_def)
      finally show ?thesis .
    qed

    txt ‹Furthermore, p› is subadditive:›

    show "p (x + y)  p x + p y"
    proof -
      have "p (x + y) = f∥‐F * x + y" by (simp only: p_def)
      also have a: "0  f∥‐F" by (rule ge_zero)
      from x y have "x + y  x + y" ..
      with a have " f∥‐F * x + y  f∥‐F * (x + y)"
        by (simp add: mult_left_mono)
      also have " = f∥‐F * x + f∥‐F * y" by (simp only: distrib_left)
      also have " = p x + p y" by (simp only: p_def)
      finally show ?thesis .
    qed
  qed

  txt f› is bounded by p›.›

  have "x  F. ¦f x¦  p x"
  proof
    fix x assume "x  F"
    with continuous F f norm and linearform
    show "¦f x¦  p x"
      unfolding p_def by (rule normed_vectorspace_with_fn_norm.fn_norm_le_cong
        [OF normed_vectorspace_with_fn_norm.intro,
         OF F_norm, folded B_def fn_norm_def])
  qed

  txt ‹Using the fact that p› is a seminorm and f› is bounded by p› we can
    apply the Hahn-Banach Theorem for real vector spaces. So f› can be
    extended in a norm-preserving way to some function g› on the whole vector
    space E›.›

  with E FE linearform q obtain g where
      linearformE: "linearform E g"
    and a: "x  F. g x = f x"
    and b: "x  E. ¦g x¦  p x"
    by (rule abs_Hahn_Banach [elim_format]) iprover

  txt ‹We furthermore have to show that g› is also continuous:›

  have g_cont: "continuous E g norm" using linearformE
  proof
    fix x assume "x  E"
    with b show "¦g x¦  f∥‐F * x"
      by (simp only: p_def)
  qed

  txt ‹To complete the proof, we show that ∥g∥ = ∥f∥›.›

  have "g∥‐E = f∥‐F"
  proof (rule order_antisym)
    txt ‹
      First we show ∥g∥ ≤ ∥f∥›. The function norm ∥g∥› is defined as the
      smallest c ∈ ℝ› such that
      \begin{center}
      \begin{tabular}{l}
      ∀x ∈ E. ¦g x¦ ≤ c ⋅ ∥x∥›
      \end{tabular}
      \end{center}
       Furthermore holds
      \begin{center}
      \begin{tabular}{l}
      ∀x ∈ E. ¦g x¦ ≤ ∥f∥ ⋅ ∥x∥›
      \end{tabular}
      \end{center}
    ›

    from g_cont _ ge_zero
    show "g∥‐E  f∥‐F"
    proof
      fix x assume "x  E"
      with b show "¦g x¦  f∥‐F * x"
        by (simp only: p_def)
    qed

    txt ‹The other direction is achieved by a similar argument.›

    show "f∥‐F  g∥‐E"
    proof (rule normed_vectorspace_with_fn_norm.fn_norm_least
        [OF normed_vectorspace_with_fn_norm.intro,
         OF F_norm, folded B_def fn_norm_def])
      fix x assume x: "x  F"
      show "¦f x¦  g∥‐E * x"
      proof -
        from a x have "g x = f x" ..
        then have "¦f x¦ = ¦g x¦" by (simp only:)
        also from g_cont have "  g∥‐E * x"
        proof (rule fn_norm_le_cong [OF _ linearformE, folded B_def fn_norm_def])
          from FE x show "x  E" ..
        qed
        finally show ?thesis .
      qed
    next
      show "0  g∥‐E"
        using g_cont by (rule fn_norm_ge_zero [of g, folded B_def fn_norm_def])
      show "continuous F f norm" by fact
    qed
  qed
  with linearformE a g_cont show ?thesis by blast
qed

end