src/HOL/Real/HahnBanach/HahnBanach_h0_lemmas.thy
author wenzelm
Wed, 29 Sep 1999 16:41:52 +0200
changeset 7656 2f18c0ffc348
parent 7566 c5a3f980a7af
child 7808 fd019ac3485f
permissions -rw-r--r--
update from Gertrud;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     1
(*  Title:      HOL/Real/HahnBanach/HahnBanach_h0_lemmas.thy
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     2
    ID:         $Id$
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     3
    Author:     Gertrud Bauer, TU Munich
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     4
*)
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     5
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     6
theory HahnBanach_h0_lemmas = FunctionNorm:;
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     7
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     8
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     9
lemma ex_xi: "[| is_vectorspace F;  (!! u v. [| u:F; v:F |] ==> a u <= b v )|]
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    10
       ==> EX xi::real. (ALL y:F. (a::'a => real) y <= xi) & (ALL y:F. xi <= b y)"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    11
proof -;
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
    12
  assume vs: "is_vectorspace F";
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    13
  assume r: "(!! u v. [| u:F; v:F |] ==> a u <= (b v::real))";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    14
  have "EX t. isLub UNIV {s::real . EX u:F. s = a u} t";  
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    15
  proof (rule reals_complete);
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
    16
    from vs; have "a <0> : {s. EX u:F. s = a u}"; by (force);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    17
    thus "EX X. X : {s. EX u:F. s = a u}"; ..;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    18
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    19
    show "EX Y. isUb UNIV {s. EX u:F. s = a u} Y"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    20
    proof; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    21
      show "isUb UNIV {s. EX u:F. s = a u} (b <0>)";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    22
      proof (intro isUbI setleI ballI, fast);
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
    23
        fix y; assume y: "y : {s. EX u:F. s = a u}"; 
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    24
        show "y <= b <0>"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    25
        proof -;
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
    26
          from y; have "EX u:F. y = a u"; by (fast);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    27
          thus ?thesis;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    28
          proof;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    29
            fix u; assume "u:F"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    30
            assume "y = a u";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    31
            also; have "a u <= b <0>"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    32
            proof (rule r);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    33
              show "<0> : F"; ..;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    34
            qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    35
            finally; show ?thesis;.;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    36
          qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    37
        qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    38
      qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    39
    qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    40
  qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    41
  thus "EX xi. (ALL y:F. a y <= xi) & (ALL y:F. xi <= b y)"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    42
  proof (elim exE);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    43
    fix t; assume "isLub UNIV {s::real . EX u:F. s = a u} t"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    44
    show ?thesis;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    45
    proof (intro exI conjI ballI); 
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
    46
      fix y; assume y: "y:F";
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    47
      show "a y <= t";    
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    48
      proof (rule isUbD);  
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    49
        show"isUb UNIV {s. EX u:F. s = a u} t"; ..;
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
    50
      qed (force simp add: y);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    51
    next;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    52
      fix y; assume "y:F";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    53
      show "t <= b y";  
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    54
      proof (intro isLub_le_isUb isUbI setleI);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    55
        show "ALL ya : {s. EX u:F. s = a u}. ya <= b y"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    56
        proof (intro ballI); 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    57
          fix au; 
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
    58
          assume au: "au : {s. EX u:F. s = a u} ";
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    59
          show "au <= b y";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    60
          proof -; 
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
    61
            from au; have "EX u:F. au = a u"; by (fast); 
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    62
            thus "au <= b y";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    63
            proof;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    64
              fix u; assume "u:F";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    65
              assume "au = a u";  
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    66
              also; have "... <= b y"; by (rule r);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    67
              finally; show ?thesis; .;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    68
            qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    69
          qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    70
        qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    71
        show "b y : UNIV"; by fast;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    72
      qed; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    73
    qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    74
  qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    75
qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    76
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    77
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    78
lemma h0_lf: 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    79
  "[| h0 = (%x. let (y, a) = @ (y, a). (x = y [+] a [*] x0 & y:H) in (h y) + a * xi);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    80
      H0 = vectorspace_sum H (lin x0); is_subspace H E; is_linearform H h; x0 ~: H; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    81
      x0 : E; x0 ~= <0>; is_vectorspace E |]
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    82
    ==> is_linearform H0 h0";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    83
proof -;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    84
  assume h0_def:  "h0 = (%x. let (y, a) = @ (y, a). (x = y [+] a [*] x0 & y:H) in (h y) + a * xi)"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    85
    and H0_def: "H0 = vectorspace_sum H (lin x0)"
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
    86
    and vs: "is_subspace H E" "is_linearform H h" "x0 ~: H" "x0 ~= <0>" "x0 : E" 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
    87
            "is_vectorspace E";
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    88
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    89
  have h0: "is_vectorspace H0"; 
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
    90
  proof (simp only: H0_def, rule vs_sum_vs);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    91
    show "is_subspace (lin x0) E"; by (rule lin_subspace); 
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
    92
  qed; 
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    93
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    94
  show ?thesis;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    95
  proof;
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
    96
    fix x1 x2; assume x1: "x1 : H0" and x2: "x2 : H0"; 
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    97
    have x1x2: "x1 [+] x2 : H0"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    98
      by (rule vs_add_closed, rule h0);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    99
  
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   100
    from x1; have ex_x1: "? y1 a1. (x1 = y1 [+] a1 [*] x0  & y1 : H)"; 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   101
      by (simp add: H0_def vectorspace_sum_def lin_def) blast;
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   102
    from x2; have ex_x2: "? y2 a2. (x2 = y2 [+] a2 [*] x0 & y2 : H)"; 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   103
      by (simp add: H0_def vectorspace_sum_def lin_def) blast;
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   104
    from x1x2; have ex_x1x2: "? y a. (x1 [+] x2 = y [+] a [*] x0  & y : H)";    
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   105
      by (simp add: H0_def vectorspace_sum_def lin_def) force;
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   106
    from ex_x1 ex_x2 ex_x1x2;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   107
    show "h0 (x1 [+] x2) = h0 x1 + h0 x2";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   108
    proof (elim exE conjE);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   109
      fix y1 y2 y a1 a2 a;
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   110
      assume y1: "x1 = y1 [+] a1 [*] x0"       and y1': "y1 : H"
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   111
         and y2: "x2 = y2 [+] a2 [*] x0"       and y2': "y2 : H" 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   112
         and y: "x1 [+] x2 = y  [+] a  [*] x0" and y': "y  : H"; 
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   113
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   114
      have ya: "y1 [+] y2 = y & a1 + a2 = a"; 
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   115
      proof (rule decomp4); 
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   116
        show "y1 [+] y2 [+] (a1 + a2) [*] x0 = y [+] a [*] x0"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   117
        proof -;
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   118
          have "y [+] a [*] x0 = x1 [+] x2"; by (rule sym);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   119
          also; from y1 y2; have "... = y1 [+] a1 [*] x0 [+] (y2 [+] a2 [*] x0)"; by simp; 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   120
          also; from vs y1' y2'; have "... = y1 [+] y2 [+] (a1 [*] x0 [+] a2 [*] x0)"; by simp;
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   121
          also; from vs y1' y2'; have "... = y1 [+] y2 [+] (a1 + a2) [*] x0"; 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   122
            by (simp add: vs_add_mult_distrib2[of E]);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   123
          finally; show ?thesis; by (rule sym);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   124
        qed;
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   125
        show "y1 [+] y2 : H"; ..;
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   126
      qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   127
      have y: "y1 [+] y2 = y"; by (rule conjunct1 [OF ya]);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   128
      have a: "a1 + a2 = a";  by (rule conjunct2 [OF ya]);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   129
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   130
      have "h0 (x1 [+] x2) = h y + a * xi";
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   131
	by (rule decomp3);
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   132
      also; have "... = h (y1 [+] y2) + (a1 + a2) * xi"; by (simp add: y a);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   133
      also; from vs y1' y2'; have  "... = h y1 + h y2 + a1 * xi + a2 * xi"; 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   134
	by (simp add: linearform_add_linear [of H] real_add_mult_distrib);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   135
      also; have "... = (h y1 + a1 * xi) + (h y2 + a2 * xi)"; by (simp);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   136
      also; have "h y1 + a1 * xi = h0 x1"; by (rule decomp3 [RS sym]);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   137
      also; have "h y2 + a2 * xi = h0 x2"; by (rule decomp3 [RS sym]);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   138
      finally; show ?thesis; .;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   139
    qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   140
 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   141
  next;  
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   142
    fix c x1; assume x1: "x1 : H0";
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   143
    
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   144
    have ax1: "c [*] x1 : H0";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   145
      by (rule vs_mult_closed, rule h0);
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   146
    from x1; have ex_x1: "? y1 a1. (x1 = y1 [+] a1 [*] x0 & y1 : H)";
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   147
      by (simp add: H0_def vectorspace_sum_def lin_def, fast);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   148
    from x1; have ex_x: "!! x. x: H0 ==> ? y a. (x = y [+] a [*] x0 & y : H)"; 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   149
      by (simp add: H0_def vectorspace_sum_def lin_def, fast);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   150
    note ex_ax1 = ex_x [of "c [*] x1", OF ax1];
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   151
    from ex_x1 ex_ax1; show "h0 (c [*] x1) = c * (h0 x1)";  
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   152
    proof (elim exE conjE);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   153
      fix y1 y a1 a; 
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   154
      assume y1: "x1 = y1 [+] a1 [*] x0"        and y1': "y1 : H"
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   155
	 and y:  "c [*] x1 = y  [+] a  [*] x0"  and y':  "y  : H"; 
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   156
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   157
      have ya: "c [*] y1 = y & c * a1 = a"; 
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   158
      proof (rule decomp4); 
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   159
	show "c [*] y1 [+] (c * a1) [*] x0 = y [+] a [*] x0"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   160
	proof -; 
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   161
          have "y [+] a [*] x0 = c [*] x1"; by (rule sym);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   162
          also; from y1; have "... = c [*] (y1 [+] a1 [*] x0)"; by simp;
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   163
          also; from vs y1'; have "... = c [*] y1 [+] c [*] (a1 [*] x0)"; 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   164
            by (simp add: vs_add_mult_distrib1);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   165
          also; from vs y1'; have "... = c [*] y1 [+] (c * a1) [*] x0"; by simp;
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   166
          finally; show ?thesis; by (rule sym);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   167
        qed;
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   168
        show "c [*] y1: H"; ..;
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   169
      qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   170
      have y: "c [*] y1 = y"; by (rule conjunct1 [OF ya]);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   171
      have a: "c * a1 = a"; by (rule conjunct2 [OF ya]);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   172
      
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   173
      have "h0 (c [*] x1) = h y + a * xi"; 
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   174
	by (rule decomp3);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   175
      also; have "... = h (c [*] y1) + (c * a1) * xi";
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   176
        by (simp add: y a);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   177
      also; from vs y1'; have  "... = c * h y1 + c * a1 * xi"; 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   178
	by (simp add: linearform_mult_linear [of H] real_add_mult_distrib);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   179
      also; from vs y1'; have "... = c * (h y1 + a1 * xi)"; 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   180
	by (simp add: real_add_mult_distrib2 real_mult_assoc);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   181
      also; have "h y1 + a1 * xi = h0 x1"; by (rule decomp3 [RS sym]);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   182
      finally; show ?thesis; .;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   183
    qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   184
  qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   185
qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   186
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   187
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   188
theorem real_linear_split:
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   189
 "[| x < a ==> Q; x = a ==> Q; a < (x::real) ==> Q |] ==> Q";
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   190
  by (rule real_linear [of x a, elimify], elim disjE, force+);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   191
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   192
theorem linorder_linear_split: 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   193
"[| x < a ==> Q; x = a ==> Q; a < (x::'a::linorder) ==> Q |] ==> Q";
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   194
  by (rule linorder_less_linear [of x a, elimify], elim disjE, force+);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   195
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   196
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   197
lemma h0_norm_pres:
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   198
    "[| h0 = (%x. let (y, a) = @ (y, a). (x = y [+] a [*] x0 & y:H) in (h y) + a * xi);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   199
      H0 = vectorspace_sum H (lin x0); x0 ~: H; x0 : E; x0 ~= <0>; is_vectorspace E; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   200
      is_subspace H E; is_quasinorm E p; is_linearform H h; ALL y:H. h y <= p y;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   201
      (ALL y:H. - p (y [+] x0) - h y <= xi) & (ALL y:H. xi <= p (y [+] x0) - h y)|]
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   202
   ==> ALL x:H0. h0 x <= p x"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   203
