src/HOL/Real/HahnBanach/HahnBanachExtLemmas.thy
changeset 9374 153853af318b
parent 9256 f9a6670427fb
child 9379 21cfeae6659d
equal deleted inserted replaced
9373:78a11a353473 9374:153853af318b
    33 \begin{matharray}{l} \All
    33 \begin{matharray}{l} \All
    34 {u\in F}{\All {v\in F}{a\ap u \leq b \ap v}} 
    34 {u\in F}{\All {v\in F}{a\ap u \leq b \ap v}} 
    35 \end{matharray} *};
    35 \end{matharray} *};
    36 
    36 
    37 lemma ex_xi: 
    37 lemma ex_xi: 
    38   "[| is_vectorspace F; !! u v. [| u:F; v:F |] ==> a u <= b v |]
    38   "[| is_vectorspace F; !! u v. [| u \<in> F; v \<in> F |] ==> a u <= b v |]
    39   ==> EX (xi::real). ALL y:F. a y <= xi & xi <= b y"; 
    39   ==> \<exists>xi::real. \<forall>y \<in> F. a y <= xi \<and> xi <= b y"; 
    40 proof -;
    40 proof -;
    41   assume vs: "is_vectorspace F";
    41   assume vs: "is_vectorspace F";
    42   assume r: "(!! u v. [| u:F; v:F |] ==> a u <= (b v::real))";
    42   assume r: "(!! u v. [| u \<in> F; v \<in> F |] ==> a u <= (b v::real))";
    43 
    43 
    44   txt {* From the completeness of the reals follows:
    44   txt {* From the completeness of the reals follows:
    45   The set $S = \{a\: u\dt\: u\in F\}$ has a supremum, if
    45   The set $S = \{a\: u\dt\: u\in F\}$ has a supremum, if
    46   it is non-empty and has an upper bound. *};
    46   it is non-empty and has an upper bound. *};
    47 
    47 
    48   let ?S = "{a u :: real | u. u:F}";
    48   let ?S = "{a u :: real | u. u \<in> F}";
    49 
    49 
    50   have "EX xi. isLub UNIV ?S xi";  
    50   have "\<exists>xi. isLub UNIV ?S xi";  
    51   proof (rule reals_complete);
    51   proof (rule reals_complete);
    52   
    52   
    53     txt {* The set $S$ is non-empty, since $a\ap\zero \in S$: *};
    53     txt {* The set $S$ is non-empty, since $a\ap\zero \in S$: *};
    54 
    54 
    55     from vs; have "a \<zero> : ?S"; by force;
    55     from vs; have "a 0 \<in> ?S"; by force;
    56     thus "EX X. X : ?S"; ..;
    56     thus "\<exists>X. X \<in> ?S"; ..;
    57 
    57 
    58     txt {* $b\ap \zero$ is an upper bound of $S$: *};
    58     txt {* $b\ap \zero$ is an upper bound of $S$: *};
    59 
    59 
    60     show "EX Y. isUb UNIV ?S Y"; 
    60     show "\<exists>Y. isUb UNIV ?S Y"; 
    61     proof; 
    61     proof; 
    62       show "isUb UNIV ?S (b \<zero>)";
    62       show "isUb UNIV ?S (b 0)";
    63       proof (intro isUbI setleI ballI);
    63       proof (intro isUbI setleI ballI);
    64         show "b \<zero> : UNIV"; ..;
    64         show "b 0 \<in> UNIV"; ..;
    65       next;
    65       next;
    66 
    66 
    67         txt {* Every element $y\in S$ is less than $b\ap \zero$: *};
    67         txt {* Every element $y\in S$ is less than $b\ap \zero$: *};
    68 
    68 
    69         fix y; assume y: "y : ?S"; 
    69         fix y; assume y: "y \<in> ?S"; 
    70         from y; have "EX u:F. y = a u"; by fast;
    70         from y; have "\<exists>u \<in> F. y = a u"; by fast;
    71         thus "y <= b \<zero>"; 
    71         thus "y <= b 0"; 
    72         proof;
    72         proof;
    73           fix u; assume "u:F"; 
    73           fix u; assume "u \<in> F"; 
    74           assume "y = a u";
    74           assume "y = a u";
    75           also; have "a u <= b \<zero>"; by (rule r) (simp!)+;
    75           also; have "a u <= b 0"; by (rule r) (simp!)+;
    76           finally; show ?thesis; .;
    76           finally; show ?thesis; .;
    77         qed;
    77         qed;
    78       qed;
    78       qed;
    79     qed;
    79     qed;
    80   qed;
    80   qed;
    81 
    81 
    82   thus "EX xi. ALL y:F. a y <= xi & xi <= b y"; 
    82   thus "\<exists>xi. \<forall>y \<in> F. a y <= xi \<and> xi <= b y"; 
    83   proof (elim exE);
    83   proof (elim exE);
    84     fix xi; assume "isLub UNIV ?S xi"; 
    84     fix xi; assume "isLub UNIV ?S xi"; 
    85     show ?thesis;
    85     show ?thesis;
    86     proof (intro exI conjI ballI); 
    86     proof (intro exI conjI ballI); 
    87    
    87    
    88       txt {* For all $y\in F$ holds $a\ap y \leq \xi$: *};
    88       txt {* For all $y\in F$ holds $a\ap y \leq \xi$: *};
    89      
    89      
    90       fix y; assume y: "y:F";
    90       fix y; assume y: "y \<in> F";
    91       show "a y <= xi";    
    91       show "a y <= xi";    
    92       proof (rule isUbD);  
    92       proof (rule isUbD);  
    93         show "isUb UNIV ?S xi"; ..;
    93         show "isUb UNIV ?S xi"; ..;
    94       qed (force!);
    94       qed (force!);
    95     next;
    95     next;
    96 
    96 
    97       txt {* For all $y\in F$ holds $\xi\leq b\ap y$: *};
    97       txt {* For all $y\in F$ holds $\xi\leq b\ap y$: *};
    98 
    98 
    99       fix y; assume "y:F";
    99       fix y; assume "y \<in> F";
   100       show "xi <= b y";  
   100       show "xi <= b y";  
   101       proof (intro isLub_le_isUb isUbI setleI);
   101       proof (intro isLub_le_isUb isUbI setleI);
   102         show "b y : UNIV"; ..;
   102         show "b y \<in> UNIV"; ..;
   103         show "ALL ya : ?S. ya <= b y"; 
   103         show "\<forall>ya \<in> ?S. ya <= b y"; 
   104         proof;
   104         proof;
   105           fix au; assume au: "au : ?S ";
   105           fix au; assume au: "au \<in> ?S ";
   106           hence "EX u:F. au = a u"; by fast;
   106           hence "\<exists>u \<in> F. au = a u"; by fast;
   107           thus "au <= b y";
   107           thus "au <= b y";
   108           proof;
   108           proof;
   109             fix u; assume "u:F"; assume "au = a u";  
   109             fix u; assume "u \<in> F"; assume "au = a u";  
   110             also; have "... <= b y"; by (rule r);
   110             also; have "... <= b y"; by (rule r);
   111             finally; show ?thesis; .;
   111             finally; show ?thesis; .;
   112           qed;
   112           qed;
   113         qed;
   113         qed;
   114       qed; 
   114       qed; 
   118 
   118 
   119 text{* \medskip The function $h_0$ is defined as a
   119 text{* \medskip The function $h_0$ is defined as a
   120 $h_0\ap x = h\ap y + a\cdot \xi$ where $x = y + a\mult \xi$
   120 $h_0\ap x = h\ap y + a\cdot \xi$ where $x = y + a\mult \xi$
   121 is a linear extension of $h$ to $H_0$. *};
   121 is a linear extension of $h$ to $H_0$. *};
   122 
   122 
   123 lemma h0_lf: 
   123 lemma h'_lf: 
   124   "[| h0 == (\\<lambda>x. let (y, a) = SOME (y, a). x = y + a \<prod> x0 & y:H 
   124   "[| h' == (\<lambda>x. let (y, a) = SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H 
   125                 in h y + a * xi);
   125                 in h y + a * xi);
   126   H0 == H + lin x0; is_subspace H E; is_linearform H h; x0 ~: H; 
   126   H' == H + lin x0; is_subspace H E; is_linearform H h; x0 \<notin> H; 
   127   x0 : E; x0 ~= \<zero>; is_vectorspace E |]
   127   x0 \<in> E; x0 \<noteq> 0; is_vectorspace E |]
   128   ==> is_linearform H0 h0";
   128   ==> is_linearform H' h'";
   129 proof -;
   129 proof -;
   130   assume h0_def: 
   130   assume h'_def: 
   131     "h0 == (\\<lambda>x. let (y, a) = SOME (y, a). x = y + a \<prod> x0 & y:H 
   131     "h' == (\<lambda>x. let (y, a) = SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H 
   132                in h y + a * xi)"
   132                in h y + a * xi)"
   133     and H0_def: "H0 == H + lin x0" 
   133     and H'_def: "H' == H + lin x0" 
   134     and vs: "is_subspace H E" "is_linearform H h" "x0 ~: H"
   134     and vs: "is_subspace H E" "is_linearform H h" "x0 \<notin> H"
   135       "x0 ~= \<zero>" "x0 : E" "is_vectorspace E";
   135       "x0 \<noteq> 0" "x0 \<in> E" "is_vectorspace E";
   136 
   136 
   137   have h0: "is_vectorspace H0"; 
   137   have h': "is_vectorspace H'"; 
   138   proof (unfold H0_def, rule vs_sum_vs);
   138   proof (unfold H'_def, rule vs_sum_vs);
   139     show "is_subspace (lin x0) E"; ..;
   139     show "is_subspace (lin x0) E"; ..;
   140   qed; 
   140   qed; 
   141 
   141 
   142   show ?thesis;
   142   show ?thesis;
   143   proof;
   143   proof;
   144     fix x1 x2; assume x1: "x1 : H0" and x2: "x2 : H0"; 
   144     fix x1 x2; assume x1: "x1 \<in> H'" and x2: "x2 \<in> H'"; 
   145 
   145 
   146     txt{* We now have to show that $h_0$ is additive, i.~e.\
   146     txt{* We now have to show that $h_0$ is additive, i.~e.\
   147     $h_0 \ap (x_1\plus x_2) = h_0\ap x_1 + h_0\ap x_2$
   147     $h_0 \ap (x_1\plus x_2) = h_0\ap x_1 + h_0\ap x_2$
   148     for $x_1, x_2\in H$. *}; 
   148     for $x_1, x_2\in H$. *}; 
   149 
   149 
   150     have x1x2: "x1 + x2 : H0"; 
   150     have x1x2: "x1 + x2 \<in> H'"; 
   151       by (rule vs_add_closed, rule h0); 
   151       by (rule vs_add_closed, rule h'); 
   152     from x1; 
   152     from x1; 
   153     have ex_x1: "EX y1 a1. x1 = y1 + a1 \<prod> x0  & y1 : H"; 
   153     have ex_x1: "\<exists> y1 a1. x1 = y1 + a1 \<cdot> x0  \<and> y1 \<in> H"; 
   154       by (unfold H0_def vs_sum_def lin_def) fast;
   154       by (unfold H'_def vs_sum_def lin_def) fast;
   155     from x2; 
   155     from x2; 
   156     have ex_x2: "EX y2 a2. x2 = y2 + a2 \<prod> x0 & y2 : H"; 
   156     have ex_x2: "\<exists> y2 a2. x2 = y2 + a2 \<cdot> x0 \<and> y2 \<in> H"; 
   157       by (unfold H0_def vs_sum_def lin_def) fast;
   157       by (unfold H'_def vs_sum_def lin_def) fast;
   158     from x1x2; 
   158     from x1x2; 
   159     have ex_x1x2: "EX y a. x1 + x2 = y + a \<prod> x0 & y : H";
   159     have ex_x1x2: "\<exists> y a. x1 + x2 = y + a \<cdot> x0 \<and> y \<in> H";
   160       by (unfold H0_def vs_sum_def lin_def) fast;
   160       by (unfold H'_def vs_sum_def lin_def) fast;
   161 
   161 
   162     from ex_x1 ex_x2 ex_x1x2;
   162     from ex_x1 ex_x2 ex_x1x2;
   163     show "h0 (x1 + x2) = h0 x1 + h0 x2";
   163     show "h' (x1 + x2) = h' x1 + h' x2";
   164     proof (elim exE conjE);
   164     proof (elim exE conjE);
   165       fix y1 y2 y a1 a2 a;
   165       fix y1 y2 y a1 a2 a;
   166       assume y1: "x1 = y1 + a1 \<prod> x0"     and y1': "y1 : H"
   166       assume y1: "x1 = y1 + a1 \<cdot> x0"     and y1': "y1 \<in> H"
   167          and y2: "x2 = y2 + a2 \<prod> x0"     and y2': "y2 : H" 
   167          and y2: "x2 = y2 + a2 \<cdot> x0"     and y2': "y2 \<in> H" 
   168          and y: "x1 + x2 = y + a \<prod> x0"   and y':  "y  : H"; 
   168          and y: "x1 + x2 = y + a \<cdot> x0"   and y':  "y  \<in> H"; 
   169 
   169 
   170       have ya: "y1 + y2 = y & a1 + a2 = a"; 
   170       have ya: "y1 + y2 = y \<and> a1 + a2 = a"; 
   171       proof (rule decomp_H0);;
   171       proof (rule decomp_H');;
   172 	txt_raw {* \label{decomp-H0-use} *};;
   172 	txt_raw {* \label{decomp-H-use} *};;
   173         show "y1 + y2 + (a1 + a2) \<prod> x0 = y + a \<prod> x0"; 
   173         show "y1 + y2 + (a1 + a2) \<cdot> x0 = y + a \<cdot> x0"; 
   174           by (simp! add: vs_add_mult_distrib2 [of E]);
   174           by (simp! add: vs_add_mult_distrib2 [of E]);
   175         show "y1 + y2 : H"; ..;
   175         show "y1 + y2 \<in> H"; ..;
   176       qed;
   176       qed;
   177 
   177 
   178       have "h0 (x1 + x2) = h y + a * xi";
   178       have "h' (x1 + x2) = h y + a * xi";
   179 	by (rule h0_definite);
   179 	by (rule h'_definite);
   180       also; have "... = h (y1 + y2) + (a1 + a2) * xi"; 
   180       also; have "... = h (y1 + y2) + (a1 + a2) * xi"; 
   181         by (simp add: ya);
   181         by (simp add: ya);
   182       also; from vs y1' y2'; 
   182       also; from vs y1' y2'; 
   183       have "... = h y1 + h y2 + a1 * xi + a2 * xi"; 
   183       have "... = h y1 + h y2 + a1 * xi + a2 * xi"; 
   184 	by (simp add: linearform_add [of H] 
   184 	by (simp add: linearform_add [of H] 
   185                       real_add_mult_distrib);
   185                       real_add_mult_distrib);
   186       also; have "... = (h y1 + a1 * xi) + (h y2 + a2 * xi)"; 
   186       also; have "... = (h y1 + a1 * xi) + (h y2 + a2 * xi)"; 
   187         by simp;
   187         by simp;
   188       also; have "h y1 + a1 * xi = h0 x1"; 
   188       also; have "h y1 + a1 * xi = h' x1"; 
   189         by (rule h0_definite [RS sym]);
   189         by (rule h'_definite [RS sym]);
   190       also; have "h y2 + a2 * xi = h0 x2"; 
   190       also; have "h y2 + a2 * xi = h' x2"; 
   191         by (rule h0_definite [RS sym]);
   191         by (rule h'_definite [RS sym]);
   192       finally; show ?thesis; .;
   192       finally; show ?thesis; .;
   193     qed;
   193     qed;
   194  
   194  
   195     txt{* We further have to show that $h_0$ is multiplicative, 
   195     txt{* We further have to show that $h_0$ is multiplicative, 
   196     i.~e.\ $h_0\ap (c \mult x_1) = c \cdot h_0\ap x_1$
   196     i.~e.\ $h_0\ap (c \mult x_1) = c \cdot h_0\ap x_1$
   197     for $x\in H$ and $c\in \bbbR$. 
   197     for $x\in H$ and $c\in \bbbR$. 
   198     *}; 
   198     *}; 
   199 
   199 
   200   next;  
   200   next;  
   201     fix c x1; assume x1: "x1 : H0";    
   201     fix c x1; assume x1: "x1 \<in> H'";    
   202     have ax1: "c \<prod> x1 : H0";
   202     have ax1: "c \<cdot> x1 \<in> H'";
   203       by (rule vs_mult_closed, rule h0);
   203       by (rule vs_mult_closed, rule h');
   204     from x1; have ex_x: "!! x. x: H0 
   204     from x1; have ex_x: "!! x. x\<in> H' 
   205                         ==> EX y a. x = y + a \<prod> x0 & y : H";
   205                         ==> \<exists> y a. x = y + a \<cdot> x0 \<and> y \<in> H";
   206       by (unfold H0_def vs_sum_def lin_def) fast;
   206       by (unfold H'_def vs_sum_def lin_def) fast;
   207 
   207 
   208     from x1; have ex_x1: "EX y1 a1. x1 = y1 + a1 \<prod> x0 & y1 : H";
   208     from x1; have ex_x1: "\<exists> y1 a1. x1 = y1 + a1 \<cdot> x0 \<and> y1 \<in> H";
   209       by (unfold H0_def vs_sum_def lin_def) fast;
   209       by (unfold H'_def vs_sum_def lin_def) fast;
   210     with ex_x [of "c \<prod> x1", OF ax1];
   210     with ex_x [of "c \<cdot> x1", OF ax1];
   211     show "h0 (c \<prod> x1) = c * (h0 x1)";  
   211     show "h' (c \<cdot> x1) = c * (h' x1)";  
   212     proof (elim exE conjE);
   212     proof (elim exE conjE);
   213       fix y1 y a1 a; 
   213       fix y1 y a1 a; 
   214       assume y1: "x1 = y1 + a1 \<prod> x0"       and y1': "y1 : H"
   214       assume y1: "x1 = y1 + a1 \<cdot> x0"     and y1': "y1 \<in> H"
   215         and y: "c \<prod> x1 = y  + a  \<prod> x0"   and y':  "y  : H"; 
   215         and y: "c \<cdot> x1 = y  + a \<cdot> x0"    and y': "y \<in> H"; 
   216 
   216 
   217       have ya: "c \<prod> y1 = y & c * a1 = a"; 
   217       have ya: "c \<cdot> y1 = y \<and> c * a1 = a"; 
   218       proof (rule decomp_H0); 
   218       proof (rule decomp_H'); 
   219 	show "c \<prod> y1 + (c * a1) \<prod> x0 = y + a \<prod> x0"; 
   219 	show "c \<cdot> y1 + (c * a1) \<cdot> x0 = y + a \<cdot> x0"; 
   220           by (simp! add: add: vs_add_mult_distrib1);
   220           by (simp! add: vs_add_mult_distrib1);
   221         show "c \<prod> y1 : H"; ..;
   221         show "c \<cdot> y1 \<in> H"; ..;
   222       qed;
   222       qed;
   223 
   223 
   224       have "h0 (c \<prod> x1) = h y + a * xi"; 
   224       have "h' (c \<cdot> x1) = h y + a * xi"; 
   225 	by (rule h0_definite);
   225 	by (rule h'_definite);
   226       also; have "... = h (c \<prod> y1) + (c * a1) * xi";
   226       also; have "... = h (c \<cdot> y1) + (c * a1) * xi";
   227         by (simp add: ya);
   227         by (simp add: ya);
   228       also; from vs y1'; have "... = c * h y1 + c * a1 * xi"; 
   228       also; from vs y1'; have "... = c * h y1 + c * a1 * xi"; 
   229 	by (simp add: linearform_mult [of H]);
   229 	by (simp add: linearform_mult [of H]);
   230       also; from vs y1'; have "... = c * (h y1 + a1 * xi)"; 
   230       also; from vs y1'; have "... = c * (h y1 + a1 * xi)"; 
   231 	by (simp add: real_add_mult_distrib2 real_mult_assoc);
   231 	by (simp add: real_add_mult_distrib2 real_mult_assoc);
   232       also; have "h y1 + a1 * xi = h0 x1"; 
   232       also; have "h y1 + a1 * xi = h' x1"; 
   233         by (rule h0_definite [RS sym]);
   233         by (rule h'_definite [RS sym]);
   234       finally; show ?thesis; .;
   234       finally; show ?thesis; .;
   235     qed;
   235     qed;
   236   qed;
   236   qed;
   237 qed;
   237 qed;
   238 
   238 
   239 text{* \medskip The linear extension $h_0$ of $h$
   239 text{* \medskip The linear extension $h_0$ of $h$
   240 is bounded by the seminorm $p$. *};
   240 is bounded by the seminorm $p$. *};
   241 
   241 
   242 lemma h0_norm_pres:
   242 lemma h'_norm_pres:
   243   "[| h0 == (\\<lambda>x. let (y, a) = SOME (y, a). x = y + a \<prod> x0 & y:H 
   243   "[| h' == (\<lambda>x. let (y, a) = SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H 
   244                 in h y + a * xi);
   244                 in h y + a * xi);
   245   H0 == H + lin x0; x0 ~: H; x0 : E; x0 ~= \<zero>; is_vectorspace E; 
   245   H' == H + lin x0; x0 \<notin> H; x0 \<in> E; x0 \<noteq> 0; is_vectorspace E; 
   246   is_subspace H E; is_seminorm E p; is_linearform H h; ALL y:H. h y <= p y; 
   246   is_subspace H E; is_seminorm E p; is_linearform H h; \<forall>y \<in> H. h y <= p y; 
   247   ALL y:H. - p (y + x0) - h y <= xi & xi <= p (y + x0) - h y |]
   247   \<forall>y \<in> H. - p (y + x0) - h y <= xi \<and> xi <= p (y + x0) - h y |]
   248    ==> ALL x:H0. h0 x <= p x"; 
   248    ==> \<forall> x \<in> H'. h' x <= p x"; 
   249 proof; 
   249 proof; 
   250   assume h0_def: 
   250   assume h'_def: 
   251     "h0 == (\\<lambda>x. let (y, a) = SOME (y, a). x = y + a \<prod> x0 & y:H 
   251     "h' == (\<lambda>x. let (y, a) = SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H 
   252                in (h y) + a * xi)"
   252                in (h y) + a * xi)"
   253     and H0_def: "H0 == H + lin x0" 
   253     and H'_def: "H' == H + lin x0" 
   254     and vs: "x0 ~: H" "x0 : E" "x0 ~= \<zero>" "is_vectorspace E" 
   254     and vs: "x0 \<notin> H" "x0 \<in> E" "x0 \<noteq> 0" "is_vectorspace E" 
   255             "is_subspace H E" "is_seminorm E p" "is_linearform H h" 
   255             "is_subspace H E" "is_seminorm E p" "is_linearform H h" 
   256     and a: "ALL y:H. h y <= p y";
   256     and a: "\<forall>y \<in> H. h y <= p y";
   257   presume a1: "ALL ya:H. - p (ya + x0) - h ya <= xi";
   257   presume a1: "\<forall>ya \<in> H. - p (ya + x0) - h ya <= xi";
   258   presume a2: "ALL ya:H. xi <= p (ya + x0) - h ya";
   258   presume a2: "\<forall>ya \<in> H. xi <= p (ya + x0) - h ya";
   259   fix x; assume "x : H0"; 
   259   fix x; assume "x \<in> H'"; 
   260   have ex_x: 
   260   have ex_x: 
   261     "!! x. x : H0 ==> EX y a. x = y + a \<prod> x0 & y : H";
   261     "!! x. x \<in> H' ==> \<exists>y a. x = y + a \<cdot> x0 \<and> y \<in> H";
   262     by (unfold H0_def vs_sum_def lin_def) fast;
   262     by (unfold H'_def vs_sum_def lin_def) fast;
   263   have "EX y a. x = y + a \<prod> x0 & y : H";
   263   have "\<exists>y a. x = y + a \<cdot> x0 \<and> y \<in> H";
   264     by (rule ex_x);
   264     by (rule ex_x);
   265   thus "h0 x <= p x";
   265   thus "h' x <= p x";
   266   proof (elim exE conjE);
   266   proof (elim exE conjE);
   267     fix y a; assume x: "x = y + a \<prod> x0" and y: "y : H";
   267     fix y a; assume x: "x = y + a \<cdot> x0" and y: "y \<in> H";
   268     have "h0 x = h y + a * xi";
   268     have "h' x = h y + a * xi";
   269       by (rule h0_definite);
   269       by (rule h'_definite);
   270 
   270 
   271     txt{* Now we show  
   271     txt{* Now we show  
   272     $h\ap y + a \cdot \xi\leq  p\ap (y\plus a \mult x_0)$ 
   272     $h\ap y + a \cdot \xi\leq  p\ap (y\plus a \mult x_0)$ 
   273     by case analysis on $a$. \label{linorder_linear_split}*};
   273     by case analysis on $a$. \label{linorder_linear_split}*};
   274 
   274 
   275     also; have "... <= p (y + a \<prod> x0)";
   275     also; have "... <= p (y + a \<cdot> x0)";
   276     proof (rule linorder_linear_split); 
   276     proof (rule linorder_linear_split); 
   277 
   277 
   278       assume z: "a = #0"; 
   278       assume z: "a = #0"; 
   279       with vs y a; show ?thesis; by simp;
   279       with vs y a; show ?thesis; by simp;
   280 
   280 
   281     txt {* In the case $a < 0$, we use $a_1$ with $\idt{ya}$ 
   281     txt {* In the case $a < 0$, we use $a_1$ with $\idt{ya}$ 
   282     taken as $y/a$: *};
   282     taken as $y/a$: *};
   283 
   283 
   284     next;
   284     next;
   285       assume lz: "a < #0"; hence nz: "a ~= #0"; by simp;
   285       assume lz: "a < #0"; hence nz: "a \<noteq> #0"; by simp;
   286       from a1; 
   286       from a1; 
   287       have "- p (rinv a \<prod> y + x0) - h (rinv a \<prod> y) <= xi";
   287       have "- p (rinv a \<cdot> y + x0) - h (rinv a \<cdot> y) <= xi";
   288         by (rule bspec) (simp!);
   288         by (rule bspec) (simp!);
   289 
   289 
   290       txt {* The thesis for this case now follows by a short  
   290       txt {* The thesis for this case now follows by a short  
   291       calculation. *};      
   291       calculation. *};      
   292       hence "a * xi 
   292       hence "a * xi 
   293             <= a * (- p (rinv a \<prod> y + x0) - h (rinv a \<prod> y))";
   293             <= a * (- p (rinv a \<cdot> y + x0) - h (rinv a \<cdot> y))";
   294         by (rule real_mult_less_le_anti [OF lz]);
   294         by (rule real_mult_less_le_anti [OF lz]);
   295       also; have "... = - a * (p (rinv a \<prod> y + x0)) 
   295       also; have "... = - a * (p (rinv a \<cdot> y + x0)) 
   296                         - a * (h (rinv a \<prod> y))";
   296                         - a * (h (rinv a \<cdot> y))";
   297         by (rule real_mult_diff_distrib);
   297         by (rule real_mult_diff_distrib);
   298       also; from lz vs y; have "- a * (p (rinv a \<prod> y + x0)) 
   298       also; from lz vs y; have "- a * (p (rinv a \<cdot> y + x0)) 
   299                                = p (a \<prod> (rinv a \<prod> y + x0))";
   299                                = p (a \<cdot> (rinv a \<cdot> y + x0))";
   300         by (simp add: seminorm_abs_homogenous abs_minus_eqI2);
   300         by (simp add: seminorm_abs_homogenous abs_minus_eqI2);
   301       also; from nz vs y; have "... = p (y + a \<prod> x0)";
   301       also; from nz vs y; have "... = p (y + a \<cdot> x0)";
   302         by (simp add: vs_add_mult_distrib1);
   302         by (simp add: vs_add_mult_distrib1);
   303       also; from nz vs y; have "a * (h (rinv a \<prod> y)) =  h y";
   303       also; from nz vs y; have "a * (h (rinv a \<cdot> y)) =  h y";
   304         by (simp add: linearform_mult [RS sym]);
   304         by (simp add: linearform_mult [RS sym]);
   305       finally; have "a * xi <= p (y + a \<prod> x0) - h y"; .;
   305       finally; have "a * xi <= p (y + a \<cdot> x0) - h y"; .;
   306 
   306 
   307       hence "h y + a * xi <= h y + p (y + a \<prod> x0) - h y";
   307       hence "h y + a * xi <= h y + p (y + a \<cdot> x0) - h y";
   308         by (simp add: real_add_left_cancel_le);
   308         by (simp add: real_add_left_cancel_le);
   309       thus ?thesis; by simp;
   309       thus ?thesis; by simp;
   310 
   310 
   311       txt {* In the case $a > 0$, we use $a_2$ with $\idt{ya}$ 
   311       txt {* In the case $a > 0$, we use $a_2$ with $\idt{ya}$ 
   312       taken as $y/a$: *};
   312       taken as $y/a$: *};
   313 
   313 
   314     next; 
   314     next; 
   315       assume gz: "#0 < a"; hence nz: "a ~= #0"; by simp;
   315       assume gz: "#0 < a"; hence nz: "a \<noteq> #0"; by simp;
   316       from a2;
   316       from a2;
   317       have "xi <= p (rinv a \<prod> y + x0) - h (rinv a \<prod> y)";
   317       have "xi <= p (rinv a \<cdot> y + x0) - h (rinv a \<cdot> y)";
   318         by (rule bspec) (simp!);
   318         by (rule bspec) (simp!);
   319 
   319 
   320       txt {* The thesis for this case follows by a short
   320       txt {* The thesis for this case follows by a short
   321       calculation: *};
   321       calculation: *};
   322 
   322 
   323       with gz; have "a * xi 
   323       with gz; have "a * xi 
   324             <= a * (p (rinv a \<prod> y + x0) - h (rinv a \<prod> y))";
   324             <= a * (p (rinv a \<cdot> y + x0) - h (rinv a \<cdot> y))";
   325         by (rule real_mult_less_le_mono);
   325         by (rule real_mult_less_le_mono);
   326       also; have "... = a * p (rinv a \<prod> y + x0) 
   326       also; have "... = a * p (rinv a \<cdot> y + x0) 
   327                         - a * h (rinv a \<prod> y)";
   327                         - a * h (rinv a \<cdot> y)";
   328         by (rule real_mult_diff_distrib2); 
   328         by (rule real_mult_diff_distrib2); 
   329       also; from gz vs y; 
   329       also; from gz vs y; 
   330       have "a * p (rinv a \<prod> y + x0) 
   330       have "a * p (rinv a \<cdot> y + x0) 
   331            = p (a \<prod> (rinv a \<prod> y + x0))";
   331            = p (a \<cdot> (rinv a \<cdot> y + x0))";
   332         by (simp add: seminorm_abs_homogenous abs_eqI2);
   332         by (simp add: seminorm_abs_homogenous abs_eqI2);
   333       also; from nz vs y; 
   333       also; from nz vs y; 
   334       have "... = p (y + a \<prod> x0)";
   334       have "... = p (y + a \<cdot> x0)";
   335         by (simp add: vs_add_mult_distrib1);
   335         by (simp add: vs_add_mult_distrib1);
   336       also; from nz vs y; have "a * h (rinv a \<prod> y) = h y";
   336       also; from nz vs y; have "a * h (rinv a \<cdot> y) = h y";
   337         by (simp add: linearform_mult [RS sym]); 
   337         by (simp add: linearform_mult [RS sym]); 
   338       finally; have "a * xi <= p (y + a \<prod> x0) - h y"; .;
   338       finally; have "a * xi <= p (y + a \<cdot> x0) - h y"; .;
   339  
   339  
   340       hence "h y + a * xi <= h y + (p (y + a \<prod> x0) - h y)";
   340       hence "h y + a * xi <= h y + (p (y + a \<cdot> x0) - h y)";
   341         by (simp add: real_add_left_cancel_le);
   341         by (simp add: real_add_left_cancel_le);
   342       thus ?thesis; by simp;
   342       thus ?thesis; by simp;
   343     qed;
   343     qed;
   344     also; from x; have "... = p x"; by simp;
   344     also; from x; have "... = p x"; by simp;
   345     finally; show ?thesis; .;
   345     finally; show ?thesis; .;