src/HOL/Hahn_Banach/Hahn_Banach.thy
changeset 61583 c2b7033fa0ba
parent 61543 de44d4fa5ef0
child 61800 f3789d5b96ca
equal deleted inserted replaced
61582:69492d32263a 61583:c2b7033fa0ba
    47 theorem Hahn_Banach:
    47 theorem Hahn_Banach:
    48   assumes E: "vectorspace E" and "subspace F E"
    48   assumes E: "vectorspace E" and "subspace F E"
    49     and "seminorm E p" and "linearform F f"
    49     and "seminorm E p" and "linearform F f"
    50   assumes fp: "\<forall>x \<in> F. f x \<le> p x"
    50   assumes fp: "\<forall>x \<in> F. f x \<le> p x"
    51   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)"
    51   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)"
    52     -- \<open>Let \<open>E\<close> be a vector space, \<open>F\<close> a subspace of \<open>E\<close>, \<open>p\<close> a seminorm on \<open>E\<close>,\<close>
    52     \<comment> \<open>Let \<open>E\<close> be a vector space, \<open>F\<close> a subspace of \<open>E\<close>, \<open>p\<close> a seminorm on \<open>E\<close>,\<close>
    53     -- \<open>and \<open>f\<close> a linear form on \<open>F\<close> such that \<open>f\<close> is bounded by \<open>p\<close>,\<close>
    53     \<comment> \<open>and \<open>f\<close> a linear form on \<open>F\<close> such that \<open>f\<close> is bounded by \<open>p\<close>,\<close>
    54     -- \<open>then \<open>f\<close> can be extended to a linear form \<open>h\<close> on \<open>E\<close> in a norm-preserving way. \<^smallskip>\<close>
    54     \<comment> \<open>then \<open>f\<close> can be extended to a linear form \<open>h\<close> on \<open>E\<close> in a norm-preserving way. \<^smallskip>\<close>
    55 proof -
    55 proof -
    56   interpret vectorspace E by fact
    56   interpret vectorspace E by fact
    57   interpret subspace F E by fact
    57   interpret subspace F E by fact
    58   interpret seminorm E p by fact
    58   interpret seminorm E p by fact
    59   interpret linearform F f by fact
    59   interpret linearform F f by fact
    62   from E have F: "vectorspace F" ..
    62   from E have F: "vectorspace F" ..
    63   note FE = \<open>F \<unlhd> E\<close>
    63   note FE = \<open>F \<unlhd> E\<close>
    64   {
    64   {
    65     fix c assume cM: "c \<in> chains M" and ex: "\<exists>x. x \<in> c"
    65     fix c assume cM: "c \<in> chains M" and ex: "\<exists>x. x \<in> c"
    66     have "\<Union>c \<in> M"
    66     have "\<Union>c \<in> M"
    67       -- \<open>Show that every non-empty chain \<open>c\<close> of \<open>M\<close> has an upper bound in \<open>M\<close>:\<close>
    67       \<comment> \<open>Show that every non-empty chain \<open>c\<close> of \<open>M\<close> has an upper bound in \<open>M\<close>:\<close>
    68       -- \<open>\<open>\<Union>c\<close> is greater than any element of the chain \<open>c\<close>, so it suffices to show \<open>\<Union>c \<in> M\<close>.\<close>
    68       \<comment> \<open>\<open>\<Union>c\<close> is greater than any element of the chain \<open>c\<close>, so it suffices to show \<open>\<Union>c \<in> M\<close>.\<close>
    69       unfolding M_def
    69       unfolding M_def
    70     proof (rule norm_pres_extensionI)
    70     proof (rule norm_pres_extensionI)
    71       let ?H = "domain (\<Union>c)"
    71       let ?H = "domain (\<Union>c)"
    72       let ?h = "funct (\<Union>c)"
    72       let ?h = "funct (\<Union>c)"
    73 
    73 
    93           \<and> graph F f \<subseteq> graph H h
    93           \<and> graph F f \<subseteq> graph H h
    94           \<and> (\<forall>x \<in> H. h x \<le> p x)" by blast
    94           \<and> (\<forall>x \<in> H. h x \<le> p x)" by blast
    95     qed
    95     qed
    96   }
    96   }
    97   then have "\<exists>g \<in> M. \<forall>x \<in> M. g \<subseteq> x \<longrightarrow> x = g"
    97   then have "\<exists>g \<in> M. \<forall>x \<in> M. g \<subseteq> x \<longrightarrow> x = g"
    98   -- \<open>With Zorn's Lemma we can conclude that there is a maximal element in \<open>M\<close>. \<^smallskip>\<close>
    98   \<comment> \<open>With Zorn's Lemma we can conclude that there is a maximal element in \<open>M\<close>. \<^smallskip>\<close>
    99   proof (rule Zorn's_Lemma)
    99   proof (rule Zorn's_Lemma)
   100       -- \<open>We show that \<open>M\<close> is non-empty:\<close>
   100       \<comment> \<open>We show that \<open>M\<close> is non-empty:\<close>
   101     show "graph F f \<in> M"
   101     show "graph F f \<in> M"
   102       unfolding M_def
   102       unfolding M_def
   103     proof (rule norm_pres_extensionI2)
   103     proof (rule norm_pres_extensionI2)
   104       show "linearform F f" by fact
   104       show "linearform F f" by fact
   105       show "F \<unlhd> E" by fact
   105       show "F \<unlhd> E" by fact
   114       g_rep: "g = graph H h"
   114       g_rep: "g = graph H h"
   115     and linearform: "linearform H h"
   115     and linearform: "linearform H h"
   116     and HE: "H \<unlhd> E" and FH: "F \<unlhd> H"
   116     and HE: "H \<unlhd> E" and FH: "F \<unlhd> H"
   117     and graphs: "graph F f \<subseteq> graph H h"
   117     and graphs: "graph F f \<subseteq> graph H h"
   118     and hp: "\<forall>x \<in> H. h x \<le> p x" unfolding M_def ..
   118     and hp: "\<forall>x \<in> H. h x \<le> p x" unfolding M_def ..
   119       -- \<open>\<open>g\<close> is a norm-preserving extension of \<open>f\<close>, in other words:\<close>
   119       \<comment> \<open>\<open>g\<close> is a norm-preserving extension of \<open>f\<close>, in other words:\<close>
   120       -- \<open>\<open>g\<close> is the graph of some linear form \<open>h\<close> defined on a subspace \<open>H\<close> of \<open>E\<close>,\<close>
   120       \<comment> \<open>\<open>g\<close> is the graph of some linear form \<open>h\<close> defined on a subspace \<open>H\<close> of \<open>E\<close>,\<close>
   121       -- \<open>and \<open>h\<close> is an extension of \<open>f\<close> that is again bounded by \<open>p\<close>. \<^smallskip>\<close>
   121       \<comment> \<open>and \<open>h\<close> is an extension of \<open>f\<close> that is again bounded by \<open>p\<close>. \<^smallskip>\<close>
   122   from HE E have H: "vectorspace H"
   122   from HE E have H: "vectorspace H"
   123     by (rule subspace.vectorspace)
   123     by (rule subspace.vectorspace)
   124 
   124 
   125   have HE_eq: "H = E"
   125   have HE_eq: "H = E"
   126     -- \<open>We show that \<open>h\<close> is defined on whole \<open>E\<close> by classical contradiction. \<^smallskip>\<close>
   126     \<comment> \<open>We show that \<open>h\<close> is defined on whole \<open>E\<close> by classical contradiction. \<^smallskip>\<close>
   127   proof (rule classical)
   127   proof (rule classical)
   128     assume neq: "H \<noteq> E"
   128     assume neq: "H \<noteq> E"
   129       -- \<open>Assume \<open>h\<close> is not defined on whole \<open>E\<close>. Then show that \<open>h\<close> can be extended\<close>
   129       \<comment> \<open>Assume \<open>h\<close> is not defined on whole \<open>E\<close>. Then show that \<open>h\<close> can be extended\<close>
   130       -- \<open>in a norm-preserving way to a function \<open>h'\<close> with the graph \<open>g'\<close>. \<^smallskip>\<close>
   130       \<comment> \<open>in a norm-preserving way to a function \<open>h'\<close> with the graph \<open>g'\<close>. \<^smallskip>\<close>
   131     have "\<exists>g' \<in> M. g \<subseteq> g' \<and> g \<noteq> g'"
   131     have "\<exists>g' \<in> M. g \<subseteq> g' \<and> g \<noteq> g'"
   132     proof -
   132     proof -
   133       from HE have "H \<subseteq> E" ..
   133       from HE have "H \<subseteq> E" ..
   134       with neq obtain x' where x'E: "x' \<in> E" and "x' \<notin> H" by blast
   134       with neq obtain x' where x'E: "x' \<in> E" and "x' \<notin> H" by blast
   135       obtain x': "x' \<noteq> 0"
   135       obtain x': "x' \<noteq> 0"
   141           with \<open>x' \<notin> H\<close> show False by contradiction
   141           with \<open>x' \<notin> H\<close> show False by contradiction
   142         qed
   142         qed
   143       qed
   143       qed
   144 
   144 
   145       def H' \<equiv> "H + lin x'"
   145       def H' \<equiv> "H + lin x'"
   146         -- \<open>Define \<open>H'\<close> as the direct sum of \<open>H\<close> and the linear closure of \<open>x'\<close>. \<^smallskip>\<close>
   146         \<comment> \<open>Define \<open>H'\<close> as the direct sum of \<open>H\<close> and the linear closure of \<open>x'\<close>. \<^smallskip>\<close>
   147       have HH': "H \<unlhd> H'"
   147       have HH': "H \<unlhd> H'"
   148       proof (unfold H'_def)
   148       proof (unfold H'_def)
   149         from x'E have "vectorspace (lin x')" ..
   149         from x'E have "vectorspace (lin x')" ..
   150         with H show "H \<unlhd> H + lin x'" ..
   150         with H show "H \<unlhd> H + lin x'" ..
   151       qed
   151       qed
   152 
   152 
   153       obtain xi where
   153       obtain xi where
   154         xi: "\<forall>y \<in> H. - p (y + x') - h y \<le> xi
   154         xi: "\<forall>y \<in> H. - p (y + x') - h y \<le> xi
   155           \<and> xi \<le> p (y + x') - h y"
   155           \<and> xi \<le> p (y + x') - h y"
   156         -- \<open>Pick a real number \<open>\<xi>\<close> that fulfills certain inequality; this will\<close>
   156         \<comment> \<open>Pick a real number \<open>\<xi>\<close> that fulfills certain inequality; this will\<close>
   157         -- \<open>be used to establish that \<open>h'\<close> is a norm-preserving extension of \<open>h\<close>.
   157         \<comment> \<open>be used to establish that \<open>h'\<close> is a norm-preserving extension of \<open>h\<close>.
   158            \label{ex-xi-use}\<^smallskip>\<close>
   158            \label{ex-xi-use}\<^smallskip>\<close>
   159       proof -
   159       proof -
   160         from H have "\<exists>xi. \<forall>y \<in> H. - p (y + x') - h y \<le> xi
   160         from H have "\<exists>xi. \<forall>y \<in> H. - p (y + x') - h y \<le> xi
   161             \<and> xi \<le> p (y + x') - h y"
   161             \<and> xi \<le> p (y + x') - h y"
   162         proof (rule ex_xi)
   162         proof (rule ex_xi)
   180         then show thesis by (blast intro: that)
   180         then show thesis by (blast intro: that)
   181       qed
   181       qed
   182 
   182 
   183       def h' \<equiv> "\<lambda>x. let (y, a) =
   183       def h' \<equiv> "\<lambda>x. let (y, a) =
   184           SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H in h y + a * xi"
   184           SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H in h y + a * xi"
   185         -- \<open>Define the extension \<open>h'\<close> of \<open>h\<close> to \<open>H'\<close> using \<open>\<xi>\<close>. \<^smallskip>\<close>
   185         \<comment> \<open>Define the extension \<open>h'\<close> of \<open>h\<close> to \<open>H'\<close> using \<open>\<xi>\<close>. \<^smallskip>\<close>
   186 
   186 
   187       have "g \<subseteq> graph H' h' \<and> g \<noteq> graph H' h'"
   187       have "g \<subseteq> graph H' h' \<and> g \<noteq> graph H' h'"
   188         -- \<open>\<open>h'\<close> is an extension of \<open>h\<close> \dots \<^smallskip>\<close>
   188         \<comment> \<open>\<open>h'\<close> is an extension of \<open>h\<close> \dots \<^smallskip>\<close>
   189       proof
   189       proof
   190         show "g \<subseteq> graph H' h'"
   190         show "g \<subseteq> graph H' h'"
   191         proof -
   191         proof -
   192           have  "graph H h \<subseteq> graph H' h'"
   192           have  "graph H h \<subseteq> graph H' h'"
   193           proof (rule graph_extI)
   193           proof (rule graph_extI)
   220           qed
   220           qed
   221           with g_rep show ?thesis by simp
   221           with g_rep show ?thesis by simp
   222         qed
   222         qed
   223       qed
   223       qed
   224       moreover have "graph H' h' \<in> M"
   224       moreover have "graph H' h' \<in> M"
   225         -- \<open>and \<open>h'\<close> is norm-preserving. \<^smallskip>\<close>
   225         \<comment> \<open>and \<open>h'\<close> is norm-preserving. \<^smallskip>\<close>
   226       proof (unfold M_def)
   226       proof (unfold M_def)
   227         show "graph H' h' \<in> norm_pres_extensions E p F f"
   227         show "graph H' h' \<in> norm_pres_extensions E p F f"
   228         proof (rule norm_pres_extensionI2)
   228         proof (rule norm_pres_extensionI2)
   229           show "linearform H' h'"
   229           show "linearform H' h'"
   230             using h'_def H'_def HE linearform \<open>x' \<notin> H\<close> \<open>x' \<in> E\<close> \<open>x' \<noteq> 0\<close> E
   230             using h'_def H'_def HE linearform \<open>x' \<notin> H\<close> \<open>x' \<in> E\<close> \<open>x' \<noteq> 0\<close> E
   268         qed
   268         qed
   269       qed
   269       qed
   270       ultimately show ?thesis ..
   270       ultimately show ?thesis ..
   271     qed
   271     qed
   272     then have "\<not> (\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x)" by simp
   272     then have "\<not> (\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x)" by simp
   273       -- \<open>So the graph \<open>g\<close> of \<open>h\<close> cannot be maximal. Contradiction! \<^smallskip>\<close>
   273       \<comment> \<open>So the graph \<open>g\<close> of \<open>h\<close> cannot be maximal. Contradiction! \<^smallskip>\<close>
   274     with gx show "H = E" by contradiction
   274     with gx show "H = E" by contradiction
   275   qed
   275   qed
   276 
   276 
   277   from HE_eq and linearform have "linearform E h"
   277   from HE_eq and linearform have "linearform E h"
   278     by (simp only:)
   278     by (simp only:)