src/HOL/Real/HahnBanach/HahnBanach.thy
author bauerg
Thu, 06 Jul 2000 10:10:10 +0200
changeset 9256 f9a6670427fb
parent 9035 371f023d3dbd
child 9261 879e0f0cd047
permissions -rw-r--r--
completed TYPES version of HahnBanach;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
     1
theory HahnBanach = HahnBanachLemmas: text_raw {* \smallskip\\ *} (* from ~/Pub/TYPES99/HB/HahnBanach.thy *)
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
     2
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
     3
theorem HahnBanach:
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
     4
  "is_vectorspace E \\<Longrightarrow> is_subspace F E \\<Longrightarrow> is_seminorm E p
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
     5
  \\<Longrightarrow> is_linearform F f \\<Longrightarrow> \\<forall>x \\<in> F. f x \\<le> p x
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
     6
  \\<Longrightarrow> \\<exists>h. is_linearform E h \\<and> (\\<forall>x \\<in> F. h x = f x) \\<and> (\\<forall>x \\<in> E. h x \\<le> p x)"   
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
     7
    -- {* 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
     8
    -- {* 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
     9
    -- {* 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
    10
proof -
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    11
  assume "is_vectorspace E" "is_subspace F E" "is_seminorm E p" 
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    12
   and "is_linearform F f" "\\<forall>x \\<in> F. f x \\<le> p x"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    13
  -- {* Assume the context of the theorem. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    14
  def M == "norm_pres_extensions E p F f"
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    15
  -- {* Define $M$ as the set of all norm-preserving extensions of $F$. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    16
  {
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    17
    fix c assume "c \\<in> chain M" "\\<exists>x. x \\<in> c"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    18
    have "\\<Union>c \\<in> M"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    19
    txt {* Show that every non-empty chain $c$ of $M$ has an upper bound in $M$: *}
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    20
    txt {* $\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
    21
    proof (unfold M_def, rule norm_pres_extensionI)
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    22
      show "EX (H::'a set) h::'a => real. graph H h = Union c
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    23
              & is_linearform H h 
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    24
              & is_subspace H E 
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    25
              & is_subspace F H 
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    26
              & graph F f <= graph H h
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    27
              & (ALL x::'a:H. h x <= p x)"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    28
      proof (intro exI conjI)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    29
        let ?H = "domain (Union c)"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    30
        let ?h = "funct (Union c)"
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    31
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    32
        show a: "graph ?H ?h = Union c" 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    33
        proof (rule graph_domain_funct)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    34
          fix x y z assume "(x, y) : Union c" "(x, z) : Union c"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    35
          show "z = y" by (rule sup_definite)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    36
        qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    37
        show "is_linearform ?H ?h" 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    38
          by (simp! add: sup_lf a)
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    39
        show "is_subspace ?H E" thm sup_subE [OF _ _ _ a]
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    40
          by (rule sup_subE [OF _ _ _ a]) (simp !)+
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    41
  (*  FIXME       by (rule sup_subE, rule a) (simp!)+; *)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    42
        show "is_subspace F ?H" 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    43
          by (rule sup_supF [OF _ _ _ a]) (simp!)+
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    44
  (* FIXME        by (rule sup_supF, rule a) (simp!)+ *)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    45
        show "graph F f <= graph ?H ?h" 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    46
          by (rule sup_ext [OF _ _ _ a]) (simp!)+
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    47
  (*  FIXME      by (rule sup_ext, rule a) (simp!)+*)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    48
        show "ALL x::'a:?H. ?h x <= p x" 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    49
          by (rule sup_norm_pres [OF _ _ a]) (simp!)+
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    50
  (* FIXME        by (rule sup_norm_pres, rule a) (simp!)+ *)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    51
      qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    52
    qed
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    53
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    54
  }
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    55
  hence "\\<exists>g \\<in> M. \\<forall>x \\<in> M. g \\<subseteq> x \\<longrightarrow> g = x" 
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    56
  txt {* 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
    57
  proof (rule Zorn's_Lemma)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    58
    txt {* We show that $M$ is non-empty: *}
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    59
    have "graph F f : norm_pres_extensions E p F f"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    60
    proof (rule norm_pres_extensionI2)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    61
      have "is_vectorspace F" ..
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    62
      thus "is_subspace F F" ..
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    63
    qed (blast!)+ 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    64
    thus "graph F f : M" by (simp!)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    65
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    66
  thus ?thesis
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    67
  proof
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    68
    fix g assume "g \\<in> M" "\\<forall>x \\<in> M. g \\<subseteq> x \\<longrightarrow> g = x"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    69
    -- {* We consider such a maximal element $g \in M$. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    70
    show ?thesis
8109
aca11f954993 obtain: renamed 'in' to 'where';
wenzelm
parents: 8108
diff changeset
    71
      obtain H h where "graph H h = g" "is_linearform H h" 
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    72
        "is_subspace H E" "is_subspace F H" "graph F f \\<subseteq> graph H h" 
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    73
        "\\<forall>x \\<in> H. h x \\<le> p x" 
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    74
        txt {* $g$ is a norm-preserving extension of $f$, in other words: *}
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    75
        txt {* $g$ is the graph of some linear form $h$ defined on a subspace $H$ of $E$, *}
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    76
        txt {* and $h$ is an extension of $f$ that is again bounded by $p$. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    77
      proof -
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    78
        have "EX H h. graph H h = g & is_linearform H h 
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    79
          & is_subspace H E & is_subspace F H
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    80
          & graph F f <= graph H h
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    81
          & (ALL x:H. h x <= p x)" by (simp! add: norm_pres_extension_D)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    82
        thus ?thesis by (elim exE conjE) rule
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    83
      qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    84
      have h: "is_vectorspace H" ..
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    85
      have "H = E"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    86
      -- {* We show that $h$ is defined on whole $E$ by classical contradiction. \skp *} 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    87
      proof (rule classical)
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    88
        assume "H \\<noteq> E"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    89
        -- {* Assume $h$ is not defined on whole $E$. Then show that $h$ can be extended *}
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    90
        -- {* in a norm-preserving way to a function $h'$ with the graph $g'$. \skp *}
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    91
        have "\\<exists>g' \\<in> M. g \\<subseteq> g' \\<and> g \\<noteq> g'" 
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    92
          obtain x' where "x' \\<in> E" "x' \\<notin> H" 
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    93
          txt {* Pick $x' \in E \setminus H$. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    94
          proof -
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
    95
            have "EX x':E. x'~:H"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    96
            proof (rule set_less_imp_diff_not_empty)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    97
              have "H <= E" ..
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    98
              thus "H < E" ..
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    99
            qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   100
            thus ?thesis by blast
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   101
          qed
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   102
          have x': "x' ~= \<zero>"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   103
          proof (rule classical)
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   104
            presume "x' = \<zero>"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   105
            with h have "x':H" by simp
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   106
            thus ?thesis by contradiction
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   107
          qed blast
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   108
          def H' == "H + lin x'"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   109
          -- {* Define $H'$ as the direct sum of $H$ and the linear closure of $x'$. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   110
          show ?thesis
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   111
            obtain xi where "\\<forall>y \\<in> H. - p (y + x') - h y \\<le> xi 
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   112
                              \\<and> xi \\<le> p (y + x') - h y" sorry
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   113
            -- {* Pick a real number $\xi$ that fulfills certain inequations; this will *}
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   114
            -- {* be used to establish that $h'$ is a norm-preserving extension of $h$. \skp *}
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   115
            def h' == "\\<lambda>x. let (y,a) = \\<epsilon>(y,a). x = y + a \<prod> x' \\<and> y \\<in> H
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   116
                           in (h y) + a * xi"
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   117
            -- {* Define the extension $h'$ of $h$ to $H'$ using $\xi$. \skp *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   118
            show ?thesis
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   119
            proof
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   120
              show "g \\<subseteq> graph H' h' \\<and> g \\<noteq> graph H' h'" 
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   121
              txt {* Show that $h'$ is an extension of $h$ \dots \skp *}
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   122
proof
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   123
		show "g <= graph H' h'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   124
		proof -
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   125
		  have  "graph H h <= graph H' h'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   126
                  proof (rule graph_extI)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   127
		    fix t assume "t:H" 
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   128
		    have "(SOME (y, a). t = y + a \<prod> x' & y : H)
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   129
                         = (t, #0)" 
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   130
		      by (rule decomp_H0_H [OF _ _ _ _ _ x'])
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   131
		    thus "h t = h' t" by (simp! add: Let_def)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   132
		  next
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   133
		    show "H <= H'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   134
		    proof (rule subspace_subset)
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   135
		      show "is_subspace H H'"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   136
		      proof (unfold H'_def, rule subspace_vs_sum1)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   137
			show "is_vectorspace H" ..
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   138
			show "is_vectorspace (lin x')" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   139
		      qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   140
		    qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   141
		  qed 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   142
		  thus ?thesis by (simp!)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   143
		qed
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   144
                show "g ~= graph H' h'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   145
		proof -
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   146
		  have "graph H h ~= graph H' h'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   147
		  proof
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   148
		    assume e: "graph H h = graph H' h'"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   149
		    have "x' : H'" 
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   150
		    proof (unfold H'_def, rule vs_sumI)
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   151
		      show "x' = \<zero> + x'" by (simp!)
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   152
		      from h show "\<zero> : H" ..
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   153
		      show "x' : lin x'" by (rule x_lin_x)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   154
		    qed
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   155
		    hence "(x', h' x') : graph H' h'" ..
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   156
		    with e have "(x', h' x') : graph H h" by simp
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   157
		    hence "x' : H" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   158
		    thus False by contradiction
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   159
		  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   160
		  thus ?thesis by (simp!)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   161
		qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   162
              qed
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   163
	      show "graph H' h' \\<in> M" 
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   164
              txt {* and $h'$ is norm-preserving. \skp *} 
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   165
              proof -
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   166
		have "graph H' h' : norm_pres_extensions E p F f"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   167
		proof (rule norm_pres_extensionI2)
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   168
		  show "is_linearform H' h'"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   169
		    by (rule h0_lf [OF _ _ _ _ _ _ x']) (simp!)+
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   170
		  show "is_subspace H' E"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   171
		    by (unfold H'_def) (rule vs_sum_subspace [OF _ lin_subspace])
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   172
		  have "is_subspace F H" .
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   173
		  also from h lin_vs 
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   174
		  have [fold H'_def]: "is_subspace H (H + lin x')" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   175
		  finally (subspace_trans [OF _ h]) 
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   176
		  show f_h': "is_subspace F H'" .
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
   177
		
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   178
		  show "graph F f <= graph H' h'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   179
		  proof (rule graph_extI)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   180
		    fix x assume "x:F"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   181
		    have "f x = h x" ..
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   182
		    also have " ... = h x + #0 * xi" by simp
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   183
		    also have "... = (let (y,a) = (x, #0) in h y + a * xi)"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   184
		      by (simp add: Let_def)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   185
		    also have 
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   186
		      "(x, #0) = (SOME (y, a). x = y + a (*) x' & y : H)"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   187
		      by (rule decomp_H0_H [RS sym, OF _ _ _ _ _ x']) (simp!)+
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   188
		    also have 
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   189
		      "(let (y,a) = (SOME (y,a). x = y + a (*) x' & y : H)
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
   190
                        in h y + a * xi) 
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   191
                      = h' x" by (simp!)
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   192
		    finally show "f x = h' x" .
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   193
		  next
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   194
		    from f_h' show "F <= H'" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   195
		  qed
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
   196
		
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   197
		  show "ALL x:H'. h' x <= p x"
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   198
		    by (rule h0_norm_pres [OF _ _ _ _ x'])
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   199
		qed
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   200
		thus "graph H' h' : M" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   201
	      qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   202
            qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   203
          qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   204
        qed
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   205
        hence "\\<not>(\\<forall>x \\<in> M. g \\<subseteq> x \\<longrightarrow> g = x)" by simp
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   206
        -- {* So the graph $g$ of $h$ cannot be maximal. Contradiction! \skp *}
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   207
        thus "H = E" by contradiction
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   208
      qed
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   209
      thus "\\<exists>h. is_linearform E h \\<and> (\\<forall>x \\<in> F. h x = f x) 
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   210
          \\<and> (\\<forall>x \\<in> E. h x \\<le> p x)" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   211
      proof (intro exI conjI)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   212
        assume eq: "H = E"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   213
	from eq show "is_linearform E h" by (simp!)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   214
	show "ALL x:F. h x = f x" 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   215
	proof (intro ballI, rule sym)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   216
	  fix x assume "x:F" show "f x = h x " ..
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   217
	qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   218
	from eq show "ALL x:E. h x <= p x" by (force!)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   219
      qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   220
    qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   221
  qed
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
   222
qed text_raw {* \smallskip\\ *}
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   223
end