proof; 
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   204
  assume h0_def: "h0 = (%x. let (y, a) = @ (y, a). (x = y [+] a [*] x0 & y:H) in (h y) + a * xi)"
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   205
     and H0_def: "H0 = vectorspace_sum H (lin x0)" 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   206
     and vs:     "x0 ~: H" "x0 : E" "x0 ~= <0>" "is_vectorspace E" 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   207
                 "is_subspace H E" "is_quasinorm E p" "is_linearform H h" 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   208
     and a:      " ALL y:H. h y <= p y";
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   209
  presume a1: "(ALL y:H. - p (y [+] x0) - h y <= xi)";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   210
  presume a2: "(ALL y:H. xi <= p (y [+] x0) - h y)";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   211
  fix x; assume "x : H0"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   212
  show "h0 x <= p x"; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   213
  proof -; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   214
    have ex_x: "!! x. x : H0 ==> ? y a. (x = y [+] a [*] x0 & y : H)";
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   215
      by (simp add: H0_def vectorspace_sum_def lin_def, fast);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   216
    have "? y a. (x = y [+] a [*] x0 & y : H)";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   217
      by (rule ex_x);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   218
    thus ?thesis;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   219
    proof (elim exE conjE);
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   220
      fix y a; assume x: "x = y [+] a [*] x0" and y: "y : H";
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   221
      show ?thesis;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   222
      proof -;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   223
        have "h0 x = h y + a * xi";
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   224
          by (rule decomp3);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   225
        also; have "... <= p (y [+] a [*] x0)";
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   226
        proof (rule real_linear_split [of a "0r"]); (*** case analysis ***)
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   227
          assume lz: "a < 0r"; hence nz: "a ~= 0r"; by force;
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   228
          show ?thesis;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   229
          proof -;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   230
            from a1; have "- p (rinv a [*] y [+] x0) - h (rinv a [*] y) <= xi";
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   231
              by (rule bspec, simp!);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   232
            
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   233
            with lz; have "a * xi <= a * (- p (rinv a [*] y [+] x0) - h (rinv a [*] y))";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   234
              by (rule real_mult_less_le_anti);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   235
            also; have "... = - a * (p (rinv a [*] y [+] x0)) - a * (h (rinv a [*] y))";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   236
              by (rule real_mult_diff_distrib);
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   237
            also; from lz vs y; 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   238
            have "- a * (p (rinv a [*] y [+] x0)) = p (a [*] (rinv a [*] y [+] x0))";
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   239
              by (simp add: quasinorm_mult_distrib rabs_minus_eqI2 [RS sym]);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   240
            also; from nz vs y; have "... = p (y [+] a [*] x0)";
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   241
              by (simp add: vs_add_mult_distrib1);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   242
            also; from nz vs y; have "a * (h (rinv a [*] y)) =  h y";
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   243
              by (simp add: linearform_mult_linear [RS sym]);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   244
            finally; have "a * xi <= p (y [+] a [*] x0) - h y"; .;
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   245
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   246
            hence "h y + a * xi <= h y + (p (y [+] a [*] x0) - h y)";
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   247
              by (rule real_add_left_cancel_le [RS iffD2]);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   248
            thus ?thesis; 
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   249
              by simp;
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   250
	  qed;
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   251
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   252
        next;
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   253
          assume z: "a = 0r"; 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   254
          with vs y a; show ?thesis; by simp;
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   255
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   256
        next; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   257
          assume gz: "0r < a"; hence nz: "a ~= 0r"; by force;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   258
          show ?thesis;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   259
          proof -;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   260
            from a2; have "xi <= p (rinv a [*] y [+] x0) - h (rinv a [*] y)";
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   261
            by (rule bspec, simp!);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   262
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   263
            with gz; have "a * xi <= a * (p (rinv a [*] y [+] x0) - h (rinv a [*] y))";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   264
              by (rule real_mult_less_le_mono);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   265
            also; have "... = a * (p (rinv a [*] y [+] x0)) - a * (h (rinv a [*] y))";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   266
              by (rule real_mult_diff_distrib2); 
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   267
            also; from gz vs y; 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   268
            have "a * (p (rinv a [*] y [+] x0)) = p (a [*] (rinv a [*] y [+] x0))";
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   269
              by (simp add: quasinorm_mult_distrib rabs_eqI2);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   270
            also; from nz vs y; 
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   271
            have "... = p (y [+] a [*] x0)";
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   272
              by (simp add: vs_add_mult_distrib1);
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   273
            also; from nz vs y; have "a * (h (rinv a [*] y)) = h y";
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   274
              by (simp add: linearform_mult_linear [RS sym]); 
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   275
            finally; have "a * xi <= p (y [+] a [*] x0) - h y"; .;
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   276
 
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   277
            hence "h y + a * xi <= h y + (p (y [+] a [*] x0) - h y)";
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   278
              by (rule real_add_left_cancel_le [RS iffD2]);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   279
            thus ?thesis; 
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   280
              by simp;
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   281
          qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   282
        qed;
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   283
        also; from x; have "... = p x"; by (simp);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   284
        finally; show ?thesis; .;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   285
      qed; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   286
    qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   287
  qed; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   288
qed blast+;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   289
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   290
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   291
end;