adapted HOL source structure to distribution layout
authorhaftmann
Mon, 29 Dec 2008 14:08:08 +0100
changeset 291976d4cb27ed19c
parent 29189 ee8572f3bb57
child 29198 418ed6411847
adapted HOL source structure to distribution layout
NEWS
src/HOL/Complex/Fundamental_Theorem_Algebra.thy
src/HOL/Complex/README.html
src/HOL/Complex/document/root.tex
src/HOL/Complex_Main.thy
src/HOL/Dense_Linear_Order.thy
src/HOL/Fundamental_Theorem_Algebra.thy
src/HOL/HahnBanach/Bounds.thy
src/HOL/HahnBanach/FunctionNorm.thy
src/HOL/HahnBanach/FunctionOrder.thy
src/HOL/HahnBanach/HahnBanach.thy
src/HOL/HahnBanach/HahnBanachExtLemmas.thy
src/HOL/HahnBanach/HahnBanachLemmas.thy
src/HOL/HahnBanach/HahnBanachSupLemmas.thy
src/HOL/HahnBanach/Linearform.thy
src/HOL/HahnBanach/NormedSpace.thy
src/HOL/HahnBanach/README.html
src/HOL/HahnBanach/ROOT.ML
src/HOL/HahnBanach/Subspace.thy
src/HOL/HahnBanach/VectorSpace.thy
src/HOL/HahnBanach/ZornLemma.thy
src/HOL/HahnBanach/document/root.bib
src/HOL/HahnBanach/document/root.tex
src/HOL/Hyperreal/SEQ.thy
src/HOL/IsaMakefile
src/HOL/Library/Dense_Linear_Order.thy
src/HOL/Library/Library.thy
src/HOL/Lim.thy
src/HOL/PReal.thy
src/HOL/Real.thy
src/HOL/Real/HahnBanach/Bounds.thy
src/HOL/Real/HahnBanach/FunctionNorm.thy
src/HOL/Real/HahnBanach/FunctionOrder.thy
src/HOL/Real/HahnBanach/HahnBanach.thy
src/HOL/Real/HahnBanach/HahnBanachExtLemmas.thy
src/HOL/Real/HahnBanach/HahnBanachLemmas.thy
src/HOL/Real/HahnBanach/HahnBanachSupLemmas.thy
src/HOL/Real/HahnBanach/Linearform.thy
src/HOL/Real/HahnBanach/NormedSpace.thy
src/HOL/Real/HahnBanach/README.html
src/HOL/Real/HahnBanach/ROOT.ML
src/HOL/Real/HahnBanach/Subspace.thy
src/HOL/Real/HahnBanach/VectorSpace.thy
src/HOL/Real/HahnBanach/ZornLemma.thy
src/HOL/Real/HahnBanach/document/root.bib
src/HOL/Real/HahnBanach/document/root.tex
src/HOL/Real/RealVector.thy
src/HOL/RealVector.thy
src/HOL/SEQ.thy
src/HOL/Series.thy
     1.1 --- a/NEWS	Mon Dec 29 13:23:53 2008 +0100
     1.2 +++ b/NEWS	Mon Dec 29 14:08:08 2008 +0100
     1.3 @@ -157,11 +157,12 @@
     1.4  
     1.5  *** HOL ***
     1.6  
     1.7 -* Made repository layout more coherent with logical
     1.8 -distribution structure:
     1.9 +* Made source layout more coherent with logical distribution
    1.10 +structure:
    1.11  
    1.12      src/HOL/Library/RType.thy ~> src/HOL/Typerep.thy
    1.13      src/HOL/Library/Code_Message.thy ~> src/HOL/
    1.14 +    src/HOL/Library/Dense_Linear_Order.thy ~> src/HOL/
    1.15      src/HOL/Library/GCD.thy ~> src/HOL/
    1.16      src/HOL/Library/Order_Relation.thy ~> src/HOL/
    1.17      src/HOL/Library/Parity.thy ~> src/HOL/
    1.18 @@ -177,6 +178,7 @@
    1.19      src/HOL/Complex/Complex_Main.thy ~> src/HOL/
    1.20      src/HOL/Complex/Complex.thy ~> src/HOL/
    1.21      src/HOL/Complex/FrechetDeriv.thy ~> src/HOL/
    1.22 +    src/HOL/Complex/Fundamental_Theorem_Algebra.thy ~> src/HOL/
    1.23      src/HOL/Hyperreal/Deriv.thy ~> src/HOL/
    1.24      src/HOL/Hyperreal/Fact.thy ~> src/HOL/
    1.25      src/HOL/Hyperreal/Integration.thy ~> src/HOL/
    1.26 @@ -186,9 +188,12 @@
    1.27      src/HOL/Hyperreal/MacLaurin.thy ~> src/HOL/
    1.28      src/HOL/Hyperreal/NthRoot.thy ~> src/HOL/
    1.29      src/HOL/Hyperreal/Series.thy ~> src/HOL/
    1.30 +    src/HOL/Hyperreal/SEQ.thy ~> src/HOL/
    1.31      src/HOL/Hyperreal/Taylor.thy ~> src/HOL/
    1.32      src/HOL/Hyperreal/Transcendental.thy ~> src/HOL/
    1.33      src/HOL/Real/Float ~> src/HOL/Library/
    1.34 +    src/HOL/Real/HahnBanach ~> src/HOL/HahnBanach
    1.35 +    src/HOL/Real/RealVector.thy ~> src/HOL/
    1.36  
    1.37      src/HOL/arith_data.ML ~> src/HOL/Tools
    1.38      src/HOL/hologic.ML ~> src/HOL/Tools
     2.1 --- a/src/HOL/Complex/Fundamental_Theorem_Algebra.thy	Mon Dec 29 13:23:53 2008 +0100
     2.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
     2.3 @@ -1,1329 +0,0 @@
     2.4 -(*  Title:       Fundamental_Theorem_Algebra.thy
     2.5 -    Author:      Amine Chaieb
     2.6 -*)
     2.7 -
     2.8 -header{*Fundamental Theorem of Algebra*}
     2.9 -
    2.10 -theory Fundamental_Theorem_Algebra
    2.11 -imports "~~/src/HOL/Univ_Poly" "~~/src/HOL/Library/Dense_Linear_Order" "~~/src/HOL/Complex"
    2.12 -begin
    2.13 -
    2.14 -subsection {* Square root of complex numbers *}
    2.15 -definition csqrt :: "complex \<Rightarrow> complex" where
    2.16 -"csqrt z = (if Im z = 0 then
    2.17 -            if 0 \<le> Re z then Complex (sqrt(Re z)) 0
    2.18 -            else Complex 0 (sqrt(- Re z))
    2.19 -           else Complex (sqrt((cmod z + Re z) /2))
    2.20 -                        ((Im z / abs(Im z)) * sqrt((cmod z - Re z) /2)))"
    2.21 -
    2.22 -lemma csqrt[algebra]: "csqrt z ^ 2 = z"
    2.23 -proof-
    2.24 -  obtain x y where xy: "z = Complex x y" by (cases z, simp_all)
    2.25 -  {assume y0: "y = 0"
    2.26 -    {assume x0: "x \<ge> 0" 
    2.27 -      then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
    2.28 -	by (simp add: csqrt_def power2_eq_square)}
    2.29 -    moreover
    2.30 -    {assume "\<not> x \<ge> 0" hence x0: "- x \<ge> 0" by arith
    2.31 -      then have ?thesis using y0 xy real_sqrt_pow2[OF x0] 
    2.32 -	by (simp add: csqrt_def power2_eq_square) }
    2.33 -    ultimately have ?thesis by blast}
    2.34 -  moreover
    2.35 -  {assume y0: "y\<noteq>0"
    2.36 -    {fix x y
    2.37 -      let ?z = "Complex x y"
    2.38 -      from abs_Re_le_cmod[of ?z] have tha: "abs x \<le> cmod ?z" by auto
    2.39 -      hence "cmod ?z - x \<ge> 0" "cmod ?z + x \<ge> 0" by arith+ 
    2.40 -      hence "(sqrt (x * x + y * y) + x) / 2 \<ge> 0" "(sqrt (x * x + y * y) - x) / 2 \<ge> 0" by (simp_all add: power2_eq_square) }
    2.41 -    note th = this
    2.42 -    have sq4: "\<And>x::real. x^2 / 4 = (x / 2) ^ 2" 
    2.43 -      by (simp add: power2_eq_square) 
    2.44 -    from th[of x y]
    2.45 -    have sq4': "sqrt (((sqrt (x * x + y * y) + x)^2 / 4)) = (sqrt (x * x + y * y) + x) / 2" "sqrt (((sqrt (x * x + y * y) - x)^2 / 4)) = (sqrt (x * x + y * y) - x) / 2" unfolding sq4 by simp_all
    2.46 -    then have th1: "sqrt ((sqrt (x * x + y * y) + x) * (sqrt (x * x + y * y) + x) / 4) - sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) - x) / 4) = x"
    2.47 -      unfolding power2_eq_square by simp 
    2.48 -    have "sqrt 4 = sqrt (2^2)" by simp 
    2.49 -    hence sqrt4: "sqrt 4 = 2" by (simp only: real_sqrt_abs)
    2.50 -    have th2: "2 *(y * sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) + x) / 4)) / \<bar>y\<bar> = y"
    2.51 -      using iffD2[OF real_sqrt_pow2_iff sum_power2_ge_zero[of x y]] y0
    2.52 -      unfolding power2_eq_square 
    2.53 -      by (simp add: ring_simps real_sqrt_divide sqrt4)
    2.54 -     from y0 xy have ?thesis  apply (simp add: csqrt_def power2_eq_square)
    2.55 -       apply (simp add: real_sqrt_sum_squares_mult_ge_zero[of x y] real_sqrt_pow2[OF th(1)[of x y], unfolded power2_eq_square] real_sqrt_pow2[OF th(2)[of x y], unfolded power2_eq_square] real_sqrt_mult[symmetric])
    2.56 -      using th1 th2  ..}
    2.57 -  ultimately show ?thesis by blast
    2.58 -qed
    2.59 -
    2.60 -
    2.61 -subsection{* More lemmas about module of complex numbers *}
    2.62 -
    2.63 -lemma complex_of_real_power: "complex_of_real x ^ n = complex_of_real (x^n)"
    2.64 -  by (rule of_real_power [symmetric])
    2.65 -
    2.66 -lemma real_down2: "(0::real) < d1 \<Longrightarrow> 0 < d2 ==> EX e. 0 < e & e < d1 & e < d2"
    2.67 -  apply ferrack apply arith done
    2.68 -
    2.69 -text{* The triangle inequality for cmod *}
    2.70 -lemma complex_mod_triangle_sub: "cmod w \<le> cmod (w + z) + norm z"
    2.71 -  using complex_mod_triangle_ineq2[of "w + z" "-z"] by auto
    2.72 -
    2.73 -subsection{* Basic lemmas about complex polynomials *}
    2.74 -
    2.75 -lemma poly_bound_exists:
    2.76 -  shows "\<exists>m. m > 0 \<and> (\<forall>z. cmod z <= r \<longrightarrow> cmod (poly p z) \<le> m)"
    2.77 -proof(induct p)
    2.78 -  case Nil thus ?case by (rule exI[where x=1], simp) 
    2.79 -next
    2.80 -  case (Cons c cs)
    2.81 -  from Cons.hyps obtain m where m: "\<forall>z. cmod z \<le> r \<longrightarrow> cmod (poly cs z) \<le> m"
    2.82 -    by blast
    2.83 -  let ?k = " 1 + cmod c + \<bar>r * m\<bar>"
    2.84 -  have kp: "?k > 0" using abs_ge_zero[of "r*m"] norm_ge_zero[of c] by arith
    2.85 -  {fix z
    2.86 -    assume H: "cmod z \<le> r"
    2.87 -    from m H have th: "cmod (poly cs z) \<le> m" by blast
    2.88 -    from H have rp: "r \<ge> 0" using norm_ge_zero[of z] by arith
    2.89 -    have "cmod (poly (c # cs) z) \<le> cmod c + cmod (z* poly cs z)"
    2.90 -      using norm_triangle_ineq[of c "z* poly cs z"] by simp
    2.91 -    also have "\<dots> \<le> cmod c + r*m" using mult_mono[OF H th rp norm_ge_zero[of "poly cs z"]] by (simp add: norm_mult)
    2.92 -    also have "\<dots> \<le> ?k" by simp
    2.93 -    finally have "cmod (poly (c # cs) z) \<le> ?k" .}
    2.94 -  with kp show ?case by blast
    2.95 -qed
    2.96 -
    2.97 -
    2.98 -text{* Offsetting the variable in a polynomial gives another of same degree *}
    2.99 -  (* FIXME : Lemma holds also in locale --- fix it later *)
   2.100 -lemma  poly_offset_lemma:
   2.101 -  shows "\<exists>b q. (length q = length p) \<and> (\<forall>x. poly (b#q) (x::complex) = (a + x) * poly p x)"
   2.102 -proof(induct p)
   2.103 -  case Nil thus ?case by simp
   2.104 -next
   2.105 -  case (Cons c cs)
   2.106 -  from Cons.hyps obtain b q where 
   2.107 -    bq: "length q = length cs" "\<forall>x. poly (b # q) x = (a + x) * poly cs x"
   2.108 -    by blast
   2.109 -  let ?b = "a*c"
   2.110 -  let ?q = "(b+c)#q"
   2.111 -  have lg: "length ?q = length (c#cs)" using bq(1) by simp
   2.112 -  {fix x
   2.113 -    from bq(2)[rule_format, of x]
   2.114 -    have "x*poly (b # q) x = x*((a + x) * poly cs x)" by simp
   2.115 -    hence "poly (?b# ?q) x = (a + x) * poly (c # cs) x"
   2.116 -      by (simp add: ring_simps)}
   2.117 -  with lg  show ?case by blast 
   2.118 -qed
   2.119 -
   2.120 -    (* FIXME : This one too*)
   2.121 -lemma poly_offset: "\<exists> q. length q = length p \<and> (\<forall>x. poly q (x::complex) = poly p (a + x))"
   2.122 -proof (induct p)
   2.123 -  case Nil thus ?case by simp
   2.124 -next
   2.125 -  case (Cons c cs)
   2.126 -  from Cons.hyps obtain q where q: "length q = length cs" "\<forall>x. poly q x = poly cs (a + x)" by blast
   2.127 -  from poly_offset_lemma[of q a] obtain b p where 
   2.128 -    bp: "length p = length q" "\<forall>x. poly (b # p) x = (a + x) * poly q x"
   2.129 -    by blast
   2.130 -  thus ?case using q bp by - (rule exI[where x="(c + b)#p"], simp)
   2.131 -qed
   2.132 -
   2.133 -text{* An alternative useful formulation of completeness of the reals *}
   2.134 -lemma real_sup_exists: assumes ex: "\<exists>x. P x" and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z"
   2.135 -  shows "\<exists>(s::real). \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s"
   2.136 -proof-
   2.137 -  from ex bz obtain x Y where x: "P x" and Y: "\<And>x. P x \<Longrightarrow> x < Y"  by blast
   2.138 -  from ex have thx:"\<exists>x. x \<in> Collect P" by blast
   2.139 -  from bz have thY: "\<exists>Y. isUb UNIV (Collect P) Y" 
   2.140 -    by(auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def order_le_less)
   2.141 -  from reals_complete[OF thx thY] obtain L where L: "isLub UNIV (Collect P) L"
   2.142 -    by blast
   2.143 -  from Y[OF x] have xY: "x < Y" .
   2.144 -  from L have L': "\<forall>x. P x \<longrightarrow> x \<le> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)  
   2.145 -  from Y have Y': "\<forall>x. P x \<longrightarrow> x \<le> Y" 
   2.146 -    apply (clarsimp, atomize (full)) by auto 
   2.147 -  from L Y' have "L \<le> Y" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
   2.148 -  {fix y
   2.149 -    {fix z assume z: "P z" "y < z"
   2.150 -      from L' z have "y < L" by auto }
   2.151 -    moreover
   2.152 -    {assume yL: "y < L" "\<forall>z. P z \<longrightarrow> \<not> y < z"
   2.153 -      hence nox: "\<forall>z. P z \<longrightarrow> y \<ge> z" by auto
   2.154 -      from nox L have "y \<ge> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def) 
   2.155 -      with yL(1) have False  by arith}
   2.156 -    ultimately have "(\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < L" by blast}
   2.157 -  thus ?thesis by blast
   2.158 -qed
   2.159 -
   2.160 -
   2.161 -subsection{* Some theorems about Sequences*}
   2.162 -text{* Given a binary function @{text "f:: nat \<Rightarrow> 'a \<Rightarrow> 'a"}, its values are uniquely determined by a function g *}
   2.163 -
   2.164 -lemma num_Axiom: "EX! g. g 0 = e \<and> (\<forall>n. g (Suc n) = f n (g n))"
   2.165 -  unfolding Ex1_def
   2.166 -  apply (rule_tac x="nat_rec e f" in exI)
   2.167 -  apply (rule conjI)+
   2.168 -apply (rule def_nat_rec_0, simp)
   2.169 -apply (rule allI, rule def_nat_rec_Suc, simp)
   2.170 -apply (rule allI, rule impI, rule ext)
   2.171 -apply (erule conjE)
   2.172 -apply (induct_tac x)
   2.173 -apply (simp add: nat_rec_0)
   2.174 -apply (erule_tac x="n" in allE)
   2.175 -apply (simp)
   2.176 -done
   2.177 -
   2.178 - text{* An equivalent formulation of monotony -- Not used here, but might be useful *}
   2.179 -lemma mono_Suc: "mono f = (\<forall>n. (f n :: 'a :: order) \<le> f (Suc n))"
   2.180 -unfolding mono_def
   2.181 -proof auto
   2.182 -  fix A B :: nat
   2.183 -  assume H: "\<forall>n. f n \<le> f (Suc n)" "A \<le> B"
   2.184 -  hence "\<exists>k. B = A + k" apply -  apply (thin_tac "\<forall>n. f n \<le> f (Suc n)") 
   2.185 -    by presburger
   2.186 -  then obtain k where k: "B = A + k" by blast
   2.187 -  {fix a k
   2.188 -    have "f a \<le> f (a + k)"
   2.189 -    proof (induct k)
   2.190 -      case 0 thus ?case by simp
   2.191 -    next
   2.192 -      case (Suc k)
   2.193 -      from Suc.hyps H(1)[rule_format, of "a + k"] show ?case by simp
   2.194 -    qed}
   2.195 -  with k show "f A \<le> f B" by blast
   2.196 -qed
   2.197 -
   2.198 -text{* for any sequence, there is a mootonic subsequence *}
   2.199 -lemma seq_monosub: "\<exists>f. subseq f \<and> monoseq (\<lambda> n. (s (f n)))"
   2.200 -proof-
   2.201 -  {assume H: "\<forall>n. \<exists>p >n. \<forall> m\<ge>p. s m \<le> s p"
   2.202 -    let ?P = "\<lambda> p n. p > n \<and> (\<forall>m \<ge> p. s m \<le> s p)"
   2.203 -    from num_Axiom[of "SOME p. ?P p 0" "\<lambda>p n. SOME p. ?P p n"]
   2.204 -    obtain f where f: "f 0 = (SOME p. ?P p 0)" "\<forall>n. f (Suc n) = (SOME p. ?P p (f n))" by blast
   2.205 -    have "?P (f 0) 0"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p 0"]
   2.206 -      using H apply - 
   2.207 -      apply (erule allE[where x=0], erule exE, rule_tac x="p" in exI) 
   2.208 -      unfolding order_le_less by blast 
   2.209 -    hence f0: "f 0 > 0" "\<forall>m \<ge> f 0. s m \<le> s (f 0)" by blast+
   2.210 -    {fix n
   2.211 -      have "?P (f (Suc n)) (f n)" 
   2.212 -	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
   2.213 -	using H apply - 
   2.214 -      apply (erule allE[where x="f n"], erule exE, rule_tac x="p" in exI) 
   2.215 -      unfolding order_le_less by blast 
   2.216 -    hence "f (Suc n) > f n" "\<forall>m \<ge> f (Suc n). s m \<le> s (f (Suc n))" by blast+}
   2.217 -  note fSuc = this
   2.218 -    {fix p q assume pq: "p \<ge> f q"
   2.219 -      have "s p \<le> s(f(q))"  using f0(2)[rule_format, of p] pq fSuc
   2.220 -	by (cases q, simp_all) }
   2.221 -    note pqth = this
   2.222 -    {fix q
   2.223 -      have "f (Suc q) > f q" apply (induct q) 
   2.224 -	using f0(1) fSuc(1)[of 0] apply simp by (rule fSuc(1))}
   2.225 -    note fss = this
   2.226 -    from fss have th1: "subseq f" unfolding subseq_Suc_iff ..
   2.227 -    {fix a b 
   2.228 -      have "f a \<le> f (a + b)"
   2.229 -      proof(induct b)
   2.230 -	case 0 thus ?case by simp
   2.231 -      next
   2.232 -	case (Suc b)
   2.233 -	from fSuc(1)[of "a + b"] Suc.hyps show ?case by simp
   2.234 -      qed}
   2.235 -    note fmon0 = this
   2.236 -    have "monoseq (\<lambda>n. s (f n))" 
   2.237 -    proof-
   2.238 -      {fix n
   2.239 -	have "s (f n) \<ge> s (f (Suc n))" 
   2.240 -	proof(cases n)
   2.241 -	  case 0
   2.242 -	  assume n0: "n = 0"
   2.243 -	  from fSuc(1)[of 0] have th0: "f 0 \<le> f (Suc 0)" by simp
   2.244 -	  from f0(2)[rule_format, OF th0] show ?thesis  using n0 by simp
   2.245 -	next
   2.246 -	  case (Suc m)
   2.247 -	  assume m: "n = Suc m"
   2.248 -	  from fSuc(1)[of n] m have th0: "f (Suc m) \<le> f (Suc (Suc m))" by simp
   2.249 -	  from m fSuc(2)[rule_format, OF th0] show ?thesis by simp 
   2.250 -	qed}
   2.251 -      thus "monoseq (\<lambda>n. s (f n))" unfolding monoseq_Suc by blast 
   2.252 -    qed
   2.253 -    with th1 have ?thesis by blast}
   2.254 -  moreover
   2.255 -  {fix N assume N: "\<forall>p >N. \<exists> m\<ge>p. s m > s p"
   2.256 -    {fix p assume p: "p \<ge> Suc N" 
   2.257 -      hence pN: "p > N" by arith with N obtain m where m: "m \<ge> p" "s m > s p" by blast
   2.258 -      have "m \<noteq> p" using m(2) by auto 
   2.259 -      with m have "\<exists>m>p. s p < s m" by - (rule exI[where x=m], auto)}
   2.260 -    note th0 = this
   2.261 -    let ?P = "\<lambda>m x. m > x \<and> s x < s m"
   2.262 -    from num_Axiom[of "SOME x. ?P x (Suc N)" "\<lambda>m x. SOME y. ?P y x"]
   2.263 -    obtain f where f: "f 0 = (SOME x. ?P x (Suc N))" 
   2.264 -      "\<forall>n. f (Suc n) = (SOME m. ?P m (f n))" by blast
   2.265 -    have "?P (f 0) (Suc N)"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p (Suc N)"]
   2.266 -      using N apply - 
   2.267 -      apply (erule allE[where x="Suc N"], clarsimp)
   2.268 -      apply (rule_tac x="m" in exI)
   2.269 -      apply auto
   2.270 -      apply (subgoal_tac "Suc N \<noteq> m")
   2.271 -      apply simp
   2.272 -      apply (rule ccontr, simp)
   2.273 -      done
   2.274 -    hence f0: "f 0 > Suc N" "s (Suc N) < s (f 0)" by blast+
   2.275 -    {fix n
   2.276 -      have "f n > N \<and> ?P (f (Suc n)) (f n)"
   2.277 -	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
   2.278 -      proof (induct n)
   2.279 -	case 0 thus ?case
   2.280 -	  using f0 N apply auto 
   2.281 -	  apply (erule allE[where x="f 0"], clarsimp) 
   2.282 -	  apply (rule_tac x="m" in exI, simp)
   2.283 -	  by (subgoal_tac "f 0 \<noteq> m", auto)
   2.284 -      next
   2.285 -	case (Suc n)
   2.286 -	from Suc.hyps have Nfn: "N < f n" by blast
   2.287 -	from Suc.hyps obtain m where m: "m > f n" "s (f n) < s m" by blast
   2.288 -	with Nfn have mN: "m > N" by arith
   2.289 -	note key = Suc.hyps[unfolded some_eq_ex[of "\<lambda>p. ?P p (f n)", symmetric] f(2)[rule_format, of n, symmetric]]
   2.290 -	
   2.291 -	from key have th0: "f (Suc n) > N" by simp
   2.292 -	from N[rule_format, OF th0]
   2.293 -	obtain m' where m': "m' \<ge> f (Suc n)" "s (f (Suc n)) < s m'" by blast
   2.294 -	have "m' \<noteq> f (Suc (n))" apply (rule ccontr) using m'(2) by auto
   2.295 -	hence "m' > f (Suc n)" using m'(1) by simp
   2.296 -	with key m'(2) show ?case by auto
   2.297 -      qed}
   2.298 -    note fSuc = this
   2.299 -    {fix n
   2.300 -      have "f n \<ge> Suc N \<and> f(Suc n) > f n \<and> s(f n) < s(f(Suc n))" using fSuc[of n] by auto 
   2.301 -      hence "f n \<ge> Suc N" "f(Suc n) > f n" "s(f n) < s(f(Suc n))" by blast+}
   2.302 -    note thf = this
   2.303 -    have sqf: "subseq f" unfolding subseq_Suc_iff using thf by simp
   2.304 -    have "monoseq (\<lambda>n. s (f n))"  unfolding monoseq_Suc using thf
   2.305 -      apply -
   2.306 -      apply (rule disjI1)
   2.307 -      apply auto
   2.308 -      apply (rule order_less_imp_le)
   2.309 -      apply blast
   2.310 -      done
   2.311 -    then have ?thesis  using sqf by blast}
   2.312 -  ultimately show ?thesis unfolding linorder_not_less[symmetric] by blast
   2.313 -qed
   2.314 -
   2.315 -lemma seq_suble: assumes sf: "subseq f" shows "n \<le> f n"
   2.316 -proof(induct n)
   2.317 -  case 0 thus ?case by simp
   2.318 -next
   2.319 -  case (Suc n)
   2.320 -  from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps
   2.321 -  have "n < f (Suc n)" by arith 
   2.322 -  thus ?case by arith
   2.323 -qed
   2.324 -
   2.325 -subsection {* Fundamental theorem of algebra *}
   2.326 -lemma  unimodular_reduce_norm:
   2.327 -  assumes md: "cmod z = 1"
   2.328 -  shows "cmod (z + 1) < 1 \<or> cmod (z - 1) < 1 \<or> cmod (z + ii) < 1 \<or> cmod (z - ii) < 1"
   2.329 -proof-
   2.330 -  obtain x y where z: "z = Complex x y " by (cases z, auto)
   2.331 -  from md z have xy: "x^2 + y^2 = 1" by (simp add: cmod_def)
   2.332 -  {assume C: "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + ii) \<ge> 1" "cmod (z - ii) \<ge> 1"
   2.333 -    from C z xy have "2*x \<le> 1" "2*x \<ge> -1" "2*y \<le> 1" "2*y \<ge> -1"
   2.334 -      by (simp_all add: cmod_def power2_eq_square ring_simps)
   2.335 -    hence "abs (2*x) \<le> 1" "abs (2*y) \<le> 1" by simp_all
   2.336 -    hence "(abs (2 * x))^2 <= 1^2" "(abs (2 * y)) ^2 <= 1^2"
   2.337 -      by - (rule power_mono, simp, simp)+
   2.338 -    hence th0: "4*x^2 \<le> 1" "4*y^2 \<le> 1" 
   2.339 -      by (simp_all  add: power2_abs power_mult_distrib)
   2.340 -    from add_mono[OF th0] xy have False by simp }
   2.341 -  thus ?thesis unfolding linorder_not_le[symmetric] by blast
   2.342 -qed
   2.343 -
   2.344 -text{* Hence we can always reduce modulus of @{text "1 + b z^n"} if nonzero *}
   2.345 -lemma reduce_poly_simple:
   2.346 - assumes b: "b \<noteq> 0" and n: "n\<noteq>0"
   2.347 -  shows "\<exists>z. cmod (1 + b * z^n) < 1"
   2.348 -using n
   2.349 -proof(induct n rule: nat_less_induct)
   2.350 -  fix n
   2.351 -  assume IH: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>z. cmod (1 + b * z ^ m) < 1)" and n: "n \<noteq> 0"
   2.352 -  let ?P = "\<lambda>z n. cmod (1 + b * z ^ n) < 1"
   2.353 -  {assume e: "even n"
   2.354 -    hence "\<exists>m. n = 2*m" by presburger
   2.355 -    then obtain m where m: "n = 2*m" by blast
   2.356 -    from n m have "m\<noteq>0" "m < n" by presburger+
   2.357 -    with IH[rule_format, of m] obtain z where z: "?P z m" by blast
   2.358 -    from z have "?P (csqrt z) n" by (simp add: m power_mult csqrt)
   2.359 -    hence "\<exists>z. ?P z n" ..}
   2.360 -  moreover
   2.361 -  {assume o: "odd n"
   2.362 -    from b have b': "b^2 \<noteq> 0" unfolding power2_eq_square by simp
   2.363 -    have "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
   2.364 -    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) = 
   2.365 -    ((Re (inverse b))^2 + (Im (inverse b))^2) * \<bar>Im b * Im b + Re b * Re b\<bar>" by algebra
   2.366 -    also have "\<dots> = cmod (inverse b) ^2 * cmod b ^ 2" 
   2.367 -      apply (simp add: cmod_def) using realpow_two_le_add_order[of "Re b" "Im b"]
   2.368 -      by (simp add: power2_eq_square)
   2.369 -    finally 
   2.370 -    have th0: "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
   2.371 -    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) =
   2.372 -    1" 
   2.373 -      apply (simp add: power2_eq_square norm_mult[symmetric] norm_inverse[symmetric])
   2.374 -      using right_inverse[OF b']
   2.375 -      by (simp add: power2_eq_square[symmetric] power_inverse[symmetric] ring_simps)
   2.376 -    have th0: "cmod (complex_of_real (cmod b) / b) = 1"
   2.377 -      apply (simp add: complex_Re_mult cmod_def power2_eq_square Re_complex_of_real Im_complex_of_real divide_inverse ring_simps )
   2.378 -      by (simp add: real_sqrt_mult[symmetric] th0)        
   2.379 -    from o have "\<exists>m. n = Suc (2*m)" by presburger+
   2.380 -    then obtain m where m: "n = Suc (2*m)" by blast
   2.381 -    from unimodular_reduce_norm[OF th0] o
   2.382 -    have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
   2.383 -      apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1", rule_tac x="1" in exI, simp)
   2.384 -      apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp add: diff_def)
   2.385 -      apply (cases "cmod (complex_of_real (cmod b) / b + ii) < 1")
   2.386 -      apply (cases "even m", rule_tac x="ii" in exI, simp add: m power_mult)
   2.387 -      apply (rule_tac x="- ii" in exI, simp add: m power_mult)
   2.388 -      apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult diff_def)
   2.389 -      apply (rule_tac x="ii" in exI, simp add: m power_mult diff_def)
   2.390 -      done
   2.391 -    then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" by blast
   2.392 -    let ?w = "v / complex_of_real (root n (cmod b))"
   2.393 -    from odd_real_root_pow[OF o, of "cmod b"]
   2.394 -    have th1: "?w ^ n = v^n / complex_of_real (cmod b)" 
   2.395 -      by (simp add: power_divide complex_of_real_power)
   2.396 -    have th2:"cmod (complex_of_real (cmod b) / b) = 1" using b by (simp add: norm_divide)
   2.397 -    hence th3: "cmod (complex_of_real (cmod b) / b) \<ge> 0" by simp
   2.398 -    have th4: "cmod (complex_of_real (cmod b) / b) *
   2.399 -   cmod (1 + b * (v ^ n / complex_of_real (cmod b)))
   2.400 -   < cmod (complex_of_real (cmod b) / b) * 1"
   2.401 -      apply (simp only: norm_mult[symmetric] right_distrib)
   2.402 -      using b v by (simp add: th2)
   2.403 -
   2.404 -    from mult_less_imp_less_left[OF th4 th3]
   2.405 -    have "?P ?w n" unfolding th1 . 
   2.406 -    hence "\<exists>z. ?P z n" .. }
   2.407 -  ultimately show "\<exists>z. ?P z n" by blast
   2.408 -qed
   2.409 -
   2.410 -
   2.411 -text{* Bolzano-Weierstrass type property for closed disc in complex plane. *}
   2.412 -
   2.413 -lemma metric_bound_lemma: "cmod (x - y) <= \<bar>Re x - Re y\<bar> + \<bar>Im x - Im y\<bar>"
   2.414 -  using real_sqrt_sum_squares_triangle_ineq[of "Re x - Re y" 0 0 "Im x - Im y" ]
   2.415 -  unfolding cmod_def by simp
   2.416 -
   2.417 -lemma bolzano_weierstrass_complex_disc:
   2.418 -  assumes r: "\<forall>n. cmod (s n) \<le> r"
   2.419 -  shows "\<exists>f z. subseq f \<and> (\<forall>e >0. \<exists>N. \<forall>n \<ge> N. cmod (s (f n) - z) < e)"
   2.420 -proof-
   2.421 -  from seq_monosub[of "Re o s"] 
   2.422 -  obtain f g where f: "subseq f" "monoseq (\<lambda>n. Re (s (f n)))" 
   2.423 -    unfolding o_def by blast
   2.424 -  from seq_monosub[of "Im o s o f"] 
   2.425 -  obtain g where g: "subseq g" "monoseq (\<lambda>n. Im (s(f(g n))))" unfolding o_def by blast  
   2.426 -  let ?h = "f o g"
   2.427 -  from r[rule_format, of 0] have rp: "r \<ge> 0" using norm_ge_zero[of "s 0"] by arith 
   2.428 -  have th:"\<forall>n. r + 1 \<ge> \<bar> Re (s n)\<bar>" 
   2.429 -  proof
   2.430 -    fix n
   2.431 -    from abs_Re_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Re (s n)\<bar> \<le> r + 1" by arith
   2.432 -  qed
   2.433 -  have conv1: "convergent (\<lambda>n. Re (s ( f n)))"
   2.434 -    apply (rule Bseq_monoseq_convergent)
   2.435 -    apply (simp add: Bseq_def)
   2.436 -    apply (rule exI[where x= "r + 1"])
   2.437 -    using th rp apply simp
   2.438 -    using f(2) .
   2.439 -  have th:"\<forall>n. r + 1 \<ge> \<bar> Im (s n)\<bar>" 
   2.440 -  proof
   2.441 -    fix n
   2.442 -    from abs_Im_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Im (s n)\<bar> \<le> r + 1" by arith
   2.443 -  qed
   2.444 -
   2.445 -  have conv2: "convergent (\<lambda>n. Im (s (f (g n))))"
   2.446 -    apply (rule Bseq_monoseq_convergent)
   2.447 -    apply (simp add: Bseq_def)
   2.448 -    apply (rule exI[where x= "r + 1"])
   2.449 -    using th rp apply simp
   2.450 -    using g(2) .
   2.451 -
   2.452 -  from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x" 
   2.453 -    by blast 
   2.454 -  hence  x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Re (s (f n)) - x \<bar> < r" 
   2.455 -    unfolding LIMSEQ_def real_norm_def .
   2.456 -
   2.457 -  from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y" 
   2.458 -    by blast 
   2.459 -  hence  y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Im (s (f (g n))) - y \<bar> < r" 
   2.460 -    unfolding LIMSEQ_def real_norm_def .
   2.461 -  let ?w = "Complex x y"
   2.462 -  from f(1) g(1) have hs: "subseq ?h" unfolding subseq_def by auto 
   2.463 -  {fix e assume ep: "e > (0::real)"
   2.464 -    hence e2: "e/2 > 0" by simp
   2.465 -    from x[rule_format, OF e2] y[rule_format, OF e2]
   2.466 -    obtain N1 N2 where N1: "\<forall>n\<ge>N1. \<bar>Re (s (f n)) - x\<bar> < e / 2" and N2: "\<forall>n\<ge>N2. \<bar>Im (s (f (g n))) - y\<bar> < e / 2" by blast
   2.467 -    {fix n assume nN12: "n \<ge> N1 + N2"
   2.468 -      hence nN1: "g n \<ge> N1" and nN2: "n \<ge> N2" using seq_suble[OF g(1), of n] by arith+
   2.469 -      from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
   2.470 -      have "cmod (s (?h n) - ?w) < e" 
   2.471 -	using metric_bound_lemma[of "s (f (g n))" ?w] by simp }
   2.472 -    hence "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" by blast }
   2.473 -  with hs show ?thesis  by blast  
   2.474 -qed
   2.475 -
   2.476 -text{* Polynomial is continuous. *}
   2.477 -
   2.478 -lemma poly_cont:
   2.479 -  assumes ep: "e > 0" 
   2.480 -  shows "\<exists>d >0. \<forall>w. 0 < cmod (w - z) \<and> cmod (w - z) < d \<longrightarrow> cmod (poly p w - poly p z) < e"
   2.481 -proof-
   2.482 -  from poly_offset[of p z] obtain q where q: "length q = length p" "\<And>x. poly q x = poly p (z + x)" by blast
   2.483 -  {fix w
   2.484 -    note q(2)[of "w - z", simplified]}
   2.485 -  note th = this
   2.486 -  show ?thesis unfolding th[symmetric]
   2.487 -  proof(induct q)
   2.488 -    case Nil thus ?case  using ep by auto
   2.489 -  next
   2.490 -    case (Cons c cs)
   2.491 -    from poly_bound_exists[of 1 "cs"] 
   2.492 -    obtain m where m: "m > 0" "\<And>z. cmod z \<le> 1 \<Longrightarrow> cmod (poly cs z) \<le> m" by blast
   2.493 -    from ep m(1) have em0: "e/m > 0" by (simp add: field_simps)
   2.494 -    have one0: "1 > (0::real)"  by arith
   2.495 -    from real_lbound_gt_zero[OF one0 em0] 
   2.496 -    obtain d where d: "d >0" "d < 1" "d < e / m" by blast
   2.497 -    from d(1,3) m(1) have dm: "d*m > 0" "d*m < e" 
   2.498 -      by (simp_all add: field_simps real_mult_order)
   2.499 -    show ?case 
   2.500 -      proof(rule ex_forward[OF real_lbound_gt_zero[OF one0 em0]], clarsimp simp add: norm_mult)
   2.501 -	fix d w
   2.502 -	assume H: "d > 0" "d < 1" "d < e/m" "w\<noteq>z" "cmod (w-z) < d"
   2.503 -	hence d1: "cmod (w-z) \<le> 1" "d \<ge> 0" by simp_all
   2.504 -	from H(3) m(1) have dme: "d*m < e" by (simp add: field_simps)
   2.505 -	from H have th: "cmod (w-z) \<le> d" by simp 
   2.506 -	from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
   2.507 -	show "cmod (w - z) * cmod (poly cs (w - z)) < e" by simp
   2.508 -      qed  
   2.509 -    qed
   2.510 -qed
   2.511 -
   2.512 -text{* Hence a polynomial attains minimum on a closed disc 
   2.513 -  in the complex plane. *}
   2.514 -lemma  poly_minimum_modulus_disc:
   2.515 -  "\<exists>z. \<forall>w. cmod w \<le> r \<longrightarrow> cmod (poly p z) \<le> cmod (poly p w)"
   2.516 -proof-
   2.517 -  {assume "\<not> r \<ge> 0" hence ?thesis unfolding linorder_not_le
   2.518 -      apply -
   2.519 -      apply (rule exI[where x=0]) 
   2.520 -      apply auto
   2.521 -      apply (subgoal_tac "cmod w < 0")
   2.522 -      apply simp
   2.523 -      apply arith
   2.524 -      done }
   2.525 -  moreover
   2.526 -  {assume rp: "r \<ge> 0"
   2.527 -    from rp have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))" by simp 
   2.528 -    hence mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x"  by blast
   2.529 -    {fix x z
   2.530 -      assume H: "cmod z \<le> r" "cmod (poly p z) = - x" "\<not>x < 1"
   2.531 -      hence "- x < 0 " by arith
   2.532 -      with H(2) norm_ge_zero[of "poly p z"]  have False by simp }
   2.533 -    then have mth2: "\<exists>z. \<forall>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<longrightarrow> x < z" by blast
   2.534 -    from real_sup_exists[OF mth1 mth2] obtain s where 
   2.535 -      s: "\<forall>y. (\<exists>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<and> y < x) \<longleftrightarrow>(y < s)" by blast
   2.536 -    let ?m = "-s"
   2.537 -    {fix y
   2.538 -      from s[rule_format, of "-y"] have 
   2.539 -    "(\<exists>z x. cmod z \<le> r \<and> -(- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y" 
   2.540 -	unfolding minus_less_iff[of y ] equation_minus_iff by blast }
   2.541 -    note s1 = this[unfolded minus_minus]
   2.542 -    from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m" 
   2.543 -      by auto
   2.544 -    {fix n::nat
   2.545 -      from s1[rule_format, of "?m + 1/real (Suc n)"] 
   2.546 -      have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)"
   2.547 -	by simp}
   2.548 -    hence th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
   2.549 -    from choice[OF th] obtain g where 
   2.550 -      g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m+1 /real(Suc n)" 
   2.551 -      by blast
   2.552 -    from bolzano_weierstrass_complex_disc[OF g(1)] 
   2.553 -    obtain f z where fz: "subseq f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
   2.554 -      by blast    
   2.555 -    {fix w 
   2.556 -      assume wr: "cmod w \<le> r"
   2.557 -      let ?e = "\<bar>cmod (poly p z) - ?m\<bar>"
   2.558 -      {assume e: "?e > 0"
   2.559 -	hence e2: "?e/2 > 0" by simp
   2.560 -	from poly_cont[OF e2, of z p] obtain d where
   2.561 -	  d: "d>0" "\<forall>w. 0<cmod (w - z)\<and> cmod(w - z) < d \<longrightarrow> cmod(poly p w - poly p z) < ?e/2" by blast
   2.562 -	{fix w assume w: "cmod (w - z) < d"
   2.563 -	  have "cmod(poly p w - poly p z) < ?e / 2"
   2.564 -	    using d(2)[rule_format, of w] w e by (cases "w=z", simp_all)}
   2.565 -	note th1 = this
   2.566 -	
   2.567 -	from fz(2)[rule_format, OF d(1)] obtain N1 where 
   2.568 -	  N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d" by blast
   2.569 -	from reals_Archimedean2[of "2/?e"] obtain N2::nat where
   2.570 -	  N2: "2/?e < real N2" by blast
   2.571 -	have th2: "cmod(poly p (g(f(N1 + N2))) - poly p z) < ?e/2"
   2.572 -	  using N1[rule_format, of "N1 + N2"] th1 by simp
   2.573 -	{fix a b e2 m :: real
   2.574 -	have "a < e2 \<Longrightarrow> abs(b - m) < e2 \<Longrightarrow> 2 * e2 <= abs(b - m) + a
   2.575 -          ==> False" by arith}
   2.576 -      note th0 = this
   2.577 -      have ath: 
   2.578 -	"\<And>m x e. m <= x \<Longrightarrow>  x < m + e ==> abs(x - m::real) < e" by arith
   2.579 -      from s1m[OF g(1)[rule_format]]
   2.580 -      have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
   2.581 -      from seq_suble[OF fz(1), of "N1+N2"]
   2.582 -      have th00: "real (Suc (N1+N2)) \<le> real (Suc (f (N1+N2)))" by simp
   2.583 -      have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1+N2)) > 0"  
   2.584 -	using N2 by auto
   2.585 -      from frac_le[OF th000 th00] have th00: "?m +1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))" by simp
   2.586 -      from g(2)[rule_format, of "f (N1 + N2)"]
   2.587 -      have th01:"cmod (poly p (g (f (N1 + N2)))) < - s + 1 / real (Suc (f (N1 + N2)))" .
   2.588 -      from order_less_le_trans[OF th01 th00]
   2.589 -      have th32: "cmod(poly p (g (f (N1 + N2)))) < ?m + (1/ real(Suc (N1 + N2)))" .
   2.590 -      from N2 have "2/?e < real (Suc (N1 + N2))" by arith
   2.591 -      with e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
   2.592 -      have "?e/2 > 1/ real (Suc (N1 + N2))" by (simp add: inverse_eq_divide)
   2.593 -      with ath[OF th31 th32]
   2.594 -      have thc1:"\<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar>< ?e/2" by arith  
   2.595 -      have ath2: "\<And>(a::real) b c m. \<bar>a - b\<bar> <= c ==> \<bar>b - m\<bar> <= \<bar>a - m\<bar> + c" 
   2.596 -	by arith
   2.597 -      have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar>
   2.598 -\<le> cmod (poly p (g (f (N1 + N2))) - poly p z)" 
   2.599 -	by (simp add: norm_triangle_ineq3)
   2.600 -      from ath2[OF th22, of ?m]
   2.601 -      have thc2: "2*(?e/2) \<le> \<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar> + cmod (poly p (g (f (N1 + N2))) - poly p z)" by simp
   2.602 -      from th0[OF th2 thc1 thc2] have False .}
   2.603 -      hence "?e = 0" by auto
   2.604 -      then have "cmod (poly p z) = ?m" by simp  
   2.605 -      with s1m[OF wr]
   2.606 -      have "cmod (poly p z) \<le> cmod (poly p w)" by simp }
   2.607 -    hence ?thesis by blast}
   2.608 -  ultimately show ?thesis by blast
   2.609 -qed
   2.610 -
   2.611 -lemma "(rcis (sqrt (abs r)) (a/2)) ^ 2 = rcis (abs r) a"
   2.612 -  unfolding power2_eq_square
   2.613 -  apply (simp add: rcis_mult)
   2.614 -  apply (simp add: power2_eq_square[symmetric])
   2.615 -  done
   2.616 -
   2.617 -lemma cispi: "cis pi = -1" 
   2.618 -  unfolding cis_def
   2.619 -  by simp
   2.620 -
   2.621 -lemma "(rcis (sqrt (abs r)) ((pi + a)/2)) ^ 2 = rcis (- abs r) a"
   2.622 -  unfolding power2_eq_square
   2.623 -  apply (simp add: rcis_mult add_divide_distrib)
   2.624 -  apply (simp add: power2_eq_square[symmetric] rcis_def cispi cis_mult[symmetric])
   2.625 -  done
   2.626 -
   2.627 -text {* Nonzero polynomial in z goes to infinity as z does. *}
   2.628 -
   2.629 -instance complex::idom_char_0 by (intro_classes)
   2.630 -instance complex :: recpower_idom_char_0 by intro_classes
   2.631 -
   2.632 -lemma poly_infinity:
   2.633 -  assumes ex: "list_ex (\<lambda>c. c \<noteq> 0) p"
   2.634 -  shows "\<exists>r. \<forall>z. r \<le> cmod z \<longrightarrow> d \<le> cmod (poly (a#p) z)"
   2.635 -using ex
   2.636 -proof(induct p arbitrary: a d)
   2.637 -  case (Cons c cs a d) 
   2.638 -  {assume H: "list_ex (\<lambda>c. c\<noteq>0) cs"
   2.639 -    with Cons.hyps obtain r where r: "\<forall>z. r \<le> cmod z \<longrightarrow> d + cmod a \<le> cmod (poly (c # cs) z)" by blast
   2.640 -    let ?r = "1 + \<bar>r\<bar>"
   2.641 -    {fix z assume h: "1 + \<bar>r\<bar> \<le> cmod z"
   2.642 -      have r0: "r \<le> cmod z" using h by arith
   2.643 -      from r[rule_format, OF r0]
   2.644 -      have th0: "d + cmod a \<le> 1 * cmod(poly (c#cs) z)" by arith
   2.645 -      from h have z1: "cmod z \<ge> 1" by arith
   2.646 -      from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (c#cs) z"]]]
   2.647 -      have th1: "d \<le> cmod(z * poly (c#cs) z) - cmod a"
   2.648 -	unfolding norm_mult by (simp add: ring_simps)
   2.649 -      from complex_mod_triangle_sub[of "z * poly (c#cs) z" a]
   2.650 -      have th2: "cmod(z * poly (c#cs) z) - cmod a \<le> cmod (poly (a#c#cs) z)" 
   2.651 -	by (simp add: diff_le_eq ring_simps) 
   2.652 -      from th1 th2 have "d \<le> cmod (poly (a#c#cs) z)"  by arith}
   2.653 -    hence ?case by blast}
   2.654 -  moreover
   2.655 -  {assume cs0: "\<not> (list_ex (\<lambda>c. c \<noteq> 0) cs)"
   2.656 -    with Cons.prems have c0: "c \<noteq> 0" by simp
   2.657 -    from cs0 have cs0': "list_all (\<lambda>c. c = 0) cs" 
   2.658 -      by (auto simp add: list_all_iff list_ex_iff)
   2.659 -    {fix z
   2.660 -      assume h: "(\<bar>d\<bar> + cmod a) / cmod c \<le> cmod z"
   2.661 -      from c0 have "cmod c > 0" by simp
   2.662 -      from h c0 have th0: "\<bar>d\<bar> + cmod a \<le> cmod (z*c)" 
   2.663 -	by (simp add: field_simps norm_mult)
   2.664 -      have ath: "\<And>mzh mazh ma. mzh <= mazh + ma ==> abs(d) + ma <= mzh ==> d <= mazh" by arith
   2.665 -      from complex_mod_triangle_sub[of "z*c" a ]
   2.666 -      have th1: "cmod (z * c) \<le> cmod (a + z * c) + cmod a"
   2.667 -	by (simp add: ring_simps)
   2.668 -      from ath[OF th1 th0] have "d \<le> cmod (poly (a # c # cs) z)" 
   2.669 -	using poly_0[OF cs0'] by simp}
   2.670 -    then have ?case  by blast}
   2.671 -  ultimately show ?case by blast
   2.672 -qed simp
   2.673 -
   2.674 -text {* Hence polynomial's modulus attains its minimum somewhere. *}
   2.675 -lemma poly_minimum_modulus:
   2.676 -  "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)"
   2.677 -proof(induct p)
   2.678 -  case (Cons c cs) 
   2.679 -  {assume cs0: "list_ex (\<lambda>c. c \<noteq> 0) cs"
   2.680 -    from poly_infinity[OF cs0, of "cmod (poly (c#cs) 0)" c]
   2.681 -    obtain r where r: "\<And>z. r \<le> cmod z \<Longrightarrow> cmod (poly (c # cs) 0) \<le> cmod (poly (c # cs) z)" by blast
   2.682 -    have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>" by arith
   2.683 -    from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "c#cs"] 
   2.684 -    obtain v where v: "\<And>w. cmod w \<le> \<bar>r\<bar> \<Longrightarrow> cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) w)" by blast
   2.685 -    {fix z assume z: "r \<le> cmod z"
   2.686 -      from v[of 0] r[OF z] 
   2.687 -      have "cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) z)"
   2.688 -	by simp }
   2.689 -    note v0 = this
   2.690 -    from v0 v ath[of r] have ?case by blast}
   2.691 -  moreover
   2.692 -  {assume cs0: "\<not> (list_ex (\<lambda>c. c\<noteq>0) cs)"
   2.693 -    hence th:"list_all (\<lambda>c. c = 0) cs" by (simp add: list_all_iff list_ex_iff)
   2.694 -    from poly_0[OF th] Cons.hyps have ?case by simp}
   2.695 -  ultimately show ?case by blast
   2.696 -qed simp
   2.697 -
   2.698 -text{* Constant function (non-syntactic characterization). *}
   2.699 -definition "constant f = (\<forall>x y. f x = f y)"
   2.700 -
   2.701 -lemma nonconstant_length: "\<not> (constant (poly p)) \<Longrightarrow> length p \<ge> 2"
   2.702 -  unfolding constant_def
   2.703 -  apply (induct p, auto)
   2.704 -  apply (unfold not_less[symmetric])
   2.705 -  apply simp
   2.706 -  apply (rule ccontr)
   2.707 -  apply auto
   2.708 -  done
   2.709 - 
   2.710 -lemma poly_replicate_append:
   2.711 -  "poly ((replicate n 0)@p) (x::'a::{recpower, comm_ring}) = x^n * poly p x"
   2.712 -  by(induct n, auto simp add: power_Suc ring_simps)
   2.713 -
   2.714 -text {* Decomposition of polynomial, skipping zero coefficients 
   2.715 -  after the first.  *}
   2.716 -
   2.717 -lemma poly_decompose_lemma:
   2.718 - assumes nz: "\<not>(\<forall>z. z\<noteq>0 \<longrightarrow> poly p z = (0::'a::{recpower,idom}))"
   2.719 -  shows "\<exists>k a q. a\<noteq>0 \<and> Suc (length q + k) = length p \<and> 
   2.720 -                 (\<forall>z. poly p z = z^k * poly (a#q) z)"
   2.721 -using nz
   2.722 -proof(induct p)
   2.723 -  case Nil thus ?case by simp
   2.724 -next
   2.725 -  case (Cons c cs)
   2.726 -  {assume c0: "c = 0"
   2.727 -    
   2.728 -    from Cons.hyps Cons.prems c0 have ?case apply auto
   2.729 -      apply (rule_tac x="k+1" in exI)
   2.730 -      apply (rule_tac x="a" in exI, clarsimp)
   2.731 -      apply (rule_tac x="q" in exI)
   2.732 -      by (auto simp add: power_Suc)}
   2.733 -  moreover
   2.734 -  {assume c0: "c\<noteq>0"
   2.735 -    hence ?case apply-
   2.736 -      apply (rule exI[where x=0])
   2.737 -      apply (rule exI[where x=c], clarsimp)
   2.738 -      apply (rule exI[where x=cs])
   2.739 -      apply auto
   2.740 -      done}
   2.741 -  ultimately show ?case by blast
   2.742 -qed
   2.743 -
   2.744 -lemma poly_decompose:
   2.745 -  assumes nc: "~constant(poly p)"
   2.746 -  shows "\<exists>k a q. a\<noteq>(0::'a::{recpower,idom}) \<and> k\<noteq>0 \<and>
   2.747 -               length q + k + 1 = length p \<and> 
   2.748 -              (\<forall>z. poly p z = poly p 0 + z^k * poly (a#q) z)"
   2.749 -using nc 
   2.750 -proof(induct p)
   2.751 -  case Nil thus ?case by (simp add: constant_def)
   2.752 -next
   2.753 -  case (Cons c cs)
   2.754 -  {assume C:"\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0"
   2.755 -    {fix x y
   2.756 -      from C have "poly (c#cs) x = poly (c#cs) y" by (cases "x=0", auto)}
   2.757 -    with Cons.prems have False by (auto simp add: constant_def)}
   2.758 -  hence th: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)" ..
   2.759 -  from poly_decompose_lemma[OF th] 
   2.760 -  show ?case 
   2.761 -    apply clarsimp    
   2.762 -    apply (rule_tac x="k+1" in exI)
   2.763 -    apply (rule_tac x="a" in exI)
   2.764 -    apply simp
   2.765 -    apply (rule_tac x="q" in exI)
   2.766 -    apply (auto simp add: power_Suc)
   2.767 -    done
   2.768 -qed
   2.769 -
   2.770 -text{* Fundamental theorem of algebral *}
   2.771 -
   2.772 -lemma fundamental_theorem_of_algebra:
   2.773 -  assumes nc: "~constant(poly p)"
   2.774 -  shows "\<exists>z::complex. poly p z = 0"
   2.775 -using nc
   2.776 -proof(induct n\<equiv> "length p" arbitrary: p rule: nat_less_induct)
   2.777 -  fix n fix p :: "complex list"
   2.778 -  let ?p = "poly p"
   2.779 -  assume H: "\<forall>m<n. \<forall>p. \<not> constant (poly p) \<longrightarrow> m = length p \<longrightarrow> (\<exists>(z::complex). poly p z = 0)" and nc: "\<not> constant ?p" and n: "n = length p"
   2.780 -  let ?ths = "\<exists>z. ?p z = 0"
   2.781 -
   2.782 -  from nonconstant_length[OF nc] have n2: "n\<ge> 2" by (simp add: n)
   2.783 -  from poly_minimum_modulus obtain c where 
   2.784 -    c: "\<forall>w. cmod (?p c) \<le> cmod (?p w)" by blast
   2.785 -  {assume pc: "?p c = 0" hence ?ths by blast}
   2.786 -  moreover
   2.787 -  {assume pc0: "?p c \<noteq> 0"
   2.788 -    from poly_offset[of p c] obtain q where
   2.789 -      q: "length q = length p" "\<forall>x. poly q x = ?p (c+x)" by blast
   2.790 -    {assume h: "constant (poly q)"
   2.791 -      from q(2) have th: "\<forall>x. poly q (x - c) = ?p x" by auto
   2.792 -      {fix x y
   2.793 -	from th have "?p x = poly q (x - c)" by auto 
   2.794 -	also have "\<dots> = poly q (y - c)" 
   2.795 -	  using h unfolding constant_def by blast
   2.796 -	also have "\<dots> = ?p y" using th by auto
   2.797 -	finally have "?p x = ?p y" .}
   2.798 -      with nc have False unfolding constant_def by blast }
   2.799 -    hence qnc: "\<not> constant (poly q)" by blast
   2.800 -    from q(2) have pqc0: "?p c = poly q 0" by simp
   2.801 -    from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)" by simp 
   2.802 -    let ?a0 = "poly q 0"
   2.803 -    from pc0 pqc0 have a00: "?a0 \<noteq> 0" by simp 
   2.804 -    from a00 
   2.805 -    have qr: "\<forall>z. poly q z = poly (map (op * (inverse ?a0)) q) z * ?a0"
   2.806 -      by (simp add: poly_cmult_map)
   2.807 -    let ?r = "map (op * (inverse ?a0)) q"
   2.808 -    have lgqr: "length q = length ?r" by simp 
   2.809 -    {assume h: "\<And>x y. poly ?r x = poly ?r y"
   2.810 -      {fix x y
   2.811 -	from qr[rule_format, of x] 
   2.812 -	have "poly q x = poly ?r x * ?a0" by auto
   2.813 -	also have "\<dots> = poly ?r y * ?a0" using h by simp
   2.814 -	also have "\<dots> = poly q y" using qr[rule_format, of y] by simp
   2.815 -	finally have "poly q x = poly q y" .} 
   2.816 -      with qnc have False unfolding constant_def by blast}
   2.817 -    hence rnc: "\<not> constant (poly ?r)" unfolding constant_def by blast
   2.818 -    from qr[rule_format, of 0] a00  have r01: "poly ?r 0 = 1" by auto
   2.819 -    {fix w 
   2.820 -      have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
   2.821 -	using qr[rule_format, of w] a00 by simp
   2.822 -      also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
   2.823 -	using a00 unfolding norm_divide by (simp add: field_simps)
   2.824 -      finally have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" .}
   2.825 -    note mrmq_eq = this
   2.826 -    from poly_decompose[OF rnc] obtain k a s where 
   2.827 -      kas: "a\<noteq>0" "k\<noteq>0" "length s + k + 1 = length ?r" 
   2.828 -      "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (a#s) z" by blast
   2.829 -    {assume "k + 1 = n"
   2.830 -      with kas(3) lgqr[symmetric] q(1) n[symmetric] have s0:"s=[]" by auto
   2.831 -      {fix w
   2.832 -	have "cmod (poly ?r w) = cmod (1 + a * w ^ k)" 
   2.833 -	  using kas(4)[rule_format, of w] s0 r01 by (simp add: ring_simps)}
   2.834 -      note hth = this [symmetric]
   2.835 -	from reduce_poly_simple[OF kas(1,2)] 
   2.836 -      have "\<exists>w. cmod (poly ?r w) < 1" unfolding hth by blast}
   2.837 -    moreover
   2.838 -    {assume kn: "k+1 \<noteq> n"
   2.839 -      from kn kas(3) q(1) n[symmetric] have k1n: "k + 1 < n" by simp
   2.840 -      have th01: "\<not> constant (poly (1#((replicate (k - 1) 0)@[a])))" 
   2.841 -	unfolding constant_def poly_Nil poly_Cons poly_replicate_append
   2.842 -	using kas(1) apply simp 
   2.843 -	by (rule exI[where x=0], rule exI[where x=1], simp)
   2.844 -      from kas(2) have th02: "k+1 = length (1#((replicate (k - 1) 0)@[a]))" 
   2.845 -	by simp
   2.846 -      from H[rule_format, OF k1n th01 th02]
   2.847 -      obtain w where w: "1 + w^k * a = 0"
   2.848 -	unfolding poly_Nil poly_Cons poly_replicate_append
   2.849 -	using kas(2) by (auto simp add: power_Suc[symmetric, of _ "k - Suc 0"] 
   2.850 -	  mult_assoc[of _ _ a, symmetric])
   2.851 -      from poly_bound_exists[of "cmod w" s] obtain m where 
   2.852 -	m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
   2.853 -      have w0: "w\<noteq>0" using kas(2) w by (auto simp add: power_0_left)
   2.854 -      from w have "(1 + w ^ k * a) - 1 = 0 - 1" by simp
   2.855 -      then have wm1: "w^k * a = - 1" by simp
   2.856 -      have inv0: "0 < inverse (cmod w ^ (k + 1) * m)" 
   2.857 -	using norm_ge_zero[of w] w0 m(1)
   2.858 -	  by (simp add: inverse_eq_divide zero_less_mult_iff)
   2.859 -      with real_down2[OF zero_less_one] obtain t where
   2.860 -	t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
   2.861 -      let ?ct = "complex_of_real t"
   2.862 -      let ?w = "?ct * w"
   2.863 -      have "1 + ?w^k * (a + ?w * poly s ?w) = 1 + ?ct^k * (w^k * a) + ?w^k * ?w * poly s ?w" using kas(1) by (simp add: ring_simps power_mult_distrib)
   2.864 -      also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
   2.865 -	unfolding wm1 by (simp)
   2.866 -      finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) = cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)" 
   2.867 -	apply -
   2.868 -	apply (rule cong[OF refl[of cmod]])
   2.869 -	apply assumption
   2.870 -	done
   2.871 -      with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"] 
   2.872 -      have th11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)" unfolding norm_of_real by simp 
   2.873 -      have ath: "\<And>x (t::real). 0\<le> x \<Longrightarrow> x < t \<Longrightarrow> t\<le>1 \<Longrightarrow> \<bar>1 - t\<bar> + x < 1" by arith
   2.874 -      have "t *cmod w \<le> 1 * cmod w" apply (rule mult_mono) using t(1,2) by auto
   2.875 -      then have tw: "cmod ?w \<le> cmod w" using t(1) by (simp add: norm_mult) 
   2.876 -      from t inv0 have "t* (cmod w ^ (k + 1) * m) < 1"
   2.877 -	by (simp add: inverse_eq_divide field_simps)
   2.878 -      with zero_less_power[OF t(1), of k] 
   2.879 -      have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1" 
   2.880 -	apply - apply (rule mult_strict_left_mono) by simp_all
   2.881 -      have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k+1) * cmod (poly s ?w)))"  using w0 t(1)
   2.882 -	by (simp add: ring_simps power_mult_distrib norm_of_real norm_power norm_mult)
   2.883 -      then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
   2.884 -	using t(1,2) m(2)[rule_format, OF tw] w0
   2.885 -	apply (simp only: )
   2.886 -	apply auto
   2.887 -	apply (rule mult_mono, simp_all add: norm_ge_zero)+
   2.888 -	apply (simp add: zero_le_mult_iff zero_le_power)
   2.889 -	done
   2.890 -      with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k" by simp 
   2.891 -      from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1" 
   2.892 -	by auto
   2.893 -      from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] th120 th121]
   2.894 -      have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" . 
   2.895 -      from th11 th12
   2.896 -      have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"  by arith 
   2.897 -      then have "cmod (poly ?r ?w) < 1" 
   2.898 -	unfolding kas(4)[rule_format, of ?w] r01 by simp 
   2.899 -      then have "\<exists>w. cmod (poly ?r w) < 1" by blast}
   2.900 -    ultimately have cr0_contr: "\<exists>w. cmod (poly ?r w) < 1" by blast
   2.901 -    from cr0_contr cq0 q(2)
   2.902 -    have ?ths unfolding mrmq_eq not_less[symmetric] by auto}
   2.903 -  ultimately show ?ths by blast
   2.904 -qed
   2.905 -
   2.906 -text {* Alternative version with a syntactic notion of constant polynomial. *}
   2.907 -
   2.908 -lemma fundamental_theorem_of_algebra_alt:
   2.909 -  assumes nc: "~(\<exists>a l. a\<noteq> 0 \<and> list_all(\<lambda>b. b = 0) l \<and> p = a#l)"
   2.910 -  shows "\<exists>z. poly p z = (0::complex)"
   2.911 -using nc
   2.912 -proof(induct p)
   2.913 -  case (Cons c cs)
   2.914 -  {assume "c=0" hence ?case by auto}
   2.915 -  moreover
   2.916 -  {assume c0: "c\<noteq>0"
   2.917 -    {assume nc: "constant (poly (c#cs))"
   2.918 -      from nc[unfolded constant_def, rule_format, of 0] 
   2.919 -      have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto 
   2.920 -      hence "list_all (\<lambda>c. c=0) cs"
   2.921 -	proof(induct cs)
   2.922 -	  case (Cons d ds)
   2.923 -	  {assume "d=0" hence ?case using Cons.prems Cons.hyps by simp}
   2.924 -	  moreover
   2.925 -	  {assume d0: "d\<noteq>0"
   2.926 -	    from poly_bound_exists[of 1 ds] obtain m where 
   2.927 -	      m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
   2.928 -	    have dm: "cmod d / m > 0" using d0 m(1) by (simp add: field_simps)
   2.929 -	    from real_down2[OF dm zero_less_one] obtain x where 
   2.930 -	      x: "x > 0" "x < cmod d / m" "x < 1" by blast
   2.931 -	    let ?x = "complex_of_real x"
   2.932 -	    from x have cx: "?x \<noteq> 0"  "cmod ?x \<le> 1" by simp_all
   2.933 -	    from Cons.prems[rule_format, OF cx(1)]
   2.934 -	    have cth: "cmod (?x*poly ds ?x) = cmod d" by (simp add: eq_diff_eq[symmetric])
   2.935 -	    from m(2)[rule_format, OF cx(2)] x(1)
   2.936 -	    have th0: "cmod (?x*poly ds ?x) \<le> x*m"
   2.937 -	      by (simp add: norm_mult)
   2.938 -	    from x(2) m(1) have "x*m < cmod d" by (simp add: field_simps)
   2.939 -	    with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d" by auto
   2.940 -	    with cth  have ?case by blast}
   2.941 -	  ultimately show ?case by blast 
   2.942 -	qed simp}
   2.943 -      then have nc: "\<not> constant (poly (c#cs))" using Cons.prems c0 
   2.944 -	by blast
   2.945 -      from fundamental_theorem_of_algebra[OF nc] have ?case .}
   2.946 -  ultimately show ?case by blast  
   2.947 -qed simp
   2.948 -
   2.949 -subsection{* Nullstellenstatz, degrees and divisibility of polynomials *}
   2.950 -
   2.951 -lemma nullstellensatz_lemma:
   2.952 -  fixes p :: "complex list"
   2.953 -  assumes "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
   2.954 -  and "degree p = n" and "n \<noteq> 0"
   2.955 -  shows "p divides (pexp q n)"
   2.956 -using prems
   2.957 -proof(induct n arbitrary: p q rule: nat_less_induct)
   2.958 -  fix n::nat fix p q :: "complex list"
   2.959 -  assume IH: "\<forall>m<n. \<forall>p q.
   2.960 -                 (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow>
   2.961 -                 degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p divides (q %^ m)"
   2.962 -    and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0" 
   2.963 -    and dpn: "degree p = n" and n0: "n \<noteq> 0"
   2.964 -  let ?ths = "p divides (q %^ n)"
   2.965 -  {fix a assume a: "poly p a = 0"
   2.966 -    {assume p0: "poly p = poly []" 
   2.967 -      hence ?ths unfolding divides_def  using pq0 n0
   2.968 -	apply - apply (rule exI[where x="[]"], rule ext)
   2.969 -	by (auto simp add: poly_mult poly_exp)}
   2.970 -    moreover
   2.971 -    {assume p0: "poly p \<noteq> poly []" 
   2.972 -      and oa: "order  a p \<noteq> 0"
   2.973 -      from p0 have pne: "p \<noteq> []" by auto
   2.974 -      let ?op = "order a p"
   2.975 -      from p0 have ap: "([- a, 1] %^ ?op) divides p" 
   2.976 -	"\<not> pexp [- a, 1] (Suc ?op) divides p" using order by blast+ 
   2.977 -      note oop = order_degree[OF p0, unfolded dpn]
   2.978 -      {assume q0: "q = []"
   2.979 -	hence ?ths using n0 unfolding divides_def 
   2.980 -	  apply simp
   2.981 -	  apply (rule exI[where x="[]"], rule ext)
   2.982 -	  by (simp add: divides_def poly_exp poly_mult)}
   2.983 -      moreover
   2.984 -      {assume q0: "q\<noteq>[]"
   2.985 -	from pq0[rule_format, OF a, unfolded poly_linear_divides] q0
   2.986 -	obtain r where r: "q = pmult [- a, 1] r" by blast
   2.987 -	from ap[unfolded divides_def] obtain s where
   2.988 -	  s: "poly p = poly (pmult (pexp [- a, 1] ?op) s)" by blast
   2.989 -	have s0: "poly s \<noteq> poly []"
   2.990 -	  using s p0 by (simp add: poly_entire)
   2.991 -	hence pns0: "poly (pnormalize s) \<noteq> poly []" and sne: "s\<noteq>[]" by auto
   2.992 -	{assume ds0: "degree s = 0"
   2.993 -	  from ds0 pns0 have "\<exists>k. pnormalize s = [k]" unfolding degree_def 
   2.994 -	    by (cases "pnormalize s", auto)
   2.995 -	  then obtain k where kpn: "pnormalize s = [k]" by blast
   2.996 -	  from pns0[unfolded poly_zero] kpn have k: "k \<noteq>0" "poly s = poly [k]"
   2.997 -	    using poly_normalize[of s] by simp_all
   2.998 -	  let ?w = "pmult (pmult [1/k] (pexp [-a,1] (n - ?op))) (pexp r n)"
   2.999 -	  from k r s oop have "poly (pexp q n) = poly (pmult p ?w)"
  2.1000 -	    by - (rule ext, simp add: poly_mult poly_exp poly_cmult poly_add power_add[symmetric] ring_simps power_mult_distrib[symmetric])
  2.1001 -	  hence ?ths unfolding divides_def by blast}
  2.1002 -	moreover
  2.1003 -	{assume ds0: "degree s \<noteq> 0"
  2.1004 -	  from ds0 s0 dpn degree_unique[OF s, unfolded linear_pow_mul_degree] oa
  2.1005 -	    have dsn: "degree s < n" by auto 
  2.1006 -	    {fix x assume h: "poly s x = 0"
  2.1007 -	      {assume xa: "x = a"
  2.1008 -		from h[unfolded xa poly_linear_divides] sne obtain u where
  2.1009 -		  u: "s = pmult [- a, 1] u" by blast
  2.1010 -		have "poly p = poly (pmult (pexp [- a, 1] (Suc ?op)) u)"
  2.1011 -		  unfolding s u
  2.1012 -		  apply (rule ext)
  2.1013 -		  by (simp add: ring_simps power_mult_distrib[symmetric] poly_mult poly_cmult poly_add poly_exp)
  2.1014 -		with ap(2)[unfolded divides_def] have False by blast}
  2.1015 -	      note xa = this
  2.1016 -	      from h s have "poly p x = 0" by (simp add: poly_mult)
  2.1017 -	      with pq0 have "poly q x = 0" by blast
  2.1018 -	      with r xa have "poly r x = 0"
  2.1019 -		by (auto simp add: poly_mult poly_add poly_cmult eq_diff_eq[symmetric])}
  2.1020 -	    note impth = this
  2.1021 -	    from IH[rule_format, OF dsn, of s r] impth ds0
  2.1022 -	    have "s divides (pexp r (degree s))" by blast
  2.1023 -	    then obtain u where u: "poly (pexp r (degree s)) = poly (pmult s u)"
  2.1024 -	      unfolding divides_def by blast
  2.1025 -	    hence u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
  2.1026 -	      by (simp add: poly_mult[symmetric] poly_exp[symmetric])
  2.1027 -	    let ?w = "pmult (pmult u (pexp [-a,1] (n - ?op))) (pexp r (n - degree s))"
  2.1028 -	    from u' s r oop[of a] dsn have "poly (pexp q n) = poly (pmult p ?w)"
  2.1029 -	      apply - apply (rule ext)
  2.1030 -	      apply (simp only:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult ring_simps)
  2.1031 -	      
  2.1032 -	      apply (simp add:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult mult_assoc[symmetric])
  2.1033 -	      done
  2.1034 -	    hence ?ths unfolding divides_def by blast}
  2.1035 -      ultimately have ?ths by blast }
  2.1036 -      ultimately have ?ths by blast}
  2.1037 -    ultimately have ?ths using a order_root by blast}
  2.1038 -  moreover
  2.1039 -  {assume exa: "\<not> (\<exists>a. poly p a = 0)"
  2.1040 -    from fundamental_theorem_of_algebra_alt[of p] exa obtain c cs where
  2.1041 -      ccs: "c\<noteq>0" "list_all (\<lambda>c. c = 0) cs" "p = c#cs" by blast
  2.1042 -    
  2.1043 -    from poly_0[OF ccs(2)] ccs(3) 
  2.1044 -    have pp: "\<And>x. poly p x =  c" by simp
  2.1045 -    let ?w = "pmult [1/c] (pexp q n)"
  2.1046 -    from pp ccs(1) 
  2.1047 -    have "poly (pexp q n) = poly (pmult p ?w) "
  2.1048 -      apply - apply (rule ext)
  2.1049 -      unfolding poly_mult_assoc[symmetric] by (simp add: poly_mult)
  2.1050 -    hence ?ths unfolding divides_def by blast}
  2.1051 -  ultimately show ?ths by blast
  2.1052 -qed
  2.1053 -
  2.1054 -lemma nullstellensatz_univariate:
  2.1055 -  "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> 
  2.1056 -    p divides (q %^ (degree p)) \<or> (poly p = poly [] \<and> poly q = poly [])"
  2.1057 -proof-
  2.1058 -  {assume pe: "poly p = poly []"
  2.1059 -    hence eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> poly q = poly []"
  2.1060 -      apply auto
  2.1061 -      by (rule ext, simp)
  2.1062 -    {assume "p divides (pexp q (degree p))"
  2.1063 -      then obtain r where r: "poly (pexp q (degree p)) = poly (pmult p r)" 
  2.1064 -	unfolding divides_def by blast
  2.1065 -      from cong[OF r refl] pe degree_unique[OF pe]
  2.1066 -      have False by (simp add: poly_mult degree_def)}
  2.1067 -    with eq pe have ?thesis by blast}
  2.1068 -  moreover
  2.1069 -  {assume pe: "poly p \<noteq> poly []"
  2.1070 -    have p0: "poly [0] = poly []" by (rule ext, simp)
  2.1071 -    {assume dp: "degree p = 0"
  2.1072 -      then obtain k where "pnormalize p = [k]" using pe poly_normalize[of p]
  2.1073 -	unfolding degree_def by (cases "pnormalize p", auto)
  2.1074 -      hence k: "pnormalize p = [k]" "poly p = poly [k]" "k\<noteq>0"
  2.1075 -	using pe poly_normalize[of p] by (auto simp add: p0)
  2.1076 -      hence th1: "\<forall>x. poly p x \<noteq> 0" by simp
  2.1077 -      from k(2,3) dp have "poly (pexp q (degree p)) = poly (pmult p [1/k]) "
  2.1078 -	by - (rule ext, simp add: poly_mult poly_exp)
  2.1079 -      hence th2: "p divides (pexp q (degree p))" unfolding divides_def by blast
  2.1080 -      from th1 th2 pe have ?thesis by blast}
  2.1081 -    moreover
  2.1082 -    {assume dp: "degree p \<noteq> 0"
  2.1083 -      then obtain n where n: "degree p = Suc n " by (cases "degree p", auto)
  2.1084 -      {assume "p divides (pexp q (Suc n))"
  2.1085 -	then obtain u where u: "poly (pexp q (Suc n)) = poly (pmult p u)"
  2.1086 -	  unfolding divides_def by blast
  2.1087 -	hence u' :"\<And>x. poly (pexp q (Suc n)) x = poly (pmult p u) x" by simp_all
  2.1088 -	{fix x assume h: "poly p x = 0" "poly q x \<noteq> 0"
  2.1089 -	  hence "poly (pexp q (Suc n)) x \<noteq> 0" by (simp only: poly_exp) simp	  
  2.1090 -	  hence False using u' h(1) by (simp only: poly_mult poly_exp) simp}}
  2.1091 -	with n nullstellensatz_lemma[of p q "degree p"] dp 
  2.1092 -	have ?thesis by auto}
  2.1093 -    ultimately have ?thesis by blast}
  2.1094 -  ultimately show ?thesis by blast
  2.1095 -qed
  2.1096 -
  2.1097 -text{* Useful lemma *}
  2.1098 -
  2.1099 -lemma (in idom_char_0) constant_degree: "constant (poly p) \<longleftrightarrow> degree p = 0" (is "?lhs = ?rhs")
  2.1100 -proof
  2.1101 -  assume l: ?lhs
  2.1102 -  from l[unfolded constant_def, rule_format, of _ "zero"]
  2.1103 -  have th: "poly p = poly [poly p 0]" apply - by (rule ext, simp)
  2.1104 -  from degree_unique[OF th] show ?rhs by (simp add: degree_def)
  2.1105 -next
  2.1106 -  assume r: ?rhs
  2.1107 -  from r have "pnormalize p = [] \<or> (\<exists>k. pnormalize p = [k])"
  2.1108 -    unfolding degree_def by (cases "pnormalize p", auto)
  2.1109 -  then show ?lhs unfolding constant_def poly_normalize[of p, symmetric]
  2.1110 -    by (auto simp del: poly_normalize)
  2.1111 -qed
  2.1112 -
  2.1113 -(* It would be nicer to prove this without using algebraic closure...        *)
  2.1114 -
  2.1115 -lemma divides_degree_lemma: assumes dpn: "degree (p::complex list) = n"
  2.1116 -  shows "n \<le> degree (p *** q) \<or> poly (p *** q) = poly []"
  2.1117 -  using dpn
  2.1118 -proof(induct n arbitrary: p q)
  2.1119 -  case 0 thus ?case by simp
  2.1120 -next
  2.1121 -  case (Suc n p q)
  2.1122 -  from Suc.prems fundamental_theorem_of_algebra[of p] constant_degree[of p]
  2.1123 -  obtain a where a: "poly p a = 0" by auto
  2.1124 -  then obtain r where r: "p = pmult [-a, 1] r" unfolding poly_linear_divides
  2.1125 -    using Suc.prems by (auto simp add: degree_def)
  2.1126 -  {assume h: "poly (pmult r q) = poly []"
  2.1127 -    hence "poly (pmult p q) = poly []" using r
  2.1128 -      apply - apply (rule ext)  by (auto simp add: poly_entire poly_mult poly_add poly_cmult) hence ?case by blast}
  2.1129 -  moreover
  2.1130 -  {assume h: "poly (pmult r q) \<noteq> poly []" 
  2.1131 -    hence r0: "poly r \<noteq> poly []" and q0: "poly q \<noteq> poly []"
  2.1132 -      by (auto simp add: poly_entire)
  2.1133 -    have eq: "poly (pmult p q) = poly (pmult [-a, 1] (pmult r q))"
  2.1134 -      apply - apply (rule ext)
  2.1135 -      by (simp add: r poly_mult poly_add poly_cmult ring_simps)
  2.1136 -    from linear_mul_degree[OF h, of "- a"]
  2.1137 -    have dqe: "degree (pmult p q) = degree (pmult r q) + 1"
  2.1138 -      unfolding degree_unique[OF eq] .
  2.1139 -    from linear_mul_degree[OF r0, of "- a", unfolded r[symmetric]] r Suc.prems 
  2.1140 -    have dr: "degree r = n" by auto
  2.1141 -    from  Suc.hyps[OF dr, of q] have "Suc n \<le> degree (pmult p q)"
  2.1142 -      unfolding dqe using h by (auto simp del: poly.simps) 
  2.1143 -    hence ?case by blast}
  2.1144 -  ultimately show ?case by blast
  2.1145 -qed
  2.1146 -
  2.1147 -lemma divides_degree: assumes pq: "p divides (q:: complex list)"
  2.1148 -  shows "degree p \<le> degree q \<or> poly q = poly []"
  2.1149 -using pq  divides_degree_lemma[OF refl, of p]
  2.1150 -apply (auto simp add: divides_def poly_entire)
  2.1151 -apply atomize
  2.1152 -apply (erule_tac x="qa" in allE, auto)
  2.1153 -apply (subgoal_tac "degree q = degree (p *** qa)", simp)
  2.1154 -apply (rule degree_unique, simp)
  2.1155 -done
  2.1156 -
  2.1157 -(* Arithmetic operations on multivariate polynomials.                        *)
  2.1158 -
  2.1159 -lemma mpoly_base_conv: 
  2.1160 -  "(0::complex) \<equiv> poly [] x" "c \<equiv> poly [c] x" "x \<equiv> poly [0,1] x" by simp_all
  2.1161 -
  2.1162 -lemma mpoly_norm_conv: 
  2.1163 -  "poly [0] (x::complex) \<equiv> poly [] x" "poly [poly [] y] x \<equiv> poly [] x" by simp_all
  2.1164 -
  2.1165 -lemma mpoly_sub_conv: 
  2.1166 -  "poly p (x::complex) - poly q x \<equiv> poly p x + -1 * poly q x"
  2.1167 -  by (simp add: diff_def)
  2.1168 -
  2.1169 -lemma poly_pad_rule: "poly p x = 0 ==> poly (0#p) x = (0::complex)" by simp
  2.1170 -
  2.1171 -lemma poly_cancel_eq_conv: "p = (0::complex) \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (q = 0) \<equiv> (a * q - b * p = 0)" apply (atomize (full)) by auto
  2.1172 -
  2.1173 -lemma resolve_eq_raw:  "poly [] x \<equiv> 0" "poly [c] x \<equiv> (c::complex)" by auto
  2.1174 -lemma  resolve_eq_then: "(P \<Longrightarrow> (Q \<equiv> Q1)) \<Longrightarrow> (\<not>P \<Longrightarrow> (Q \<equiv> Q2))
  2.1175 -  \<Longrightarrow> Q \<equiv> P \<and> Q1 \<or> \<not>P\<and> Q2" apply (atomize (full)) by blast 
  2.1176 -lemma expand_ex_beta_conv: "list_ex P [c] \<equiv> P c" by simp
  2.1177 -
  2.1178 -lemma poly_divides_pad_rule: 
  2.1179 -  fixes p q :: "complex list"
  2.1180 -  assumes pq: "p divides q"
  2.1181 -  shows "p divides ((0::complex)#q)"
  2.1182 -proof-
  2.1183 -  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
  2.1184 -  hence "poly (0#q) = poly (p *** ([0,1] *** r))" 
  2.1185 -    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
  2.1186 -  thus ?thesis unfolding divides_def by blast
  2.1187 -qed
  2.1188 -
  2.1189 -lemma poly_divides_pad_const_rule: 
  2.1190 -  fixes p q :: "complex list"
  2.1191 -  assumes pq: "p divides q"
  2.1192 -  shows "p divides (a %* q)"
  2.1193 -proof-
  2.1194 -  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
  2.1195 -  hence "poly (a %* q) = poly (p *** (a %* r))" 
  2.1196 -    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
  2.1197 -  thus ?thesis unfolding divides_def by blast
  2.1198 -qed
  2.1199 -
  2.1200 -
  2.1201 -lemma poly_divides_conv0:  
  2.1202 -  fixes p :: "complex list"
  2.1203 -  assumes lgpq: "length q < length p" and lq:"last p \<noteq> 0"
  2.1204 -  shows "p divides q \<equiv> (\<not> (list_ex (\<lambda>c. c \<noteq> 0) q))" (is "?lhs \<equiv> ?rhs")
  2.1205 -proof-
  2.1206 -  {assume r: ?rhs 
  2.1207 -    hence eq: "poly q = poly []" unfolding poly_zero 
  2.1208 -      by (simp add: list_all_iff list_ex_iff)
  2.1209 -    hence "poly q = poly (p *** [])" by - (rule ext, simp add: poly_mult)
  2.1210 -    hence ?lhs unfolding divides_def  by blast}
  2.1211 -  moreover
  2.1212 -  {assume l: ?lhs
  2.1213 -    have ath: "\<And>lq lp dq::nat. lq < lp ==> lq \<noteq> 0 \<Longrightarrow> dq <= lq - 1 ==> dq < lp - 1"
  2.1214 -      by arith
  2.1215 -    {assume q0: "length q = 0"
  2.1216 -      hence "q = []" by simp
  2.1217 -      hence ?rhs by simp}
  2.1218 -    moreover
  2.1219 -    {assume lgq0: "length q \<noteq> 0"
  2.1220 -      from pnormalize_length[of q] have dql: "degree q \<le> length q - 1" 
  2.1221 -	unfolding degree_def by simp
  2.1222 -      from ath[OF lgpq lgq0 dql, unfolded pnormal_degree[OF lq, symmetric]] divides_degree[OF l] have "poly q = poly []" by auto
  2.1223 -      hence ?rhs unfolding poly_zero by (simp add: list_all_iff list_ex_iff)}
  2.1224 -    ultimately have ?rhs by blast }
  2.1225 -  ultimately show "?lhs \<equiv> ?rhs" by - (atomize (full), blast) 
  2.1226 -qed
  2.1227 -
  2.1228 -lemma poly_divides_conv1: 
  2.1229 -  assumes a0: "a\<noteq> (0::complex)" and pp': "(p::complex list) divides p'"
  2.1230 -  and qrp': "\<And>x. a * poly q x - poly p' x \<equiv> poly r x"
  2.1231 -  shows "p divides q \<equiv> p divides (r::complex list)" (is "?lhs \<equiv> ?rhs")
  2.1232 -proof-
  2.1233 -  {
  2.1234 -  from pp' obtain t where t: "poly p' = poly (p *** t)" 
  2.1235 -    unfolding divides_def by blast
  2.1236 -  {assume l: ?lhs
  2.1237 -    then obtain u where u: "poly q = poly (p *** u)" unfolding divides_def by blast
  2.1238 -     have "poly r = poly (p *** ((a %* u) +++ (-- t)))"
  2.1239 -       using u qrp' t
  2.1240 -       by - (rule ext, 
  2.1241 -	 simp add: poly_add poly_mult poly_cmult poly_minus ring_simps)
  2.1242 -     then have ?rhs unfolding divides_def by blast}
  2.1243 -  moreover
  2.1244 -  {assume r: ?rhs
  2.1245 -    then obtain u where u: "poly r = poly (p *** u)" unfolding divides_def by blast
  2.1246 -    from u t qrp' a0 have "poly q = poly (p *** ((1/a) %* (u +++ t)))"
  2.1247 -      by - (rule ext, atomize (full), simp add: poly_mult poly_add poly_cmult field_simps)
  2.1248 -    hence ?lhs  unfolding divides_def by blast}
  2.1249 -  ultimately have "?lhs = ?rhs" by blast }
  2.1250 -thus "?lhs \<equiv> ?rhs"  by - (atomize(full), blast) 
  2.1251 -qed
  2.1252 -
  2.1253 -lemma basic_cqe_conv1:
  2.1254 -  "(\<exists>x. poly p x = 0 \<and> poly [] x \<noteq> 0) \<equiv> False"
  2.1255 -  "(\<exists>x. poly [] x \<noteq> 0) \<equiv> False"
  2.1256 -  "(\<exists>x. poly [c] x \<noteq> 0) \<equiv> c\<noteq>0"
  2.1257 -  "(\<exists>x. poly [] x = 0) \<equiv> True"
  2.1258 -  "(\<exists>x. poly [c] x = 0) \<equiv> c = 0" by simp_all
  2.1259 -
  2.1260 -lemma basic_cqe_conv2: 
  2.1261 -  assumes l:"last (a#b#p) \<noteq> 0" 
  2.1262 -  shows "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True"
  2.1263 -proof-
  2.1264 -  {fix h t
  2.1265 -    assume h: "h\<noteq>0" "list_all (\<lambda>c. c=(0::complex)) t"  "a#b#p = h#t"
  2.1266 -    hence "list_all (\<lambda>c. c= 0) (b#p)" by simp
  2.1267 -    moreover have "last (b#p) \<in> set (b#p)" by simp
  2.1268 -    ultimately have "last (b#p) = 0" by (simp add: list_all_iff)
  2.1269 -    with l have False by simp}
  2.1270 -  hence th: "\<not> (\<exists> h t. h\<noteq>0 \<and> list_all (\<lambda>c. c=0) t \<and> a#b#p = h#t)"
  2.1271 -    by blast
  2.1272 -  from fundamental_theorem_of_algebra_alt[OF th] 
  2.1273 -  show "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True" by auto
  2.1274 -qed
  2.1275 -
  2.1276 -lemma  basic_cqe_conv_2b: "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
  2.1277 -proof-
  2.1278 -  have "\<not> (list_ex (\<lambda>c. c \<noteq> 0) p) \<longleftrightarrow> poly p = poly []" 
  2.1279 -    by (simp add: poly_zero list_all_iff list_ex_iff)
  2.1280 -  also have "\<dots> \<longleftrightarrow> (\<not> (\<exists>x. poly p x \<noteq> 0))" by (auto intro: ext)
  2.1281 -  finally show "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
  2.1282 -    by - (atomize (full), blast)
  2.1283 -qed
  2.1284 -
  2.1285 -lemma basic_cqe_conv3:
  2.1286 -  fixes p q :: "complex list"
  2.1287 -  assumes l: "last (a#p) \<noteq> 0" 
  2.1288 -  shows "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
  2.1289 -proof-
  2.1290 -  note np = pnormalize_eq[OF l]
  2.1291 -  {assume "poly (a#p) = poly []" hence False using l
  2.1292 -      unfolding poly_zero apply (auto simp add: list_all_iff del: last.simps)
  2.1293 -      apply (cases p, simp_all) done}
  2.1294 -  then have p0: "poly (a#p) \<noteq> poly []"  by blast
  2.1295 -  from np have dp:"degree (a#p) = length p" by (simp add: degree_def)
  2.1296 -  from nullstellensatz_univariate[of "a#p" q] p0 dp
  2.1297 -  show "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
  2.1298 -    by - (atomize (full), auto)
  2.1299 -qed
  2.1300 -
  2.1301 -lemma basic_cqe_conv4:
  2.1302 -  fixes p q :: "complex list"
  2.1303 -  assumes h: "\<And>x. poly (q %^ n) x \<equiv> poly r x"
  2.1304 -  shows "p divides (q %^ n) \<equiv> p divides r"
  2.1305 -proof-
  2.1306 -  from h have "poly (q %^ n) = poly r" by (auto intro: ext)  
  2.1307 -  thus "p divides (q %^ n) \<equiv> p divides r" unfolding divides_def by simp
  2.1308 -qed
  2.1309 -
  2.1310 -lemma pmult_Cons_Cons: "((a::complex)#b#p) *** q = (a %*q) +++ (0#((b#p) *** q))"
  2.1311 -  by simp
  2.1312 -
  2.1313 -lemma elim_neg_conv: "- z \<equiv> (-1) * (z::complex)" by simp
  2.1314 -lemma eqT_intr: "PROP P \<Longrightarrow> (True \<Longrightarrow> PROP P )" "PROP P \<Longrightarrow> True" by blast+
  2.1315 -lemma negate_negate_rule: "Trueprop P \<equiv> \<not> P \<equiv> False" by (atomize (full), auto)
  2.1316 -lemma last_simps: "last [x] = x" "last (x#y#ys) = last (y#ys)" by simp_all
  2.1317 -lemma length_simps: "length [] = 0" "length (x#y#xs) = length xs + 2" "length [x] = 1" by simp_all
  2.1318 -
  2.1319 -lemma complex_entire: "(z::complex) \<noteq> 0 \<and> w \<noteq> 0 \<equiv> z*w \<noteq> 0" by simp
  2.1320 -lemma resolve_eq_ne: "(P \<equiv> True) \<equiv> (\<not>P \<equiv> False)" "(P \<equiv> False) \<equiv> (\<not>P \<equiv> True)" 
  2.1321 -  by (atomize (full)) simp_all
  2.1322 -lemma cqe_conv1: "poly [] x = 0 \<longleftrightarrow> True"  by simp
  2.1323 -lemma cqe_conv2: "(p \<Longrightarrow> (q \<equiv> r)) \<equiv> ((p \<and> q) \<equiv> (p \<and> r))"  (is "?l \<equiv> ?r")
  2.1324 -proof
  2.1325 -  assume "p \<Longrightarrow> q \<equiv> r" thus "p \<and> q \<equiv> p \<and> r" apply - apply (atomize (full)) by blast
  2.1326 -next
  2.1327 -  assume "p \<and> q \<equiv> p \<and> r" "p"
  2.1328 -  thus "q \<equiv> r" apply - apply (atomize (full)) apply blast done
  2.1329 -qed
  2.1330 -lemma poly_const_conv: "poly [c] (x::complex) = y \<longleftrightarrow> c = y" by simp
  2.1331 -
  2.1332 -end
  2.1333 \ No newline at end of file
     3.1 --- a/src/HOL/Complex/README.html	Mon Dec 29 13:23:53 2008 +0100
     3.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
     3.3 @@ -1,67 +0,0 @@
     3.4 -<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
     3.5 -
     3.6 -<!-- $Id$ -->
     3.7 -
     3.8 -<HTML>
     3.9 -
    3.10 -<HEAD>
    3.11 -  <meta http-equiv="content-type" content="text/html;charset=iso-8859-1">
    3.12 -  <TITLE>HOL/Complex/README</TITLE>
    3.13 -</HEAD>
    3.14 -
    3.15 -<BODY>
    3.16 -
    3.17 -<H1>Complex: The Complex Numbers</H1>
    3.18 -		<P>This directory defines the type <KBD>complex</KBD> of the complex numbers,
    3.19 -with numeric constants and some complex analysis.  The development includes
    3.20 -nonstandard analysis for the complex numbers.  Note that the image
    3.21 -<KBD>HOL-Complex</KBD> includes theories from the directories 
    3.22 -<KBD><a href="#Anchor-Real">HOL/Real</a></KBD>  and <KBD><a href="#Anchor-Hyperreal">HOL/Hyperreal</a></KBD>. They define other types including <kbd>real</kbd> (the real numbers) and <kbd>hypreal</kbd> (the hyperreal or non-standard reals).
    3.23 -
    3.24 -<ul>
    3.25 -<li><a href="CLim.html">CLim</a> Limits, continuous functions, and derivatives for the complex numbers
    3.26 -<li><a href="CSeries.html">CSeries</a> Finite summation and infinite series for the complex numbers
    3.27 -<li><a href="CStar.html">CStar</a> Star-transforms for the complex numbers, to form non-standard extensions of sets and functions
    3.28 -<li><a href="Complex.html">Complex</a> The complex numbers
    3.29 -<li><a href="NSCA.html">NSCA</a> Nonstandard complex analysis
    3.30 -<li><a href="NSComplex.html">NSComplex</a> Ultrapower construction of the nonstandard complex numbers
    3.31 -</ul>
    3.32 -
    3.33 -<h2><a name="Anchor-Real" id="Anchor-Real"></a>Real: Dedekind Cut Construction of the Real Line</h2>
    3.34 -
    3.35 -<ul>
    3.36 -<li><a href="Lubs.html">Lubs</a> Definition of upper bounds, lubs and so on, to support completeness proofs.
    3.37 -<li><a href="PReal.html">PReal</a> The positive reals constructed using Dedekind cuts
    3.38 -<li><a href="Rational.html">Rational</a> The rational numbers constructed as equivalence classes of integers
    3.39 -<li><a href="RComplete.html">RComplete</a> The reals are complete: they satisfy the supremum property. They also have the Archimedean property.
    3.40 -<li><a href="RealDef.html">RealDef</a> The real numbers, their ordering properties, and embedding of the integers and the natural numbers
    3.41 -<li><a href="RealPow.html">RealPow</a> Real numbers raised to natural number powers
    3.42 -</ul>
    3.43 -<h2><a name="Anchor-Hyperreal" id="Anchor-Hyperreal"></a>Hyperreal: Ultrafilter Construction of the Non-Standard Reals</h2>
    3.44 -See J. D. Fleuriot and L. C. Paulson. Mechanizing Nonstandard Real Analysis. LMS J. Computation and Mathematics 3 (2000), 140-190.
    3.45 -<ul>
    3.46 -<li><a href="Filter.html">Filter</a> Theory of Filters and Ultrafilters. Main result is a version of the Ultrafilter Theorem proved using Zorn's Lemma.
    3.47 -<li><a href="HLog.html">HLog</a> Non-standard logarithms
    3.48 -<li><a href="HSeries.html">HSeries</a> Non-standard theory of finite summation and infinite series
    3.49 -<li><a href="HTranscendental.html">HTranscendental</a> Non-standard extensions of transcendental functions
    3.50 -<li><a href="HyperDef.html">HyperDef</a> Ultrapower construction of the hyperreals
    3.51 -<li><a href="HyperNat.html">HyperNat</a> Ultrapower construction of the hypernaturals
    3.52 -<li><a href="HyperPow.html">HyperPow</a> Powers theory for the hyperreals
    3.53 -<!-- <li><a href="IntFloor.html">IntFloor</a> Floor and Ceiling functions relating the reals and integers -->
    3.54 -<li><a href="Integration.html">Integration</a> Gage integrals
    3.55 -<li><a href="Lim.html">Lim</a> Theory of limits, continuous functions, and derivatives
    3.56 -<li><a href="Log.html">Log</a> Logarithms for the reals
    3.57 -<li><a href="MacLaurin.html">MacLaurin</a> MacLaurin series
    3.58 -<li><a href="NatStar.html">NatStar</a> Star-transforms for the hypernaturals, to form non-standard extensions of sets and functions involving the naturals or reals
    3.59 -<li><a href="NthRoot.html">NthRoot</a> Existence of n-th roots of real numbers
    3.60 -<li><a href="NSA.html">NSA</a> Theory defining sets of infinite numbers, infinitesimals, the infinitely close relation, and their various algebraic properties.
    3.61 -<li><a href="Poly.html">Poly</a> Univariate real polynomials
    3.62 -<li><a href="SEQ.html">SEQ</a> Convergence of sequences and series using standard and nonstandard analysis
    3.63 -<li><a href="Series.html">Series</a> Finite summation and infinite series for the reals
    3.64 -<li><a href="Star.html">Star</a> Nonstandard extensions of real sets and real functions
    3.65 -<li><a href="Transcendental.html">Transcendental</a> Power series and transcendental functions
    3.66 -</ul>
    3.67 -<HR>
    3.68 -<P>Last modified $Date$
    3.69 -</BODY>
    3.70 -</HTML>
     4.1 --- a/src/HOL/Complex/document/root.tex	Mon Dec 29 13:23:53 2008 +0100
     4.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
     4.3 @@ -1,32 +0,0 @@
     4.4 -
     4.5 -% $Id$
     4.6 -
     4.7 -\documentclass[11pt,a4paper]{article}
     4.8 -\usepackage{graphicx,isabelle,isabellesym,latexsym}
     4.9 -\usepackage[latin1]{inputenc}
    4.10 -\usepackage{pdfsetup}
    4.11 -
    4.12 -\urlstyle{rm}
    4.13 -\isabellestyle{it}
    4.14 -\pagestyle{myheadings}
    4.15 -
    4.16 -\begin{document}
    4.17 -
    4.18 -\title{Isabelle/HOL-Complex --- Higher-Order Logic with Complex Numbers}
    4.19 -\maketitle
    4.20 -
    4.21 -\tableofcontents
    4.22 -
    4.23 -\begin{center}
    4.24 -  \includegraphics[width=\textwidth,height=\textheight,keepaspectratio]{session_graph}
    4.25 -\end{center}
    4.26 -
    4.27 -\newpage
    4.28 -
    4.29 -\renewcommand{\isamarkupheader}[1]%
    4.30 -{\section{\isabellecontext: #1}\markright{THEORY~``\isabellecontext''}}
    4.31 -
    4.32 -\parindent 0pt\parskip 0.5ex
    4.33 -\input{session}
    4.34 -
    4.35 -\end{document}
     5.1 --- a/src/HOL/Complex_Main.thy	Mon Dec 29 13:23:53 2008 +0100
     5.2 +++ b/src/HOL/Complex_Main.thy	Mon Dec 29 14:08:08 2008 +0100
     5.3 @@ -9,7 +9,7 @@
     5.4  imports
     5.5    Main
     5.6    Real
     5.7 -  "~~/src/HOL/Complex/Fundamental_Theorem_Algebra"
     5.8 +  Fundamental_Theorem_Algebra
     5.9    Log
    5.10    Ln
    5.11    Taylor
     6.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     6.2 +++ b/src/HOL/Dense_Linear_Order.thy	Mon Dec 29 14:08:08 2008 +0100
     6.3 @@ -0,0 +1,877 @@
     6.4 +(* Author: Amine Chaieb, TU Muenchen *)
     6.5 +
     6.6 +header {* Dense linear order without endpoints
     6.7 +  and a quantifier elimination procedure in Ferrante and Rackoff style *}
     6.8 +
     6.9 +theory Dense_Linear_Order
    6.10 +imports Plain Groebner_Basis
    6.11 +uses
    6.12 +  "~~/src/HOL/Tools/Qelim/langford_data.ML"
    6.13 +  "~~/src/HOL/Tools/Qelim/ferrante_rackoff_data.ML"
    6.14 +  ("~~/src/HOL/Tools/Qelim/langford.ML")
    6.15 +  ("~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML")
    6.16 +begin
    6.17 +
    6.18 +setup {* Langford_Data.setup #> Ferrante_Rackoff_Data.setup *}
    6.19 +
    6.20 +context linorder
    6.21 +begin
    6.22 +
    6.23 +lemma less_not_permute: "\<not> (x < y \<and> y < x)" by (simp add: not_less linear)
    6.24 +
    6.25 +lemma gather_simps: 
    6.26 +  shows 
    6.27 +  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y) \<and> P x)"
    6.28 +  and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y) \<and> P x)"
    6.29 +  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y))"
    6.30 +  and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y))"  by auto
    6.31 +
    6.32 +lemma 
    6.33 +  gather_start: "(\<exists>x. P x) \<equiv> (\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y\<in> {}. x < y) \<and> P x)" 
    6.34 +  by simp
    6.35 +
    6.36 +text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"}*}
    6.37 +lemma minf_lt:  "\<exists>z . \<forall>x. x < z \<longrightarrow> (x < t \<longleftrightarrow> True)" by auto
    6.38 +lemma minf_gt: "\<exists>z . \<forall>x. x < z \<longrightarrow>  (t < x \<longleftrightarrow>  False)"
    6.39 +  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
    6.40 +
    6.41 +lemma minf_le: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<le> t \<longleftrightarrow> True)" by (auto simp add: less_le)
    6.42 +lemma minf_ge: "\<exists>z. \<forall>x. x < z \<longrightarrow> (t \<le> x \<longleftrightarrow> False)"
    6.43 +  by (auto simp add: less_le not_less not_le)
    6.44 +lemma minf_eq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
    6.45 +lemma minf_neq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
    6.46 +lemma minf_P: "\<exists>z. \<forall>x. x < z \<longrightarrow> (P \<longleftrightarrow> P)" by blast
    6.47 +
    6.48 +text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"}*}
    6.49 +lemma pinf_gt:  "\<exists>z . \<forall>x. z < x \<longrightarrow> (t < x \<longleftrightarrow> True)" by auto
    6.50 +lemma pinf_lt: "\<exists>z . \<forall>x. z < x \<longrightarrow>  (x < t \<longleftrightarrow>  False)"
    6.51 +  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
    6.52 +
    6.53 +lemma pinf_ge: "\<exists>z. \<forall>x. z < x \<longrightarrow> (t \<le> x \<longleftrightarrow> True)" by (auto simp add: less_le)
    6.54 +lemma pinf_le: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<le> t \<longleftrightarrow> False)"
    6.55 +  by (auto simp add: less_le not_less not_le)
    6.56 +lemma pinf_eq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
    6.57 +lemma pinf_neq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
    6.58 +lemma pinf_P: "\<exists>z. \<forall>x. z < x \<longrightarrow> (P \<longleftrightarrow> P)" by blast
    6.59 +
    6.60 +lemma nmi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x < t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
    6.61 +lemma nmi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t < x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)"
    6.62 +  by (auto simp add: le_less)
    6.63 +lemma  nmi_le: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x\<le> t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
    6.64 +lemma  nmi_ge: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t\<le> x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
    6.65 +lemma  nmi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x = t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
    6.66 +lemma  nmi_neq: "t \<in> U \<Longrightarrow>\<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
    6.67 +lemma  nmi_P: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
    6.68 +lemma  nmi_conj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x) ;
    6.69 +  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
    6.70 +  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
    6.71 +lemma  nmi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x) ;
    6.72 +  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
    6.73 +  \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
    6.74 +
    6.75 +lemma  npi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x < t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by (auto simp add: le_less)
    6.76 +lemma  npi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t < x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
    6.77 +lemma  npi_le: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x \<le> t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
    6.78 +lemma  npi_ge: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t \<le> x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
    6.79 +lemma  npi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x = t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
    6.80 +lemma  npi_neq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u )" by auto
    6.81 +lemma  npi_P: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
    6.82 +lemma  npi_conj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u) ;  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)\<rbrakk>
    6.83 +  \<Longrightarrow>  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
    6.84 +lemma  npi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u) ; \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)\<rbrakk>
    6.85 +  \<Longrightarrow> \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
    6.86 +
    6.87 +lemma lin_dense_lt: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t < u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x < t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y < t)"
    6.88 +proof(clarsimp)
    6.89 +  fix x l u y  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x"
    6.90 +    and xu: "x<u"  and px: "x < t" and ly: "l<y" and yu:"y < u"
    6.91 +  from tU noU ly yu have tny: "t\<noteq>y" by auto
    6.92 +  {assume H: "t < y"
    6.93 +    from less_trans[OF lx px] less_trans[OF H yu]
    6.94 +    have "l < t \<and> t < u"  by simp
    6.95 +    with tU noU have "False" by auto}
    6.96 +  hence "\<not> t < y"  by auto hence "y \<le> t" by (simp add: not_less)
    6.97 +  thus "y < t" using tny by (simp add: less_le)
    6.98 +qed
    6.99 +
   6.100 +lemma lin_dense_gt: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l < x \<and> x < u \<and> t < x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t < y)"
   6.101 +proof(clarsimp)
   6.102 +  fix x l u y
   6.103 +  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
   6.104 +  and px: "t < x" and ly: "l<y" and yu:"y < u"
   6.105 +  from tU noU ly yu have tny: "t\<noteq>y" by auto
   6.106 +  {assume H: "y< t"
   6.107 +    from less_trans[OF ly H] less_trans[OF px xu] have "l < t \<and> t < u" by simp
   6.108 +    with tU noU have "False" by auto}
   6.109 +  hence "\<not> y<t"  by auto hence "t \<le> y" by (auto simp add: not_less)
   6.110 +  thus "t < y" using tny by (simp add:less_le)
   6.111 +qed
   6.112 +
   6.113 +lemma lin_dense_le: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<le> t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<le> t)"
   6.114 +proof(clarsimp)
   6.115 +  fix x l u y
   6.116 +  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
   6.117 +  and px: "x \<le> t" and ly: "l<y" and yu:"y < u"
   6.118 +  from tU noU ly yu have tny: "t\<noteq>y" by auto
   6.119 +  {assume H: "t < y"
   6.120 +    from less_le_trans[OF lx px] less_trans[OF H yu]
   6.121 +    have "l < t \<and> t < u" by simp
   6.122 +    with tU noU have "False" by auto}
   6.123 +  hence "\<not> t < y"  by auto thus "y \<le> t" by (simp add: not_less)
   6.124 +qed
   6.125 +
   6.126 +lemma lin_dense_ge: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> t \<le> x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t \<le> y)"
   6.127 +proof(clarsimp)
   6.128 +  fix x l u y
   6.129 +  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
   6.130 +  and px: "t \<le> x" and ly: "l<y" and yu:"y < u"
   6.131 +  from tU noU ly yu have tny: "t\<noteq>y" by auto
   6.132 +  {assume H: "y< t"
   6.133 +    from less_trans[OF ly H] le_less_trans[OF px xu]
   6.134 +    have "l < t \<and> t < u" by simp
   6.135 +    with tU noU have "False" by auto}
   6.136 +  hence "\<not> y<t"  by auto thus "t \<le> y" by (simp add: not_less)
   6.137 +qed
   6.138 +lemma lin_dense_eq: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x = t   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y= t)"  by auto
   6.139 +lemma lin_dense_neq: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<noteq> t   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<noteq> t)"  by auto
   6.140 +lemma lin_dense_P: "\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P)"  by auto
   6.141 +
   6.142 +lemma lin_dense_conj:
   6.143 +  "\<lbrakk>\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P1 x
   6.144 +  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
   6.145 +  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
   6.146 +  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
   6.147 +  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> (P1 x \<and> P2 x)
   6.148 +  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<and> P2 y))"
   6.149 +  by blast
   6.150 +lemma lin_dense_disj:
   6.151 +  "\<lbrakk>\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P1 x
   6.152 +  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
   6.153 +  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
   6.154 +  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
   6.155 +  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> (P1 x \<or> P2 x)
   6.156 +  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<or> P2 y))"
   6.157 +  by blast
   6.158 +
   6.159 +lemma npmibnd: "\<lbrakk>\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<le> x); \<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)\<rbrakk>
   6.160 +  \<Longrightarrow> \<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<le> x \<and> x \<le> u')"
   6.161 +by auto
   6.162 +
   6.163 +lemma finite_set_intervals:
   6.164 +  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
   6.165 +  and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
   6.166 +  shows "\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a \<le> x \<and> x \<le> b \<and> P x"
   6.167 +proof-
   6.168 +  let ?Mx = "{y. y\<in> S \<and> y \<le> x}"
   6.169 +  let ?xM = "{y. y\<in> S \<and> x \<le> y}"
   6.170 +  let ?a = "Max ?Mx"
   6.171 +  let ?b = "Min ?xM"
   6.172 +  have MxS: "?Mx \<subseteq> S" by blast
   6.173 +  hence fMx: "finite ?Mx" using fS finite_subset by auto
   6.174 +  from lx linS have linMx: "l \<in> ?Mx" by blast
   6.175 +  hence Mxne: "?Mx \<noteq> {}" by blast
   6.176 +  have xMS: "?xM \<subseteq> S" by blast
   6.177 +  hence fxM: "finite ?xM" using fS finite_subset by auto
   6.178 +  from xu uinS have linxM: "u \<in> ?xM" by blast
   6.179 +  hence xMne: "?xM \<noteq> {}" by blast
   6.180 +  have ax:"?a \<le> x" using Mxne fMx by auto
   6.181 +  have xb:"x \<le> ?b" using xMne fxM by auto
   6.182 +  have "?a \<in> ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \<in> S" using MxS by blast
   6.183 +  have "?b \<in> ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \<in> S" using xMS by blast
   6.184 +  have noy:"\<forall> y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S"
   6.185 +  proof(clarsimp)
   6.186 +    fix y   assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S"
   6.187 +    from yS have "y\<in> ?Mx \<or> y\<in> ?xM" by (auto simp add: linear)
   6.188 +    moreover {assume "y \<in> ?Mx" hence "y \<le> ?a" using Mxne fMx by auto with ay have "False" by (simp add: not_le[symmetric])}
   6.189 +    moreover {assume "y \<in> ?xM" hence "?b \<le> y" using xMne fxM by auto with yb have "False" by (simp add: not_le[symmetric])}
   6.190 +    ultimately show "False" by blast
   6.191 +  qed
   6.192 +  from ainS binS noy ax xb px show ?thesis by blast
   6.193 +qed
   6.194 +
   6.195 +lemma finite_set_intervals2:
   6.196 +  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
   6.197 +  and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
   6.198 +  shows "(\<exists> s\<in> S. P s) \<or> (\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a < x \<and> x < b \<and> P x)"
   6.199 +proof-
   6.200 +  from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su]
   6.201 +  obtain a and b where
   6.202 +    as: "a\<in> S" and bs: "b\<in> S" and noS:"\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S"
   6.203 +    and axb: "a \<le> x \<and> x \<le> b \<and> P x"  by auto
   6.204 +  from axb have "x= a \<or> x= b \<or> (a < x \<and> x < b)" by (auto simp add: le_less)
   6.205 +  thus ?thesis using px as bs noS by blast
   6.206 +qed
   6.207 +
   6.208 +end
   6.209 +
   6.210 +section {* The classical QE after Langford for dense linear orders *}
   6.211 +
   6.212 +context dense_linear_order
   6.213 +begin
   6.214 +
   6.215 +lemma interval_empty_iff:
   6.216 +  "{y. x < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
   6.217 +  by (auto dest: dense)
   6.218 +
   6.219 +lemma dlo_qe_bnds: 
   6.220 +  assumes ne: "L \<noteq> {}" and neU: "U \<noteq> {}" and fL: "finite L" and fU: "finite U"
   6.221 +  shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> (\<forall> l \<in> L. \<forall>u \<in> U. l < u)"
   6.222 +proof (simp only: atomize_eq, rule iffI)
   6.223 +  assume H: "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)"
   6.224 +  then obtain x where xL: "\<forall>y\<in>L. y < x" and xU: "\<forall>y\<in>U. x < y" by blast
   6.225 +  {fix l u assume l: "l \<in> L" and u: "u \<in> U"
   6.226 +    have "l < x" using xL l by blast
   6.227 +    also have "x < u" using xU u by blast
   6.228 +    finally (less_trans) have "l < u" .}
   6.229 +  thus "\<forall>l\<in>L. \<forall>u\<in>U. l < u" by blast
   6.230 +next
   6.231 +  assume H: "\<forall>l\<in>L. \<forall>u\<in>U. l < u"
   6.232 +  let ?ML = "Max L"
   6.233 +  let ?MU = "Min U"  
   6.234 +  from fL ne have th1: "?ML \<in> L" and th1': "\<forall>l\<in>L. l \<le> ?ML" by auto
   6.235 +  from fU neU have th2: "?MU \<in> U" and th2': "\<forall>u\<in>U. ?MU \<le> u" by auto
   6.236 +  from th1 th2 H have "?ML < ?MU" by auto
   6.237 +  with dense obtain w where th3: "?ML < w" and th4: "w < ?MU" by blast
   6.238 +  from th3 th1' have "\<forall>l \<in> L. l < w" by auto
   6.239 +  moreover from th4 th2' have "\<forall>u \<in> U. w < u" by auto
   6.240 +  ultimately show "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)" by auto
   6.241 +qed
   6.242 +
   6.243 +lemma dlo_qe_noub: 
   6.244 +  assumes ne: "L \<noteq> {}" and fL: "finite L"
   6.245 +  shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> {}. x < y)) \<equiv> True"
   6.246 +proof(simp add: atomize_eq)
   6.247 +  from gt_ex[of "Max L"] obtain M where M: "Max L < M" by blast
   6.248 +  from ne fL have "\<forall>x \<in> L. x \<le> Max L" by simp
   6.249 +  with M have "\<forall>x\<in>L. x < M" by (auto intro: le_less_trans)
   6.250 +  thus "\<exists>x. \<forall>y\<in>L. y < x" by blast
   6.251 +qed
   6.252 +
   6.253 +lemma dlo_qe_nolb: 
   6.254 +  assumes ne: "U \<noteq> {}" and fU: "finite U"
   6.255 +  shows "(\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> True"
   6.256 +proof(simp add: atomize_eq)
   6.257 +  from lt_ex[of "Min U"] obtain M where M: "M < Min U" by blast
   6.258 +  from ne fU have "\<forall>x \<in> U. Min U \<le> x" by simp
   6.259 +  with M have "\<forall>x\<in>U. M < x" by (auto intro: less_le_trans)
   6.260 +  thus "\<exists>x. \<forall>y\<in>U. x < y" by blast
   6.261 +qed
   6.262 +
   6.263 +lemma exists_neq: "\<exists>(x::'a). x \<noteq> t" "\<exists>(x::'a). t \<noteq> x" 
   6.264 +  using gt_ex[of t] by auto
   6.265 +
   6.266 +lemmas dlo_simps = order_refl less_irrefl not_less not_le exists_neq 
   6.267 +  le_less neq_iff linear less_not_permute
   6.268 +
   6.269 +lemma axiom: "dense_linear_order (op \<le>) (op <)" by (rule dense_linear_order_axioms)
   6.270 +lemma atoms:
   6.271 +  shows "TERM (less :: 'a \<Rightarrow> _)"
   6.272 +    and "TERM (less_eq :: 'a \<Rightarrow> _)"
   6.273 +    and "TERM (op = :: 'a \<Rightarrow> _)" .
   6.274 +
   6.275 +declare axiom[langford qe: dlo_qe_bnds dlo_qe_nolb dlo_qe_noub gather: gather_start gather_simps atoms: atoms]
   6.276 +declare dlo_simps[langfordsimp]
   6.277 +
   6.278 +end
   6.279 +
   6.280 +(* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *)
   6.281 +lemma dnf:
   6.282 +  "(P & (Q | R)) = ((P&Q) | (P&R))" 
   6.283 +  "((Q | R) & P) = ((Q&P) | (R&P))"
   6.284 +  by blast+
   6.285 +
   6.286 +lemmas weak_dnf_simps = simp_thms dnf
   6.287 +
   6.288 +lemma nnf_simps:
   6.289 +    "(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)" "(P \<longrightarrow> Q) = (\<not>P \<or> Q)"
   6.290 +    "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
   6.291 +  by blast+
   6.292 +
   6.293 +lemma ex_distrib: "(\<exists>x. P x \<or> Q x) \<longleftrightarrow> ((\<exists>x. P x) \<or> (\<exists>x. Q x))" by blast
   6.294 +
   6.295 +lemmas dnf_simps = weak_dnf_simps nnf_simps ex_distrib
   6.296 +
   6.297 +use "~~/src/HOL/Tools/Qelim/langford.ML"
   6.298 +method_setup dlo = {*
   6.299 +  Method.ctxt_args (Method.SIMPLE_METHOD' o LangfordQE.dlo_tac)
   6.300 +*} "Langford's algorithm for quantifier elimination in dense linear orders"
   6.301 +
   6.302 +
   6.303 +section {* Contructive dense linear orders yield QE for linear arithmetic over ordered Fields -- see @{text "Arith_Tools.thy"} *}
   6.304 +
   6.305 +text {* Linear order without upper bounds *}
   6.306 +
   6.307 +locale linorder_stupid_syntax = linorder
   6.308 +begin
   6.309 +notation
   6.310 +  less_eq  ("op \<sqsubseteq>") and
   6.311 +  less_eq  ("(_/ \<sqsubseteq> _)" [51, 51] 50) and
   6.312 +  less  ("op \<sqsubset>") and
   6.313 +  less  ("(_/ \<sqsubset> _)"  [51, 51] 50)
   6.314 +
   6.315 +end
   6.316 +
   6.317 +locale linorder_no_ub = linorder_stupid_syntax +
   6.318 +  assumes gt_ex: "\<exists>y. less x y"
   6.319 +begin
   6.320 +lemma ge_ex: "\<exists>y. x \<sqsubseteq> y" using gt_ex by auto
   6.321 +
   6.322 +text {* Theorems for @{text "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"} *}
   6.323 +lemma pinf_conj:
   6.324 +  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
   6.325 +  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
   6.326 +  shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
   6.327 +proof-
   6.328 +  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
   6.329 +     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
   6.330 +  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
   6.331 +  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
   6.332 +  {fix x assume H: "z \<sqsubset> x"
   6.333 +    from less_trans[OF zz1 H] less_trans[OF zz2 H]
   6.334 +    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
   6.335 +  }
   6.336 +  thus ?thesis by blast
   6.337 +qed
   6.338 +
   6.339 +lemma pinf_disj:
   6.340 +  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
   6.341 +  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
   6.342 +  shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
   6.343 +proof-
   6.344 +  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
   6.345 +     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
   6.346 +  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
   6.347 +  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
   6.348 +  {fix x assume H: "z \<sqsubset> x"
   6.349 +    from less_trans[OF zz1 H] less_trans[OF zz2 H]
   6.350 +    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
   6.351 +  }
   6.352 +  thus ?thesis by blast
   6.353 +qed
   6.354 +
   6.355 +lemma pinf_ex: assumes ex:"\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x"
   6.356 +proof-
   6.357 +  from ex obtain z where z: "\<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
   6.358 +  from gt_ex obtain x where x: "z \<sqsubset> x" by blast
   6.359 +  from z x p1 show ?thesis by blast
   6.360 +qed
   6.361 +
   6.362 +end
   6.363 +
   6.364 +text {* Linear order without upper bounds *}
   6.365 +
   6.366 +locale linorder_no_lb = linorder_stupid_syntax +
   6.367 +  assumes lt_ex: "\<exists>y. less y x"
   6.368 +begin
   6.369 +lemma le_ex: "\<exists>y. y \<sqsubseteq> x" using lt_ex by auto
   6.370 +
   6.371 +
   6.372 +text {* Theorems for @{text "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"} *}
   6.373 +lemma minf_conj:
   6.374 +  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
   6.375 +  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
   6.376 +  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
   6.377 +proof-
   6.378 +  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
   6.379 +  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
   6.380 +  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
   6.381 +  {fix x assume H: "x \<sqsubset> z"
   6.382 +    from less_trans[OF H zz1] less_trans[OF H zz2]
   6.383 +    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
   6.384 +  }
   6.385 +  thus ?thesis by blast
   6.386 +qed
   6.387 +
   6.388 +lemma minf_disj:
   6.389 +  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
   6.390 +  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
   6.391 +  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
   6.392 +proof-
   6.393 +  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
   6.394 +  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
   6.395 +  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
   6.396 +  {fix x assume H: "x \<sqsubset> z"
   6.397 +    from less_trans[OF H zz1] less_trans[OF H zz2]
   6.398 +    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
   6.399 +  }
   6.400 +  thus ?thesis by blast
   6.401 +qed
   6.402 +
   6.403 +lemma minf_ex: assumes ex:"\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x"
   6.404 +proof-
   6.405 +  from ex obtain z where z: "\<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
   6.406 +  from lt_ex obtain x where x: "x \<sqsubset> z" by blast
   6.407 +  from z x p1 show ?thesis by blast
   6.408 +qed
   6.409 +
   6.410 +end
   6.411 +
   6.412 +
   6.413 +locale constr_dense_linear_order = linorder_no_lb + linorder_no_ub +
   6.414 +  fixes between
   6.415 +  assumes between_less: "less x y \<Longrightarrow> less x (between x y) \<and> less (between x y) y"
   6.416 +     and  between_same: "between x x = x"
   6.417 +
   6.418 +interpretation  constr_dense_linear_order < dense_linear_order 
   6.419 +  apply unfold_locales
   6.420 +  using gt_ex lt_ex between_less
   6.421 +    by (auto, rule_tac x="between x y" in exI, simp)
   6.422 +
   6.423 +context  constr_dense_linear_order
   6.424 +begin
   6.425 +
   6.426 +lemma rinf_U:
   6.427 +  assumes fU: "finite U"
   6.428 +  and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
   6.429 +  \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
   6.430 +  and nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')"
   6.431 +  and nmi: "\<not> MP"  and npi: "\<not> PP"  and ex: "\<exists> x.  P x"
   6.432 +  shows "\<exists> u\<in> U. \<exists> u' \<in> U. P (between u u')"
   6.433 +proof-
   6.434 +  from ex obtain x where px: "P x" by blast
   6.435 +  from px nmi npi nmpiU have "\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u'" by auto
   6.436 +  then obtain u and u' where uU:"u\<in> U" and uU': "u' \<in> U" and ux:"u \<sqsubseteq> x" and xu':"x \<sqsubseteq> u'" by auto
   6.437 +  from uU have Une: "U \<noteq> {}" by auto
   6.438 +  term "linorder.Min less_eq"
   6.439 +  let ?l = "linorder.Min less_eq U"
   6.440 +  let ?u = "linorder.Max less_eq U"
   6.441 +  have linM: "?l \<in> U" using fU Une by simp
   6.442 +  have uinM: "?u \<in> U" using fU Une by simp
   6.443 +  have lM: "\<forall> t\<in> U. ?l \<sqsubseteq> t" using Une fU by auto
   6.444 +  have Mu: "\<forall> t\<in> U. t \<sqsubseteq> ?u" using Une fU by auto
   6.445 +  have th:"?l \<sqsubseteq> u" using uU Une lM by auto
   6.446 +  from order_trans[OF th ux] have lx: "?l \<sqsubseteq> x" .
   6.447 +  have th: "u' \<sqsubseteq> ?u" using uU' Une Mu by simp
   6.448 +  from order_trans[OF xu' th] have xu: "x \<sqsubseteq> ?u" .
   6.449 +  from finite_set_intervals2[where P="P",OF px lx xu linM uinM fU lM Mu]
   6.450 +  have "(\<exists> s\<in> U. P s) \<or>
   6.451 +      (\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x)" .
   6.452 +  moreover { fix u assume um: "u\<in>U" and pu: "P u"
   6.453 +    have "between u u = u" by (simp add: between_same)
   6.454 +    with um pu have "P (between u u)" by simp
   6.455 +    with um have ?thesis by blast}
   6.456 +  moreover{
   6.457 +    assume "\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x"
   6.458 +      then obtain t1 and t2 where t1M: "t1 \<in> U" and t2M: "t2\<in> U"
   6.459 +        and noM: "\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U" and t1x: "t1 \<sqsubset> x" and xt2: "x \<sqsubset> t2" and px: "P x"
   6.460 +        by blast
   6.461 +      from less_trans[OF t1x xt2] have t1t2: "t1 \<sqsubset> t2" .
   6.462 +      let ?u = "between t1 t2"
   6.463 +      from between_less t1t2 have t1lu: "t1 \<sqsubset> ?u" and ut2: "?u \<sqsubset> t2" by auto
   6.464 +      from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast
   6.465 +      with t1M t2M have ?thesis by blast}
   6.466 +    ultimately show ?thesis by blast
   6.467 +  qed
   6.468 +
   6.469 +theorem fr_eq:
   6.470 +  assumes fU: "finite U"
   6.471 +  and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
   6.472 +   \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
   6.473 +  and nmibnd: "\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<sqsubseteq> x)"
   6.474 +  and npibnd: "\<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<sqsubseteq> u)"
   6.475 +  and mi: "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x = MP)"  and pi: "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x = PP)"
   6.476 +  shows "(\<exists> x. P x) \<equiv> (MP \<or> PP \<or> (\<exists> u \<in> U. \<exists> u'\<in> U. P (between u u')))"
   6.477 +  (is "_ \<equiv> (_ \<or> _ \<or> ?F)" is "?E \<equiv> ?D")
   6.478 +proof-
   6.479 + {
   6.480 +   assume px: "\<exists> x. P x"
   6.481 +   have "MP \<or> PP \<or> (\<not> MP \<and> \<not> PP)" by blast
   6.482 +   moreover {assume "MP \<or> PP" hence "?D" by blast}
   6.483 +   moreover {assume nmi: "\<not> MP" and npi: "\<not> PP"
   6.484 +     from npmibnd[OF nmibnd npibnd]
   6.485 +     have nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')" .
   6.486 +     from rinf_U[OF fU lin_dense nmpiU nmi npi px] have "?D" by blast}
   6.487 +   ultimately have "?D" by blast}
   6.488 + moreover
   6.489 + { assume "?D"
   6.490 +   moreover {assume m:"MP" from minf_ex[OF mi m] have "?E" .}
   6.491 +   moreover {assume p: "PP" from pinf_ex[OF pi p] have "?E" . }
   6.492 +   moreover {assume f:"?F" hence "?E" by blast}
   6.493 +   ultimately have "?E" by blast}
   6.494 + ultimately have "?E = ?D" by blast thus "?E \<equiv> ?D" by simp
   6.495 +qed
   6.496 +
   6.497 +lemmas minf_thms = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P
   6.498 +lemmas pinf_thms = pinf_conj pinf_disj pinf_eq pinf_neq pinf_lt pinf_le pinf_gt pinf_ge pinf_P
   6.499 +
   6.500 +lemmas nmi_thms = nmi_conj nmi_disj nmi_eq nmi_neq nmi_lt nmi_le nmi_gt nmi_ge nmi_P
   6.501 +lemmas npi_thms = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P
   6.502 +lemmas lin_dense_thms = lin_dense_conj lin_dense_disj lin_dense_eq lin_dense_neq lin_dense_lt lin_dense_le lin_dense_gt lin_dense_ge lin_dense_P
   6.503 +
   6.504 +lemma ferrack_axiom: "constr_dense_linear_order less_eq less between"
   6.505 +  by (rule constr_dense_linear_order_axioms)
   6.506 +lemma atoms:
   6.507 +  shows "TERM (less :: 'a \<Rightarrow> _)"
   6.508 +    and "TERM (less_eq :: 'a \<Rightarrow> _)"
   6.509 +    and "TERM (op = :: 'a \<Rightarrow> _)" .
   6.510 +
   6.511 +declare ferrack_axiom [ferrack minf: minf_thms pinf: pinf_thms
   6.512 +    nmi: nmi_thms npi: npi_thms lindense:
   6.513 +    lin_dense_thms qe: fr_eq atoms: atoms]
   6.514 +
   6.515 +declaration {*
   6.516 +let
   6.517 +fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
   6.518 +fun generic_whatis phi =
   6.519 + let
   6.520 +  val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
   6.521 +  fun h x t =
   6.522 +   case term_of t of
   6.523 +     Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
   6.524 +                            else Ferrante_Rackoff_Data.Nox
   6.525 +   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
   6.526 +                            else Ferrante_Rackoff_Data.Nox
   6.527 +   | b$y$z => if Term.could_unify (b, lt) then
   6.528 +                 if term_of x aconv y then Ferrante_Rackoff_Data.Lt
   6.529 +                 else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
   6.530 +                 else Ferrante_Rackoff_Data.Nox
   6.531 +             else if Term.could_unify (b, le) then
   6.532 +                 if term_of x aconv y then Ferrante_Rackoff_Data.Le
   6.533 +                 else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
   6.534 +                 else Ferrante_Rackoff_Data.Nox
   6.535 +             else Ferrante_Rackoff_Data.Nox
   6.536 +   | _ => Ferrante_Rackoff_Data.Nox
   6.537 + in h end
   6.538 + fun ss phi = HOL_ss addsimps (simps phi)
   6.539 +in
   6.540 + Ferrante_Rackoff_Data.funs  @{thm "ferrack_axiom"}
   6.541 +  {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss}
   6.542 +end
   6.543 +*}
   6.544 +
   6.545 +end
   6.546 +
   6.547 +use "~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML"
   6.548 +
   6.549 +method_setup ferrack = {*
   6.550 +  Method.ctxt_args (Method.SIMPLE_METHOD' o FerranteRackoff.dlo_tac)
   6.551 +*} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders"
   6.552 +
   6.553 +subsection {* Ferrante and Rackoff algorithm over ordered fields *}
   6.554 +
   6.555 +lemma neg_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x < 0) == (x > 0))"
   6.556 +proof-
   6.557 +  assume H: "c < 0"
   6.558 +  have "c*x < 0 = (0/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
   6.559 +  also have "\<dots> = (0 < x)" by simp
   6.560 +  finally show  "(c*x < 0) == (x > 0)" by simp
   6.561 +qed
   6.562 +
   6.563 +lemma pos_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x < 0) == (x < 0))"
   6.564 +proof-
   6.565 +  assume H: "c > 0"
   6.566 +  hence "c*x < 0 = (0/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
   6.567 +  also have "\<dots> = (0 > x)" by simp
   6.568 +  finally show  "(c*x < 0) == (x < 0)" by simp
   6.569 +qed
   6.570 +
   6.571 +lemma neg_prod_sum_lt: "(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x + t< 0) == (x > (- 1/c)*t))"
   6.572 +proof-
   6.573 +  assume H: "c < 0"
   6.574 +  have "c*x + t< 0 = (c*x < -t)" by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
   6.575 +  also have "\<dots> = (-t/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
   6.576 +  also have "\<dots> = ((- 1/c)*t < x)" by simp
   6.577 +  finally show  "(c*x + t < 0) == (x > (- 1/c)*t)" by simp
   6.578 +qed
   6.579 +
   6.580 +lemma pos_prod_sum_lt:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x + t < 0) == (x < (- 1/c)*t))"
   6.581 +proof-
   6.582 +  assume H: "c > 0"
   6.583 +  have "c*x + t< 0 = (c*x < -t)"  by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
   6.584 +  also have "\<dots> = (-t/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
   6.585 +  also have "\<dots> = ((- 1/c)*t > x)" by simp
   6.586 +  finally show  "(c*x + t < 0) == (x < (- 1/c)*t)" by simp
   6.587 +qed
   6.588 +
   6.589 +lemma sum_lt:"((x::'a::pordered_ab_group_add) + t < 0) == (x < - t)"
   6.590 +  using less_diff_eq[where a= x and b=t and c=0] by simp
   6.591 +
   6.592 +lemma neg_prod_le:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x <= 0) == (x >= 0))"
   6.593 +proof-
   6.594 +  assume H: "c < 0"
   6.595 +  have "c*x <= 0 = (0/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
   6.596 +  also have "\<dots> = (0 <= x)" by simp
   6.597 +  finally show  "(c*x <= 0) == (x >= 0)" by simp
   6.598 +qed
   6.599 +
   6.600 +lemma pos_prod_le:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x <= 0) == (x <= 0))"
   6.601 +proof-
   6.602 +  assume H: "c > 0"
   6.603 +  hence "c*x <= 0 = (0/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
   6.604 +  also have "\<dots> = (0 >= x)" by simp
   6.605 +  finally show  "(c*x <= 0) == (x <= 0)" by simp
   6.606 +qed
   6.607 +
   6.608 +lemma neg_prod_sum_le: "(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x + t <= 0) == (x >= (- 1/c)*t))"
   6.609 +proof-
   6.610 +  assume H: "c < 0"
   6.611 +  have "c*x + t <= 0 = (c*x <= -t)"  by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
   6.612 +  also have "\<dots> = (-t/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
   6.613 +  also have "\<dots> = ((- 1/c)*t <= x)" by simp
   6.614 +  finally show  "(c*x + t <= 0) == (x >= (- 1/c)*t)" by simp
   6.615 +qed
   6.616 +
   6.617 +lemma pos_prod_sum_le:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x + t <= 0) == (x <= (- 1/c)*t))"
   6.618 +proof-
   6.619 +  assume H: "c > 0"
   6.620 +  have "c*x + t <= 0 = (c*x <= -t)" by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
   6.621 +  also have "\<dots> = (-t/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
   6.622 +  also have "\<dots> = ((- 1/c)*t >= x)" by simp
   6.623 +  finally show  "(c*x + t <= 0) == (x <= (- 1/c)*t)" by simp
   6.624 +qed
   6.625 +
   6.626 +lemma sum_le:"((x::'a::pordered_ab_group_add) + t <= 0) == (x <= - t)"
   6.627 +  using le_diff_eq[where a= x and b=t and c=0] by simp
   6.628 +
   6.629 +lemma nz_prod_eq:"(c\<Colon>'a\<Colon>ordered_field) \<noteq> 0 \<Longrightarrow> ((c*x = 0) == (x = 0))" by simp
   6.630 +lemma nz_prod_sum_eq: "(c\<Colon>'a\<Colon>ordered_field) \<noteq> 0 \<Longrightarrow> ((c*x + t = 0) == (x = (- 1/c)*t))"
   6.631 +proof-
   6.632 +  assume H: "c \<noteq> 0"
   6.633 +  have "c*x + t = 0 = (c*x = -t)" by (subst eq_iff_diff_eq_0 [of "c*x" "-t"], simp)
   6.634 +  also have "\<dots> = (x = -t/c)" by (simp only: nonzero_eq_divide_eq[OF H] ring_simps)
   6.635 +  finally show  "(c*x + t = 0) == (x = (- 1/c)*t)" by simp
   6.636 +qed
   6.637 +lemma sum_eq:"((x::'a::pordered_ab_group_add) + t = 0) == (x = - t)"
   6.638 +  using eq_diff_eq[where a= x and b=t and c=0] by simp
   6.639 +
   6.640 +
   6.641 +interpretation class_ordered_field_dense_linear_order: constr_dense_linear_order
   6.642 + ["op <=" "op <"
   6.643 +   "\<lambda> x y. 1/2 * ((x::'a::{ordered_field,recpower,number_ring}) + y)"]
   6.644 +proof (unfold_locales, dlo, dlo, auto)
   6.645 +  fix x y::'a assume lt: "x < y"
   6.646 +  from  less_half_sum[OF lt] show "x < (x + y) /2" by simp
   6.647 +next
   6.648 +  fix x y::'a assume lt: "x < y"
   6.649 +  from  gt_half_sum[OF lt] show "(x + y) /2 < y" by simp
   6.650 +qed
   6.651 +
   6.652 +declaration{*
   6.653 +let
   6.654 +fun earlier [] x y = false
   6.655 +        | earlier (h::t) x y =
   6.656 +    if h aconvc y then false else if h aconvc x then true else earlier t x y;
   6.657 +
   6.658 +fun dest_frac ct = case term_of ct of
   6.659 +   Const (@{const_name "HOL.divide"},_) $ a $ b=>
   6.660 +    Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b))
   6.661 + | t => Rat.rat_of_int (snd (HOLogic.dest_number t))
   6.662 +
   6.663 +fun mk_frac phi cT x =
   6.664 + let val (a, b) = Rat.quotient_of_rat x
   6.665 + in if b = 1 then Numeral.mk_cnumber cT a
   6.666 +    else Thm.capply
   6.667 +         (Thm.capply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"})
   6.668 +                     (Numeral.mk_cnumber cT a))
   6.669 +         (Numeral.mk_cnumber cT b)
   6.670 + end
   6.671 +
   6.672 +fun whatis x ct = case term_of ct of
   6.673 +  Const(@{const_name "HOL.plus"}, _)$(Const(@{const_name "HOL.times"},_)$_$y)$_ =>
   6.674 +     if y aconv term_of x then ("c*x+t",[(funpow 2 Thm.dest_arg1) ct, Thm.dest_arg ct])
   6.675 +     else ("Nox",[])
   6.676 +| Const(@{const_name "HOL.plus"}, _)$y$_ =>
   6.677 +     if y aconv term_of x then ("x+t",[Thm.dest_arg ct])
   6.678 +     else ("Nox",[])
   6.679 +| Const(@{const_name "HOL.times"}, _)$_$y =>
   6.680 +     if y aconv term_of x then ("c*x",[Thm.dest_arg1 ct])
   6.681 +     else ("Nox",[])
   6.682 +| t => if t aconv term_of x then ("x",[]) else ("Nox",[]);
   6.683 +
   6.684 +fun xnormalize_conv ctxt [] ct = reflexive ct
   6.685 +| xnormalize_conv ctxt (vs as (x::_)) ct =
   6.686 +   case term_of ct of
   6.687 +   Const(@{const_name HOL.less},_)$_$Const(@{const_name "HOL.zero"},_) =>
   6.688 +    (case whatis x (Thm.dest_arg1 ct) of
   6.689 +    ("c*x+t",[c,t]) =>
   6.690 +       let
   6.691 +        val cr = dest_frac c
   6.692 +        val clt = Thm.dest_fun2 ct
   6.693 +        val cz = Thm.dest_arg ct
   6.694 +        val neg = cr </ Rat.zero
   6.695 +        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
   6.696 +               (Thm.capply @{cterm "Trueprop"}
   6.697 +                  (if neg then Thm.capply (Thm.capply clt c) cz
   6.698 +                    else Thm.capply (Thm.capply clt cz) c))
   6.699 +        val cth = equal_elim (symmetric cthp) TrueI
   6.700 +        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x,t])
   6.701 +             (if neg then @{thm neg_prod_sum_lt} else @{thm pos_prod_sum_lt})) cth
   6.702 +        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
   6.703 +                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
   6.704 +      in rth end
   6.705 +    | ("x+t",[t]) =>
   6.706 +       let
   6.707 +        val T = ctyp_of_term x
   6.708 +        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_lt"}
   6.709 +        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
   6.710 +              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
   6.711 +       in  rth end
   6.712 +    | ("c*x",[c]) =>
   6.713 +       let
   6.714 +        val cr = dest_frac c
   6.715 +        val clt = Thm.dest_fun2 ct
   6.716 +        val cz = Thm.dest_arg ct
   6.717 +        val neg = cr </ Rat.zero
   6.718 +        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
   6.719 +               (Thm.capply @{cterm "Trueprop"}
   6.720 +                  (if neg then Thm.capply (Thm.capply clt c) cz
   6.721 +                    else Thm.capply (Thm.capply clt cz) c))
   6.722 +        val cth = equal_elim (symmetric cthp) TrueI
   6.723 +        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
   6.724 +             (if neg then @{thm neg_prod_lt} else @{thm pos_prod_lt})) cth
   6.725 +        val rth = th
   6.726 +      in rth end
   6.727 +    | _ => reflexive ct)
   6.728 +
   6.729 +
   6.730 +|  Const(@{const_name HOL.less_eq},_)$_$Const(@{const_name "HOL.zero"},_) =>
   6.731 +   (case whatis x (Thm.dest_arg1 ct) of
   6.732 +    ("c*x+t",[c,t]) =>
   6.733 +       let
   6.734 +        val T = ctyp_of_term x
   6.735 +        val cr = dest_frac c
   6.736 +        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
   6.737 +        val cz = Thm.dest_arg ct
   6.738 +        val neg = cr </ Rat.zero
   6.739 +        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
   6.740 +               (Thm.capply @{cterm "Trueprop"}
   6.741 +                  (if neg then Thm.capply (Thm.capply clt c) cz
   6.742 +                    else Thm.capply (Thm.capply clt cz) c))
   6.743 +        val cth = equal_elim (symmetric cthp) TrueI
   6.744 +        val th = implies_elim (instantiate' [SOME T] (map SOME [c,x,t])
   6.745 +             (if neg then @{thm neg_prod_sum_le} else @{thm pos_prod_sum_le})) cth
   6.746 +        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
   6.747 +                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
   6.748 +      in rth end
   6.749 +    | ("x+t",[t]) =>
   6.750 +       let
   6.751 +        val T = ctyp_of_term x
   6.752 +        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_le"}
   6.753 +        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
   6.754 +              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
   6.755 +       in  rth end
   6.756 +    | ("c*x",[c]) =>
   6.757 +       let
   6.758 +        val T = ctyp_of_term x
   6.759 +        val cr = dest_frac c
   6.760 +        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
   6.761 +        val cz = Thm.dest_arg ct
   6.762 +        val neg = cr </ Rat.zero
   6.763 +        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
   6.764 +               (Thm.capply @{cterm "Trueprop"}
   6.765 +                  (if neg then Thm.capply (Thm.capply clt c) cz
   6.766 +                    else Thm.capply (Thm.capply clt cz) c))
   6.767 +        val cth = equal_elim (symmetric cthp) TrueI
   6.768 +        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
   6.769 +             (if neg then @{thm neg_prod_le} else @{thm pos_prod_le})) cth
   6.770 +        val rth = th
   6.771 +      in rth end
   6.772 +    | _ => reflexive ct)
   6.773 +
   6.774 +|  Const("op =",_)$_$Const(@{const_name "HOL.zero"},_) =>
   6.775 +   (case whatis x (Thm.dest_arg1 ct) of
   6.776 +    ("c*x+t",[c,t]) =>
   6.777 +       let
   6.778 +        val T = ctyp_of_term x
   6.779 +        val cr = dest_frac c
   6.780 +        val ceq = Thm.dest_fun2 ct
   6.781 +        val cz = Thm.dest_arg ct
   6.782 +        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
   6.783 +            (Thm.capply @{cterm "Trueprop"}
   6.784 +             (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
   6.785 +        val cth = equal_elim (symmetric cthp) TrueI
   6.786 +        val th = implies_elim
   6.787 +                 (instantiate' [SOME T] (map SOME [c,x,t]) @{thm nz_prod_sum_eq}) cth
   6.788 +        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
   6.789 +                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
   6.790 +      in rth end
   6.791 +    | ("x+t",[t]) =>
   6.792 +       let
   6.793 +        val T = ctyp_of_term x
   6.794 +        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_eq"}
   6.795 +        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
   6.796 +              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
   6.797 +       in  rth end
   6.798 +    | ("c*x",[c]) =>
   6.799 +       let
   6.800 +        val T = ctyp_of_term x
   6.801 +        val cr = dest_frac c
   6.802 +        val ceq = Thm.dest_fun2 ct
   6.803 +        val cz = Thm.dest_arg ct
   6.804 +        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
   6.805 +            (Thm.capply @{cterm "Trueprop"}
   6.806 +             (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
   6.807 +        val cth = equal_elim (symmetric cthp) TrueI
   6.808 +        val rth = implies_elim
   6.809 +                 (instantiate' [SOME T] (map SOME [c,x]) @{thm nz_prod_eq}) cth
   6.810 +      in rth end
   6.811 +    | _ => reflexive ct);
   6.812 +
   6.813 +local
   6.814 +  val less_iff_diff_less_0 = mk_meta_eq @{thm "less_iff_diff_less_0"}
   6.815 +  val le_iff_diff_le_0 = mk_meta_eq @{thm "le_iff_diff_le_0"}
   6.816 +  val eq_iff_diff_eq_0 = mk_meta_eq @{thm "eq_iff_diff_eq_0"}
   6.817 +in
   6.818 +fun field_isolate_conv phi ctxt vs ct = case term_of ct of
   6.819 +  Const(@{const_name HOL.less},_)$a$b =>
   6.820 +   let val (ca,cb) = Thm.dest_binop ct
   6.821 +       val T = ctyp_of_term ca
   6.822 +       val th = instantiate' [SOME T] [SOME ca, SOME cb] less_iff_diff_less_0
   6.823 +       val nth = Conv.fconv_rule
   6.824 +         (Conv.arg_conv (Conv.arg1_conv
   6.825 +              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
   6.826 +       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
   6.827 +   in rth end
   6.828 +| Const(@{const_name HOL.less_eq},_)$a$b =>
   6.829 +   let val (ca,cb) = Thm.dest_binop ct
   6.830 +       val T = ctyp_of_term ca
   6.831 +       val th = instantiate' [SOME T] [SOME ca, SOME cb] le_iff_diff_le_0
   6.832 +       val nth = Conv.fconv_rule
   6.833 +         (Conv.arg_conv (Conv.arg1_conv
   6.834 +              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
   6.835 +       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
   6.836 +   in rth end
   6.837 +
   6.838 +| Const("op =",_)$a$b =>
   6.839 +   let val (ca,cb) = Thm.dest_binop ct
   6.840 +       val T = ctyp_of_term ca
   6.841 +       val th = instantiate' [SOME T] [SOME ca, SOME cb] eq_iff_diff_eq_0
   6.842 +       val nth = Conv.fconv_rule
   6.843 +         (Conv.arg_conv (Conv.arg1_conv
   6.844 +              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
   6.845 +       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
   6.846 +   in rth end
   6.847 +| @{term "Not"} $(Const("op =",_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
   6.848 +| _ => reflexive ct
   6.849 +end;
   6.850 +
   6.851 +fun classfield_whatis phi =
   6.852 + let
   6.853 +  fun h x t =
   6.854 +   case term_of t of
   6.855 +     Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
   6.856 +                            else Ferrante_Rackoff_Data.Nox
   6.857 +   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
   6.858 +                            else Ferrante_Rackoff_Data.Nox
   6.859 +   | Const(@{const_name HOL.less},_)$y$z =>
   6.860 +       if term_of x aconv y then Ferrante_Rackoff_Data.Lt
   6.861 +        else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
   6.862 +        else Ferrante_Rackoff_Data.Nox
   6.863 +   | Const (@{const_name HOL.less_eq},_)$y$z =>
   6.864 +         if term_of x aconv y then Ferrante_Rackoff_Data.Le
   6.865 +         else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
   6.866 +         else Ferrante_Rackoff_Data.Nox
   6.867 +   | _ => Ferrante_Rackoff_Data.Nox
   6.868 + in h end;
   6.869 +fun class_field_ss phi =
   6.870 +   HOL_basic_ss addsimps ([@{thm "linorder_not_less"}, @{thm "linorder_not_le"}])
   6.871 +   addsplits [@{thm "abs_split"},@{thm "split_max"}, @{thm "split_min"}]
   6.872 +
   6.873 +in
   6.874 +Ferrante_Rackoff_Data.funs @{thm "class_ordered_field_dense_linear_order.ferrack_axiom"}
   6.875 +  {isolate_conv = field_isolate_conv, whatis = classfield_whatis, simpset = class_field_ss}
   6.876 +end
   6.877 +*}
   6.878 +
   6.879 +
   6.880 +end 
     7.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     7.2 +++ b/src/HOL/Fundamental_Theorem_Algebra.thy	Mon Dec 29 14:08:08 2008 +0100
     7.3 @@ -0,0 +1,1327 @@
     7.4 +(* Author: Amine Chaieb, TU Muenchen *)
     7.5 +
     7.6 +header{*Fundamental Theorem of Algebra*}
     7.7 +
     7.8 +theory Fundamental_Theorem_Algebra
     7.9 +imports Univ_Poly Dense_Linear_Order Complex
    7.10 +begin
    7.11 +
    7.12 +subsection {* Square root of complex numbers *}
    7.13 +definition csqrt :: "complex \<Rightarrow> complex" where
    7.14 +"csqrt z = (if Im z = 0 then
    7.15 +            if 0 \<le> Re z then Complex (sqrt(Re z)) 0
    7.16 +            else Complex 0 (sqrt(- Re z))
    7.17 +           else Complex (sqrt((cmod z + Re z) /2))
    7.18 +                        ((Im z / abs(Im z)) * sqrt((cmod z - Re z) /2)))"
    7.19 +
    7.20 +lemma csqrt[algebra]: "csqrt z ^ 2 = z"
    7.21 +proof-
    7.22 +  obtain x y where xy: "z = Complex x y" by (cases z, simp_all)
    7.23 +  {assume y0: "y = 0"
    7.24 +    {assume x0: "x \<ge> 0" 
    7.25 +      then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
    7.26 +	by (simp add: csqrt_def power2_eq_square)}
    7.27 +    moreover
    7.28 +    {assume "\<not> x \<ge> 0" hence x0: "- x \<ge> 0" by arith
    7.29 +      then have ?thesis using y0 xy real_sqrt_pow2[OF x0] 
    7.30 +	by (simp add: csqrt_def power2_eq_square) }
    7.31 +    ultimately have ?thesis by blast}
    7.32 +  moreover
    7.33 +  {assume y0: "y\<noteq>0"
    7.34 +    {fix x y
    7.35 +      let ?z = "Complex x y"
    7.36 +      from abs_Re_le_cmod[of ?z] have tha: "abs x \<le> cmod ?z" by auto
    7.37 +      hence "cmod ?z - x \<ge> 0" "cmod ?z + x \<ge> 0" by arith+ 
    7.38 +      hence "(sqrt (x * x + y * y) + x) / 2 \<ge> 0" "(sqrt (x * x + y * y) - x) / 2 \<ge> 0" by (simp_all add: power2_eq_square) }
    7.39 +    note th = this
    7.40 +    have sq4: "\<And>x::real. x^2 / 4 = (x / 2) ^ 2" 
    7.41 +      by (simp add: power2_eq_square) 
    7.42 +    from th[of x y]
    7.43 +    have sq4': "sqrt (((sqrt (x * x + y * y) + x)^2 / 4)) = (sqrt (x * x + y * y) + x) / 2" "sqrt (((sqrt (x * x + y * y) - x)^2 / 4)) = (sqrt (x * x + y * y) - x) / 2" unfolding sq4 by simp_all
    7.44 +    then have th1: "sqrt ((sqrt (x * x + y * y) + x) * (sqrt (x * x + y * y) + x) / 4) - sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) - x) / 4) = x"
    7.45 +      unfolding power2_eq_square by simp 
    7.46 +    have "sqrt 4 = sqrt (2^2)" by simp 
    7.47 +    hence sqrt4: "sqrt 4 = 2" by (simp only: real_sqrt_abs)
    7.48 +    have th2: "2 *(y * sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) + x) / 4)) / \<bar>y\<bar> = y"
    7.49 +      using iffD2[OF real_sqrt_pow2_iff sum_power2_ge_zero[of x y]] y0
    7.50 +      unfolding power2_eq_square 
    7.51 +      by (simp add: ring_simps real_sqrt_divide sqrt4)
    7.52 +     from y0 xy have ?thesis  apply (simp add: csqrt_def power2_eq_square)
    7.53 +       apply (simp add: real_sqrt_sum_squares_mult_ge_zero[of x y] real_sqrt_pow2[OF th(1)[of x y], unfolded power2_eq_square] real_sqrt_pow2[OF th(2)[of x y], unfolded power2_eq_square] real_sqrt_mult[symmetric])
    7.54 +      using th1 th2  ..}
    7.55 +  ultimately show ?thesis by blast
    7.56 +qed
    7.57 +
    7.58 +
    7.59 +subsection{* More lemmas about module of complex numbers *}
    7.60 +
    7.61 +lemma complex_of_real_power: "complex_of_real x ^ n = complex_of_real (x^n)"
    7.62 +  by (rule of_real_power [symmetric])
    7.63 +
    7.64 +lemma real_down2: "(0::real) < d1 \<Longrightarrow> 0 < d2 ==> EX e. 0 < e & e < d1 & e < d2"
    7.65 +  apply ferrack apply arith done
    7.66 +
    7.67 +text{* The triangle inequality for cmod *}
    7.68 +lemma complex_mod_triangle_sub: "cmod w \<le> cmod (w + z) + norm z"
    7.69 +  using complex_mod_triangle_ineq2[of "w + z" "-z"] by auto
    7.70 +
    7.71 +subsection{* Basic lemmas about complex polynomials *}
    7.72 +
    7.73 +lemma poly_bound_exists:
    7.74 +  shows "\<exists>m. m > 0 \<and> (\<forall>z. cmod z <= r \<longrightarrow> cmod (poly p z) \<le> m)"
    7.75 +proof(induct p)
    7.76 +  case Nil thus ?case by (rule exI[where x=1], simp) 
    7.77 +next
    7.78 +  case (Cons c cs)
    7.79 +  from Cons.hyps obtain m where m: "\<forall>z. cmod z \<le> r \<longrightarrow> cmod (poly cs z) \<le> m"
    7.80 +    by blast
    7.81 +  let ?k = " 1 + cmod c + \<bar>r * m\<bar>"
    7.82 +  have kp: "?k > 0" using abs_ge_zero[of "r*m"] norm_ge_zero[of c] by arith
    7.83 +  {fix z
    7.84 +    assume H: "cmod z \<le> r"
    7.85 +    from m H have th: "cmod (poly cs z) \<le> m" by blast
    7.86 +    from H have rp: "r \<ge> 0" using norm_ge_zero[of z] by arith
    7.87 +    have "cmod (poly (c # cs) z) \<le> cmod c + cmod (z* poly cs z)"
    7.88 +      using norm_triangle_ineq[of c "z* poly cs z"] by simp
    7.89 +    also have "\<dots> \<le> cmod c + r*m" using mult_mono[OF H th rp norm_ge_zero[of "poly cs z"]] by (simp add: norm_mult)
    7.90 +    also have "\<dots> \<le> ?k" by simp
    7.91 +    finally have "cmod (poly (c # cs) z) \<le> ?k" .}
    7.92 +  with kp show ?case by blast
    7.93 +qed
    7.94 +
    7.95 +
    7.96 +text{* Offsetting the variable in a polynomial gives another of same degree *}
    7.97 +  (* FIXME : Lemma holds also in locale --- fix it later *)
    7.98 +lemma  poly_offset_lemma:
    7.99 +  shows "\<exists>b q. (length q = length p) \<and> (\<forall>x. poly (b#q) (x::complex) = (a + x) * poly p x)"
   7.100 +proof(induct p)
   7.101 +  case Nil thus ?case by simp
   7.102 +next
   7.103 +  case (Cons c cs)
   7.104 +  from Cons.hyps obtain b q where 
   7.105 +    bq: "length q = length cs" "\<forall>x. poly (b # q) x = (a + x) * poly cs x"
   7.106 +    by blast
   7.107 +  let ?b = "a*c"
   7.108 +  let ?q = "(b+c)#q"
   7.109 +  have lg: "length ?q = length (c#cs)" using bq(1) by simp
   7.110 +  {fix x
   7.111 +    from bq(2)[rule_format, of x]
   7.112 +    have "x*poly (b # q) x = x*((a + x) * poly cs x)" by simp
   7.113 +    hence "poly (?b# ?q) x = (a + x) * poly (c # cs) x"
   7.114 +      by (simp add: ring_simps)}
   7.115 +  with lg  show ?case by blast 
   7.116 +qed
   7.117 +
   7.118 +    (* FIXME : This one too*)
   7.119 +lemma poly_offset: "\<exists> q. length q = length p \<and> (\<forall>x. poly q (x::complex) = poly p (a + x))"
   7.120 +proof (induct p)
   7.121 +  case Nil thus ?case by simp
   7.122 +next
   7.123 +  case (Cons c cs)
   7.124 +  from Cons.hyps obtain q where q: "length q = length cs" "\<forall>x. poly q x = poly cs (a + x)" by blast
   7.125 +  from poly_offset_lemma[of q a] obtain b p where 
   7.126 +    bp: "length p = length q" "\<forall>x. poly (b # p) x = (a + x) * poly q x"
   7.127 +    by blast
   7.128 +  thus ?case using q bp by - (rule exI[where x="(c + b)#p"], simp)
   7.129 +qed
   7.130 +
   7.131 +text{* An alternative useful formulation of completeness of the reals *}
   7.132 +lemma real_sup_exists: assumes ex: "\<exists>x. P x" and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z"
   7.133 +  shows "\<exists>(s::real). \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s"
   7.134 +proof-
   7.135 +  from ex bz obtain x Y where x: "P x" and Y: "\<And>x. P x \<Longrightarrow> x < Y"  by blast
   7.136 +  from ex have thx:"\<exists>x. x \<in> Collect P" by blast
   7.137 +  from bz have thY: "\<exists>Y. isUb UNIV (Collect P) Y" 
   7.138 +    by(auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def order_le_less)
   7.139 +  from reals_complete[OF thx thY] obtain L where L: "isLub UNIV (Collect P) L"
   7.140 +    by blast
   7.141 +  from Y[OF x] have xY: "x < Y" .
   7.142 +  from L have L': "\<forall>x. P x \<longrightarrow> x \<le> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)  
   7.143 +  from Y have Y': "\<forall>x. P x \<longrightarrow> x \<le> Y" 
   7.144 +    apply (clarsimp, atomize (full)) by auto 
   7.145 +  from L Y' have "L \<le> Y" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
   7.146 +  {fix y
   7.147 +    {fix z assume z: "P z" "y < z"
   7.148 +      from L' z have "y < L" by auto }
   7.149 +    moreover
   7.150 +    {assume yL: "y < L" "\<forall>z. P z \<longrightarrow> \<not> y < z"
   7.151 +      hence nox: "\<forall>z. P z \<longrightarrow> y \<ge> z" by auto
   7.152 +      from nox L have "y \<ge> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def) 
   7.153 +      with yL(1) have False  by arith}
   7.154 +    ultimately have "(\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < L" by blast}
   7.155 +  thus ?thesis by blast
   7.156 +qed
   7.157 +
   7.158 +
   7.159 +subsection{* Some theorems about Sequences*}
   7.160 +text{* Given a binary function @{text "f:: nat \<Rightarrow> 'a \<Rightarrow> 'a"}, its values are uniquely determined by a function g *}
   7.161 +
   7.162 +lemma num_Axiom: "EX! g. g 0 = e \<and> (\<forall>n. g (Suc n) = f n (g n))"
   7.163 +  unfolding Ex1_def
   7.164 +  apply (rule_tac x="nat_rec e f" in exI)
   7.165 +  apply (rule conjI)+
   7.166 +apply (rule def_nat_rec_0, simp)
   7.167 +apply (rule allI, rule def_nat_rec_Suc, simp)
   7.168 +apply (rule allI, rule impI, rule ext)
   7.169 +apply (erule conjE)
   7.170 +apply (induct_tac x)
   7.171 +apply (simp add: nat_rec_0)
   7.172 +apply (erule_tac x="n" in allE)
   7.173 +apply (simp)
   7.174 +done
   7.175 +
   7.176 + text{* An equivalent formulation of monotony -- Not used here, but might be useful *}
   7.177 +lemma mono_Suc: "mono f = (\<forall>n. (f n :: 'a :: order) \<le> f (Suc n))"
   7.178 +unfolding mono_def
   7.179 +proof auto
   7.180 +  fix A B :: nat
   7.181 +  assume H: "\<forall>n. f n \<le> f (Suc n)" "A \<le> B"
   7.182 +  hence "\<exists>k. B = A + k" apply -  apply (thin_tac "\<forall>n. f n \<le> f (Suc n)") 
   7.183 +    by presburger
   7.184 +  then obtain k where k: "B = A + k" by blast
   7.185 +  {fix a k
   7.186 +    have "f a \<le> f (a + k)"
   7.187 +    proof (induct k)
   7.188 +      case 0 thus ?case by simp
   7.189 +    next
   7.190 +      case (Suc k)
   7.191 +      from Suc.hyps H(1)[rule_format, of "a + k"] show ?case by simp
   7.192 +    qed}
   7.193 +  with k show "f A \<le> f B" by blast
   7.194 +qed
   7.195 +
   7.196 +text{* for any sequence, there is a mootonic subsequence *}
   7.197 +lemma seq_monosub: "\<exists>f. subseq f \<and> monoseq (\<lambda> n. (s (f n)))"
   7.198 +proof-
   7.199 +  {assume H: "\<forall>n. \<exists>p >n. \<forall> m\<ge>p. s m \<le> s p"
   7.200 +    let ?P = "\<lambda> p n. p > n \<and> (\<forall>m \<ge> p. s m \<le> s p)"
   7.201 +    from num_Axiom[of "SOME p. ?P p 0" "\<lambda>p n. SOME p. ?P p n"]
   7.202 +    obtain f where f: "f 0 = (SOME p. ?P p 0)" "\<forall>n. f (Suc n) = (SOME p. ?P p (f n))" by blast
   7.203 +    have "?P (f 0) 0"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p 0"]
   7.204 +      using H apply - 
   7.205 +      apply (erule allE[where x=0], erule exE, rule_tac x="p" in exI) 
   7.206 +      unfolding order_le_less by blast 
   7.207 +    hence f0: "f 0 > 0" "\<forall>m \<ge> f 0. s m \<le> s (f 0)" by blast+
   7.208 +    {fix n
   7.209 +      have "?P (f (Suc n)) (f n)" 
   7.210 +	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
   7.211 +	using H apply - 
   7.212 +      apply (erule allE[where x="f n"], erule exE, rule_tac x="p" in exI) 
   7.213 +      unfolding order_le_less by blast 
   7.214 +    hence "f (Suc n) > f n" "\<forall>m \<ge> f (Suc n). s m \<le> s (f (Suc n))" by blast+}
   7.215 +  note fSuc = this
   7.216 +    {fix p q assume pq: "p \<ge> f q"
   7.217 +      have "s p \<le> s(f(q))"  using f0(2)[rule_format, of p] pq fSuc
   7.218 +	by (cases q, simp_all) }
   7.219 +    note pqth = this
   7.220 +    {fix q
   7.221 +      have "f (Suc q) > f q" apply (induct q) 
   7.222 +	using f0(1) fSuc(1)[of 0] apply simp by (rule fSuc(1))}
   7.223 +    note fss = this
   7.224 +    from fss have th1: "subseq f" unfolding subseq_Suc_iff ..
   7.225 +    {fix a b 
   7.226 +      have "f a \<le> f (a + b)"
   7.227 +      proof(induct b)
   7.228 +	case 0 thus ?case by simp
   7.229 +      next
   7.230 +	case (Suc b)
   7.231 +	from fSuc(1)[of "a + b"] Suc.hyps show ?case by simp
   7.232 +      qed}
   7.233 +    note fmon0 = this
   7.234 +    have "monoseq (\<lambda>n. s (f n))" 
   7.235 +    proof-
   7.236 +      {fix n
   7.237 +	have "s (f n) \<ge> s (f (Suc n))" 
   7.238 +	proof(cases n)
   7.239 +	  case 0
   7.240 +	  assume n0: "n = 0"
   7.241 +	  from fSuc(1)[of 0] have th0: "f 0 \<le> f (Suc 0)" by simp
   7.242 +	  from f0(2)[rule_format, OF th0] show ?thesis  using n0 by simp
   7.243 +	next
   7.244 +	  case (Suc m)
   7.245 +	  assume m: "n = Suc m"
   7.246 +	  from fSuc(1)[of n] m have th0: "f (Suc m) \<le> f (Suc (Suc m))" by simp
   7.247 +	  from m fSuc(2)[rule_format, OF th0] show ?thesis by simp 
   7.248 +	qed}
   7.249 +      thus "monoseq (\<lambda>n. s (f n))" unfolding monoseq_Suc by blast 
   7.250 +    qed
   7.251 +    with th1 have ?thesis by blast}
   7.252 +  moreover
   7.253 +  {fix N assume N: "\<forall>p >N. \<exists> m\<ge>p. s m > s p"
   7.254 +    {fix p assume p: "p \<ge> Suc N" 
   7.255 +      hence pN: "p > N" by arith with N obtain m where m: "m \<ge> p" "s m > s p" by blast
   7.256 +      have "m \<noteq> p" using m(2) by auto 
   7.257 +      with m have "\<exists>m>p. s p < s m" by - (rule exI[where x=m], auto)}
   7.258 +    note th0 = this
   7.259 +    let ?P = "\<lambda>m x. m > x \<and> s x < s m"
   7.260 +    from num_Axiom[of "SOME x. ?P x (Suc N)" "\<lambda>m x. SOME y. ?P y x"]
   7.261 +    obtain f where f: "f 0 = (SOME x. ?P x (Suc N))" 
   7.262 +      "\<forall>n. f (Suc n) = (SOME m. ?P m (f n))" by blast
   7.263 +    have "?P (f 0) (Suc N)"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p (Suc N)"]
   7.264 +      using N apply - 
   7.265 +      apply (erule allE[where x="Suc N"], clarsimp)
   7.266 +      apply (rule_tac x="m" in exI)
   7.267 +      apply auto
   7.268 +      apply (subgoal_tac "Suc N \<noteq> m")
   7.269 +      apply simp
   7.270 +      apply (rule ccontr, simp)
   7.271 +      done
   7.272 +    hence f0: "f 0 > Suc N" "s (Suc N) < s (f 0)" by blast+
   7.273 +    {fix n
   7.274 +      have "f n > N \<and> ?P (f (Suc n)) (f n)"
   7.275 +	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
   7.276 +      proof (induct n)
   7.277 +	case 0 thus ?case
   7.278 +	  using f0 N apply auto 
   7.279 +	  apply (erule allE[where x="f 0"], clarsimp) 
   7.280 +	  apply (rule_tac x="m" in exI, simp)
   7.281 +	  by (subgoal_tac "f 0 \<noteq> m", auto)
   7.282 +      next
   7.283 +	case (Suc n)
   7.284 +	from Suc.hyps have Nfn: "N < f n" by blast
   7.285 +	from Suc.hyps obtain m where m: "m > f n" "s (f n) < s m" by blast
   7.286 +	with Nfn have mN: "m > N" by arith
   7.287 +	note key = Suc.hyps[unfolded some_eq_ex[of "\<lambda>p. ?P p (f n)", symmetric] f(2)[rule_format, of n, symmetric]]
   7.288 +	
   7.289 +	from key have th0: "f (Suc n) > N" by simp
   7.290 +	from N[rule_format, OF th0]
   7.291 +	obtain m' where m': "m' \<ge> f (Suc n)" "s (f (Suc n)) < s m'" by blast
   7.292 +	have "m' \<noteq> f (Suc (n))" apply (rule ccontr) using m'(2) by auto
   7.293 +	hence "m' > f (Suc n)" using m'(1) by simp
   7.294 +	with key m'(2) show ?case by auto
   7.295 +      qed}
   7.296 +    note fSuc = this
   7.297 +    {fix n
   7.298 +      have "f n \<ge> Suc N \<and> f(Suc n) > f n \<and> s(f n) < s(f(Suc n))" using fSuc[of n] by auto 
   7.299 +      hence "f n \<ge> Suc N" "f(Suc n) > f n" "s(f n) < s(f(Suc n))" by blast+}
   7.300 +    note thf = this
   7.301 +    have sqf: "subseq f" unfolding subseq_Suc_iff using thf by simp
   7.302 +    have "monoseq (\<lambda>n. s (f n))"  unfolding monoseq_Suc using thf
   7.303 +      apply -
   7.304 +      apply (rule disjI1)
   7.305 +      apply auto
   7.306 +      apply (rule order_less_imp_le)
   7.307 +      apply blast
   7.308 +      done
   7.309 +    then have ?thesis  using sqf by blast}
   7.310 +  ultimately show ?thesis unfolding linorder_not_less[symmetric] by blast
   7.311 +qed
   7.312 +
   7.313 +lemma seq_suble: assumes sf: "subseq f" shows "n \<le> f n"
   7.314 +proof(induct n)
   7.315 +  case 0 thus ?case by simp
   7.316 +next
   7.317 +  case (Suc n)
   7.318 +  from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps
   7.319 +  have "n < f (Suc n)" by arith 
   7.320 +  thus ?case by arith
   7.321 +qed
   7.322 +
   7.323 +subsection {* Fundamental theorem of algebra *}
   7.324 +lemma  unimodular_reduce_norm:
   7.325 +  assumes md: "cmod z = 1"
   7.326 +  shows "cmod (z + 1) < 1 \<or> cmod (z - 1) < 1 \<or> cmod (z + ii) < 1 \<or> cmod (z - ii) < 1"
   7.327 +proof-
   7.328 +  obtain x y where z: "z = Complex x y " by (cases z, auto)
   7.329 +  from md z have xy: "x^2 + y^2 = 1" by (simp add: cmod_def)
   7.330 +  {assume C: "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + ii) \<ge> 1" "cmod (z - ii) \<ge> 1"
   7.331 +    from C z xy have "2*x \<le> 1" "2*x \<ge> -1" "2*y \<le> 1" "2*y \<ge> -1"
   7.332 +      by (simp_all add: cmod_def power2_eq_square ring_simps)
   7.333 +    hence "abs (2*x) \<le> 1" "abs (2*y) \<le> 1" by simp_all
   7.334 +    hence "(abs (2 * x))^2 <= 1^2" "(abs (2 * y)) ^2 <= 1^2"
   7.335 +      by - (rule power_mono, simp, simp)+
   7.336 +    hence th0: "4*x^2 \<le> 1" "4*y^2 \<le> 1" 
   7.337 +      by (simp_all  add: power2_abs power_mult_distrib)
   7.338 +    from add_mono[OF th0] xy have False by simp }
   7.339 +  thus ?thesis unfolding linorder_not_le[symmetric] by blast
   7.340 +qed
   7.341 +
   7.342 +text{* Hence we can always reduce modulus of @{text "1 + b z^n"} if nonzero *}
   7.343 +lemma reduce_poly_simple:
   7.344 + assumes b: "b \<noteq> 0" and n: "n\<noteq>0"
   7.345 +  shows "\<exists>z. cmod (1 + b * z^n) < 1"
   7.346 +using n
   7.347 +proof(induct n rule: nat_less_induct)
   7.348 +  fix n
   7.349 +  assume IH: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>z. cmod (1 + b * z ^ m) < 1)" and n: "n \<noteq> 0"
   7.350 +  let ?P = "\<lambda>z n. cmod (1 + b * z ^ n) < 1"
   7.351 +  {assume e: "even n"
   7.352 +    hence "\<exists>m. n = 2*m" by presburger
   7.353 +    then obtain m where m: "n = 2*m" by blast
   7.354 +    from n m have "m\<noteq>0" "m < n" by presburger+
   7.355 +    with IH[rule_format, of m] obtain z where z: "?P z m" by blast
   7.356 +    from z have "?P (csqrt z) n" by (simp add: m power_mult csqrt)
   7.357 +    hence "\<exists>z. ?P z n" ..}
   7.358 +  moreover
   7.359 +  {assume o: "odd n"
   7.360 +    from b have b': "b^2 \<noteq> 0" unfolding power2_eq_square by simp
   7.361 +    have "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
   7.362 +    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) = 
   7.363 +    ((Re (inverse b))^2 + (Im (inverse b))^2) * \<bar>Im b * Im b + Re b * Re b\<bar>" by algebra
   7.364 +    also have "\<dots> = cmod (inverse b) ^2 * cmod b ^ 2" 
   7.365 +      apply (simp add: cmod_def) using realpow_two_le_add_order[of "Re b" "Im b"]
   7.366 +      by (simp add: power2_eq_square)
   7.367 +    finally 
   7.368 +    have th0: "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
   7.369 +    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) =
   7.370 +    1" 
   7.371 +      apply (simp add: power2_eq_square norm_mult[symmetric] norm_inverse[symmetric])
   7.372 +      using right_inverse[OF b']
   7.373 +      by (simp add: power2_eq_square[symmetric] power_inverse[symmetric] ring_simps)
   7.374 +    have th0: "cmod (complex_of_real (cmod b) / b) = 1"
   7.375 +      apply (simp add: complex_Re_mult cmod_def power2_eq_square Re_complex_of_real Im_complex_of_real divide_inverse ring_simps )
   7.376 +      by (simp add: real_sqrt_mult[symmetric] th0)        
   7.377 +    from o have "\<exists>m. n = Suc (2*m)" by presburger+
   7.378 +    then obtain m where m: "n = Suc (2*m)" by blast
   7.379 +    from unimodular_reduce_norm[OF th0] o
   7.380 +    have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
   7.381 +      apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1", rule_tac x="1" in exI, simp)
   7.382 +      apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp add: diff_def)
   7.383 +      apply (cases "cmod (complex_of_real (cmod b) / b + ii) < 1")
   7.384 +      apply (cases "even m", rule_tac x="ii" in exI, simp add: m power_mult)
   7.385 +      apply (rule_tac x="- ii" in exI, simp add: m power_mult)
   7.386 +      apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult diff_def)
   7.387 +      apply (rule_tac x="ii" in exI, simp add: m power_mult diff_def)
   7.388 +      done
   7.389 +    then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" by blast
   7.390 +    let ?w = "v / complex_of_real (root n (cmod b))"
   7.391 +    from odd_real_root_pow[OF o, of "cmod b"]
   7.392 +    have th1: "?w ^ n = v^n / complex_of_real (cmod b)" 
   7.393 +      by (simp add: power_divide complex_of_real_power)
   7.394 +    have th2:"cmod (complex_of_real (cmod b) / b) = 1" using b by (simp add: norm_divide)
   7.395 +    hence th3: "cmod (complex_of_real (cmod b) / b) \<ge> 0" by simp
   7.396 +    have th4: "cmod (complex_of_real (cmod b) / b) *
   7.397 +   cmod (1 + b * (v ^ n / complex_of_real (cmod b)))
   7.398 +   < cmod (complex_of_real (cmod b) / b) * 1"
   7.399 +      apply (simp only: norm_mult[symmetric] right_distrib)
   7.400 +      using b v by (simp add: th2)
   7.401 +
   7.402 +    from mult_less_imp_less_left[OF th4 th3]
   7.403 +    have "?P ?w n" unfolding th1 . 
   7.404 +    hence "\<exists>z. ?P z n" .. }
   7.405 +  ultimately show "\<exists>z. ?P z n" by blast
   7.406 +qed
   7.407 +
   7.408 +
   7.409 +text{* Bolzano-Weierstrass type property for closed disc in complex plane. *}
   7.410 +
   7.411 +lemma metric_bound_lemma: "cmod (x - y) <= \<bar>Re x - Re y\<bar> + \<bar>Im x - Im y\<bar>"
   7.412 +  using real_sqrt_sum_squares_triangle_ineq[of "Re x - Re y" 0 0 "Im x - Im y" ]
   7.413 +  unfolding cmod_def by simp
   7.414 +
   7.415 +lemma bolzano_weierstrass_complex_disc:
   7.416 +  assumes r: "\<forall>n. cmod (s n) \<le> r"
   7.417 +  shows "\<exists>f z. subseq f \<and> (\<forall>e >0. \<exists>N. \<forall>n \<ge> N. cmod (s (f n) - z) < e)"
   7.418 +proof-
   7.419 +  from seq_monosub[of "Re o s"] 
   7.420 +  obtain f g where f: "subseq f" "monoseq (\<lambda>n. Re (s (f n)))" 
   7.421 +    unfolding o_def by blast
   7.422 +  from seq_monosub[of "Im o s o f"] 
   7.423 +  obtain g where g: "subseq g" "monoseq (\<lambda>n. Im (s(f(g n))))" unfolding o_def by blast  
   7.424 +  let ?h = "f o g"
   7.425 +  from r[rule_format, of 0] have rp: "r \<ge> 0" using norm_ge_zero[of "s 0"] by arith 
   7.426 +  have th:"\<forall>n. r + 1 \<ge> \<bar> Re (s n)\<bar>" 
   7.427 +  proof
   7.428 +    fix n
   7.429 +    from abs_Re_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Re (s n)\<bar> \<le> r + 1" by arith
   7.430 +  qed
   7.431 +  have conv1: "convergent (\<lambda>n. Re (s ( f n)))"
   7.432 +    apply (rule Bseq_monoseq_convergent)
   7.433 +    apply (simp add: Bseq_def)
   7.434 +    apply (rule exI[where x= "r + 1"])
   7.435 +    using th rp apply simp
   7.436 +    using f(2) .
   7.437 +  have th:"\<forall>n. r + 1 \<ge> \<bar> Im (s n)\<bar>" 
   7.438 +  proof
   7.439 +    fix n
   7.440 +    from abs_Im_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Im (s n)\<bar> \<le> r + 1" by arith
   7.441 +  qed
   7.442 +
   7.443 +  have conv2: "convergent (\<lambda>n. Im (s (f (g n))))"
   7.444 +    apply (rule Bseq_monoseq_convergent)
   7.445 +    apply (simp add: Bseq_def)
   7.446 +    apply (rule exI[where x= "r + 1"])
   7.447 +    using th rp apply simp
   7.448 +    using g(2) .
   7.449 +
   7.450 +  from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x" 
   7.451 +    by blast 
   7.452 +  hence  x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Re (s (f n)) - x \<bar> < r" 
   7.453 +    unfolding LIMSEQ_def real_norm_def .
   7.454 +
   7.455 +  from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y" 
   7.456 +    by blast 
   7.457 +  hence  y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Im (s (f (g n))) - y \<bar> < r" 
   7.458 +    unfolding LIMSEQ_def real_norm_def .
   7.459 +  let ?w = "Complex x y"
   7.460 +  from f(1) g(1) have hs: "subseq ?h" unfolding subseq_def by auto 
   7.461 +  {fix e assume ep: "e > (0::real)"
   7.462 +    hence e2: "e/2 > 0" by simp
   7.463 +    from x[rule_format, OF e2] y[rule_format, OF e2]
   7.464 +    obtain N1 N2 where N1: "\<forall>n\<ge>N1. \<bar>Re (s (f n)) - x\<bar> < e / 2" and N2: "\<forall>n\<ge>N2. \<bar>Im (s (f (g n))) - y\<bar> < e / 2" by blast
   7.465 +    {fix n assume nN12: "n \<ge> N1 + N2"
   7.466 +      hence nN1: "g n \<ge> N1" and nN2: "n \<ge> N2" using seq_suble[OF g(1), of n] by arith+
   7.467 +      from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
   7.468 +      have "cmod (s (?h n) - ?w) < e" 
   7.469 +	using metric_bound_lemma[of "s (f (g n))" ?w] by simp }
   7.470 +    hence "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" by blast }
   7.471 +  with hs show ?thesis  by blast  
   7.472 +qed
   7.473 +
   7.474 +text{* Polynomial is continuous. *}
   7.475 +
   7.476 +lemma poly_cont:
   7.477 +  assumes ep: "e > 0" 
   7.478 +  shows "\<exists>d >0. \<forall>w. 0 < cmod (w - z) \<and> cmod (w - z) < d \<longrightarrow> cmod (poly p w - poly p z) < e"
   7.479 +proof-
   7.480 +  from poly_offset[of p z] obtain q where q: "length q = length p" "\<And>x. poly q x = poly p (z + x)" by blast
   7.481 +  {fix w
   7.482 +    note q(2)[of "w - z", simplified]}
   7.483 +  note th = this
   7.484 +  show ?thesis unfolding th[symmetric]
   7.485 +  proof(induct q)
   7.486 +    case Nil thus ?case  using ep by auto
   7.487 +  next
   7.488 +    case (Cons c cs)
   7.489 +    from poly_bound_exists[of 1 "cs"] 
   7.490 +    obtain m where m: "m > 0" "\<And>z. cmod z \<le> 1 \<Longrightarrow> cmod (poly cs z) \<le> m" by blast
   7.491 +    from ep m(1) have em0: "e/m > 0" by (simp add: field_simps)
   7.492 +    have one0: "1 > (0::real)"  by arith
   7.493 +    from real_lbound_gt_zero[OF one0 em0] 
   7.494 +    obtain d where d: "d >0" "d < 1" "d < e / m" by blast
   7.495 +    from d(1,3) m(1) have dm: "d*m > 0" "d*m < e" 
   7.496 +      by (simp_all add: field_simps real_mult_order)
   7.497 +    show ?case 
   7.498 +      proof(rule ex_forward[OF real_lbound_gt_zero[OF one0 em0]], clarsimp simp add: norm_mult)
   7.499 +	fix d w
   7.500 +	assume H: "d > 0" "d < 1" "d < e/m" "w\<noteq>z" "cmod (w-z) < d"
   7.501 +	hence d1: "cmod (w-z) \<le> 1" "d \<ge> 0" by simp_all
   7.502 +	from H(3) m(1) have dme: "d*m < e" by (simp add: field_simps)
   7.503 +	from H have th: "cmod (w-z) \<le> d" by simp 
   7.504 +	from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
   7.505 +	show "cmod (w - z) * cmod (poly cs (w - z)) < e" by simp
   7.506 +      qed  
   7.507 +    qed
   7.508 +qed
   7.509 +
   7.510 +text{* Hence a polynomial attains minimum on a closed disc 
   7.511 +  in the complex plane. *}
   7.512 +lemma  poly_minimum_modulus_disc:
   7.513 +  "\<exists>z. \<forall>w. cmod w \<le> r \<longrightarrow> cmod (poly p z) \<le> cmod (poly p w)"
   7.514 +proof-
   7.515 +  {assume "\<not> r \<ge> 0" hence ?thesis unfolding linorder_not_le
   7.516 +      apply -
   7.517 +      apply (rule exI[where x=0]) 
   7.518 +      apply auto
   7.519 +      apply (subgoal_tac "cmod w < 0")
   7.520 +      apply simp
   7.521 +      apply arith
   7.522 +      done }
   7.523 +  moreover
   7.524 +  {assume rp: "r \<ge> 0"
   7.525 +    from rp have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))" by simp 
   7.526 +    hence mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x"  by blast
   7.527 +    {fix x z
   7.528 +      assume H: "cmod z \<le> r" "cmod (poly p z) = - x" "\<not>x < 1"
   7.529 +      hence "- x < 0 " by arith
   7.530 +      with H(2) norm_ge_zero[of "poly p z"]  have False by simp }
   7.531 +    then have mth2: "\<exists>z. \<forall>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<longrightarrow> x < z" by blast
   7.532 +    from real_sup_exists[OF mth1 mth2] obtain s where 
   7.533 +      s: "\<forall>y. (\<exists>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<and> y < x) \<longleftrightarrow>(y < s)" by blast
   7.534 +    let ?m = "-s"
   7.535 +    {fix y
   7.536 +      from s[rule_format, of "-y"] have 
   7.537 +    "(\<exists>z x. cmod z \<le> r \<and> -(- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y" 
   7.538 +	unfolding minus_less_iff[of y ] equation_minus_iff by blast }
   7.539 +    note s1 = this[unfolded minus_minus]
   7.540 +    from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m" 
   7.541 +      by auto
   7.542 +    {fix n::nat
   7.543 +      from s1[rule_format, of "?m + 1/real (Suc n)"] 
   7.544 +      have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)"
   7.545 +	by simp}
   7.546 +    hence th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
   7.547 +    from choice[OF th] obtain g where 
   7.548 +      g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m+1 /real(Suc n)" 
   7.549 +      by blast
   7.550 +    from bolzano_weierstrass_complex_disc[OF g(1)] 
   7.551 +    obtain f z where fz: "subseq f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
   7.552 +      by blast    
   7.553 +    {fix w 
   7.554 +      assume wr: "cmod w \<le> r"
   7.555 +      let ?e = "\<bar>cmod (poly p z) - ?m\<bar>"
   7.556 +      {assume e: "?e > 0"
   7.557 +	hence e2: "?e/2 > 0" by simp
   7.558 +	from poly_cont[OF e2, of z p] obtain d where
   7.559 +	  d: "d>0" "\<forall>w. 0<cmod (w - z)\<and> cmod(w - z) < d \<longrightarrow> cmod(poly p w - poly p z) < ?e/2" by blast
   7.560 +	{fix w assume w: "cmod (w - z) < d"
   7.561 +	  have "cmod(poly p w - poly p z) < ?e / 2"
   7.562 +	    using d(2)[rule_format, of w] w e by (cases "w=z", simp_all)}
   7.563 +	note th1 = this
   7.564 +	
   7.565 +	from fz(2)[rule_format, OF d(1)] obtain N1 where 
   7.566 +	  N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d" by blast
   7.567 +	from reals_Archimedean2[of "2/?e"] obtain N2::nat where
   7.568 +	  N2: "2/?e < real N2" by blast
   7.569 +	have th2: "cmod(poly p (g(f(N1 + N2))) - poly p z) < ?e/2"
   7.570 +	  using N1[rule_format, of "N1 + N2"] th1 by simp
   7.571 +	{fix a b e2 m :: real
   7.572 +	have "a < e2 \<Longrightarrow> abs(b - m) < e2 \<Longrightarrow> 2 * e2 <= abs(b - m) + a
   7.573 +          ==> False" by arith}
   7.574 +      note th0 = this
   7.575 +      have ath: 
   7.576 +	"\<And>m x e. m <= x \<Longrightarrow>  x < m + e ==> abs(x - m::real) < e" by arith
   7.577 +      from s1m[OF g(1)[rule_format]]
   7.578 +      have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
   7.579 +      from seq_suble[OF fz(1), of "N1+N2"]
   7.580 +      have th00: "real (Suc (N1+N2)) \<le> real (Suc (f (N1+N2)))" by simp
   7.581 +      have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1+N2)) > 0"  
   7.582 +	using N2 by auto
   7.583 +      from frac_le[OF th000 th00] have th00: "?m +1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))" by simp
   7.584 +      from g(2)[rule_format, of "f (N1 + N2)"]
   7.585 +      have th01:"cmod (poly p (g (f (N1 + N2)))) < - s + 1 / real (Suc (f (N1 + N2)))" .
   7.586 +      from order_less_le_trans[OF th01 th00]
   7.587 +      have th32: "cmod(poly p (g (f (N1 + N2)))) < ?m + (1/ real(Suc (N1 + N2)))" .
   7.588 +      from N2 have "2/?e < real (Suc (N1 + N2))" by arith
   7.589 +      with e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
   7.590 +      have "?e/2 > 1/ real (Suc (N1 + N2))" by (simp add: inverse_eq_divide)
   7.591 +      with ath[OF th31 th32]
   7.592 +      have thc1:"\<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar>< ?e/2" by arith  
   7.593 +      have ath2: "\<And>(a::real) b c m. \<bar>a - b\<bar> <= c ==> \<bar>b - m\<bar> <= \<bar>a - m\<bar> + c" 
   7.594 +	by arith
   7.595 +      have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar>
   7.596 +\<le> cmod (poly p (g (f (N1 + N2))) - poly p z)" 
   7.597 +	by (simp add: norm_triangle_ineq3)
   7.598 +      from ath2[OF th22, of ?m]
   7.599 +      have thc2: "2*(?e/2) \<le> \<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar> + cmod (poly p (g (f (N1 + N2))) - poly p z)" by simp
   7.600 +      from th0[OF th2 thc1 thc2] have False .}
   7.601 +      hence "?e = 0" by auto
   7.602 +      then have "cmod (poly p z) = ?m" by simp  
   7.603 +      with s1m[OF wr]
   7.604 +      have "cmod (poly p z) \<le> cmod (poly p w)" by simp }
   7.605 +    hence ?thesis by blast}
   7.606 +  ultimately show ?thesis by blast
   7.607 +qed
   7.608 +
   7.609 +lemma "(rcis (sqrt (abs r)) (a/2)) ^ 2 = rcis (abs r) a"
   7.610 +  unfolding power2_eq_square
   7.611 +  apply (simp add: rcis_mult)
   7.612 +  apply (simp add: power2_eq_square[symmetric])
   7.613 +  done
   7.614 +
   7.615 +lemma cispi: "cis pi = -1" 
   7.616 +  unfolding cis_def
   7.617 +  by simp
   7.618 +
   7.619 +lemma "(rcis (sqrt (abs r)) ((pi + a)/2)) ^ 2 = rcis (- abs r) a"
   7.620 +  unfolding power2_eq_square
   7.621 +  apply (simp add: rcis_mult add_divide_distrib)
   7.622 +  apply (simp add: power2_eq_square[symmetric] rcis_def cispi cis_mult[symmetric])
   7.623 +  done
   7.624 +
   7.625 +text {* Nonzero polynomial in z goes to infinity as z does. *}
   7.626 +
   7.627 +instance complex::idom_char_0 by (intro_classes)
   7.628 +instance complex :: recpower_idom_char_0 by intro_classes
   7.629 +
   7.630 +lemma poly_infinity:
   7.631 +  assumes ex: "list_ex (\<lambda>c. c \<noteq> 0) p"
   7.632 +  shows "\<exists>r. \<forall>z. r \<le> cmod z \<longrightarrow> d \<le> cmod (poly (a#p) z)"
   7.633 +using ex
   7.634 +proof(induct p arbitrary: a d)
   7.635 +  case (Cons c cs a d) 
   7.636 +  {assume H: "list_ex (\<lambda>c. c\<noteq>0) cs"
   7.637 +    with Cons.hyps obtain r where r: "\<forall>z. r \<le> cmod z \<longrightarrow> d + cmod a \<le> cmod (poly (c # cs) z)" by blast
   7.638 +    let ?r = "1 + \<bar>r\<bar>"
   7.639 +    {fix z assume h: "1 + \<bar>r\<bar> \<le> cmod z"
   7.640 +      have r0: "r \<le> cmod z" using h by arith
   7.641 +      from r[rule_format, OF r0]
   7.642 +      have th0: "d + cmod a \<le> 1 * cmod(poly (c#cs) z)" by arith
   7.643 +      from h have z1: "cmod z \<ge> 1" by arith
   7.644 +      from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (c#cs) z"]]]
   7.645 +      have th1: "d \<le> cmod(z * poly (c#cs) z) - cmod a"
   7.646 +	unfolding norm_mult by (simp add: ring_simps)
   7.647 +      from complex_mod_triangle_sub[of "z * poly (c#cs) z" a]
   7.648 +      have th2: "cmod(z * poly (c#cs) z) - cmod a \<le> cmod (poly (a#c#cs) z)" 
   7.649 +	by (simp add: diff_le_eq ring_simps) 
   7.650 +      from th1 th2 have "d \<le> cmod (poly (a#c#cs) z)"  by arith}
   7.651 +    hence ?case by blast}
   7.652 +  moreover
   7.653 +  {assume cs0: "\<not> (list_ex (\<lambda>c. c \<noteq> 0) cs)"
   7.654 +    with Cons.prems have c0: "c \<noteq> 0" by simp
   7.655 +    from cs0 have cs0': "list_all (\<lambda>c. c = 0) cs" 
   7.656 +      by (auto simp add: list_all_iff list_ex_iff)
   7.657 +    {fix z
   7.658 +      assume h: "(\<bar>d\<bar> + cmod a) / cmod c \<le> cmod z"
   7.659 +      from c0 have "cmod c > 0" by simp
   7.660 +      from h c0 have th0: "\<bar>d\<bar> + cmod a \<le> cmod (z*c)" 
   7.661 +	by (simp add: field_simps norm_mult)
   7.662 +      have ath: "\<And>mzh mazh ma. mzh <= mazh + ma ==> abs(d) + ma <= mzh ==> d <= mazh" by arith
   7.663 +      from complex_mod_triangle_sub[of "z*c" a ]
   7.664 +      have th1: "cmod (z * c) \<le> cmod (a + z * c) + cmod a"
   7.665 +	by (simp add: ring_simps)
   7.666 +      from ath[OF th1 th0] have "d \<le> cmod (poly (a # c # cs) z)" 
   7.667 +	using poly_0[OF cs0'] by simp}
   7.668 +    then have ?case  by blast}
   7.669 +  ultimately show ?case by blast
   7.670 +qed simp
   7.671 +
   7.672 +text {* Hence polynomial's modulus attains its minimum somewhere. *}
   7.673 +lemma poly_minimum_modulus:
   7.674 +  "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)"
   7.675 +proof(induct p)
   7.676 +  case (Cons c cs) 
   7.677 +  {assume cs0: "list_ex (\<lambda>c. c \<noteq> 0) cs"
   7.678 +    from poly_infinity[OF cs0, of "cmod (poly (c#cs) 0)" c]
   7.679 +    obtain r where r: "\<And>z. r \<le> cmod z \<Longrightarrow> cmod (poly (c # cs) 0) \<le> cmod (poly (c # cs) z)" by blast
   7.680 +    have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>" by arith
   7.681 +    from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "c#cs"] 
   7.682 +    obtain v where v: "\<And>w. cmod w \<le> \<bar>r\<bar> \<Longrightarrow> cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) w)" by blast
   7.683 +    {fix z assume z: "r \<le> cmod z"
   7.684 +      from v[of 0] r[OF z] 
   7.685 +      have "cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) z)"
   7.686 +	by simp }
   7.687 +    note v0 = this
   7.688 +    from v0 v ath[of r] have ?case by blast}
   7.689 +  moreover
   7.690 +  {assume cs0: "\<not> (list_ex (\<lambda>c. c\<noteq>0) cs)"
   7.691 +    hence th:"list_all (\<lambda>c. c = 0) cs" by (simp add: list_all_iff list_ex_iff)
   7.692 +    from poly_0[OF th] Cons.hyps have ?case by simp}
   7.693 +  ultimately show ?case by blast
   7.694 +qed simp
   7.695 +
   7.696 +text{* Constant function (non-syntactic characterization). *}
   7.697 +definition "constant f = (\<forall>x y. f x = f y)"
   7.698 +
   7.699 +lemma nonconstant_length: "\<not> (constant (poly p)) \<Longrightarrow> length p \<ge> 2"
   7.700 +  unfolding constant_def
   7.701 +  apply (induct p, auto)
   7.702 +  apply (unfold not_less[symmetric])
   7.703 +  apply simp
   7.704 +  apply (rule ccontr)
   7.705 +  apply auto
   7.706 +  done
   7.707 + 
   7.708 +lemma poly_replicate_append:
   7.709 +  "poly ((replicate n 0)@p) (x::'a::{recpower, comm_ring}) = x^n * poly p x"
   7.710 +  by(induct n, auto simp add: power_Suc ring_simps)
   7.711 +
   7.712 +text {* Decomposition of polynomial, skipping zero coefficients 
   7.713 +  after the first.  *}
   7.714 +
   7.715 +lemma poly_decompose_lemma:
   7.716 + assumes nz: "\<not>(\<forall>z. z\<noteq>0 \<longrightarrow> poly p z = (0::'a::{recpower,idom}))"
   7.717 +  shows "\<exists>k a q. a\<noteq>0 \<and> Suc (length q + k) = length p \<and> 
   7.718 +                 (\<forall>z. poly p z = z^k * poly (a#q) z)"
   7.719 +using nz
   7.720 +proof(induct p)
   7.721 +  case Nil thus ?case by simp
   7.722 +next
   7.723 +  case (Cons c cs)
   7.724 +  {assume c0: "c = 0"
   7.725 +    
   7.726 +    from Cons.hyps Cons.prems c0 have ?case apply auto
   7.727 +      apply (rule_tac x="k+1" in exI)
   7.728 +      apply (rule_tac x="a" in exI, clarsimp)
   7.729 +      apply (rule_tac x="q" in exI)
   7.730 +      by (auto simp add: power_Suc)}
   7.731 +  moreover
   7.732 +  {assume c0: "c\<noteq>0"
   7.733 +    hence ?case apply-
   7.734 +      apply (rule exI[where x=0])
   7.735 +      apply (rule exI[where x=c], clarsimp)
   7.736 +      apply (rule exI[where x=cs])
   7.737 +      apply auto
   7.738 +      done}
   7.739 +  ultimately show ?case by blast
   7.740 +qed
   7.741 +
   7.742 +lemma poly_decompose:
   7.743 +  assumes nc: "~constant(poly p)"
   7.744 +  shows "\<exists>k a q. a\<noteq>(0::'a::{recpower,idom}) \<and> k\<noteq>0 \<and>
   7.745 +               length q + k + 1 = length p \<and> 
   7.746 +              (\<forall>z. poly p z = poly p 0 + z^k * poly (a#q) z)"
   7.747 +using nc 
   7.748 +proof(induct p)
   7.749 +  case Nil thus ?case by (simp add: constant_def)
   7.750 +next
   7.751 +  case (Cons c cs)
   7.752 +  {assume C:"\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0"
   7.753 +    {fix x y
   7.754 +      from C have "poly (c#cs) x = poly (c#cs) y" by (cases "x=0", auto)}
   7.755 +    with Cons.prems have False by (auto simp add: constant_def)}
   7.756 +  hence th: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)" ..
   7.757 +  from poly_decompose_lemma[OF th] 
   7.758 +  show ?case 
   7.759 +    apply clarsimp    
   7.760 +    apply (rule_tac x="k+1" in exI)
   7.761 +    apply (rule_tac x="a" in exI)
   7.762 +    apply simp
   7.763 +    apply (rule_tac x="q" in exI)
   7.764 +    apply (auto simp add: power_Suc)
   7.765 +    done
   7.766 +qed
   7.767 +
   7.768 +text{* Fundamental theorem of algebral *}
   7.769 +
   7.770 +lemma fundamental_theorem_of_algebra:
   7.771 +  assumes nc: "~constant(poly p)"
   7.772 +  shows "\<exists>z::complex. poly p z = 0"
   7.773 +using nc
   7.774 +proof(induct n\<equiv> "length p" arbitrary: p rule: nat_less_induct)
   7.775 +  fix n fix p :: "complex list"
   7.776 +  let ?p = "poly p"
   7.777 +  assume H: "\<forall>m<n. \<forall>p. \<not> constant (poly p) \<longrightarrow> m = length p \<longrightarrow> (\<exists>(z::complex). poly p z = 0)" and nc: "\<not> constant ?p" and n: "n = length p"
   7.778 +  let ?ths = "\<exists>z. ?p z = 0"
   7.779 +
   7.780 +  from nonconstant_length[OF nc] have n2: "n\<ge> 2" by (simp add: n)
   7.781 +  from poly_minimum_modulus obtain c where 
   7.782 +    c: "\<forall>w. cmod (?p c) \<le> cmod (?p w)" by blast
   7.783 +  {assume pc: "?p c = 0" hence ?ths by blast}
   7.784 +  moreover
   7.785 +  {assume pc0: "?p c \<noteq> 0"
   7.786 +    from poly_offset[of p c] obtain q where
   7.787 +      q: "length q = length p" "\<forall>x. poly q x = ?p (c+x)" by blast
   7.788 +    {assume h: "constant (poly q)"
   7.789 +      from q(2) have th: "\<forall>x. poly q (x - c) = ?p x" by auto
   7.790 +      {fix x y
   7.791 +	from th have "?p x = poly q (x - c)" by auto 
   7.792 +	also have "\<dots> = poly q (y - c)" 
   7.793 +	  using h unfolding constant_def by blast
   7.794 +	also have "\<dots> = ?p y" using th by auto
   7.795 +	finally have "?p x = ?p y" .}
   7.796 +      with nc have False unfolding constant_def by blast }
   7.797 +    hence qnc: "\<not> constant (poly q)" by blast
   7.798 +    from q(2) have pqc0: "?p c = poly q 0" by simp
   7.799 +    from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)" by simp 
   7.800 +    let ?a0 = "poly q 0"
   7.801 +    from pc0 pqc0 have a00: "?a0 \<noteq> 0" by simp 
   7.802 +    from a00 
   7.803 +    have qr: "\<forall>z. poly q z = poly (map (op * (inverse ?a0)) q) z * ?a0"
   7.804 +      by (simp add: poly_cmult_map)
   7.805 +    let ?r = "map (op * (inverse ?a0)) q"
   7.806 +    have lgqr: "length q = length ?r" by simp 
   7.807 +    {assume h: "\<And>x y. poly ?r x = poly ?r y"
   7.808 +      {fix x y
   7.809 +	from qr[rule_format, of x] 
   7.810 +	have "poly q x = poly ?r x * ?a0" by auto
   7.811 +	also have "\<dots> = poly ?r y * ?a0" using h by simp
   7.812 +	also have "\<dots> = poly q y" using qr[rule_format, of y] by simp
   7.813 +	finally have "poly q x = poly q y" .} 
   7.814 +      with qnc have False unfolding constant_def by blast}
   7.815 +    hence rnc: "\<not> constant (poly ?r)" unfolding constant_def by blast
   7.816 +    from qr[rule_format, of 0] a00  have r01: "poly ?r 0 = 1" by auto
   7.817 +    {fix w 
   7.818 +      have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
   7.819 +	using qr[rule_format, of w] a00 by simp
   7.820 +      also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
   7.821 +	using a00 unfolding norm_divide by (simp add: field_simps)
   7.822 +      finally have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" .}
   7.823 +    note mrmq_eq = this
   7.824 +    from poly_decompose[OF rnc] obtain k a s where 
   7.825 +      kas: "a\<noteq>0" "k\<noteq>0" "length s + k + 1 = length ?r" 
   7.826 +      "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (a#s) z" by blast
   7.827 +    {assume "k + 1 = n"
   7.828 +      with kas(3) lgqr[symmetric] q(1) n[symmetric] have s0:"s=[]" by auto
   7.829 +      {fix w
   7.830 +	have "cmod (poly ?r w) = cmod (1 + a * w ^ k)" 
   7.831 +	  using kas(4)[rule_format, of w] s0 r01 by (simp add: ring_simps)}
   7.832 +      note hth = this [symmetric]
   7.833 +	from reduce_poly_simple[OF kas(1,2)] 
   7.834 +      have "\<exists>w. cmod (poly ?r w) < 1" unfolding hth by blast}
   7.835 +    moreover
   7.836 +    {assume kn: "k+1 \<noteq> n"
   7.837 +      from kn kas(3) q(1) n[symmetric] have k1n: "k + 1 < n" by simp
   7.838 +      have th01: "\<not> constant (poly (1#((replicate (k - 1) 0)@[a])))" 
   7.839 +	unfolding constant_def poly_Nil poly_Cons poly_replicate_append
   7.840 +	using kas(1) apply simp 
   7.841 +	by (rule exI[where x=0], rule exI[where x=1], simp)
   7.842 +      from kas(2) have th02: "k+1 = length (1#((replicate (k - 1) 0)@[a]))" 
   7.843 +	by simp
   7.844 +      from H[rule_format, OF k1n th01 th02]
   7.845 +      obtain w where w: "1 + w^k * a = 0"
   7.846 +	unfolding poly_Nil poly_Cons poly_replicate_append
   7.847 +	using kas(2) by (auto simp add: power_Suc[symmetric, of _ "k - Suc 0"] 
   7.848 +	  mult_assoc[of _ _ a, symmetric])
   7.849 +      from poly_bound_exists[of "cmod w" s] obtain m where 
   7.850 +	m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
   7.851 +      have w0: "w\<noteq>0" using kas(2) w by (auto simp add: power_0_left)
   7.852 +      from w have "(1 + w ^ k * a) - 1 = 0 - 1" by simp
   7.853 +      then have wm1: "w^k * a = - 1" by simp
   7.854 +      have inv0: "0 < inverse (cmod w ^ (k + 1) * m)" 
   7.855 +	using norm_ge_zero[of w] w0 m(1)
   7.856 +	  by (simp add: inverse_eq_divide zero_less_mult_iff)
   7.857 +      with real_down2[OF zero_less_one] obtain t where
   7.858 +	t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
   7.859 +      let ?ct = "complex_of_real t"
   7.860 +      let ?w = "?ct * w"
   7.861 +      have "1 + ?w^k * (a + ?w * poly s ?w) = 1 + ?ct^k * (w^k * a) + ?w^k * ?w * poly s ?w" using kas(1) by (simp add: ring_simps power_mult_distrib)
   7.862 +      also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
   7.863 +	unfolding wm1 by (simp)
   7.864 +      finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) = cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)" 
   7.865 +	apply -
   7.866 +	apply (rule cong[OF refl[of cmod]])
   7.867 +	apply assumption
   7.868 +	done
   7.869 +      with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"] 
   7.870 +      have th11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)" unfolding norm_of_real by simp 
   7.871 +      have ath: "\<And>x (t::real). 0\<le> x \<Longrightarrow> x < t \<Longrightarrow> t\<le>1 \<Longrightarrow> \<bar>1 - t\<bar> + x < 1" by arith
   7.872 +      have "t *cmod w \<le> 1 * cmod w" apply (rule mult_mono) using t(1,2) by auto
   7.873 +      then have tw: "cmod ?w \<le> cmod w" using t(1) by (simp add: norm_mult) 
   7.874 +      from t inv0 have "t* (cmod w ^ (k + 1) * m) < 1"
   7.875 +	by (simp add: inverse_eq_divide field_simps)
   7.876 +      with zero_less_power[OF t(1), of k] 
   7.877 +      have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1" 
   7.878 +	apply - apply (rule mult_strict_left_mono) by simp_all
   7.879 +      have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k+1) * cmod (poly s ?w)))"  using w0 t(1)
   7.880 +	by (simp add: ring_simps power_mult_distrib norm_of_real norm_power norm_mult)
   7.881 +      then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
   7.882 +	using t(1,2) m(2)[rule_format, OF tw] w0
   7.883 +	apply (simp only: )
   7.884 +	apply auto
   7.885 +	apply (rule mult_mono, simp_all add: norm_ge_zero)+
   7.886 +	apply (simp add: zero_le_mult_iff zero_le_power)
   7.887 +	done
   7.888 +      with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k" by simp 
   7.889 +      from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1" 
   7.890 +	by auto
   7.891 +      from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] th120 th121]
   7.892 +      have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" . 
   7.893 +      from th11 th12
   7.894 +      have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"  by arith 
   7.895 +      then have "cmod (poly ?r ?w) < 1" 
   7.896 +	unfolding kas(4)[rule_format, of ?w] r01 by simp 
   7.897 +      then have "\<exists>w. cmod (poly ?r w) < 1" by blast}
   7.898 +    ultimately have cr0_contr: "\<exists>w. cmod (poly ?r w) < 1" by blast
   7.899 +    from cr0_contr cq0 q(2)
   7.900 +    have ?ths unfolding mrmq_eq not_less[symmetric] by auto}
   7.901 +  ultimately show ?ths by blast
   7.902 +qed
   7.903 +
   7.904 +text {* Alternative version with a syntactic notion of constant polynomial. *}
   7.905 +
   7.906 +lemma fundamental_theorem_of_algebra_alt:
   7.907 +  assumes nc: "~(\<exists>a l. a\<noteq> 0 \<and> list_all(\<lambda>b. b = 0) l \<and> p = a#l)"
   7.908 +  shows "\<exists>z. poly p z = (0::complex)"
   7.909 +using nc
   7.910 +proof(induct p)
   7.911 +  case (Cons c cs)
   7.912 +  {assume "c=0" hence ?case by auto}
   7.913 +  moreover
   7.914 +  {assume c0: "c\<noteq>0"
   7.915 +    {assume nc: "constant (poly (c#cs))"
   7.916 +      from nc[unfolded constant_def, rule_format, of 0] 
   7.917 +      have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto 
   7.918 +      hence "list_all (\<lambda>c. c=0) cs"
   7.919 +	proof(induct cs)
   7.920 +	  case (Cons d ds)
   7.921 +	  {assume "d=0" hence ?case using Cons.prems Cons.hyps by simp}
   7.922 +	  moreover
   7.923 +	  {assume d0: "d\<noteq>0"
   7.924 +	    from poly_bound_exists[of 1 ds] obtain m where 
   7.925 +	      m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
   7.926 +	    have dm: "cmod d / m > 0" using d0 m(1) by (simp add: field_simps)
   7.927 +	    from real_down2[OF dm zero_less_one] obtain x where 
   7.928 +	      x: "x > 0" "x < cmod d / m" "x < 1" by blast
   7.929 +	    let ?x = "complex_of_real x"
   7.930 +	    from x have cx: "?x \<noteq> 0"  "cmod ?x \<le> 1" by simp_all
   7.931 +	    from Cons.prems[rule_format, OF cx(1)]
   7.932 +	    have cth: "cmod (?x*poly ds ?x) = cmod d" by (simp add: eq_diff_eq[symmetric])
   7.933 +	    from m(2)[rule_format, OF cx(2)] x(1)
   7.934 +	    have th0: "cmod (?x*poly ds ?x) \<le> x*m"
   7.935 +	      by (simp add: norm_mult)
   7.936 +	    from x(2) m(1) have "x*m < cmod d" by (simp add: field_simps)
   7.937 +	    with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d" by auto
   7.938 +	    with cth  have ?case by blast}
   7.939 +	  ultimately show ?case by blast 
   7.940 +	qed simp}
   7.941 +      then have nc: "\<not> constant (poly (c#cs))" using Cons.prems c0 
   7.942 +	by blast
   7.943 +      from fundamental_theorem_of_algebra[OF nc] have ?case .}
   7.944 +  ultimately show ?case by blast  
   7.945 +qed simp
   7.946 +
   7.947 +subsection{* Nullstellenstatz, degrees and divisibility of polynomials *}
   7.948 +
   7.949 +lemma nullstellensatz_lemma:
   7.950 +  fixes p :: "complex list"
   7.951 +  assumes "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
   7.952 +  and "degree p = n" and "n \<noteq> 0"
   7.953 +  shows "p divides (pexp q n)"
   7.954 +using prems
   7.955 +proof(induct n arbitrary: p q rule: nat_less_induct)
   7.956 +  fix n::nat fix p q :: "complex list"
   7.957 +  assume IH: "\<forall>m<n. \<forall>p q.
   7.958 +                 (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow>
   7.959 +                 degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p divides (q %^ m)"
   7.960 +    and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0" 
   7.961 +    and dpn: "degree p = n" and n0: "n \<noteq> 0"
   7.962 +  let ?ths = "p divides (q %^ n)"
   7.963 +  {fix a assume a: "poly p a = 0"
   7.964 +    {assume p0: "poly p = poly []" 
   7.965 +      hence ?ths unfolding divides_def  using pq0 n0
   7.966 +	apply - apply (rule exI[where x="[]"], rule ext)
   7.967 +	by (auto simp add: poly_mult poly_exp)}
   7.968 +    moreover
   7.969 +    {assume p0: "poly p \<noteq> poly []" 
   7.970 +      and oa: "order  a p \<noteq> 0"
   7.971 +      from p0 have pne: "p \<noteq> []" by auto
   7.972 +      let ?op = "order a p"
   7.973 +      from p0 have ap: "([- a, 1] %^ ?op) divides p" 
   7.974 +	"\<not> pexp [- a, 1] (Suc ?op) divides p" using order by blast+ 
   7.975 +      note oop = order_degree[OF p0, unfolded dpn]
   7.976 +      {assume q0: "q = []"
   7.977 +	hence ?ths using n0 unfolding divides_def 
   7.978 +	  apply simp
   7.979 +	  apply (rule exI[where x="[]"], rule ext)
   7.980 +	  by (simp add: divides_def poly_exp poly_mult)}
   7.981 +      moreover
   7.982 +      {assume q0: "q\<noteq>[]"
   7.983 +	from pq0[rule_format, OF a, unfolded poly_linear_divides] q0
   7.984 +	obtain r where r: "q = pmult [- a, 1] r" by blast
   7.985 +	from ap[unfolded divides_def] obtain s where
   7.986 +	  s: "poly p = poly (pmult (pexp [- a, 1] ?op) s)" by blast
   7.987 +	have s0: "poly s \<noteq> poly []"
   7.988 +	  using s p0 by (simp add: poly_entire)
   7.989 +	hence pns0: "poly (pnormalize s) \<noteq> poly []" and sne: "s\<noteq>[]" by auto
   7.990 +	{assume ds0: "degree s = 0"
   7.991 +	  from ds0 pns0 have "\<exists>k. pnormalize s = [k]" unfolding degree_def 
   7.992 +	    by (cases "pnormalize s", auto)
   7.993 +	  then obtain k where kpn: "pnormalize s = [k]" by blast
   7.994 +	  from pns0[unfolded poly_zero] kpn have k: "k \<noteq>0" "poly s = poly [k]"
   7.995 +	    using poly_normalize[of s] by simp_all
   7.996 +	  let ?w = "pmult (pmult [1/k] (pexp [-a,1] (n - ?op))) (pexp r n)"
   7.997 +	  from k r s oop have "poly (pexp q n) = poly (pmult p ?w)"
   7.998 +	    by - (rule ext, simp add: poly_mult poly_exp poly_cmult poly_add power_add[symmetric] ring_simps power_mult_distrib[symmetric])
   7.999 +	  hence ?ths unfolding divides_def by blast}
  7.1000 +	moreover
  7.1001 +	{assume ds0: "degree s \<noteq> 0"
  7.1002 +	  from ds0 s0 dpn degree_unique[OF s, unfolded linear_pow_mul_degree] oa
  7.1003 +	    have dsn: "degree s < n" by auto 
  7.1004 +	    {fix x assume h: "poly s x = 0"
  7.1005 +	      {assume xa: "x = a"
  7.1006 +		from h[unfolded xa poly_linear_divides] sne obtain u where
  7.1007 +		  u: "s = pmult [- a, 1] u" by blast
  7.1008 +		have "poly p = poly (pmult (pexp [- a, 1] (Suc ?op)) u)"
  7.1009 +		  unfolding s u
  7.1010 +		  apply (rule ext)
  7.1011 +		  by (simp add: ring_simps power_mult_distrib[symmetric] poly_mult poly_cmult poly_add poly_exp)
  7.1012 +		with ap(2)[unfolded divides_def] have False by blast}
  7.1013 +	      note xa = this
  7.1014 +	      from h s have "poly p x = 0" by (simp add: poly_mult)
  7.1015 +	      with pq0 have "poly q x = 0" by blast
  7.1016 +	      with r xa have "poly r x = 0"
  7.1017 +		by (auto simp add: poly_mult poly_add poly_cmult eq_diff_eq[symmetric])}
  7.1018 +	    note impth = this
  7.1019 +	    from IH[rule_format, OF dsn, of s r] impth ds0
  7.1020 +	    have "s divides (pexp r (degree s))" by blast
  7.1021 +	    then obtain u where u: "poly (pexp r (degree s)) = poly (pmult s u)"
  7.1022 +	      unfolding divides_def by blast
  7.1023 +	    hence u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
  7.1024 +	      by (simp add: poly_mult[symmetric] poly_exp[symmetric])
  7.1025 +	    let ?w = "pmult (pmult u (pexp [-a,1] (n - ?op))) (pexp r (n - degree s))"
  7.1026 +	    from u' s r oop[of a] dsn have "poly (pexp q n) = poly (pmult p ?w)"
  7.1027 +	      apply - apply (rule ext)
  7.1028 +	      apply (simp only:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult ring_simps)
  7.1029 +	      
  7.1030 +	      apply (simp add:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult mult_assoc[symmetric])
  7.1031 +	      done
  7.1032 +	    hence ?ths unfolding divides_def by blast}
  7.1033 +      ultimately have ?ths by blast }
  7.1034 +      ultimately have ?ths by blast}
  7.1035 +    ultimately have ?ths using a order_root by blast}
  7.1036 +  moreover
  7.1037 +  {assume exa: "\<not> (\<exists>a. poly p a = 0)"
  7.1038 +    from fundamental_theorem_of_algebra_alt[of p] exa obtain c cs where
  7.1039 +      ccs: "c\<noteq>0" "list_all (\<lambda>c. c = 0) cs" "p = c#cs" by blast
  7.1040 +    
  7.1041 +    from poly_0[OF ccs(2)] ccs(3) 
  7.1042 +    have pp: "\<And>x. poly p x =  c" by simp
  7.1043 +    let ?w = "pmult [1/c] (pexp q n)"
  7.1044 +    from pp ccs(1) 
  7.1045 +    have "poly (pexp q n) = poly (pmult p ?w) "
  7.1046 +      apply - apply (rule ext)
  7.1047 +      unfolding poly_mult_assoc[symmetric] by (simp add: poly_mult)
  7.1048 +    hence ?ths unfolding divides_def by blast}
  7.1049 +  ultimately show ?ths by blast
  7.1050 +qed
  7.1051 +
  7.1052 +lemma nullstellensatz_univariate:
  7.1053 +  "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> 
  7.1054 +    p divides (q %^ (degree p)) \<or> (poly p = poly [] \<and> poly q = poly [])"
  7.1055 +proof-
  7.1056 +  {assume pe: "poly p = poly []"
  7.1057 +    hence eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> poly q = poly []"
  7.1058 +      apply auto
  7.1059 +      by (rule ext, simp)
  7.1060 +    {assume "p divides (pexp q (degree p))"
  7.1061 +      then obtain r where r: "poly (pexp q (degree p)) = poly (pmult p r)" 
  7.1062 +	unfolding divides_def by blast
  7.1063 +      from cong[OF r refl] pe degree_unique[OF pe]
  7.1064 +      have False by (simp add: poly_mult degree_def)}
  7.1065 +    with eq pe have ?thesis by blast}
  7.1066 +  moreover
  7.1067 +  {assume pe: "poly p \<noteq> poly []"
  7.1068 +    have p0: "poly [0] = poly []" by (rule ext, simp)
  7.1069 +    {assume dp: "degree p = 0"
  7.1070 +      then obtain k where "pnormalize p = [k]" using pe poly_normalize[of p]
  7.1071 +	unfolding degree_def by (cases "pnormalize p", auto)
  7.1072 +      hence k: "pnormalize p = [k]" "poly p = poly [k]" "k\<noteq>0"
  7.1073 +	using pe poly_normalize[of p] by (auto simp add: p0)
  7.1074 +      hence th1: "\<forall>x. poly p x \<noteq> 0" by simp
  7.1075 +      from k(2,3) dp have "poly (pexp q (degree p)) = poly (pmult p [1/k]) "
  7.1076 +	by - (rule ext, simp add: poly_mult poly_exp)
  7.1077 +      hence th2: "p divides (pexp q (degree p))" unfolding divides_def by blast
  7.1078 +      from th1 th2 pe have ?thesis by blast}
  7.1079 +    moreover
  7.1080 +    {assume dp: "degree p \<noteq> 0"
  7.1081 +      then obtain n where n: "degree p = Suc n " by (cases "degree p", auto)
  7.1082 +      {assume "p divides (pexp q (Suc n))"
  7.1083 +	then obtain u where u: "poly (pexp q (Suc n)) = poly (pmult p u)"
  7.1084 +	  unfolding divides_def by blast
  7.1085 +	hence u' :"\<And>x. poly (pexp q (Suc n)) x = poly (pmult p u) x" by simp_all
  7.1086 +	{fix x assume h: "poly p x = 0" "poly q x \<noteq> 0"
  7.1087 +	  hence "poly (pexp q (Suc n)) x \<noteq> 0" by (simp only: poly_exp) simp	  
  7.1088 +	  hence False using u' h(1) by (simp only: poly_mult poly_exp) simp}}
  7.1089 +	with n nullstellensatz_lemma[of p q "degree p"] dp 
  7.1090 +	have ?thesis by auto}
  7.1091 +    ultimately have ?thesis by blast}
  7.1092 +  ultimately show ?thesis by blast
  7.1093 +qed
  7.1094 +
  7.1095 +text{* Useful lemma *}
  7.1096 +
  7.1097 +lemma (in idom_char_0) constant_degree: "constant (poly p) \<longleftrightarrow> degree p = 0" (is "?lhs = ?rhs")
  7.1098 +proof
  7.1099 +  assume l: ?lhs
  7.1100 +  from l[unfolded constant_def, rule_format, of _ "zero"]
  7.1101 +  have th: "poly p = poly [poly p 0]" apply - by (rule ext, simp)
  7.1102 +  from degree_unique[OF th] show ?rhs by (simp add: degree_def)
  7.1103 +next
  7.1104 +  assume r: ?rhs
  7.1105 +  from r have "pnormalize p = [] \<or> (\<exists>k. pnormalize p = [k])"
  7.1106 +    unfolding degree_def by (cases "pnormalize p", auto)
  7.1107 +  then show ?lhs unfolding constant_def poly_normalize[of p, symmetric]
  7.1108 +    by (auto simp del: poly_normalize)
  7.1109 +qed
  7.1110 +
  7.1111 +(* It would be nicer to prove this without using algebraic closure...        *)
  7.1112 +
  7.1113 +lemma divides_degree_lemma: assumes dpn: "degree (p::complex list) = n"
  7.1114 +  shows "n \<le> degree (p *** q) \<or> poly (p *** q) = poly []"
  7.1115 +  using dpn
  7.1116 +proof(induct n arbitrary: p q)
  7.1117 +  case 0 thus ?case by simp
  7.1118 +next
  7.1119 +  case (Suc n p q)
  7.1120 +  from Suc.prems fundamental_theorem_of_algebra[of p] constant_degree[of p]
  7.1121 +  obtain a where a: "poly p a = 0" by auto
  7.1122 +  then obtain r where r: "p = pmult [-a, 1] r" unfolding poly_linear_divides
  7.1123 +    using Suc.prems by (auto simp add: degree_def)
  7.1124 +  {assume h: "poly (pmult r q) = poly []"
  7.1125 +    hence "poly (pmult p q) = poly []" using r
  7.1126 +      apply - apply (rule ext)  by (auto simp add: poly_entire poly_mult poly_add poly_cmult) hence ?case by blast}
  7.1127 +  moreover
  7.1128 +  {assume h: "poly (pmult r q) \<noteq> poly []" 
  7.1129 +    hence r0: "poly r \<noteq> poly []" and q0: "poly q \<noteq> poly []"
  7.1130 +      by (auto simp add: poly_entire)
  7.1131 +    have eq: "poly (pmult p q) = poly (pmult [-a, 1] (pmult r q))"
  7.1132 +      apply - apply (rule ext)
  7.1133 +      by (simp add: r poly_mult poly_add poly_cmult ring_simps)
  7.1134 +    from linear_mul_degree[OF h, of "- a"]
  7.1135 +    have dqe: "degree (pmult p q) = degree (pmult r q) + 1"
  7.1136 +      unfolding degree_unique[OF eq] .
  7.1137 +    from linear_mul_degree[OF r0, of "- a", unfolded r[symmetric]] r Suc.prems 
  7.1138 +    have dr: "degree r = n" by auto
  7.1139 +    from  Suc.hyps[OF dr, of q] have "Suc n \<le> degree (pmult p q)"
  7.1140 +      unfolding dqe using h by (auto simp del: poly.simps) 
  7.1141 +    hence ?case by blast}
  7.1142 +  ultimately show ?case by blast
  7.1143 +qed
  7.1144 +
  7.1145 +lemma divides_degree: assumes pq: "p divides (q:: complex list)"
  7.1146 +  shows "degree p \<le> degree q \<or> poly q = poly []"
  7.1147 +using pq  divides_degree_lemma[OF refl, of p]
  7.1148 +apply (auto simp add: divides_def poly_entire)
  7.1149 +apply atomize
  7.1150 +apply (erule_tac x="qa" in allE, auto)
  7.1151 +apply (subgoal_tac "degree q = degree (p *** qa)", simp)
  7.1152 +apply (rule degree_unique, simp)
  7.1153 +done
  7.1154 +
  7.1155 +(* Arithmetic operations on multivariate polynomials.                        *)
  7.1156 +
  7.1157 +lemma mpoly_base_conv: 
  7.1158 +  "(0::complex) \<equiv> poly [] x" "c \<equiv> poly [c] x" "x \<equiv> poly [0,1] x" by simp_all
  7.1159 +
  7.1160 +lemma mpoly_norm_conv: 
  7.1161 +  "poly [0] (x::complex) \<equiv> poly [] x" "poly [poly [] y] x \<equiv> poly [] x" by simp_all
  7.1162 +
  7.1163 +lemma mpoly_sub_conv: 
  7.1164 +  "poly p (x::complex) - poly q x \<equiv> poly p x + -1 * poly q x"
  7.1165 +  by (simp add: diff_def)
  7.1166 +
  7.1167 +lemma poly_pad_rule: "poly p x = 0 ==> poly (0#p) x = (0::complex)" by simp
  7.1168 +
  7.1169 +lemma poly_cancel_eq_conv: "p = (0::complex) \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (q = 0) \<equiv> (a * q - b * p = 0)" apply (atomize (full)) by auto
  7.1170 +
  7.1171 +lemma resolve_eq_raw:  "poly [] x \<equiv> 0" "poly [c] x \<equiv> (c::complex)" by auto
  7.1172 +lemma  resolve_eq_then: "(P \<Longrightarrow> (Q \<equiv> Q1)) \<Longrightarrow> (\<not>P \<Longrightarrow> (Q \<equiv> Q2))
  7.1173 +  \<Longrightarrow> Q \<equiv> P \<and> Q1 \<or> \<not>P\<and> Q2" apply (atomize (full)) by blast 
  7.1174 +lemma expand_ex_beta_conv: "list_ex P [c] \<equiv> P c" by simp
  7.1175 +
  7.1176 +lemma poly_divides_pad_rule: 
  7.1177 +  fixes p q :: "complex list"
  7.1178 +  assumes pq: "p divides q"
  7.1179 +  shows "p divides ((0::complex)#q)"
  7.1180 +proof-
  7.1181 +  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
  7.1182 +  hence "poly (0#q) = poly (p *** ([0,1] *** r))" 
  7.1183 +    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
  7.1184 +  thus ?thesis unfolding divides_def by blast
  7.1185 +qed
  7.1186 +
  7.1187 +lemma poly_divides_pad_const_rule: 
  7.1188 +  fixes p q :: "complex list"
  7.1189 +  assumes pq: "p divides q"
  7.1190 +  shows "p divides (a %* q)"
  7.1191 +proof-
  7.1192 +  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
  7.1193 +  hence "poly (a %* q) = poly (p *** (a %* r))" 
  7.1194 +    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
  7.1195 +  thus ?thesis unfolding divides_def by blast
  7.1196 +qed
  7.1197 +
  7.1198 +
  7.1199 +lemma poly_divides_conv0:  
  7.1200 +  fixes p :: "complex list"
  7.1201 +  assumes lgpq: "length q < length p" and lq:"last p \<noteq> 0"
  7.1202 +  shows "p divides q \<equiv> (\<not> (list_ex (\<lambda>c. c \<noteq> 0) q))" (is "?lhs \<equiv> ?rhs")
  7.1203 +proof-
  7.1204 +  {assume r: ?rhs 
  7.1205 +    hence eq: "poly q = poly []" unfolding poly_zero 
  7.1206 +      by (simp add: list_all_iff list_ex_iff)
  7.1207 +    hence "poly q = poly (p *** [])" by - (rule ext, simp add: poly_mult)
  7.1208 +    hence ?lhs unfolding divides_def  by blast}
  7.1209 +  moreover
  7.1210 +  {assume l: ?lhs
  7.1211 +    have ath: "\<And>lq lp dq::nat. lq < lp ==> lq \<noteq> 0 \<Longrightarrow> dq <= lq - 1 ==> dq < lp - 1"
  7.1212 +      by arith
  7.1213 +    {assume q0: "length q = 0"
  7.1214 +      hence "q = []" by simp
  7.1215 +      hence ?rhs by simp}
  7.1216 +    moreover
  7.1217 +    {assume lgq0: "length q \<noteq> 0"
  7.1218 +      from pnormalize_length[of q] have dql: "degree q \<le> length q - 1" 
  7.1219 +	unfolding degree_def by simp
  7.1220 +      from ath[OF lgpq lgq0 dql, unfolded pnormal_degree[OF lq, symmetric]] divides_degree[OF l] have "poly q = poly []" by auto
  7.1221 +      hence ?rhs unfolding poly_zero by (simp add: list_all_iff list_ex_iff)}
  7.1222 +    ultimately have ?rhs by blast }
  7.1223 +  ultimately show "?lhs \<equiv> ?rhs" by - (atomize (full), blast) 
  7.1224 +qed
  7.1225 +
  7.1226 +lemma poly_divides_conv1: 
  7.1227 +  assumes a0: "a\<noteq> (0::complex)" and pp': "(p::complex list) divides p'"
  7.1228 +  and qrp': "\<And>x. a * poly q x - poly p' x \<equiv> poly r x"
  7.1229 +  shows "p divides q \<equiv> p divides (r::complex list)" (is "?lhs \<equiv> ?rhs")
  7.1230 +proof-
  7.1231 +  {
  7.1232 +  from pp' obtain t where t: "poly p' = poly (p *** t)" 
  7.1233 +    unfolding divides_def by blast
  7.1234 +  {assume l: ?lhs
  7.1235 +    then obtain u where u: "poly q = poly (p *** u)" unfolding divides_def by blast
  7.1236 +     have "poly r = poly (p *** ((a %* u) +++ (-- t)))"
  7.1237 +       using u qrp' t
  7.1238 +       by - (rule ext, 
  7.1239 +	 simp add: poly_add poly_mult poly_cmult poly_minus ring_simps)
  7.1240 +     then have ?rhs unfolding divides_def by blast}
  7.1241 +  moreover
  7.1242 +  {assume r: ?rhs
  7.1243 +    then obtain u where u: "poly r = poly (p *** u)" unfolding divides_def by blast
  7.1244 +    from u t qrp' a0 have "poly q = poly (p *** ((1/a) %* (u +++ t)))"
  7.1245 +      by - (rule ext, atomize (full), simp add: poly_mult poly_add poly_cmult field_simps)
  7.1246 +    hence ?lhs  unfolding divides_def by blast}
  7.1247 +  ultimately have "?lhs = ?rhs" by blast }
  7.1248 +thus "?lhs \<equiv> ?rhs"  by - (atomize(full), blast) 
  7.1249 +qed
  7.1250 +
  7.1251 +lemma basic_cqe_conv1:
  7.1252 +  "(\<exists>x. poly p x = 0 \<and> poly [] x \<noteq> 0) \<equiv> False"
  7.1253 +  "(\<exists>x. poly [] x \<noteq> 0) \<equiv> False"
  7.1254 +  "(\<exists>x. poly [c] x \<noteq> 0) \<equiv> c\<noteq>0"
  7.1255 +  "(\<exists>x. poly [] x = 0) \<equiv> True"
  7.1256 +  "(\<exists>x. poly [c] x = 0) \<equiv> c = 0" by simp_all
  7.1257 +
  7.1258 +lemma basic_cqe_conv2: 
  7.1259 +  assumes l:"last (a#b#p) \<noteq> 0" 
  7.1260 +  shows "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True"
  7.1261 +proof-
  7.1262 +  {fix h t
  7.1263 +    assume h: "h\<noteq>0" "list_all (\<lambda>c. c=(0::complex)) t"  "a#b#p = h#t"
  7.1264 +    hence "list_all (\<lambda>c. c= 0) (b#p)" by simp
  7.1265 +    moreover have "last (b#p) \<in> set (b#p)" by simp
  7.1266 +    ultimately have "last (b#p) = 0" by (simp add: list_all_iff)
  7.1267 +    with l have False by simp}
  7.1268 +  hence th: "\<not> (\<exists> h t. h\<noteq>0 \<and> list_all (\<lambda>c. c=0) t \<and> a#b#p = h#t)"
  7.1269 +    by blast
  7.1270 +  from fundamental_theorem_of_algebra_alt[OF th] 
  7.1271 +  show "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True" by auto
  7.1272 +qed
  7.1273 +
  7.1274 +lemma  basic_cqe_conv_2b: "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
  7.1275 +proof-
  7.1276 +  have "\<not> (list_ex (\<lambda>c. c \<noteq> 0) p) \<longleftrightarrow> poly p = poly []" 
  7.1277 +    by (simp add: poly_zero list_all_iff list_ex_iff)
  7.1278 +  also have "\<dots> \<longleftrightarrow> (\<not> (\<exists>x. poly p x \<noteq> 0))" by (auto intro: ext)
  7.1279 +  finally show "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
  7.1280 +    by - (atomize (full), blast)
  7.1281 +qed
  7.1282 +
  7.1283 +lemma basic_cqe_conv3:
  7.1284 +  fixes p q :: "complex list"
  7.1285 +  assumes l: "last (a#p) \<noteq> 0" 
  7.1286 +  shows "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
  7.1287 +proof-
  7.1288 +  note np = pnormalize_eq[OF l]
  7.1289 +  {assume "poly (a#p) = poly []" hence False using l
  7.1290 +      unfolding poly_zero apply (auto simp add: list_all_iff del: last.simps)
  7.1291 +      apply (cases p, simp_all) done}
  7.1292 +  then have p0: "poly (a#p) \<noteq> poly []"  by blast
  7.1293 +  from np have dp:"degree (a#p) = length p" by (simp add: degree_def)
  7.1294 +  from nullstellensatz_univariate[of "a#p" q] p0 dp
  7.1295 +  show "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
  7.1296 +    by - (atomize (full), auto)
  7.1297 +qed
  7.1298 +
  7.1299 +lemma basic_cqe_conv4:
  7.1300 +  fixes p q :: "complex list"
  7.1301 +  assumes h: "\<And>x. poly (q %^ n) x \<equiv> poly r x"
  7.1302 +  shows "p divides (q %^ n) \<equiv> p divides r"
  7.1303 +proof-
  7.1304 +  from h have "poly (q %^ n) = poly r" by (auto intro: ext)  
  7.1305 +  thus "p divides (q %^ n) \<equiv> p divides r" unfolding divides_def by simp
  7.1306 +qed
  7.1307 +
  7.1308 +lemma pmult_Cons_Cons: "((a::complex)#b#p) *** q = (a %*q) +++ (0#((b#p) *** q))"
  7.1309 +  by simp
  7.1310 +
  7.1311 +lemma elim_neg_conv: "- z \<equiv> (-1) * (z::complex)" by simp
  7.1312 +lemma eqT_intr: "PROP P \<Longrightarrow> (True \<Longrightarrow> PROP P )" "PROP P \<Longrightarrow> True" by blast+
  7.1313 +lemma negate_negate_rule: "Trueprop P \<equiv> \<not> P \<equiv> False" by (atomize (full), auto)
  7.1314 +lemma last_simps: "last [x] = x" "last (x#y#ys) = last (y#ys)" by simp_all
  7.1315 +lemma length_simps: "length [] = 0" "length (x#y#xs) = length xs + 2" "length [x] = 1" by simp_all
  7.1316 +
  7.1317 +lemma complex_entire: "(z::complex) \<noteq> 0 \<and> w \<noteq> 0 \<equiv> z*w \<noteq> 0" by simp
  7.1318 +lemma resolve_eq_ne: "(P \<equiv> True) \<equiv> (\<not>P \<equiv> False)" "(P \<equiv> False) \<equiv> (\<not>P \<equiv> True)" 
  7.1319 +  by (atomize (full)) simp_all
  7.1320 +lemma cqe_conv1: "poly [] x = 0 \<longleftrightarrow> True"  by simp
  7.1321 +lemma cqe_conv2: "(p \<Longrightarrow> (q \<equiv> r)) \<equiv> ((p \<and> q) \<equiv> (p \<and> r))"  (is "?l \<equiv> ?r")
  7.1322 +proof
  7.1323 +  assume "p \<Longrightarrow> q \<equiv> r" thus "p \<and> q \<equiv> p \<and> r" apply - apply (atomize (full)) by blast
  7.1324 +next
  7.1325 +  assume "p \<and> q \<equiv> p \<and> r" "p"
  7.1326 +  thus "q \<equiv> r" apply - apply (atomize (full)) apply blast done
  7.1327 +qed
  7.1328 +lemma poly_const_conv: "poly [c] (x::complex) = y \<longleftrightarrow> c = y" by simp
  7.1329 +
  7.1330 +end
  7.1331 \ No newline at end of file
     8.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     8.2 +++ b/src/HOL/HahnBanach/Bounds.thy	Mon Dec 29 14:08:08 2008 +0100
     8.3 @@ -0,0 +1,83 @@
     8.4 +(*  Title:      HOL/Real/HahnBanach/Bounds.thy
     8.5 +    ID:         $Id$
     8.6 +    Author:     Gertrud Bauer, TU Munich
     8.7 +*)
     8.8 +
     8.9 +header {* Bounds *}
    8.10 +
    8.11 +theory Bounds
    8.12 +imports Main ContNotDenum
    8.13 +begin
    8.14 +
    8.15 +locale lub =
    8.16 +  fixes A and x
    8.17 +  assumes least [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> a \<le> b) \<Longrightarrow> x \<le> b"
    8.18 +    and upper [intro?]: "a \<in> A \<Longrightarrow> a \<le> x"
    8.19 +
    8.20 +lemmas [elim?] = lub.least lub.upper
    8.21 +
    8.22 +definition
    8.23 +  the_lub :: "'a::order set \<Rightarrow> 'a" where
    8.24 +  "the_lub A = The (lub A)"
    8.25 +
    8.26 +notation (xsymbols)
    8.27 +  the_lub  ("\<Squnion>_" [90] 90)
    8.28 +
    8.29 +lemma the_lub_equality [elim?]:
    8.30 +  assumes "lub A x"
    8.31 +  shows "\<Squnion>A = (x::'a::order)"
    8.32 +proof -
    8.33 +  interpret lub [A x] by fact
    8.34 +  show ?thesis
    8.35 +  proof (unfold the_lub_def)
    8.36 +    from `lub A x` show "The (lub A) = x"
    8.37 +    proof
    8.38 +      fix x' assume lub': "lub A x'"
    8.39 +      show "x' = x"
    8.40 +      proof (rule order_antisym)
    8.41 +	from lub' show "x' \<le> x"
    8.42 +	proof
    8.43 +          fix a assume "a \<in> A"
    8.44 +          then show "a \<le> x" ..
    8.45 +	qed
    8.46 +	show "x \<le> x'"
    8.47 +	proof
    8.48 +          fix a assume "a \<in> A"
    8.49 +          with lub' show "a \<le> x'" ..
    8.50 +	qed
    8.51 +      qed
    8.52 +    qed
    8.53 +  qed
    8.54 +qed
    8.55 +
    8.56 +lemma the_lubI_ex:
    8.57 +  assumes ex: "\<exists>x. lub A x"
    8.58 +  shows "lub A (\<Squnion>A)"
    8.59 +proof -
    8.60 +  from ex obtain x where x: "lub A x" ..
    8.61 +  also from x have [symmetric]: "\<Squnion>A = x" ..
    8.62 +  finally show ?thesis .
    8.63 +qed
    8.64 +
    8.65 +lemma lub_compat: "lub A x = isLub UNIV A x"
    8.66 +proof -
    8.67 +  have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)"
    8.68 +    by (rule ext) (simp only: isUb_def)
    8.69 +  then show ?thesis
    8.70 +    by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast
    8.71 +qed
    8.72 +
    8.73 +lemma real_complete:
    8.74 +  fixes A :: "real set"
    8.75 +  assumes nonempty: "\<exists>a. a \<in> A"
    8.76 +    and ex_upper: "\<exists>y. \<forall>a \<in> A. a \<le> y"
    8.77 +  shows "\<exists>x. lub A x"
    8.78 +proof -
    8.79 +  from ex_upper have "\<exists>y. isUb UNIV A y"
    8.80 +    unfolding isUb_def setle_def by blast
    8.81 +  with nonempty have "\<exists>x. isLub UNIV A x"
    8.82 +    by (rule reals_complete)
    8.83 +  then show ?thesis by (simp only: lub_compat)
    8.84 +qed
    8.85 +
    8.86 +end
     9.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     9.2 +++ b/src/HOL/HahnBanach/FunctionNorm.thy	Mon Dec 29 14:08:08 2008 +0100
     9.3 @@ -0,0 +1,279 @@
     9.4 +(*  Title:      HOL/Real/HahnBanach/FunctionNorm.thy
     9.5 +    ID:         $Id$
     9.6 +    Author:     Gertrud Bauer, TU Munich
     9.7 +*)
     9.8 +
     9.9 +header {* The norm of a function *}
    9.10 +
    9.11 +theory FunctionNorm
    9.12 +imports NormedSpace FunctionOrder
    9.13 +begin
    9.14 +
    9.15 +subsection {* Continuous linear forms*}
    9.16 +
    9.17 +text {*
    9.18 +  A linear form @{text f} on a normed vector space @{text "(V, \<parallel>\<cdot>\<parallel>)"}
    9.19 +  is \emph{continuous}, iff it is bounded, i.e.
    9.20 +  \begin{center}
    9.21 +  @{text "\<exists>c \<in> R. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
    9.22 +  \end{center}
    9.23 +  In our application no other functions than linear forms are
    9.24 +  considered, so we can define continuous linear forms as bounded
    9.25 +  linear forms:
    9.26 +*}
    9.27 +
    9.28 +locale continuous = var V + norm_syntax + linearform +
    9.29 +  assumes bounded: "\<exists>c. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
    9.30 +
    9.31 +declare continuous.intro [intro?] continuous_axioms.intro [intro?]
    9.32 +
    9.33 +lemma continuousI [intro]:
    9.34 +  fixes norm :: "_ \<Rightarrow> real"  ("\<parallel>_\<parallel>")
    9.35 +  assumes "linearform V f"
    9.36 +  assumes r: "\<And>x. x \<in> V \<Longrightarrow> \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
    9.37 +  shows "continuous V norm f"
    9.38 +proof
    9.39 +  show "linearform V f" by fact
    9.40 +  from r have "\<exists>c. \<forall>x\<in>V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" by blast
    9.41 +  then show "continuous_axioms V norm f" ..
    9.42 +qed
    9.43 +
    9.44 +
    9.45 +subsection {* The norm of a linear form *}
    9.46 +
    9.47 +text {*
    9.48 +  The least real number @{text c} for which holds
    9.49 +  \begin{center}
    9.50 +  @{text "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
    9.51 +  \end{center}
    9.52 +  is called the \emph{norm} of @{text f}.
    9.53 +
    9.54 +  For non-trivial vector spaces @{text "V \<noteq> {0}"} the norm can be
    9.55 +  defined as
    9.56 +  \begin{center}
    9.57 +  @{text "\<parallel>f\<parallel> = \<sup>x \<noteq> 0. \<bar>f x\<bar> / \<parallel>x\<parallel>"}
    9.58 +  \end{center}
    9.59 +
    9.60 +  For the case @{text "V = {0}"} the supremum would be taken from an
    9.61 +  empty set. Since @{text \<real>} is unbounded, there would be no supremum.
    9.62 +  To avoid this situation it must be guaranteed that there is an
    9.63 +  element in this set. This element must be @{text "{} \<ge> 0"} so that
    9.64 +  @{text fn_norm} has the norm properties. Furthermore it does not
    9.65 +  have to change the norm in all other cases, so it must be @{text 0},
    9.66 +  as all other elements are @{text "{} \<ge> 0"}.
    9.67 +
    9.68 +  Thus we define the set @{text B} where the supremum is taken from as
    9.69 +  follows:
    9.70 +  \begin{center}
    9.71 +  @{text "{0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel>. x \<noteq> 0 \<and> x \<in> F}"}
    9.72 +  \end{center}
    9.73 +
    9.74 +  @{text fn_norm} is equal to the supremum of @{text B}, if the
    9.75 +  supremum exists (otherwise it is undefined).
    9.76 +*}
    9.77 +
    9.78 +locale fn_norm = norm_syntax +
    9.79 +  fixes B defines "B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
    9.80 +  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
    9.81 +  defines "\<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
    9.82 +
    9.83 +locale normed_vectorspace_with_fn_norm = normed_vectorspace + fn_norm
    9.84 +
    9.85 +lemma (in fn_norm) B_not_empty [intro]: "0 \<in> B V f"
    9.86 +  by (simp add: B_def)
    9.87 +
    9.88 +text {*
    9.89 +  The following lemma states that every continuous linear form on a
    9.90 +  normed space @{text "(V, \<parallel>\<cdot>\<parallel>)"} has a function norm.
    9.91 +*}
    9.92 +
    9.93 +lemma (in normed_vectorspace_with_fn_norm) fn_norm_works:
    9.94 +  assumes "continuous V norm f"
    9.95 +  shows "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
    9.96 +proof -
    9.97 +  interpret continuous [V norm f] by fact
    9.98 +  txt {* The existence of the supremum is shown using the
    9.99 +    completeness of the reals. Completeness means, that every
   9.100 +    non-empty bounded set of reals has a supremum. *}
   9.101 +  have "\<exists>a. lub (B V f) a"
   9.102 +  proof (rule real_complete)
   9.103 +    txt {* First we have to show that @{text B} is non-empty: *}
   9.104 +    have "0 \<in> B V f" ..
   9.105 +    then show "\<exists>x. x \<in> B V f" ..
   9.106 +
   9.107 +    txt {* Then we have to show that @{text B} is bounded: *}
   9.108 +    show "\<exists>c. \<forall>y \<in> B V f. y \<le> c"
   9.109 +    proof -
   9.110 +      txt {* We know that @{text f} is bounded by some value @{text c}. *}
   9.111 +      from bounded obtain c where c: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
   9.112 +
   9.113 +      txt {* To prove the thesis, we have to show that there is some
   9.114 +        @{text b}, such that @{text "y \<le> b"} for all @{text "y \<in>
   9.115 +        B"}. Due to the definition of @{text B} there are two cases. *}
   9.116 +
   9.117 +      def b \<equiv> "max c 0"
   9.118 +      have "\<forall>y \<in> B V f. y \<le> b"
   9.119 +      proof
   9.120 +        fix y assume y: "y \<in> B V f"
   9.121 +        show "y \<le> b"
   9.122 +        proof cases
   9.123 +          assume "y = 0"
   9.124 +          then show ?thesis unfolding b_def by arith
   9.125 +        next
   9.126 +          txt {* The second case is @{text "y = \<bar>f x\<bar> / \<parallel>x\<parallel>"} for some
   9.127 +            @{text "x \<in> V"} with @{text "x \<noteq> 0"}. *}
   9.128 +          assume "y \<noteq> 0"
   9.129 +          with y obtain x where y_rep: "y = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
   9.130 +              and x: "x \<in> V" and neq: "x \<noteq> 0"
   9.131 +            by (auto simp add: B_def real_divide_def)
   9.132 +          from x neq have gt: "0 < \<parallel>x\<parallel>" ..
   9.133 +
   9.134 +          txt {* The thesis follows by a short calculation using the
   9.135 +            fact that @{text f} is bounded. *}
   9.136 +
   9.137 +          note y_rep
   9.138 +          also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
   9.139 +          proof (rule mult_right_mono)
   9.140 +            from c x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
   9.141 +            from gt have "0 < inverse \<parallel>x\<parallel>" 
   9.142 +              by (rule positive_imp_inverse_positive)
   9.143 +            then show "0 \<le> inverse \<parallel>x\<parallel>" by (rule order_less_imp_le)
   9.144 +          qed
   9.145 +          also have "\<dots> = c * (\<parallel>x\<parallel> * inverse \<parallel>x\<parallel>)"
   9.146 +            by (rule real_mult_assoc)
   9.147 +          also
   9.148 +          from gt have "\<parallel>x\<parallel> \<noteq> 0" by simp
   9.149 +          then have "\<parallel>x\<parallel> * inverse \<parallel>x\<parallel> = 1" by simp 
   9.150 +          also have "c * 1 \<le> b" by (simp add: b_def le_maxI1)
   9.151 +          finally show "y \<le> b" .
   9.152 +        qed
   9.153 +      qed
   9.154 +      then show ?thesis ..
   9.155 +    qed
   9.156 +  qed
   9.157 +  then show ?thesis unfolding fn_norm_def by (rule the_lubI_ex)
   9.158 +qed
   9.159 +
   9.160 +lemma (in normed_vectorspace_with_fn_norm) fn_norm_ub [iff?]:
   9.161 +  assumes "continuous V norm f"
   9.162 +  assumes b: "b \<in> B V f"
   9.163 +  shows "b \<le> \<parallel>f\<parallel>\<hyphen>V"
   9.164 +proof -
   9.165 +  interpret continuous [V norm f] by fact
   9.166 +  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
   9.167 +    using `continuous V norm f` by (rule fn_norm_works)
   9.168 +  from this and b show ?thesis ..
   9.169 +qed
   9.170 +
   9.171 +lemma (in normed_vectorspace_with_fn_norm) fn_norm_leastB:
   9.172 +  assumes "continuous V norm f"
   9.173 +  assumes b: "\<And>b. b \<in> B V f \<Longrightarrow> b \<le> y"
   9.174 +  shows "\<parallel>f\<parallel>\<hyphen>V \<le> y"
   9.175 +proof -
   9.176 +  interpret continuous [V norm f] by fact
   9.177 +  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
   9.178 +    using `continuous V norm f` by (rule fn_norm_works)
   9.179 +  from this and b show ?thesis ..
   9.180 +qed
   9.181 +
   9.182 +text {* The norm of a continuous function is always @{text "\<ge> 0"}. *}
   9.183 +
   9.184 +lemma (in normed_vectorspace_with_fn_norm) fn_norm_ge_zero [iff]:
   9.185 +  assumes "continuous V norm f"
   9.186 +  shows "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
   9.187 +proof -
   9.188 +  interpret continuous [V norm f] by fact
   9.189 +  txt {* The function norm is defined as the supremum of @{text B}.
   9.190 +    So it is @{text "\<ge> 0"} if all elements in @{text B} are @{text "\<ge>
   9.191 +    0"}, provided the supremum exists and @{text B} is not empty. *}
   9.192 +  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
   9.193 +    using `continuous V norm f` by (rule fn_norm_works)
   9.194 +  moreover have "0 \<in> B V f" ..
   9.195 +  ultimately show ?thesis ..
   9.196 +qed
   9.197 +
   9.198 +text {*
   9.199 +  \medskip The fundamental property of function norms is:
   9.200 +  \begin{center}
   9.201 +  @{text "\<bar>f x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
   9.202 +  \end{center}
   9.203 +*}
   9.204 +
   9.205 +lemma (in normed_vectorspace_with_fn_norm) fn_norm_le_cong:
   9.206 +  assumes "continuous V norm f" "linearform V f"
   9.207 +  assumes x: "x \<in> V"
   9.208 +  shows "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
   9.209 +proof -
   9.210 +  interpret continuous [V norm f] by fact
   9.211 +  interpret linearform [V f] .
   9.212 +  show ?thesis
   9.213 +  proof cases
   9.214 +    assume "x = 0"
   9.215 +    then have "\<bar>f x\<bar> = \<bar>f 0\<bar>" by simp
   9.216 +    also have "f 0 = 0" by rule unfold_locales
   9.217 +    also have "\<bar>\<dots>\<bar> = 0" by simp
   9.218 +    also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
   9.219 +      using `continuous V norm f` by (rule fn_norm_ge_zero)
   9.220 +    from x have "0 \<le> norm x" ..
   9.221 +    with a have "0 \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" by (simp add: zero_le_mult_iff)
   9.222 +    finally show "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" .
   9.223 +  next
   9.224 +    assume "x \<noteq> 0"
   9.225 +    with x have neq: "\<parallel>x\<parallel> \<noteq> 0" by simp
   9.226 +    then have "\<bar>f x\<bar> = (\<bar>f x\<bar> * inverse \<parallel>x\<parallel>) * \<parallel>x\<parallel>" by simp
   9.227 +    also have "\<dots> \<le>  \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
   9.228 +    proof (rule mult_right_mono)
   9.229 +      from x show "0 \<le> \<parallel>x\<parallel>" ..
   9.230 +      from x and neq have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<in> B V f"
   9.231 +	by (auto simp add: B_def real_divide_def)
   9.232 +      with `continuous V norm f` show "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>V"
   9.233 +	by (rule fn_norm_ub)
   9.234 +    qed
   9.235 +    finally show ?thesis .
   9.236 +  qed
   9.237 +qed
   9.238 +
   9.239 +text {*
   9.240 +  \medskip The function norm is the least positive real number for
   9.241 +  which the following inequation holds:
   9.242 +  \begin{center}
   9.243 +    @{text "\<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
   9.244 +  \end{center}
   9.245 +*}
   9.246 +
   9.247 +lemma (in normed_vectorspace_with_fn_norm) fn_norm_least [intro?]:
   9.248 +  assumes "continuous V norm f"
   9.249 +  assumes ineq: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" and ge: "0 \<le> c"
   9.250 +  shows "\<parallel>f\<parallel>\<hyphen>V \<le> c"
   9.251 +proof -
   9.252 +  interpret continuous [V norm f] by fact
   9.253 +  show ?thesis
   9.254 +  proof (rule fn_norm_leastB [folded B_def fn_norm_def])
   9.255 +    fix b assume b: "b \<in> B V f"
   9.256 +    show "b \<le> c"
   9.257 +    proof cases
   9.258 +      assume "b = 0"
   9.259 +      with ge show ?thesis by simp
   9.260 +    next
   9.261 +      assume "b \<noteq> 0"
   9.262 +      with b obtain x where b_rep: "b = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
   9.263 +        and x_neq: "x \<noteq> 0" and x: "x \<in> V"
   9.264 +	by (auto simp add: B_def real_divide_def)
   9.265 +      note b_rep
   9.266 +      also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
   9.267 +      proof (rule mult_right_mono)
   9.268 +	have "0 < \<parallel>x\<parallel>" using x x_neq ..
   9.269 +	then show "0 \<le> inverse \<parallel>x\<parallel>" by simp
   9.270 +	from ineq and x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
   9.271 +      qed
   9.272 +      also have "\<dots> = c"
   9.273 +      proof -
   9.274 +	from x_neq and x have "\<parallel>x\<parallel> \<noteq> 0" by simp
   9.275 +	then show ?thesis by simp
   9.276 +      qed
   9.277 +      finally show ?thesis .
   9.278 +    qed
   9.279 +  qed (insert `continuous V norm f`, simp_all add: continuous_def)
   9.280 +qed
   9.281 +
   9.282 +end
    10.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    10.2 +++ b/src/HOL/HahnBanach/FunctionOrder.thy	Mon Dec 29 14:08:08 2008 +0100
    10.3 @@ -0,0 +1,142 @@
    10.4 +(*  Title:      HOL/Real/HahnBanach/FunctionOrder.thy
    10.5 +    ID:         $Id$
    10.6 +    Author:     Gertrud Bauer, TU Munich
    10.7 +*)
    10.8 +
    10.9 +header {* An order on functions *}
   10.10 +
   10.11 +theory FunctionOrder
   10.12 +imports Subspace Linearform
   10.13 +begin
   10.14 +
   10.15 +subsection {* The graph of a function *}
   10.16 +
   10.17 +text {*
   10.18 +  We define the \emph{graph} of a (real) function @{text f} with
   10.19 +  domain @{text F} as the set
   10.20 +  \begin{center}
   10.21 +  @{text "{(x, f x). x \<in> F}"}
   10.22 +  \end{center}
   10.23 +  So we are modeling partial functions by specifying the domain and
   10.24 +  the mapping function. We use the term ``function'' also for its
   10.25 +  graph.
   10.26 +*}
   10.27 +
   10.28 +types 'a graph = "('a \<times> real) set"
   10.29 +
   10.30 +definition
   10.31 +  graph :: "'a set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a graph" where
   10.32 +  "graph F f = {(x, f x) | x. x \<in> F}"
   10.33 +
   10.34 +lemma graphI [intro]: "x \<in> F \<Longrightarrow> (x, f x) \<in> graph F f"
   10.35 +  unfolding graph_def by blast
   10.36 +
   10.37 +lemma graphI2 [intro?]: "x \<in> F \<Longrightarrow> \<exists>t \<in> graph F f. t = (x, f x)"
   10.38 +  unfolding graph_def by blast
   10.39 +
   10.40 +lemma graphE [elim?]:
   10.41 +    "(x, y) \<in> graph F f \<Longrightarrow> (x \<in> F \<Longrightarrow> y = f x \<Longrightarrow> C) \<Longrightarrow> C"
   10.42 +  unfolding graph_def by blast
   10.43 +
   10.44 +
   10.45 +subsection {* Functions ordered by domain extension *}
   10.46 +
   10.47 +text {*
   10.48 +  A function @{text h'} is an extension of @{text h}, iff the graph of
   10.49 +  @{text h} is a subset of the graph of @{text h'}.
   10.50 +*}
   10.51 +
   10.52 +lemma graph_extI:
   10.53 +  "(\<And>x. x \<in> H \<Longrightarrow> h x = h' x) \<Longrightarrow> H \<subseteq> H'
   10.54 +    \<Longrightarrow> graph H h \<subseteq> graph H' h'"
   10.55 +  unfolding graph_def by blast
   10.56 +
   10.57 +lemma graph_extD1 [dest?]:
   10.58 +  "graph H h \<subseteq> graph H' h' \<Longrightarrow> x \<in> H \<Longrightarrow> h x = h' x"
   10.59 +  unfolding graph_def by blast
   10.60 +
   10.61 +lemma graph_extD2 [dest?]:
   10.62 +  "graph H h \<subseteq> graph H' h' \<Longrightarrow> H \<subseteq> H'"
   10.63 +  unfolding graph_def by blast
   10.64 +
   10.65 +
   10.66 +subsection {* Domain and function of a graph *}
   10.67 +
   10.68 +text {*
   10.69 +  The inverse functions to @{text graph} are @{text domain} and @{text
   10.70 +  funct}.
   10.71 +*}
   10.72 +
   10.73 +definition
   10.74 +  "domain" :: "'a graph \<Rightarrow> 'a set" where
   10.75 +  "domain g = {x. \<exists>y. (x, y) \<in> g}"
   10.76 +
   10.77 +definition
   10.78 +  funct :: "'a graph \<Rightarrow> ('a \<Rightarrow> real)" where
   10.79 +  "funct g = (\<lambda>x. (SOME y. (x, y) \<in> g))"
   10.80 +
   10.81 +text {*
   10.82 +  The following lemma states that @{text g} is the graph of a function
   10.83 +  if the relation induced by @{text g} is unique.
   10.84 +*}
   10.85 +
   10.86 +lemma graph_domain_funct:
   10.87 +  assumes uniq: "\<And>x y z. (x, y) \<in> g \<Longrightarrow> (x, z) \<in> g \<Longrightarrow> z = y"
   10.88 +  shows "graph (domain g) (funct g) = g"
   10.89 +  unfolding domain_def funct_def graph_def
   10.90 +proof auto  (* FIXME !? *)
   10.91 +  fix a b assume g: "(a, b) \<in> g"
   10.92 +  from g show "(a, SOME y. (a, y) \<in> g) \<in> g" by (rule someI2)
   10.93 +  from g show "\<exists>y. (a, y) \<in> g" ..
   10.94 +  from g show "b = (SOME y. (a, y) \<in> g)"
   10.95 +  proof (rule some_equality [symmetric])
   10.96 +    fix y assume "(a, y) \<in> g"
   10.97 +    with g show "y = b" by (rule uniq)
   10.98 +  qed
   10.99 +qed
  10.100 +
  10.101 +
  10.102 +subsection {* Norm-preserving extensions of a function *}
  10.103 +
  10.104 +text {*
  10.105 +  Given a linear form @{text f} on the space @{text F} and a seminorm
  10.106 +  @{text p} on @{text E}. The set of all linear extensions of @{text
  10.107 +  f}, to superspaces @{text H} of @{text F}, which are bounded by
  10.108 +  @{text p}, is defined as follows.
  10.109 +*}
  10.110 +
  10.111 +definition
  10.112 +  norm_pres_extensions ::
  10.113 +    "'a::{plus, minus, uminus, zero} set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> real)
  10.114 +      \<Rightarrow> 'a graph set" where
  10.115 +    "norm_pres_extensions E p F f
  10.116 +      = {g. \<exists>H h. g = graph H h
  10.117 +          \<and> linearform H h
  10.118 +          \<and> H \<unlhd> E
  10.119 +          \<and> F \<unlhd> H
  10.120 +          \<and> graph F f \<subseteq> graph H h
  10.121 +          \<and> (\<forall>x \<in> H. h x \<le> p x)}"
  10.122 +
  10.123 +lemma norm_pres_extensionE [elim]:
  10.124 +  "g \<in> norm_pres_extensions E p F f
  10.125 +  \<Longrightarrow> (\<And>H h. g = graph H h \<Longrightarrow> linearform H h
  10.126 +        \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H \<Longrightarrow> graph F f \<subseteq> graph H h
  10.127 +        \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x \<Longrightarrow> C) \<Longrightarrow> C"
  10.128 +  unfolding norm_pres_extensions_def by blast
  10.129 +
  10.130 +lemma norm_pres_extensionI2 [intro]:
  10.131 +  "linearform H h \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H
  10.132 +    \<Longrightarrow> graph F f \<subseteq> graph H h \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x
  10.133 +    \<Longrightarrow> graph H h \<in> norm_pres_extensions E p F f"
  10.134 +  unfolding norm_pres_extensions_def by blast
  10.135 +
  10.136 +lemma norm_pres_extensionI:  (* FIXME ? *)
  10.137 +  "\<exists>H h. g = graph H h
  10.138 +    \<and> linearform H h
  10.139 +    \<and> H \<unlhd> E
  10.140 +    \<and> F \<unlhd> H
  10.141 +    \<and> graph F f \<subseteq> graph H h
  10.142 +    \<and> (\<forall>x \<in> H. h x \<le> p x) \<Longrightarrow> g \<in> norm_pres_extensions E p F f"
  10.143 +  unfolding norm_pres_extensions_def by blast
  10.144 +
  10.145 +end
    11.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    11.2 +++ b/src/HOL/HahnBanach/HahnBanach.thy	Mon Dec 29 14:08:08 2008 +0100
    11.3 @@ -0,0 +1,510 @@
    11.4 +(*  Title:      HOL/Real/HahnBanach/HahnBanach.thy
    11.5 +    ID:         $Id$
    11.6 +    Author:     Gertrud Bauer, TU Munich
    11.7 +*)
    11.8 +
    11.9 +header {* The Hahn-Banach Theorem *}
   11.10 +
   11.11 +theory HahnBanach
   11.12 +imports HahnBanachLemmas
   11.13 +begin
   11.14 +
   11.15 +text {*
   11.16 +  We present the proof of two different versions of the Hahn-Banach
   11.17 +  Theorem, closely following \cite[\S36]{Heuser:1986}.
   11.18 +*}
   11.19 +
   11.20 +subsection {* The Hahn-Banach Theorem for vector spaces *}
   11.21 +
   11.22 +text {*
   11.23 +  \textbf{Hahn-Banach Theorem.} Let @{text F} be a subspace of a real
   11.24 +  vector space @{text E}, let @{text p} be a semi-norm on @{text E},
   11.25 +  and @{text f} be a linear form defined on @{text F} such that @{text
   11.26 +  f} is bounded by @{text p}, i.e.  @{text "\<forall>x \<in> F. f x \<le> p x"}.  Then
   11.27 +  @{text f} can be extended to a linear form @{text h} on @{text E}
   11.28 +  such that @{text h} is norm-preserving, i.e. @{text h} is also
   11.29 +  bounded by @{text p}.
   11.30 +
   11.31 +  \bigskip
   11.32 +  \textbf{Proof Sketch.}
   11.33 +  \begin{enumerate}
   11.34 +
   11.35 +  \item Define @{text M} as the set of norm-preserving extensions of
   11.36 +  @{text f} to subspaces of @{text E}. The linear forms in @{text M}
   11.37 +  are ordered by domain extension.
   11.38 +
   11.39 +  \item We show that every non-empty chain in @{text M} has an upper
   11.40 +  bound in @{text M}.
   11.41 +
   11.42 +  \item With Zorn's Lemma we conclude that there is a maximal function
   11.43 +  @{text g} in @{text M}.
   11.44 +
   11.45 +  \item The domain @{text H} of @{text g} is the whole space @{text
   11.46 +  E}, as shown by classical contradiction:
   11.47 +
   11.48 +  \begin{itemize}
   11.49 +
   11.50 +  \item Assuming @{text g} is not defined on whole @{text E}, it can
   11.51 +  still be extended in a norm-preserving way to a super-space @{text
   11.52 +  H'} of @{text H}.
   11.53 +
   11.54 +  \item Thus @{text g} can not be maximal. Contradiction!
   11.55 +
   11.56 +  \end{itemize}
   11.57 +  \end{enumerate}
   11.58 +*}
   11.59 +
   11.60 +theorem HahnBanach:
   11.61 +  assumes E: "vectorspace E" and "subspace F E"
   11.62 +    and "seminorm E p" and "linearform F f"
   11.63 +  assumes fp: "\<forall>x \<in> F. f x \<le> p x"
   11.64 +  shows "\<exists>h. linearform E h \<and> (\<forall>x \<in> F. h x = f x) \<and> (\<forall>x \<in> E. h x \<le> p x)"
   11.65 +    -- {* Let @{text E} be a vector space, @{text F} a subspace of @{text E}, @{text p} a seminorm on @{text E}, *}
   11.66 +    -- {* and @{text f} a linear form on @{text F} such that @{text f} is bounded by @{text p}, *}
   11.67 +    -- {* then @{text f} can be extended to a linear form @{text h} on @{text E} in a norm-preserving way. \skp *}
   11.68 +proof -
   11.69 +  interpret vectorspace [E] by fact
   11.70 +  interpret subspace [F E] by fact
   11.71 +  interpret seminorm [E p] by fact
   11.72 +  interpret linearform [F f] by fact
   11.73 +  def M \<equiv> "norm_pres_extensions E p F f"
   11.74 +  then have M: "M = \<dots>" by (simp only:)
   11.75 +  from E have F: "vectorspace F" ..
   11.76 +  note FE = `F \<unlhd> E`
   11.77 +  {
   11.78 +    fix c assume cM: "c \<in> chain M" and ex: "\<exists>x. x \<in> c"
   11.79 +    have "\<Union>c \<in> M"
   11.80 +      -- {* Show that every non-empty chain @{text c} of @{text M} has an upper bound in @{text M}: *}
   11.81 +      -- {* @{text "\<Union>c"} is greater than any element of the chain @{text c}, so it suffices to show @{text "\<Union>c \<in> M"}. *}
   11.82 +      unfolding M_def
   11.83 +    proof (rule norm_pres_extensionI)
   11.84 +      let ?H = "domain (\<Union>c)"
   11.85 +      let ?h = "funct (\<Union>c)"
   11.86 +
   11.87 +      have a: "graph ?H ?h = \<Union>c"
   11.88 +      proof (rule graph_domain_funct)
   11.89 +        fix x y z assume "(x, y) \<in> \<Union>c" and "(x, z) \<in> \<Union>c"
   11.90 +        with M_def cM show "z = y" by (rule sup_definite)
   11.91 +      qed
   11.92 +      moreover from M cM a have "linearform ?H ?h"
   11.93 +        by (rule sup_lf)
   11.94 +      moreover from a M cM ex FE E have "?H \<unlhd> E"
   11.95 +        by (rule sup_subE)
   11.96 +      moreover from a M cM ex FE have "F \<unlhd> ?H"
   11.97 +        by (rule sup_supF)
   11.98 +      moreover from a M cM ex have "graph F f \<subseteq> graph ?H ?h"
   11.99 +        by (rule sup_ext)
  11.100 +      moreover from a M cM have "\<forall>x \<in> ?H. ?h x \<le> p x"
  11.101 +        by (rule sup_norm_pres)
  11.102 +      ultimately show "\<exists>H h. \<Union>c = graph H h
  11.103 +          \<and> linearform H h
  11.104 +          \<and> H \<unlhd> E
  11.105 +          \<and> F \<unlhd> H
  11.106 +          \<and> graph F f \<subseteq> graph H h
  11.107 +          \<and> (\<forall>x \<in> H. h x \<le> p x)" by blast
  11.108 +    qed
  11.109 +  }
  11.110 +  then have "\<exists>g \<in> M. \<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
  11.111 +  -- {* With Zorn's Lemma we can conclude that there is a maximal element in @{text M}. \skp *}
  11.112 +  proof (rule Zorn's_Lemma)
  11.113 +      -- {* We show that @{text M} is non-empty: *}
  11.114 +    show "graph F f \<in> M"
  11.115 +      unfolding M_def
  11.116 +    proof (rule norm_pres_extensionI2)
  11.117 +      show "linearform F f" by fact
  11.118 +      show "F \<unlhd> E" by fact
  11.119 +      from F show "F \<unlhd> F" by (rule vectorspace.subspace_refl)
  11.120 +      show "graph F f \<subseteq> graph F f" ..
  11.121 +      show "\<forall>x\<in>F. f x \<le> p x" by fact
  11.122 +    qed
  11.123 +  qed
  11.124 +  then obtain g where gM: "g \<in> M" and gx: "\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
  11.125 +    by blast
  11.126 +  from gM obtain H h where
  11.127 +      g_rep: "g = graph H h"
  11.128 +    and linearform: "linearform H h"
  11.129 +    and HE: "H \<unlhd> E" and FH: "F \<unlhd> H"
  11.130 +    and graphs: "graph F f \<subseteq> graph H h"
  11.131 +    and hp: "\<forall>x \<in> H. h x \<le> p x" unfolding M_def ..
  11.132 +      -- {* @{text g} is a norm-preserving extension of @{text f}, in other words: *}
  11.133 +      -- {* @{text g} is the graph of some linear form @{text h} defined on a subspace @{text H} of @{text E}, *}
  11.134 +      -- {* and @{text h} is an extension of @{text f} that is again bounded by @{text p}. \skp *}
  11.135 +  from HE E have H: "vectorspace H"
  11.136 +    by (rule subspace.vectorspace)
  11.137 +
  11.138 +  have HE_eq: "H = E"
  11.139 +    -- {* We show that @{text h} is defined on whole @{text E} by classical contradiction. \skp *}
  11.140 +  proof (rule classical)
  11.141 +    assume neq: "H \<noteq> E"
  11.142 +      -- {* Assume @{text h} is not defined on whole @{text E}. Then show that @{text h} can be extended *}
  11.143 +      -- {* in a norm-preserving way to a function @{text h'} with the graph @{text g'}. \skp *}
  11.144 +    have "\<exists>g' \<in> M. g \<subseteq> g' \<and> g \<noteq> g'"
  11.145 +    proof -
  11.146 +      from HE have "H \<subseteq> E" ..
  11.147 +      with neq obtain x' where x'E: "x' \<in> E" and "x' \<notin> H" by blast
  11.148 +      obtain x': "x' \<noteq> 0"
  11.149 +      proof
  11.150 +        show "x' \<noteq> 0"
  11.151 +        proof
  11.152 +          assume "x' = 0"
  11.153 +          with H have "x' \<in> H" by (simp only: vectorspace.zero)
  11.154 +          with `x' \<notin> H` show False by contradiction
  11.155 +        qed
  11.156 +      qed
  11.157 +
  11.158 +      def H' \<equiv> "H + lin x'"
  11.159 +        -- {* Define @{text H'} as the direct sum of @{text H} and the linear closure of @{text x'}. \skp *}
  11.160 +      have HH': "H \<unlhd> H'"
  11.161 +      proof (unfold H'_def)
  11.162 +        from x'E have "vectorspace (lin x')" ..
  11.163 +        with H show "H \<unlhd> H + lin x'" ..
  11.164 +      qed
  11.165 +
  11.166 +      obtain xi where
  11.167 +        xi: "\<forall>y \<in> H. - p (y + x') - h y \<le> xi
  11.168 +          \<and> xi \<le> p (y + x') - h y"
  11.169 +        -- {* Pick a real number @{text \<xi>} that fulfills certain inequations; this will *}
  11.170 +        -- {* be used to establish that @{text h'} is a norm-preserving extension of @{text h}.
  11.171 +           \label{ex-xi-use}\skp *}
  11.172 +      proof -
  11.173 +        from H have "\<exists>xi. \<forall>y \<in> H. - p (y + x') - h y \<le> xi
  11.174 +            \<and> xi \<le> p (y + x') - h y"
  11.175 +        proof (rule ex_xi)
  11.176 +          fix u v assume u: "u \<in> H" and v: "v \<in> H"
  11.177 +          with HE have uE: "u \<in> E" and vE: "v \<in> E" by auto
  11.178 +          from H u v linearform have "h v - h u = h (v - u)"
  11.179 +            by (simp add: linearform.diff)
  11.180 +          also from hp and H u v have "\<dots> \<le> p (v - u)"
  11.181 +            by (simp only: vectorspace.diff_closed)
  11.182 +          also from x'E uE vE have "v - u = x' + - x' + v + - u"
  11.183 +            by (simp add: diff_eq1)
  11.184 +          also from x'E uE vE have "\<dots> = v + x' + - (u + x')"
  11.185 +            by (simp add: add_ac)
  11.186 +          also from x'E uE vE have "\<dots> = (v + x') - (u + x')"
  11.187 +            by (simp add: diff_eq1)
  11.188 +          also from x'E uE vE E have "p \<dots> \<le> p (v + x') + p (u + x')"
  11.189 +            by (simp add: diff_subadditive)
  11.190 +          finally have "h v - h u \<le> p (v + x') + p (u + x')" .
  11.191 +          then show "- p (u + x') - h u \<le> p (v + x') - h v" by simp
  11.192 +        qed
  11.193 +        then show thesis by (blast intro: that)
  11.194 +      qed
  11.195 +
  11.196 +      def h' \<equiv> "\<lambda>x. let (y, a) =
  11.197 +          SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H in h y + a * xi"
  11.198 +        -- {* Define the extension @{text h'} of @{text h} to @{text H'} using @{text \<xi>}. \skp *}
  11.199 +
  11.200 +      have "g \<subseteq> graph H' h' \<and> g \<noteq> graph H' h'"
  11.201 +        -- {* @{text h'} is an extension of @{text h} \dots \skp *}
  11.202 +      proof
  11.203 +        show "g \<subseteq> graph H' h'"
  11.204 +        proof -
  11.205 +          have  "graph H h \<subseteq> graph H' h'"
  11.206 +          proof (rule graph_extI)
  11.207 +            fix t assume t: "t \<in> H"
  11.208 +            from E HE t have "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
  11.209 +	      using `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` by (rule decomp_H'_H)
  11.210 +            with h'_def show "h t = h' t" by (simp add: Let_def)
  11.211 +          next
  11.212 +            from HH' show "H \<subseteq> H'" ..
  11.213 +          qed
  11.214 +          with g_rep show ?thesis by (simp only:)
  11.215 +        qed
  11.216 +
  11.217 +        show "g \<noteq> graph H' h'"
  11.218 +        proof -
  11.219 +          have "graph H h \<noteq> graph H' h'"
  11.220 +          proof
  11.221 +            assume eq: "graph H h = graph H' h'"
  11.222 +            have "x' \<in> H'"
  11.223 +	      unfolding H'_def
  11.224 +            proof
  11.225 +              from H show "0 \<in> H" by (rule vectorspace.zero)
  11.226 +              from x'E show "x' \<in> lin x'" by (rule x_lin_x)
  11.227 +              from x'E show "x' = 0 + x'" by simp
  11.228 +            qed
  11.229 +            then have "(x', h' x') \<in> graph H' h'" ..
  11.230 +            with eq have "(x', h' x') \<in> graph H h" by (simp only:)
  11.231 +            then have "x' \<in> H" ..
  11.232 +            with `x' \<notin> H` show False by contradiction
  11.233 +          qed
  11.234 +          with g_rep show ?thesis by simp
  11.235 +        qed
  11.236 +      qed
  11.237 +      moreover have "graph H' h' \<in> M"
  11.238 +        -- {* and @{text h'} is norm-preserving. \skp *}
  11.239 +      proof (unfold M_def)
  11.240 +        show "graph H' h' \<in> norm_pres_extensions E p F f"
  11.241 +        proof (rule norm_pres_extensionI2)
  11.242 +          show "linearform H' h'"
  11.243 +	    using h'_def H'_def HE linearform `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E
  11.244 +	    by (rule h'_lf)
  11.245 +          show "H' \<unlhd> E"
  11.246 +	  unfolding H'_def
  11.247 +          proof
  11.248 +            show "H \<unlhd> E" by fact
  11.249 +            show "vectorspace E" by fact
  11.250 +            from x'E show "lin x' \<unlhd> E" ..
  11.251 +          qed
  11.252 +          from H `F \<unlhd> H` HH' show FH': "F \<unlhd> H'"
  11.253 +            by (rule vectorspace.subspace_trans)
  11.254 +          show "graph F f \<subseteq> graph H' h'"
  11.255 +          proof (rule graph_extI)
  11.256 +            fix x assume x: "x \<in> F"
  11.257 +            with graphs have "f x = h x" ..
  11.258 +            also have "\<dots> = h x + 0 * xi" by simp
  11.259 +            also have "\<dots> = (let (y, a) = (x, 0) in h y + a * xi)"
  11.260 +              by (simp add: Let_def)
  11.261 +            also have "(x, 0) =
  11.262 +                (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)"
  11.263 +	      using E HE
  11.264 +            proof (rule decomp_H'_H [symmetric])
  11.265 +              from FH x show "x \<in> H" ..
  11.266 +              from x' show "x' \<noteq> 0" .
  11.267 +	      show "x' \<notin> H" by fact
  11.268 +	      show "x' \<in> E" by fact
  11.269 +            qed
  11.270 +            also have
  11.271 +              "(let (y, a) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)
  11.272 +              in h y + a * xi) = h' x" by (simp only: h'_def)
  11.273 +            finally show "f x = h' x" .
  11.274 +          next
  11.275 +            from FH' show "F \<subseteq> H'" ..
  11.276 +          qed
  11.277 +          show "\<forall>x \<in> H'. h' x \<le> p x"
  11.278 +	    using h'_def H'_def `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E HE
  11.279 +	      `seminorm E p` linearform and hp xi
  11.280 +	    by (rule h'_norm_pres)
  11.281 +        qed
  11.282 +      qed
  11.283 +      ultimately show ?thesis ..
  11.284 +    qed
  11.285 +    then have "\<not> (\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x)" by simp
  11.286 +      -- {* So the graph @{text g} of @{text h} cannot be maximal. Contradiction! \skp *}
  11.287 +    with gx show "H = E" by contradiction
  11.288 +  qed
  11.289 +
  11.290 +  from HE_eq and linearform have "linearform E h"
  11.291 +    by (simp only:)
  11.292 +  moreover have "\<forall>x \<in> F. h x = f x"
  11.293 +  proof
  11.294 +    fix x assume "x \<in> F"
  11.295 +    with graphs have "f x = h x" ..
  11.296 +    then show "h x = f x" ..
  11.297 +  qed
  11.298 +  moreover from HE_eq and hp have "\<forall>x \<in> E. h x \<le> p x"
  11.299 +    by (simp only:)
  11.300 +  ultimately show ?thesis by blast
  11.301 +qed
  11.302 +
  11.303 +
  11.304 +subsection  {* Alternative formulation *}
  11.305 +
  11.306 +text {*
  11.307 +  The following alternative formulation of the Hahn-Banach
  11.308 +  Theorem\label{abs-HahnBanach} uses the fact that for a real linear
  11.309 +  form @{text f} and a seminorm @{text p} the following inequations
  11.310 +  are equivalent:\footnote{This was shown in lemma @{thm [source]
  11.311 +  abs_ineq_iff} (see page \pageref{abs-ineq-iff}).}
  11.312 +  \begin{center}
  11.313 +  \begin{tabular}{lll}
  11.314 +  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
  11.315 +  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
  11.316 +  \end{tabular}
  11.317 +  \end{center}
  11.318 +*}
  11.319 +
  11.320 +theorem abs_HahnBanach:
  11.321 +  assumes E: "vectorspace E" and FE: "subspace F E"
  11.322 +    and lf: "linearform F f" and sn: "seminorm E p"
  11.323 +  assumes fp: "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
  11.324 +  shows "\<exists>g. linearform E g
  11.325 +    \<and> (\<forall>x \<in> F. g x = f x)
  11.326 +    \<and> (\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x)"
  11.327 +proof -
  11.328 +  interpret vectorspace [E] by fact
  11.329 +  interpret subspace [F E] by fact
  11.330 +  interpret linearform [F f] by fact
  11.331 +  interpret seminorm [E p] by fact
  11.332 +  have "\<exists>g. linearform E g \<and> (\<forall>x \<in> F. g x = f x) \<and> (\<forall>x \<in> E. g x \<le> p x)"
  11.333 +    using E FE sn lf
  11.334 +  proof (rule HahnBanach)
  11.335 +    show "\<forall>x \<in> F. f x \<le> p x"
  11.336 +      using FE E sn lf and fp by (rule abs_ineq_iff [THEN iffD1])
  11.337 +  qed
  11.338 +  then obtain g where lg: "linearform E g" and *: "\<forall>x \<in> F. g x = f x"
  11.339 +      and **: "\<forall>x \<in> E. g x \<le> p x" by blast
  11.340 +  have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
  11.341 +    using _ E sn lg **
  11.342 +  proof (rule abs_ineq_iff [THEN iffD2])
  11.343 +    show "E \<unlhd> E" ..
  11.344 +  qed
  11.345 +  with lg * show ?thesis by blast
  11.346 +qed
  11.347 +
  11.348 +
  11.349 +subsection {* The Hahn-Banach Theorem for normed spaces *}
  11.350 +
  11.351 +text {*
  11.352 +  Every continuous linear form @{text f} on a subspace @{text F} of a
  11.353 +  norm space @{text E}, can be extended to a continuous linear form
  11.354 +  @{text g} on @{text E} such that @{text "\<parallel>f\<parallel> = \<parallel>g\<parallel>"}.
  11.355 +*}
  11.356 +
  11.357 +theorem norm_HahnBanach:
  11.358 +  fixes V and norm ("\<parallel>_\<parallel>")
  11.359 +  fixes B defines "\<And>V f. B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
  11.360 +  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
  11.361 +  defines "\<And>V f. \<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
  11.362 +  assumes E_norm: "normed_vectorspace E norm" and FE: "subspace F E"
  11.363 +    and linearform: "linearform F f" and "continuous F norm f"
  11.364 +  shows "\<exists>g. linearform E g
  11.365 +     \<and> continuous E norm g
  11.366 +     \<and> (\<forall>x \<in> F. g x = f x)
  11.367 +     \<and> \<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
  11.368 +proof -
  11.369 +  interpret normed_vectorspace [E norm] by fact
  11.370 +  interpret normed_vectorspace_with_fn_norm [E norm B fn_norm]
  11.371 +    by (auto simp: B_def fn_norm_def) intro_locales
  11.372 +  interpret subspace [F E] by fact
  11.373 +  interpret linearform [F f] by fact
  11.374 +  interpret continuous [F norm f] by fact
  11.375 +  have E: "vectorspace E" by intro_locales
  11.376 +  have F: "vectorspace F" by rule intro_locales
  11.377 +  have F_norm: "normed_vectorspace F norm"
  11.378 +    using FE E_norm by (rule subspace_normed_vs)
  11.379 +  have ge_zero: "0 \<le> \<parallel>f\<parallel>\<hyphen>F"
  11.380 +    by (rule normed_vectorspace_with_fn_norm.fn_norm_ge_zero
  11.381 +      [OF normed_vectorspace_with_fn_norm.intro,
  11.382 +       OF F_norm `continuous F norm f` , folded B_def fn_norm_def])
  11.383 +  txt {* We define a function @{text p} on @{text E} as follows:
  11.384 +    @{text "p x = \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"} *}
  11.385 +  def p \<equiv> "\<lambda>x. \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
  11.386 +
  11.387 +  txt {* @{text p} is a seminorm on @{text E}: *}
  11.388 +  have q: "seminorm E p"
  11.389 +  proof
  11.390 +    fix x y a assume x: "x \<in> E" and y: "y \<in> E"
  11.391 +    
  11.392 +    txt {* @{text p} is positive definite: *}
  11.393 +    have "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
  11.394 +    moreover from x have "0 \<le> \<parallel>x\<parallel>" ..
  11.395 +    ultimately show "0 \<le> p x"  
  11.396 +      by (simp add: p_def zero_le_mult_iff)
  11.397 +
  11.398 +    txt {* @{text p} is absolutely homogenous: *}
  11.399 +
  11.400 +    show "p (a \<cdot> x) = \<bar>a\<bar> * p x"
  11.401 +    proof -
  11.402 +      have "p (a \<cdot> x) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>a \<cdot> x\<parallel>" by (simp only: p_def)
  11.403 +      also from x have "\<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>" by (rule abs_homogenous)
  11.404 +      also have "\<parallel>f\<parallel>\<hyphen>F * (\<bar>a\<bar> * \<parallel>x\<parallel>) = \<bar>a\<bar> * (\<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>)" by simp
  11.405 +      also have "\<dots> = \<bar>a\<bar> * p x" by (simp only: p_def)
  11.406 +      finally show ?thesis .
  11.407 +    qed
  11.408 +
  11.409 +    txt {* Furthermore, @{text p} is subadditive: *}
  11.410 +
  11.411 +    show "p (x + y) \<le> p x + p y"
  11.412 +    proof -
  11.413 +      have "p (x + y) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel>" by (simp only: p_def)
  11.414 +      also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
  11.415 +      from x y have "\<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>" ..
  11.416 +      with a have " \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>F * (\<parallel>x\<parallel> + \<parallel>y\<parallel>)"
  11.417 +        by (simp add: mult_left_mono)
  11.418 +      also have "\<dots> = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel> + \<parallel>f\<parallel>\<hyphen>F * \<parallel>y\<parallel>" by (simp only: right_distrib)
  11.419 +      also have "\<dots> = p x + p y" by (simp only: p_def)
  11.420 +      finally show ?thesis .
  11.421 +    qed
  11.422 +  qed
  11.423 +
  11.424 +  txt {* @{text f} is bounded by @{text p}. *}
  11.425 +
  11.426 +  have "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
  11.427 +  proof
  11.428 +    fix x assume "x \<in> F"
  11.429 +    with `continuous F norm f` and linearform
  11.430 +    show "\<bar>f x\<bar> \<le> p x"
  11.431 +      unfolding p_def by (rule normed_vectorspace_with_fn_norm.fn_norm_le_cong
  11.432 +        [OF normed_vectorspace_with_fn_norm.intro,
  11.433 +         OF F_norm, folded B_def fn_norm_def])
  11.434 +  qed
  11.435 +
  11.436 +  txt {* Using the fact that @{text p} is a seminorm and @{text f} is bounded
  11.437 +    by @{text p} we can apply the Hahn-Banach Theorem for real vector
  11.438 +    spaces. So @{text f} can be extended in a norm-preserving way to
  11.439 +    some function @{text g} on the whole vector space @{text E}. *}
  11.440 +
  11.441 +  with E FE linearform q obtain g where
  11.442 +      linearformE: "linearform E g"
  11.443 +    and a: "\<forall>x \<in> F. g x = f x"
  11.444 +    and b: "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
  11.445 +    by (rule abs_HahnBanach [elim_format]) iprover
  11.446 +
  11.447 +  txt {* We furthermore have to show that @{text g} is also continuous: *}
  11.448 +
  11.449 +  have g_cont: "continuous E norm g" using linearformE
  11.450 +  proof
  11.451 +    fix x assume "x \<in> E"
  11.452 +    with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
  11.453 +      by (simp only: p_def)
  11.454 +  qed
  11.455 +
  11.456 +  txt {* To complete the proof, we show that @{text "\<parallel>g\<parallel> = \<parallel>f\<parallel>"}. *}
  11.457 +
  11.458 +  have "\<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
  11.459 +  proof (rule order_antisym)
  11.460 +    txt {*
  11.461 +      First we show @{text "\<parallel>g\<parallel> \<le> \<parallel>f\<parallel>"}.  The function norm @{text
  11.462 +      "\<parallel>g\<parallel>"} is defined as the smallest @{text "c \<in> \<real>"} such that
  11.463 +      \begin{center}
  11.464 +      \begin{tabular}{l}
  11.465 +      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
  11.466 +      \end{tabular}
  11.467 +      \end{center}
  11.468 +      \noindent Furthermore holds
  11.469 +      \begin{center}
  11.470 +      \begin{tabular}{l}
  11.471 +      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
  11.472 +      \end{tabular}
  11.473 +      \end{center}
  11.474 +    *}
  11.475 +
  11.476 +    have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
  11.477 +    proof
  11.478 +      fix x assume "x \<in> E"
  11.479 +      with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
  11.480 +        by (simp only: p_def)
  11.481 +    qed
  11.482 +    from g_cont this ge_zero
  11.483 +    show "\<parallel>g\<parallel>\<hyphen>E \<le> \<parallel>f\<parallel>\<hyphen>F"
  11.484 +      by (rule fn_norm_least [of g, folded B_def fn_norm_def])
  11.485 +
  11.486 +    txt {* The other direction is achieved by a similar argument. *}
  11.487 +
  11.488 +    show "\<parallel>f\<parallel>\<hyphen>F \<le> \<parallel>g\<parallel>\<hyphen>E"
  11.489 +    proof (rule normed_vectorspace_with_fn_norm.fn_norm_least
  11.490 +	[OF normed_vectorspace_with_fn_norm.intro,
  11.491 +	 OF F_norm, folded B_def fn_norm_def])
  11.492 +      show "\<forall>x \<in> F. \<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
  11.493 +      proof
  11.494 +	fix x assume x: "x \<in> F"
  11.495 +	from a x have "g x = f x" ..
  11.496 +	then have "\<bar>f x\<bar> = \<bar>g x\<bar>" by (simp only:)
  11.497 +	also from g_cont
  11.498 +	have "\<dots> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
  11.499 +	proof (rule fn_norm_le_cong [of g, folded B_def fn_norm_def])
  11.500 +	  from FE x show "x \<in> E" ..
  11.501 +	qed
  11.502 +	finally show "\<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>" .
  11.503 +      qed
  11.504 +      show "0 \<le> \<parallel>g\<parallel>\<hyphen>E"
  11.505 +	using g_cont
  11.506 +	by (rule fn_norm_ge_zero [of g, folded B_def fn_norm_def])
  11.507 +      show "continuous F norm f" by fact
  11.508 +    qed
  11.509 +  qed
  11.510 +  with linearformE a g_cont show ?thesis by blast
  11.511 +qed
  11.512 +
  11.513 +end
    12.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    12.2 +++ b/src/HOL/HahnBanach/HahnBanachExtLemmas.thy	Mon Dec 29 14:08:08 2008 +0100
    12.3 @@ -0,0 +1,281 @@
    12.4 +(*  Title:      HOL/Real/HahnBanach/HahnBanachExtLemmas.thy
    12.5 +    ID:         $Id$
    12.6 +    Author:     Gertrud Bauer, TU Munich
    12.7 +*)
    12.8 +
    12.9 +header {* Extending non-maximal functions *}
   12.10 +
   12.11 +theory HahnBanachExtLemmas
   12.12 +imports FunctionNorm
   12.13 +begin
   12.14 +
   12.15 +text {*
   12.16 +  In this section the following context is presumed.  Let @{text E} be
   12.17 +  a real vector space with a seminorm @{text q} on @{text E}. @{text
   12.18 +  F} is a subspace of @{text E} and @{text f} a linear function on
   12.19 +  @{text F}. We consider a subspace @{text H} of @{text E} that is a
   12.20 +  superspace of @{text F} and a linear form @{text h} on @{text
   12.21 +  H}. @{text H} is a not equal to @{text E} and @{text "x\<^sub>0"} is
   12.22 +  an element in @{text "E - H"}.  @{text H} is extended to the direct
   12.23 +  sum @{text "H' = H + lin x\<^sub>0"}, so for any @{text "x \<in> H'"}
   12.24 +  the decomposition of @{text "x = y + a \<cdot> x"} with @{text "y \<in> H"} is
   12.25 +  unique. @{text h'} is defined on @{text H'} by @{text "h' x = h y +
   12.26 +  a \<cdot> \<xi>"} for a certain @{text \<xi>}.
   12.27 +
   12.28 +  Subsequently we show some properties of this extension @{text h'} of
   12.29 +  @{text h}.
   12.30 +
   12.31 +  \medskip This lemma will be used to show the existence of a linear
   12.32 +  extension of @{text f} (see page \pageref{ex-xi-use}). It is a
   12.33 +  consequence of the completeness of @{text \<real>}. To show
   12.34 +  \begin{center}
   12.35 +  \begin{tabular}{l}
   12.36 +  @{text "\<exists>\<xi>. \<forall>y \<in> F. a y \<le> \<xi> \<and> \<xi> \<le> b y"}
   12.37 +  \end{tabular}
   12.38 +  \end{center}
   12.39 +  \noindent it suffices to show that
   12.40 +  \begin{center}
   12.41 +  \begin{tabular}{l}
   12.42 +  @{text "\<forall>u \<in> F. \<forall>v \<in> F. a u \<le> b v"}
   12.43 +  \end{tabular}
   12.44 +  \end{center}
   12.45 +*}
   12.46 +
   12.47 +lemma ex_xi:
   12.48 +  assumes "vectorspace F"
   12.49 +  assumes r: "\<And>u v. u \<in> F \<Longrightarrow> v \<in> F \<Longrightarrow> a u \<le> b v"
   12.50 +  shows "\<exists>xi::real. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y"
   12.51 +proof -
   12.52 +  interpret vectorspace [F] by fact
   12.53 +  txt {* From the completeness of the reals follows:
   12.54 +    The set @{text "S = {a u. u \<in> F}"} has a supremum, if it is
   12.55 +    non-empty and has an upper bound. *}
   12.56 +
   12.57 +  let ?S = "{a u | u. u \<in> F}"
   12.58 +  have "\<exists>xi. lub ?S xi"
   12.59 +  proof (rule real_complete)
   12.60 +    have "a 0 \<in> ?S" by blast
   12.61 +    then show "\<exists>X. X \<in> ?S" ..
   12.62 +    have "\<forall>y \<in> ?S. y \<le> b 0"
   12.63 +    proof
   12.64 +      fix y assume y: "y \<in> ?S"
   12.65 +      then obtain u where u: "u \<in> F" and y: "y = a u" by blast
   12.66 +      from u and zero have "a u \<le> b 0" by (rule r)
   12.67 +      with y show "y \<le> b 0" by (simp only:)
   12.68 +    qed
   12.69 +    then show "\<exists>u. \<forall>y \<in> ?S. y \<le> u" ..
   12.70 +  qed
   12.71 +  then obtain xi where xi: "lub ?S xi" ..
   12.72 +  {
   12.73 +    fix y assume "y \<in> F"
   12.74 +    then have "a y \<in> ?S" by blast
   12.75 +    with xi have "a y \<le> xi" by (rule lub.upper)
   12.76 +  } moreover {
   12.77 +    fix y assume y: "y \<in> F"
   12.78 +    from xi have "xi \<le> b y"
   12.79 +    proof (rule lub.least)
   12.80 +      fix au assume "au \<in> ?S"
   12.81 +      then obtain u where u: "u \<in> F" and au: "au = a u" by blast
   12.82 +      from u y have "a u \<le> b y" by (rule r)
   12.83 +      with au show "au \<le> b y" by (simp only:)
   12.84 +    qed
   12.85 +  } ultimately show "\<exists>xi. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y" by blast
   12.86 +qed
   12.87 +
   12.88 +text {*
   12.89 +  \medskip The function @{text h'} is defined as a @{text "h' x = h y
   12.90 +  + a \<cdot> \<xi>"} where @{text "x = y + a \<cdot> \<xi>"} is a linear extension of
   12.91 +  @{text h} to @{text H'}.
   12.92 +*}
   12.93 +
   12.94 +lemma h'_lf:
   12.95 +  assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
   12.96 +      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
   12.97 +    and H'_def: "H' \<equiv> H + lin x0"
   12.98 +    and HE: "H \<unlhd> E"
   12.99 +  assumes "linearform H h"
  12.100 +  assumes x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
  12.101 +  assumes E: "vectorspace E"
  12.102 +  shows "linearform H' h'"
  12.103 +proof -
  12.104 +  interpret linearform [H h] by fact
  12.105 +  interpret vectorspace [E] by fact
  12.106 +  show ?thesis
  12.107 +  proof
  12.108 +    note E = `vectorspace E`
  12.109 +    have H': "vectorspace H'"
  12.110 +    proof (unfold H'_def)
  12.111 +      from `x0 \<in> E`
  12.112 +      have "lin x0 \<unlhd> E" ..
  12.113 +      with HE show "vectorspace (H + lin x0)" using E ..
  12.114 +    qed
  12.115 +    {
  12.116 +      fix x1 x2 assume x1: "x1 \<in> H'" and x2: "x2 \<in> H'"
  12.117 +      show "h' (x1 + x2) = h' x1 + h' x2"
  12.118 +      proof -
  12.119 +	from H' x1 x2 have "x1 + x2 \<in> H'"
  12.120 +          by (rule vectorspace.add_closed)
  12.121 +	with x1 x2 obtain y y1 y2 a a1 a2 where
  12.122 +          x1x2: "x1 + x2 = y + a \<cdot> x0" and y: "y \<in> H"
  12.123 +          and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
  12.124 +          and x2_rep: "x2 = y2 + a2 \<cdot> x0" and y2: "y2 \<in> H"
  12.125 +          unfolding H'_def sum_def lin_def by blast
  12.126 +	
  12.127 +	have ya: "y1 + y2 = y \<and> a1 + a2 = a" using E HE _ y x0
  12.128 +	proof (rule decomp_H') txt_raw {* \label{decomp-H-use} *}
  12.129 +          from HE y1 y2 show "y1 + y2 \<in> H"
  12.130 +            by (rule subspace.add_closed)
  12.131 +          from x0 and HE y y1 y2
  12.132 +          have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E"  "y2 \<in> E" by auto
  12.133 +          with x1_rep x2_rep have "(y1 + y2) + (a1 + a2) \<cdot> x0 = x1 + x2"
  12.134 +            by (simp add: add_ac add_mult_distrib2)
  12.135 +          also note x1x2
  12.136 +          finally show "(y1 + y2) + (a1 + a2) \<cdot> x0 = y + a \<cdot> x0" .
  12.137 +	qed
  12.138 +	
  12.139 +	from h'_def x1x2 E HE y x0
  12.140 +	have "h' (x1 + x2) = h y + a * xi"
  12.141 +          by (rule h'_definite)
  12.142 +	also have "\<dots> = h (y1 + y2) + (a1 + a2) * xi"
  12.143 +          by (simp only: ya)
  12.144 +	also from y1 y2 have "h (y1 + y2) = h y1 + h y2"
  12.145 +          by simp
  12.146 +	also have "\<dots> + (a1 + a2) * xi = (h y1 + a1 * xi) + (h y2 + a2 * xi)"
  12.147 +          by (simp add: left_distrib)
  12.148 +	also from h'_def x1_rep E HE y1 x0
  12.149 +	have "h y1 + a1 * xi = h' x1"
  12.150 +          by (rule h'_definite [symmetric])
  12.151 +	also from h'_def x2_rep E HE y2 x0
  12.152 +	have "h y2 + a2 * xi = h' x2"
  12.153 +          by (rule h'_definite [symmetric])
  12.154 +	finally show ?thesis .
  12.155 +      qed
  12.156 +    next
  12.157 +      fix x1 c assume x1: "x1 \<in> H'"
  12.158 +      show "h' (c \<cdot> x1) = c * (h' x1)"
  12.159 +      proof -
  12.160 +	from H' x1 have ax1: "c \<cdot> x1 \<in> H'"
  12.161 +          by (rule vectorspace.mult_closed)
  12.162 +	with x1 obtain y a y1 a1 where
  12.163 +            cx1_rep: "c \<cdot> x1 = y + a \<cdot> x0" and y: "y \<in> H"
  12.164 +          and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
  12.165 +          unfolding H'_def sum_def lin_def by blast
  12.166 +	
  12.167 +	have ya: "c \<cdot> y1 = y \<and> c * a1 = a" using E HE _ y x0
  12.168 +	proof (rule decomp_H')
  12.169 +          from HE y1 show "c \<cdot> y1 \<in> H"
  12.170 +            by (rule subspace.mult_closed)
  12.171 +          from x0 and HE y y1
  12.172 +          have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E" by auto
  12.173 +          with x1_rep have "c \<cdot> y1 + (c * a1) \<cdot> x0 = c \<cdot> x1"
  12.174 +            by (simp add: mult_assoc add_mult_distrib1)
  12.175 +          also note cx1_rep
  12.176 +          finally show "c \<cdot> y1 + (c * a1) \<cdot> x0 = y + a \<cdot> x0" .
  12.177 +	qed
  12.178 +	
  12.179 +	from h'_def cx1_rep E HE y x0 have "h' (c \<cdot> x1) = h y + a * xi"
  12.180 +          by (rule h'_definite)
  12.181 +	also have "\<dots> = h (c \<cdot> y1) + (c * a1) * xi"
  12.182 +          by (simp only: ya)
  12.183 +	also from y1 have "h (c \<cdot> y1) = c * h y1"
  12.184 +          by simp
  12.185 +	also have "\<dots> + (c * a1) * xi = c * (h y1 + a1 * xi)"
  12.186 +          by (simp only: right_distrib)
  12.187 +	also from h'_def x1_rep E HE y1 x0 have "h y1 + a1 * xi = h' x1"
  12.188 +          by (rule h'_definite [symmetric])
  12.189 +	finally show ?thesis .
  12.190 +      qed
  12.191 +    }
  12.192 +  qed
  12.193 +qed
  12.194 +
  12.195 +text {* \medskip The linear extension @{text h'} of @{text h}
  12.196 +  is bounded by the seminorm @{text p}. *}
  12.197 +
  12.198 +lemma h'_norm_pres:
  12.199 +  assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
  12.200 +      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
  12.201 +    and H'_def: "H' \<equiv> H + lin x0"
  12.202 +    and x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
  12.203 +  assumes E: "vectorspace E" and HE: "subspace H E"
  12.204 +    and "seminorm E p" and "linearform H h"
  12.205 +  assumes a: "\<forall>y \<in> H. h y \<le> p y"
  12.206 +    and a': "\<forall>y \<in> H. - p (y + x0) - h y \<le> xi \<and> xi \<le> p (y + x0) - h y"
  12.207 +  shows "\<forall>x \<in> H'. h' x \<le> p x"
  12.208 +proof -
  12.209 +  interpret vectorspace [E] by fact
  12.210 +  interpret subspace [H E] by fact
  12.211 +  interpret seminorm [E p] by fact
  12.212 +  interpret linearform [H h] by fact
  12.213 +  show ?thesis
  12.214 +  proof
  12.215 +    fix x assume x': "x \<in> H'"
  12.216 +    show "h' x \<le> p x"
  12.217 +    proof -
  12.218 +      from a' have a1: "\<forall>ya \<in> H. - p (ya + x0) - h ya \<le> xi"
  12.219 +	and a2: "\<forall>ya \<in> H. xi \<le> p (ya + x0) - h ya" by auto
  12.220 +      from x' obtain y a where
  12.221 +          x_rep: "x = y + a \<cdot> x0" and y: "y \<in> H"
  12.222 +	unfolding H'_def sum_def lin_def by blast
  12.223 +      from y have y': "y \<in> E" ..
  12.224 +      from y have ay: "inverse a \<cdot> y \<in> H" by simp
  12.225 +      
  12.226 +      from h'_def x_rep E HE y x0 have "h' x = h y + a * xi"
  12.227 +	by (rule h'_definite)
  12.228 +      also have "\<dots> \<le> p (y + a \<cdot> x0)"
  12.229 +      proof (rule linorder_cases)
  12.230 +	assume z: "a = 0"
  12.231 +	then have "h y + a * xi = h y" by simp
  12.232 +	also from a y have "\<dots> \<le> p y" ..
  12.233 +	also from x0 y' z have "p y = p (y + a \<cdot> x0)" by simp
  12.234 +	finally show ?thesis .
  12.235 +      next
  12.236 +	txt {* In the case @{text "a < 0"}, we use @{text "a\<^sub>1"}
  12.237 +          with @{text ya} taken as @{text "y / a"}: *}
  12.238 +	assume lz: "a < 0" then have nz: "a \<noteq> 0" by simp
  12.239 +	from a1 ay
  12.240 +	have "- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y) \<le> xi" ..
  12.241 +	with lz have "a * xi \<le>
  12.242 +          a * (- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
  12.243 +          by (simp add: mult_left_mono_neg order_less_imp_le)
  12.244 +	
  12.245 +	also have "\<dots> =
  12.246 +          - a * (p (inverse a \<cdot> y + x0)) - a * (h (inverse a \<cdot> y))"
  12.247 +	  by (simp add: right_diff_distrib)
  12.248 +	also from lz x0 y' have "- a * (p (inverse a \<cdot> y + x0)) =
  12.249 +          p (a \<cdot> (inverse a \<cdot> y + x0))"
  12.250 +          by (simp add: abs_homogenous)
  12.251 +	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
  12.252 +          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
  12.253 +	also from nz y have "a * (h (inverse a \<cdot> y)) =  h y"
  12.254 +          by simp
  12.255 +	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
  12.256 +	then show ?thesis by simp
  12.257 +      next
  12.258 +	txt {* In the case @{text "a > 0"}, we use @{text "a\<^sub>2"}
  12.259 +          with @{text ya} taken as @{text "y / a"}: *}
  12.260 +	assume gz: "0 < a" then have nz: "a \<noteq> 0" by simp
  12.261 +	from a2 ay
  12.262 +	have "xi \<le> p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y)" ..
  12.263 +	with gz have "a * xi \<le>
  12.264 +          a * (p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
  12.265 +          by simp
  12.266 +	also have "\<dots> = a * p (inverse a \<cdot> y + x0) - a * h (inverse a \<cdot> y)"
  12.267 +	  by (simp add: right_diff_distrib)
  12.268 +	also from gz x0 y'
  12.269 +	have "a * p (inverse a \<cdot> y + x0) = p (a \<cdot> (inverse a \<cdot> y + x0))"
  12.270 +          by (simp add: abs_homogenous)
  12.271 +	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
  12.272 +          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
  12.273 +	also from nz y have "a * h (inverse a \<cdot> y) = h y"
  12.274 +          by simp
  12.275 +	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
  12.276 +	then show ?thesis by simp
  12.277 +      qed
  12.278 +      also from x_rep have "\<dots> = p x" by (simp only:)
  12.279 +      finally show ?thesis .
  12.280 +    qed
  12.281 +  qed
  12.282 +qed
  12.283 +
  12.284 +end
    13.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    13.2 +++ b/src/HOL/HahnBanach/HahnBanachLemmas.thy	Mon Dec 29 14:08:08 2008 +0100
    13.3 @@ -0,0 +1,4 @@
    13.4 +(*<*)
    13.5 +theory HahnBanachLemmas imports HahnBanachSupLemmas HahnBanachExtLemmas begin
    13.6 +end
    13.7 +(*>*)
    13.8 \ No newline at end of file
    14.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    14.2 +++ b/src/HOL/HahnBanach/HahnBanachSupLemmas.thy	Mon Dec 29 14:08:08 2008 +0100
    14.3 @@ -0,0 +1,446 @@
    14.4 +(*  Title:      HOL/Real/HahnBanach/HahnBanachSupLemmas.thy
    14.5 +    ID:         $Id$
    14.6 +    Author:     Gertrud Bauer, TU Munich
    14.7 +*)
    14.8 +
    14.9 +header {* The supremum w.r.t.~the function order *}
   14.10 +
   14.11 +theory HahnBanachSupLemmas
   14.12 +imports FunctionNorm ZornLemma
   14.13 +begin
   14.14 +
   14.15 +text {*
   14.16 +  This section contains some lemmas that will be used in the proof of
   14.17 +  the Hahn-Banach Theorem.  In this section the following context is
   14.18 +  presumed.  Let @{text E} be a real vector space with a seminorm
   14.19 +  @{text p} on @{text E}.  @{text F} is a subspace of @{text E} and
   14.20 +  @{text f} a linear form on @{text F}. We consider a chain @{text c}
   14.21 +  of norm-preserving extensions of @{text f}, such that @{text "\<Union>c =
   14.22 +  graph H h"}.  We will show some properties about the limit function
   14.23 +  @{text h}, i.e.\ the supremum of the chain @{text c}.
   14.24 +
   14.25 +  \medskip Let @{text c} be a chain of norm-preserving extensions of
   14.26 +  the function @{text f} and let @{text "graph H h"} be the supremum
   14.27 +  of @{text c}.  Every element in @{text H} is member of one of the
   14.28 +  elements of the chain.
   14.29 +*}
   14.30 +lemmas [dest?] = chainD
   14.31 +lemmas chainE2 [elim?] = chainD2 [elim_format, standard]
   14.32 +
   14.33 +lemma some_H'h't:
   14.34 +  assumes M: "M = norm_pres_extensions E p F f"
   14.35 +    and cM: "c \<in> chain M"
   14.36 +    and u: "graph H h = \<Union>c"
   14.37 +    and x: "x \<in> H"
   14.38 +  shows "\<exists>H' h'. graph H' h' \<in> c
   14.39 +    \<and> (x, h x) \<in> graph H' h'
   14.40 +    \<and> linearform H' h' \<and> H' \<unlhd> E
   14.41 +    \<and> F \<unlhd> H' \<and> graph F f \<subseteq> graph H' h'
   14.42 +    \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
   14.43 +proof -
   14.44 +  from x have "(x, h x) \<in> graph H h" ..
   14.45 +  also from u have "\<dots> = \<Union>c" .
   14.46 +  finally obtain g where gc: "g \<in> c" and gh: "(x, h x) \<in> g" by blast
   14.47 +
   14.48 +  from cM have "c \<subseteq> M" ..
   14.49 +  with gc have "g \<in> M" ..
   14.50 +  also from M have "\<dots> = norm_pres_extensions E p F f" .
   14.51 +  finally obtain H' and h' where g: "g = graph H' h'"
   14.52 +    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
   14.53 +      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x" ..
   14.54 +
   14.55 +  from gc and g have "graph H' h' \<in> c" by (simp only:)
   14.56 +  moreover from gh and g have "(x, h x) \<in> graph H' h'" by (simp only:)
   14.57 +  ultimately show ?thesis using * by blast
   14.58 +qed
   14.59 +
   14.60 +text {*
   14.61 +  \medskip Let @{text c} be a chain of norm-preserving extensions of
   14.62 +  the function @{text f} and let @{text "graph H h"} be the supremum
   14.63 +  of @{text c}.  Every element in the domain @{text H} of the supremum
   14.64 +  function is member of the domain @{text H'} of some function @{text
   14.65 +  h'}, such that @{text h} extends @{text h'}.
   14.66 +*}
   14.67 +
   14.68 +lemma some_H'h':
   14.69 +  assumes M: "M = norm_pres_extensions E p F f"
   14.70 +    and cM: "c \<in> chain M"
   14.71 +    and u: "graph H h = \<Union>c"
   14.72 +    and x: "x \<in> H"
   14.73 +  shows "\<exists>H' h'. x \<in> H' \<and> graph H' h' \<subseteq> graph H h
   14.74 +    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
   14.75 +    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
   14.76 +proof -
   14.77 +  from M cM u x obtain H' h' where
   14.78 +      x_hx: "(x, h x) \<in> graph H' h'"
   14.79 +    and c: "graph H' h' \<in> c"
   14.80 +    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
   14.81 +      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
   14.82 +    by (rule some_H'h't [elim_format]) blast
   14.83 +  from x_hx have "x \<in> H'" ..
   14.84 +  moreover from cM u c have "graph H' h' \<subseteq> graph H h"
   14.85 +    by (simp only: chain_ball_Union_upper)
   14.86 +  ultimately show ?thesis using * by blast
   14.87 +qed
   14.88 +
   14.89 +text {*
   14.90 +  \medskip Any two elements @{text x} and @{text y} in the domain
   14.91 +  @{text H} of the supremum function @{text h} are both in the domain
   14.92 +  @{text H'} of some function @{text h'}, such that @{text h} extends
   14.93 +  @{text h'}.
   14.94 +*}
   14.95 +
   14.96 +lemma some_H'h'2:
   14.97 +  assumes M: "M = norm_pres_extensions E p F f"
   14.98 +    and cM: "c \<in> chain M"
   14.99 +    and u: "graph H h = \<Union>c"
  14.100 +    and x: "x \<in> H"
  14.101 +    and y: "y \<in> H"
  14.102 +  shows "\<exists>H' h'. x \<in> H' \<and> y \<in> H'
  14.103 +    \<and> graph H' h' \<subseteq> graph H h
  14.104 +    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
  14.105 +    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
  14.106 +proof -
  14.107 +  txt {* @{text y} is in the domain @{text H''} of some function @{text h''},
  14.108 +  such that @{text h} extends @{text h''}. *}
  14.109 +
  14.110 +  from M cM u and y obtain H' h' where
  14.111 +      y_hy: "(y, h y) \<in> graph H' h'"
  14.112 +    and c': "graph H' h' \<in> c"
  14.113 +    and * :
  14.114 +      "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
  14.115 +      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
  14.116 +    by (rule some_H'h't [elim_format]) blast
  14.117 +
  14.118 +  txt {* @{text x} is in the domain @{text H'} of some function @{text h'},
  14.119 +    such that @{text h} extends @{text h'}. *}
  14.120 +
  14.121 +  from M cM u and x obtain H'' h'' where
  14.122 +      x_hx: "(x, h x) \<in> graph H'' h''"
  14.123 +    and c'': "graph H'' h'' \<in> c"
  14.124 +    and ** :
  14.125 +      "linearform H'' h''"  "H'' \<unlhd> E"  "F \<unlhd> H''"
  14.126 +      "graph F f \<subseteq> graph H'' h''"  "\<forall>x \<in> H''. h'' x \<le> p x"
  14.127 +    by (rule some_H'h't [elim_format]) blast
  14.128 +
  14.129 +  txt {* Since both @{text h'} and @{text h''} are elements of the chain,
  14.130 +    @{text h''} is an extension of @{text h'} or vice versa. Thus both
  14.131 +    @{text x} and @{text y} are contained in the greater
  14.132 +    one. \label{cases1}*}
  14.133 +
  14.134 +  from cM c'' c' have "graph H'' h'' \<subseteq> graph H' h' \<or> graph H' h' \<subseteq> graph H'' h''"
  14.135 +    (is "?case1 \<or> ?case2") ..
  14.136 +  then show ?thesis
  14.137 +  proof
  14.138 +    assume ?case1
  14.139 +    have "(x, h x) \<in> graph H'' h''" by fact
  14.140 +    also have "\<dots> \<subseteq> graph H' h'" by fact
  14.141 +    finally have xh:"(x, h x) \<in> graph H' h'" .
  14.142 +    then have "x \<in> H'" ..
  14.143 +    moreover from y_hy have "y \<in> H'" ..
  14.144 +    moreover from cM u and c' have "graph H' h' \<subseteq> graph H h"
  14.145 +      by (simp only: chain_ball_Union_upper)
  14.146 +    ultimately show ?thesis using * by blast
  14.147 +  next
  14.148 +    assume ?case2
  14.149 +    from x_hx have "x \<in> H''" ..
  14.150 +    moreover {
  14.151 +      have "(y, h y) \<in> graph H' h'" by (rule y_hy)
  14.152 +      also have "\<dots> \<subseteq> graph H'' h''" by fact
  14.153 +      finally have "(y, h y) \<in> graph H'' h''" .
  14.154 +    } then have "y \<in> H''" ..
  14.155 +    moreover from cM u and c'' have "graph H'' h'' \<subseteq> graph H h"
  14.156 +      by (simp only: chain_ball_Union_upper)
  14.157 +    ultimately show ?thesis using ** by blast
  14.158 +  qed
  14.159 +qed
  14.160 +
  14.161 +text {*
  14.162 +  \medskip The relation induced by the graph of the supremum of a
  14.163 +  chain @{text c} is definite, i.~e.~t is the graph of a function.
  14.164 +*}
  14.165 +
  14.166 +lemma sup_definite:
  14.167 +  assumes M_def: "M \<equiv> norm_pres_extensions E p F f"
  14.168 +    and cM: "c \<in> chain M"
  14.169 +    and xy: "(x, y) \<in> \<Union>c"
  14.170 +    and xz: "(x, z) \<in> \<Union>c"
  14.171 +  shows "z = y"
  14.172 +proof -
  14.173 +  from cM have c: "c \<subseteq> M" ..
  14.174 +  from xy obtain G1 where xy': "(x, y) \<in> G1" and G1: "G1 \<in> c" ..
  14.175 +  from xz obtain G2 where xz': "(x, z) \<in> G2" and G2: "G2 \<in> c" ..
  14.176 +
  14.177 +  from G1 c have "G1 \<in> M" ..
  14.178 +  then obtain H1 h1 where G1_rep: "G1 = graph H1 h1"
  14.179 +    unfolding M_def by blast
  14.180 +
  14.181 +  from G2 c have "G2 \<in> M" ..
  14.182 +  then obtain H2 h2 where G2_rep: "G2 = graph H2 h2"
  14.183 +    unfolding M_def by blast
  14.184 +
  14.185 +  txt {* @{text "G\<^sub>1"} is contained in @{text "G\<^sub>2"}
  14.186 +    or vice versa, since both @{text "G\<^sub>1"} and @{text
  14.187 +    "G\<^sub>2"} are members of @{text c}. \label{cases2}*}
  14.188 +
  14.189 +  from cM G1 G2 have "G1 \<subseteq> G2 \<or> G2 \<subseteq> G1" (is "?case1 \<or> ?case2") ..
  14.190 +  then show ?thesis
  14.191 +  proof
  14.192 +    assume ?case1
  14.193 +    with xy' G2_rep have "(x, y) \<in> graph H2 h2" by blast
  14.194 +    then have "y = h2 x" ..
  14.195 +    also
  14.196 +    from xz' G2_rep have "(x, z) \<in> graph H2 h2" by (simp only:)
  14.197 +    then have "z = h2 x" ..
  14.198 +    finally show ?thesis .
  14.199 +  next
  14.200 +    assume ?case2
  14.201 +    with xz' G1_rep have "(x, z) \<in> graph H1 h1" by blast
  14.202 +    then have "z = h1 x" ..
  14.203 +    also
  14.204 +    from xy' G1_rep have "(x, y) \<in> graph H1 h1" by (simp only:)
  14.205 +    then have "y = h1 x" ..
  14.206 +    finally show ?thesis ..
  14.207 +  qed
  14.208 +qed
  14.209 +
  14.210 +text {*
  14.211 +  \medskip The limit function @{text h} is linear. Every element
  14.212 +  @{text x} in the domain of @{text h} is in the domain of a function
  14.213 +  @{text h'} in the chain of norm preserving extensions.  Furthermore,
  14.214 +  @{text h} is an extension of @{text h'} so the function values of
  14.215 +  @{text x} are identical for @{text h'} and @{text h}.  Finally, the
  14.216 +  function @{text h'} is linear by construction of @{text M}.
  14.217 +*}
  14.218 +
  14.219 +lemma sup_lf:
  14.220 +  assumes M: "M = norm_pres_extensions E p F f"
  14.221 +    and cM: "c \<in> chain M"
  14.222 +    and u: "graph H h = \<Union>c"
  14.223 +  shows "linearform H h"
  14.224 +proof
  14.225 +  fix x y assume x: "x \<in> H" and y: "y \<in> H"
  14.226 +  with M cM u obtain H' h' where
  14.227 +        x': "x \<in> H'" and y': "y \<in> H'"
  14.228 +      and b: "graph H' h' \<subseteq> graph H h"
  14.229 +      and linearform: "linearform H' h'"
  14.230 +      and subspace: "H' \<unlhd> E"
  14.231 +    by (rule some_H'h'2 [elim_format]) blast
  14.232 +
  14.233 +  show "h (x + y) = h x + h y"
  14.234 +  proof -
  14.235 +    from linearform x' y' have "h' (x + y) = h' x + h' y"
  14.236 +      by (rule linearform.add)
  14.237 +    also from b x' have "h' x = h x" ..
  14.238 +    also from b y' have "h' y = h y" ..
  14.239 +    also from subspace x' y' have "x + y \<in> H'"
  14.240 +      by (rule subspace.add_closed)
  14.241 +    with b have "h' (x + y) = h (x + y)" ..
  14.242 +    finally show ?thesis .
  14.243 +  qed
  14.244 +next
  14.245 +  fix x a assume x: "x \<in> H"
  14.246 +  with M cM u obtain H' h' where
  14.247 +        x': "x \<in> H'"
  14.248 +      and b: "graph H' h' \<subseteq> graph H h"
  14.249 +      and linearform: "linearform H' h'"
  14.250 +      and subspace: "H' \<unlhd> E"
  14.251 +    by (rule some_H'h' [elim_format]) blast
  14.252 +
  14.253 +  show "h (a \<cdot> x) = a * h x"
  14.254 +  proof -
  14.255 +    from linearform x' have "h' (a \<cdot> x) = a * h' x"
  14.256 +      by (rule linearform.mult)
  14.257 +    also from b x' have "h' x = h x" ..
  14.258 +    also from subspace x' have "a \<cdot> x \<in> H'"
  14.259 +      by (rule subspace.mult_closed)
  14.260 +    with b have "h' (a \<cdot> x) = h (a \<cdot> x)" ..
  14.261 +    finally show ?thesis .
  14.262 +  qed
  14.263 +qed
  14.264 +
  14.265 +text {*
  14.266 +  \medskip The limit of a non-empty chain of norm preserving
  14.267 +  extensions of @{text f} is an extension of @{text f}, since every
  14.268 +  element of the chain is an extension of @{text f} and the supremum
  14.269 +  is an extension for every element of the chain.
  14.270 +*}
  14.271 +
  14.272 +lemma sup_ext:
  14.273 +  assumes graph: "graph H h = \<Union>c"
  14.274 +    and M: "M = norm_pres_extensions E p F f"
  14.275 +    and cM: "c \<in> chain M"
  14.276 +    and ex: "\<exists>x. x \<in> c"
  14.277 +  shows "graph F f \<subseteq> graph H h"
  14.278 +proof -
  14.279 +  from ex obtain x where xc: "x \<in> c" ..
  14.280 +  from cM have "c \<subseteq> M" ..
  14.281 +  with xc have "x \<in> M" ..
  14.282 +  with M have "x \<in> norm_pres_extensions E p F f"
  14.283 +    by (simp only:)
  14.284 +  then obtain G g where "x = graph G g" and "graph F f \<subseteq> graph G g" ..
  14.285 +  then have "graph F f \<subseteq> x" by (simp only:)
  14.286 +  also from xc have "\<dots> \<subseteq> \<Union>c" by blast
  14.287 +  also from graph have "\<dots> = graph H h" ..
  14.288 +  finally show ?thesis .
  14.289 +qed
  14.290 +
  14.291 +text {*
  14.292 +  \medskip The domain @{text H} of the limit function is a superspace
  14.293 +  of @{text F}, since @{text F} is a subset of @{text H}. The
  14.294 +  existence of the @{text 0} element in @{text F} and the closure
  14.295 +  properties follow from the fact that @{text F} is a vector space.
  14.296 +*}
  14.297 +
  14.298 +lemma sup_supF:
  14.299 +  assumes graph: "graph H h = \<Union>c"
  14.300 +    and M: "M = norm_pres_extensions E p F f"
  14.301 +    and cM: "c \<in> chain M"
  14.302 +    and ex: "\<exists>x. x \<in> c"
  14.303 +    and FE: "F \<unlhd> E"
  14.304 +  shows "F \<unlhd> H"
  14.305 +proof
  14.306 +  from FE show "F \<noteq> {}" by (rule subspace.non_empty)
  14.307 +  from graph M cM ex have "graph F f \<subseteq> graph H h" by (rule sup_ext)
  14.308 +  then show "F \<subseteq> H" ..
  14.309 +  fix x y assume "x \<in> F" and "y \<in> F"
  14.310 +  with FE show "x + y \<in> F" by (rule subspace.add_closed)
  14.311 +next
  14.312 +  fix x a assume "x \<in> F"
  14.313 +  with FE show "a \<cdot> x \<in> F" by (rule subspace.mult_closed)
  14.314 +qed
  14.315 +
  14.316 +text {*
  14.317 +  \medskip The domain @{text H} of the limit function is a subspace of
  14.318 +  @{text E}.
  14.319 +*}
  14.320 +
  14.321 +lemma sup_subE:
  14.322 +  assumes graph: "graph H h = \<Union>c"
  14.323 +    and M: "M = norm_pres_extensions E p F f"
  14.324 +    and cM: "c \<in> chain M"
  14.325 +    and ex: "\<exists>x. x \<in> c"
  14.326 +    and FE: "F \<unlhd> E"
  14.327 +    and E: "vectorspace E"
  14.328 +  shows "H \<unlhd> E"
  14.329 +proof
  14.330 +  show "H \<noteq> {}"
  14.331 +  proof -
  14.332 +    from FE E have "0 \<in> F" by (rule subspace.zero)
  14.333 +    also from graph M cM ex FE have "F \<unlhd> H" by (rule sup_supF)
  14.334 +    then have "F \<subseteq> H" ..
  14.335 +    finally show ?thesis by blast
  14.336 +  qed
  14.337 +  show "H \<subseteq> E"
  14.338 +  proof
  14.339 +    fix x assume "x \<in> H"
  14.340 +    with M cM graph
  14.341 +    obtain H' h' where x: "x \<in> H'" and H'E: "H' \<unlhd> E"
  14.342 +      by (rule some_H'h' [elim_format]) blast
  14.343 +    from H'E have "H' \<subseteq> E" ..
  14.344 +    with x show "x \<in> E" ..
  14.345 +  qed
  14.346 +  fix x y assume x: "x \<in> H" and y: "y \<in> H"
  14.347 +  show "x + y \<in> H"
  14.348 +  proof -
  14.349 +    from M cM graph x y obtain H' h' where
  14.350 +          x': "x \<in> H'" and y': "y \<in> H'" and H'E: "H' \<unlhd> E"
  14.351 +        and graphs: "graph H' h' \<subseteq> graph H h"
  14.352 +      by (rule some_H'h'2 [elim_format]) blast
  14.353 +    from H'E x' y' have "x + y \<in> H'"
  14.354 +      by (rule subspace.add_closed)
  14.355 +    also from graphs have "H' \<subseteq> H" ..
  14.356 +    finally show ?thesis .
  14.357 +  qed
  14.358 +next
  14.359 +  fix x a assume x: "x \<in> H"
  14.360 +  show "a \<cdot> x \<in> H"
  14.361 +  proof -
  14.362 +    from M cM graph x
  14.363 +    obtain H' h' where x': "x \<in> H'" and H'E: "H' \<unlhd> E"
  14.364 +        and graphs: "graph H' h' \<subseteq> graph H h"
  14.365 +      by (rule some_H'h' [elim_format]) blast
  14.366 +    from H'E x' have "a \<cdot> x \<in> H'" by (rule subspace.mult_closed)
  14.367 +    also from graphs have "H' \<subseteq> H" ..
  14.368 +    finally show ?thesis .
  14.369 +  qed
  14.370 +qed
  14.371 +
  14.372 +text {*
  14.373 +  \medskip The limit function is bounded by the norm @{text p} as
  14.374 +  well, since all elements in the chain are bounded by @{text p}.
  14.375 +*}
  14.376 +
  14.377 +lemma sup_norm_pres:
  14.378 +  assumes graph: "graph H h = \<Union>c"
  14.379 +    and M: "M = norm_pres_extensions E p F f"
  14.380 +    and cM: "c \<in> chain M"
  14.381 +  shows "\<forall>x \<in> H. h x \<le> p x"
  14.382 +proof
  14.383 +  fix x assume "x \<in> H"
  14.384 +  with M cM graph obtain H' h' where x': "x \<in> H'"
  14.385 +      and graphs: "graph H' h' \<subseteq> graph H h"
  14.386 +      and a: "\<forall>x \<in> H'. h' x \<le> p x"
  14.387 +    by (rule some_H'h' [elim_format]) blast
  14.388 +  from graphs x' have [symmetric]: "h' x = h x" ..
  14.389 +  also from a x' have "h' x \<le> p x " ..
  14.390 +  finally show "h x \<le> p x" .
  14.391 +qed
  14.392 +
  14.393 +text {*
  14.394 +  \medskip The following lemma is a property of linear forms on real
  14.395 +  vector spaces. It will be used for the lemma @{text abs_HahnBanach}
  14.396 +  (see page \pageref{abs-HahnBanach}). \label{abs-ineq-iff} For real
  14.397 +  vector spaces the following inequations are equivalent:
  14.398 +  \begin{center}
  14.399 +  \begin{tabular}{lll}
  14.400 +  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
  14.401 +  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
  14.402 +  \end{tabular}
  14.403 +  \end{center}
  14.404 +*}
  14.405 +
  14.406 +lemma abs_ineq_iff:
  14.407 +  assumes "subspace H E" and "vectorspace E" and "seminorm E p"
  14.408 +    and "linearform H h"
  14.409 +  shows "(\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x) = (\<forall>x \<in> H. h x \<le> p x)" (is "?L = ?R")
  14.410 +proof
  14.411 +  interpret subspace [H E] by fact
  14.412 +  interpret vectorspace [E] by fact
  14.413 +  interpret seminorm [E p] by fact
  14.414 +  interpret linearform [H h] by fact
  14.415 +  have H: "vectorspace H" using `vectorspace E` ..
  14.416 +  {
  14.417 +    assume l: ?L
  14.418 +    show ?R
  14.419 +    proof
  14.420 +      fix x assume x: "x \<in> H"
  14.421 +      have "h x \<le> \<bar>h x\<bar>" by arith
  14.422 +      also from l x have "\<dots> \<le> p x" ..
  14.423 +      finally show "h x \<le> p x" .
  14.424 +    qed
  14.425 +  next
  14.426 +    assume r: ?R
  14.427 +    show ?L
  14.428 +    proof
  14.429 +      fix x assume x: "x \<in> H"
  14.430 +      show "\<And>a b :: real. - a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> \<bar>b\<bar> \<le> a"
  14.431 +        by arith
  14.432 +      from `linearform H h` and H x
  14.433 +      have "- h x = h (- x)" by (rule linearform.neg [symmetric])
  14.434 +      also
  14.435 +      from H x have "- x \<in> H" by (rule vectorspace.neg_closed)
  14.436 +      with r have "h (- x) \<le> p (- x)" ..
  14.437 +      also have "\<dots> = p x"
  14.438 +	using `seminorm E p` `vectorspace E`
  14.439 +      proof (rule seminorm.minus)
  14.440 +        from x show "x \<in> E" ..
  14.441 +      qed
  14.442 +      finally have "- h x \<le> p x" .
  14.443 +      then show "- p x \<le> h x" by simp
  14.444 +      from r x show "h x \<le> p x" ..
  14.445 +    qed
  14.446 +  }
  14.447 +qed
  14.448 +
  14.449 +end
    15.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    15.2 +++ b/src/HOL/HahnBanach/Linearform.thy	Mon Dec 29 14:08:08 2008 +0100
    15.3 @@ -0,0 +1,61 @@
    15.4 +(*  Title:      HOL/Real/HahnBanach/Linearform.thy
    15.5 +    ID:         $Id$
    15.6 +    Author:     Gertrud Bauer, TU Munich
    15.7 +*)
    15.8 +
    15.9 +header {* Linearforms *}
   15.10 +
   15.11 +theory Linearform
   15.12 +imports VectorSpace
   15.13 +begin
   15.14 +
   15.15 +text {*
   15.16 +  A \emph{linear form} is a function on a vector space into the reals
   15.17 +  that is additive and multiplicative.
   15.18 +*}
   15.19 +
   15.20 +locale linearform = var V + var f +
   15.21 +  constrains V :: "'a\<Colon>{minus, plus, zero, uminus} set"
   15.22 +  assumes add [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x + y) = f x + f y"
   15.23 +    and mult [iff]: "x \<in> V \<Longrightarrow> f (a \<cdot> x) = a * f x"
   15.24 +
   15.25 +declare linearform.intro [intro?]
   15.26 +
   15.27 +lemma (in linearform) neg [iff]:
   15.28 +  assumes "vectorspace V"
   15.29 +  shows "x \<in> V \<Longrightarrow> f (- x) = - f x"
   15.30 +proof -
   15.31 +  interpret vectorspace [V] by fact
   15.32 +  assume x: "x \<in> V"
   15.33 +  then have "f (- x) = f ((- 1) \<cdot> x)" by (simp add: negate_eq1)
   15.34 +  also from x have "\<dots> = (- 1) * (f x)" by (rule mult)
   15.35 +  also from x have "\<dots> = - (f x)" by simp
   15.36 +  finally show ?thesis .
   15.37 +qed
   15.38 +
   15.39 +lemma (in linearform) diff [iff]:
   15.40 +  assumes "vectorspace V"
   15.41 +  shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x - y) = f x - f y"
   15.42 +proof -
   15.43 +  interpret vectorspace [V] by fact
   15.44 +  assume x: "x \<in> V" and y: "y \<in> V"
   15.45 +  then have "x - y = x + - y" by (rule diff_eq1)
   15.46 +  also have "f \<dots> = f x + f (- y)" by (rule add) (simp_all add: x y)
   15.47 +  also have "f (- y) = - f y" using `vectorspace V` y by (rule neg)
   15.48 +  finally show ?thesis by simp
   15.49 +qed
   15.50 +
   15.51 +text {* Every linear form yields @{text 0} for the @{text 0} vector. *}
   15.52 +
   15.53 +lemma (in linearform) zero [iff]:
   15.54 +  assumes "vectorspace V"
   15.55 +  shows "f 0 = 0"
   15.56 +proof -
   15.57 +  interpret vectorspace [V] by fact
   15.58 +  have "f 0 = f (0 - 0)" by simp
   15.59 +  also have "\<dots> = f 0 - f 0" using `vectorspace V` by (rule diff) simp_all
   15.60 +  also have "\<dots> = 0" by simp
   15.61 +  finally show ?thesis .
   15.62 +qed
   15.63 +
   15.64 +end
    16.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    16.2 +++ b/src/HOL/HahnBanach/NormedSpace.thy	Mon Dec 29 14:08:08 2008 +0100
    16.3 @@ -0,0 +1,118 @@
    16.4 +(*  Title:      HOL/Real/HahnBanach/NormedSpace.thy
    16.5 +    ID:         $Id$
    16.6 +    Author:     Gertrud Bauer, TU Munich
    16.7 +*)
    16.8 +
    16.9 +header {* Normed vector spaces *}
   16.10 +
   16.11 +theory NormedSpace
   16.12 +imports Subspace
   16.13 +begin
   16.14 +
   16.15 +subsection {* Quasinorms *}
   16.16 +
   16.17 +text {*
   16.18 +  A \emph{seminorm} @{text "\<parallel>\<cdot>\<parallel>"} is a function on a real vector space
   16.19 +  into the reals that has the following properties: it is positive
   16.20 +  definite, absolute homogenous and subadditive.
   16.21 +*}
   16.22 +
   16.23 +locale norm_syntax =
   16.24 +  fixes norm :: "'a \<Rightarrow> real"    ("\<parallel>_\<parallel>")
   16.25 +
   16.26 +locale seminorm = var V + norm_syntax +
   16.27 +  constrains V :: "'a\<Colon>{minus, plus, zero, uminus} set"
   16.28 +  assumes ge_zero [iff?]: "x \<in> V \<Longrightarrow> 0 \<le> \<parallel>x\<parallel>"
   16.29 +    and abs_homogenous [iff?]: "x \<in> V \<Longrightarrow> \<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>"
   16.30 +    and subadditive [iff?]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> \<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>"
   16.31 +
   16.32 +declare seminorm.intro [intro?]
   16.33 +
   16.34 +lemma (in seminorm) diff_subadditive:
   16.35 +  assumes "vectorspace V"
   16.36 +  shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> \<parallel>x - y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>"
   16.37 +proof -
   16.38 +  interpret vectorspace [V] by fact
   16.39 +  assume x: "x \<in> V" and y: "y \<in> V"
   16.40 +  then have "x - y = x + - 1 \<cdot> y"
   16.41 +    by (simp add: diff_eq2 negate_eq2a)
   16.42 +  also from x y have "\<parallel>\<dots>\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>- 1 \<cdot> y\<parallel>"
   16.43 +    by (simp add: subadditive)
   16.44 +  also from y have "\<parallel>- 1 \<cdot> y\<parallel> = \<bar>- 1\<bar> * \<parallel>y\<parallel>"
   16.45 +    by (rule abs_homogenous)
   16.46 +  also have "\<dots> = \<parallel>y\<parallel>" by simp
   16.47 +  finally show ?thesis .
   16.48 +qed
   16.49 +
   16.50 +lemma (in seminorm) minus:
   16.51 +  assumes "vectorspace V"
   16.52 +  shows "x \<in> V \<Longrightarrow> \<parallel>- x\<parallel> = \<parallel>x\<parallel>"
   16.53 +proof -
   16.54 +  interpret vectorspace [V] by fact
   16.55 +  assume x: "x \<in> V"
   16.56 +  then have "- x = - 1 \<cdot> x" by (simp only: negate_eq1)
   16.57 +  also from x have "\<parallel>\<dots>\<parallel> = \<bar>- 1\<bar> * \<parallel>x\<parallel>"
   16.58 +    by (rule abs_homogenous)
   16.59 +  also have "\<dots> = \<parallel>x\<parallel>" by simp
   16.60 +  finally show ?thesis .
   16.61 +qed
   16.62 +
   16.63 +
   16.64 +subsection {* Norms *}
   16.65 +
   16.66 +text {*
   16.67 +  A \emph{norm} @{text "\<parallel>\<cdot>\<parallel>"} is a seminorm that maps only the
   16.68 +  @{text 0} vector to @{text 0}.
   16.69 +*}
   16.70 +
   16.71 +locale norm = seminorm +
   16.72 +  assumes zero_iff [iff]: "x \<in> V \<Longrightarrow> (\<parallel>x\<parallel> = 0) = (x = 0)"
   16.73 +
   16.74 +
   16.75 +subsection {* Normed vector spaces *}
   16.76 +
   16.77 +text {*
   16.78 +  A vector space together with a norm is called a \emph{normed
   16.79 +  space}.
   16.80 +*}
   16.81 +
   16.82 +locale normed_vectorspace = vectorspace + norm
   16.83 +
   16.84 +declare normed_vectorspace.intro [intro?]
   16.85 +
   16.86 +lemma (in normed_vectorspace) gt_zero [intro?]:
   16.87 +  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> 0 < \<parallel>x\<parallel>"
   16.88 +proof -
   16.89 +  assume x: "x \<in> V" and neq: "x \<noteq> 0"
   16.90 +  from x have "0 \<le> \<parallel>x\<parallel>" ..
   16.91 +  also have [symmetric]: "\<dots> \<noteq> 0"
   16.92 +  proof
   16.93 +    assume "\<parallel>x\<parallel> = 0"
   16.94 +    with x have "x = 0" by simp
   16.95 +    with neq show False by contradiction
   16.96 +  qed
   16.97 +  finally show ?thesis .
   16.98 +qed
   16.99 +
  16.100 +text {*
  16.101 +  Any subspace of a normed vector space is again a normed vectorspace.
  16.102 +*}
  16.103 +
  16.104 +lemma subspace_normed_vs [intro?]:
  16.105 +  fixes F E norm
  16.106 +  assumes "subspace F E" "normed_vectorspace E norm"
  16.107 +  shows "normed_vectorspace F norm"
  16.108 +proof -
  16.109 +  interpret subspace [F E] by fact
  16.110 +  interpret normed_vectorspace [E norm] by fact
  16.111 +  show ?thesis
  16.112 +  proof
  16.113 +    show "vectorspace F" by (rule vectorspace) unfold_locales
  16.114 +  next
  16.115 +    have "NormedSpace.norm E norm" ..
  16.116 +    with subset show "NormedSpace.norm F norm"
  16.117 +      by (simp add: norm_def seminorm_def norm_axioms_def)
  16.118 +  qed
  16.119 +qed
  16.120 +
  16.121 +end
    17.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    17.2 +++ b/src/HOL/HahnBanach/README.html	Mon Dec 29 14:08:08 2008 +0100
    17.3 @@ -0,0 +1,38 @@
    17.4 +<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
    17.5 +
    17.6 +<!-- $Id$ -->
    17.7 +
    17.8 +<HTML>
    17.9 +
   17.10 +<HEAD>
   17.11 +  <meta http-equiv="Content-Type" content="text/html; charset=iso-8859-1">
   17.12 +  <TITLE>HOL/Real/HahnBanach/README</TITLE>
   17.13 +</HEAD>
   17.14 +
   17.15 +<BODY>
   17.16 +
   17.17 +<H3>The Hahn-Banach Theorem for Real Vector Spaces (Isabelle/Isar)</H3>
   17.18 +
   17.19 +Author: Gertrud Bauer, Technische Universit&auml;t M&uuml;nchen<P>
   17.20 +
   17.21 +This directory contains the proof of the Hahn-Banach theorem for real vectorspaces,
   17.22 +following H. Heuser, Funktionalanalysis, p. 228 -232.
   17.23 +The Hahn-Banach theorem is one of the fundamental theorems of functioal analysis.
   17.24 +It is a conclusion of Zorn's lemma.<P>
   17.25 +
   17.26 +Two different formaulations of the theorem are presented, one for general real vectorspaces
   17.27 +and its application to normed vectorspaces. <P>
   17.28 +
   17.29 +The theorem says, that every continous linearform, defined on arbitrary subspaces
   17.30 +(not only one-dimensional subspaces), can be extended to a continous linearform on
   17.31 +the whole vectorspace.
   17.32 +
   17.33 +
   17.34 +<HR>
   17.35 +
   17.36 +<ADDRESS>
   17.37 +<A NAME="bauerg@in.tum.de" HREF="mailto:bauerg@in.tum.de">bauerg@in.tum.de</A>
   17.38 +</ADDRESS>
   17.39 +
   17.40 +</BODY>
   17.41 +</HTML>
    18.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    18.2 +++ b/src/HOL/HahnBanach/ROOT.ML	Mon Dec 29 14:08:08 2008 +0100
    18.3 @@ -0,0 +1,8 @@
    18.4 +(*  Title:      HOL/Real/HahnBanach/ROOT.ML
    18.5 +    ID:         $Id$
    18.6 +    Author:     Gertrud Bauer, TU Munich
    18.7 +
    18.8 +The Hahn-Banach theorem for real vector spaces (Isabelle/Isar).
    18.9 +*)
   18.10 +
   18.11 +time_use_thy "HahnBanach";
    19.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    19.2 +++ b/src/HOL/HahnBanach/Subspace.thy	Mon Dec 29 14:08:08 2008 +0100
    19.3 @@ -0,0 +1,514 @@
    19.4 +(*  Title:      HOL/Real/HahnBanach/Subspace.thy
    19.5 +    ID:         $Id$
    19.6 +    Author:     Gertrud Bauer, TU Munich
    19.7 +*)
    19.8 +
    19.9 +header {* Subspaces *}
   19.10 +
   19.11 +theory Subspace
   19.12 +imports VectorSpace
   19.13 +begin
   19.14 +
   19.15 +subsection {* Definition *}
   19.16 +
   19.17 +text {*
   19.18 +  A non-empty subset @{text U} of a vector space @{text V} is a
   19.19 +  \emph{subspace} of @{text V}, iff @{text U} is closed under addition
   19.20 +  and scalar multiplication.
   19.21 +*}
   19.22 +
   19.23 +locale subspace = var U + var V +
   19.24 +  constrains U :: "'a\<Colon>{minus, plus, zero, uminus} set"
   19.25 +  assumes non_empty [iff, intro]: "U \<noteq> {}"
   19.26 +    and subset [iff]: "U \<subseteq> V"
   19.27 +    and add_closed [iff]: "x \<in> U \<Longrightarrow> y \<in> U \<Longrightarrow> x + y \<in> U"
   19.28 +    and mult_closed [iff]: "x \<in> U \<Longrightarrow> a \<cdot> x \<in> U"
   19.29 +
   19.30 +notation (symbols)
   19.31 +  subspace  (infix "\<unlhd>" 50)
   19.32 +
   19.33 +declare vectorspace.intro [intro?] subspace.intro [intro?]
   19.34 +
   19.35 +lemma subspace_subset [elim]: "U \<unlhd> V \<Longrightarrow> U \<subseteq> V"
   19.36 +  by (rule subspace.subset)
   19.37 +
   19.38 +lemma (in subspace) subsetD [iff]: "x \<in> U \<Longrightarrow> x \<in> V"
   19.39 +  using subset by blast
   19.40 +
   19.41 +lemma subspaceD [elim]: "U \<unlhd> V \<Longrightarrow> x \<in> U \<Longrightarrow> x \<in> V"
   19.42 +  by (rule subspace.subsetD)
   19.43 +
   19.44 +lemma rev_subspaceD [elim?]: "x \<in> U \<Longrightarrow> U \<unlhd> V \<Longrightarrow> x \<in> V"
   19.45 +  by (rule subspace.subsetD)
   19.46 +
   19.47 +lemma (in subspace) diff_closed [iff]:
   19.48 +  assumes "vectorspace V"
   19.49 +  assumes x: "x \<in> U" and y: "y \<in> U"
   19.50 +  shows "x - y \<in> U"
   19.51 +proof -
   19.52 +  interpret vectorspace [V] by fact
   19.53 +  from x y show ?thesis by (simp add: diff_eq1 negate_eq1)
   19.54 +qed
   19.55 +
   19.56 +text {*
   19.57 +  \medskip Similar as for linear spaces, the existence of the zero
   19.58 +  element in every subspace follows from the non-emptiness of the
   19.59 +  carrier set and by vector space laws.
   19.60 +*}
   19.61 +
   19.62 +lemma (in subspace) zero [intro]:
   19.63 +  assumes "vectorspace V"
   19.64 +  shows "0 \<in> U"
   19.65 +proof -
   19.66 +  interpret vectorspace [V] by fact
   19.67 +  have "U \<noteq> {}" by (rule U_V.non_empty)
   19.68 +  then obtain x where x: "x \<in> U" by blast
   19.69 +  then have "x \<in> V" .. then have "0 = x - x" by simp
   19.70 +  also from `vectorspace V` x x have "\<dots> \<in> U" by (rule U_V.diff_closed)
   19.71 +  finally show ?thesis .
   19.72 +qed
   19.73 +
   19.74 +lemma (in subspace) neg_closed [iff]:
   19.75 +  assumes "vectorspace V"
   19.76 +  assumes x: "x \<in> U"
   19.77 +  shows "- x \<in> U"
   19.78 +proof -
   19.79 +  interpret vectorspace [V] by fact
   19.80 +  from x show ?thesis by (simp add: negate_eq1)
   19.81 +qed
   19.82 +
   19.83 +text {* \medskip Further derived laws: every subspace is a vector space. *}
   19.84 +
   19.85 +lemma (in subspace) vectorspace [iff]:
   19.86 +  assumes "vectorspace V"
   19.87 +  shows "vectorspace U"
   19.88 +proof -
   19.89 +  interpret vectorspace [V] by fact
   19.90 +  show ?thesis
   19.91 +  proof
   19.92 +    show "U \<noteq> {}" ..
   19.93 +    fix x y z assume x: "x \<in> U" and y: "y \<in> U" and z: "z \<in> U"
   19.94 +    fix a b :: real
   19.95 +    from x y show "x + y \<in> U" by simp
   19.96 +    from x show "a \<cdot> x \<in> U" by simp
   19.97 +    from x y z show "(x + y) + z = x + (y + z)" by (simp add: add_ac)
   19.98 +    from x y show "x + y = y + x" by (simp add: add_ac)
   19.99 +    from x show "x - x = 0" by simp
  19.100 +    from x show "0 + x = x" by simp
  19.101 +    from x y show "a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y" by (simp add: distrib)
  19.102 +    from x show "(a + b) \<cdot> x = a \<cdot> x + b \<cdot> x" by (simp add: distrib)
  19.103 +    from x show "(a * b) \<cdot> x = a \<cdot> b \<cdot> x" by (simp add: mult_assoc)
  19.104 +    from x show "1 \<cdot> x = x" by simp
  19.105 +    from x show "- x = - 1 \<cdot> x" by (simp add: negate_eq1)
  19.106 +    from x y show "x - y = x + - y" by (simp add: diff_eq1)
  19.107 +  qed
  19.108 +qed
  19.109 +
  19.110 +
  19.111 +text {* The subspace relation is reflexive. *}
  19.112 +
  19.113 +lemma (in vectorspace) subspace_refl [intro]: "V \<unlhd> V"
  19.114 +proof
  19.115 +  show "V \<noteq> {}" ..
  19.116 +  show "V \<subseteq> V" ..
  19.117 +  fix x y assume x: "x \<in> V" and y: "y \<in> V"
  19.118 +  fix a :: real
  19.119 +  from x y show "x + y \<in> V" by simp
  19.120 +  from x show "a \<cdot> x \<in> V" by simp
  19.121 +qed
  19.122 +
  19.123 +text {* The subspace relation is transitive. *}
  19.124 +
  19.125 +lemma (in vectorspace) subspace_trans [trans]:
  19.126 +  "U \<unlhd> V \<Longrightarrow> V \<unlhd> W \<Longrightarrow> U \<unlhd> W"
  19.127 +proof
  19.128 +  assume uv: "U \<unlhd> V" and vw: "V \<unlhd> W"
  19.129 +  from uv show "U \<noteq> {}" by (rule subspace.non_empty)
  19.130 +  show "U \<subseteq> W"
  19.131 +  proof -
  19.132 +    from uv have "U \<subseteq> V" by (rule subspace.subset)
  19.133 +    also from vw have "V \<subseteq> W" by (rule subspace.subset)
  19.134 +    finally show ?thesis .
  19.135 +  qed
  19.136 +  fix x y assume x: "x \<in> U" and y: "y \<in> U"
  19.137 +  from uv and x y show "x + y \<in> U" by (rule subspace.add_closed)
  19.138 +  from uv and x show "\<And>a. a \<cdot> x \<in> U" by (rule subspace.mult_closed)
  19.139 +qed
  19.140 +
  19.141 +
  19.142 +subsection {* Linear closure *}
  19.143 +
  19.144 +text {*
  19.145 +  The \emph{linear closure} of a vector @{text x} is the set of all
  19.146 +  scalar multiples of @{text x}.
  19.147 +*}
  19.148 +
  19.149 +definition
  19.150 +  lin :: "('a::{minus, plus, zero}) \<Rightarrow> 'a set" where
  19.151 +  "lin x = {a \<cdot> x | a. True}"
  19.152 +
  19.153 +lemma linI [intro]: "y = a \<cdot> x \<Longrightarrow> y \<in> lin x"
  19.154 +  unfolding lin_def by blast
  19.155 +
  19.156 +lemma linI' [iff]: "a \<cdot> x \<in> lin x"
  19.157 +  unfolding lin_def by blast
  19.158 +
  19.159 +lemma linE [elim]: "x \<in> lin v \<Longrightarrow> (\<And>a::real. x = a \<cdot> v \<Longrightarrow> C) \<Longrightarrow> C"
  19.160 +  unfolding lin_def by blast
  19.161 +
  19.162 +
  19.163 +text {* Every vector is contained in its linear closure. *}
  19.164 +
  19.165 +lemma (in vectorspace) x_lin_x [iff]: "x \<in> V \<Longrightarrow> x \<in> lin x"
  19.166 +proof -
  19.167 +  assume "x \<in> V"
  19.168 +  then have "x = 1 \<cdot> x" by simp
  19.169 +  also have "\<dots> \<in> lin x" ..
  19.170 +  finally show ?thesis .
  19.171 +qed
  19.172 +
  19.173 +lemma (in vectorspace) "0_lin_x" [iff]: "x \<in> V \<Longrightarrow> 0 \<in> lin x"
  19.174 +proof
  19.175 +  assume "x \<in> V"
  19.176 +  then show "0 = 0 \<cdot> x" by simp
  19.177 +qed
  19.178 +
  19.179 +text {* Any linear closure is a subspace. *}
  19.180 +
  19.181 +lemma (in vectorspace) lin_subspace [intro]:
  19.182 +  "x \<in> V \<Longrightarrow> lin x \<unlhd> V"
  19.183 +proof
  19.184 +  assume x: "x \<in> V"
  19.185 +  then show "lin x \<noteq> {}" by (auto simp add: x_lin_x)
  19.186 +  show "lin x \<subseteq> V"
  19.187 +  proof
  19.188 +    fix x' assume "x' \<in> lin x"
  19.189 +    then obtain a where "x' = a \<cdot> x" ..
  19.190 +    with x show "x' \<in> V" by simp
  19.191 +  qed
  19.192 +  fix x' x'' assume x': "x' \<in> lin x" and x'': "x'' \<in> lin x"
  19.193 +  show "x' + x'' \<in> lin x"
  19.194 +  proof -
  19.195 +    from x' obtain a' where "x' = a' \<cdot> x" ..
  19.196 +    moreover from x'' obtain a'' where "x'' = a'' \<cdot> x" ..
  19.197 +    ultimately have "x' + x'' = (a' + a'') \<cdot> x"
  19.198 +      using x by (simp add: distrib)
  19.199 +    also have "\<dots> \<in> lin x" ..
  19.200 +    finally show ?thesis .
  19.201 +  qed
  19.202 +  fix a :: real
  19.203 +  show "a \<cdot> x' \<in> lin x"
  19.204 +  proof -
  19.205 +    from x' obtain a' where "x' = a' \<cdot> x" ..
  19.206 +    with x have "a \<cdot> x' = (a * a') \<cdot> x" by (simp add: mult_assoc)
  19.207 +    also have "\<dots> \<in> lin x" ..
  19.208 +    finally show ?thesis .
  19.209 +  qed
  19.210 +qed
  19.211 +
  19.212 +
  19.213 +text {* Any linear closure is a vector space. *}
  19.214 +
  19.215 +lemma (in vectorspace) lin_vectorspace [intro]:
  19.216 +  assumes "x \<in> V"
  19.217 +  shows "vectorspace (lin x)"
  19.218 +proof -
  19.219 +  from `x \<in> V` have "subspace (lin x) V"
  19.220 +    by (rule lin_subspace)
  19.221 +  from this and vectorspace_axioms show ?thesis
  19.222 +    by (rule subspace.vectorspace)
  19.223 +qed
  19.224 +
  19.225 +
  19.226 +subsection {* Sum of two vectorspaces *}
  19.227 +
  19.228 +text {*
  19.229 +  The \emph{sum} of two vectorspaces @{text U} and @{text V} is the
  19.230 +  set of all sums of elements from @{text U} and @{text V}.
  19.231 +*}
  19.232 +
  19.233 +instantiation "fun" :: (type, type) plus
  19.234 +begin
  19.235 +
  19.236 +definition
  19.237 +  sum_def: "plus_fun U V = {u + v | u v. u \<in> U \<and> v \<in> V}"  (* FIXME not fully general!? *)
  19.238 +
  19.239 +instance ..
  19.240 +
  19.241 +end
  19.242 +
  19.243 +lemma sumE [elim]:
  19.244 +    "x \<in> U + V \<Longrightarrow> (\<And>u v. x = u + v \<Longrightarrow> u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> C) \<Longrightarrow> C"
  19.245 +  unfolding sum_def by blast
  19.246 +
  19.247 +lemma sumI [intro]:
  19.248 +    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> x = u + v \<Longrightarrow> x \<in> U + V"
  19.249 +  unfolding sum_def by blast
  19.250 +
  19.251 +lemma sumI' [intro]:
  19.252 +    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> u + v \<in> U + V"
  19.253 +  unfolding sum_def by blast
  19.254 +
  19.255 +text {* @{text U} is a subspace of @{text "U + V"}. *}
  19.256 +
  19.257 +lemma subspace_sum1 [iff]:
  19.258 +  assumes "vectorspace U" "vectorspace V"
  19.259 +  shows "U \<unlhd> U + V"
  19.260 +proof -
  19.261 +  interpret vectorspace [U] by fact
  19.262 +  interpret vectorspace [V] by fact
  19.263 +  show ?thesis
  19.264 +  proof
  19.265 +    show "U \<noteq> {}" ..
  19.266 +    show "U \<subseteq> U + V"
  19.267 +    proof
  19.268 +      fix x assume x: "x \<in> U"
  19.269 +      moreover have "0 \<in> V" ..
  19.270 +      ultimately have "x + 0 \<in> U + V" ..
  19.271 +      with x show "x \<in> U + V" by simp
  19.272 +    qed
  19.273 +    fix x y assume x: "x \<in> U" and "y \<in> U"
  19.274 +    then show "x + y \<in> U" by simp
  19.275 +    from x show "\<And>a. a \<cdot> x \<in> U" by simp
  19.276 +  qed
  19.277 +qed
  19.278 +
  19.279 +text {* The sum of two subspaces is again a subspace. *}
  19.280 +
  19.281 +lemma sum_subspace [intro?]:
  19.282 +  assumes "subspace U E" "vectorspace E" "subspace V E"
  19.283 +  shows "U + V \<unlhd> E"
  19.284 +proof -
  19.285 +  interpret subspace [U E] by fact
  19.286 +  interpret vectorspace [E] by fact
  19.287 +  interpret subspace [V E] by fact
  19.288 +  show ?thesis
  19.289 +  proof
  19.290 +    have "0 \<in> U + V"
  19.291 +    proof
  19.292 +      show "0 \<in> U" using `vectorspace E` ..
  19.293 +      show "0 \<in> V" using `vectorspace E` ..
  19.294 +      show "(0::'a) = 0 + 0" by simp
  19.295 +    qed
  19.296 +    then show "U + V \<noteq> {}" by blast
  19.297 +    show "U + V \<subseteq> E"
  19.298 +    proof
  19.299 +      fix x assume "x \<in> U + V"
  19.300 +      then obtain u v where "x = u + v" and
  19.301 +	"u \<in> U" and "v \<in> V" ..
  19.302 +      then show "x \<in> E" by simp
  19.303 +    qed
  19.304 +    fix x y assume x: "x \<in> U + V" and y: "y \<in> U + V"
  19.305 +    show "x + y \<in> U + V"
  19.306 +    proof -
  19.307 +      from x obtain ux vx where "x = ux + vx" and "ux \<in> U" and "vx \<in> V" ..
  19.308 +      moreover
  19.309 +      from y obtain uy vy where "y = uy + vy" and "uy \<in> U" and "vy \<in> V" ..
  19.310 +      ultimately
  19.311 +      have "ux + uy \<in> U"
  19.312 +	and "vx + vy \<in> V"
  19.313 +	and "x + y = (ux + uy) + (vx + vy)"
  19.314 +	using x y by (simp_all add: add_ac)
  19.315 +      then show ?thesis ..
  19.316 +    qed
  19.317 +    fix a show "a \<cdot> x \<in> U + V"
  19.318 +    proof -
  19.319 +      from x obtain u v where "x = u + v" and "u \<in> U" and "v \<in> V" ..
  19.320 +      then have "a \<cdot> u \<in> U" and "a \<cdot> v \<in> V"
  19.321 +	and "a \<cdot> x = (a \<cdot> u) + (a \<cdot> v)" by (simp_all add: distrib)
  19.322 +      then show ?thesis ..
  19.323 +    qed
  19.324 +  qed
  19.325 +qed
  19.326 +
  19.327 +text{* The sum of two subspaces is a vectorspace. *}
  19.328 +
  19.329 +lemma sum_vs [intro?]:
  19.330 +    "U \<unlhd> E \<Longrightarrow> V \<unlhd> E \<Longrightarrow> vectorspace E \<Longrightarrow> vectorspace (U + V)"
  19.331 +  by (rule subspace.vectorspace) (rule sum_subspace)
  19.332 +
  19.333 +
  19.334 +subsection {* Direct sums *}
  19.335 +
  19.336 +text {*
  19.337 +  The sum of @{text U} and @{text V} is called \emph{direct}, iff the
  19.338 +  zero element is the only common element of @{text U} and @{text
  19.339 +  V}. For every element @{text x} of the direct sum of @{text U} and
  19.340 +  @{text V} the decomposition in @{text "x = u + v"} with
  19.341 +  @{text "u \<in> U"} and @{text "v \<in> V"} is unique.
  19.342 +*}
  19.343 +
  19.344 +lemma decomp:
  19.345 +  assumes "vectorspace E" "subspace U E" "subspace V E"
  19.346 +  assumes direct: "U \<inter> V = {0}"
  19.347 +    and u1: "u1 \<in> U" and u2: "u2 \<in> U"
  19.348 +    and v1: "v1 \<in> V" and v2: "v2 \<in> V"
  19.349 +    and sum: "u1 + v1 = u2 + v2"
  19.350 +  shows "u1 = u2 \<and> v1 = v2"
  19.351 +proof -
  19.352 +  interpret vectorspace [E] by fact
  19.353 +  interpret subspace [U E] by fact
  19.354 +  interpret subspace [V E] by fact
  19.355 +  show ?thesis
  19.356 +  proof
  19.357 +    have U: "vectorspace U"  (* FIXME: use interpret *)
  19.358 +      using `subspace U E` `vectorspace E` by (rule subspace.vectorspace)
  19.359 +    have V: "vectorspace V"
  19.360 +      using `subspace V E` `vectorspace E` by (rule subspace.vectorspace)
  19.361 +    from u1 u2 v1 v2 and sum have eq: "u1 - u2 = v2 - v1"
  19.362 +      by (simp add: add_diff_swap)
  19.363 +    from u1 u2 have u: "u1 - u2 \<in> U"
  19.364 +      by (rule vectorspace.diff_closed [OF U])
  19.365 +    with eq have v': "v2 - v1 \<in> U" by (simp only:)
  19.366 +    from v2 v1 have v: "v2 - v1 \<in> V"
  19.367 +      by (rule vectorspace.diff_closed [OF V])
  19.368 +    with eq have u': " u1 - u2 \<in> V" by (simp only:)
  19.369 +    
  19.370 +    show "u1 = u2"
  19.371 +    proof (rule add_minus_eq)
  19.372 +      from u1 show "u1 \<in> E" ..
  19.373 +      from u2 show "u2 \<in> E" ..
  19.374 +      from u u' and direct show "u1 - u2 = 0" by blast
  19.375 +    qed
  19.376 +    show "v1 = v2"
  19.377 +    proof (rule add_minus_eq [symmetric])
  19.378 +      from v1 show "v1 \<in> E" ..
  19.379 +      from v2 show "v2 \<in> E" ..
  19.380 +      from v v' and direct show "v2 - v1 = 0" by blast
  19.381 +    qed
  19.382 +  qed
  19.383 +qed
  19.384 +
  19.385 +text {*
  19.386 +  An application of the previous lemma will be used in the proof of
  19.387 +  the Hahn-Banach Theorem (see page \pageref{decomp-H-use}): for any
  19.388 +  element @{text "y + a \<cdot> x\<^sub>0"} of the direct sum of a
  19.389 +  vectorspace @{text H} and the linear closure of @{text "x\<^sub>0"}
  19.390 +  the components @{text "y \<in> H"} and @{text a} are uniquely
  19.391 +  determined.
  19.392 +*}
  19.393 +
  19.394 +lemma decomp_H':
  19.395 +  assumes "vectorspace E" "subspace H E"
  19.396 +  assumes y1: "y1 \<in> H" and y2: "y2 \<in> H"
  19.397 +    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
  19.398 +    and eq: "y1 + a1 \<cdot> x' = y2 + a2 \<cdot> x'"
  19.399 +  shows "y1 = y2 \<and> a1 = a2"
  19.400 +proof -
  19.401 +  interpret vectorspace [E] by fact
  19.402 +  interpret subspace [H E] by fact
  19.403 +  show ?thesis
  19.404 +  proof
  19.405 +    have c: "y1 = y2 \<and> a1 \<cdot> x' = a2 \<cdot> x'"
  19.406 +    proof (rule decomp)
  19.407 +      show "a1 \<cdot> x' \<in> lin x'" ..
  19.408 +      show "a2 \<cdot> x' \<in> lin x'" ..
  19.409 +      show "H \<inter> lin x' = {0}"
  19.410 +      proof
  19.411 +	show "H \<inter> lin x' \<subseteq> {0}"
  19.412 +	proof
  19.413 +          fix x assume x: "x \<in> H \<inter> lin x'"
  19.414 +          then obtain a where xx': "x = a \<cdot> x'"
  19.415 +            by blast
  19.416 +          have "x = 0"
  19.417 +          proof cases
  19.418 +            assume "a = 0"
  19.419 +            with xx' and x' show ?thesis by simp
  19.420 +          next
  19.421 +            assume a: "a \<noteq> 0"
  19.422 +            from x have "x \<in> H" ..
  19.423 +            with xx' have "inverse a \<cdot> a \<cdot> x' \<in> H" by simp
  19.424 +            with a and x' have "x' \<in> H" by (simp add: mult_assoc2)
  19.425 +            with `x' \<notin> H` show ?thesis by contradiction
  19.426 +          qed
  19.427 +          then show "x \<in> {0}" ..
  19.428 +	qed
  19.429 +	show "{0} \<subseteq> H \<inter> lin x'"
  19.430 +	proof -
  19.431 +          have "0 \<in> H" using `vectorspace E` ..
  19.432 +          moreover have "0 \<in> lin x'" using `x' \<in> E` ..
  19.433 +          ultimately show ?thesis by blast
  19.434 +	qed
  19.435 +      qed
  19.436 +      show "lin x' \<unlhd> E" using `x' \<in> E` ..
  19.437 +    qed (rule `vectorspace E`, rule `subspace H E`, rule y1, rule y2, rule eq)
  19.438 +    then show "y1 = y2" ..
  19.439 +    from c have "a1 \<cdot> x' = a2 \<cdot> x'" ..
  19.440 +    with x' show "a1 = a2" by (simp add: mult_right_cancel)
  19.441 +  qed
  19.442 +qed
  19.443 +
  19.444 +text {*
  19.445 +  Since for any element @{text "y + a \<cdot> x'"} of the direct sum of a
  19.446 +  vectorspace @{text H} and the linear closure of @{text x'} the
  19.447 +  components @{text "y \<in> H"} and @{text a} are unique, it follows from
  19.448 +  @{text "y \<in> H"} that @{text "a = 0"}.
  19.449 +*}
  19.450 +
  19.451 +lemma decomp_H'_H:
  19.452 +  assumes "vectorspace E" "subspace H E"
  19.453 +  assumes t: "t \<in> H"
  19.454 +    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
  19.455 +  shows "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
  19.456 +proof -
  19.457 +  interpret vectorspace [E] by fact
  19.458 +  interpret subspace [H E] by fact
  19.459 +  show ?thesis
  19.460 +  proof (rule, simp_all only: split_paired_all split_conv)
  19.461 +    from t x' show "t = t + 0 \<cdot> x' \<and> t \<in> H" by simp
  19.462 +    fix y and a assume ya: "t = y + a \<cdot> x' \<and> y \<in> H"
  19.463 +    have "y = t \<and> a = 0"
  19.464 +    proof (rule decomp_H')
  19.465 +      from ya x' show "y + a \<cdot> x' = t + 0 \<cdot> x'" by simp
  19.466 +      from ya show "y \<in> H" ..
  19.467 +    qed (rule `vectorspace E`, rule `subspace H E`, rule t, (rule x')+)
  19.468 +    with t x' show "(y, a) = (y + a \<cdot> x', 0)" by simp
  19.469 +  qed
  19.470 +qed
  19.471 +
  19.472 +text {*
  19.473 +  The components @{text "y \<in> H"} and @{text a} in @{text "y + a \<cdot> x'"}
  19.474 +  are unique, so the function @{text h'} defined by
  19.475 +  @{text "h' (y + a \<cdot> x') = h y + a \<cdot> \<xi>"} is definite.
  19.476 +*}
  19.477 +
  19.478 +lemma h'_definite:
  19.479 +  fixes H
  19.480 +  assumes h'_def:
  19.481 +    "h' \<equiv> (\<lambda>x. let (y, a) = SOME (y, a). (x = y + a \<cdot> x' \<and> y \<in> H)
  19.482 +                in (h y) + a * xi)"
  19.483 +    and x: "x = y + a \<cdot> x'"
  19.484 +  assumes "vectorspace E" "subspace H E"
  19.485 +  assumes y: "y \<in> H"
  19.486 +    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
  19.487 +  shows "h' x = h y + a * xi"
  19.488 +proof -
  19.489 +  interpret vectorspace [E] by fact
  19.490 +  interpret subspace [H E] by fact
  19.491 +  from x y x' have "x \<in> H + lin x'" by auto
  19.492 +  have "\<exists>!p. (\<lambda>(y, a). x = y + a \<cdot> x' \<and> y \<in> H) p" (is "\<exists>!p. ?P p")
  19.493 +  proof (rule ex_ex1I)
  19.494 +    from x y show "\<exists>p. ?P p" by blast
  19.495 +    fix p q assume p: "?P p" and q: "?P q"
  19.496 +    show "p = q"
  19.497 +    proof -
  19.498 +      from p have xp: "x = fst p + snd p \<cdot> x' \<and> fst p \<in> H"
  19.499 +        by (cases p) simp
  19.500 +      from q have xq: "x = fst q + snd q \<cdot> x' \<and> fst q \<in> H"
  19.501 +        by (cases q) simp
  19.502 +      have "fst p = fst q \<and> snd p = snd q"
  19.503 +      proof (rule decomp_H')
  19.504 +        from xp show "fst p \<in> H" ..
  19.505 +        from xq show "fst q \<in> H" ..
  19.506 +        from xp and xq show "fst p + snd p \<cdot> x' = fst q + snd q \<cdot> x'"
  19.507 +          by simp
  19.508 +      qed (rule `vectorspace E`, rule `subspace H E`, (rule x')+)
  19.509 +      then show ?thesis by (cases p, cases q) simp
  19.510 +    qed
  19.511 +  qed
  19.512 +  then have eq: "(SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H) = (y, a)"
  19.513 +    by (rule some1_equality) (simp add: x y)
  19.514 +  with h'_def show "h' x = h y + a * xi" by (simp add: Let_def)
  19.515 +qed
  19.516 +
  19.517 +end
    20.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    20.2 +++ b/src/HOL/HahnBanach/VectorSpace.thy	Mon Dec 29 14:08:08 2008 +0100
    20.3 @@ -0,0 +1,417 @@
    20.4 +(*  Title:      HOL/Real/HahnBanach/VectorSpace.thy
    20.5 +    ID:         $Id$
    20.6 +    Author:     Gertrud Bauer, TU Munich
    20.7 +*)
    20.8 +
    20.9 +header {* Vector spaces *}
   20.10 +
   20.11 +theory VectorSpace
   20.12 +imports Real Bounds Zorn
   20.13 +begin
   20.14 +
   20.15 +subsection {* Signature *}
   20.16 +
   20.17 +text {*
   20.18 +  For the definition of real vector spaces a type @{typ 'a} of the
   20.19 +  sort @{text "{plus, minus, zero}"} is considered, on which a real
   20.20 +  scalar multiplication @{text \<cdot>} is declared.
   20.21 +*}
   20.22 +
   20.23 +consts
   20.24 +  prod  :: "real \<Rightarrow> 'a::{plus, minus, zero} \<Rightarrow> 'a"     (infixr "'(*')" 70)
   20.25 +
   20.26 +notation (xsymbols)
   20.27 +  prod  (infixr "\<cdot>" 70)
   20.28 +notation (HTML output)
   20.29 +  prod  (infixr "\<cdot>" 70)
   20.30 +
   20.31 +
   20.32 +subsection {* Vector space laws *}
   20.33 +
   20.34 +text {*
   20.35 +  A \emph{vector space} is a non-empty set @{text V} of elements from
   20.36 +  @{typ 'a} with the following vector space laws: The set @{text V} is
   20.37 +  closed under addition and scalar multiplication, addition is
   20.38 +  associative and commutative; @{text "- x"} is the inverse of @{text
   20.39 +  x} w.~r.~t.~addition and @{text 0} is the neutral element of
   20.40 +  addition.  Addition and multiplication are distributive; scalar
   20.41 +  multiplication is associative and the real number @{text "1"} is
   20.42 +  the neutral element of scalar multiplication.
   20.43 +*}
   20.44 +
   20.45 +locale vectorspace = var V +
   20.46 +  assumes non_empty [iff, intro?]: "V \<noteq> {}"
   20.47 +    and add_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y \<in> V"
   20.48 +    and mult_closed [iff]: "x \<in> V \<Longrightarrow> a \<cdot> x \<in> V"
   20.49 +    and add_assoc: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y) + z = x + (y + z)"
   20.50 +    and add_commute: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = y + x"
   20.51 +    and diff_self [simp]: "x \<in> V \<Longrightarrow> x - x = 0"
   20.52 +    and add_zero_left [simp]: "x \<in> V \<Longrightarrow> 0 + x = x"
   20.53 +    and add_mult_distrib1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y"
   20.54 +    and add_mult_distrib2: "x \<in> V \<Longrightarrow> (a + b) \<cdot> x = a \<cdot> x + b \<cdot> x"
   20.55 +    and mult_assoc: "x \<in> V \<Longrightarrow> (a * b) \<cdot> x = a \<cdot> (b \<cdot> x)"
   20.56 +    and mult_1 [simp]: "x \<in> V \<Longrightarrow> 1 \<cdot> x = x"
   20.57 +    and negate_eq1: "x \<in> V \<Longrightarrow> - x = (- 1) \<cdot> x"
   20.58 +    and diff_eq1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = x + - y"
   20.59 +
   20.60 +lemma (in vectorspace) negate_eq2: "x \<in> V \<Longrightarrow> (- 1) \<cdot> x = - x"
   20.61 +  by (rule negate_eq1 [symmetric])
   20.62 +
   20.63 +lemma (in vectorspace) negate_eq2a: "x \<in> V \<Longrightarrow> -1 \<cdot> x = - x"
   20.64 +  by (simp add: negate_eq1)
   20.65 +
   20.66 +lemma (in vectorspace) diff_eq2: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + - y = x - y"
   20.67 +  by (rule diff_eq1 [symmetric])
   20.68 +
   20.69 +lemma (in vectorspace) diff_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y \<in> V"
   20.70 +  by (simp add: diff_eq1 negate_eq1)
   20.71 +
   20.72 +lemma (in vectorspace) neg_closed [iff]: "x \<in> V \<Longrightarrow> - x \<in> V"
   20.73 +  by (simp add: negate_eq1)
   20.74 +
   20.75 +lemma (in vectorspace) add_left_commute:
   20.76 +  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> x + (y + z) = y + (x + z)"
   20.77 +proof -
   20.78 +  assume xyz: "x \<in> V"  "y \<in> V"  "z \<in> V"
   20.79 +  then have "x + (y + z) = (x + y) + z"
   20.80 +    by (simp only: add_assoc)
   20.81 +  also from xyz have "\<dots> = (y + x) + z" by (simp only: add_commute)
   20.82 +  also from xyz have "\<dots> = y + (x + z)" by (simp only: add_assoc)
   20.83 +  finally show ?thesis .
   20.84 +qed
   20.85 +
   20.86 +theorems (in vectorspace) add_ac =
   20.87 +  add_assoc add_commute add_left_commute
   20.88 +
   20.89 +
   20.90 +text {* The existence of the zero element of a vector space
   20.91 +  follows from the non-emptiness of carrier set. *}
   20.92 +
   20.93 +lemma (in vectorspace) zero [iff]: "0 \<in> V"
   20.94 +proof -
   20.95 +  from non_empty obtain x where x: "x \<in> V" by blast
   20.96 +  then have "0 = x - x" by (rule diff_self [symmetric])
   20.97 +  also from x x have "\<dots> \<in> V" by (rule diff_closed)
   20.98 +  finally show ?thesis .
   20.99 +qed
  20.100 +
  20.101 +lemma (in vectorspace) add_zero_right [simp]:
  20.102 +  "x \<in> V \<Longrightarrow>  x + 0 = x"
  20.103 +proof -
  20.104 +  assume x: "x \<in> V"
  20.105 +  from this and zero have "x + 0 = 0 + x" by (rule add_commute)
  20.106 +  also from x have "\<dots> = x" by (rule add_zero_left)
  20.107 +  finally show ?thesis .
  20.108 +qed
  20.109 +
  20.110 +lemma (in vectorspace) mult_assoc2:
  20.111 +    "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = (a * b) \<cdot> x"
  20.112 +  by (simp only: mult_assoc)
  20.113 +
  20.114 +lemma (in vectorspace) diff_mult_distrib1:
  20.115 +    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x - y) = a \<cdot> x - a \<cdot> y"
  20.116 +  by (simp add: diff_eq1 negate_eq1 add_mult_distrib1 mult_assoc2)
  20.117 +
  20.118 +lemma (in vectorspace) diff_mult_distrib2:
  20.119 +  "x \<in> V \<Longrightarrow> (a - b) \<cdot> x = a \<cdot> x - (b \<cdot> x)"
  20.120 +proof -
  20.121 +  assume x: "x \<in> V"
  20.122 +  have " (a - b) \<cdot> x = (a + - b) \<cdot> x"
  20.123 +    by (simp add: real_diff_def)
  20.124 +  also from x have "\<dots> = a \<cdot> x + (- b) \<cdot> x"
  20.125 +    by (rule add_mult_distrib2)
  20.126 +  also from x have "\<dots> = a \<cdot> x + - (b \<cdot> x)"
  20.127 +    by (simp add: negate_eq1 mult_assoc2)
  20.128 +  also from x have "\<dots> = a \<cdot> x - (b \<cdot> x)"
  20.129 +    by (simp add: diff_eq1)
  20.130 +  finally show ?thesis .
  20.131 +qed
  20.132 +
  20.133 +lemmas (in vectorspace) distrib =
  20.134 +  add_mult_distrib1 add_mult_distrib2
  20.135 +  diff_mult_distrib1 diff_mult_distrib2
  20.136 +
  20.137 +
  20.138 +text {* \medskip Further derived laws: *}
  20.139 +
  20.140 +lemma (in vectorspace) mult_zero_left [simp]:
  20.141 +  "x \<in> V \<Longrightarrow> 0 \<cdot> x = 0"
  20.142 +proof -
  20.143 +  assume x: "x \<in> V"
  20.144 +  have "0 \<cdot> x = (1 - 1) \<cdot> x" by simp
  20.145 +  also have "\<dots> = (1 + - 1) \<cdot> x" by simp
  20.146 +  also from x have "\<dots> =  1 \<cdot> x + (- 1) \<cdot> x"
  20.147 +    by (rule add_mult_distrib2)
  20.148 +  also from x have "\<dots> = x + (- 1) \<cdot> x" by simp
  20.149 +  also from x have "\<dots> = x + - x" by (simp add: negate_eq2a)
  20.150 +  also from x have "\<dots> = x - x" by (simp add: diff_eq2)
  20.151 +  also from x have "\<dots> = 0" by simp
  20.152 +  finally show ?thesis .
  20.153 +qed
  20.154 +
  20.155 +lemma (in vectorspace) mult_zero_right [simp]:
  20.156 +  "a \<cdot> 0 = (0::'a)"
  20.157 +proof -
  20.158 +  have "a \<cdot> 0 = a \<cdot> (0 - (0::'a))" by simp
  20.159 +  also have "\<dots> =  a \<cdot> 0 - a \<cdot> 0"
  20.160 +    by (rule diff_mult_distrib1) simp_all
  20.161 +  also have "\<dots> = 0" by simp
  20.162 +  finally show ?thesis .
  20.163 +qed
  20.164 +
  20.165 +lemma (in vectorspace) minus_mult_cancel [simp]:
  20.166 +    "x \<in> V \<Longrightarrow> (- a) \<cdot> - x = a \<cdot> x"
  20.167 +  by (simp add: negate_eq1 mult_assoc2)
  20.168 +
  20.169 +lemma (in vectorspace) add_minus_left_eq_diff:
  20.170 +  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + y = y - x"
  20.171 +proof -
  20.172 +  assume xy: "x \<in> V"  "y \<in> V"
  20.173 +  then have "- x + y = y + - x" by (simp add: add_commute)
  20.174 +  also from xy have "\<dots> = y - x" by (simp add: diff_eq1)
  20.175 +  finally show ?thesis .
  20.176 +qed
  20.177 +
  20.178 +lemma (in vectorspace) add_minus [simp]:
  20.179 +    "x \<in> V \<Longrightarrow> x + - x = 0"
  20.180 +  by (simp add: diff_eq2)
  20.181 +
  20.182 +lemma (in vectorspace) add_minus_left [simp]:
  20.183 +    "x \<in> V \<Longrightarrow> - x + x = 0"
  20.184 +  by (simp add: diff_eq2 add_commute)
  20.185 +
  20.186 +lemma (in vectorspace) minus_minus [simp]:
  20.187 +    "x \<in> V \<Longrightarrow> - (- x) = x"
  20.188 +  by (simp add: negate_eq1 mult_assoc2)
  20.189 +
  20.190 +lemma (in vectorspace) minus_zero [simp]:
  20.191 +    "- (0::'a) = 0"
  20.192 +  by (simp add: negate_eq1)
  20.193 +
  20.194 +lemma (in vectorspace) minus_zero_iff [simp]:
  20.195 +  "x \<in> V \<Longrightarrow> (- x = 0) = (x = 0)"
  20.196 +proof
  20.197 +  assume x: "x \<in> V"
  20.198 +  {
  20.199 +    from x have "x = - (- x)" by (simp add: minus_minus)
  20.200 +    also assume "- x = 0"
  20.201 +    also have "- \<dots> = 0" by (rule minus_zero)
  20.202 +    finally show "x = 0" .
  20.203 +  next
  20.204 +    assume "x = 0"
  20.205 +    then show "- x = 0" by simp
  20.206 +  }
  20.207 +qed
  20.208 +
  20.209 +lemma (in vectorspace) add_minus_cancel [simp]:
  20.210 +    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + (- x + y) = y"
  20.211 +  by (simp add: add_assoc [symmetric] del: add_commute)
  20.212 +
  20.213 +lemma (in vectorspace) minus_add_cancel [simp]:
  20.214 +    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + (x + y) = y"
  20.215 +  by (simp add: add_assoc [symmetric] del: add_commute)
  20.216 +
  20.217 +lemma (in vectorspace) minus_add_distrib [simp]:
  20.218 +    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - (x + y) = - x + - y"
  20.219 +  by (simp add: negate_eq1 add_mult_distrib1)
  20.220 +
  20.221 +lemma (in vectorspace) diff_zero [simp]:
  20.222 +    "x \<in> V \<Longrightarrow> x - 0 = x"
  20.223 +  by (simp add: diff_eq1)
  20.224 +
  20.225 +lemma (in vectorspace) diff_zero_right [simp]:
  20.226 +    "x \<in> V \<Longrightarrow> 0 - x = - x"
  20.227 +  by (simp add: diff_eq1)
  20.228 +
  20.229 +lemma (in vectorspace) add_left_cancel:
  20.230 +  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y = x + z) = (y = z)"
  20.231 +proof
  20.232 +  assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V"
  20.233 +  {
  20.234 +    from y have "y = 0 + y" by simp
  20.235 +    also from x y have "\<dots> = (- x + x) + y" by simp
  20.236 +    also from x y have "\<dots> = - x + (x + y)"
  20.237 +      by (simp add: add_assoc neg_closed)
  20.238 +    also assume "x + y = x + z"
  20.239 +    also from x z have "- x + (x + z) = - x + x + z"
  20.240 +      by (simp add: add_assoc [symmetric] neg_closed)
  20.241 +    also from x z have "\<dots> = z" by simp
  20.242 +    finally show "y = z" .
  20.243 +  next
  20.244 +    assume "y = z"
  20.245 +    then show "x + y = x + z" by (simp only:)
  20.246 +  }
  20.247 +qed
  20.248 +
  20.249 +lemma (in vectorspace) add_right_cancel:
  20.250 +    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (y + x = z + x) = (y = z)"
  20.251 +  by (simp only: add_commute add_left_cancel)
  20.252 +
  20.253 +lemma (in vectorspace) add_assoc_cong:
  20.254 +  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x' \<in> V \<Longrightarrow> y' \<in> V \<Longrightarrow> z \<in> V
  20.255 +    \<Longrightarrow> x + y = x' + y' \<Longrightarrow> x + (y + z) = x' + (y' + z)"
  20.256 +  by (simp only: add_assoc [symmetric])
  20.257 +
  20.258 +lemma (in vectorspace) mult_left_commute:
  20.259 +    "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = b \<cdot> a \<cdot> x"
  20.260 +  by (simp add: real_mult_commute mult_assoc2)
  20.261 +
  20.262 +lemma (in vectorspace) mult_zero_uniq:
  20.263 +  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> a \<cdot> x = 0 \<Longrightarrow> a = 0"
  20.264 +proof (rule classical)
  20.265 +  assume a: "a \<noteq> 0"
  20.266 +  assume x: "x \<in> V"  "x \<noteq> 0" and ax: "a \<cdot> x = 0"
  20.267 +  from x a have "x = (inverse a * a) \<cdot> x" by simp
  20.268 +  also from `x \<in> V` have "\<dots> = inverse a \<cdot> (a \<cdot> x)" by (rule mult_assoc)
  20.269 +  also from ax have "\<dots> = inverse a \<cdot> 0" by simp
  20.270 +  also have "\<dots> = 0" by simp
  20.271 +  finally have "x = 0" .
  20.272 +  with `x \<noteq> 0` show "a = 0" by contradiction
  20.273 +qed
  20.274 +
  20.275 +lemma (in vectorspace) mult_left_cancel:
  20.276 +  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (a \<cdot> x = a \<cdot> y) = (x = y)"
  20.277 +proof
  20.278 +  assume x: "x \<in> V" and y: "y \<in> V" and a: "a \<noteq> 0"
  20.279 +  from x have "x = 1 \<cdot> x" by simp
  20.280 +  also from a have "\<dots> = (inverse a * a) \<cdot> x" by simp
  20.281 +  also from x have "\<dots> = inverse a \<cdot> (a \<cdot> x)"
  20.282 +    by (simp only: mult_assoc)
  20.283 +  also assume "a \<cdot> x = a \<cdot> y"
  20.284 +  also from a y have "inverse a \<cdot> \<dots> = y"
  20.285 +    by (simp add: mult_assoc2)
  20.286 +  finally show "x = y" .
  20.287 +next
  20.288 +  assume "x = y"
  20.289 +  then show "a \<cdot> x = a \<cdot> y" by (simp only:)
  20.290 +qed
  20.291 +
  20.292 +lemma (in vectorspace) mult_right_cancel:
  20.293 +  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> (a \<cdot> x = b \<cdot> x) = (a = b)"
  20.294 +proof
  20.295 +  assume x: "x \<in> V" and neq: "x \<noteq> 0"
  20.296 +  {
  20.297 +    from x have "(a - b) \<cdot> x = a \<cdot> x - b \<cdot> x"
  20.298 +      by (simp add: diff_mult_distrib2)
  20.299 +    also assume "a \<cdot> x = b \<cdot> x"
  20.300 +    with x have "a \<cdot> x - b \<cdot> x = 0" by simp
  20.301 +    finally have "(a - b) \<cdot> x = 0" .
  20.302 +    with x neq have "a - b = 0" by (rule mult_zero_uniq)
  20.303 +    then show "a = b" by simp
  20.304 +  next
  20.305 +    assume "a = b"
  20.306 +    then show "a \<cdot> x = b \<cdot> x" by (simp only:)
  20.307 +  }
  20.308 +qed
  20.309 +
  20.310 +lemma (in vectorspace) eq_diff_eq:
  20.311 +  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x = z - y) = (x + y = z)"
  20.312 +proof
  20.313 +  assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V"
  20.314 +  {
  20.315 +    assume "x = z - y"
  20.316 +    then have "x + y = z - y + y" by simp
  20.317 +    also from y z have "\<dots> = z + - y + y"
  20.318 +      by (simp add: diff_eq1)
  20.319 +    also have "\<dots> = z + (- y + y)"
  20.320 +      by (rule add_assoc) (simp_all add: y z)
  20.321 +    also from y z have "\<dots> = z + 0"
  20.322 +      by (simp only: add_minus_left)
  20.323 +    also from z have "\<dots> = z"
  20.324 +      by (simp only: add_zero_right)
  20.325 +    finally show "x + y = z" .
  20.326 +  next
  20.327 +    assume "x + y = z"
  20.328 +    then have "z - y = (x + y) - y" by simp
  20.329 +    also from x y have "\<dots> = x + y + - y"
  20.330 +      by (simp add: diff_eq1)
  20.331 +    also have "\<dots> = x + (y + - y)"
  20.332 +      by (rule add_assoc) (simp_all add: x y)
  20.333 +    also from x y have "\<dots> = x" by simp
  20.334 +    finally show "x = z - y" ..
  20.335 +  }
  20.336 +qed
  20.337 +
  20.338 +lemma (in vectorspace) add_minus_eq_minus:
  20.339 +  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = 0 \<Longrightarrow> x = - y"
  20.340 +proof -
  20.341 +  assume x: "x \<in> V" and y: "y \<in> V"
  20.342 +  from x y have "x = (- y + y) + x" by simp
  20.343 +  also from x y have "\<dots> = - y + (x + y)" by (simp add: add_ac)
  20.344 +  also assume "x + y = 0"
  20.345 +  also from y have "- y + 0 = - y" by simp
  20.346 +  finally show "x = - y" .
  20.347 +qed
  20.348 +
  20.349 +lemma (in vectorspace) add_minus_eq:
  20.350 +  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = 0 \<Longrightarrow> x = y"
  20.351 +proof -
  20.352 +  assume x: "x \<in> V" and y: "y \<in> V"
  20.353 +  assume "x - y = 0"
  20.354 +  with x y have eq: "x + - y = 0" by (simp add: diff_eq1)
  20.355 +  with _ _ have "x = - (- y)"
  20.356 +    by (rule add_minus_eq_minus) (simp_all add: x y)
  20.357 +  with x y show "x = y" by simp
  20.358 +qed
  20.359 +
  20.360 +lemma (in vectorspace) add_diff_swap:
  20.361 +  "a \<in> V \<Longrightarrow> b \<in> V \<Longrightarrow> c \<in> V \<Longrightarrow> d \<in> V \<Longrightarrow> a + b = c + d
  20.362 +    \<Longrightarrow> a - c = d - b"
  20.363 +proof -
  20.364 +  assume vs: "a \<in> V"  "b \<in> V"  "c \<in> V"  "d \<in> V"
  20.365 +    and eq: "a + b = c + d"
  20.366 +  then have "- c + (a + b) = - c + (c + d)"
  20.367 +    by (simp add: add_left_cancel)
  20.368 +  also have "\<dots> = d" using `c \<in> V` `d \<in> V` by (rule minus_add_cancel)
  20.369 +  finally have eq: "- c + (a + b) = d" .
  20.370 +  from vs have "a - c = (- c + (a + b)) + - b"
  20.371 +    by (simp add: add_ac diff_eq1)
  20.372 +  also from vs eq have "\<dots>  = d + - b"
  20.373 +    by (simp add: add_right_cancel)
  20.374 +  also from vs have "\<dots> = d - b" by (simp add: diff_eq2)
  20.375 +  finally show "a - c = d - b" .
  20.376 +qed
  20.377 +
  20.378 +lemma (in vectorspace) vs_add_cancel_21:
  20.379 +  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> u \<in> V
  20.380 +    \<Longrightarrow> (x + (y + z) = y + u) = (x + z = u)"
  20.381 +proof
  20.382 +  assume vs: "x \<in> V"  "y \<in> V"  "z \<in> V"  "u \<in> V"
  20.383 +  {
  20.384 +    from vs have "x + z = - y + y + (x + z)" by simp
  20.385 +    also have "\<dots> = - y + (y + (x + z))"
  20.386 +      by (rule add_assoc) (simp_all add: vs)
  20.387 +    also from vs have "y + (x + z) = x + (y + z)"
  20.388 +      by (simp add: add_ac)
  20.389 +    also assume "x + (y + z) = y + u"
  20.390 +    also from vs have "- y + (y + u) = u" by simp
  20.391 +    finally show "x + z = u" .
  20.392 +  next
  20.393 +    assume "x + z = u"
  20.394 +    with vs show "x + (y + z) = y + u"
  20.395 +      by (simp only: add_left_commute [of x])
  20.396 +  }
  20.397 +qed
  20.398 +
  20.399 +lemma (in vectorspace) add_cancel_end:
  20.400 +  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + (y + z) = y) = (x = - z)"
  20.401 +proof
  20.402 +  assume vs: "x \<in> V"  "y \<in> V"  "z \<in> V"
  20.403 +  {
  20.404 +    assume "x + (y + z) = y"
  20.405 +    with vs have "(x + z) + y = 0 + y"
  20.406 +      by (simp add: add_ac)
  20.407 +    with vs have "x + z = 0"
  20.408 +      by (simp only: add_right_cancel add_closed zero)
  20.409 +    with vs show "x = - z" by (simp add: add_minus_eq_minus)
  20.410 +  next
  20.411 +    assume eq: "x = - z"
  20.412 +    then have "x + (y + z) = - z + (y + z)" by simp
  20.413 +    also have "\<dots> = y + (- z + z)"
  20.414 +      by (rule add_left_commute) (simp_all add: vs)
  20.415 +    also from vs have "\<dots> = y"  by simp
  20.416 +    finally show "x + (y + z) = y" .
  20.417 +  }
  20.418 +qed
  20.419 +
  20.420 +end
    21.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    21.2 +++ b/src/HOL/HahnBanach/ZornLemma.thy	Mon Dec 29 14:08:08 2008 +0100
    21.3 @@ -0,0 +1,58 @@
    21.4 +(*  Title:      HOL/Real/HahnBanach/ZornLemma.thy
    21.5 +    ID:         $Id$
    21.6 +    Author:     Gertrud Bauer, TU Munich
    21.7 +*)
    21.8 +
    21.9 +header {* Zorn's Lemma *}
   21.10 +
   21.11 +theory ZornLemma
   21.12 +imports Zorn
   21.13 +begin
   21.14 +
   21.15 +text {*
   21.16 +  Zorn's Lemmas states: if every linear ordered subset of an ordered
   21.17 +  set @{text S} has an upper bound in @{text S}, then there exists a
   21.18 +  maximal element in @{text S}.  In our application, @{text S} is a
   21.19 +  set of sets ordered by set inclusion. Since the union of a chain of
   21.20 +  sets is an upper bound for all elements of the chain, the conditions
   21.21 +  of Zorn's lemma can be modified: if @{text S} is non-empty, it
   21.22 +  suffices to show that for every non-empty chain @{text c} in @{text
   21.23 +  S} the union of @{text c} also lies in @{text S}.
   21.24 +*}
   21.25 +
   21.26 +theorem Zorn's_Lemma:
   21.27 +  assumes r: "\<And>c. c \<in> chain S \<Longrightarrow> \<exists>x. x \<in> c \<Longrightarrow> \<Union>c \<in> S"
   21.28 +    and aS: "a \<in> S"
   21.29 +  shows "\<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z \<longrightarrow> y = z"
   21.30 +proof (rule Zorn_Lemma2)
   21.31 +  show "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
   21.32 +  proof
   21.33 +    fix c assume "c \<in> chain S"
   21.34 +    show "\<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
   21.35 +    proof cases
   21.36 +
   21.37 +      txt {* If @{text c} is an empty chain, then every element in
   21.38 +	@{text S} is an upper bound of @{text c}. *}
   21.39 +
   21.40 +      assume "c = {}"
   21.41 +      with aS show ?thesis by fast
   21.42 +
   21.43 +      txt {* If @{text c} is non-empty, then @{text "\<Union>c"} is an upper
   21.44 +	bound of @{text c}, lying in @{text S}. *}
   21.45 +
   21.46 +    next
   21.47 +      assume "c \<noteq> {}"
   21.48 +      show ?thesis
   21.49 +      proof
   21.50 +        show "\<forall>z \<in> c. z \<subseteq> \<Union>c" by fast
   21.51 +        show "\<Union>c \<in> S"
   21.52 +        proof (rule r)
   21.53 +          from `c \<noteq> {}` show "\<exists>x. x \<in> c" by fast
   21.54 +	  show "c \<in> chain S" by fact
   21.55 +        qed
   21.56 +      qed
   21.57 +    qed
   21.58 +  qed
   21.59 +qed
   21.60 +
   21.61 +end
    22.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    22.2 +++ b/src/HOL/HahnBanach/document/root.bib	Mon Dec 29 14:08:08 2008 +0100
    22.3 @@ -0,0 +1,27 @@
    22.4 +
    22.5 +@Book{Heuser:1986,
    22.6 +  author = 	 {H. Heuser},
    22.7 +  title = 	 {Funktionalanalysis: Theorie und Anwendung},
    22.8 +  publisher = 	 {Teubner},
    22.9 +  year = 	 1986
   22.10 +}
   22.11 +
   22.12 +@InCollection{Narici:1996,
   22.13 +  author = 	 {L. Narici and E. Beckenstein},
   22.14 +  title = 	 {The {Hahn-Banach Theorem}: The Life and Times},
   22.15 +  booktitle = 	 {Topology Atlas},
   22.16 +  publisher =	 {York University, Toronto, Ontario, Canada},
   22.17 +  year =	 1996,
   22.18 +  note =	 {\url{http://at.yorku.ca/topology/preprint.htm} and
   22.19 +                  \url{http://at.yorku.ca/p/a/a/a/16.htm}}
   22.20 +}
   22.21 +
   22.22 +@Article{Nowak:1993,
   22.23 +  author =       {B. Nowak and A. Trybulec},
   22.24 +  title =	 {{Hahn-Banach} Theorem},
   22.25 +  journal =      {Journal of Formalized Mathematics},
   22.26 +  year =         {1993},
   22.27 +  volume =       {5},
   22.28 +  institution =  {University of Bialystok},
   22.29 +  note =         {\url{http://mizar.uwb.edu.pl/JFM/Vol5/hahnban.html}}
   22.30 +}
    23.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    23.2 +++ b/src/HOL/HahnBanach/document/root.tex	Mon Dec 29 14:08:08 2008 +0100
    23.3 @@ -0,0 +1,83 @@
    23.4 +\documentclass[10pt,a4paper,twoside]{article}
    23.5 +\usepackage{graphicx}
    23.6 +\usepackage{latexsym,theorem}
    23.7 +\usepackage{isabelle,isabellesym}
    23.8 +\usepackage{pdfsetup} %last one!
    23.9 +
   23.10 +\isabellestyle{it}
   23.11 +\urlstyle{rm}
   23.12 +
   23.13 +\newcommand{\isasymsup}{\isamath{\sup\,}}
   23.14 +\newcommand{\skp}{\smallskip}
   23.15 +
   23.16 +
   23.17 +\begin{document}
   23.18 +
   23.19 +\pagestyle{headings}
   23.20 +\pagenumbering{arabic}
   23.21 +
   23.22 +\title{The Hahn-Banach Theorem \\ for Real Vector Spaces}
   23.23 +\author{Gertrud Bauer \\ \url{http://www.in.tum.de/~bauerg/}}
   23.24 +\maketitle
   23.25 +
   23.26 +\begin{abstract}
   23.27 +  The Hahn-Banach Theorem is one of the most fundamental results in functional
   23.28 +  analysis. We present a fully formal proof of two versions of the theorem,
   23.29 +  one for general linear spaces and another for normed spaces.  This
   23.30 +  development is based on simply-typed classical set-theory, as provided by
   23.31 +  Isabelle/HOL.
   23.32 +\end{abstract}
   23.33 +
   23.34 +
   23.35 +\tableofcontents
   23.36 +\parindent 0pt \parskip 0.5ex
   23.37 +
   23.38 +\clearpage
   23.39 +\section{Preface}
   23.40 +
   23.41 +This is a fully formal proof of the Hahn-Banach Theorem. It closely follows
   23.42 +the informal presentation given in Heuser's textbook \cite[{\S} 36]{Heuser:1986}.
   23.43 +Another formal proof of the same theorem has been done in Mizar
   23.44 +\cite{Nowak:1993}.  A general overview of the relevance and history of the
   23.45 +Hahn-Banach Theorem is given by Narici and Beckenstein \cite{Narici:1996}.
   23.46 +
   23.47 +\medskip The document is structured as follows.  The first part contains
   23.48 +definitions of basic notions of linear algebra: vector spaces, subspaces,
   23.49 +normed spaces, continuous linear-forms, norm of functions and an order on
   23.50 +functions by domain extension.  The second part contains some lemmas about the
   23.51 +supremum (w.r.t.\ the function order) and extension of non-maximal functions.
   23.52 +With these preliminaries, the main proof of the theorem (in its two versions)
   23.53 +is conducted in the third part.  The dependencies of individual theories are
   23.54 +as follows.
   23.55 +
   23.56 +\begin{center}
   23.57 +  \includegraphics[scale=0.5]{session_graph}  
   23.58 +\end{center}
   23.59 +
   23.60 +\clearpage
   23.61 +\part {Basic Notions}
   23.62 +
   23.63 +\input{Bounds}
   23.64 +\input{VectorSpace}
   23.65 +\input{Subspace}
   23.66 +\input{NormedSpace}
   23.67 +\input{Linearform}
   23.68 +\input{FunctionOrder}
   23.69 +\input{FunctionNorm}
   23.70 +\input{ZornLemma}
   23.71 +
   23.72 +\clearpage
   23.73 +\part {Lemmas for the Proof}
   23.74 +
   23.75 +\input{HahnBanachSupLemmas}
   23.76 +\input{HahnBanachExtLemmas}
   23.77 +\input{HahnBanachLemmas}
   23.78 +
   23.79 +\clearpage
   23.80 +\part {The Main Proof}
   23.81 +
   23.82 +\input{HahnBanach}
   23.83 +\bibliographystyle{abbrv}
   23.84 +\bibliography{root}
   23.85 +
   23.86 +\end{document}
    24.1 --- a/src/HOL/Hyperreal/SEQ.thy	Mon Dec 29 13:23:53 2008 +0100
    24.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    24.3 @@ -1,1136 +0,0 @@
    24.4 -(*  Title       : SEQ.thy
    24.5 -    Author      : Jacques D. Fleuriot
    24.6 -    Copyright   : 1998  University of Cambridge
    24.7 -    Description : Convergence of sequences and series
    24.8 -    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
    24.9 -    Additional contributions by Jeremy Avigad and Brian Huffman
   24.10 -*)
   24.11 -
   24.12 -header {* Sequences and Convergence *}
   24.13 -
   24.14 -theory SEQ
   24.15 -imports "../Real/RealVector" "../RComplete"
   24.16 -begin
   24.17 -
   24.18 -definition
   24.19 -  Zseq :: "[nat \<Rightarrow> 'a::real_normed_vector] \<Rightarrow> bool" where
   24.20 -    --{*Standard definition of sequence converging to zero*}
   24.21 -  [code del]: "Zseq X = (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. norm (X n) < r)"
   24.22 -
   24.23 -definition
   24.24 -  LIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool"
   24.25 -    ("((_)/ ----> (_))" [60, 60] 60) where
   24.26 -    --{*Standard definition of convergence of sequence*}
   24.27 -  [code del]: "X ----> L = (\<forall>r. 0 < r --> (\<exists>no. \<forall>n. no \<le> n --> norm (X n - L) < r))"
   24.28 -
   24.29 -definition
   24.30 -  lim :: "(nat => 'a::real_normed_vector) => 'a" where
   24.31 -    --{*Standard definition of limit using choice operator*}
   24.32 -  "lim X = (THE L. X ----> L)"
   24.33 -
   24.34 -definition
   24.35 -  convergent :: "(nat => 'a::real_normed_vector) => bool" where
   24.36 -    --{*Standard definition of convergence*}
   24.37 -  "convergent X = (\<exists>L. X ----> L)"
   24.38 -
   24.39 -definition
   24.40 -  Bseq :: "(nat => 'a::real_normed_vector) => bool" where
   24.41 -    --{*Standard definition for bounded sequence*}
   24.42 -  [code del]: "Bseq X = (\<exists>K>0.\<forall>n. norm (X n) \<le> K)"
   24.43 -
   24.44 -definition
   24.45 -  monoseq :: "(nat=>real)=>bool" where
   24.46 -    --{*Definition for monotonicity*}
   24.47 -  [code del]: "monoseq X = ((\<forall>m. \<forall>n\<ge>m. X m \<le> X n) | (\<forall>m. \<forall>n\<ge>m. X n \<le> X m))"
   24.48 -
   24.49 -definition
   24.50 -  subseq :: "(nat => nat) => bool" where
   24.51 -    --{*Definition of subsequence*}
   24.52 -  [code del]:   "subseq f = (\<forall>m. \<forall>n>m. (f m) < (f n))"
   24.53 -
   24.54 -definition
   24.55 -  Cauchy :: "(nat => 'a::real_normed_vector) => bool" where
   24.56 -    --{*Standard definition of the Cauchy condition*}
   24.57 -  [code del]: "Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. norm (X m - X n) < e)"
   24.58 -
   24.59 -
   24.60 -subsection {* Bounded Sequences *}
   24.61 -
   24.62 -lemma BseqI': assumes K: "\<And>n. norm (X n) \<le> K" shows "Bseq X"
   24.63 -unfolding Bseq_def
   24.64 -proof (intro exI conjI allI)
   24.65 -  show "0 < max K 1" by simp
   24.66 -next
   24.67 -  fix n::nat
   24.68 -  have "norm (X n) \<le> K" by (rule K)
   24.69 -  thus "norm (X n) \<le> max K 1" by simp
   24.70 -qed
   24.71 -
   24.72 -lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
   24.73 -unfolding Bseq_def by auto
   24.74 -
   24.75 -lemma BseqI2': assumes K: "\<forall>n\<ge>N. norm (X n) \<le> K" shows "Bseq X"
   24.76 -proof (rule BseqI')
   24.77 -  let ?A = "norm ` X ` {..N}"
   24.78 -  have 1: "finite ?A" by simp
   24.79 -  fix n::nat
   24.80 -  show "norm (X n) \<le> max K (Max ?A)"
   24.81 -  proof (cases rule: linorder_le_cases)
   24.82 -    assume "n \<ge> N"
   24.83 -    hence "norm (X n) \<le> K" using K by simp
   24.84 -    thus "norm (X n) \<le> max K (Max ?A)" by simp
   24.85 -  next
   24.86 -    assume "n \<le> N"
   24.87 -    hence "norm (X n) \<in> ?A" by simp
   24.88 -    with 1 have "norm (X n) \<le> Max ?A" by (rule Max_ge)
   24.89 -    thus "norm (X n) \<le> max K (Max ?A)" by simp
   24.90 -  qed
   24.91 -qed
   24.92 -
   24.93 -lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))"
   24.94 -unfolding Bseq_def by auto
   24.95 -
   24.96 -lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X"
   24.97 -apply (erule BseqE)
   24.98 -apply (rule_tac N="k" and K="K" in BseqI2')
   24.99 -apply clarify
  24.100 -apply (drule_tac x="n - k" in spec, simp)
  24.101 -done
  24.102 -
  24.103 -
  24.104 -subsection {* Sequences That Converge to Zero *}
  24.105 -
  24.106 -lemma ZseqI:
  24.107 -  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r) \<Longrightarrow> Zseq X"
  24.108 -unfolding Zseq_def by simp
  24.109 -
  24.110 -lemma ZseqD:
  24.111 -  "\<lbrakk>Zseq X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r"
  24.112 -unfolding Zseq_def by simp
  24.113 -
  24.114 -lemma Zseq_zero: "Zseq (\<lambda>n. 0)"
  24.115 -unfolding Zseq_def by simp
  24.116 -
  24.117 -lemma Zseq_const_iff: "Zseq (\<lambda>n. k) = (k = 0)"
  24.118 -unfolding Zseq_def by force
  24.119 -
  24.120 -lemma Zseq_norm_iff: "Zseq (\<lambda>n. norm (X n)) = Zseq (\<lambda>n. X n)"
  24.121 -unfolding Zseq_def by simp
  24.122 -
  24.123 -lemma Zseq_imp_Zseq:
  24.124 -  assumes X: "Zseq X"
  24.125 -  assumes Y: "\<And>n. norm (Y n) \<le> norm (X n) * K"
  24.126 -  shows "Zseq (\<lambda>n. Y n)"
  24.127 -proof (cases)
  24.128 -  assume K: "0 < K"
  24.129 -  show ?thesis
  24.130 -  proof (rule ZseqI)
  24.131 -    fix r::real assume "0 < r"
  24.132 -    hence "0 < r / K"
  24.133 -      using K by (rule divide_pos_pos)
  24.134 -    then obtain N where "\<forall>n\<ge>N. norm (X n) < r / K"
  24.135 -      using ZseqD [OF X] by fast
  24.136 -    hence "\<forall>n\<ge>N. norm (X n) * K < r"
  24.137 -      by (simp add: pos_less_divide_eq K)
  24.138 -    hence "\<forall>n\<ge>N. norm (Y n) < r"
  24.139 -      by (simp add: order_le_less_trans [OF Y])
  24.140 -    thus "\<exists>N. \<forall>n\<ge>N. norm (Y n) < r" ..
  24.141 -  qed
  24.142 -next
  24.143 -  assume "\<not> 0 < K"
  24.144 -  hence K: "K \<le> 0" by (simp only: linorder_not_less)
  24.145 -  {
  24.146 -    fix n::nat
  24.147 -    have "norm (Y n) \<le> norm (X n) * K" by (rule Y)
  24.148 -    also have "\<dots> \<le> norm (X n) * 0"
  24.149 -      using K norm_ge_zero by (rule mult_left_mono)
  24.150 -    finally have "norm (Y n) = 0" by simp
  24.151 -  }
  24.152 -  thus ?thesis by (simp add: Zseq_zero)
  24.153 -qed
  24.154 -
  24.155 -lemma Zseq_le: "\<lbrakk>Zseq Y; \<forall>n. norm (X n) \<le> norm (Y n)\<rbrakk> \<Longrightarrow> Zseq X"
  24.156 -by (erule_tac K="1" in Zseq_imp_Zseq, simp)
  24.157 -
  24.158 -lemma Zseq_add:
  24.159 -  assumes X: "Zseq X"
  24.160 -  assumes Y: "Zseq Y"
  24.161 -  shows "Zseq (\<lambda>n. X n + Y n)"
  24.162 -proof (rule ZseqI)
  24.163 -  fix r::real assume "0 < r"
  24.164 -  hence r: "0 < r / 2" by simp
  24.165 -  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r/2"
  24.166 -    using ZseqD [OF X r] by fast
  24.167 -  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < r/2"
  24.168 -    using ZseqD [OF Y r] by fast
  24.169 -  show "\<exists>N. \<forall>n\<ge>N. norm (X n + Y n) < r"
  24.170 -  proof (intro exI allI impI)
  24.171 -    fix n assume n: "max M N \<le> n"
  24.172 -    have "norm (X n + Y n) \<le> norm (X n) + norm (Y n)"
  24.173 -      by (rule norm_triangle_ineq)
  24.174 -    also have "\<dots> < r/2 + r/2"
  24.175 -    proof (rule add_strict_mono)
  24.176 -      from M n show "norm (X n) < r/2" by simp
  24.177 -      from N n show "norm (Y n) < r/2" by simp
  24.178 -    qed
  24.179 -    finally show "norm (X n + Y n) < r" by simp
  24.180 -  qed
  24.181 -qed
  24.182 -
  24.183 -lemma Zseq_minus: "Zseq X \<Longrightarrow> Zseq (\<lambda>n. - X n)"
  24.184 -unfolding Zseq_def by simp
  24.185 -
  24.186 -lemma Zseq_diff: "\<lbrakk>Zseq X; Zseq Y\<rbrakk> \<Longrightarrow> Zseq (\<lambda>n. X n - Y n)"
  24.187 -by (simp only: diff_minus Zseq_add Zseq_minus)
  24.188 -
  24.189 -lemma (in bounded_linear) Zseq:
  24.190 -  assumes X: "Zseq X"
  24.191 -  shows "Zseq (\<lambda>n. f (X n))"
  24.192 -proof -
  24.193 -  obtain K where "\<And>x. norm (f x) \<le> norm x * K"
  24.194 -    using bounded by fast
  24.195 -  with X show ?thesis
  24.196 -    by (rule Zseq_imp_Zseq)
  24.197 -qed
  24.198 -
  24.199 -lemma (in bounded_bilinear) Zseq:
  24.200 -  assumes X: "Zseq X"
  24.201 -  assumes Y: "Zseq Y"
  24.202 -  shows "Zseq (\<lambda>n. X n ** Y n)"
  24.203 -proof (rule ZseqI)
  24.204 -  fix r::real assume r: "0 < r"
  24.205 -  obtain K where K: "0 < K"
  24.206 -    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
  24.207 -    using pos_bounded by fast
  24.208 -  from K have K': "0 < inverse K"
  24.209 -    by (rule positive_imp_inverse_positive)
  24.210 -  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r"
  24.211 -    using ZseqD [OF X r] by fast
  24.212 -  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < inverse K"
  24.213 -    using ZseqD [OF Y K'] by fast
  24.214 -  show "\<exists>N. \<forall>n\<ge>N. norm (X n ** Y n) < r"
  24.215 -  proof (intro exI allI impI)
  24.216 -    fix n assume n: "max M N \<le> n"
  24.217 -    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
  24.218 -      by (rule norm_le)
  24.219 -    also have "norm (X n) * norm (Y n) * K < r * inverse K * K"
  24.220 -    proof (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero K)
  24.221 -      from M n show Xn: "norm (X n) < r" by simp
  24.222 -      from N n show Yn: "norm (Y n) < inverse K" by simp
  24.223 -    qed
  24.224 -    also from K have "r * inverse K * K = r" by simp
  24.225 -    finally show "norm (X n ** Y n) < r" .
  24.226 -  qed
  24.227 -qed
  24.228 -
  24.229 -lemma (in bounded_bilinear) Zseq_prod_Bseq:
  24.230 -  assumes X: "Zseq X"
  24.231 -  assumes Y: "Bseq Y"
  24.232 -  shows "Zseq (\<lambda>n. X n ** Y n)"
  24.233 -proof -
  24.234 -  obtain K where K: "0 \<le> K"
  24.235 -    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
  24.236 -    using nonneg_bounded by fast
  24.237 -  obtain B where B: "0 < B"
  24.238 -    and norm_Y: "\<And>n. norm (Y n) \<le> B"
  24.239 -    using Y [unfolded Bseq_def] by fast
  24.240 -  from X show ?thesis
  24.241 -  proof (rule Zseq_imp_Zseq)
  24.242 -    fix n::nat
  24.243 -    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
  24.244 -      by (rule norm_le)
  24.245 -    also have "\<dots> \<le> norm (X n) * B * K"
  24.246 -      by (intro mult_mono' order_refl norm_Y norm_ge_zero
  24.247 -                mult_nonneg_nonneg K)
  24.248 -    also have "\<dots> = norm (X n) * (B * K)"
  24.249 -      by (rule mult_assoc)
  24.250 -    finally show "norm (X n ** Y n) \<le> norm (X n) * (B * K)" .
  24.251 -  qed
  24.252 -qed
  24.253 -
  24.254 -lemma (in bounded_bilinear) Bseq_prod_Zseq:
  24.255 -  assumes X: "Bseq X"
  24.256 -  assumes Y: "Zseq Y"
  24.257 -  shows "Zseq (\<lambda>n. X n ** Y n)"
  24.258 -proof -
  24.259 -  obtain K where K: "0 \<le> K"
  24.260 -    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
  24.261 -    using nonneg_bounded by fast
  24.262 -  obtain B where B: "0 < B"
  24.263 -    and norm_X: "\<And>n. norm (X n) \<le> B"
  24.264 -    using X [unfolded Bseq_def] by fast
  24.265 -  from Y show ?thesis
  24.266 -  proof (rule Zseq_imp_Zseq)
  24.267 -    fix n::nat
  24.268 -    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
  24.269 -      by (rule norm_le)
  24.270 -    also have "\<dots> \<le> B * norm (Y n) * K"
  24.271 -      by (intro mult_mono' order_refl norm_X norm_ge_zero
  24.272 -                mult_nonneg_nonneg K)
  24.273 -    also have "\<dots> = norm (Y n) * (B * K)"
  24.274 -      by (simp only: mult_ac)
  24.275 -    finally show "norm (X n ** Y n) \<le> norm (Y n) * (B * K)" .
  24.276 -  qed
  24.277 -qed
  24.278 -
  24.279 -lemma (in bounded_bilinear) Zseq_left:
  24.280 -  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. X n ** a)"
  24.281 -by (rule bounded_linear_left [THEN bounded_linear.Zseq])
  24.282 -
  24.283 -lemma (in bounded_bilinear) Zseq_right:
  24.284 -  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. a ** X n)"
  24.285 -by (rule bounded_linear_right [THEN bounded_linear.Zseq])
  24.286 -
  24.287 -lemmas Zseq_mult = mult.Zseq
  24.288 -lemmas Zseq_mult_right = mult.Zseq_right
  24.289 -lemmas Zseq_mult_left = mult.Zseq_left
  24.290 -
  24.291 -
  24.292 -subsection {* Limits of Sequences *}
  24.293 -
  24.294 -lemma LIMSEQ_iff:
  24.295 -      "(X ----> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)"
  24.296 -by (rule LIMSEQ_def)
  24.297 -
  24.298 -lemma LIMSEQ_Zseq_iff: "((\<lambda>n. X n) ----> L) = Zseq (\<lambda>n. X n - L)"
  24.299 -by (simp only: LIMSEQ_def Zseq_def)
  24.300 -
  24.301 -lemma LIMSEQ_I:
  24.302 -  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X ----> L"
  24.303 -by (simp add: LIMSEQ_def)
  24.304 -
  24.305 -lemma LIMSEQ_D:
  24.306 -  "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
  24.307 -by (simp add: LIMSEQ_def)
  24.308 -
  24.309 -lemma LIMSEQ_const: "(\<lambda>n. k) ----> k"
  24.310 -by (simp add: LIMSEQ_def)
  24.311 -
  24.312 -lemma LIMSEQ_const_iff: "(\<lambda>n. k) ----> l = (k = l)"
  24.313 -by (simp add: LIMSEQ_Zseq_iff Zseq_const_iff)
  24.314 -
  24.315 -lemma LIMSEQ_norm: "X ----> a \<Longrightarrow> (\<lambda>n. norm (X n)) ----> norm a"
  24.316 -apply (simp add: LIMSEQ_def, safe)
  24.317 -apply (drule_tac x="r" in spec, safe)
  24.318 -apply (rule_tac x="no" in exI, safe)
  24.319 -apply (drule_tac x="n" in spec, safe)
  24.320 -apply (erule order_le_less_trans [OF norm_triangle_ineq3])
  24.321 -done
  24.322 -
  24.323 -lemma LIMSEQ_ignore_initial_segment:
  24.324 -  "f ----> a \<Longrightarrow> (\<lambda>n. f (n + k)) ----> a"
  24.325 -apply (rule LIMSEQ_I)
  24.326 -apply (drule (1) LIMSEQ_D)
  24.327 -apply (erule exE, rename_tac N)
  24.328 -apply (rule_tac x=N in exI)
  24.329 -apply simp
  24.330 -done
  24.331 -
  24.332 -lemma LIMSEQ_offset:
  24.333 -  "(\<lambda>n. f (n + k)) ----> a \<Longrightarrow> f ----> a"
  24.334 -apply (rule LIMSEQ_I)
  24.335 -apply (drule (1) LIMSEQ_D)
  24.336 -apply (erule exE, rename_tac N)
  24.337 -apply (rule_tac x="N + k" in exI)
  24.338 -apply clarify
  24.339 -apply (drule_tac x="n - k" in spec)
  24.340 -apply (simp add: le_diff_conv2)
  24.341 -done
  24.342 -
  24.343 -lemma LIMSEQ_Suc: "f ----> l \<Longrightarrow> (\<lambda>n. f (Suc n)) ----> l"
  24.344 -by (drule_tac k="1" in LIMSEQ_ignore_initial_segment, simp)
  24.345 -
  24.346 -lemma LIMSEQ_imp_Suc: "(\<lambda>n. f (Suc n)) ----> l \<Longrightarrow> f ----> l"
  24.347 -by (rule_tac k="1" in LIMSEQ_offset, simp)
  24.348 -
  24.349 -lemma LIMSEQ_Suc_iff: "(\<lambda>n. f (Suc n)) ----> l = f ----> l"
  24.350 -by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc)
  24.351 -
  24.352 -lemma add_diff_add:
  24.353 -  fixes a b c d :: "'a::ab_group_add"
  24.354 -  shows "(a + c) - (b + d) = (a - b) + (c - d)"
  24.355 -by simp
  24.356 -
  24.357 -lemma minus_diff_minus:
  24.358 -  fixes a b :: "'a::ab_group_add"
  24.359 -  shows "(- a) - (- b) = - (a - b)"
  24.360 -by simp
  24.361 -
  24.362 -lemma LIMSEQ_add: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) ----> a + b"
  24.363 -by (simp only: LIMSEQ_Zseq_iff add_diff_add Zseq_add)
  24.364 -
  24.365 -lemma LIMSEQ_minus: "X ----> a \<Longrightarrow> (\<lambda>n. - X n) ----> - a"
  24.366 -by (simp only: LIMSEQ_Zseq_iff minus_diff_minus Zseq_minus)
  24.367 -
  24.368 -lemma LIMSEQ_minus_cancel: "(\<lambda>n. - X n) ----> - a \<Longrightarrow> X ----> a"
  24.369 -by (drule LIMSEQ_minus, simp)
  24.370 -
  24.371 -lemma LIMSEQ_diff: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) ----> a - b"
  24.372 -by (simp add: diff_minus LIMSEQ_add LIMSEQ_minus)
  24.373 -
  24.374 -lemma LIMSEQ_unique: "\<lbrakk>X ----> a; X ----> b\<rbrakk> \<Longrightarrow> a = b"
  24.375 -by (drule (1) LIMSEQ_diff, simp add: LIMSEQ_const_iff)
  24.376 -
  24.377 -lemma (in bounded_linear) LIMSEQ:
  24.378 -  "X ----> a \<Longrightarrow> (\<lambda>n. f (X n)) ----> f a"
  24.379 -by (simp only: LIMSEQ_Zseq_iff diff [symmetric] Zseq)
  24.380 -
  24.381 -lemma (in bounded_bilinear) LIMSEQ:
  24.382 -  "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n ** Y n) ----> a ** b"
  24.383 -by (simp only: LIMSEQ_Zseq_iff prod_diff_prod
  24.384 -               Zseq_add Zseq Zseq_left Zseq_right)
  24.385 -
  24.386 -lemma LIMSEQ_mult:
  24.387 -  fixes a b :: "'a::real_normed_algebra"
  24.388 -  shows "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b"
  24.389 -by (rule mult.LIMSEQ)
  24.390 -
  24.391 -lemma inverse_diff_inverse:
  24.392 -  "\<lbrakk>(a::'a::division_ring) \<noteq> 0; b \<noteq> 0\<rbrakk>
  24.393 -   \<Longrightarrow> inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
  24.394 -by (simp add: ring_simps)
  24.395 -
  24.396 -lemma Bseq_inverse_lemma:
  24.397 -  fixes x :: "'a::real_normed_div_algebra"
  24.398 -  shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
  24.399 -apply (subst nonzero_norm_inverse, clarsimp)
  24.400 -apply (erule (1) le_imp_inverse_le)
  24.401 -done
  24.402 -
  24.403 -lemma Bseq_inverse:
  24.404 -  fixes a :: "'a::real_normed_div_algebra"
  24.405 -  assumes X: "X ----> a"
  24.406 -  assumes a: "a \<noteq> 0"
  24.407 -  shows "Bseq (\<lambda>n. inverse (X n))"
  24.408 -proof -
  24.409 -  from a have "0 < norm a" by simp
  24.410 -  hence "\<exists>r>0. r < norm a" by (rule dense)
  24.411 -  then obtain r where r1: "0 < r" and r2: "r < norm a" by fast
  24.412 -  obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> norm (X n - a) < r"
  24.413 -    using LIMSEQ_D [OF X r1] by fast
  24.414 -  show ?thesis
  24.415 -  proof (rule BseqI2' [rule_format])
  24.416 -    fix n assume n: "N \<le> n"
  24.417 -    hence 1: "norm (X n - a) < r" by (rule N)
  24.418 -    hence 2: "X n \<noteq> 0" using r2 by auto
  24.419 -    hence "norm (inverse (X n)) = inverse (norm (X n))"
  24.420 -      by (rule nonzero_norm_inverse)
  24.421 -    also have "\<dots> \<le> inverse (norm a - r)"
  24.422 -    proof (rule le_imp_inverse_le)
  24.423 -      show "0 < norm a - r" using r2 by simp
  24.424 -    next
  24.425 -      have "norm a - norm (X n) \<le> norm (a - X n)"
  24.426 -        by (rule norm_triangle_ineq2)
  24.427 -      also have "\<dots> = norm (X n - a)"
  24.428 -        by (rule norm_minus_commute)
  24.429 -      also have "\<dots> < r" using 1 .
  24.430 -      finally show "norm a - r \<le> norm (X n)" by simp
  24.431 -    qed
  24.432 -    finally show "norm (inverse (X n)) \<le> inverse (norm a - r)" .
  24.433 -  qed
  24.434 -qed
  24.435 -
  24.436 -lemma LIMSEQ_inverse_lemma:
  24.437 -  fixes a :: "'a::real_normed_div_algebra"
  24.438 -  shows "\<lbrakk>X ----> a; a \<noteq> 0; \<forall>n. X n \<noteq> 0\<rbrakk>
  24.439 -         \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> inverse a"
  24.440 -apply (subst LIMSEQ_Zseq_iff)
  24.441 -apply (simp add: inverse_diff_inverse nonzero_imp_inverse_nonzero)
  24.442 -apply (rule Zseq_minus)
  24.443 -apply (rule Zseq_mult_left)
  24.444 -apply (rule mult.Bseq_prod_Zseq)
  24.445 -apply (erule (1) Bseq_inverse)
  24.446 -apply (simp add: LIMSEQ_Zseq_iff)
  24.447 -done
  24.448 -
  24.449 -lemma LIMSEQ_inverse:
  24.450 -  fixes a :: "'a::real_normed_div_algebra"
  24.451 -  assumes X: "X ----> a"
  24.452 -  assumes a: "a \<noteq> 0"
  24.453 -  shows "(\<lambda>n. inverse (X n)) ----> inverse a"
  24.454 -proof -
  24.455 -  from a have "0 < norm a" by simp
  24.456 -  then obtain k where "\<forall>n\<ge>k. norm (X n - a) < norm a"
  24.457 -    using LIMSEQ_D [OF X] by fast
  24.458 -  hence "\<forall>n\<ge>k. X n \<noteq> 0" by auto
  24.459 -  hence k: "\<forall>n. X (n + k) \<noteq> 0" by simp
  24.460 -
  24.461 -  from X have "(\<lambda>n. X (n + k)) ----> a"
  24.462 -    by (rule LIMSEQ_ignore_initial_segment)
  24.463 -  hence "(\<lambda>n. inverse (X (n + k))) ----> inverse a"
  24.464 -    using a k by (rule LIMSEQ_inverse_lemma)
  24.465 -  thus "(\<lambda>n. inverse (X n)) ----> inverse a"
  24.466 -    by (rule LIMSEQ_offset)
  24.467 -qed
  24.468 -
  24.469 -lemma LIMSEQ_divide:
  24.470 -  fixes a b :: "'a::real_normed_field"
  24.471 -  shows "\<lbrakk>X ----> a; Y ----> b; b \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>n. X n / Y n) ----> a / b"
  24.472 -by (simp add: LIMSEQ_mult LIMSEQ_inverse divide_inverse)
  24.473 -
  24.474 -lemma LIMSEQ_pow:
  24.475 -  fixes a :: "'a::{real_normed_algebra,recpower}"
  24.476 -  shows "X ----> a \<Longrightarrow> (\<lambda>n. (X n) ^ m) ----> a ^ m"
  24.477 -by (induct m) (simp_all add: power_Suc LIMSEQ_const LIMSEQ_mult)
  24.478 -
  24.479 -lemma LIMSEQ_setsum:
  24.480 -  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
  24.481 -  shows "(\<lambda>m. \<Sum>n\<in>S. X n m) ----> (\<Sum>n\<in>S. L n)"
  24.482 -proof (cases "finite S")
  24.483 -  case True
  24.484 -  thus ?thesis using n
  24.485 -  proof (induct)
  24.486 -    case empty
  24.487 -    show ?case
  24.488 -      by (simp add: LIMSEQ_const)
  24.489 -  next
  24.490 -    case insert
  24.491 -    thus ?case
  24.492 -      by (simp add: LIMSEQ_add)
  24.493 -  qed
  24.494 -next
  24.495 -  case False
  24.496 -  thus ?thesis
  24.497 -    by (simp add: LIMSEQ_const)
  24.498 -qed
  24.499 -
  24.500 -lemma LIMSEQ_setprod:
  24.501 -  fixes L :: "'a \<Rightarrow> 'b::{real_normed_algebra,comm_ring_1}"
  24.502 -  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
  24.503 -  shows "(\<lambda>m. \<Prod>n\<in>S. X n m) ----> (\<Prod>n\<in>S. L n)"
  24.504 -proof (cases "finite S")
  24.505 -  case True
  24.506 -  thus ?thesis using n
  24.507 -  proof (induct)
  24.508 -    case empty
  24.509 -    show ?case
  24.510 -      by (simp add: LIMSEQ_const)
  24.511 -  next
  24.512 -    case insert
  24.513 -    thus ?case
  24.514 -      by (simp add: LIMSEQ_mult)
  24.515 -  qed
  24.516 -next
  24.517 -  case False
  24.518 -  thus ?thesis
  24.519 -    by (simp add: setprod_def LIMSEQ_const)
  24.520 -qed
  24.521 -
  24.522 -lemma LIMSEQ_add_const: "f ----> a ==> (%n.(f n + b)) ----> a + b"
  24.523 -by (simp add: LIMSEQ_add LIMSEQ_const)
  24.524 -
  24.525 -(* FIXME: delete *)
  24.526 -lemma LIMSEQ_add_minus:
  24.527 -     "[| X ----> a; Y ----> b |] ==> (%n. X n + -Y n) ----> a + -b"
  24.528 -by (simp only: LIMSEQ_add LIMSEQ_minus)
  24.529 -
  24.530 -lemma LIMSEQ_diff_const: "f ----> a ==> (%n.(f n  - b)) ----> a - b"
  24.531 -by (simp add: LIMSEQ_diff LIMSEQ_const)
  24.532 -
  24.533 -lemma LIMSEQ_diff_approach_zero: 
  24.534 -  "g ----> L ==> (%x. f x - g x) ----> 0  ==>
  24.535 -     f ----> L"
  24.536 -  apply (drule LIMSEQ_add)
  24.537 -  apply assumption
  24.538 -  apply simp
  24.539 -done
  24.540 -
  24.541 -lemma LIMSEQ_diff_approach_zero2: 
  24.542 -  "f ----> L ==> (%x. f x - g x) ----> 0  ==>
  24.543 -     g ----> L";
  24.544 -  apply (drule LIMSEQ_diff)
  24.545 -  apply assumption
  24.546 -  apply simp
  24.547 -done
  24.548 -
  24.549 -text{*A sequence tends to zero iff its abs does*}
  24.550 -lemma LIMSEQ_norm_zero: "((\<lambda>n. norm (X n)) ----> 0) = (X ----> 0)"
  24.551 -by (simp add: LIMSEQ_def)
  24.552 -
  24.553 -lemma LIMSEQ_rabs_zero: "((%n. \<bar>f n\<bar>) ----> 0) = (f ----> (0::real))"
  24.554 -by (simp add: LIMSEQ_def)
  24.555 -
  24.556 -lemma LIMSEQ_imp_rabs: "f ----> (l::real) ==> (%n. \<bar>f n\<bar>) ----> \<bar>l\<bar>"
  24.557 -by (drule LIMSEQ_norm, simp)
  24.558 -
  24.559 -text{*An unbounded sequence's inverse tends to 0*}
  24.560 -
  24.561 -lemma LIMSEQ_inverse_zero:
  24.562 -  "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
  24.563 -apply (rule LIMSEQ_I)
  24.564 -apply (drule_tac x="inverse r" in spec, safe)
  24.565 -apply (rule_tac x="N" in exI, safe)
  24.566 -apply (drule_tac x="n" in spec, safe)
  24.567 -apply (frule positive_imp_inverse_positive)
  24.568 -apply (frule (1) less_imp_inverse_less)
  24.569 -apply (subgoal_tac "0 < X n", simp)
  24.570 -apply (erule (1) order_less_trans)
  24.571 -done
  24.572 -
  24.573 -text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}
  24.574 -
  24.575 -lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
  24.576 -apply (rule LIMSEQ_inverse_zero, safe)
  24.577 -apply (cut_tac x = r in reals_Archimedean2)
  24.578 -apply (safe, rule_tac x = n in exI)
  24.579 -apply (auto simp add: real_of_nat_Suc)
  24.580 -done
  24.581 -
  24.582 -text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
  24.583 -infinity is now easily proved*}
  24.584 -
  24.585 -lemma LIMSEQ_inverse_real_of_nat_add:
  24.586 -     "(%n. r + inverse(real(Suc n))) ----> r"
  24.587 -by (cut_tac LIMSEQ_add [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
  24.588 -
  24.589 -lemma LIMSEQ_inverse_real_of_nat_add_minus:
  24.590 -     "(%n. r + -inverse(real(Suc n))) ----> r"
  24.591 -by (cut_tac LIMSEQ_add_minus [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
  24.592 -
  24.593 -lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
  24.594 -     "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
  24.595 -by (cut_tac b=1 in
  24.596 -        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat_add_minus], auto)
  24.597 -
  24.598 -lemma LIMSEQ_le_const:
  24.599 -  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. a \<le> X n\<rbrakk> \<Longrightarrow> a \<le> x"
  24.600 -apply (rule ccontr, simp only: linorder_not_le)
  24.601 -apply (drule_tac r="a - x" in LIMSEQ_D, simp)
  24.602 -apply clarsimp
  24.603 -apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI1)
  24.604 -apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI2)
  24.605 -apply simp
  24.606 -done
  24.607 -
  24.608 -lemma LIMSEQ_le_const2:
  24.609 -  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. X n \<le> a\<rbrakk> \<Longrightarrow> x \<le> a"
  24.610 -apply (subgoal_tac "- a \<le> - x", simp)
  24.611 -apply (rule LIMSEQ_le_const)
  24.612 -apply (erule LIMSEQ_minus)
  24.613 -apply simp
  24.614 -done
  24.615 -
  24.616 -lemma LIMSEQ_le:
  24.617 -  "\<lbrakk>X ----> x; Y ----> y; \<exists>N. \<forall>n\<ge>N. X n \<le> Y n\<rbrakk> \<Longrightarrow> x \<le> (y::real)"
  24.618 -apply (subgoal_tac "0 \<le> y - x", simp)
  24.619 -apply (rule LIMSEQ_le_const)
  24.620 -apply (erule (1) LIMSEQ_diff)
  24.621 -apply (simp add: le_diff_eq)
  24.622 -done
  24.623 -
  24.624 -
  24.625 -subsection {* Convergence *}
  24.626 -
  24.627 -lemma limI: "X ----> L ==> lim X = L"
  24.628 -apply (simp add: lim_def)
  24.629 -apply (blast intro: LIMSEQ_unique)
  24.630 -done
  24.631 -
  24.632 -lemma convergentD: "convergent X ==> \<exists>L. (X ----> L)"
  24.633 -by (simp add: convergent_def)
  24.634 -
  24.635 -lemma convergentI: "(X ----> L) ==> convergent X"
  24.636 -by (auto simp add: convergent_def)
  24.637 -
  24.638 -lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)"
  24.639 -by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def)
  24.640 -
  24.641 -lemma convergent_minus_iff: "(convergent X) = (convergent (%n. -(X n)))"
  24.642 -apply (simp add: convergent_def)
  24.643 -apply (auto dest: LIMSEQ_minus)
  24.644 -apply (drule LIMSEQ_minus, auto)
  24.645 -done
  24.646 -
  24.647 -
  24.648 -subsection {* Bounded Monotonic Sequences *}
  24.649 -
  24.650 -text{*Subsequence (alternative definition, (e.g. Hoskins)*}
  24.651 -
  24.652 -lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))"
  24.653 -apply (simp add: subseq_def)
  24.654 -apply (auto dest!: less_imp_Suc_add)
  24.655 -apply (induct_tac k)
  24.656 -apply (auto intro: less_trans)
  24.657 -done
  24.658 -
  24.659 -lemma monoseq_Suc:
  24.660 -   "monoseq X = ((\<forall>n. X n \<le> X (Suc n))
  24.661 -                 | (\<forall>n. X (Suc n) \<le> X n))"
  24.662 -apply (simp add: monoseq_def)
  24.663 -apply (auto dest!: le_imp_less_or_eq)
  24.664 -apply (auto intro!: lessI [THEN less_imp_le] dest!: less_imp_Suc_add)
  24.665 -apply (induct_tac "ka")
  24.666 -apply (auto intro: order_trans)
  24.667 -apply (erule contrapos_np)
  24.668 -apply (induct_tac "k")
  24.669 -apply (auto intro: order_trans)
  24.670 -done
  24.671 -
  24.672 -lemma monoI1: "\<forall>m. \<forall> n \<ge> m. X m \<le> X n ==> monoseq X"
  24.673 -by (simp add: monoseq_def)
  24.674 -
  24.675 -lemma monoI2: "\<forall>m. \<forall> n \<ge> m. X n \<le> X m ==> monoseq X"
  24.676 -by (simp add: monoseq_def)
  24.677 -
  24.678 -lemma mono_SucI1: "\<forall>n. X n \<le> X (Suc n) ==> monoseq X"
  24.679 -by (simp add: monoseq_Suc)
  24.680 -
  24.681 -lemma mono_SucI2: "\<forall>n. X (Suc n) \<le> X n ==> monoseq X"
  24.682 -by (simp add: monoseq_Suc)
  24.683 -
  24.684 -text{*Bounded Sequence*}
  24.685 -
  24.686 -lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)"
  24.687 -by (simp add: Bseq_def)
  24.688 -
  24.689 -lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
  24.690 -by (auto simp add: Bseq_def)
  24.691 -
  24.692 -lemma lemma_NBseq_def:
  24.693 -     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) =
  24.694 -      (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
  24.695 -apply auto
  24.696 - prefer 2 apply force
  24.697 -apply (cut_tac x = K in reals_Archimedean2, clarify)
  24.698 -apply (rule_tac x = n in exI, clarify)
  24.699 -apply (drule_tac x = na in spec)
  24.700 -apply (auto simp add: real_of_nat_Suc)
  24.701 -done
  24.702 -
  24.703 -text{* alternative definition for Bseq *}
  24.704 -lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
  24.705 -apply (simp add: Bseq_def)
  24.706 -apply (simp (no_asm) add: lemma_NBseq_def)
  24.707 -done
  24.708 -
  24.709 -lemma lemma_NBseq_def2:
  24.710 -     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
  24.711 -apply (subst lemma_NBseq_def, auto)
  24.712 -apply (rule_tac x = "Suc N" in exI)
  24.713 -apply (rule_tac [2] x = N in exI)
  24.714 -apply (auto simp add: real_of_nat_Suc)
  24.715 - prefer 2 apply (blast intro: order_less_imp_le)
  24.716 -apply (drule_tac x = n in spec, simp)
  24.717 -done
  24.718 -
  24.719 -(* yet another definition for Bseq *)
  24.720 -lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
  24.721 -by (simp add: Bseq_def lemma_NBseq_def2)
  24.722 -
  24.723 -subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
  24.724 -
  24.725 -lemma Bseq_isUb:
  24.726 -  "!!(X::nat=>real). Bseq X ==> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
  24.727 -by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
  24.728 -
  24.729 -
  24.730 -text{* Use completeness of reals (supremum property)
  24.731 -   to show that any bounded sequence has a least upper bound*}
  24.732 -
  24.733 -lemma Bseq_isLub:
  24.734 -  "!!(X::nat=>real). Bseq X ==>
  24.735 -   \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
  24.736 -by (blast intro: reals_complete Bseq_isUb)
  24.737 -
  24.738 -subsubsection{*A Bounded and Monotonic Sequence Converges*}
  24.739 -
  24.740 -lemma lemma_converg1:
  24.741 -     "!!(X::nat=>real). [| \<forall>m. \<forall> n \<ge> m. X m \<le> X n;
  24.742 -                  isLub (UNIV::real set) {x. \<exists>n. X n = x} (X ma)
  24.743 -               |] ==> \<forall>n \<ge> ma. X n = X ma"
  24.744 -apply safe
  24.745 -apply (drule_tac y = "X n" in isLubD2)
  24.746 -apply (blast dest: order_antisym)+
  24.747 -done
  24.748 -
  24.749 -text{* The best of both worlds: Easier to prove this result as a standard
  24.750 -   theorem and then use equivalence to "transfer" it into the
  24.751 -   equivalent nonstandard form if needed!*}
  24.752 -
  24.753 -lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)"
  24.754 -apply (simp add: LIMSEQ_def)
  24.755 -apply (rule_tac x = "X m" in exI, safe)
  24.756 -apply (rule_tac x = m in exI, safe)
  24.757 -apply (drule spec, erule impE, auto)
  24.758 -done
  24.759 -
  24.760 -lemma lemma_converg2:
  24.761 -   "!!(X::nat=>real).
  24.762 -    [| \<forall>m. X m ~= U;  isLub UNIV {x. \<exists>n. X n = x} U |] ==> \<forall>m. X m < U"
  24.763 -apply safe
  24.764 -apply (drule_tac y = "X m" in isLubD2)
  24.765 -apply (auto dest!: order_le_imp_less_or_eq)
  24.766 -done
  24.767 -
  24.768 -lemma lemma_converg3: "!!(X ::nat=>real). \<forall>m. X m \<le> U ==> isUb UNIV {x. \<exists>n. X n = x} U"
  24.769 -by (rule setleI [THEN isUbI], auto)
  24.770 -
  24.771 -text{* FIXME: @{term "U - T < U"} is redundant *}
  24.772 -lemma lemma_converg4: "!!(X::nat=> real).
  24.773 -               [| \<forall>m. X m ~= U;
  24.774 -                  isLub UNIV {x. \<exists>n. X n = x} U;
  24.775 -                  0 < T;
  24.776 -                  U + - T < U
  24.777 -               |] ==> \<exists>m. U + -T < X m & X m < U"
  24.778 -apply (drule lemma_converg2, assumption)
  24.779 -apply (rule ccontr, simp)
  24.780 -apply (simp add: linorder_not_less)
  24.781 -apply (drule lemma_converg3)
  24.782 -apply (drule isLub_le_isUb, assumption)
  24.783 -apply (auto dest: order_less_le_trans)
  24.784 -done
  24.785 -
  24.786 -text{*A standard proof of the theorem for monotone increasing sequence*}
  24.787 -
  24.788 -lemma Bseq_mono_convergent:
  24.789 -     "[| Bseq X; \<forall>m. \<forall>n \<ge> m. X m \<le> X n |] ==> convergent (X::nat=>real)"
  24.790 -apply (simp add: convergent_def)
  24.791 -apply (frule Bseq_isLub, safe)
  24.792 -apply (case_tac "\<exists>m. X m = U", auto)
  24.793 -apply (blast dest: lemma_converg1 Bmonoseq_LIMSEQ)
  24.794 -(* second case *)
  24.795 -apply (rule_tac x = U in exI)
  24.796 -apply (subst LIMSEQ_iff, safe)
  24.797 -apply (frule lemma_converg2, assumption)
  24.798 -apply (drule lemma_converg4, auto)
  24.799 -apply (rule_tac x = m in exI, safe)
  24.800 -apply (subgoal_tac "X m \<le> X n")
  24.801 - prefer 2 apply blast
  24.802 -apply (drule_tac x=n and P="%m. X m < U" in spec, arith)
  24.803 -done
  24.804 -
  24.805 -lemma Bseq_minus_iff: "Bseq (%n. -(X n)) = Bseq X"
  24.806 -by (simp add: Bseq_def)
  24.807 -
  24.808 -text{*Main monotonicity theorem*}
  24.809 -lemma Bseq_monoseq_convergent: "[| Bseq X; monoseq X |] ==> convergent X"
  24.810 -apply (simp add: monoseq_def, safe)
  24.811 -apply (rule_tac [2] convergent_minus_iff [THEN ssubst])
  24.812 -apply (drule_tac [2] Bseq_minus_iff [THEN ssubst])
  24.813 -apply (auto intro!: Bseq_mono_convergent)
  24.814 -done
  24.815 -
  24.816 -subsubsection{*A Few More Equivalence Theorems for Boundedness*}
  24.817 -
  24.818 -text{*alternative formulation for boundedness*}
  24.819 -lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
  24.820 -apply (unfold Bseq_def, safe)
  24.821 -apply (rule_tac [2] x = "k + norm x" in exI)
  24.822 -apply (rule_tac x = K in exI, simp)
  24.823 -apply (rule exI [where x = 0], auto)
  24.824 -apply (erule order_less_le_trans, simp)
  24.825 -apply (drule_tac x=n in spec, fold diff_def)
  24.826 -apply (drule order_trans [OF norm_triangle_ineq2])
  24.827 -apply simp
  24.828 -done
  24.829 -
  24.830 -text{*alternative formulation for boundedness*}
  24.831 -lemma Bseq_iff3: "Bseq X = (\<exists>k > 0. \<exists>N. \<forall>n. norm(X(n) + -X(N)) \<le> k)"
  24.832 -apply safe
  24.833 -apply (simp add: Bseq_def, safe)
  24.834 -apply (rule_tac x = "K + norm (X N)" in exI)
  24.835 -apply auto
  24.836 -apply (erule order_less_le_trans, simp)
  24.837 -apply (rule_tac x = N in exI, safe)
  24.838 -apply (drule_tac x = n in spec)
  24.839 -apply (rule order_trans [OF norm_triangle_ineq], simp)
  24.840 -apply (auto simp add: Bseq_iff2)
  24.841 -done
  24.842 -
  24.843 -lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
  24.844 -apply (simp add: Bseq_def)
  24.845 -apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
  24.846 -apply (drule_tac x = n in spec, arith)
  24.847 -done
  24.848 -
  24.849 -
  24.850 -subsection {* Cauchy Sequences *}
  24.851 -
  24.852 -lemma CauchyI:
  24.853 -  "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
  24.854 -by (simp add: Cauchy_def)
  24.855 -
  24.856 -lemma CauchyD:
  24.857 -  "\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
  24.858 -by (simp add: Cauchy_def)
  24.859 -
  24.860 -subsubsection {* Cauchy Sequences are Bounded *}
  24.861 -
  24.862 -text{*A Cauchy sequence is bounded -- this is the standard
  24.863 -  proof mechanization rather than the nonstandard proof*}
  24.864 -
  24.865 -lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
  24.866 -          ==>  \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
  24.867 -apply (clarify, drule spec, drule (1) mp)
  24.868 -apply (simp only: norm_minus_commute)
  24.869 -apply (drule order_le_less_trans [OF norm_triangle_ineq2])
  24.870 -apply simp
  24.871 -done
  24.872 -
  24.873 -lemma Cauchy_Bseq: "Cauchy X ==> Bseq X"
  24.874 -apply (simp add: Cauchy_def)
  24.875 -apply (drule spec, drule mp, rule zero_less_one, safe)
  24.876 -apply (drule_tac x="M" in spec, simp)
  24.877 -apply (drule lemmaCauchy)
  24.878 -apply (rule_tac k="M" in Bseq_offset)
  24.879 -apply (simp add: Bseq_def)
  24.880 -apply (rule_tac x="1 + norm (X M)" in exI)
  24.881 -apply (rule conjI, rule order_less_le_trans [OF zero_less_one], simp)
  24.882 -apply (simp add: order_less_imp_le)
  24.883 -done
  24.884 -
  24.885 -subsubsection {* Cauchy Sequences are Convergent *}
  24.886 -
  24.887 -axclass banach \<subseteq> real_normed_vector
  24.888 -  Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X"
  24.889 -
  24.890 -theorem LIMSEQ_imp_Cauchy:
  24.891 -  assumes X: "X ----> a" shows "Cauchy X"
  24.892 -proof (rule CauchyI)
  24.893 -  fix e::real assume "0 < e"
  24.894 -  hence "0 < e/2" by simp
  24.895 -  with X have "\<exists>N. \<forall>n\<ge>N. norm (X n - a) < e/2" by (rule LIMSEQ_D)
  24.896 -  then obtain N where N: "\<forall>n\<ge>N. norm (X n - a) < e/2" ..
  24.897 -  show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < e"
  24.898 -  proof (intro exI allI impI)
  24.899 -    fix m assume "N \<le> m"
  24.900 -    hence m: "norm (X m - a) < e/2" using N by fast
  24.901 -    fix n assume "N \<le> n"
  24.902 -    hence n: "norm (X n - a) < e/2" using N by fast
  24.903 -    have "norm (X m - X n) = norm ((X m - a) - (X n - a))" by simp
  24.904 -    also have "\<dots> \<le> norm (X m - a) + norm (X n - a)"
  24.905 -      by (rule norm_triangle_ineq4)
  24.906 -    also from m n have "\<dots> < e" by(simp add:field_simps)
  24.907 -    finally show "norm (X m - X n) < e" .
  24.908 -  qed
  24.909 -qed
  24.910 -
  24.911 -lemma convergent_Cauchy: "convergent X \<Longrightarrow> Cauchy X"
  24.912 -unfolding convergent_def
  24.913 -by (erule exE, erule LIMSEQ_imp_Cauchy)
  24.914 -
  24.915 -text {*
  24.916 -Proof that Cauchy sequences converge based on the one from
  24.917 -http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html
  24.918 -*}
  24.919 -
  24.920 -text {*
  24.921 -  If sequence @{term "X"} is Cauchy, then its limit is the lub of
  24.922 -  @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
  24.923 -*}
  24.924 -
  24.925 -lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u"
  24.926 -by (simp add: isUbI setleI)
  24.927 -
  24.928 -lemma real_abs_diff_less_iff:
  24.929 -  "(\<bar>x - a\<bar> < (r::real)) = (a - r < x \<and> x < a + r)"
  24.930 -by auto
  24.931 -
  24.932 -locale real_Cauchy =
  24.933 -  fixes X :: "nat \<Rightarrow> real"
  24.934 -  assumes X: "Cauchy X"
  24.935 -  fixes S :: "real set"
  24.936 -  defines S_def: "S \<equiv> {x::real. \<exists>N. \<forall>n\<ge>N. x < X n}"
  24.937 -
  24.938 -lemma real_CauchyI:
  24.939 -  assumes "Cauchy X"
  24.940 -  shows "real_Cauchy X"
  24.941 -  proof qed (fact assms)
  24.942 -
  24.943 -lemma (in real_Cauchy) mem_S: "\<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S"
  24.944 -by (unfold S_def, auto)
  24.945 -
  24.946 -lemma (in real_Cauchy) bound_isUb:
  24.947 -  assumes N: "\<forall>n\<ge>N. X n < x"
  24.948 -  shows "isUb UNIV S x"
  24.949 -proof (rule isUb_UNIV_I)
  24.950 -  fix y::real assume "y \<in> S"
  24.951 -  hence "\<exists>M. \<forall>n\<ge>M. y < X n"
  24.952 -    by (simp add: S_def)
  24.953 -  then obtain M where "\<forall>n\<ge>M. y < X n" ..
  24.954 -  hence "y < X (max M N)" by simp
  24.955 -  also have "\<dots> < x" using N by simp
  24.956 -  finally show "y \<le> x"
  24.957 -    by (rule order_less_imp_le)
  24.958 -qed
  24.959 -
  24.960 -lemma (in real_Cauchy) isLub_ex: "\<exists>u. isLub UNIV S u"
  24.961 -proof (rule reals_complete)
  24.962 -  obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < 1"
  24.963 -    using CauchyD [OF X zero_less_one] by fast
  24.964 -  hence N: "\<forall>n\<ge>N. norm (X n - X N) < 1" by simp
  24.965 -  show "\<exists>x. x \<in> S"
  24.966 -  proof
  24.967 -    from N have "\<forall>n\<ge>N. X N - 1 < X n"
  24.968 -      by (simp add: real_abs_diff_less_iff)
  24.969 -    thus "X N - 1 \<in> S" by (rule mem_S)
  24.970 -  qed
  24.971 -  show "\<exists>u. isUb UNIV S u"
  24.972 -  proof
  24.973 -    from N have "\<forall>n\<ge>N. X n < X N + 1"
  24.974 -      by (simp add: real_abs_diff_less_iff)
  24.975 -    thus "isUb UNIV S (X N + 1)"
  24.976 -      by (rule bound_isUb)
  24.977 -  qed
  24.978 -qed
  24.979 -
  24.980 -lemma (in real_Cauchy) isLub_imp_LIMSEQ:
  24.981 -  assumes x: "isLub UNIV S x"
  24.982 -  shows "X ----> x"
  24.983 -proof (rule LIMSEQ_I)
  24.984 -  fix r::real assume "0 < r"
  24.985 -  hence r: "0 < r/2" by simp
  24.986 -  obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. norm (X n - X m) < r/2"
  24.987 -    using CauchyD [OF X r] by fast
  24.988 -  hence "\<forall>n\<ge>N. norm (X n - X N) < r/2" by simp
  24.989 -  hence N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2"
  24.990 -    by (simp only: real_norm_def real_abs_diff_less_iff)
  24.991 -
  24.992 -  from N have "\<forall>n\<ge>N. X N - r/2 < X n" by fast
  24.993 -  hence "X N - r/2 \<in> S" by (rule mem_S)
  24.994 -  hence 1: "X N - r/2 \<le> x" using x isLub_isUb isUbD by fast
  24.995 -
  24.996 -  from N have "\<forall>n\<ge>N. X n < X N + r/2" by fast
  24.997 -  hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb)
  24.998 -  hence 2: "x \<le> X N + r/2" using x isLub_le_isUb by fast
  24.999 -
 24.1000 -  show "\<exists>N. \<forall>n\<ge>N. norm (X n - x) < r"
 24.1001 -  proof (intro exI allI impI)
 24.1002 -    fix n assume n: "N \<le> n"
 24.1003 -    from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
 24.1004 -    thus "norm (X n - x) < r" using 1 2
 24.1005 -      by (simp add: real_abs_diff_less_iff)
 24.1006 -  qed
 24.1007 -qed
 24.1008 -
 24.1009 -lemma (in real_Cauchy) LIMSEQ_ex: "\<exists>x. X ----> x"
 24.1010 -proof -
 24.1011 -  obtain x where "isLub UNIV S x"
 24.1012 -    using isLub_ex by fast
 24.1013 -  hence "X ----> x"
 24.1014 -    by (rule isLub_imp_LIMSEQ)
 24.1015 -  thus ?thesis ..
 24.1016 -qed
 24.1017 -
 24.1018 -lemma real_Cauchy_convergent:
 24.1019 -  fixes X :: "nat \<Rightarrow> real"
 24.1020 -  shows "Cauchy X \<Longrightarrow> convergent X"
 24.1021 -unfolding convergent_def
 24.1022 -by (rule real_Cauchy.LIMSEQ_ex)
 24.1023 - (rule real_CauchyI)
 24.1024 -
 24.1025 -instance real :: banach
 24.1026 -by intro_classes (rule real_Cauchy_convergent)
 24.1027 -
 24.1028 -lemma Cauchy_convergent_iff:
 24.1029 -  fixes X :: "nat \<Rightarrow> 'a::banach"
 24.1030 -  shows "Cauchy X = convergent X"
 24.1031 -by (fast intro: Cauchy_convergent convergent_Cauchy)
 24.1032 -
 24.1033 -
 24.1034 -subsection {* Power Sequences *}
 24.1035 -
 24.1036 -text{*The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
 24.1037 -"x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
 24.1038 -  also fact that bounded and monotonic sequence converges.*}
 24.1039 -
 24.1040 -lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
 24.1041 -apply (simp add: Bseq_def)
 24.1042 -apply (rule_tac x = 1 in exI)
 24.1043 -apply (simp add: power_abs)
 24.1044 -apply (auto dest: power_mono)
 24.1045 -done
 24.1046 -
 24.1047 -lemma monoseq_realpow: "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
 24.1048 -apply (clarify intro!: mono_SucI2)
 24.1049 -apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
 24.1050 -done
 24.1051 -
 24.1052 -lemma convergent_realpow:
 24.1053 -  "[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)"
 24.1054 -by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
 24.1055 -
 24.1056 -lemma LIMSEQ_inverse_realpow_zero_lemma:
 24.1057 -  fixes x :: real
 24.1058 -  assumes x: "0 \<le> x"
 24.1059 -  shows "real n * x + 1 \<le> (x + 1) ^ n"
 24.1060 -apply (induct n)
 24.1061 -apply simp
 24.1062 -apply simp
 24.1063 -apply (rule order_trans)
 24.1064 -prefer 2
 24.1065 -apply (erule mult_left_mono)
 24.1066 -apply (rule add_increasing [OF x], simp)
 24.1067 -apply (simp add: real_of_nat_Suc)
 24.1068 -apply (simp add: ring_distribs)
 24.1069 -apply (simp add: mult_nonneg_nonneg x)
 24.1070 -done
 24.1071 -
 24.1072 -lemma LIMSEQ_inverse_realpow_zero:
 24.1073 -  "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) ----> 0"
 24.1074 -proof (rule LIMSEQ_inverse_zero [rule_format])
 24.1075 -  fix y :: real
 24.1076 -  assume x: "1 < x"
 24.1077 -  hence "0 < x - 1" by simp
 24.1078 -  hence "\<forall>y. \<exists>N::nat. y < real N * (x - 1)"
 24.1079 -    by (rule reals_Archimedean3)
 24.1080 -  hence "\<exists>N::nat. y < real N * (x - 1)" ..
 24.1081 -  then obtain N::nat where "y < real N * (x - 1)" ..
 24.1082 -  also have "\<dots> \<le> real N * (x - 1) + 1" by simp
 24.1083 -  also have "\<dots> \<le> (x - 1 + 1) ^ N"
 24.1084 -    by (rule LIMSEQ_inverse_realpow_zero_lemma, cut_tac x, simp)
 24.1085 -  also have "\<dots> = x ^ N" by simp
 24.1086 -  finally have "y < x ^ N" .
 24.1087 -  hence "\<forall>n\<ge>N. y < x ^ n"
 24.1088 -    apply clarify
 24.1089 -    apply (erule order_less_le_trans)
 24.1090 -    apply (erule power_increasing)
 24.1091 -    apply (rule order_less_imp_le [OF x])
 24.1092 -    done
 24.1093 -  thus "\<exists>N. \<forall>n\<ge>N. y < x ^ n" ..
 24.1094 -qed
 24.1095 -
 24.1096 -lemma LIMSEQ_realpow_zero:
 24.1097 -  "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
 24.1098 -proof (cases)
 24.1099 -  assume "x = 0"
 24.1100 -  hence "(\<lambda>n. x ^ Suc n) ----> 0" by (simp add: LIMSEQ_const)
 24.1101 -  thus ?thesis by (rule LIMSEQ_imp_Suc)
 24.1102 -next
 24.1103 -  assume "0 \<le> x" and "x \<noteq> 0"
 24.1104 -  hence x0: "0 < x" by simp
 24.1105 -  assume x1: "x < 1"
 24.1106 -  from x0 x1 have "1 < inverse x"
 24.1107 -    by (rule real_inverse_gt_one)
 24.1108 -  hence "(\<lambda>n. inverse (inverse x ^ n)) ----> 0"
 24.1109 -    by (rule LIMSEQ_inverse_realpow_zero)
 24.1110 -  thus ?thesis by (simp add: power_inverse)
 24.1111 -qed
 24.1112 -
 24.1113 -lemma LIMSEQ_power_zero:
 24.1114 -  fixes x :: "'a::{real_normed_algebra_1,recpower}"
 24.1115 -  shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
 24.1116 -apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
 24.1117 -apply (simp only: LIMSEQ_Zseq_iff, erule Zseq_le)
 24.1118 -apply (simp add: power_abs norm_power_ineq)
 24.1119 -done
 24.1120 -
 24.1121 -lemma LIMSEQ_divide_realpow_zero:
 24.1122 -  "1 < (x::real) ==> (%n. a / (x ^ n)) ----> 0"
 24.1123 -apply (cut_tac a = a and x1 = "inverse x" in
 24.1124 -        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_realpow_zero])
 24.1125 -apply (auto simp add: divide_inverse power_inverse)
 24.1126 -apply (simp add: inverse_eq_divide pos_divide_less_eq)
 24.1127 -done
 24.1128 -
 24.1129 -text{*Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}*}
 24.1130 -
 24.1131 -lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < (1::real) ==> (%n. \<bar>c\<bar> ^ n) ----> 0"
 24.1132 -by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
 24.1133 -
 24.1134 -lemma LIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < (1::real) ==> (%n. c ^ n) ----> 0"
 24.1135 -apply (rule LIMSEQ_rabs_zero [THEN iffD1])
 24.1136 -apply (auto intro: LIMSEQ_rabs_realpow_zero simp add: power_abs)
 24.1137 -done
 24.1138 -
 24.1139 -end
    25.1 --- a/src/HOL/IsaMakefile	Mon Dec 29 13:23:53 2008 +0100
    25.2 +++ b/src/HOL/IsaMakefile	Mon Dec 29 14:08:08 2008 +0100
    25.3 @@ -261,7 +261,7 @@
    25.4  $(OUT)/HOL: ROOT.ML $(MAIN_DEPENDENCIES) \
    25.5    Complex_Main.thy \
    25.6    Complex.thy \
    25.7 -  Complex/Fundamental_Theorem_Algebra.thy \
    25.8 +  Fundamental_Theorem_Algebra.thy \
    25.9    Deriv.thy \
   25.10    Fact.thy \
   25.11    FrechetDeriv.thy \
   25.12 @@ -271,11 +271,11 @@
   25.13    Log.thy \
   25.14    MacLaurin.thy \
   25.15    NthRoot.thy \
   25.16 -  Hyperreal/SEQ.thy \
   25.17 +  SEQ.thy \
   25.18    Series.thy \
   25.19    Taylor.thy \
   25.20    Transcendental.thy \
   25.21 -  Library/Dense_Linear_Order.thy \
   25.22 +  Dense_Linear_Order.thy \
   25.23    GCD.thy \
   25.24    Order_Relation.thy \
   25.25    Parity.thy \
   25.26 @@ -287,7 +287,7 @@
   25.27    RealDef.thy \
   25.28    RealPow.thy \
   25.29    Real.thy \
   25.30 -  Real/RealVector.thy \
   25.31 +  RealVector.thy \
   25.32    Tools/float_syntax.ML \
   25.33    Tools/rat_arith.ML \
   25.34    Tools/real_arith.ML \
   25.35 @@ -337,16 +337,16 @@
   25.36  HOL-HahnBanach: HOL $(LOG)/HOL-HahnBanach.gz
   25.37  
   25.38  $(LOG)/HOL-HahnBanach.gz: $(OUT)/HOL			\
   25.39 -  Real/HahnBanach/Bounds.thy Real/HahnBanach/FunctionNorm.thy		\
   25.40 -  Real/HahnBanach/FunctionOrder.thy Real/HahnBanach/HahnBanach.thy	\
   25.41 -  Real/HahnBanach/HahnBanachExtLemmas.thy				\
   25.42 -  Real/HahnBanach/HahnBanachSupLemmas.thy				\
   25.43 -  Real/HahnBanach/Linearform.thy Real/HahnBanach/NormedSpace.thy	\
   25.44 -  Real/HahnBanach/README.html Real/HahnBanach/ROOT.ML			\
   25.45 -  Real/HahnBanach/Subspace.thy Real/HahnBanach/VectorSpace.thy		\
   25.46 -  Real/HahnBanach/ZornLemma.thy Real/HahnBanach/document/root.bib	\
   25.47 -  Real/HahnBanach/document/root.tex
   25.48 -	@cd Real; $(ISABELLE_TOOL) usedir -g true $(OUT)/HOL HahnBanach
   25.49 +  HahnBanach/Bounds.thy HahnBanach/FunctionNorm.thy		\
   25.50 +  HahnBanach/FunctionOrder.thy HahnBanach/HahnBanach.thy	\
   25.51 +  HahnBanach/HahnBanachExtLemmas.thy				\
   25.52 +  HahnBanach/HahnBanachSupLemmas.thy				\
   25.53 +  HahnBanach/Linearform.thy HahnBanach/NormedSpace.thy	\
   25.54 +  HahnBanach/README.html HahnBanach/ROOT.ML			\
   25.55 +  HahnBanach/Subspace.thy HahnBanach/VectorSpace.thy		\
   25.56 +  HahnBanach/ZornLemma.thy HahnBanach/document/root.bib	\
   25.57 +  HahnBanach/document/root.tex
   25.58 +	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL HahnBanach
   25.59  
   25.60  
   25.61  ## HOL-Subst
    26.1 --- a/src/HOL/Library/Dense_Linear_Order.thy	Mon Dec 29 13:23:53 2008 +0100
    26.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    26.3 @@ -1,880 +0,0 @@
    26.4 -(*
    26.5 -    ID:         $Id$
    26.6 -    Author:     Amine Chaieb, TU Muenchen
    26.7 -*)
    26.8 -
    26.9 -header {* Dense linear order without endpoints
   26.10 -  and a quantifier elimination procedure in Ferrante and Rackoff style *}
   26.11 -
   26.12 -theory Dense_Linear_Order
   26.13 -imports Plain "~~/src/HOL/Groebner_Basis"
   26.14 -uses
   26.15 -  "~~/src/HOL/Tools/Qelim/langford_data.ML"
   26.16 -  "~~/src/HOL/Tools/Qelim/ferrante_rackoff_data.ML"
   26.17 -  ("~~/src/HOL/Tools/Qelim/langford.ML")
   26.18 -  ("~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML")
   26.19 -begin
   26.20 -
   26.21 -setup {* Langford_Data.setup #> Ferrante_Rackoff_Data.setup *}
   26.22 -
   26.23 -context linorder
   26.24 -begin
   26.25 -
   26.26 -lemma less_not_permute: "\<not> (x < y \<and> y < x)" by (simp add: not_less linear)
   26.27 -
   26.28 -lemma gather_simps: 
   26.29 -  shows 
   26.30 -  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y) \<and> P x)"
   26.31 -  and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x \<and> P x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y) \<and> P x)"
   26.32 -  "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> x < u) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> (insert u U). x < y))"
   26.33 -  and "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y) \<and> l < x) \<longleftrightarrow> (\<exists>x. (\<forall>y \<in> (insert l L). y < x) \<and> (\<forall>y \<in> U. x < y))"  by auto
   26.34 -
   26.35 -lemma 
   26.36 -  gather_start: "(\<exists>x. P x) \<equiv> (\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y\<in> {}. x < y) \<and> P x)" 
   26.37 -  by simp
   26.38 -
   26.39 -text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"}*}
   26.40 -lemma minf_lt:  "\<exists>z . \<forall>x. x < z \<longrightarrow> (x < t \<longleftrightarrow> True)" by auto
   26.41 -lemma minf_gt: "\<exists>z . \<forall>x. x < z \<longrightarrow>  (t < x \<longleftrightarrow>  False)"
   26.42 -  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
   26.43 -
   26.44 -lemma minf_le: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<le> t \<longleftrightarrow> True)" by (auto simp add: less_le)
   26.45 -lemma minf_ge: "\<exists>z. \<forall>x. x < z \<longrightarrow> (t \<le> x \<longleftrightarrow> False)"
   26.46 -  by (auto simp add: less_le not_less not_le)
   26.47 -lemma minf_eq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
   26.48 -lemma minf_neq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
   26.49 -lemma minf_P: "\<exists>z. \<forall>x. x < z \<longrightarrow> (P \<longleftrightarrow> P)" by blast
   26.50 -
   26.51 -text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"}*}
   26.52 -lemma pinf_gt:  "\<exists>z . \<forall>x. z < x \<longrightarrow> (t < x \<longleftrightarrow> True)" by auto
   26.53 -lemma pinf_lt: "\<exists>z . \<forall>x. z < x \<longrightarrow>  (x < t \<longleftrightarrow>  False)"
   26.54 -  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
   26.55 -
   26.56 -lemma pinf_ge: "\<exists>z. \<forall>x. z < x \<longrightarrow> (t \<le> x \<longleftrightarrow> True)" by (auto simp add: less_le)
   26.57 -lemma pinf_le: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<le> t \<longleftrightarrow> False)"
   26.58 -  by (auto simp add: less_le not_less not_le)
   26.59 -lemma pinf_eq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
   26.60 -lemma pinf_neq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
   26.61 -lemma pinf_P: "\<exists>z. \<forall>x. z < x \<longrightarrow> (P \<longleftrightarrow> P)" by blast
   26.62 -
   26.63 -lemma nmi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x < t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
   26.64 -lemma nmi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t < x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)"
   26.65 -  by (auto simp add: le_less)
   26.66 -lemma  nmi_le: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x\<le> t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
   26.67 -lemma  nmi_ge: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t\<le> x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
   26.68 -lemma  nmi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x = t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
   26.69 -lemma  nmi_neq: "t \<in> U \<Longrightarrow>\<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
   26.70 -lemma  nmi_P: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
   26.71 -lemma  nmi_conj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x) ;
   26.72 -  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
   26.73 -  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
   26.74 -lemma  nmi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x) ;
   26.75 -  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
   26.76 -  \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
   26.77 -
   26.78 -lemma  npi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x < t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by (auto simp add: le_less)
   26.79 -lemma  npi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t < x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
   26.80 -lemma  npi_le: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x \<le> t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
   26.81 -lemma  npi_ge: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t \<le> x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
   26.82 -lemma  npi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x = t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
   26.83 -lemma  npi_neq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x \<noteq> t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u )" by auto
   26.84 -lemma  npi_P: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
   26.85 -lemma  npi_conj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u) ;  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)\<rbrakk>
   26.86 -  \<Longrightarrow>  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
   26.87 -lemma  npi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u) ; \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)\<rbrakk>
   26.88 -  \<Longrightarrow> \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
   26.89 -
   26.90 -lemma lin_dense_lt: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t < u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x < t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y < t)"
   26.91 -proof(clarsimp)
   26.92 -  fix x l u y  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x"
   26.93 -    and xu: "x<u"  and px: "x < t" and ly: "l<y" and yu:"y < u"
   26.94 -  from tU noU ly yu have tny: "t\<noteq>y" by auto
   26.95 -  {assume H: "t < y"
   26.96 -    from less_trans[OF lx px] less_trans[OF H yu]
   26.97 -    have "l < t \<and> t < u"  by simp
   26.98 -    with tU noU have "False" by auto}
   26.99 -  hence "\<not> t < y"  by auto hence "y \<le> t" by (simp add: not_less)
  26.100 -  thus "y < t" using tny by (simp add: less_le)
  26.101 -qed
  26.102 -
  26.103 -lemma lin_dense_gt: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l < x \<and> x < u \<and> t < x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t < y)"
  26.104 -proof(clarsimp)
  26.105 -  fix x l u y
  26.106 -  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
  26.107 -  and px: "t < x" and ly: "l<y" and yu:"y < u"
  26.108 -  from tU noU ly yu have tny: "t\<noteq>y" by auto
  26.109 -  {assume H: "y< t"
  26.110 -    from less_trans[OF ly H] less_trans[OF px xu] have "l < t \<and> t < u" by simp
  26.111 -    with tU noU have "False" by auto}
  26.112 -  hence "\<not> y<t"  by auto hence "t \<le> y" by (auto simp add: not_less)
  26.113 -  thus "t < y" using tny by (simp add:less_le)
  26.114 -qed
  26.115 -
  26.116 -lemma lin_dense_le: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<le> t \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<le> t)"
  26.117 -proof(clarsimp)
  26.118 -  fix x l u y
  26.119 -  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
  26.120 -  and px: "x \<le> t" and ly: "l<y" and yu:"y < u"
  26.121 -  from tU noU ly yu have tny: "t\<noteq>y" by auto
  26.122 -  {assume H: "t < y"
  26.123 -    from less_le_trans[OF lx px] less_trans[OF H yu]
  26.124 -    have "l < t \<and> t < u" by simp
  26.125 -    with tU noU have "False" by auto}
  26.126 -  hence "\<not> t < y"  by auto thus "y \<le> t" by (simp add: not_less)
  26.127 -qed
  26.128 -
  26.129 -lemma lin_dense_ge: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> t \<le> x \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> t \<le> y)"
  26.130 -proof(clarsimp)
  26.131 -  fix x l u y
  26.132 -  assume tU: "t \<in> U" and noU: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<notin> U" and lx: "l < x" and xu: "x<u"
  26.133 -  and px: "t \<le> x" and ly: "l<y" and yu:"y < u"
  26.134 -  from tU noU ly yu have tny: "t\<noteq>y" by auto
  26.135 -  {assume H: "y< t"
  26.136 -    from less_trans[OF ly H] le_less_trans[OF px xu]
  26.137 -    have "l < t \<and> t < u" by simp
  26.138 -    with tU noU have "False" by auto}
  26.139 -  hence "\<not> y<t"  by auto thus "t \<le> y" by (simp add: not_less)
  26.140 -qed
  26.141 -lemma lin_dense_eq: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x = t   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y= t)"  by auto
  26.142 -lemma lin_dense_neq: "t \<in> U \<Longrightarrow> \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> x \<noteq> t   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> y\<noteq> t)"  by auto
  26.143 -lemma lin_dense_P: "\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P   \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P)"  by auto
  26.144 -
  26.145 -lemma lin_dense_conj:
  26.146 -  "\<lbrakk>\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P1 x
  26.147 -  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
  26.148 -  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
  26.149 -  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
  26.150 -  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> (P1 x \<and> P2 x)
  26.151 -  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<and> P2 y))"
  26.152 -  by blast
  26.153 -lemma lin_dense_disj:
  26.154 -  "\<lbrakk>\<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P1 x
  26.155 -  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
  26.156 -  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
  26.157 -  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
  26.158 -  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> (P1 x \<or> P2 x)
  26.159 -  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<or> P2 y))"
  26.160 -  by blast
  26.161 -
  26.162 -lemma npmibnd: "\<lbrakk>\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<le> x); \<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<le> u)\<rbrakk>
  26.163 -  \<Longrightarrow> \<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<le> x \<and> x \<le> u')"
  26.164 -by auto
  26.165 -
  26.166 -lemma finite_set_intervals:
  26.167 -  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
  26.168 -  and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
  26.169 -  shows "\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a \<le> x \<and> x \<le> b \<and> P x"
  26.170 -proof-
  26.171 -  let ?Mx = "{y. y\<in> S \<and> y \<le> x}"
  26.172 -  let ?xM = "{y. y\<in> S \<and> x \<le> y}"
  26.173 -  let ?a = "Max ?Mx"
  26.174 -  let ?b = "Min ?xM"
  26.175 -  have MxS: "?Mx \<subseteq> S" by blast
  26.176 -  hence fMx: "finite ?Mx" using fS finite_subset by auto
  26.177 -  from lx linS have linMx: "l \<in> ?Mx" by blast
  26.178 -  hence Mxne: "?Mx \<noteq> {}" by blast
  26.179 -  have xMS: "?xM \<subseteq> S" by blast
  26.180 -  hence fxM: "finite ?xM" using fS finite_subset by auto
  26.181 -  from xu uinS have linxM: "u \<in> ?xM" by blast
  26.182 -  hence xMne: "?xM \<noteq> {}" by blast
  26.183 -  have ax:"?a \<le> x" using Mxne fMx by auto
  26.184 -  have xb:"x \<le> ?b" using xMne fxM by auto
  26.185 -  have "?a \<in> ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \<in> S" using MxS by blast
  26.186 -  have "?b \<in> ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \<in> S" using xMS by blast
  26.187 -  have noy:"\<forall> y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S"
  26.188 -  proof(clarsimp)
  26.189 -    fix y   assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S"
  26.190 -    from yS have "y\<in> ?Mx \<or> y\<in> ?xM" by (auto simp add: linear)
  26.191 -    moreover {assume "y \<in> ?Mx" hence "y \<le> ?a" using Mxne fMx by auto with ay have "False" by (simp add: not_le[symmetric])}
  26.192 -    moreover {assume "y \<in> ?xM" hence "?b \<le> y" using xMne fxM by auto with yb have "False" by (simp add: not_le[symmetric])}
  26.193 -    ultimately show "False" by blast
  26.194 -  qed
  26.195 -  from ainS binS noy ax xb px show ?thesis by blast
  26.196 -qed
  26.197 -
  26.198 -lemma finite_set_intervals2:
  26.199 -  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
  26.200 -  and uinS: "u \<in> S" and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
  26.201 -  shows "(\<exists> s\<in> S. P s) \<or> (\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a < x \<and> x < b \<and> P x)"
  26.202 -proof-
  26.203 -  from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su]
  26.204 -  obtain a and b where
  26.205 -    as: "a\<in> S" and bs: "b\<in> S" and noS:"\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S"
  26.206 -    and axb: "a \<le> x \<and> x \<le> b \<and> P x"  by auto
  26.207 -  from axb have "x= a \<or> x= b \<or> (a < x \<and> x < b)" by (auto simp add: le_less)
  26.208 -  thus ?thesis using px as bs noS by blast
  26.209 -qed
  26.210 -
  26.211 -end
  26.212 -
  26.213 -section {* The classical QE after Langford for dense linear orders *}
  26.214 -
  26.215 -context dense_linear_order
  26.216 -begin
  26.217 -
  26.218 -lemma interval_empty_iff:
  26.219 -  "{y. x < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
  26.220 -  by (auto dest: dense)
  26.221 -
  26.222 -lemma dlo_qe_bnds: 
  26.223 -  assumes ne: "L \<noteq> {}" and neU: "U \<noteq> {}" and fL: "finite L" and fU: "finite U"
  26.224 -  shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> (\<forall> l \<in> L. \<forall>u \<in> U. l < u)"
  26.225 -proof (simp only: atomize_eq, rule iffI)
  26.226 -  assume H: "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)"
  26.227 -  then obtain x where xL: "\<forall>y\<in>L. y < x" and xU: "\<forall>y\<in>U. x < y" by blast
  26.228 -  {fix l u assume l: "l \<in> L" and u: "u \<in> U"
  26.229 -    have "l < x" using xL l by blast
  26.230 -    also have "x < u" using xU u by blast
  26.231 -    finally (less_trans) have "l < u" .}
  26.232 -  thus "\<forall>l\<in>L. \<forall>u\<in>U. l < u" by blast
  26.233 -next
  26.234 -  assume H: "\<forall>l\<in>L. \<forall>u\<in>U. l < u"
  26.235 -  let ?ML = "Max L"
  26.236 -  let ?MU = "Min U"  
  26.237 -  from fL ne have th1: "?ML \<in> L" and th1': "\<forall>l\<in>L. l \<le> ?ML" by auto
  26.238 -  from fU neU have th2: "?MU \<in> U" and th2': "\<forall>u\<in>U. ?MU \<le> u" by auto
  26.239 -  from th1 th2 H have "?ML < ?MU" by auto
  26.240 -  with dense obtain w where th3: "?ML < w" and th4: "w < ?MU" by blast
  26.241 -  from th3 th1' have "\<forall>l \<in> L. l < w" by auto
  26.242 -  moreover from th4 th2' have "\<forall>u \<in> U. w < u" by auto
  26.243 -  ultimately show "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)" by auto
  26.244 -qed
  26.245 -
  26.246 -lemma dlo_qe_noub: 
  26.247 -  assumes ne: "L \<noteq> {}" and fL: "finite L"
  26.248 -  shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> {}. x < y)) \<equiv> True"
  26.249 -proof(simp add: atomize_eq)
  26.250 -  from gt_ex[of "Max L"] obtain M where M: "Max L < M" by blast
  26.251 -  from ne fL have "\<forall>x \<in> L. x \<le> Max L" by simp
  26.252 -  with M have "\<forall>x\<in>L. x < M" by (auto intro: le_less_trans)
  26.253 -  thus "\<exists>x. \<forall>y\<in>L. y < x" by blast
  26.254 -qed
  26.255 -
  26.256 -lemma dlo_qe_nolb: 
  26.257 -  assumes ne: "U \<noteq> {}" and fU: "finite U"
  26.258 -  shows "(\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> True"
  26.259 -proof(simp add: atomize_eq)
  26.260 -  from lt_ex[of "Min U"] obtain M where M: "M < Min U" by blast
  26.261 -  from ne fU have "\<forall>x \<in> U. Min U \<le> x" by simp
  26.262 -  with M have "\<forall>x\<in>U. M < x" by (auto intro: less_le_trans)
  26.263 -  thus "\<exists>x. \<forall>y\<in>U. x < y" by blast
  26.264 -qed
  26.265 -
  26.266 -lemma exists_neq: "\<exists>(x::'a). x \<noteq> t" "\<exists>(x::'a). t \<noteq> x" 
  26.267 -  using gt_ex[of t] by auto
  26.268 -
  26.269 -lemmas dlo_simps = order_refl less_irrefl not_less not_le exists_neq 
  26.270 -  le_less neq_iff linear less_not_permute
  26.271 -
  26.272 -lemma axiom: "dense_linear_order (op \<le>) (op <)" by (rule dense_linear_order_axioms)
  26.273 -lemma atoms:
  26.274 -  shows "TERM (less :: 'a \<Rightarrow> _)"
  26.275 -    and "TERM (less_eq :: 'a \<Rightarrow> _)"
  26.276 -    and "TERM (op = :: 'a \<Rightarrow> _)" .
  26.277 -
  26.278 -declare axiom[langford qe: dlo_qe_bnds dlo_qe_nolb dlo_qe_noub gather: gather_start gather_simps atoms: atoms]
  26.279 -declare dlo_simps[langfordsimp]
  26.280 -
  26.281 -end
  26.282 -
  26.283 -(* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *)
  26.284 -lemma dnf:
  26.285 -  "(P & (Q | R)) = ((P&Q) | (P&R))" 
  26.286 -  "((Q | R) & P) = ((Q&P) | (R&P))"
  26.287 -  by blast+
  26.288 -
  26.289 -lemmas weak_dnf_simps = simp_thms dnf
  26.290 -
  26.291 -lemma nnf_simps:
  26.292 -    "(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)" "(P \<longrightarrow> Q) = (\<not>P \<or> Q)"
  26.293 -    "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
  26.294 -  by blast+
  26.295 -
  26.296 -lemma ex_distrib: "(\<exists>x. P x \<or> Q x) \<longleftrightarrow> ((\<exists>x. P x) \<or> (\<exists>x. Q x))" by blast
  26.297 -
  26.298 -lemmas dnf_simps = weak_dnf_simps nnf_simps ex_distrib
  26.299 -
  26.300 -use "~~/src/HOL/Tools/Qelim/langford.ML"
  26.301 -method_setup dlo = {*
  26.302 -  Method.ctxt_args (Method.SIMPLE_METHOD' o LangfordQE.dlo_tac)
  26.303 -*} "Langford's algorithm for quantifier elimination in dense linear orders"
  26.304 -
  26.305 -
  26.306 -section {* Contructive dense linear orders yield QE for linear arithmetic over ordered Fields -- see @{text "Arith_Tools.thy"} *}
  26.307 -
  26.308 -text {* Linear order without upper bounds *}
  26.309 -
  26.310 -locale linorder_stupid_syntax = linorder
  26.311 -begin
  26.312 -notation
  26.313 -  less_eq  ("op \<sqsubseteq>") and
  26.314 -  less_eq  ("(_/ \<sqsubseteq> _)" [51, 51] 50) and
  26.315 -  less  ("op \<sqsubset>") and
  26.316 -  less  ("(_/ \<sqsubset> _)"  [51, 51] 50)
  26.317 -
  26.318 -end
  26.319 -
  26.320 -locale linorder_no_ub = linorder_stupid_syntax +
  26.321 -  assumes gt_ex: "\<exists>y. less x y"
  26.322 -begin
  26.323 -lemma ge_ex: "\<exists>y. x \<sqsubseteq> y" using gt_ex by auto
  26.324 -
  26.325 -text {* Theorems for @{text "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"} *}
  26.326 -lemma pinf_conj:
  26.327 -  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
  26.328 -  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
  26.329 -  shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
  26.330 -proof-
  26.331 -  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
  26.332 -     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
  26.333 -  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
  26.334 -  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
  26.335 -  {fix x assume H: "z \<sqsubset> x"
  26.336 -    from less_trans[OF zz1 H] less_trans[OF zz2 H]
  26.337 -    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
  26.338 -  }
  26.339 -  thus ?thesis by blast
  26.340 -qed
  26.341 -
  26.342 -lemma pinf_disj:
  26.343 -  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
  26.344 -  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
  26.345 -  shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
  26.346 -proof-
  26.347 -  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
  26.348 -     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
  26.349 -  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
  26.350 -  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
  26.351 -  {fix x assume H: "z \<sqsubset> x"
  26.352 -    from less_trans[OF zz1 H] less_trans[OF zz2 H]
  26.353 -    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
  26.354 -  }
  26.355 -  thus ?thesis by blast
  26.356 -qed
  26.357 -
  26.358 -lemma pinf_ex: assumes ex:"\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x"
  26.359 -proof-
  26.360 -  from ex obtain z where z: "\<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
  26.361 -  from gt_ex obtain x where x: "z \<sqsubset> x" by blast
  26.362 -  from z x p1 show ?thesis by blast
  26.363 -qed
  26.364 -
  26.365 -end
  26.366 -
  26.367 -text {* Linear order without upper bounds *}
  26.368 -
  26.369 -locale linorder_no_lb = linorder_stupid_syntax +
  26.370 -  assumes lt_ex: "\<exists>y. less y x"
  26.371 -begin
  26.372 -lemma le_ex: "\<exists>y. y \<sqsubseteq> x" using lt_ex by auto
  26.373 -
  26.374 -
  26.375 -text {* Theorems for @{text "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"} *}
  26.376 -lemma minf_conj:
  26.377 -  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
  26.378 -  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
  26.379 -  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
  26.380 -proof-
  26.381 -  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
  26.382 -  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
  26.383 -  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
  26.384 -  {fix x assume H: "x \<sqsubset> z"
  26.385 -    from less_trans[OF H zz1] less_trans[OF H zz2]
  26.386 -    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
  26.387 -  }
  26.388 -  thus ?thesis by blast
  26.389 -qed
  26.390 -
  26.391 -lemma minf_disj:
  26.392 -  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
  26.393 -  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
  26.394 -  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
  26.395 -proof-
  26.396 -  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"and z2: "\<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
  26.397 -  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
  26.398 -  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
  26.399 -  {fix x assume H: "x \<sqsubset> z"
  26.400 -    from less_trans[OF H zz1] less_trans[OF H zz2]
  26.401 -    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
  26.402 -  }
  26.403 -  thus ?thesis by blast
  26.404 -qed
  26.405 -
  26.406 -lemma minf_ex: assumes ex:"\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" and p1: P1 shows "\<exists> x. P x"
  26.407 -proof-
  26.408 -  from ex obtain z where z: "\<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
  26.409 -  from lt_ex obtain x where x: "x \<sqsubset> z" by blast
  26.410 -  from z x p1 show ?thesis by blast
  26.411 -qed
  26.412 -
  26.413 -end
  26.414 -
  26.415 -
  26.416 -locale constr_dense_linear_order = linorder_no_lb + linorder_no_ub +
  26.417 -  fixes between
  26.418 -  assumes between_less: "less x y \<Longrightarrow> less x (between x y) \<and> less (between x y) y"
  26.419 -     and  between_same: "between x x = x"
  26.420 -
  26.421 -interpretation  constr_dense_linear_order < dense_linear_order 
  26.422 -  apply unfold_locales
  26.423 -  using gt_ex lt_ex between_less
  26.424 -    by (auto, rule_tac x="between x y" in exI, simp)
  26.425 -
  26.426 -context  constr_dense_linear_order
  26.427 -begin
  26.428 -
  26.429 -lemma rinf_U:
  26.430 -  assumes fU: "finite U"
  26.431 -  and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
  26.432 -  \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
  26.433 -  and nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')"
  26.434 -  and nmi: "\<not> MP"  and npi: "\<not> PP"  and ex: "\<exists> x.  P x"
  26.435 -  shows "\<exists> u\<in> U. \<exists> u' \<in> U. P (between u u')"
  26.436 -proof-
  26.437 -  from ex obtain x where px: "P x" by blast
  26.438 -  from px nmi npi nmpiU have "\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u'" by auto
  26.439 -  then obtain u and u' where uU:"u\<in> U" and uU': "u' \<in> U" and ux:"u \<sqsubseteq> x" and xu':"x \<sqsubseteq> u'" by auto
  26.440 -  from uU have Une: "U \<noteq> {}" by auto
  26.441 -  term "linorder.Min less_eq"
  26.442 -  let ?l = "linorder.Min less_eq U"
  26.443 -  let ?u = "linorder.Max less_eq U"
  26.444 -  have linM: "?l \<in> U" using fU Une by simp
  26.445 -  have uinM: "?u \<in> U" using fU Une by simp
  26.446 -  have lM: "\<forall> t\<in> U. ?l \<sqsubseteq> t" using Une fU by auto
  26.447 -  have Mu: "\<forall> t\<in> U. t \<sqsubseteq> ?u" using Une fU by auto
  26.448 -  have th:"?l \<sqsubseteq> u" using uU Une lM by auto
  26.449 -  from order_trans[OF th ux] have lx: "?l \<sqsubseteq> x" .
  26.450 -  have th: "u' \<sqsubseteq> ?u" using uU' Une Mu by simp
  26.451 -  from order_trans[OF xu' th] have xu: "x \<sqsubseteq> ?u" .
  26.452 -  from finite_set_intervals2[where P="P",OF px lx xu linM uinM fU lM Mu]
  26.453 -  have "(\<exists> s\<in> U. P s) \<or>
  26.454 -      (\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x)" .
  26.455 -  moreover { fix u assume um: "u\<in>U" and pu: "P u"
  26.456 -    have "between u u = u" by (simp add: between_same)
  26.457 -    with um pu have "P (between u u)" by simp
  26.458 -    with um have ?thesis by blast}
  26.459 -  moreover{
  26.460 -    assume "\<exists> t1\<in> U. \<exists> t2 \<in> U. (\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U) \<and> t1 \<sqsubset> x \<and> x \<sqsubset> t2 \<and> P x"
  26.461 -      then obtain t1 and t2 where t1M: "t1 \<in> U" and t2M: "t2\<in> U"
  26.462 -        and noM: "\<forall> y. t1 \<sqsubset> y \<and> y \<sqsubset> t2 \<longrightarrow> y \<notin> U" and t1x: "t1 \<sqsubset> x" and xt2: "x \<sqsubset> t2" and px: "P x"
  26.463 -        by blast
  26.464 -      from less_trans[OF t1x xt2] have t1t2: "t1 \<sqsubset> t2" .
  26.465 -      let ?u = "between t1 t2"
  26.466 -      from between_less t1t2 have t1lu: "t1 \<sqsubset> ?u" and ut2: "?u \<sqsubset> t2" by auto
  26.467 -      from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast
  26.468 -      with t1M t2M have ?thesis by blast}
  26.469 -    ultimately show ?thesis by blast
  26.470 -  qed
  26.471 -
  26.472 -theorem fr_eq:
  26.473 -  assumes fU: "finite U"
  26.474 -  and lin_dense: "\<forall>x l u. (\<forall> t. l \<sqsubset> t \<and> t\<sqsubset> u \<longrightarrow> t \<notin> U) \<and> l\<sqsubset> x \<and> x \<sqsubset> u \<and> P x
  26.475 -   \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
  26.476 -  and nmibnd: "\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<sqsubseteq> x)"
  26.477 -  and npibnd: "\<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<sqsubseteq> u)"
  26.478 -  and mi: "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x = MP)"  and pi: "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x = PP)"
  26.479 -  shows "(\<exists> x. P x) \<equiv> (MP \<or> PP \<or> (\<exists> u \<in> U. \<exists> u'\<in> U. P (between u u')))"
  26.480 -  (is "_ \<equiv> (_ \<or> _ \<or> ?F)" is "?E \<equiv> ?D")
  26.481 -proof-
  26.482 - {
  26.483 -   assume px: "\<exists> x. P x"
  26.484 -   have "MP \<or> PP \<or> (\<not> MP \<and> \<not> PP)" by blast
  26.485 -   moreover {assume "MP \<or> PP" hence "?D" by blast}
  26.486 -   moreover {assume nmi: "\<not> MP" and npi: "\<not> PP"
  26.487 -     from npmibnd[OF nmibnd npibnd]
  26.488 -     have nmpiU: "\<forall>x. \<not> MP \<and> \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u')" .
  26.489 -     from rinf_U[OF fU lin_dense nmpiU nmi npi px] have "?D" by blast}
  26.490 -   ultimately have "?D" by blast}
  26.491 - moreover
  26.492 - { assume "?D"
  26.493 -   moreover {assume m:"MP" from minf_ex[OF mi m] have "?E" .}
  26.494 -   moreover {assume p: "PP" from pinf_ex[OF pi p] have "?E" . }
  26.495 -   moreover {assume f:"?F" hence "?E" by blast}
  26.496 -   ultimately have "?E" by blast}
  26.497 - ultimately have "?E = ?D" by blast thus "?E \<equiv> ?D" by simp
  26.498 -qed
  26.499 -
  26.500 -lemmas minf_thms = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P
  26.501 -lemmas pinf_thms = pinf_conj pinf_disj pinf_eq pinf_neq pinf_lt pinf_le pinf_gt pinf_ge pinf_P
  26.502 -
  26.503 -lemmas nmi_thms = nmi_conj nmi_disj nmi_eq nmi_neq nmi_lt nmi_le nmi_gt nmi_ge nmi_P
  26.504 -lemmas npi_thms = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P
  26.505 -lemmas lin_dense_thms = lin_dense_conj lin_dense_disj lin_dense_eq lin_dense_neq lin_dense_lt lin_dense_le lin_dense_gt lin_dense_ge lin_dense_P
  26.506 -
  26.507 -lemma ferrack_axiom: "constr_dense_linear_order less_eq less between"
  26.508 -  by (rule constr_dense_linear_order_axioms)
  26.509 -lemma atoms:
  26.510 -  shows "TERM (less :: 'a \<Rightarrow> _)"
  26.511 -    and "TERM (less_eq :: 'a \<Rightarrow> _)"
  26.512 -    and "TERM (op = :: 'a \<Rightarrow> _)" .
  26.513 -
  26.514 -declare ferrack_axiom [ferrack minf: minf_thms pinf: pinf_thms
  26.515 -    nmi: nmi_thms npi: npi_thms lindense:
  26.516 -    lin_dense_thms qe: fr_eq atoms: atoms]
  26.517 -
  26.518 -declaration {*
  26.519 -let
  26.520 -fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
  26.521 -fun generic_whatis phi =
  26.522 - let
  26.523 -  val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
  26.524 -  fun h x t =
  26.525 -   case term_of t of
  26.526 -     Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
  26.527 -                            else Ferrante_Rackoff_Data.Nox
  26.528 -   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
  26.529 -                            else Ferrante_Rackoff_Data.Nox
  26.530 -   | b$y$z => if Term.could_unify (b, lt) then
  26.531 -                 if term_of x aconv y then Ferrante_Rackoff_Data.Lt
  26.532 -                 else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
  26.533 -                 else Ferrante_Rackoff_Data.Nox
  26.534 -             else if Term.could_unify (b, le) then
  26.535 -                 if term_of x aconv y then Ferrante_Rackoff_Data.Le
  26.536 -                 else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
  26.537 -                 else Ferrante_Rackoff_Data.Nox
  26.538 -             else Ferrante_Rackoff_Data.Nox
  26.539 -   | _ => Ferrante_Rackoff_Data.Nox
  26.540 - in h end
  26.541 - fun ss phi = HOL_ss addsimps (simps phi)
  26.542 -in
  26.543 - Ferrante_Rackoff_Data.funs  @{thm "ferrack_axiom"}
  26.544 -  {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss}
  26.545 -end
  26.546 -*}
  26.547 -
  26.548 -end
  26.549 -
  26.550 -use "~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML"
  26.551 -
  26.552 -method_setup ferrack = {*
  26.553 -  Method.ctxt_args (Method.SIMPLE_METHOD' o FerranteRackoff.dlo_tac)
  26.554 -*} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders"
  26.555 -
  26.556 -subsection {* Ferrante and Rackoff algorithm over ordered fields *}
  26.557 -
  26.558 -lemma neg_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x < 0) == (x > 0))"
  26.559 -proof-
  26.560 -  assume H: "c < 0"
  26.561 -  have "c*x < 0 = (0/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
  26.562 -  also have "\<dots> = (0 < x)" by simp
  26.563 -  finally show  "(c*x < 0) == (x > 0)" by simp
  26.564 -qed
  26.565 -
  26.566 -lemma pos_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x < 0) == (x < 0))"
  26.567 -proof-
  26.568 -  assume H: "c > 0"
  26.569 -  hence "c*x < 0 = (0/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
  26.570 -  also have "\<dots> = (0 > x)" by simp
  26.571 -  finally show  "(c*x < 0) == (x < 0)" by simp
  26.572 -qed
  26.573 -
  26.574 -lemma neg_prod_sum_lt: "(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x + t< 0) == (x > (- 1/c)*t))"
  26.575 -proof-
  26.576 -  assume H: "c < 0"
  26.577 -  have "c*x + t< 0 = (c*x < -t)" by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
  26.578 -  also have "\<dots> = (-t/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
  26.579 -  also have "\<dots> = ((- 1/c)*t < x)" by simp
  26.580 -  finally show  "(c*x + t < 0) == (x > (- 1/c)*t)" by simp
  26.581 -qed
  26.582 -
  26.583 -lemma pos_prod_sum_lt:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x + t < 0) == (x < (- 1/c)*t))"
  26.584 -proof-
  26.585 -  assume H: "c > 0"
  26.586 -  have "c*x + t< 0 = (c*x < -t)"  by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
  26.587 -  also have "\<dots> = (-t/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
  26.588 -  also have "\<dots> = ((- 1/c)*t > x)" by simp
  26.589 -  finally show  "(c*x + t < 0) == (x < (- 1/c)*t)" by simp
  26.590 -qed
  26.591 -
  26.592 -lemma sum_lt:"((x::'a::pordered_ab_group_add) + t < 0) == (x < - t)"
  26.593 -  using less_diff_eq[where a= x and b=t and c=0] by simp
  26.594 -
  26.595 -lemma neg_prod_le:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x <= 0) == (x >= 0))"
  26.596 -proof-
  26.597 -  assume H: "c < 0"
  26.598 -  have "c*x <= 0 = (0/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
  26.599 -  also have "\<dots> = (0 <= x)" by simp
  26.600 -  finally show  "(c*x <= 0) == (x >= 0)" by simp
  26.601 -qed
  26.602 -
  26.603 -lemma pos_prod_le:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x <= 0) == (x <= 0))"
  26.604 -proof-
  26.605 -  assume H: "c > 0"
  26.606 -  hence "c*x <= 0 = (0/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
  26.607 -  also have "\<dots> = (0 >= x)" by simp
  26.608 -  finally show  "(c*x <= 0) == (x <= 0)" by simp
  26.609 -qed
  26.610 -
  26.611 -lemma neg_prod_sum_le: "(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x + t <= 0) == (x >= (- 1/c)*t))"
  26.612 -proof-
  26.613 -  assume H: "c < 0"
  26.614 -  have "c*x + t <= 0 = (c*x <= -t)"  by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
  26.615 -  also have "\<dots> = (-t/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
  26.616 -  also have "\<dots> = ((- 1/c)*t <= x)" by simp
  26.617 -  finally show  "(c*x + t <= 0) == (x >= (- 1/c)*t)" by simp
  26.618 -qed
  26.619 -
  26.620 -lemma pos_prod_sum_le:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x + t <= 0) == (x <= (- 1/c)*t))"
  26.621 -proof-
  26.622 -  assume H: "c > 0"
  26.623 -  have "c*x + t <= 0 = (c*x <= -t)" by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
  26.624 -  also have "\<dots> = (-t/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
  26.625 -  also have "\<dots> = ((- 1/c)*t >= x)" by simp
  26.626 -  finally show  "(c*x + t <= 0) == (x <= (- 1/c)*t)" by simp
  26.627 -qed
  26.628 -
  26.629 -lemma sum_le:"((x::'a::pordered_ab_group_add) + t <= 0) == (x <= - t)"
  26.630 -  using le_diff_eq[where a= x and b=t and c=0] by simp
  26.631 -
  26.632 -lemma nz_prod_eq:"(c\<Colon>'a\<Colon>ordered_field) \<noteq> 0 \<Longrightarrow> ((c*x = 0) == (x = 0))" by simp
  26.633 -lemma nz_prod_sum_eq: "(c\<Colon>'a\<Colon>ordered_field) \<noteq> 0 \<Longrightarrow> ((c*x + t = 0) == (x = (- 1/c)*t))"
  26.634 -proof-
  26.635 -  assume H: "c \<noteq> 0"
  26.636 -  have "c*x + t = 0 = (c*x = -t)" by (subst eq_iff_diff_eq_0 [of "c*x" "-t"], simp)
  26.637 -  also have "\<dots> = (x = -t/c)" by (simp only: nonzero_eq_divide_eq[OF H] ring_simps)
  26.638 -  finally show  "(c*x + t = 0) == (x = (- 1/c)*t)" by simp
  26.639 -qed
  26.640 -lemma sum_eq:"((x::'a::pordered_ab_group_add) + t = 0) == (x = - t)"
  26.641 -  using eq_diff_eq[where a= x and b=t and c=0] by simp
  26.642 -
  26.643 -
  26.644 -interpretation class_ordered_field_dense_linear_order: constr_dense_linear_order
  26.645 - ["op <=" "op <"
  26.646 -   "\<lambda> x y. 1/2 * ((x::'a::{ordered_field,recpower,number_ring}) + y)"]
  26.647 -proof (unfold_locales, dlo, dlo, auto)
  26.648 -  fix x y::'a assume lt: "x < y"
  26.649 -  from  less_half_sum[OF lt] show "x < (x + y) /2" by simp
  26.650 -next
  26.651 -  fix x y::'a assume lt: "x < y"
  26.652 -  from  gt_half_sum[OF lt] show "(x + y) /2 < y" by simp
  26.653 -qed
  26.654 -
  26.655 -declaration{*
  26.656 -let
  26.657 -fun earlier [] x y = false
  26.658 -        | earlier (h::t) x y =
  26.659 -    if h aconvc y then false else if h aconvc x then true else earlier t x y;
  26.660 -
  26.661 -fun dest_frac ct = case term_of ct of
  26.662 -   Const (@{const_name "HOL.divide"},_) $ a $ b=>
  26.663 -    Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b))
  26.664 - | t => Rat.rat_of_int (snd (HOLogic.dest_number t))
  26.665 -
  26.666 -fun mk_frac phi cT x =
  26.667 - let val (a, b) = Rat.quotient_of_rat x
  26.668 - in if b = 1 then Numeral.mk_cnumber cT a
  26.669 -    else Thm.capply
  26.670 -         (Thm.capply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"})
  26.671 -                     (Numeral.mk_cnumber cT a))
  26.672 -         (Numeral.mk_cnumber cT b)
  26.673 - end
  26.674 -
  26.675 -fun whatis x ct = case term_of ct of
  26.676 -  Const(@{const_name "HOL.plus"}, _)$(Const(@{const_name "HOL.times"},_)$_$y)$_ =>
  26.677 -     if y aconv term_of x then ("c*x+t",[(funpow 2 Thm.dest_arg1) ct, Thm.dest_arg ct])
  26.678 -     else ("Nox",[])
  26.679 -| Const(@{const_name "HOL.plus"}, _)$y$_ =>
  26.680 -     if y aconv term_of x then ("x+t",[Thm.dest_arg ct])
  26.681 -     else ("Nox",[])
  26.682 -| Const(@{const_name "HOL.times"}, _)$_$y =>
  26.683 -     if y aconv term_of x then ("c*x",[Thm.dest_arg1 ct])
  26.684 -     else ("Nox",[])
  26.685 -| t => if t aconv term_of x then ("x",[]) else ("Nox",[]);
  26.686 -
  26.687 -fun xnormalize_conv ctxt [] ct = reflexive ct
  26.688 -| xnormalize_conv ctxt (vs as (x::_)) ct =
  26.689 -   case term_of ct of
  26.690 -   Const(@{const_name HOL.less},_)$_$Const(@{const_name "HOL.zero"},_) =>
  26.691 -    (case whatis x (Thm.dest_arg1 ct) of
  26.692 -    ("c*x+t",[c,t]) =>
  26.693 -       let
  26.694 -        val cr = dest_frac c
  26.695 -        val clt = Thm.dest_fun2 ct
  26.696 -        val cz = Thm.dest_arg ct
  26.697 -        val neg = cr </ Rat.zero
  26.698 -        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
  26.699 -               (Thm.capply @{cterm "Trueprop"}
  26.700 -                  (if neg then Thm.capply (Thm.capply clt c) cz
  26.701 -                    else Thm.capply (Thm.capply clt cz) c))
  26.702 -        val cth = equal_elim (symmetric cthp) TrueI
  26.703 -        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x,t])
  26.704 -             (if neg then @{thm neg_prod_sum_lt} else @{thm pos_prod_sum_lt})) cth
  26.705 -        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
  26.706 -                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
  26.707 -      in rth end
  26.708 -    | ("x+t",[t]) =>
  26.709 -       let
  26.710 -        val T = ctyp_of_term x
  26.711 -        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_lt"}
  26.712 -        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
  26.713 -              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
  26.714 -       in  rth end
  26.715 -    | ("c*x",[c]) =>
  26.716 -       let
  26.717 -        val cr = dest_frac c
  26.718 -        val clt = Thm.dest_fun2 ct
  26.719 -        val cz = Thm.dest_arg ct
  26.720 -        val neg = cr </ Rat.zero
  26.721 -        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
  26.722 -               (Thm.capply @{cterm "Trueprop"}
  26.723 -                  (if neg then Thm.capply (Thm.capply clt c) cz
  26.724 -                    else Thm.capply (Thm.capply clt cz) c))
  26.725 -        val cth = equal_elim (symmetric cthp) TrueI
  26.726 -        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
  26.727 -             (if neg then @{thm neg_prod_lt} else @{thm pos_prod_lt})) cth
  26.728 -        val rth = th
  26.729 -      in rth end
  26.730 -    | _ => reflexive ct)
  26.731 -
  26.732 -
  26.733 -|  Const(@{const_name HOL.less_eq},_)$_$Const(@{const_name "HOL.zero"},_) =>
  26.734 -   (case whatis x (Thm.dest_arg1 ct) of
  26.735 -    ("c*x+t",[c,t]) =>
  26.736 -       let
  26.737 -        val T = ctyp_of_term x
  26.738 -        val cr = dest_frac c
  26.739 -        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
  26.740 -        val cz = Thm.dest_arg ct
  26.741 -        val neg = cr </ Rat.zero
  26.742 -        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
  26.743 -               (Thm.capply @{cterm "Trueprop"}
  26.744 -                  (if neg then Thm.capply (Thm.capply clt c) cz
  26.745 -                    else Thm.capply (Thm.capply clt cz) c))
  26.746 -        val cth = equal_elim (symmetric cthp) TrueI
  26.747 -        val th = implies_elim (instantiate' [SOME T] (map SOME [c,x,t])
  26.748 -             (if neg then @{thm neg_prod_sum_le} else @{thm pos_prod_sum_le})) cth
  26.749 -        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
  26.750 -                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
  26.751 -      in rth end
  26.752 -    | ("x+t",[t]) =>
  26.753 -       let
  26.754 -        val T = ctyp_of_term x
  26.755 -        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_le"}
  26.756 -        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
  26.757 -              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
  26.758 -       in  rth end
  26.759 -    | ("c*x",[c]) =>
  26.760 -       let
  26.761 -        val T = ctyp_of_term x
  26.762 -        val cr = dest_frac c
  26.763 -        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
  26.764 -        val cz = Thm.dest_arg ct
  26.765 -        val neg = cr </ Rat.zero
  26.766 -        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
  26.767 -               (Thm.capply @{cterm "Trueprop"}
  26.768 -                  (if neg then Thm.capply (Thm.capply clt c) cz
  26.769 -                    else Thm.capply (Thm.capply clt cz) c))
  26.770 -        val cth = equal_elim (symmetric cthp) TrueI
  26.771 -        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
  26.772 -             (if neg then @{thm neg_prod_le} else @{thm pos_prod_le})) cth
  26.773 -        val rth = th
  26.774 -      in rth end
  26.775 -    | _ => reflexive ct)
  26.776 -
  26.777 -|  Const("op =",_)$_$Const(@{const_name "HOL.zero"},_) =>
  26.778 -   (case whatis x (Thm.dest_arg1 ct) of
  26.779 -    ("c*x+t",[c,t]) =>
  26.780 -       let
  26.781 -        val T = ctyp_of_term x
  26.782 -        val cr = dest_frac c
  26.783 -        val ceq = Thm.dest_fun2 ct
  26.784 -        val cz = Thm.dest_arg ct
  26.785 -        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
  26.786 -            (Thm.capply @{cterm "Trueprop"}
  26.787 -             (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
  26.788 -        val cth = equal_elim (symmetric cthp) TrueI
  26.789 -        val th = implies_elim
  26.790 -                 (instantiate' [SOME T] (map SOME [c,x,t]) @{thm nz_prod_sum_eq}) cth
  26.791 -        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
  26.792 -                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
  26.793 -      in rth end
  26.794 -    | ("x+t",[t]) =>
  26.795 -       let
  26.796 -        val T = ctyp_of_term x
  26.797 -        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_eq"}
  26.798 -        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
  26.799 -              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
  26.800 -       in  rth end
  26.801 -    | ("c*x",[c]) =>
  26.802 -       let
  26.803 -        val T = ctyp_of_term x
  26.804 -        val cr = dest_frac c
  26.805 -        val ceq = Thm.dest_fun2 ct
  26.806 -        val cz = Thm.dest_arg ct
  26.807 -        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
  26.808 -            (Thm.capply @{cterm "Trueprop"}
  26.809 -             (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
  26.810 -        val cth = equal_elim (symmetric cthp) TrueI
  26.811 -        val rth = implies_elim
  26.812 -                 (instantiate' [SOME T] (map SOME [c,x]) @{thm nz_prod_eq}) cth
  26.813 -      in rth end
  26.814 -    | _ => reflexive ct);
  26.815 -
  26.816 -local
  26.817 -  val less_iff_diff_less_0 = mk_meta_eq @{thm "less_iff_diff_less_0"}
  26.818 -  val le_iff_diff_le_0 = mk_meta_eq @{thm "le_iff_diff_le_0"}
  26.819 -  val eq_iff_diff_eq_0 = mk_meta_eq @{thm "eq_iff_diff_eq_0"}
  26.820 -in
  26.821 -fun field_isolate_conv phi ctxt vs ct = case term_of ct of
  26.822 -  Const(@{const_name HOL.less},_)$a$b =>
  26.823 -   let val (ca,cb) = Thm.dest_binop ct
  26.824 -       val T = ctyp_of_term ca
  26.825 -       val th = instantiate' [SOME T] [SOME ca, SOME cb] less_iff_diff_less_0
  26.826 -       val nth = Conv.fconv_rule
  26.827 -         (Conv.arg_conv (Conv.arg1_conv
  26.828 -              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
  26.829 -       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
  26.830 -   in rth end
  26.831 -| Const(@{const_name HOL.less_eq},_)$a$b =>
  26.832 -   let val (ca,cb) = Thm.dest_binop ct
  26.833 -       val T = ctyp_of_term ca
  26.834 -       val th = instantiate' [SOME T] [SOME ca, SOME cb] le_iff_diff_le_0
  26.835 -       val nth = Conv.fconv_rule
  26.836 -         (Conv.arg_conv (Conv.arg1_conv
  26.837 -              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
  26.838 -       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
  26.839 -   in rth end
  26.840 -
  26.841 -| Const("op =",_)$a$b =>
  26.842 -   let val (ca,cb) = Thm.dest_binop ct
  26.843 -       val T = ctyp_of_term ca
  26.844 -       val th = instantiate' [SOME T] [SOME ca, SOME cb] eq_iff_diff_eq_0
  26.845 -       val nth = Conv.fconv_rule
  26.846 -         (Conv.arg_conv (Conv.arg1_conv
  26.847 -              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
  26.848 -       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
  26.849 -   in rth end
  26.850 -| @{term "Not"} $(Const("op =",_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
  26.851 -| _ => reflexive ct
  26.852 -end;
  26.853 -
  26.854 -fun classfield_whatis phi =
  26.855 - let
  26.856 -  fun h x t =
  26.857 -   case term_of t of
  26.858 -     Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
  26.859 -                            else Ferrante_Rackoff_Data.Nox
  26.860 -   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
  26.861 -                            else Ferrante_Rackoff_Data.Nox
  26.862 -   | Const(@{const_name HOL.less},_)$y$z =>
  26.863 -       if term_of x aconv y then Ferrante_Rackoff_Data.Lt
  26.864 -        else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
  26.865 -        else Ferrante_Rackoff_Data.Nox
  26.866 -   | Const (@{const_name HOL.less_eq},_)$y$z =>
  26.867 -         if term_of x aconv y then Ferrante_Rackoff_Data.Le
  26.868 -         else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
  26.869 -         else Ferrante_Rackoff_Data.Nox
  26.870 -   | _ => Ferrante_Rackoff_Data.Nox
  26.871 - in h end;
  26.872 -fun class_field_ss phi =
  26.873 -   HOL_basic_ss addsimps ([@{thm "linorder_not_less"}, @{thm "linorder_not_le"}])
  26.874 -   addsplits [@{thm "abs_split"},@{thm "split_max"}, @{thm "split_min"}]
  26.875 -
  26.876 -in
  26.877 -Ferrante_Rackoff_Data.funs @{thm "class_ordered_field_dense_linear_order.ferrack_axiom"}
  26.878 -  {isolate_conv = field_isolate_conv, whatis = classfield_whatis, simpset = class_field_ss}
  26.879 -end
  26.880 -*}
  26.881 -
  26.882 -
  26.883 -end 
    27.1 --- a/src/HOL/Library/Library.thy	Mon Dec 29 13:23:53 2008 +0100
    27.2 +++ b/src/HOL/Library/Library.thy	Mon Dec 29 14:08:08 2008 +0100
    27.3 @@ -16,7 +16,6 @@
    27.4    Continuity
    27.5    ContNotDenum
    27.6    Countable
    27.7 -  Dense_Linear_Order
    27.8    Efficient_Nat
    27.9    Enum
   27.10    Eval_Witness
    28.1 --- a/src/HOL/Lim.thy	Mon Dec 29 13:23:53 2008 +0100
    28.2 +++ b/src/HOL/Lim.thy	Mon Dec 29 14:08:08 2008 +0100
    28.3 @@ -7,7 +7,7 @@
    28.4  header{* Limits and Continuity *}
    28.5  
    28.6  theory Lim
    28.7 -imports "~~/src/HOL/Hyperreal/SEQ"
    28.8 +imports SEQ
    28.9  begin
   28.10  
   28.11  text{*Standard Definitions*}
    29.1 --- a/src/HOL/PReal.thy	Mon Dec 29 13:23:53 2008 +0100
    29.2 +++ b/src/HOL/PReal.thy	Mon Dec 29 14:08:08 2008 +0100
    29.3 @@ -9,7 +9,7 @@
    29.4  header {* Positive real numbers *}
    29.5  
    29.6  theory PReal
    29.7 -imports Rational "~~/src/HOL/Library/Dense_Linear_Order"
    29.8 +imports Rational Dense_Linear_Order
    29.9  begin
   29.10  
   29.11  text{*Could be generalized and moved to @{text Ring_and_Field}*}
    30.1 --- a/src/HOL/Real.thy	Mon Dec 29 13:23:53 2008 +0100
    30.2 +++ b/src/HOL/Real.thy	Mon Dec 29 14:08:08 2008 +0100
    30.3 @@ -1,5 +1,5 @@
    30.4  theory Real
    30.5 -imports RComplete "~~/src/HOL/Real/RealVector"
    30.6 +imports RComplete RealVector
    30.7  begin
    30.8  
    30.9  end
    31.1 --- a/src/HOL/Real/HahnBanach/Bounds.thy	Mon Dec 29 13:23:53 2008 +0100
    31.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    31.3 @@ -1,83 +0,0 @@
    31.4 -(*  Title:      HOL/Real/HahnBanach/Bounds.thy
    31.5 -    ID:         $Id$
    31.6 -    Author:     Gertrud Bauer, TU Munich
    31.7 -*)
    31.8 -
    31.9 -header {* Bounds *}
   31.10 -
   31.11 -theory Bounds
   31.12 -imports Main ContNotDenum
   31.13 -begin
   31.14 -
   31.15 -locale lub =
   31.16 -  fixes A and x
   31.17 -  assumes least [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> a \<le> b) \<Longrightarrow> x \<le> b"
   31.18 -    and upper [intro?]: "a \<in> A \<Longrightarrow> a \<le> x"
   31.19 -
   31.20 -lemmas [elim?] = lub.least lub.upper
   31.21 -
   31.22 -definition
   31.23 -  the_lub :: "'a::order set \<Rightarrow> 'a" where
   31.24 -  "the_lub A = The (lub A)"
   31.25 -
   31.26 -notation (xsymbols)
   31.27 -  the_lub  ("\<Squnion>_" [90] 90)
   31.28 -
   31.29 -lemma the_lub_equality [elim?]:
   31.30 -  assumes "lub A x"
   31.31 -  shows "\<Squnion>A = (x::'a::order)"
   31.32 -proof -
   31.33 -  interpret lub [A x] by fact
   31.34 -  show ?thesis
   31.35 -  proof (unfold the_lub_def)
   31.36 -    from `lub A x` show "The (lub A) = x"
   31.37 -    proof
   31.38 -      fix x' assume lub': "lub A x'"
   31.39 -      show "x' = x"
   31.40 -      proof (rule order_antisym)
   31.41 -	from lub' show "x' \<le> x"
   31.42 -	proof
   31.43 -          fix a assume "a \<in> A"
   31.44 -          then show "a \<le> x" ..
   31.45 -	qed
   31.46 -	show "x \<le> x'"
   31.47 -	proof
   31.48 -          fix a assume "a \<in> A"
   31.49 -          with lub' show "a \<le> x'" ..
   31.50 -	qed
   31.51 -      qed
   31.52 -    qed
   31.53 -  qed
   31.54 -qed
   31.55 -
   31.56 -lemma the_lubI_ex:
   31.57 -  assumes ex: "\<exists>x. lub A x"
   31.58 -  shows "lub A (\<Squnion>A)"
   31.59 -proof -
   31.60 -  from ex obtain x where x: "lub A x" ..
   31.61 -  also from x have [symmetric]: "\<Squnion>A = x" ..
   31.62 -  finally show ?thesis .
   31.63 -qed
   31.64 -
   31.65 -lemma lub_compat: "lub A x = isLub UNIV A x"
   31.66 -proof -
   31.67 -  have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)"
   31.68 -    by (rule ext) (simp only: isUb_def)
   31.69 -  then show ?thesis
   31.70 -    by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast
   31.71 -qed
   31.72 -
   31.73 -lemma real_complete:
   31.74 -  fixes A :: "real set"
   31.75 -  assumes nonempty: "\<exists>a. a \<in> A"
   31.76 -    and ex_upper: "\<exists>y. \<forall>a \<in> A. a \<le> y"
   31.77 -  shows "\<exists>x. lub A x"
   31.78 -proof -
   31.79 -  from ex_upper have "\<exists>y. isUb UNIV A y"
   31.80 -    unfolding isUb_def setle_def by blast
   31.81 -  with nonempty have "\<exists>x. isLub UNIV A x"
   31.82 -    by (rule reals_complete)
   31.83 -  then show ?thesis by (simp only: lub_compat)
   31.84 -qed
   31.85 -
   31.86 -end
    32.1 --- a/src/HOL/Real/HahnBanach/FunctionNorm.thy	Mon Dec 29 13:23:53 2008 +0100
    32.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    32.3 @@ -1,279 +0,0 @@
    32.4 -(*  Title:      HOL/Real/HahnBanach/FunctionNorm.thy
    32.5 -    ID:         $Id$
    32.6 -    Author:     Gertrud Bauer, TU Munich
    32.7 -*)
    32.8 -
    32.9 -header {* The norm of a function *}
   32.10 -
   32.11 -theory FunctionNorm
   32.12 -imports NormedSpace FunctionOrder
   32.13 -begin
   32.14 -
   32.15 -subsection {* Continuous linear forms*}
   32.16 -
   32.17 -text {*
   32.18 -  A linear form @{text f} on a normed vector space @{text "(V, \<parallel>\<cdot>\<parallel>)"}
   32.19 -  is \emph{continuous}, iff it is bounded, i.e.
   32.20 -  \begin{center}
   32.21 -  @{text "\<exists>c \<in> R. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
   32.22 -  \end{center}
   32.23 -  In our application no other functions than linear forms are
   32.24 -  considered, so we can define continuous linear forms as bounded
   32.25 -  linear forms:
   32.26 -*}
   32.27 -
   32.28 -locale continuous = var V + norm_syntax + linearform +
   32.29 -  assumes bounded: "\<exists>c. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
   32.30 -
   32.31 -declare continuous.intro [intro?] continuous_axioms.intro [intro?]
   32.32 -
   32.33 -lemma continuousI [intro]:
   32.34 -  fixes norm :: "_ \<Rightarrow> real"  ("\<parallel>_\<parallel>")
   32.35 -  assumes "linearform V f"
   32.36 -  assumes r: "\<And>x. x \<in> V \<Longrightarrow> \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
   32.37 -  shows "continuous V norm f"
   32.38 -proof
   32.39 -  show "linearform V f" by fact
   32.40 -  from r have "\<exists>c. \<forall>x\<in>V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" by blast
   32.41 -  then show "continuous_axioms V norm f" ..
   32.42 -qed
   32.43 -
   32.44 -
   32.45 -subsection {* The norm of a linear form *}
   32.46 -
   32.47 -text {*
   32.48 -  The least real number @{text c} for which holds
   32.49 -  \begin{center}
   32.50 -  @{text "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
   32.51 -  \end{center}
   32.52 -  is called the \emph{norm} of @{text f}.
   32.53 -
   32.54 -  For non-trivial vector spaces @{text "V \<noteq> {0}"} the norm can be
   32.55 -  defined as
   32.56 -  \begin{center}
   32.57 -  @{text "\<parallel>f\<parallel> = \<sup>x \<noteq> 0. \<bar>f x\<bar> / \<parallel>x\<parallel>"}
   32.58 -  \end{center}
   32.59 -
   32.60 -  For the case @{text "V = {0}"} the supremum would be taken from an
   32.61 -  empty set. Since @{text \<real>} is unbounded, there would be no supremum.
   32.62 -  To avoid this situation it must be guaranteed that there is an
   32.63 -  element in this set. This element must be @{text "{} \<ge> 0"} so that
   32.64 -  @{text fn_norm} has the norm properties. Furthermore it does not
   32.65 -  have to change the norm in all other cases, so it must be @{text 0},
   32.66 -  as all other elements are @{text "{} \<ge> 0"}.
   32.67 -
   32.68 -  Thus we define the set @{text B} where the supremum is taken from as
   32.69 -  follows:
   32.70 -  \begin{center}
   32.71 -  @{text "{0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel>. x \<noteq> 0 \<and> x \<in> F}"}
   32.72 -  \end{center}
   32.73 -
   32.74 -  @{text fn_norm} is equal to the supremum of @{text B}, if the
   32.75 -  supremum exists (otherwise it is undefined).
   32.76 -*}
   32.77 -
   32.78 -locale fn_norm = norm_syntax +
   32.79 -  fixes B defines "B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
   32.80 -  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
   32.81 -  defines "\<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
   32.82 -
   32.83 -locale normed_vectorspace_with_fn_norm = normed_vectorspace + fn_norm
   32.84 -
   32.85 -lemma (in fn_norm) B_not_empty [intro]: "0 \<in> B V f"
   32.86 -  by (simp add: B_def)
   32.87 -
   32.88 -text {*
   32.89 -  The following lemma states that every continuous linear form on a
   32.90 -  normed space @{text "(V, \<parallel>\<cdot>\<parallel>)"} has a function norm.
   32.91 -*}
   32.92 -
   32.93 -lemma (in normed_vectorspace_with_fn_norm) fn_norm_works:
   32.94 -  assumes "continuous V norm f"
   32.95 -  shows "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
   32.96 -proof -
   32.97 -  interpret continuous [V norm f] by fact
   32.98 -  txt {* The existence of the supremum is shown using the
   32.99 -    completeness of the reals. Completeness means, that every
  32.100 -    non-empty bounded set of reals has a supremum. *}
  32.101 -  have "\<exists>a. lub (B V f) a"
  32.102 -  proof (rule real_complete)
  32.103 -    txt {* First we have to show that @{text B} is non-empty: *}
  32.104 -    have "0 \<in> B V f" ..
  32.105 -    then show "\<exists>x. x \<in> B V f" ..
  32.106 -
  32.107 -    txt {* Then we have to show that @{text B} is bounded: *}
  32.108 -    show "\<exists>c. \<forall>y \<in> B V f. y \<le> c"
  32.109 -    proof -
  32.110 -      txt {* We know that @{text f} is bounded by some value @{text c}. *}
  32.111 -      from bounded obtain c where c: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
  32.112 -
  32.113 -      txt {* To prove the thesis, we have to show that there is some
  32.114 -        @{text b}, such that @{text "y \<le> b"} for all @{text "y \<in>
  32.115 -        B"}. Due to the definition of @{text B} there are two cases. *}
  32.116 -
  32.117 -      def b \<equiv> "max c 0"
  32.118 -      have "\<forall>y \<in> B V f. y \<le> b"
  32.119 -      proof
  32.120 -        fix y assume y: "y \<in> B V f"
  32.121 -        show "y \<le> b"
  32.122 -        proof cases
  32.123 -          assume "y = 0"
  32.124 -          then show ?thesis unfolding b_def by arith
  32.125 -        next
  32.126 -          txt {* The second case is @{text "y = \<bar>f x\<bar> / \<parallel>x\<parallel>"} for some
  32.127 -            @{text "x \<in> V"} with @{text "x \<noteq> 0"}. *}
  32.128 -          assume "y \<noteq> 0"
  32.129 -          with y obtain x where y_rep: "y = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
  32.130 -              and x: "x \<in> V" and neq: "x \<noteq> 0"
  32.131 -            by (auto simp add: B_def real_divide_def)
  32.132 -          from x neq have gt: "0 < \<parallel>x\<parallel>" ..
  32.133 -
  32.134 -          txt {* The thesis follows by a short calculation using the
  32.135 -            fact that @{text f} is bounded. *}
  32.136 -
  32.137 -          note y_rep
  32.138 -          also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
  32.139 -          proof (rule mult_right_mono)
  32.140 -            from c x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
  32.141 -            from gt have "0 < inverse \<parallel>x\<parallel>" 
  32.142 -              by (rule positive_imp_inverse_positive)
  32.143 -            then show "0 \<le> inverse \<parallel>x\<parallel>" by (rule order_less_imp_le)
  32.144 -          qed
  32.145 -          also have "\<dots> = c * (\<parallel>x\<parallel> * inverse \<parallel>x\<parallel>)"
  32.146 -            by (rule real_mult_assoc)
  32.147 -          also
  32.148 -          from gt have "\<parallel>x\<parallel> \<noteq> 0" by simp
  32.149 -          then have "\<parallel>x\<parallel> * inverse \<parallel>x\<parallel> = 1" by simp 
  32.150 -          also have "c * 1 \<le> b" by (simp add: b_def le_maxI1)
  32.151 -          finally show "y \<le> b" .
  32.152 -        qed
  32.153 -      qed
  32.154 -      then show ?thesis ..
  32.155 -    qed
  32.156 -  qed
  32.157 -  then show ?thesis unfolding fn_norm_def by (rule the_lubI_ex)
  32.158 -qed
  32.159 -
  32.160 -lemma (in normed_vectorspace_with_fn_norm) fn_norm_ub [iff?]:
  32.161 -  assumes "continuous V norm f"
  32.162 -  assumes b: "b \<in> B V f"
  32.163 -  shows "b \<le> \<parallel>f\<parallel>\<hyphen>V"
  32.164 -proof -
  32.165 -  interpret continuous [V norm f] by fact
  32.166 -  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
  32.167 -    using `continuous V norm f` by (rule fn_norm_works)
  32.168 -  from this and b show ?thesis ..
  32.169 -qed
  32.170 -
  32.171 -lemma (in normed_vectorspace_with_fn_norm) fn_norm_leastB:
  32.172 -  assumes "continuous V norm f"
  32.173 -  assumes b: "\<And>b. b \<in> B V f \<Longrightarrow> b \<le> y"
  32.174 -  shows "\<parallel>f\<parallel>\<hyphen>V \<le> y"
  32.175 -proof -
  32.176 -  interpret continuous [V norm f] by fact
  32.177 -  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
  32.178 -    using `continuous V norm f` by (rule fn_norm_works)
  32.179 -  from this and b show ?thesis ..
  32.180 -qed
  32.181 -
  32.182 -text {* The norm of a continuous function is always @{text "\<ge> 0"}. *}
  32.183 -
  32.184 -lemma (in normed_vectorspace_with_fn_norm) fn_norm_ge_zero [iff]:
  32.185 -  assumes "continuous V norm f"
  32.186 -  shows "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
  32.187 -proof -
  32.188 -  interpret continuous [V norm f] by fact
  32.189 -  txt {* The function norm is defined as the supremum of @{text B}.
  32.190 -    So it is @{text "\<ge> 0"} if all elements in @{text B} are @{text "\<ge>
  32.191 -    0"}, provided the supremum exists and @{text B} is not empty. *}
  32.192 -  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
  32.193 -    using `continuous V norm f` by (rule fn_norm_works)
  32.194 -  moreover have "0 \<in> B V f" ..
  32.195 -  ultimately show ?thesis ..
  32.196 -qed
  32.197 -
  32.198 -text {*
  32.199 -  \medskip The fundamental property of function norms is:
  32.200 -  \begin{center}
  32.201 -  @{text "\<bar>f x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
  32.202 -  \end{center}
  32.203 -*}
  32.204 -
  32.205 -lemma (in normed_vectorspace_with_fn_norm) fn_norm_le_cong:
  32.206 -  assumes "continuous V norm f" "linearform V f"
  32.207 -  assumes x: "x \<in> V"
  32.208 -  shows "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
  32.209 -proof -
  32.210 -  interpret continuous [V norm f] by fact
  32.211 -  interpret linearform [V f] .
  32.212 -  show ?thesis
  32.213 -  proof cases
  32.214 -    assume "x = 0"
  32.215 -    then have "\<bar>f x\<bar> = \<bar>f 0\<bar>" by simp
  32.216 -    also have "f 0 = 0" by rule unfold_locales
  32.217 -    also have "\<bar>\<dots>\<bar> = 0" by simp
  32.218 -    also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
  32.219 -      using `continuous V norm f` by (rule fn_norm_ge_zero)
  32.220 -    from x have "0 \<le> norm x" ..
  32.221 -    with a have "0 \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" by (simp add: zero_le_mult_iff)
  32.222 -    finally show "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" .
  32.223 -  next
  32.224 -    assume "x \<noteq> 0"
  32.225 -    with x have neq: "\<parallel>x\<parallel> \<noteq> 0" by simp
  32.226 -    then have "\<bar>f x\<bar> = (\<bar>f x\<bar> * inverse \<parallel>x\<parallel>) * \<parallel>x\<parallel>" by simp
  32.227 -    also have "\<dots> \<le>  \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
  32.228 -    proof (rule mult_right_mono)
  32.229 -      from x show "0 \<le> \<parallel>x\<parallel>" ..
  32.230 -      from x and neq have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<in> B V f"
  32.231 -	by (auto simp add: B_def real_divide_def)
  32.232 -      with `continuous V norm f` show "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>V"
  32.233 -	by (rule fn_norm_ub)
  32.234 -    qed
  32.235 -    finally show ?thesis .
  32.236 -  qed
  32.237 -qed
  32.238 -
  32.239 -text {*
  32.240 -  \medskip The function norm is the least positive real number for
  32.241 -  which the following inequation holds:
  32.242 -  \begin{center}
  32.243 -    @{text "\<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
  32.244 -  \end{center}
  32.245 -*}
  32.246 -
  32.247 -lemma (in normed_vectorspace_with_fn_norm) fn_norm_least [intro?]:
  32.248 -  assumes "continuous V norm f"
  32.249 -  assumes ineq: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" and ge: "0 \<le> c"
  32.250 -  shows "\<parallel>f\<parallel>\<hyphen>V \<le> c"
  32.251 -proof -
  32.252 -  interpret continuous [V norm f] by fact
  32.253 -  show ?thesis
  32.254 -  proof (rule fn_norm_leastB [folded B_def fn_norm_def])
  32.255 -    fix b assume b: "b \<in> B V f"
  32.256 -    show "b \<le> c"
  32.257 -    proof cases
  32.258 -      assume "b = 0"
  32.259 -      with ge show ?thesis by simp
  32.260 -    next
  32.261 -      assume "b \<noteq> 0"
  32.262 -      with b obtain x where b_rep: "b = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
  32.263 -        and x_neq: "x \<noteq> 0" and x: "x \<in> V"
  32.264 -	by (auto simp add: B_def real_divide_def)
  32.265 -      note b_rep
  32.266 -      also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
  32.267 -      proof (rule mult_right_mono)
  32.268 -	have "0 < \<parallel>x\<parallel>" using x x_neq ..
  32.269 -	then show "0 \<le> inverse \<parallel>x\<parallel>" by simp
  32.270 -	from ineq and x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
  32.271 -      qed
  32.272 -      also have "\<dots> = c"
  32.273 -      proof -
  32.274 -	from x_neq and x have "\<parallel>x\<parallel> \<noteq> 0" by simp
  32.275 -	then show ?thesis by simp
  32.276 -      qed
  32.277 -      finally show ?thesis .
  32.278 -    qed
  32.279 -  qed (insert `continuous V norm f`, simp_all add: continuous_def)
  32.280 -qed
  32.281 -
  32.282 -end
    33.1 --- a/src/HOL/Real/HahnBanach/FunctionOrder.thy	Mon Dec 29 13:23:53 2008 +0100
    33.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    33.3 @@ -1,142 +0,0 @@
    33.4 -(*  Title:      HOL/Real/HahnBanach/FunctionOrder.thy
    33.5 -    ID:         $Id$
    33.6 -    Author:     Gertrud Bauer, TU Munich
    33.7 -*)
    33.8 -
    33.9 -header {* An order on functions *}
   33.10 -
   33.11 -theory FunctionOrder
   33.12 -imports Subspace Linearform
   33.13 -begin
   33.14 -
   33.15 -subsection {* The graph of a function *}
   33.16 -
   33.17 -text {*
   33.18 -  We define the \emph{graph} of a (real) function @{text f} with
   33.19 -  domain @{text F} as the set
   33.20 -  \begin{center}
   33.21 -  @{text "{(x, f x). x \<in> F}"}
   33.22 -  \end{center}
   33.23 -  So we are modeling partial functions by specifying the domain and
   33.24 -  the mapping function. We use the term ``function'' also for its
   33.25 -  graph.
   33.26 -*}
   33.27 -
   33.28 -types 'a graph = "('a \<times> real) set"
   33.29 -
   33.30 -definition
   33.31 -  graph :: "'a set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a graph" where
   33.32 -  "graph F f = {(x, f x) | x. x \<in> F}"
   33.33 -
   33.34 -lemma graphI [intro]: "x \<in> F \<Longrightarrow> (x, f x) \<in> graph F f"
   33.35 -  unfolding graph_def by blast
   33.36 -
   33.37 -lemma graphI2 [intro?]: "x \<in> F \<Longrightarrow> \<exists>t \<in> graph F f. t = (x, f x)"
   33.38 -  unfolding graph_def by blast
   33.39 -
   33.40 -lemma graphE [elim?]:
   33.41 -    "(x, y) \<in> graph F f \<Longrightarrow> (x \<in> F \<Longrightarrow> y = f x \<Longrightarrow> C) \<Longrightarrow> C"
   33.42 -  unfolding graph_def by blast
   33.43 -
   33.44 -
   33.45 -subsection {* Functions ordered by domain extension *}
   33.46 -
   33.47 -text {*
   33.48 -  A function @{text h'} is an extension of @{text h}, iff the graph of
   33.49 -  @{text h} is a subset of the graph of @{text h'}.
   33.50 -*}
   33.51 -
   33.52 -lemma graph_extI:
   33.53 -  "(\<And>x. x \<in> H \<Longrightarrow> h x = h' x) \<Longrightarrow> H \<subseteq> H'
   33.54 -    \<Longrightarrow> graph H h \<subseteq> graph H' h'"
   33.55 -  unfolding graph_def by blast
   33.56 -
   33.57 -lemma graph_extD1 [dest?]:
   33.58 -  "graph H h \<subseteq> graph H' h' \<Longrightarrow> x \<in> H \<Longrightarrow> h x = h' x"
   33.59 -  unfolding graph_def by blast
   33.60 -
   33.61 -lemma graph_extD2 [dest?]:
   33.62 -  "graph H h \<subseteq> graph H' h' \<Longrightarrow> H \<subseteq> H'"
   33.63 -  unfolding graph_def by blast
   33.64 -
   33.65 -
   33.66 -subsection {* Domain and function of a graph *}
   33.67 -
   33.68 -text {*
   33.69 -  The inverse functions to @{text graph} are @{text domain} and @{text
   33.70 -  funct}.
   33.71 -*}
   33.72 -
   33.73 -definition
   33.74 -  "domain" :: "'a graph \<Rightarrow> 'a set" where
   33.75 -  "domain g = {x. \<exists>y. (x, y) \<in> g}"
   33.76 -
   33.77 -definition
   33.78 -  funct :: "'a graph \<Rightarrow> ('a \<Rightarrow> real)" where
   33.79 -  "funct g = (\<lambda>x. (SOME y. (x, y) \<in> g))"
   33.80 -
   33.81 -text {*
   33.82 -  The following lemma states that @{text g} is the graph of a function
   33.83 -  if the relation induced by @{text g} is unique.
   33.84 -*}
   33.85 -
   33.86 -lemma graph_domain_funct:
   33.87 -  assumes uniq: "\<And>x y z. (x, y) \<in> g \<Longrightarrow> (x, z) \<in> g \<Longrightarrow> z = y"
   33.88 -  shows "graph (domain g) (funct g) = g"
   33.89 -  unfolding domain_def funct_def graph_def
   33.90 -proof auto  (* FIXME !? *)
   33.91 -  fix a b assume g: "(a, b) \<in> g"
   33.92 -  from g show "(a, SOME y. (a, y) \<in> g) \<in> g" by (rule someI2)
   33.93 -  from g show "\<exists>y. (a, y) \<in> g" ..
   33.94 -  from g show "b = (SOME y. (a, y) \<in> g)"
   33.95 -  proof (rule some_equality [symmetric])
   33.96 -    fix y assume "(a, y) \<in> g"
   33.97 -    with g show "y = b" by (rule uniq)
   33.98 -  qed
   33.99 -qed
  33.100 -
  33.101 -
  33.102 -subsection {* Norm-preserving extensions of a function *}
  33.103 -
  33.104 -text {*
  33.105 -  Given a linear form @{text f} on the space @{text F} and a seminorm
  33.106 -  @{text p} on @{text E}. The set of all linear extensions of @{text
  33.107 -  f}, to superspaces @{text H} of @{text F}, which are bounded by
  33.108 -  @{text p}, is defined as follows.
  33.109 -*}
  33.110 -
  33.111 -definition
  33.112 -  norm_pres_extensions ::
  33.113 -    "'a::{plus, minus, uminus, zero} set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> real)
  33.114 -      \<Rightarrow> 'a graph set" where
  33.115 -    "norm_pres_extensions E p F f
  33.116 -      = {g. \<exists>H h. g = graph H h
  33.117 -          \<and> linearform H h
  33.118 -          \<and> H \<unlhd> E
  33.119 -          \<and> F \<unlhd> H
  33.120 -          \<and> graph F f \<subseteq> graph H h
  33.121 -          \<and> (\<forall>x \<in> H. h x \<le> p x)}"
  33.122 -
  33.123 -lemma norm_pres_extensionE [elim]:
  33.124 -  "g \<in> norm_pres_extensions E p F f
  33.125 -  \<Longrightarrow> (\<And>H h. g = graph H h \<Longrightarrow> linearform H h
  33.126 -        \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H \<Longrightarrow> graph F f \<subseteq> graph H h
  33.127 -        \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x \<Longrightarrow> C) \<Longrightarrow> C"
  33.128 -  unfolding norm_pres_extensions_def by blast
  33.129 -
  33.130 -lemma norm_pres_extensionI2 [intro]:
  33.131 -  "linearform H h \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H
  33.132 -    \<Longrightarrow> graph F f \<subseteq> graph H h \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x
  33.133 -    \<Longrightarrow> graph H h \<in> norm_pres_extensions E p F f"
  33.134 -  unfolding norm_pres_extensions_def by blast
  33.135 -
  33.136 -lemma norm_pres_extensionI:  (* FIXME ? *)
  33.137 -  "\<exists>H h. g = graph H h
  33.138 -    \<and> linearform H h
  33.139 -    \<and> H \<unlhd> E
  33.140 -    \<and> F \<unlhd> H
  33.141 -    \<and> graph F f \<subseteq> graph H h
  33.142 -    \<and> (\<forall>x \<in> H. h x \<le> p x) \<Longrightarrow> g \<in> norm_pres_extensions E p F f"
  33.143 -  unfolding norm_pres_extensions_def by blast
  33.144 -
  33.145 -end
    34.1 --- a/src/HOL/Real/HahnBanach/HahnBanach.thy	Mon Dec 29 13:23:53 2008 +0100
    34.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    34.3 @@ -1,510 +0,0 @@
    34.4 -(*  Title:      HOL/Real/HahnBanach/HahnBanach.thy
    34.5 -    ID:         $Id$
    34.6 -    Author:     Gertrud Bauer, TU Munich
    34.7 -*)
    34.8 -
    34.9 -header {* The Hahn-Banach Theorem *}
   34.10 -
   34.11 -theory HahnBanach
   34.12 -imports HahnBanachLemmas
   34.13 -begin
   34.14 -
   34.15 -text {*
   34.16 -  We present the proof of two different versions of the Hahn-Banach
   34.17 -  Theorem, closely following \cite[\S36]{Heuser:1986}.
   34.18 -*}
   34.19 -
   34.20 -subsection {* The Hahn-Banach Theorem for vector spaces *}
   34.21 -
   34.22 -text {*
   34.23 -  \textbf{Hahn-Banach Theorem.} Let @{text F} be a subspace of a real
   34.24 -  vector space @{text E}, let @{text p} be a semi-norm on @{text E},
   34.25 -  and @{text f} be a linear form defined on @{text F} such that @{text
   34.26 -  f} is bounded by @{text p}, i.e.  @{text "\<forall>x \<in> F. f x \<le> p x"}.  Then
   34.27 -  @{text f} can be extended to a linear form @{text h} on @{text E}
   34.28 -  such that @{text h} is norm-preserving, i.e. @{text h} is also
   34.29 -  bounded by @{text p}.
   34.30 -
   34.31 -  \bigskip
   34.32 -  \textbf{Proof Sketch.}
   34.33 -  \begin{enumerate}
   34.34 -
   34.35 -  \item Define @{text M} as the set of norm-preserving extensions of
   34.36 -  @{text f} to subspaces of @{text E}. The linear forms in @{text M}
   34.37 -  are ordered by domain extension.
   34.38 -
   34.39 -  \item We show that every non-empty chain in @{text M} has an upper
   34.40 -  bound in @{text M}.
   34.41 -
   34.42 -  \item With Zorn's Lemma we conclude that there is a maximal function
   34.43 -  @{text g} in @{text M}.
   34.44 -
   34.45 -  \item The domain @{text H} of @{text g} is the whole space @{text
   34.46 -  E}, as shown by classical contradiction:
   34.47 -
   34.48 -  \begin{itemize}
   34.49 -
   34.50 -  \item Assuming @{text g} is not defined on whole @{text E}, it can
   34.51 -  still be extended in a norm-preserving way to a super-space @{text
   34.52 -  H'} of @{text H}.
   34.53 -
   34.54 -  \item Thus @{text g} can not be maximal. Contradiction!
   34.55 -
   34.56 -  \end{itemize}
   34.57 -  \end{enumerate}
   34.58 -*}
   34.59 -
   34.60 -theorem HahnBanach:
   34.61 -  assumes E: "vectorspace E" and "subspace F E"
   34.62 -    and "seminorm E p" and "linearform F f"
   34.63 -  assumes fp: "\<forall>x \<in> F. f x \<le> p x"
   34.64 -  shows "\<exists>h. linearform E h \<and> (\<forall>x \<in> F. h x = f x) \<and> (\<forall>x \<in> E. h x \<le> p x)"
   34.65 -    -- {* Let @{text E} be a vector space, @{text F} a subspace of @{text E}, @{text p} a seminorm on @{text E}, *}
   34.66 -    -- {* and @{text f} a linear form on @{text F} such that @{text f} is bounded by @{text p}, *}
   34.67 -    -- {* then @{text f} can be extended to a linear form @{text h} on @{text E} in a norm-preserving way. \skp *}
   34.68 -proof -
   34.69 -  interpret vectorspace [E] by fact
   34.70 -  interpret subspace [F E] by fact
   34.71 -  interpret seminorm [E p] by fact
   34.72 -  interpret linearform [F f] by fact
   34.73 -  def M \<equiv> "norm_pres_extensions E p F f"
   34.74 -  then have M: "M = \<dots>" by (simp only:)
   34.75 -  from E have F: "vectorspace F" ..
   34.76 -  note FE = `F \<unlhd> E`
   34.77 -  {
   34.78 -    fix c assume cM: "c \<in> chain M" and ex: "\<exists>x. x \<in> c"
   34.79 -    have "\<Union>c \<in> M"
   34.80 -      -- {* Show that every non-empty chain @{text c} of @{text M} has an upper bound in @{text M}: *}
   34.81 -      -- {* @{text "\<Union>c"} is greater than any element of the chain @{text c}, so it suffices to show @{text "\<Union>c \<in> M"}. *}
   34.82 -      unfolding M_def
   34.83 -    proof (rule norm_pres_extensionI)
   34.84 -      let ?H = "domain (\<Union>c)"
   34.85 -      let ?h = "funct (\<Union>c)"
   34.86 -
   34.87 -      have a: "graph ?H ?h = \<Union>c"
   34.88 -      proof (rule graph_domain_funct)
   34.89 -        fix x y z assume "(x, y) \<in> \<Union>c" and "(x, z) \<in> \<Union>c"
   34.90 -        with M_def cM show "z = y" by (rule sup_definite)
   34.91 -      qed
   34.92 -      moreover from M cM a have "linearform ?H ?h"
   34.93 -        by (rule sup_lf)
   34.94 -      moreover from a M cM ex FE E have "?H \<unlhd> E"
   34.95 -        by (rule sup_subE)
   34.96 -      moreover from a M cM ex FE have "F \<unlhd> ?H"
   34.97 -        by (rule sup_supF)
   34.98 -      moreover from a M cM ex have "graph F f \<subseteq> graph ?H ?h"
   34.99 -        by (rule sup_ext)
  34.100 -      moreover from a M cM have "\<forall>x \<in> ?H. ?h x \<le> p x"
  34.101 -        by (rule sup_norm_pres)
  34.102 -      ultimately show "\<exists>H h. \<Union>c = graph H h
  34.103 -          \<and> linearform H h
  34.104 -          \<and> H \<unlhd> E
  34.105 -          \<and> F \<unlhd> H
  34.106 -          \<and> graph F f \<subseteq> graph H h
  34.107 -          \<and> (\<forall>x \<in> H. h x \<le> p x)" by blast
  34.108 -    qed
  34.109 -  }
  34.110 -  then have "\<exists>g \<in> M. \<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
  34.111 -  -- {* With Zorn's Lemma we can conclude that there is a maximal element in @{text M}. \skp *}
  34.112 -  proof (rule Zorn's_Lemma)
  34.113 -      -- {* We show that @{text M} is non-empty: *}
  34.114 -    show "graph F f \<in> M"
  34.115 -      unfolding M_def
  34.116 -    proof (rule norm_pres_extensionI2)
  34.117 -      show "linearform F f" by fact
  34.118 -      show "F \<unlhd> E" by fact
  34.119 -      from F show "F \<unlhd> F" by (rule vectorspace.subspace_refl)
  34.120 -      show "graph F f \<subseteq> graph F f" ..
  34.121 -      show "\<forall>x\<in>F. f x \<le> p x" by fact
  34.122 -    qed
  34.123 -  qed
  34.124 -  then obtain g where gM: "g \<in> M" and gx: "\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
  34.125 -    by blast
  34.126 -  from gM obtain H h where
  34.127 -      g_rep: "g = graph H h"
  34.128 -    and linearform: "linearform H h"
  34.129 -    and HE: "H \<unlhd> E" and FH: "F \<unlhd> H"
  34.130 -    and graphs: "graph F f \<subseteq> graph H h"
  34.131 -    and hp: "\<forall>x \<in> H. h x \<le> p x" unfolding M_def ..
  34.132 -      -- {* @{text g} is a norm-preserving extension of @{text f}, in other words: *}
  34.133 -      -- {* @{text g} is the graph of some linear form @{text h} defined on a subspace @{text H} of @{text E}, *}
  34.134 -      -- {* and @{text h} is an extension of @{text f} that is again bounded by @{text p}. \skp *}
  34.135 -  from HE E have H: "vectorspace H"
  34.136 -    by (rule subspace.vectorspace)
  34.137 -
  34.138 -  have HE_eq: "H = E"
  34.139 -    -- {* We show that @{text h} is defined on whole @{text E} by classical contradiction. \skp *}
  34.140 -  proof (rule classical)
  34.141 -    assume neq: "H \<noteq> E"
  34.142 -      -- {* Assume @{text h} is not defined on whole @{text E}. Then show that @{text h} can be extended *}
  34.143 -      -- {* in a norm-preserving way to a function @{text h'} with the graph @{text g'}. \skp *}
  34.144 -    have "\<exists>g' \<in> M. g \<subseteq> g' \<and> g \<noteq> g'"
  34.145 -    proof -
  34.146 -      from HE have "H \<subseteq> E" ..
  34.147 -      with neq obtain x' where x'E: "x' \<in> E" and "x' \<notin> H" by blast
  34.148 -      obtain x': "x' \<noteq> 0"
  34.149 -      proof
  34.150 -        show "x' \<noteq> 0"
  34.151 -        proof
  34.152 -          assume "x' = 0"
  34.153 -          with H have "x' \<in> H" by (simp only: vectorspace.zero)
  34.154 -          with `x' \<notin> H` show False by contradiction
  34.155 -        qed
  34.156 -      qed
  34.157 -
  34.158 -      def H' \<equiv> "H + lin x'"
  34.159 -        -- {* Define @{text H'} as the direct sum of @{text H} and the linear closure of @{text x'}. \skp *}
  34.160 -      have HH': "H \<unlhd> H'"
  34.161 -      proof (unfold H'_def)
  34.162 -        from x'E have "vectorspace (lin x')" ..
  34.163 -        with H show "H \<unlhd> H + lin x'" ..
  34.164 -      qed
  34.165 -
  34.166 -      obtain xi where
  34.167 -        xi: "\<forall>y \<in> H. - p (y + x') - h y \<le> xi
  34.168 -          \<and> xi \<le> p (y + x') - h y"
  34.169 -        -- {* Pick a real number @{text \<xi>} that fulfills certain inequations; this will *}
  34.170 -        -- {* be used to establish that @{text h'} is a norm-preserving extension of @{text h}.
  34.171 -           \label{ex-xi-use}\skp *}
  34.172 -      proof -
  34.173 -        from H have "\<exists>xi. \<forall>y \<in> H. - p (y + x') - h y \<le> xi
  34.174 -            \<and> xi \<le> p (y + x') - h y"
  34.175 -        proof (rule ex_xi)
  34.176 -          fix u v assume u: "u \<in> H" and v: "v \<in> H"
  34.177 -          with HE have uE: "u \<in> E" and vE: "v \<in> E" by auto
  34.178 -          from H u v linearform have "h v - h u = h (v - u)"
  34.179 -            by (simp add: linearform.diff)
  34.180 -          also from hp and H u v have "\<dots> \<le> p (v - u)"
  34.181 -            by (simp only: vectorspace.diff_closed)
  34.182 -          also from x'E uE vE have "v - u = x' + - x' + v + - u"
  34.183 -            by (simp add: diff_eq1)
  34.184 -          also from x'E uE vE have "\<dots> = v + x' + - (u + x')"
  34.185 -            by (simp add: add_ac)
  34.186 -          also from x'E uE vE have "\<dots> = (v + x') - (u + x')"
  34.187 -            by (simp add: diff_eq1)
  34.188 -          also from x'E uE vE E have "p \<dots> \<le> p (v + x') + p (u + x')"
  34.189 -            by (simp add: diff_subadditive)
  34.190 -          finally have "h v - h u \<le> p (v + x') + p (u + x')" .
  34.191 -          then show "- p (u + x') - h u \<le> p (v + x') - h v" by simp
  34.192 -        qed
  34.193 -        then show thesis by (blast intro: that)
  34.194 -      qed
  34.195 -
  34.196 -      def h' \<equiv> "\<lambda>x. let (y, a) =
  34.197 -          SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H in h y + a * xi"
  34.198 -        -- {* Define the extension @{text h'} of @{text h} to @{text H'} using @{text \<xi>}. \skp *}
  34.199 -
  34.200 -      have "g \<subseteq> graph H' h' \<and> g \<noteq> graph H' h'"
  34.201 -        -- {* @{text h'} is an extension of @{text h} \dots \skp *}
  34.202 -      proof
  34.203 -        show "g \<subseteq> graph H' h'"
  34.204 -        proof -
  34.205 -          have  "graph H h \<subseteq> graph H' h'"
  34.206 -          proof (rule graph_extI)
  34.207 -            fix t assume t: "t \<in> H"
  34.208 -            from E HE t have "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
  34.209 -	      using `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` by (rule decomp_H'_H)
  34.210 -            with h'_def show "h t = h' t" by (simp add: Let_def)
  34.211 -          next
  34.212 -            from HH' show "H \<subseteq> H'" ..
  34.213 -          qed
  34.214 -          with g_rep show ?thesis by (simp only:)
  34.215 -        qed
  34.216 -
  34.217 -        show "g \<noteq> graph H' h'"
  34.218 -        proof -
  34.219 -          have "graph H h \<noteq> graph H' h'"
  34.220 -          proof
  34.221 -            assume eq: "graph H h = graph H' h'"
  34.222 -            have "x' \<in> H'"
  34.223 -	      unfolding H'_def
  34.224 -            proof
  34.225 -              from H show "0 \<in> H" by (rule vectorspace.zero)
  34.226 -              from x'E show "x' \<in> lin x'" by (rule x_lin_x)
  34.227 -              from x'E show "x' = 0 + x'" by simp
  34.228 -            qed
  34.229 -            then have "(x', h' x') \<in> graph H' h'" ..
  34.230 -            with eq have "(x', h' x') \<in> graph H h" by (simp only:)
  34.231 -            then have "x' \<in> H" ..
  34.232 -            with `x' \<notin> H` show False by contradiction
  34.233 -          qed
  34.234 -          with g_rep show ?thesis by simp
  34.235 -        qed
  34.236 -      qed
  34.237 -      moreover have "graph H' h' \<in> M"
  34.238 -        -- {* and @{text h'} is norm-preserving. \skp *}
  34.239 -      proof (unfold M_def)
  34.240 -        show "graph H' h' \<in> norm_pres_extensions E p F f"
  34.241 -        proof (rule norm_pres_extensionI2)
  34.242 -          show "linearform H' h'"
  34.243 -	    using h'_def H'_def HE linearform `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E
  34.244 -	    by (rule h'_lf)
  34.245 -          show "H' \<unlhd> E"
  34.246 -	  unfolding H'_def
  34.247 -          proof
  34.248 -            show "H \<unlhd> E" by fact
  34.249 -            show "vectorspace E" by fact
  34.250 -            from x'E show "lin x' \<unlhd> E" ..
  34.251 -          qed
  34.252 -          from H `F \<unlhd> H` HH' show FH': "F \<unlhd> H'"
  34.253 -            by (rule vectorspace.subspace_trans)
  34.254 -          show "graph F f \<subseteq> graph H' h'"
  34.255 -          proof (rule graph_extI)
  34.256 -            fix x assume x: "x \<in> F"
  34.257 -            with graphs have "f x = h x" ..
  34.258 -            also have "\<dots> = h x + 0 * xi" by simp
  34.259 -            also have "\<dots> = (let (y, a) = (x, 0) in h y + a * xi)"
  34.260 -              by (simp add: Let_def)
  34.261 -            also have "(x, 0) =
  34.262 -                (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)"
  34.263 -	      using E HE
  34.264 -            proof (rule decomp_H'_H [symmetric])
  34.265 -              from FH x show "x \<in> H" ..
  34.266 -              from x' show "x' \<noteq> 0" .
  34.267 -	      show "x' \<notin> H" by fact
  34.268 -	      show "x' \<in> E" by fact
  34.269 -            qed
  34.270 -            also have
  34.271 -              "(let (y, a) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)
  34.272 -              in h y + a * xi) = h' x" by (simp only: h'_def)
  34.273 -            finally show "f x = h' x" .
  34.274 -          next
  34.275 -            from FH' show "F \<subseteq> H'" ..
  34.276 -          qed
  34.277 -          show "\<forall>x \<in> H'. h' x \<le> p x"
  34.278 -	    using h'_def H'_def `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E HE
  34.279 -	      `seminorm E p` linearform and hp xi
  34.280 -	    by (rule h'_norm_pres)
  34.281 -        qed
  34.282 -      qed
  34.283 -      ultimately show ?thesis ..
  34.284 -    qed
  34.285 -    then have "\<not> (\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x)" by simp
  34.286 -      -- {* So the graph @{text g} of @{text h} cannot be maximal. Contradiction! \skp *}
  34.287 -    with gx show "H = E" by contradiction
  34.288 -  qed
  34.289 -
  34.290 -  from HE_eq and linearform have "linearform E h"
  34.291 -    by (simp only:)
  34.292 -  moreover have "\<forall>x \<in> F. h x = f x"
  34.293 -  proof
  34.294 -    fix x assume "x \<in> F"
  34.295 -    with graphs have "f x = h x" ..
  34.296 -    then show "h x = f x" ..
  34.297 -  qed
  34.298 -  moreover from HE_eq and hp have "\<forall>x \<in> E. h x \<le> p x"
  34.299 -    by (simp only:)
  34.300 -  ultimately show ?thesis by blast
  34.301 -qed
  34.302 -
  34.303 -
  34.304 -subsection  {* Alternative formulation *}
  34.305 -
  34.306 -text {*
  34.307 -  The following alternative formulation of the Hahn-Banach
  34.308 -  Theorem\label{abs-HahnBanach} uses the fact that for a real linear
  34.309 -  form @{text f} and a seminorm @{text p} the following inequations
  34.310 -  are equivalent:\footnote{This was shown in lemma @{thm [source]
  34.311 -  abs_ineq_iff} (see page \pageref{abs-ineq-iff}).}
  34.312 -  \begin{center}
  34.313 -  \begin{tabular}{lll}
  34.314 -  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
  34.315 -  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
  34.316 -  \end{tabular}
  34.317 -  \end{center}
  34.318 -*}
  34.319 -
  34.320 -theorem abs_HahnBanach:
  34.321 -  assumes E: "vectorspace E" and FE: "subspace F E"
  34.322 -    and lf: "linearform F f" and sn: "seminorm E p"
  34.323 -  assumes fp: "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
  34.324 -  shows "\<exists>g. linearform E g
  34.325 -    \<and> (\<forall>x \<in> F. g x = f x)
  34.326 -    \<and> (\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x)"
  34.327 -proof -
  34.328 -  interpret vectorspace [E] by fact
  34.329 -  interpret subspace [F E] by fact
  34.330 -  interpret linearform [F f] by fact
  34.331 -  interpret seminorm [E p] by fact
  34.332 -  have "\<exists>g. linearform E g \<and> (\<forall>x \<in> F. g x = f x) \<and> (\<forall>x \<in> E. g x \<le> p x)"
  34.333 -    using E FE sn lf
  34.334 -  proof (rule HahnBanach)
  34.335 -    show "\<forall>x \<in> F. f x \<le> p x"
  34.336 -      using FE E sn lf and fp by (rule abs_ineq_iff [THEN iffD1])
  34.337 -  qed
  34.338 -  then obtain g where lg: "linearform E g" and *: "\<forall>x \<in> F. g x = f x"
  34.339 -      and **: "\<forall>x \<in> E. g x \<le> p x" by blast
  34.340 -  have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
  34.341 -    using _ E sn lg **
  34.342 -  proof (rule abs_ineq_iff [THEN iffD2])
  34.343 -    show "E \<unlhd> E" ..
  34.344 -  qed
  34.345 -  with lg * show ?thesis by blast
  34.346 -qed
  34.347 -
  34.348 -
  34.349 -subsection {* The Hahn-Banach Theorem for normed spaces *}
  34.350 -
  34.351 -text {*
  34.352 -  Every continuous linear form @{text f} on a subspace @{text F} of a
  34.353 -  norm space @{text E}, can be extended to a continuous linear form
  34.354 -  @{text g} on @{text E} such that @{text "\<parallel>f\<parallel> = \<parallel>g\<parallel>"}.
  34.355 -*}
  34.356 -
  34.357 -theorem norm_HahnBanach:
  34.358 -  fixes V and norm ("\<parallel>_\<parallel>")
  34.359 -  fixes B defines "\<And>V f. B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
  34.360 -  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
  34.361 -  defines "\<And>V f. \<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
  34.362 -  assumes E_norm: "normed_vectorspace E norm" and FE: "subspace F E"
  34.363 -    and linearform: "linearform F f" and "continuous F norm f"
  34.364 -  shows "\<exists>g. linearform E g
  34.365 -     \<and> continuous E norm g
  34.366 -     \<and> (\<forall>x \<in> F. g x = f x)
  34.367 -     \<and> \<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
  34.368 -proof -
  34.369 -  interpret normed_vectorspace [E norm] by fact
  34.370 -  interpret normed_vectorspace_with_fn_norm [E norm B fn_norm]
  34.371 -    by (auto simp: B_def fn_norm_def) intro_locales
  34.372 -  interpret subspace [F E] by fact
  34.373 -  interpret linearform [F f] by fact
  34.374 -  interpret continuous [F norm f] by fact
  34.375 -  have E: "vectorspace E" by intro_locales
  34.376 -  have F: "vectorspace F" by rule intro_locales
  34.377 -  have F_norm: "normed_vectorspace F norm"
  34.378 -    using FE E_norm by (rule subspace_normed_vs)
  34.379 -  have ge_zero: "0 \<le> \<parallel>f\<parallel>\<hyphen>F"
  34.380 -    by (rule normed_vectorspace_with_fn_norm.fn_norm_ge_zero
  34.381 -      [OF normed_vectorspace_with_fn_norm.intro,
  34.382 -       OF F_norm `continuous F norm f` , folded B_def fn_norm_def])
  34.383 -  txt {* We define a function @{text p} on @{text E} as follows:
  34.384 -    @{text "p x = \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"} *}
  34.385 -  def p \<equiv> "\<lambda>x. \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
  34.386 -
  34.387 -  txt {* @{text p} is a seminorm on @{text E}: *}
  34.388 -  have q: "seminorm E p"
  34.389 -  proof
  34.390 -    fix x y a assume x: "x \<in> E" and y: "y \<in> E"
  34.391 -    
  34.392 -    txt {* @{text p} is positive definite: *}
  34.393 -    have "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
  34.394 -    moreover from x have "0 \<le> \<parallel>x\<parallel>" ..
  34.395 -    ultimately show "0 \<le> p x"  
  34.396 -      by (simp add: p_def zero_le_mult_iff)
  34.397 -
  34.398 -    txt {* @{text p} is absolutely homogenous: *}
  34.399 -
  34.400 -    show "p (a \<cdot> x) = \<bar>a\<bar> * p x"
  34.401 -    proof -
  34.402 -      have "p (a \<cdot> x) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>a \<cdot> x\<parallel>" by (simp only: p_def)
  34.403 -      also from x have "\<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>" by (rule abs_homogenous)
  34.404 -      also have "\<parallel>f\<parallel>\<hyphen>F * (\<bar>a\<bar> * \<parallel>x\<parallel>) = \<bar>a\<bar> * (\<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>)" by simp
  34.405 -      also have "\<dots> = \<bar>a\<bar> * p x" by (simp only: p_def)
  34.406 -      finally show ?thesis .
  34.407 -    qed
  34.408 -
  34.409 -    txt {* Furthermore, @{text p} is subadditive: *}
  34.410 -
  34.411 -    show "p (x + y) \<le> p x + p y"
  34.412 -    proof -
  34.413 -      have "p (x + y) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel>" by (simp only: p_def)
  34.414 -      also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
  34.415 -      from x y have "\<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>" ..
  34.416 -      with a have " \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>F * (\<parallel>x\<parallel> + \<parallel>y\<parallel>)"
  34.417 -        by (simp add: mult_left_mono)
  34.418 -      also have "\<dots> = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel> + \<parallel>f\<parallel>\<hyphen>F * \<parallel>y\<parallel>" by (simp only: right_distrib)
  34.419 -      also have "\<dots> = p x + p y" by (simp only: p_def)
  34.420 -      finally show ?thesis .
  34.421 -    qed
  34.422 -  qed
  34.423 -
  34.424 -  txt {* @{text f} is bounded by @{text p}. *}
  34.425 -
  34.426 -  have "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
  34.427 -  proof
  34.428 -    fix x assume "x \<in> F"
  34.429 -    with `continuous F norm f` and linearform
  34.430 -    show "\<bar>f x\<bar> \<le> p x"
  34.431 -      unfolding p_def by (rule normed_vectorspace_with_fn_norm.fn_norm_le_cong
  34.432 -        [OF normed_vectorspace_with_fn_norm.intro,
  34.433 -         OF F_norm, folded B_def fn_norm_def])
  34.434 -  qed
  34.435 -
  34.436 -  txt {* Using the fact that @{text p} is a seminorm and @{text f} is bounded
  34.437 -    by @{text p} we can apply the Hahn-Banach Theorem for real vector
  34.438 -    spaces. So @{text f} can be extended in a norm-preserving way to
  34.439 -    some function @{text g} on the whole vector space @{text E}. *}
  34.440 -
  34.441 -  with E FE linearform q obtain g where
  34.442 -      linearformE: "linearform E g"
  34.443 -    and a: "\<forall>x \<in> F. g x = f x"
  34.444 -    and b: "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
  34.445 -    by (rule abs_HahnBanach [elim_format]) iprover
  34.446 -
  34.447 -  txt {* We furthermore have to show that @{text g} is also continuous: *}
  34.448 -
  34.449 -  have g_cont: "continuous E norm g" using linearformE
  34.450 -  proof
  34.451 -    fix x assume "x \<in> E"
  34.452 -    with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
  34.453 -      by (simp only: p_def)
  34.454 -  qed
  34.455 -
  34.456 -  txt {* To complete the proof, we show that @{text "\<parallel>g\<parallel> = \<parallel>f\<parallel>"}. *}
  34.457 -
  34.458 -  have "\<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
  34.459 -  proof (rule order_antisym)
  34.460 -    txt {*
  34.461 -      First we show @{text "\<parallel>g\<parallel> \<le> \<parallel>f\<parallel>"}.  The function norm @{text
  34.462 -      "\<parallel>g\<parallel>"} is defined as the smallest @{text "c \<in> \<real>"} such that
  34.463 -      \begin{center}
  34.464 -      \begin{tabular}{l}
  34.465 -      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
  34.466 -      \end{tabular}
  34.467 -      \end{center}
  34.468 -      \noindent Furthermore holds
  34.469 -      \begin{center}
  34.470 -      \begin{tabular}{l}
  34.471 -      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
  34.472 -      \end{tabular}
  34.473 -      \end{center}
  34.474 -    *}
  34.475 -
  34.476 -    have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
  34.477 -    proof
  34.478 -      fix x assume "x \<in> E"
  34.479 -      with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
  34.480 -        by (simp only: p_def)
  34.481 -    qed
  34.482 -    from g_cont this ge_zero
  34.483 -    show "\<parallel>g\<parallel>\<hyphen>E \<le> \<parallel>f\<parallel>\<hyphen>F"
  34.484 -      by (rule fn_norm_least [of g, folded B_def fn_norm_def])
  34.485 -
  34.486 -    txt {* The other direction is achieved by a similar argument. *}
  34.487 -
  34.488 -    show "\<parallel>f\<parallel>\<hyphen>F \<le> \<parallel>g\<parallel>\<hyphen>E"
  34.489 -    proof (rule normed_vectorspace_with_fn_norm.fn_norm_least
  34.490 -	[OF normed_vectorspace_with_fn_norm.intro,
  34.491 -	 OF F_norm, folded B_def fn_norm_def])
  34.492 -      show "\<forall>x \<in> F. \<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
  34.493 -      proof
  34.494 -	fix x assume x: "x \<in> F"
  34.495 -	from a x have "g x = f x" ..
  34.496 -	then have "\<bar>f x\<bar> = \<bar>g x\<bar>" by (simp only:)
  34.497 -	also from g_cont
  34.498 -	have "\<dots> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
  34.499 -	proof (rule fn_norm_le_cong [of g, folded B_def fn_norm_def])
  34.500 -	  from FE x show "x \<in> E" ..
  34.501 -	qed
  34.502 -	finally show "\<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>" .
  34.503 -      qed
  34.504 -      show "0 \<le> \<parallel>g\<parallel>\<hyphen>E"
  34.505 -	using g_cont
  34.506 -	by (rule fn_norm_ge_zero [of g, folded B_def fn_norm_def])
  34.507 -      show "continuous F norm f" by fact
  34.508 -    qed
  34.509 -  qed
  34.510 -  with linearformE a g_cont show ?thesis by blast
  34.511 -qed
  34.512 -
  34.513 -end
    35.1 --- a/src/HOL/Real/HahnBanach/HahnBanachExtLemmas.thy	Mon Dec 29 13:23:53 2008 +0100
    35.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    35.3 @@ -1,281 +0,0 @@
    35.4 -(*  Title:      HOL/Real/HahnBanach/HahnBanachExtLemmas.thy
    35.5 -    ID:         $Id$
    35.6 -    Author:     Gertrud Bauer, TU Munich
    35.7 -*)
    35.8 -
    35.9 -header {* Extending non-maximal functions *}
   35.10 -
   35.11 -theory HahnBanachExtLemmas
   35.12 -imports FunctionNorm
   35.13 -begin
   35.14 -
   35.15 -text {*
   35.16 -  In this section the following context is presumed.  Let @{text E} be
   35.17 -  a real vector space with a seminorm @{text q} on @{text E}. @{text
   35.18 -  F} is a subspace of @{text E} and @{text f} a linear function on
   35.19 -  @{text F}. We consider a subspace @{text H} of @{text E} that is a
   35.20 -  superspace of @{text F} and a linear form @{text h} on @{text
   35.21 -  H}. @{text H} is a not equal to @{text E} and @{text "x\<^sub>0"} is
   35.22 -  an element in @{text "E - H"}.  @{text H} is extended to the direct
   35.23 -  sum @{text "H' = H + lin x\<^sub>0"}, so for any @{text "x \<in> H'"}
   35.24 -  the decomposition of @{text "x = y + a \<cdot> x"} with @{text "y \<in> H"} is
   35.25 -  unique. @{text h'} is defined on @{text H'} by @{text "h' x = h y +
   35.26 -  a \<cdot> \<xi>"} for a certain @{text \<xi>}.
   35.27 -
   35.28 -  Subsequently we show some properties of this extension @{text h'} of
   35.29 -  @{text h}.
   35.30 -
   35.31 -  \medskip This lemma will be used to show the existence of a linear
   35.32 -  extension of @{text f} (see page \pageref{ex-xi-use}). It is a
   35.33 -  consequence of the completeness of @{text \<real>}. To show
   35.34 -  \begin{center}
   35.35 -  \begin{tabular}{l}
   35.36 -  @{text "\<exists>\<xi>. \<forall>y \<in> F. a y \<le> \<xi> \<and> \<xi> \<le> b y"}
   35.37 -  \end{tabular}
   35.38 -  \end{center}
   35.39 -  \noindent it suffices to show that
   35.40 -  \begin{center}
   35.41 -  \begin{tabular}{l}
   35.42 -  @{text "\<forall>u \<in> F. \<forall>v \<in> F. a u \<le> b v"}
   35.43 -  \end{tabular}
   35.44 -  \end{center}
   35.45 -*}
   35.46 -
   35.47 -lemma ex_xi:
   35.48 -  assumes "vectorspace F"
   35.49 -  assumes r: "\<And>u v. u \<in> F \<Longrightarrow> v \<in> F \<Longrightarrow> a u \<le> b v"
   35.50 -  shows "\<exists>xi::real. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y"
   35.51 -proof -
   35.52 -  interpret vectorspace [F] by fact
   35.53 -  txt {* From the completeness of the reals follows:
   35.54 -    The set @{text "S = {a u. u \<in> F}"} has a supremum, if it is
   35.55 -    non-empty and has an upper bound. *}
   35.56 -
   35.57 -  let ?S = "{a u | u. u \<in> F}"
   35.58 -  have "\<exists>xi. lub ?S xi"
   35.59 -  proof (rule real_complete)
   35.60 -    have "a 0 \<in> ?S" by blast
   35.61 -    then show "\<exists>X. X \<in> ?S" ..
   35.62 -    have "\<forall>y \<in> ?S. y \<le> b 0"
   35.63 -    proof
   35.64 -      fix y assume y: "y \<in> ?S"
   35.65 -      then obtain u where u: "u \<in> F" and y: "y = a u" by blast
   35.66 -      from u and zero have "a u \<le> b 0" by (rule r)
   35.67 -      with y show "y \<le> b 0" by (simp only:)
   35.68 -    qed
   35.69 -    then show "\<exists>u. \<forall>y \<in> ?S. y \<le> u" ..
   35.70 -  qed
   35.71 -  then obtain xi where xi: "lub ?S xi" ..
   35.72 -  {
   35.73 -    fix y assume "y \<in> F"
   35.74 -    then have "a y \<in> ?S" by blast
   35.75 -    with xi have "a y \<le> xi" by (rule lub.upper)
   35.76 -  } moreover {
   35.77 -    fix y assume y: "y \<in> F"
   35.78 -    from xi have "xi \<le> b y"
   35.79 -    proof (rule lub.least)
   35.80 -      fix au assume "au \<in> ?S"
   35.81 -      then obtain u where u: "u \<in> F" and au: "au = a u" by blast
   35.82 -      from u y have "a u \<le> b y" by (rule r)
   35.83 -      with au show "au \<le> b y" by (simp only:)
   35.84 -    qed
   35.85 -  } ultimately show "\<exists>xi. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y" by blast
   35.86 -qed
   35.87 -
   35.88 -text {*
   35.89 -  \medskip The function @{text h'} is defined as a @{text "h' x = h y
   35.90 -  + a \<cdot> \<xi>"} where @{text "x = y + a \<cdot> \<xi>"} is a linear extension of
   35.91 -  @{text h} to @{text H'}.
   35.92 -*}
   35.93 -
   35.94 -lemma h'_lf:
   35.95 -  assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
   35.96 -      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
   35.97 -    and H'_def: "H' \<equiv> H + lin x0"
   35.98 -    and HE: "H \<unlhd> E"
   35.99 -  assumes "linearform H h"
  35.100 -  assumes x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
  35.101 -  assumes E: "vectorspace E"
  35.102 -  shows "linearform H' h'"
  35.103 -proof -
  35.104 -  interpret linearform [H h] by fact
  35.105 -  interpret vectorspace [E] by fact
  35.106 -  show ?thesis
  35.107 -  proof
  35.108 -    note E = `vectorspace E`
  35.109 -    have H': "vectorspace H'"
  35.110 -    proof (unfold H'_def)
  35.111 -      from `x0 \<in> E`
  35.112 -      have "lin x0 \<unlhd> E" ..
  35.113 -      with HE show "vectorspace (H + lin x0)" using E ..
  35.114 -    qed
  35.115 -    {
  35.116 -      fix x1 x2 assume x1: "x1 \<in> H'" and x2: "x2 \<in> H'"
  35.117 -      show "h' (x1 + x2) = h' x1 + h' x2"
  35.118 -      proof -
  35.119 -	from H' x1 x2 have "x1 + x2 \<in> H'"
  35.120 -          by (rule vectorspace.add_closed)
  35.121 -	with x1 x2 obtain y y1 y2 a a1 a2 where
  35.122 -          x1x2: "x1 + x2 = y + a \<cdot> x0" and y: "y \<in> H"
  35.123 -          and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
  35.124 -          and x2_rep: "x2 = y2 + a2 \<cdot> x0" and y2: "y2 \<in> H"
  35.125 -          unfolding H'_def sum_def lin_def by blast
  35.126 -	
  35.127 -	have ya: "y1 + y2 = y \<and> a1 + a2 = a" using E HE _ y x0
  35.128 -	proof (rule decomp_H') txt_raw {* \label{decomp-H-use} *}
  35.129 -          from HE y1 y2 show "y1 + y2 \<in> H"
  35.130 -            by (rule subspace.add_closed)
  35.131 -          from x0 and HE y y1 y2
  35.132 -          have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E"  "y2 \<in> E" by auto
  35.133 -          with x1_rep x2_rep have "(y1 + y2) + (a1 + a2) \<cdot> x0 = x1 + x2"
  35.134 -            by (simp add: add_ac add_mult_distrib2)
  35.135 -          also note x1x2
  35.136 -          finally show "(y1 + y2) + (a1 + a2) \<cdot> x0 = y + a \<cdot> x0" .
  35.137 -	qed
  35.138 -	
  35.139 -	from h'_def x1x2 E HE y x0
  35.140 -	have "h' (x1 + x2) = h y + a * xi"
  35.141 -          by (rule h'_definite)
  35.142 -	also have "\<dots> = h (y1 + y2) + (a1 + a2) * xi"
  35.143 -          by (simp only: ya)
  35.144 -	also from y1 y2 have "h (y1 + y2) = h y1 + h y2"
  35.145 -          by simp
  35.146 -	also have "\<dots> + (a1 + a2) * xi = (h y1 + a1 * xi) + (h y2 + a2 * xi)"
  35.147 -          by (simp add: left_distrib)
  35.148 -	also from h'_def x1_rep E HE y1 x0
  35.149 -	have "h y1 + a1 * xi = h' x1"
  35.150 -          by (rule h'_definite [symmetric])
  35.151 -	also from h'_def x2_rep E HE y2 x0
  35.152 -	have "h y2 + a2 * xi = h' x2"
  35.153 -          by (rule h'_definite [symmetric])
  35.154 -	finally show ?thesis .
  35.155 -      qed
  35.156 -    next
  35.157 -      fix x1 c assume x1: "x1 \<in> H'"
  35.158 -      show "h' (c \<cdot> x1) = c * (h' x1)"
  35.159 -      proof -
  35.160 -	from H' x1 have ax1: "c \<cdot> x1 \<in> H'"
  35.161 -          by (rule vectorspace.mult_closed)
  35.162 -	with x1 obtain y a y1 a1 where
  35.163 -            cx1_rep: "c \<cdot> x1 = y + a \<cdot> x0" and y: "y \<in> H"
  35.164 -          and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
  35.165 -          unfolding H'_def sum_def lin_def by blast
  35.166 -	
  35.167 -	have ya: "c \<cdot> y1 = y \<and> c * a1 = a" using E HE _ y x0
  35.168 -	proof (rule decomp_H')
  35.169 -          from HE y1 show "c \<cdot> y1 \<in> H"
  35.170 -            by (rule subspace.mult_closed)
  35.171 -          from x0 and HE y y1
  35.172 -          have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E" by auto
  35.173 -          with x1_rep have "c \<cdot> y1 + (c * a1) \<cdot> x0 = c \<cdot> x1"
  35.174 -            by (simp add: mult_assoc add_mult_distrib1)
  35.175 -          also note cx1_rep
  35.176 -          finally show "c \<cdot> y1 + (c * a1) \<cdot> x0 = y + a \<cdot> x0" .
  35.177 -	qed
  35.178 -	
  35.179 -	from h'_def cx1_rep E HE y x0 have "h' (c \<cdot> x1) = h y + a * xi"
  35.180 -          by (rule h'_definite)
  35.181 -	also have "\<dots> = h (c \<cdot> y1) + (c * a1) * xi"
  35.182 -          by (simp only: ya)
  35.183 -	also from y1 have "h (c \<cdot> y1) = c * h y1"
  35.184 -          by simp
  35.185 -	also have "\<dots> + (c * a1) * xi = c * (h y1 + a1 * xi)"
  35.186 -          by (simp only: right_distrib)
  35.187 -	also from h'_def x1_rep E HE y1 x0 have "h y1 + a1 * xi = h' x1"
  35.188 -          by (rule h'_definite [symmetric])
  35.189 -	finally show ?thesis .
  35.190 -      qed
  35.191 -    }
  35.192 -  qed
  35.193 -qed
  35.194 -
  35.195 -text {* \medskip The linear extension @{text h'} of @{text h}
  35.196 -  is bounded by the seminorm @{text p}. *}
  35.197 -
  35.198 -lemma h'_norm_pres:
  35.199 -  assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
  35.200 -      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
  35.201 -    and H'_def: "H' \<equiv> H + lin x0"
  35.202 -    and x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
  35.203 -  assumes E: "vectorspace E" and HE: "subspace H E"
  35.204 -    and "seminorm E p" and "linearform H h"
  35.205 -  assumes a: "\<forall>y \<in> H. h y \<le> p y"
  35.206 -    and a': "\<forall>y \<in> H. - p (y + x0) - h y \<le> xi \<and> xi \<le> p (y + x0) - h y"
  35.207 -  shows "\<forall>x \<in> H'. h' x \<le> p x"
  35.208 -proof -
  35.209 -  interpret vectorspace [E] by fact
  35.210 -  interpret subspace [H E] by fact
  35.211 -  interpret seminorm [E p] by fact
  35.212 -  interpret linearform [H h] by fact
  35.213 -  show ?thesis
  35.214 -  proof
  35.215 -    fix x assume x': "x \<in> H'"
  35.216 -    show "h' x \<le> p x"
  35.217 -    proof -
  35.218 -      from a' have a1: "\<forall>ya \<in> H. - p (ya + x0) - h ya \<le> xi"
  35.219 -	and a2: "\<forall>ya \<in> H. xi \<le> p (ya + x0) - h ya" by auto
  35.220 -      from x' obtain y a where
  35.221 -          x_rep: "x = y + a \<cdot> x0" and y: "y \<in> H"
  35.222 -	unfolding H'_def sum_def lin_def by blast
  35.223 -      from y have y': "y \<in> E" ..
  35.224 -      from y have ay: "inverse a \<cdot> y \<in> H" by simp
  35.225 -      
  35.226 -      from h'_def x_rep E HE y x0 have "h' x = h y + a * xi"
  35.227 -	by (rule h'_definite)
  35.228 -      also have "\<dots> \<le> p (y + a \<cdot> x0)"
  35.229 -      proof (rule linorder_cases)
  35.230 -	assume z: "a = 0"
  35.231 -	then have "h y + a * xi = h y" by simp
  35.232 -	also from a y have "\<dots> \<le> p y" ..
  35.233 -	also from x0 y' z have "p y = p (y + a \<cdot> x0)" by simp
  35.234 -	finally show ?thesis .
  35.235 -      next
  35.236 -	txt {* In the case @{text "a < 0"}, we use @{text "a\<^sub>1"}
  35.237 -          with @{text ya} taken as @{text "y / a"}: *}
  35.238 -	assume lz: "a < 0" then have nz: "a \<noteq> 0" by simp
  35.239 -	from a1 ay
  35.240 -	have "- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y) \<le> xi" ..
  35.241 -	with lz have "a * xi \<le>
  35.242 -          a * (- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
  35.243 -          by (simp add: mult_left_mono_neg order_less_imp_le)
  35.244 -	
  35.245 -	also have "\<dots> =
  35.246 -          - a * (p (inverse a \<cdot> y + x0)) - a * (h (inverse a \<cdot> y))"
  35.247 -	  by (simp add: right_diff_distrib)
  35.248 -	also from lz x0 y' have "- a * (p (inverse a \<cdot> y + x0)) =
  35.249 -          p (a \<cdot> (inverse a \<cdot> y + x0))"
  35.250 -          by (simp add: abs_homogenous)
  35.251 -	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
  35.252 -          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
  35.253 -	also from nz y have "a * (h (inverse a \<cdot> y)) =  h y"
  35.254 -          by simp
  35.255 -	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
  35.256 -	then show ?thesis by simp
  35.257 -      next
  35.258 -	txt {* In the case @{text "a > 0"}, we use @{text "a\<^sub>2"}
  35.259 -          with @{text ya} taken as @{text "y / a"}: *}
  35.260 -	assume gz: "0 < a" then have nz: "a \<noteq> 0" by simp
  35.261 -	from a2 ay
  35.262 -	have "xi \<le> p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y)" ..
  35.263 -	with gz have "a * xi \<le>
  35.264 -          a * (p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
  35.265 -          by simp
  35.266 -	also have "\<dots> = a * p (inverse a \<cdot> y + x0) - a * h (inverse a \<cdot> y)"
  35.267 -	  by (simp add: right_diff_distrib)
  35.268 -	also from gz x0 y'
  35.269 -	have "a * p (inverse a \<cdot> y + x0) = p (a \<cdot> (inverse a \<cdot> y + x0))"
  35.270 -          by (simp add: abs_homogenous)
  35.271 -	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
  35.272 -          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
  35.273 -	also from nz y have "a * h (inverse a \<cdot> y) = h y"
  35.274 -          by simp
  35.275 -	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
  35.276 -	then show ?thesis by simp
  35.277 -      qed
  35.278 -      also from x_rep have "\<dots> = p x" by (simp only:)
  35.279 -      finally show ?thesis .
  35.280 -    qed
  35.281 -  qed
  35.282 -qed
  35.283 -
  35.284 -end
    36.1 --- a/src/HOL/Real/HahnBanach/HahnBanachLemmas.thy	Mon Dec 29 13:23:53 2008 +0100
    36.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    36.3 @@ -1,4 +0,0 @@
    36.4 -(*<*)
    36.5 -theory HahnBanachLemmas imports HahnBanachSupLemmas HahnBanachExtLemmas begin
    36.6 -end
    36.7 -(*>*)
    36.8 \ No newline at end of file
    37.1 --- a/src/HOL/Real/HahnBanach/HahnBanachSupLemmas.thy	Mon Dec 29 13:23:53 2008 +0100
    37.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    37.3 @@ -1,446 +0,0 @@
    37.4 -(*  Title:      HOL/Real/HahnBanach/HahnBanachSupLemmas.thy
    37.5 -    ID:         $Id$
    37.6 -    Author:     Gertrud Bauer, TU Munich
    37.7 -*)
    37.8 -
    37.9 -header {* The supremum w.r.t.~the function order *}
   37.10 -
   37.11 -theory HahnBanachSupLemmas
   37.12 -imports FunctionNorm ZornLemma
   37.13 -begin
   37.14 -
   37.15 -text {*
   37.16 -  This section contains some lemmas that will be used in the proof of
   37.17 -  the Hahn-Banach Theorem.  In this section the following context is
   37.18 -  presumed.  Let @{text E} be a real vector space with a seminorm
   37.19 -  @{text p} on @{text E}.  @{text F} is a subspace of @{text E} and
   37.20 -  @{text f} a linear form on @{text F}. We consider a chain @{text c}
   37.21 -  of norm-preserving extensions of @{text f}, such that @{text "\<Union>c =
   37.22 -  graph H h"}.  We will show some properties about the limit function
   37.23 -  @{text h}, i.e.\ the supremum of the chain @{text c}.
   37.24 -
   37.25 -  \medskip Let @{text c} be a chain of norm-preserving extensions of
   37.26 -  the function @{text f} and let @{text "graph H h"} be the supremum
   37.27 -  of @{text c}.  Every element in @{text H} is member of one of the
   37.28 -  elements of the chain.
   37.29 -*}
   37.30 -lemmas [dest?] = chainD
   37.31 -lemmas chainE2 [elim?] = chainD2 [elim_format, standard]
   37.32 -
   37.33 -lemma some_H'h't:
   37.34 -  assumes M: "M = norm_pres_extensions E p F f"
   37.35 -    and cM: "c \<in> chain M"
   37.36 -    and u: "graph H h = \<Union>c"
   37.37 -    and x: "x \<in> H"
   37.38 -  shows "\<exists>H' h'. graph H' h' \<in> c
   37.39 -    \<and> (x, h x) \<in> graph H' h'
   37.40 -    \<and> linearform H' h' \<and> H' \<unlhd> E
   37.41 -    \<and> F \<unlhd> H' \<and> graph F f \<subseteq> graph H' h'
   37.42 -    \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
   37.43 -proof -
   37.44 -  from x have "(x, h x) \<in> graph H h" ..
   37.45 -  also from u have "\<dots> = \<Union>c" .
   37.46 -  finally obtain g where gc: "g \<in> c" and gh: "(x, h x) \<in> g" by blast
   37.47 -
   37.48 -  from cM have "c \<subseteq> M" ..
   37.49 -  with gc have "g \<in> M" ..
   37.50 -  also from M have "\<dots> = norm_pres_extensions E p F f" .
   37.51 -  finally obtain H' and h' where g: "g = graph H' h'"
   37.52 -    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
   37.53 -      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x" ..
   37.54 -
   37.55 -  from gc and g have "graph H' h' \<in> c" by (simp only:)
   37.56 -  moreover from gh and g have "(x, h x) \<in> graph H' h'" by (simp only:)
   37.57 -  ultimately show ?thesis using * by blast
   37.58 -qed
   37.59 -
   37.60 -text {*
   37.61 -  \medskip Let @{text c} be a chain of norm-preserving extensions of
   37.62 -  the function @{text f} and let @{text "graph H h"} be the supremum
   37.63 -  of @{text c}.  Every element in the domain @{text H} of the supremum
   37.64 -  function is member of the domain @{text H'} of some function @{text
   37.65 -  h'}, such that @{text h} extends @{text h'}.
   37.66 -*}
   37.67 -
   37.68 -lemma some_H'h':
   37.69 -  assumes M: "M = norm_pres_extensions E p F f"
   37.70 -    and cM: "c \<in> chain M"
   37.71 -    and u: "graph H h = \<Union>c"
   37.72 -    and x: "x \<in> H"
   37.73 -  shows "\<exists>H' h'. x \<in> H' \<and> graph H' h' \<subseteq> graph H h
   37.74 -    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
   37.75 -    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
   37.76 -proof -
   37.77 -  from M cM u x obtain H' h' where
   37.78 -      x_hx: "(x, h x) \<in> graph H' h'"
   37.79 -    and c: "graph H' h' \<in> c"
   37.80 -    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
   37.81 -      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
   37.82 -    by (rule some_H'h't [elim_format]) blast
   37.83 -  from x_hx have "x \<in> H'" ..
   37.84 -  moreover from cM u c have "graph H' h' \<subseteq> graph H h"
   37.85 -    by (simp only: chain_ball_Union_upper)
   37.86 -  ultimately show ?thesis using * by blast
   37.87 -qed
   37.88 -
   37.89 -text {*
   37.90 -  \medskip Any two elements @{text x} and @{text y} in the domain
   37.91 -  @{text H} of the supremum function @{text h} are both in the domain
   37.92 -  @{text H'} of some function @{text h'}, such that @{text h} extends
   37.93 -  @{text h'}.
   37.94 -*}
   37.95 -
   37.96 -lemma some_H'h'2:
   37.97 -  assumes M: "M = norm_pres_extensions E p F f"
   37.98 -    and cM: "c \<in> chain M"
   37.99 -    and u: "graph H h = \<Union>c"
  37.100 -    and x: "x \<in> H"
  37.101 -    and y: "y \<in> H"
  37.102 -  shows "\<exists>H' h'. x \<in> H' \<and> y \<in> H'
  37.103 -    \<and> graph H' h' \<subseteq> graph H h
  37.104 -    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
  37.105 -    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
  37.106 -proof -
  37.107 -  txt {* @{text y} is in the domain @{text H''} of some function @{text h''},
  37.108 -  such that @{text h} extends @{text h''}. *}
  37.109 -
  37.110 -  from M cM u and y obtain H' h' where
  37.111 -      y_hy: "(y, h y) \<in> graph H' h'"
  37.112 -    and c': "graph H' h' \<in> c"
  37.113 -    and * :
  37.114 -      "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
  37.115 -      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
  37.116 -    by (rule some_H'h't [elim_format]) blast
  37.117 -
  37.118 -  txt {* @{text x} is in the domain @{text H'} of some function @{text h'},
  37.119 -    such that @{text h} extends @{text h'}. *}
  37.120 -
  37.121 -  from M cM u and x obtain H'' h'' where
  37.122 -      x_hx: "(x, h x) \<in> graph H'' h''"
  37.123 -    and c'': "graph H'' h'' \<in> c"
  37.124 -    and ** :
  37.125 -      "linearform H'' h''"  "H'' \<unlhd> E"  "F \<unlhd> H''"
  37.126 -      "graph F f \<subseteq> graph H'' h''"  "\<forall>x \<in> H''. h'' x \<le> p x"
  37.127 -    by (rule some_H'h't [elim_format]) blast
  37.128 -
  37.129 -  txt {* Since both @{text h'} and @{text h''} are elements of the chain,
  37.130 -    @{text h''} is an extension of @{text h'} or vice versa. Thus both
  37.131 -    @{text x} and @{text y} are contained in the greater
  37.132 -    one. \label{cases1}*}
  37.133 -
  37.134 -  from cM c'' c' have "graph H'' h'' \<subseteq> graph H' h' \<or> graph H' h' \<subseteq> graph H'' h''"
  37.135 -    (is "?case1 \<or> ?case2") ..
  37.136 -  then show ?thesis
  37.137 -  proof
  37.138 -    assume ?case1
  37.139 -    have "(x, h x) \<in> graph H'' h''" by fact
  37.140 -    also have "\<dots> \<subseteq> graph H' h'" by fact
  37.141 -    finally have xh:"(x, h x) \<in> graph H' h'" .
  37.142 -    then have "x \<in> H'" ..
  37.143 -    moreover from y_hy have "y \<in> H'" ..
  37.144 -    moreover from cM u and c' have "graph H' h' \<subseteq> graph H h"
  37.145 -      by (simp only: chain_ball_Union_upper)
  37.146 -    ultimately show ?thesis using * by blast
  37.147 -  next
  37.148 -    assume ?case2
  37.149 -    from x_hx have "x \<in> H''" ..
  37.150 -    moreover {
  37.151 -      have "(y, h y) \<in> graph H' h'" by (rule y_hy)
  37.152 -      also have "\<dots> \<subseteq> graph H'' h''" by fact
  37.153 -      finally have "(y, h y) \<in> graph H'' h''" .
  37.154 -    } then have "y \<in> H''" ..
  37.155 -    moreover from cM u and c'' have "graph H'' h'' \<subseteq> graph H h"
  37.156 -      by (simp only: chain_ball_Union_upper)
  37.157 -    ultimately show ?thesis using ** by blast
  37.158 -  qed
  37.159 -qed
  37.160 -
  37.161 -text {*
  37.162 -  \medskip The relation induced by the graph of the supremum of a
  37.163 -  chain @{text c} is definite, i.~e.~t is the graph of a function.
  37.164 -*}
  37.165 -
  37.166 -lemma sup_definite:
  37.167 -  assumes M_def: "M \<equiv> norm_pres_extensions E p F f"
  37.168 -    and cM: "c \<in> chain M"
  37.169 -    and xy: "(x, y) \<in> \<Union>c"
  37.170 -    and xz: "(x, z) \<in> \<Union>c"
  37.171 -  shows "z = y"
  37.172 -proof -
  37.173 -  from cM have c: "c \<subseteq> M" ..
  37.174 -  from xy obtain G1 where xy': "(x, y) \<in> G1" and G1: "G1 \<in> c" ..
  37.175 -  from xz obtain G2 where xz': "(x, z) \<in> G2" and G2: "G2 \<in> c" ..
  37.176 -
  37.177 -  from G1 c have "G1 \<in> M" ..
  37.178 -  then obtain H1 h1 where G1_rep: "G1 = graph H1 h1"
  37.179 -    unfolding M_def by blast
  37.180 -
  37.181 -  from G2 c have "G2 \<in> M" ..
  37.182 -  then obtain H2 h2 where G2_rep: "G2 = graph H2 h2"
  37.183 -    unfolding M_def by blast
  37.184 -
  37.185 -  txt {* @{text "G\<^sub>1"} is contained in @{text "G\<^sub>2"}
  37.186 -    or vice versa, since both @{text "G\<^sub>1"} and @{text
  37.187 -    "G\<^sub>2"} are members of @{text c}. \label{cases2}*}
  37.188 -
  37.189 -  from cM G1 G2 have "G1 \<subseteq> G2 \<or> G2 \<subseteq> G1" (is "?case1 \<or> ?case2") ..
  37.190 -  then show ?thesis
  37.191 -  proof
  37.192 -    assume ?case1
  37.193 -    with xy' G2_rep have "(x, y) \<in> graph H2 h2" by blast
  37.194 -    then have "y = h2 x" ..
  37.195 -    also
  37.196 -    from xz' G2_rep have "(x, z) \<in> graph H2 h2" by (simp only:)
  37.197 -    then have "z = h2 x" ..
  37.198 -    finally show ?thesis .
  37.199 -  next
  37.200 -    assume ?case2
  37.201 -    with xz' G1_rep have "(x, z) \<in> graph H1 h1" by blast
  37.202 -    then have "z = h1 x" ..
  37.203 -    also
  37.204 -    from xy' G1_rep have "(x, y) \<in> graph H1 h1" by (simp only:)
  37.205 -    then have "y = h1 x" ..
  37.206 -    finally show ?thesis ..
  37.207 -  qed
  37.208 -qed
  37.209 -
  37.210 -text {*
  37.211 -  \medskip The limit function @{text h} is linear. Every element
  37.212 -  @{text x} in the domain of @{text h} is in the domain of a function
  37.213 -  @{text h'} in the chain of norm preserving extensions.  Furthermore,
  37.214 -  @{text h} is an extension of @{text h'} so the function values of
  37.215 -  @{text x} are identical for @{text h'} and @{text h}.  Finally, the
  37.216 -  function @{text h'} is linear by construction of @{text M}.
  37.217 -*}
  37.218 -
  37.219 -lemma sup_lf:
  37.220 -  assumes M: "M = norm_pres_extensions E p F f"
  37.221 -    and cM: "c \<in> chain M"
  37.222 -    and u: "graph H h = \<Union>c"
  37.223 -  shows "linearform H h"
  37.224 -proof
  37.225 -  fix x y assume x: "x \<in> H" and y: "y \<in> H"
  37.226 -  with M cM u obtain H' h' where
  37.227 -        x': "x \<in> H'" and y': "y \<in> H'"
  37.228 -      and b: "graph H' h' \<subseteq> graph H h"
  37.229 -      and linearform: "linearform H' h'"
  37.230 -      and subspace: "H' \<unlhd> E"
  37.231 -    by (rule some_H'h'2 [elim_format]) blast
  37.232 -
  37.233 -  show "h (x + y) = h x + h y"
  37.234 -  proof -
  37.235 -    from linearform x' y' have "h' (x + y) = h' x + h' y"
  37.236 -      by (rule linearform.add)
  37.237 -    also from b x' have "h' x = h x" ..
  37.238 -    also from b y' have "h' y = h y" ..
  37.239 -    also from subspace x' y' have "x + y \<in> H'"
  37.240 -      by (rule subspace.add_closed)
  37.241 -    with b have "h' (x + y) = h (x + y)" ..
  37.242 -    finally show ?thesis .
  37.243 -  qed
  37.244 -next
  37.245 -  fix x a assume x: "x \<in> H"
  37.246 -  with M cM u obtain H' h' where
  37.247 -        x': "x \<in> H'"
  37.248 -      and b: "graph H' h' \<subseteq> graph H h"
  37.249 -      and linearform: "linearform H' h'"
  37.250 -      and subspace: "H' \<unlhd> E"
  37.251 -    by (rule some_H'h' [elim_format]) blast
  37.252 -
  37.253 -  show "h (a \<cdot> x) = a * h x"
  37.254 -  proof -
  37.255 -    from linearform x' have "h' (a \<cdot> x) = a * h' x"
  37.256 -      by (rule linearform.mult)
  37.257 -    also from b x' have "h' x = h x" ..
  37.258 -    also from subspace x' have "a \<cdot> x \<in> H'"
  37.259 -      by (rule subspace.mult_closed)
  37.260 -    with b have "h' (a \<cdot> x) = h (a \<cdot> x)" ..
  37.261 -    finally show ?thesis .
  37.262 -  qed
  37.263 -qed
  37.264 -
  37.265 -text {*
  37.266 -  \medskip The limit of a non-empty chain of norm preserving
  37.267 -  extensions of @{text f} is an extension of @{text f}, since every
  37.268 -  element of the chain is an extension of @{text f} and the supremum
  37.269 -  is an extension for every element of the chain.
  37.270 -*}
  37.271 -
  37.272 -lemma sup_ext:
  37.273 -  assumes graph: "graph H h = \<Union>c"
  37.274 -    and M: "M = norm_pres_extensions E p F f"
  37.275 -    and cM: "c \<in> chain M"
  37.276 -    and ex: "\<exists>x. x \<in> c"
  37.277 -  shows "graph F f \<subseteq> graph H h"
  37.278 -proof -
  37.279 -  from ex obtain x where xc: "x \<in> c" ..
  37.280 -  from cM have "c \<subseteq> M" ..
  37.281 -  with xc have "x \<in> M" ..
  37.282 -  with M have "x \<in> norm_pres_extensions E p F f"
  37.283 -    by (simp only:)
  37.284 -  then obtain G g where "x = graph G g" and "graph F f \<subseteq> graph G g" ..
  37.285 -  then have "graph F f \<subseteq> x" by (simp only:)
  37.286 -  also from xc have "\<dots> \<subseteq> \<Union>c" by blast
  37.287 -  also from graph have "\<dots> = graph H h" ..
  37.288 -  finally show ?thesis .
  37.289 -qed
  37.290 -
  37.291 -text {*
  37.292 -  \medskip The domain @{text H} of the limit function is a superspace
  37.293 -  of @{text F}, since @{text F} is a subset of @{text H}. The
  37.294 -  existence of the @{text 0} element in @{text F} and the closure
  37.295 -  properties follow from the fact that @{text F} is a vector space.
  37.296 -*}
  37.297 -
  37.298 -lemma sup_supF:
  37.299 -  assumes graph: "graph H h = \<Union>c"
  37.300 -    and M: "M = norm_pres_extensions E p F f"
  37.301 -    and cM: "c \<in> chain M"
  37.302 -    and ex: "\<exists>x. x \<in> c"
  37.303 -    and FE: "F \<unlhd> E"
  37.304 -  shows "F \<unlhd> H"
  37.305 -proof
  37.306 -  from FE show "F \<noteq> {}" by (rule subspace.non_empty)
  37.307 -  from graph M cM ex have "graph F f \<subseteq> graph H h" by (rule sup_ext)
  37.308 -  then show "F \<subseteq> H" ..
  37.309 -  fix x y assume "x \<in> F" and "y \<in> F"
  37.310 -  with FE show "x + y \<in> F" by (rule subspace.add_closed)
  37.311 -next
  37.312 -  fix x a assume "x \<in> F"
  37.313 -  with FE show "a \<cdot> x \<in> F" by (rule subspace.mult_closed)
  37.314 -qed
  37.315 -
  37.316 -text {*
  37.317 -  \medskip The domain @{text H} of the limit function is a subspace of
  37.318 -  @{text E}.
  37.319 -*}
  37.320 -
  37.321 -lemma sup_subE:
  37.322 -  assumes graph: "graph H h = \<Union>c"
  37.323 -    and M: "M = norm_pres_extensions E p F f"
  37.324 -    and cM: "c \<in> chain M"
  37.325 -    and ex: "\<exists>x. x \<in> c"
  37.326 -    and FE: "F \<unlhd> E"
  37.327 -    and E: "vectorspace E"
  37.328 -  shows "H \<unlhd> E"
  37.329 -proof
  37.330 -  show "H \<noteq> {}"
  37.331 -  proof -
  37.332 -    from FE E have "0 \<in> F" by (rule subspace.zero)
  37.333 -    also from graph M cM ex FE have "F \<unlhd> H" by (rule sup_supF)
  37.334 -    then have "F \<subseteq> H" ..
  37.335 -    finally show ?thesis by blast
  37.336 -  qed
  37.337 -  show "H \<subseteq> E"
  37.338 -  proof
  37.339 -    fix x assume "x \<in> H"
  37.340 -    with M cM graph
  37.341 -    obtain H' h' where x: "x \<in> H'" and H'E: "H' \<unlhd> E"
  37.342 -      by (rule some_H'h' [elim_format]) blast
  37.343 -    from H'E have "H' \<subseteq> E" ..
  37.344 -    with x show "x \<in> E" ..
  37.345 -  qed
  37.346 -  fix x y assume x: "x \<in> H" and y: "y \<in> H"
  37.347 -  show "x + y \<in> H"
  37.348 -  proof -
  37.349 -    from M cM graph x y obtain H' h' where
  37.350 -          x': "x \<in> H'" and y': "y \<in> H'" and H'E: "H' \<unlhd> E"
  37.351 -        and graphs: "graph H' h' \<subseteq> graph H h"
  37.352 -      by (rule some_H'h'2 [elim_format]) blast
  37.353 -    from H'E x' y' have "x + y \<in> H'"
  37.354 -      by (rule subspace.add_closed)
  37.355 -    also from graphs have "H' \<subseteq> H" ..
  37.356 -    finally show ?thesis .
  37.357 -  qed
  37.358 -next
  37.359 -  fix x a assume x: "x \<in> H"
  37.360 -  show "a \<cdot> x \<in> H"
  37.361 -  proof -
  37.362 -    from M cM graph x
  37.363 -    obtain H' h' where x': "x \<in> H'" and H'E: "H' \<unlhd> E"
  37.364 -        and graphs: "graph H' h' \<subseteq> graph H h"
  37.365 -      by (rule some_H'h' [elim_format]) blast
  37.366 -    from H'E x' have "a \<cdot> x \<in> H'" by (rule subspace.mult_closed)
  37.367 -    also from graphs have "H' \<subseteq> H" ..
  37.368 -    finally show ?thesis .
  37.369 -  qed
  37.370 -qed
  37.371 -
  37.372 -text {*
  37.373 -  \medskip The limit function is bounded by the norm @{text p} as
  37.374 -  well, since all elements in the chain are bounded by @{text p}.
  37.375 -*}
  37.376 -
  37.377 -lemma sup_norm_pres:
  37.378 -  assumes graph: "graph H h = \<Union>c"
  37.379 -    and M: "M = norm_pres_extensions E p F f"
  37.380 -    and cM: "c \<in> chain M"
  37.381 -  shows "\<forall>x \<in> H. h x \<le> p x"
  37.382 -proof
  37.383 -  fix x assume "x \<in> H"
  37.384 -  with M cM graph obtain H' h' where x': "x \<in> H'"
  37.385 -      and graphs: "graph H' h' \<subseteq> graph H h"
  37.386 -      and a: "\<forall>x \<in> H'. h' x \<le> p x"
  37.387 -    by (rule some_H'h' [elim_format]) blast
  37.388 -  from graphs x' have [symmetric]: "h' x = h x" ..
  37.389 -  also from a x' have "h' x \<le> p x " ..
  37.390 -  finally show "h x \<le> p x" .
  37.391 -qed
  37.392 -
  37.393 -text {*
  37.394 -  \medskip The following lemma is a property of linear forms on real
  37.395 -  vector spaces. It will be used for the lemma @{text abs_HahnBanach}
  37.396 -  (see page \pageref{abs-HahnBanach}). \label{abs-ineq-iff} For real
  37.397 -  vector spaces the following inequations are equivalent:
  37.398 -  \begin{center}
  37.399 -  \begin{tabular}{lll}
  37.400 -  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
  37.401 -  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
  37.402 -  \end{tabular}
  37.403 -  \end{center}
  37.404 -*}
  37.405 -
  37.406 -lemma abs_ineq_iff:
  37.407 -  assumes "subspace H E" and "vectorspace E" and "seminorm E p"
  37.408 -    and "linearform H h"
  37.409 -  shows "(\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x) = (\<forall>x \<in> H. h x \<le> p x)" (is "?L = ?R")
  37.410 -proof
  37.411 -  interpret subspace [H E] by fact
  37.412 -  interpret vectorspace [E] by fact
  37.413 -  interpret seminorm [E p] by fact
  37.414 -  interpret linearform [H h] by fact
  37.415 -  have H: "vectorspace H" using `vectorspace E` ..
  37.416 -  {
  37.417 -    assume l: ?L
  37.418 -    show ?R
  37.419 -    proof
  37.420 -      fix x assume x: "x \<in> H"
  37.421 -      have "h x \<le> \<bar>h x\<bar>" by arith
  37.422 -      also from l x have "\<dots> \<le> p x" ..
  37.423 -      finally show "h x \<le> p x" .
  37.424 -    qed
  37.425 -  next
  37.426 -    assume r: ?R
  37.427 -    show ?L
  37.428 -    proof
  37.429 -      fix x assume x: "x \<in> H"
  37.430 -      show "\<And>a b :: real. - a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> \<bar>b\<bar> \<le> a"
  37.431 -        by arith
  37.432 -      from `linearform H h` and H x
  37.433 -      have "- h x = h (- x)" by (rule linearform.neg [symmetric])
  37.434 -      also
  37.435 -      from H x have "- x \<in> H" by (rule vectorspace.neg_closed)
  37.436 -      with r have "h (- x) \<le> p (- x)" ..
  37.437 -      also have "\<dots> = p x"
  37.438 -	using `seminorm E p` `vectorspace E`
  37.439 -      proof (rule seminorm.minus)
  37.440 -        from x show "x \<in> E" ..
  37.441 -      qed
  37.442 -      finally have "- h x \<le> p x" .
  37.443 -      then show "- p x \<le> h x" by simp
  37.444 -      from r x show "h x \<le> p x" ..
  37.445 -    qed
  37.446 -  }
  37.447 -qed
  37.448 -
  37.449 -end
    38.1 --- a/src/HOL/Real/HahnBanach/Linearform.thy	Mon Dec 29 13:23:53 2008 +0100
    38.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    38.3 @@ -1,61 +0,0 @@
    38.4 -(*  Title:      HOL/Real/HahnBanach/Linearform.thy
    38.5 -    ID:         $Id$
    38.6 -    Author:     Gertrud Bauer, TU Munich
    38.7 -*)
    38.8 -
    38.9 -header {* Linearforms *}
   38.10 -
   38.11 -theory Linearform
   38.12 -imports VectorSpace
   38.13 -begin
   38.14 -
   38.15 -text {*
   38.16 -  A \emph{linear form} is a function on a vector space into the reals
   38.17 -  that is additive and multiplicative.
   38.18 -*}
   38.19 -
   38.20 -locale linearform = var V + var f +
   38.21 -  constrains V :: "'a\<Colon>{minus, plus, zero, uminus} set"
   38.22 -  assumes add [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x + y) = f x + f y"
   38.23 -    and mult [iff]: "x \<in> V \<Longrightarrow> f (a \<cdot> x) = a * f x"
   38.24 -
   38.25 -declare linearform.intro [intro?]
   38.26 -
   38.27 -lemma (in linearform) neg [iff]:
   38.28 -  assumes "vectorspace V"
   38.29 -  shows "x \<in> V \<Longrightarrow> f (- x) = - f x"
   38.30 -proof -
   38.31 -  interpret vectorspace [V] by fact
   38.32 -  assume x: "x \<in> V"
   38.33 -  then have "f (- x) = f ((- 1) \<cdot> x)" by (simp add: negate_eq1)
   38.34 -  also from x have "\<dots> = (- 1) * (f x)" by (rule mult)
   38.35 -  also from x have "\<dots> = - (f x)" by simp
   38.36 -  finally show ?thesis .
   38.37 -qed
   38.38 -
   38.39 -lemma (in linearform) diff [iff]:
   38.40 -  assumes "vectorspace V"
   38.41 -  shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x - y) = f x - f y"
   38.42 -proof -
   38.43 -  interpret vectorspace [V] by fact
   38.44 -  assume x: "x \<in> V" and y: "y \<in> V"
   38.45 -  then have "x - y = x + - y" by (rule diff_eq1)
   38.46 -  also have "f \<dots> = f x + f (- y)" by (rule add) (simp_all add: x y)
   38.47 -  also have "f (- y) = - f y" using `vectorspace V` y by (rule neg)
   38.48 -  finally show ?thesis by simp
   38.49 -qed
   38.50 -
   38.51 -text {* Every linear form yields @{text 0} for the @{text 0} vector. *}
   38.52 -
   38.53 -lemma (in linearform) zero [iff]:
   38.54 -  assumes "vectorspace V"
   38.55 -  shows "f 0 = 0"
   38.56 -proof -
   38.57 -  interpret vectorspace [V] by fact
   38.58 -  have "f 0 = f (0 - 0)" by simp
   38.59 -  also have "\<dots> = f 0 - f 0" using `vectorspace V` by (rule diff) simp_all
   38.60 -  also have "\<dots> = 0" by simp
   38.61 -  finally show ?thesis .
   38.62 -qed
   38.63 -
   38.64 -end
    39.1 --- a/src/HOL/Real/HahnBanach/NormedSpace.thy	Mon Dec 29 13:23:53 2008 +0100
    39.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    39.3 @@ -1,118 +0,0 @@
    39.4 -(*  Title:      HOL/Real/HahnBanach/NormedSpace.thy
    39.5 -    ID:         $Id$
    39.6 -    Author:     Gertrud Bauer, TU Munich
    39.7 -*)
    39.8 -
    39.9 -header {* Normed vector spaces *}
   39.10 -
   39.11 -theory NormedSpace
   39.12 -imports Subspace
   39.13 -begin
   39.14 -
   39.15 -subsection {* Quasinorms *}
   39.16 -
   39.17 -text {*
   39.18 -  A \emph{seminorm} @{text "\<parallel>\<cdot>\<parallel>"} is a function on a real vector space
   39.19 -  into the reals that has the following properties: it is positive
   39.20 -  definite, absolute homogenous and subadditive.
   39.21 -*}
   39.22 -
   39.23 -locale norm_syntax =
   39.24 -  fixes norm :: "'a \<Rightarrow> real"    ("\<parallel>_\<parallel>")
   39.25 -
   39.26 -locale seminorm = var V + norm_syntax +
   39.27 -  constrains V :: "'a\<Colon>{minus, plus, zero, uminus} set"
   39.28 -  assumes ge_zero [iff?]: "x \<in> V \<Longrightarrow> 0 \<le> \<parallel>x\<parallel>"
   39.29 -    and abs_homogenous [iff?]: "x \<in> V \<Longrightarrow> \<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>"
   39.30 -    and subadditive [iff?]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> \<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>"
   39.31 -
   39.32 -declare seminorm.intro [intro?]
   39.33 -
   39.34 -lemma (in seminorm) diff_subadditive:
   39.35 -  assumes "vectorspace V"
   39.36 -  shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> \<parallel>x - y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>"
   39.37 -proof -
   39.38 -  interpret vectorspace [V] by fact
   39.39 -  assume x: "x \<in> V" and y: "y \<in> V"
   39.40 -  then have "x - y = x + - 1 \<cdot> y"
   39.41 -    by (simp add: diff_eq2 negate_eq2a)
   39.42 -  also from x y have "\<parallel>\<dots>\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>- 1 \<cdot> y\<parallel>"
   39.43 -    by (simp add: subadditive)
   39.44 -  also from y have "\<parallel>- 1 \<cdot> y\<parallel> = \<bar>- 1\<bar> * \<parallel>y\<parallel>"
   39.45 -    by (rule abs_homogenous)
   39.46 -  also have "\<dots> = \<parallel>y\<parallel>" by simp
   39.47 -  finally show ?thesis .
   39.48 -qed
   39.49 -
   39.50 -lemma (in seminorm) minus:
   39.51 -  assumes "vectorspace V"
   39.52 -  shows "x \<in> V \<Longrightarrow> \<parallel>- x\<parallel> = \<parallel>x\<parallel>"
   39.53 -proof -
   39.54 -  interpret vectorspace [V] by fact
   39.55 -  assume x: "x \<in> V"
   39.56 -  then have "- x = - 1 \<cdot> x" by (simp only: negate_eq1)
   39.57 -  also from x have "\<parallel>\<dots>\<parallel> = \<bar>- 1\<bar> * \<parallel>x\<parallel>"
   39.58 -    by (rule abs_homogenous)
   39.59 -  also have "\<dots> = \<parallel>x\<parallel>" by simp
   39.60 -  finally show ?thesis .
   39.61 -qed
   39.62 -
   39.63 -
   39.64 -subsection {* Norms *}
   39.65 -
   39.66 -text {*
   39.67 -  A \emph{norm} @{text "\<parallel>\<cdot>\<parallel>"} is a seminorm that maps only the
   39.68 -  @{text 0} vector to @{text 0}.
   39.69 -*}
   39.70 -
   39.71 -locale norm = seminorm +
   39.72 -  assumes zero_iff [iff]: "x \<in> V \<Longrightarrow> (\<parallel>x\<parallel> = 0) = (x = 0)"
   39.73 -
   39.74 -
   39.75 -subsection {* Normed vector spaces *}
   39.76 -
   39.77 -text {*
   39.78 -  A vector space together with a norm is called a \emph{normed
   39.79 -  space}.
   39.80 -*}
   39.81 -
   39.82 -locale normed_vectorspace = vectorspace + norm
   39.83 -
   39.84 -declare normed_vectorspace.intro [intro?]
   39.85 -
   39.86 -lemma (in normed_vectorspace) gt_zero [intro?]:
   39.87 -  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> 0 < \<parallel>x\<parallel>"
   39.88 -proof -
   39.89 -  assume x: "x \<in> V" and neq: "x \<noteq> 0"
   39.90 -  from x have "0 \<le> \<parallel>x\<parallel>" ..
   39.91 -  also have [symmetric]: "\<dots> \<noteq> 0"
   39.92 -  proof
   39.93 -    assume "\<parallel>x\<parallel> = 0"
   39.94 -    with x have "x = 0" by simp
   39.95 -    with neq show False by contradiction
   39.96 -  qed
   39.97 -  finally show ?thesis .
   39.98 -qed
   39.99 -
  39.100 -text {*
  39.101 -  Any subspace of a normed vector space is again a normed vectorspace.
  39.102 -*}
  39.103 -
  39.104 -lemma subspace_normed_vs [intro?]:
  39.105 -  fixes F E norm
  39.106 -  assumes "subspace F E" "normed_vectorspace E norm"
  39.107 -  shows "normed_vectorspace F norm"
  39.108 -proof -
  39.109 -  interpret subspace [F E] by fact
  39.110 -  interpret normed_vectorspace [E norm] by fact
  39.111 -  show ?thesis
  39.112 -  proof
  39.113 -    show "vectorspace F" by (rule vectorspace) unfold_locales
  39.114 -  next
  39.115 -    have "NormedSpace.norm E norm" ..
  39.116 -    with subset show "NormedSpace.norm F norm"
  39.117 -      by (simp add: norm_def seminorm_def norm_axioms_def)
  39.118 -  qed
  39.119 -qed
  39.120 -
  39.121 -end
    40.1 --- a/src/HOL/Real/HahnBanach/README.html	Mon Dec 29 13:23:53 2008 +0100
    40.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    40.3 @@ -1,38 +0,0 @@
    40.4 -<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
    40.5 -
    40.6 -<!-- $Id$ -->
    40.7 -
    40.8 -<HTML>
    40.9 -
   40.10 -<HEAD>
   40.11 -  <meta http-equiv="Content-Type" content="text/html; charset=iso-8859-1">
   40.12 -  <TITLE>HOL/Real/HahnBanach/README</TITLE>
   40.13 -</HEAD>
   40.14 -
   40.15 -<BODY>
   40.16 -
   40.17 -<H3>The Hahn-Banach Theorem for Real Vector Spaces (Isabelle/Isar)</H3>
   40.18 -
   40.19 -Author: Gertrud Bauer, Technische Universit&auml;t M&uuml;nchen<P>
   40.20 -
   40.21 -This directory contains the proof of the Hahn-Banach theorem for real vectorspaces,
   40.22 -following H. Heuser, Funktionalanalysis, p. 228 -232.
   40.23 -The Hahn-Banach theorem is one of the fundamental theorems of functioal analysis.
   40.24 -It is a conclusion of Zorn's lemma.<P>
   40.25 -
   40.26 -Two different formaulations of the theorem are presented, one for general real vectorspaces
   40.27 -and its application to normed vectorspaces. <P>
   40.28 -
   40.29 -The theorem says, that every continous linearform, defined on arbitrary subspaces
   40.30 -(not only one-dimensional subspaces), can be extended to a continous linearform on
   40.31 -the whole vectorspace.
   40.32 -
   40.33 -
   40.34 -<HR>
   40.35 -
   40.36 -<ADDRESS>
   40.37 -<A NAME="bauerg@in.tum.de" HREF="mailto:bauerg@in.tum.de">bauerg@in.tum.de</A>
   40.38 -</ADDRESS>
   40.39 -
   40.40 -</BODY>
   40.41 -</HTML>
    41.1 --- a/src/HOL/Real/HahnBanach/ROOT.ML	Mon Dec 29 13:23:53 2008 +0100
    41.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    41.3 @@ -1,8 +0,0 @@
    41.4 -(*  Title:      HOL/Real/HahnBanach/ROOT.ML
    41.5 -    ID:         $Id$
    41.6 -    Author:     Gertrud Bauer, TU Munich
    41.7 -
    41.8 -The Hahn-Banach theorem for real vector spaces (Isabelle/Isar).
    41.9 -*)
   41.10 -
   41.11 -time_use_thy "HahnBanach";
    42.1 --- a/src/HOL/Real/HahnBanach/Subspace.thy	Mon Dec 29 13:23:53 2008 +0100
    42.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    42.3 @@ -1,514 +0,0 @@
    42.4 -(*  Title:      HOL/Real/HahnBanach/Subspace.thy
    42.5 -    ID:         $Id$
    42.6 -    Author:     Gertrud Bauer, TU Munich
    42.7 -*)
    42.8 -
    42.9 -header {* Subspaces *}
   42.10 -
   42.11 -theory Subspace
   42.12 -imports VectorSpace
   42.13 -begin
   42.14 -
   42.15 -subsection {* Definition *}
   42.16 -
   42.17 -text {*
   42.18 -  A non-empty subset @{text U} of a vector space @{text V} is a
   42.19 -  \emph{subspace} of @{text V}, iff @{text U} is closed under addition
   42.20 -  and scalar multiplication.
   42.21 -*}
   42.22 -
   42.23 -locale subspace = var U + var V +
   42.24 -  constrains U :: "'a\<Colon>{minus, plus, zero, uminus} set"
   42.25 -  assumes non_empty [iff, intro]: "U \<noteq> {}"
   42.26 -    and subset [iff]: "U \<subseteq> V"
   42.27 -    and add_closed [iff]: "x \<in> U \<Longrightarrow> y \<in> U \<Longrightarrow> x + y \<in> U"
   42.28 -    and mult_closed [iff]: "x \<in> U \<Longrightarrow> a \<cdot> x \<in> U"
   42.29 -
   42.30 -notation (symbols)
   42.31 -  subspace  (infix "\<unlhd>" 50)
   42.32 -
   42.33 -declare vectorspace.intro [intro?] subspace.intro [intro?]
   42.34 -
   42.35 -lemma subspace_subset [elim]: "U \<unlhd> V \<Longrightarrow> U \<subseteq> V"
   42.36 -  by (rule subspace.subset)
   42.37 -
   42.38 -lemma (in subspace) subsetD [iff]: "x \<in> U \<Longrightarrow> x \<in> V"
   42.39 -  using subset by blast
   42.40 -
   42.41 -lemma subspaceD [elim]: "U \<unlhd> V \<Longrightarrow> x \<in> U \<Longrightarrow> x \<in> V"
   42.42 -  by (rule subspace.subsetD)
   42.43 -
   42.44 -lemma rev_subspaceD [elim?]: "x \<in> U \<Longrightarrow> U \<unlhd> V \<Longrightarrow> x \<in> V"
   42.45 -  by (rule subspace.subsetD)
   42.46 -
   42.47 -lemma (in subspace) diff_closed [iff]:
   42.48 -  assumes "vectorspace V"
   42.49 -  assumes x: "x \<in> U" and y: "y \<in> U"
   42.50 -  shows "x - y \<in> U"
   42.51 -proof -
   42.52 -  interpret vectorspace [V] by fact
   42.53 -  from x y show ?thesis by (simp add: diff_eq1 negate_eq1)
   42.54 -qed
   42.55 -
   42.56 -text {*
   42.57 -  \medskip Similar as for linear spaces, the existence of the zero
   42.58 -  element in every subspace follows from the non-emptiness of the
   42.59 -  carrier set and by vector space laws.
   42.60 -*}
   42.61 -
   42.62 -lemma (in subspace) zero [intro]:
   42.63 -  assumes "vectorspace V"
   42.64 -  shows "0 \<in> U"
   42.65 -proof -
   42.66 -  interpret vectorspace [V] by fact
   42.67 -  have "U \<noteq> {}" by (rule U_V.non_empty)
   42.68 -  then obtain x where x: "x \<in> U" by blast
   42.69 -  then have "x \<in> V" .. then have "0 = x - x" by simp
   42.70 -  also from `vectorspace V` x x have "\<dots> \<in> U" by (rule U_V.diff_closed)
   42.71 -  finally show ?thesis .
   42.72 -qed
   42.73 -
   42.74 -lemma (in subspace) neg_closed [iff]:
   42.75 -  assumes "vectorspace V"
   42.76 -  assumes x: "x \<in> U"
   42.77 -  shows "- x \<in> U"
   42.78 -proof -
   42.79 -  interpret vectorspace [V] by fact
   42.80 -  from x show ?thesis by (simp add: negate_eq1)
   42.81 -qed
   42.82 -
   42.83 -text {* \medskip Further derived laws: every subspace is a vector space. *}
   42.84 -
   42.85 -lemma (in subspace) vectorspace [iff]:
   42.86 -  assumes "vectorspace V"
   42.87 -  shows "vectorspace U"
   42.88 -proof -
   42.89 -  interpret vectorspace [V] by fact
   42.90 -  show ?thesis
   42.91 -  proof
   42.92 -    show "U \<noteq> {}" ..
   42.93 -    fix x y z assume x: "x \<in> U" and y: "y \<in> U" and z: "z \<in> U"
   42.94 -    fix a b :: real
   42.95 -    from x y show "x + y \<in> U" by simp
   42.96 -    from x show "a \<cdot> x \<in> U" by simp
   42.97 -    from x y z show "(x + y) + z = x + (y + z)" by (simp add: add_ac)
   42.98 -    from x y show "x + y = y + x" by (simp add: add_ac)
   42.99 -    from x show "x - x = 0" by simp
  42.100 -    from x show "0 + x = x" by simp
  42.101 -    from x y show "a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y" by (simp add: distrib)
  42.102 -    from x show "(a + b) \<cdot> x = a \<cdot> x + b \<cdot> x" by (simp add: distrib)
  42.103 -    from x show "(a * b) \<cdot> x = a \<cdot> b \<cdot> x" by (simp add: mult_assoc)
  42.104 -    from x show "1 \<cdot> x = x" by simp
  42.105 -    from x show "- x = - 1 \<cdot> x" by (simp add: negate_eq1)
  42.106 -    from x y show "x - y = x + - y" by (simp add: diff_eq1)
  42.107 -  qed
  42.108 -qed
  42.109 -
  42.110 -
  42.111 -text {* The subspace relation is reflexive. *}
  42.112 -
  42.113 -lemma (in vectorspace) subspace_refl [intro]: "V \<unlhd> V"
  42.114 -proof
  42.115 -  show "V \<noteq> {}" ..
  42.116 -  show "V \<subseteq> V" ..
  42.117 -  fix x y assume x: "x \<in> V" and y: "y \<in> V"
  42.118 -  fix a :: real
  42.119 -  from x y show "x + y \<in> V" by simp
  42.120 -  from x show "a \<cdot> x \<in> V" by simp
  42.121 -qed
  42.122 -
  42.123 -text {* The subspace relation is transitive. *}
  42.124 -
  42.125 -lemma (in vectorspace) subspace_trans [trans]:
  42.126 -  "U \<unlhd> V \<Longrightarrow> V \<unlhd> W \<Longrightarrow> U \<unlhd> W"
  42.127 -proof
  42.128 -  assume uv: "U \<unlhd> V" and vw: "V \<unlhd> W"
  42.129 -  from uv show "U \<noteq> {}" by (rule subspace.non_empty)
  42.130 -  show "U \<subseteq> W"
  42.131 -  proof -
  42.132 -    from uv have "U \<subseteq> V" by (rule subspace.subset)
  42.133 -    also from vw have "V \<subseteq> W" by (rule subspace.subset)
  42.134 -    finally show ?thesis .
  42.135 -  qed
  42.136 -  fix x y assume x: "x \<in> U" and y: "y \<in> U"
  42.137 -  from uv and x y show "x + y \<in> U" by (rule subspace.add_closed)
  42.138 -  from uv and x show "\<And>a. a \<cdot> x \<in> U" by (rule subspace.mult_closed)
  42.139 -qed
  42.140 -
  42.141 -
  42.142 -subsection {* Linear closure *}
  42.143 -
  42.144 -text {*
  42.145 -  The \emph{linear closure} of a vector @{text x} is the set of all
  42.146 -  scalar multiples of @{text x}.
  42.147 -*}
  42.148 -
  42.149 -definition
  42.150 -  lin :: "('a::{minus, plus, zero}) \<Rightarrow> 'a set" where
  42.151 -  "lin x = {a \<cdot> x | a. True}"
  42.152 -
  42.153 -lemma linI [intro]: "y = a \<cdot> x \<Longrightarrow> y \<in> lin x"
  42.154 -  unfolding lin_def by blast
  42.155 -
  42.156 -lemma linI' [iff]: "a \<cdot> x \<in> lin x"
  42.157 -  unfolding lin_def by blast
  42.158 -
  42.159 -lemma linE [elim]: "x \<in> lin v \<Longrightarrow> (\<And>a::real. x = a \<cdot> v \<Longrightarrow> C) \<Longrightarrow> C"
  42.160 -  unfolding lin_def by blast
  42.161 -
  42.162 -
  42.163 -text {* Every vector is contained in its linear closure. *}
  42.164 -
  42.165 -lemma (in vectorspace) x_lin_x [iff]: "x \<in> V \<Longrightarrow> x \<in> lin x"
  42.166 -proof -
  42.167 -  assume "x \<in> V"
  42.168 -  then have "x = 1 \<cdot> x" by simp
  42.169 -  also have "\<dots> \<in> lin x" ..
  42.170 -  finally show ?thesis .
  42.171 -qed
  42.172 -
  42.173 -lemma (in vectorspace) "0_lin_x" [iff]: "x \<in> V \<Longrightarrow> 0 \<in> lin x"
  42.174 -proof
  42.175 -  assume "x \<in> V"
  42.176 -  then show "0 = 0 \<cdot> x" by simp
  42.177 -qed
  42.178 -
  42.179 -text {* Any linear closure is a subspace. *}
  42.180 -
  42.181 -lemma (in vectorspace) lin_subspace [intro]:
  42.182 -  "x \<in> V \<Longrightarrow> lin x \<unlhd> V"
  42.183 -proof
  42.184 -  assume x: "x \<in> V"
  42.185 -  then show "lin x \<noteq> {}" by (auto simp add: x_lin_x)
  42.186 -  show "lin x \<subseteq> V"
  42.187 -  proof
  42.188 -    fix x' assume "x' \<in> lin x"
  42.189 -    then obtain a where "x' = a \<cdot> x" ..
  42.190 -    with x show "x' \<in> V" by simp
  42.191 -  qed
  42.192 -  fix x' x'' assume x': "x' \<in> lin x" and x'': "x'' \<in> lin x"
  42.193 -  show "x' + x'' \<in> lin x"
  42.194 -  proof -
  42.195 -    from x' obtain a' where "x' = a' \<cdot> x" ..
  42.196 -    moreover from x'' obtain a'' where "x'' = a'' \<cdot> x" ..
  42.197 -    ultimately have "x' + x'' = (a' + a'') \<cdot> x"
  42.198 -      using x by (simp add: distrib)
  42.199 -    also have "\<dots> \<in> lin x" ..
  42.200 -    finally show ?thesis .
  42.201 -  qed
  42.202 -  fix a :: real
  42.203 -  show "a \<cdot> x' \<in> lin x"
  42.204 -  proof -
  42.205 -    from x' obtain a' where "x' = a' \<cdot> x" ..
  42.206 -    with x have "a \<cdot> x' = (a * a') \<cdot> x" by (simp add: mult_assoc)
  42.207 -    also have "\<dots> \<in> lin x" ..
  42.208 -    finally show ?thesis .
  42.209 -  qed
  42.210 -qed
  42.211 -
  42.212 -
  42.213 -text {* Any linear closure is a vector space. *}
  42.214 -
  42.215 -lemma (in vectorspace) lin_vectorspace [intro]:
  42.216 -  assumes "x \<in> V"
  42.217 -  shows "vectorspace (lin x)"
  42.218 -proof -
  42.219 -  from `x \<in> V` have "subspace (lin x) V"
  42.220 -    by (rule lin_subspace)
  42.221 -  from this and vectorspace_axioms show ?thesis
  42.222 -    by (rule subspace.vectorspace)
  42.223 -qed
  42.224 -
  42.225 -
  42.226 -subsection {* Sum of two vectorspaces *}
  42.227 -
  42.228 -text {*
  42.229 -  The \emph{sum} of two vectorspaces @{text U} and @{text V} is the
  42.230 -  set of all sums of elements from @{text U} and @{text V}.
  42.231 -*}
  42.232 -
  42.233 -instantiation "fun" :: (type, type) plus
  42.234 -begin
  42.235 -
  42.236 -definition
  42.237 -  sum_def: "plus_fun U V = {u + v | u v. u \<in> U \<and> v \<in> V}"  (* FIXME not fully general!? *)
  42.238 -
  42.239 -instance ..
  42.240 -
  42.241 -end
  42.242 -
  42.243 -lemma sumE [elim]:
  42.244 -    "x \<in> U + V \<Longrightarrow> (\<And>u v. x = u + v \<Longrightarrow> u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> C) \<Longrightarrow> C"
  42.245 -  unfolding sum_def by blast
  42.246 -
  42.247 -lemma sumI [intro]:
  42.248 -    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> x = u + v \<Longrightarrow> x \<in> U + V"
  42.249 -  unfolding sum_def by blast
  42.250 -
  42.251 -lemma sumI' [intro]:
  42.252 -    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> u + v \<in> U + V"
  42.253 -  unfolding sum_def by blast
  42.254 -
  42.255 -text {* @{text U} is a subspace of @{text "U + V"}. *}
  42.256 -
  42.257 -lemma subspace_sum1 [iff]:
  42.258 -  assumes "vectorspace U" "vectorspace V"
  42.259 -  shows "U \<unlhd> U + V"
  42.260 -proof -
  42.261 -  interpret vectorspace [U] by fact
  42.262 -  interpret vectorspace [V] by fact
  42.263 -  show ?thesis
  42.264 -  proof
  42.265 -    show "U \<noteq> {}" ..
  42.266 -    show "U \<subseteq> U + V"
  42.267 -    proof
  42.268 -      fix x assume x: "x \<in> U"
  42.269 -      moreover have "0 \<in> V" ..
  42.270 -      ultimately have "x + 0 \<in> U + V" ..
  42.271 -      with x show "x \<in> U + V" by simp
  42.272 -    qed
  42.273 -    fix x y assume x: "x \<in> U" and "y \<in> U"
  42.274 -    then show "x + y \<in> U" by simp
  42.275 -    from x show "\<And>a. a \<cdot> x \<in> U" by simp
  42.276 -  qed
  42.277 -qed
  42.278 -
  42.279 -text {* The sum of two subspaces is again a subspace. *}
  42.280 -
  42.281 -lemma sum_subspace [intro?]:
  42.282 -  assumes "subspace U E" "vectorspace E" "subspace V E"
  42.283 -  shows "U + V \<unlhd> E"
  42.284 -proof -
  42.285 -  interpret subspace [U E] by fact
  42.286 -  interpret vectorspace [E] by fact
  42.287 -  interpret subspace [V E] by fact
  42.288 -  show ?thesis
  42.289 -  proof
  42.290 -    have "0 \<in> U + V"
  42.291 -    proof
  42.292 -      show "0 \<in> U" using `vectorspace E` ..
  42.293 -      show "0 \<in> V" using `vectorspace E` ..
  42.294 -      show "(0::'a) = 0 + 0" by simp
  42.295 -    qed
  42.296 -    then show "U + V \<noteq> {}" by blast
  42.297 -    show "U + V \<subseteq> E"
  42.298 -    proof
  42.299 -      fix x assume "x \<in> U + V"
  42.300 -      then obtain u v where "x = u + v" and
  42.301 -	"u \<in> U" and "v \<in> V" ..
  42.302 -      then show "x \<in> E" by simp
  42.303 -    qed
  42.304 -    fix x y assume x: "x \<in> U + V" and y: "y \<in> U + V"
  42.305 -    show "x + y \<in> U + V"
  42.306 -    proof -
  42.307 -      from x obtain ux vx where "x = ux + vx" and "ux \<in> U" and "vx \<in> V" ..
  42.308 -      moreover
  42.309 -      from y obtain uy vy where "y = uy + vy" and "uy \<in> U" and "vy \<in> V" ..
  42.310 -      ultimately
  42.311 -      have "ux + uy \<in> U"
  42.312 -	and "vx + vy \<in> V"
  42.313 -	and "x + y = (ux + uy) + (vx + vy)"
  42.314 -	using x y by (simp_all add: add_ac)
  42.315 -      then show ?thesis ..
  42.316 -    qed
  42.317 -    fix a show "a \<cdot> x \<in> U + V"
  42.318 -    proof -
  42.319 -      from x obtain u v where "x = u + v" and "u \<in> U" and "v \<in> V" ..
  42.320 -      then have "a \<cdot> u \<in> U" and "a \<cdot> v \<in> V"
  42.321 -	and "a \<cdot> x = (a \<cdot> u) + (a \<cdot> v)" by (simp_all add: distrib)
  42.322 -      then show ?thesis ..
  42.323 -    qed
  42.324 -  qed
  42.325 -qed
  42.326 -
  42.327 -text{* The sum of two subspaces is a vectorspace. *}
  42.328 -
  42.329 -lemma sum_vs [intro?]:
  42.330 -    "U \<unlhd> E \<Longrightarrow> V \<unlhd> E \<Longrightarrow> vectorspace E \<Longrightarrow> vectorspace (U + V)"
  42.331 -  by (rule subspace.vectorspace) (rule sum_subspace)
  42.332 -
  42.333 -
  42.334 -subsection {* Direct sums *}
  42.335 -
  42.336 -text {*
  42.337 -  The sum of @{text U} and @{text V} is called \emph{direct}, iff the
  42.338 -  zero element is the only common element of @{text U} and @{text
  42.339 -  V}. For every element @{text x} of the direct sum of @{text U} and
  42.340 -  @{text V} the decomposition in @{text "x = u + v"} with
  42.341 -  @{text "u \<in> U"} and @{text "v \<in> V"} is unique.
  42.342 -*}
  42.343 -
  42.344 -lemma decomp:
  42.345 -  assumes "vectorspace E" "subspace U E" "subspace V E"
  42.346 -  assumes direct: "U \<inter> V = {0}"
  42.347 -    and u1: "u1 \<in> U" and u2: "u2 \<in> U"
  42.348 -    and v1: "v1 \<in> V" and v2: "v2 \<in> V"
  42.349 -    and sum: "u1 + v1 = u2 + v2"
  42.350 -  shows "u1 = u2 \<and> v1 = v2"
  42.351 -proof -
  42.352 -  interpret vectorspace [E] by fact
  42.353 -  interpret subspace [U E] by fact
  42.354 -  interpret subspace [V E] by fact
  42.355 -  show ?thesis
  42.356 -  proof
  42.357 -    have U: "vectorspace U"  (* FIXME: use interpret *)
  42.358 -      using `subspace U E` `vectorspace E` by (rule subspace.vectorspace)
  42.359 -    have V: "vectorspace V"
  42.360 -      using `subspace V E` `vectorspace E` by (rule subspace.vectorspace)
  42.361 -    from u1 u2 v1 v2 and sum have eq: "u1 - u2 = v2 - v1"
  42.362 -      by (simp add: add_diff_swap)
  42.363 -    from u1 u2 have u: "u1 - u2 \<in> U"
  42.364 -      by (rule vectorspace.diff_closed [OF U])
  42.365 -    with eq have v': "v2 - v1 \<in> U" by (simp only:)
  42.366 -    from v2 v1 have v: "v2 - v1 \<in> V"
  42.367 -      by (rule vectorspace.diff_closed [OF V])
  42.368 -    with eq have u': " u1 - u2 \<in> V" by (simp only:)
  42.369 -    
  42.370 -    show "u1 = u2"
  42.371 -    proof (rule add_minus_eq)
  42.372 -      from u1 show "u1 \<in> E" ..
  42.373 -      from u2 show "u2 \<in> E" ..
  42.374 -      from u u' and direct show "u1 - u2 = 0" by blast
  42.375 -    qed
  42.376 -    show "v1 = v2"
  42.377 -    proof (rule add_minus_eq [symmetric])
  42.378 -      from v1 show "v1 \<in> E" ..
  42.379 -      from v2 show "v2 \<in> E" ..
  42.380 -      from v v' and direct show "v2 - v1 = 0" by blast
  42.381 -    qed
  42.382 -  qed
  42.383 -qed
  42.384 -
  42.385 -text {*
  42.386 -  An application of the previous lemma will be used in the proof of
  42.387 -  the Hahn-Banach Theorem (see page \pageref{decomp-H-use}): for any
  42.388 -  element @{text "y + a \<cdot> x\<^sub>0"} of the direct sum of a
  42.389 -  vectorspace @{text H} and the linear closure of @{text "x\<^sub>0"}
  42.390 -  the components @{text "y \<in> H"} and @{text a} are uniquely
  42.391 -  determined.
  42.392 -*}
  42.393 -
  42.394 -lemma decomp_H':
  42.395 -  assumes "vectorspace E" "subspace H E"
  42.396 -  assumes y1: "y1 \<in> H" and y2: "y2 \<in> H"
  42.397 -    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
  42.398 -    and eq: "y1 + a1 \<cdot> x' = y2 + a2 \<cdot> x'"
  42.399 -  shows "y1 = y2 \<and> a1 = a2"
  42.400 -proof -
  42.401 -  interpret vectorspace [E] by fact
  42.402 -  interpret subspace [H E] by fact
  42.403 -  show ?thesis
  42.404 -  proof
  42.405 -    have c: "y1 = y2 \<and> a1 \<cdot> x' = a2 \<cdot> x'"
  42.406 -    proof (rule decomp)
  42.407 -      show "a1 \<cdot> x' \<in> lin x'" ..
  42.408 -      show "a2 \<cdot> x' \<in> lin x'" ..
  42.409 -      show "H \<inter> lin x' = {0}"
  42.410 -      proof
  42.411 -	show "H \<inter> lin x' \<subseteq> {0}"
  42.412 -	proof
  42.413 -          fix x assume x: "x \<in> H \<inter> lin x'"
  42.414 -          then obtain a where xx': "x = a \<cdot> x'"
  42.415 -            by blast
  42.416 -          have "x = 0"
  42.417 -          proof cases
  42.418 -            assume "a = 0"
  42.419 -            with xx' and x' show ?thesis by simp
  42.420 -          next
  42.421 -            assume a: "a \<noteq> 0"
  42.422 -            from x have "x \<in> H" ..
  42.423 -            with xx' have "inverse a \<cdot> a \<cdot> x' \<in> H" by simp
  42.424 -            with a and x' have "x' \<in> H" by (simp add: mult_assoc2)
  42.425 -            with `x' \<notin> H` show ?thesis by contradiction
  42.426 -          qed
  42.427 -          then show "x \<in> {0}" ..
  42.428 -	qed
  42.429 -	show "{0} \<subseteq> H \<inter> lin x'"
  42.430 -	proof -
  42.431 -          have "0 \<in> H" using `vectorspace E` ..
  42.432 -          moreover have "0 \<in> lin x'" using `x' \<in> E` ..
  42.433 -          ultimately show ?thesis by blast
  42.434 -	qed
  42.435 -      qed
  42.436 -      show "lin x' \<unlhd> E" using `x' \<in> E` ..
  42.437 -    qed (rule `vectorspace E`, rule `subspace H E`, rule y1, rule y2, rule eq)
  42.438 -    then show "y1 = y2" ..
  42.439 -    from c have "a1 \<cdot> x' = a2 \<cdot> x'" ..
  42.440 -    with x' show "a1 = a2" by (simp add: mult_right_cancel)
  42.441 -  qed
  42.442 -qed
  42.443 -
  42.444 -text {*
  42.445 -  Since for any element @{text "y + a \<cdot> x'"} of the direct sum of a
  42.446 -  vectorspace @{text H} and the linear closure of @{text x'} the
  42.447 -  components @{text "y \<in> H"} and @{text a} are unique, it follows from
  42.448 -  @{text "y \<in> H"} that @{text "a = 0"}.
  42.449 -*}
  42.450 -
  42.451 -lemma decomp_H'_H:
  42.452 -  assumes "vectorspace E" "subspace H E"
  42.453 -  assumes t: "t \<in> H"
  42.454 -    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
  42.455 -  shows "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
  42.456 -proof -
  42.457 -  interpret vectorspace [E] by fact
  42.458 -  interpret subspace [H E] by fact
  42.459 -  show ?thesis
  42.460 -  proof (rule, simp_all only: split_paired_all split_conv)
  42.461 -    from t x' show "t = t + 0 \<cdot> x' \<and> t \<in> H" by simp
  42.462 -    fix y and a assume ya: "t = y + a \<cdot> x' \<and> y \<in> H"
  42.463 -    have "y = t \<and> a = 0"
  42.464 -    proof (rule decomp_H')
  42.465 -      from ya x' show "y + a \<cdot> x' = t + 0 \<cdot> x'" by simp
  42.466 -      from ya show "y \<in> H" ..
  42.467 -    qed (rule `vectorspace E`, rule `subspace H E`, rule t, (rule x')+)
  42.468 -    with t x' show "(y, a) = (y + a \<cdot> x', 0)" by simp
  42.469 -  qed
  42.470 -qed
  42.471 -
  42.472 -text {*
  42.473 -  The components @{text "y \<in> H"} and @{text a} in @{text "y + a \<cdot> x'"}
  42.474 -  are unique, so the function @{text h'} defined by
  42.475 -  @{text "h' (y + a \<cdot> x') = h y + a \<cdot> \<xi>"} is definite.
  42.476 -*}
  42.477 -
  42.478 -lemma h'_definite:
  42.479 -  fixes H
  42.480 -  assumes h'_def:
  42.481 -    "h' \<equiv> (\<lambda>x. let (y, a) = SOME (y, a). (x = y + a \<cdot> x' \<and> y \<in> H)
  42.482 -                in (h y) + a * xi)"
  42.483 -    and x: "x = y + a \<cdot> x'"
  42.484 -  assumes "vectorspace E" "subspace H E"
  42.485 -  assumes y: "y \<in> H"
  42.486 -    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
  42.487 -  shows "h' x = h y + a * xi"
  42.488 -proof -
  42.489 -  interpret vectorspace [E] by fact
  42.490 -  interpret subspace [H E] by fact
  42.491 -  from x y x' have "x \<in> H + lin x'" by auto
  42.492 -  have "\<exists>!p. (\<lambda>(y, a). x = y + a \<cdot> x' \<and> y \<in> H) p" (is "\<exists>!p. ?P p")
  42.493 -  proof (rule ex_ex1I)
  42.494 -    from x y show "\<exists>p. ?P p" by blast
  42.495 -    fix p q assume p: "?P p" and q: "?P q"
  42.496 -    show "p = q"
  42.497 -    proof -
  42.498 -      from p have xp: "x = fst p + snd p \<cdot> x' \<and> fst p \<in> H"
  42.499 -        by (cases p) simp
  42.500 -      from q have xq: "x = fst q + snd q \<cdot> x' \<and> fst q \<in> H"
  42.501 -        by (cases q) simp
  42.502 -      have "fst p = fst q \<and> snd p = snd q"
  42.503 -      proof (rule decomp_H')
  42.504 -        from xp show "fst p \<in> H" ..
  42.505 -        from xq show "fst q \<in> H" ..
  42.506 -        from xp and xq show "fst p + snd p \<cdot> x' = fst q + snd q \<cdot> x'"
  42.507 -          by simp
  42.508 -      qed (rule `vectorspace E`, rule `subspace H E`, (rule x')+)
  42.509 -      then show ?thesis by (cases p, cases q) simp
  42.510 -    qed
  42.511 -  qed
  42.512 -  then have eq: "(SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H) = (y, a)"
  42.513 -    by (rule some1_equality) (simp add: x y)
  42.514 -  with h'_def show "h' x = h y + a * xi" by (simp add: Let_def)
  42.515 -qed
  42.516 -
  42.517 -end
    43.1 --- a/src/HOL/Real/HahnBanach/VectorSpace.thy	Mon Dec 29 13:23:53 2008 +0100
    43.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    43.3 @@ -1,417 +0,0 @@
    43.4 -(*  Title:      HOL/Real/HahnBanach/VectorSpace.thy
    43.5 -    ID:         $Id$
    43.6 -    Author:     Gertrud Bauer, TU Munich
    43.7 -*)
    43.8 -
    43.9 -header {* Vector spaces *}
   43.10 -
   43.11 -theory VectorSpace
   43.12 -imports Real Bounds Zorn
   43.13 -begin
   43.14 -
   43.15 -subsection {* Signature *}
   43.16 -
   43.17 -text {*
   43.18 -  For the definition of real vector spaces a type @{typ 'a} of the
   43.19 -  sort @{text "{plus, minus, zero}"} is considered, on which a real
   43.20 -  scalar multiplication @{text \<cdot>} is declared.
   43.21 -*}
   43.22 -
   43.23 -consts
   43.24 -  prod  :: "real \<Rightarrow> 'a::{plus, minus, zero} \<Rightarrow> 'a"     (infixr "'(*')" 70)
   43.25 -
   43.26 -notation (xsymbols)
   43.27 -  prod  (infixr "\<cdot>" 70)
   43.28 -notation (HTML output)
   43.29 -  prod  (infixr "\<cdot>" 70)
   43.30 -
   43.31 -
   43.32 -subsection {* Vector space laws *}
   43.33 -
   43.34 -text {*
   43.35 -  A \emph{vector space} is a non-empty set @{text V} of elements from
   43.36 -  @{typ 'a} with the following vector space laws: The set @{text V} is
   43.37 -  closed under addition and scalar multiplication, addition is
   43.38 -  associative and commutative; @{text "- x"} is the inverse of @{text
   43.39 -  x} w.~r.~t.~addition and @{text 0} is the neutral element of
   43.40 -  addition.  Addition and multiplication are distributive; scalar
   43.41 -  multiplication is associative and the real number @{text "1"} is
   43.42 -  the neutral element of scalar multiplication.
   43.43 -*}
   43.44 -
   43.45 -locale vectorspace = var V +
   43.46 -  assumes non_empty [iff, intro?]: "V \<noteq> {}"
   43.47 -    and add_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y \<in> V"
   43.48 -    and mult_closed [iff]: "x \<in> V \<Longrightarrow> a \<cdot> x \<in> V"
   43.49 -    and add_assoc: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y) + z = x + (y + z)"
   43.50 -    and add_commute: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = y + x"
   43.51 -    and diff_self [simp]: "x \<in> V \<Longrightarrow> x - x = 0"
   43.52 -    and add_zero_left [simp]: "x \<in> V \<Longrightarrow> 0 + x = x"
   43.53 -    and add_mult_distrib1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y"
   43.54 -    and add_mult_distrib2: "x \<in> V \<Longrightarrow> (a + b) \<cdot> x = a \<cdot> x + b \<cdot> x"
   43.55 -    and mult_assoc: "x \<in> V \<Longrightarrow> (a * b) \<cdot> x = a \<cdot> (b \<cdot> x)"
   43.56 -    and mult_1 [simp]: "x \<in> V \<Longrightarrow> 1 \<cdot> x = x"
   43.57 -    and negate_eq1: "x \<in> V \<Longrightarrow> - x = (- 1) \<cdot> x"
   43.58 -    and diff_eq1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = x + - y"
   43.59 -
   43.60 -lemma (in vectorspace) negate_eq2: "x \<in> V \<Longrightarrow> (- 1) \<cdot> x = - x"
   43.61 -  by (rule negate_eq1 [symmetric])
   43.62 -
   43.63 -lemma (in vectorspace) negate_eq2a: "x \<in> V \<Longrightarrow> -1 \<cdot> x = - x"
   43.64 -  by (simp add: negate_eq1)
   43.65 -
   43.66 -lemma (in vectorspace) diff_eq2: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + - y = x - y"
   43.67 -  by (rule diff_eq1 [symmetric])
   43.68 -
   43.69 -lemma (in vectorspace) diff_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y \<in> V"
   43.70 -  by (simp add: diff_eq1 negate_eq1)
   43.71 -
   43.72 -lemma (in vectorspace) neg_closed [iff]: "x \<in> V \<Longrightarrow> - x \<in> V"
   43.73 -  by (simp add: negate_eq1)
   43.74 -
   43.75 -lemma (in vectorspace) add_left_commute:
   43.76 -  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> x + (y + z) = y + (x + z)"
   43.77 -proof -
   43.78 -  assume xyz: "x \<in> V"  "y \<in> V"  "z \<in> V"
   43.79 -  then have "x + (y + z) = (x + y) + z"
   43.80 -    by (simp only: add_assoc)
   43.81 -  also from xyz have "\<dots> = (y + x) + z" by (simp only: add_commute)
   43.82 -  also from xyz have "\<dots> = y + (x + z)" by (simp only: add_assoc)
   43.83 -  finally show ?thesis .
   43.84 -qed
   43.85 -
   43.86 -theorems (in vectorspace) add_ac =
   43.87 -  add_assoc add_commute add_left_commute
   43.88 -
   43.89 -
   43.90 -text {* The existence of the zero element of a vector space
   43.91 -  follows from the non-emptiness of carrier set. *}
   43.92 -
   43.93 -lemma (in vectorspace) zero [iff]: "0 \<in> V"
   43.94 -proof -
   43.95 -  from non_empty obtain x where x: "x \<in> V" by blast
   43.96 -  then have "0 = x - x" by (rule diff_self [symmetric])
   43.97 -  also from x x have "\<dots> \<in> V" by (rule diff_closed)
   43.98 -  finally show ?thesis .
   43.99 -qed
  43.100 -
  43.101 -lemma (in vectorspace) add_zero_right [simp]:
  43.102 -  "x \<in> V \<Longrightarrow>  x + 0 = x"
  43.103 -proof -
  43.104 -  assume x: "x \<in> V"
  43.105 -  from this and zero have "x + 0 = 0 + x" by (rule add_commute)
  43.106 -  also from x have "\<dots> = x" by (rule add_zero_left)
  43.107 -  finally show ?thesis .
  43.108 -qed
  43.109 -
  43.110 -lemma (in vectorspace) mult_assoc2:
  43.111 -    "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = (a * b) \<cdot> x"
  43.112 -  by (simp only: mult_assoc)
  43.113 -
  43.114 -lemma (in vectorspace) diff_mult_distrib1:
  43.115 -    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x - y) = a \<cdot> x - a \<cdot> y"
  43.116 -  by (simp add: diff_eq1 negate_eq1 add_mult_distrib1 mult_assoc2)
  43.117 -
  43.118 -lemma (in vectorspace) diff_mult_distrib2:
  43.119 -  "x \<in> V \<Longrightarrow> (a - b) \<cdot> x = a \<cdot> x - (b \<cdot> x)"
  43.120 -proof -
  43.121 -  assume x: "x \<in> V"
  43.122 -  have " (a - b) \<cdot> x = (a + - b) \<cdot> x"
  43.123 -    by (simp add: real_diff_def)
  43.124 -  also from x have "\<dots> = a \<cdot> x + (- b) \<cdot> x"
  43.125 -    by (rule add_mult_distrib2)
  43.126 -  also from x have "\<dots> = a \<cdot> x + - (b \<cdot> x)"
  43.127 -    by (simp add: negate_eq1 mult_assoc2)
  43.128 -  also from x have "\<dots> = a \<cdot> x - (b \<cdot> x)"
  43.129 -    by (simp add: diff_eq1)
  43.130 -  finally show ?thesis .
  43.131 -qed
  43.132 -
  43.133 -lemmas (in vectorspace) distrib =
  43.134 -  add_mult_distrib1 add_mult_distrib2
  43.135 -  diff_mult_distrib1 diff_mult_distrib2
  43.136 -
  43.137 -
  43.138 -text {* \medskip Further derived laws: *}
  43.139 -
  43.140 -lemma (in vectorspace) mult_zero_left [simp]:
  43.141 -  "x \<in> V \<Longrightarrow> 0 \<cdot> x = 0"
  43.142 -proof -
  43.143 -  assume x: "x \<in> V"
  43.144 -  have "0 \<cdot> x = (1 - 1) \<cdot> x" by simp
  43.145 -  also have "\<dots> = (1 + - 1) \<cdot> x" by simp
  43.146 -  also from x have "\<dots> =  1 \<cdot> x + (- 1) \<cdot> x"
  43.147 -    by (rule add_mult_distrib2)
  43.148 -  also from x have "\<dots> = x + (- 1) \<cdot> x" by simp
  43.149 -  also from x have "\<dots> = x + - x" by (simp add: negate_eq2a)
  43.150 -  also from x have "\<dots> = x - x" by (simp add: diff_eq2)
  43.151 -  also from x have "\<dots> = 0" by simp
  43.152 -  finally show ?thesis .
  43.153 -qed
  43.154 -
  43.155 -lemma (in vectorspace) mult_zero_right [simp]:
  43.156 -  "a \<cdot> 0 = (0::'a)"
  43.157 -proof -
  43.158 -  have "a \<cdot> 0 = a \<cdot> (0 - (0::'a))" by simp
  43.159 -  also have "\<dots> =  a \<cdot> 0 - a \<cdot> 0"
  43.160 -    by (rule diff_mult_distrib1) simp_all
  43.161 -  also have "\<dots> = 0" by simp
  43.162 -  finally show ?thesis .
  43.163 -qed
  43.164 -
  43.165 -lemma (in vectorspace) minus_mult_cancel [simp]:
  43.166 -    "x \<in> V \<Longrightarrow> (- a) \<cdot> - x = a \<cdot> x"
  43.167 -  by (simp add: negate_eq1 mult_assoc2)
  43.168 -
  43.169 -lemma (in vectorspace) add_minus_left_eq_diff:
  43.170 -  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + y = y - x"
  43.171 -proof -
  43.172 -  assume xy: "x \<in> V"  "y \<in> V"
  43.173 -  then have "- x + y = y + - x" by (simp add: add_commute)
  43.174 -  also from xy have "\<dots> = y - x" by (simp add: diff_eq1)
  43.175 -  finally show ?thesis .
  43.176 -qed
  43.177 -
  43.178 -lemma (in vectorspace) add_minus [simp]:
  43.179 -    "x \<in> V \<Longrightarrow> x + - x = 0"
  43.180 -  by (simp add: diff_eq2)
  43.181 -
  43.182 -lemma (in vectorspace) add_minus_left [simp]:
  43.183 -    "x \<in> V \<Longrightarrow> - x + x = 0"
  43.184 -  by (simp add: diff_eq2 add_commute)
  43.185 -
  43.186 -lemma (in vectorspace) minus_minus [simp]:
  43.187 -    "x \<in> V \<Longrightarrow> - (- x) = x"
  43.188 -  by (simp add: negate_eq1 mult_assoc2)
  43.189 -
  43.190 -lemma (in vectorspace) minus_zero [simp]:
  43.191 -    "- (0::'a) = 0"
  43.192 -  by (simp add: negate_eq1)
  43.193 -
  43.194 -lemma (in vectorspace) minus_zero_iff [simp]:
  43.195 -  "x \<in> V \<Longrightarrow> (- x = 0) = (x = 0)"
  43.196 -proof
  43.197 -  assume x: "x \<in> V"
  43.198 -  {
  43.199 -    from x have "x = - (- x)" by (simp add: minus_minus)
  43.200 -    also assume "- x = 0"
  43.201 -    also have "- \<dots> = 0" by (rule minus_zero)
  43.202 -    finally show "x = 0" .
  43.203 -  next
  43.204 -    assume "x = 0"
  43.205 -    then show "- x = 0" by simp
  43.206 -  }
  43.207 -qed
  43.208 -
  43.209 -lemma (in vectorspace) add_minus_cancel [simp]:
  43.210 -    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + (- x + y) = y"
  43.211 -  by (simp add: add_assoc [symmetric] del: add_commute)
  43.212 -
  43.213 -lemma (in vectorspace) minus_add_cancel [simp]:
  43.214 -    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + (x + y) = y"
  43.215 -  by (simp add: add_assoc [symmetric] del: add_commute)
  43.216 -
  43.217 -lemma (in vectorspace) minus_add_distrib [simp]:
  43.218 -    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - (x + y) = - x + - y"
  43.219 -  by (simp add: negate_eq1 add_mult_distrib1)
  43.220 -
  43.221 -lemma (in vectorspace) diff_zero [simp]:
  43.222 -    "x \<in> V \<Longrightarrow> x - 0 = x"
  43.223 -  by (simp add: diff_eq1)
  43.224 -
  43.225 -lemma (in vectorspace) diff_zero_right [simp]:
  43.226 -    "x \<in> V \<Longrightarrow> 0 - x = - x"
  43.227 -  by (simp add: diff_eq1)
  43.228 -
  43.229 -lemma (in vectorspace) add_left_cancel:
  43.230 -  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y = x + z) = (y = z)"
  43.231 -proof
  43.232 -  assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V"
  43.233 -  {
  43.234 -    from y have "y = 0 + y" by simp
  43.235 -    also from x y have "\<dots> = (- x + x) + y" by simp
  43.236 -    also from x y have "\<dots> = - x + (x + y)"
  43.237 -      by (simp add: add_assoc neg_closed)
  43.238 -    also assume "x + y = x + z"
  43.239 -    also from x z have "- x + (x + z) = - x + x + z"
  43.240 -      by (simp add: add_assoc [symmetric] neg_closed)
  43.241 -    also from x z have "\<dots> = z" by simp
  43.242 -    finally show "y = z" .
  43.243 -  next
  43.244 -    assume "y = z"
  43.245 -    then show "x + y = x + z" by (simp only:)
  43.246 -  }
  43.247 -qed
  43.248 -
  43.249 -lemma (in vectorspace) add_right_cancel:
  43.250 -    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (y + x = z + x) = (y = z)"
  43.251 -  by (simp only: add_commute add_left_cancel)
  43.252 -
  43.253 -lemma (in vectorspace) add_assoc_cong:
  43.254 -  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x' \<in> V \<Longrightarrow> y' \<in> V \<Longrightarrow> z \<in> V
  43.255 -    \<Longrightarrow> x + y = x' + y' \<Longrightarrow> x + (y + z) = x' + (y' + z)"
  43.256 -  by (simp only: add_assoc [symmetric])
  43.257 -
  43.258 -lemma (in vectorspace) mult_left_commute:
  43.259 -    "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = b \<cdot> a \<cdot> x"
  43.260 -  by (simp add: real_mult_commute mult_assoc2)
  43.261 -
  43.262 -lemma (in vectorspace) mult_zero_uniq:
  43.263 -  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> a \<cdot> x = 0 \<Longrightarrow> a = 0"
  43.264 -proof (rule classical)
  43.265 -  assume a: "a \<noteq> 0"
  43.266 -  assume x: "x \<in> V"  "x \<noteq> 0" and ax: "a \<cdot> x = 0"
  43.267 -  from x a have "x = (inverse a * a) \<cdot> x" by simp
  43.268 -  also from `x \<in> V` have "\<dots> = inverse a \<cdot> (a \<cdot> x)" by (rule mult_assoc)
  43.269 -  also from ax have "\<dots> = inverse a \<cdot> 0" by simp
  43.270 -  also have "\<dots> = 0" by simp
  43.271 -  finally have "x = 0" .
  43.272 -  with `x \<noteq> 0` show "a = 0" by contradiction
  43.273 -qed
  43.274 -
  43.275 -lemma (in vectorspace) mult_left_cancel:
  43.276 -  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (a \<cdot> x = a \<cdot> y) = (x = y)"
  43.277 -proof
  43.278 -  assume x: "x \<in> V" and y: "y \<in> V" and a: "a \<noteq> 0"
  43.279 -  from x have "x = 1 \<cdot> x" by simp
  43.280 -  also from a have "\<dots> = (inverse a * a) \<cdot> x" by simp
  43.281 -  also from x have "\<dots> = inverse a \<cdot> (a \<cdot> x)"
  43.282 -    by (simp only: mult_assoc)
  43.283 -  also assume "a \<cdot> x = a \<cdot> y"
  43.284 -  also from a y have "inverse a \<cdot> \<dots> = y"
  43.285 -    by (simp add: mult_assoc2)
  43.286 -  finally show "x = y" .
  43.287 -next
  43.288 -  assume "x = y"
  43.289 -  then show "a \<cdot> x = a \<cdot> y" by (simp only:)
  43.290 -qed
  43.291 -
  43.292 -lemma (in vectorspace) mult_right_cancel:
  43.293 -  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> (a \<cdot> x = b \<cdot> x) = (a = b)"
  43.294 -proof
  43.295 -  assume x: "x \<in> V" and neq: "x \<noteq> 0"
  43.296 -  {
  43.297 -    from x have "(a - b) \<cdot> x = a \<cdot> x - b \<cdot> x"
  43.298 -      by (simp add: diff_mult_distrib2)
  43.299 -    also assume "a \<cdot> x = b \<cdot> x"
  43.300 -    with x have "a \<cdot> x - b \<cdot> x = 0" by simp
  43.301 -    finally have "(a - b) \<cdot> x = 0" .
  43.302 -    with x neq have "a - b = 0" by (rule mult_zero_uniq)
  43.303 -    then show "a = b" by simp
  43.304 -  next
  43.305 -    assume "a = b"
  43.306 -    then show "a \<cdot> x = b \<cdot> x" by (simp only:)
  43.307 -  }
  43.308 -qed
  43.309 -
  43.310 -lemma (in vectorspace) eq_diff_eq:
  43.311 -  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x = z - y) = (x + y = z)"
  43.312 -proof
  43.313 -  assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V"
  43.314 -  {
  43.315 -    assume "x = z - y"
  43.316 -    then have "x + y = z - y + y" by simp
  43.317 -    also from y z have "\<dots> = z + - y + y"
  43.318 -      by (simp add: diff_eq1)
  43.319 -    also have "\<dots> = z + (- y + y)"
  43.320 -      by (rule add_assoc) (simp_all add: y z)
  43.321 -    also from y z have "\<dots> = z + 0"
  43.322 -      by (simp only: add_minus_left)
  43.323 -    also from z have "\<dots> = z"
  43.324 -      by (simp only: add_zero_right)
  43.325 -    finally show "x + y = z" .
  43.326 -  next
  43.327 -    assume "x + y = z"
  43.328 -    then have "z - y = (x + y) - y" by simp
  43.329 -    also from x y have "\<dots> = x + y + - y"
  43.330 -      by (simp add: diff_eq1)
  43.331 -    also have "\<dots> = x + (y + - y)"
  43.332 -      by (rule add_assoc) (simp_all add: x y)
  43.333 -    also from x y have "\<dots> = x" by simp
  43.334 -    finally show "x = z - y" ..
  43.335 -  }
  43.336 -qed
  43.337 -
  43.338 -lemma (in vectorspace) add_minus_eq_minus:
  43.339 -  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = 0 \<Longrightarrow> x = - y"
  43.340 -proof -
  43.341 -  assume x: "x \<in> V" and y: "y \<in> V"
  43.342 -  from x y have "x = (- y + y) + x" by simp
  43.343 -  also from x y have "\<dots> = - y + (x + y)" by (simp add: add_ac)
  43.344 -  also assume "x + y = 0"
  43.345 -  also from y have "- y + 0 = - y" by simp
  43.346 -  finally show "x = - y" .
  43.347 -qed
  43.348 -
  43.349 -lemma (in vectorspace) add_minus_eq:
  43.350 -  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = 0 \<Longrightarrow> x = y"
  43.351 -proof -
  43.352 -  assume x: "x \<in> V" and y: "y \<in> V"
  43.353 -  assume "x - y = 0"
  43.354 -  with x y have eq: "x + - y = 0" by (simp add: diff_eq1)
  43.355 -  with _ _ have "x = - (- y)"
  43.356 -    by (rule add_minus_eq_minus) (simp_all add: x y)
  43.357 -  with x y show "x = y" by simp
  43.358 -qed
  43.359 -
  43.360 -lemma (in vectorspace) add_diff_swap:
  43.361 -  "a \<in> V \<Longrightarrow> b \<in> V \<Longrightarrow> c \<in> V \<Longrightarrow> d \<in> V \<Longrightarrow> a + b = c + d
  43.362 -    \<Longrightarrow> a - c = d - b"
  43.363 -proof -
  43.364 -  assume vs: "a \<in> V"  "b \<in> V"  "c \<in> V"  "d \<in> V"
  43.365 -    and eq: "a + b = c + d"
  43.366 -  then have "- c + (a + b) = - c + (c + d)"
  43.367 -    by (simp add: add_left_cancel)
  43.368 -  also have "\<dots> = d" using `c \<in> V` `d \<in> V` by (rule minus_add_cancel)
  43.369 -  finally have eq: "- c + (a + b) = d" .
  43.370 -  from vs have "a - c = (- c + (a + b)) + - b"
  43.371 -    by (simp add: add_ac diff_eq1)
  43.372 -  also from vs eq have "\<dots>  = d + - b"
  43.373 -    by (simp add: add_right_cancel)
  43.374 -  also from vs have "\<dots> = d - b" by (simp add: diff_eq2)
  43.375 -  finally show "a - c = d - b" .
  43.376 -qed
  43.377 -
  43.378 -lemma (in vectorspace) vs_add_cancel_21:
  43.379 -  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> u \<in> V
  43.380 -    \<Longrightarrow> (x + (y + z) = y + u) = (x + z = u)"
  43.381 -proof
  43.382 -  assume vs: "x \<in> V"  "y \<in> V"  "z \<in> V"  "u \<in> V"
  43.383 -  {
  43.384 -    from vs have "x + z = - y + y + (x + z)" by simp
  43.385 -    also have "\<dots> = - y + (y + (x + z))"
  43.386 -      by (rule add_assoc) (simp_all add: vs)
  43.387 -    also from vs have "y + (x + z) = x + (y + z)"
  43.388 -      by (simp add: add_ac)
  43.389 -    also assume "x + (y + z) = y + u"
  43.390 -    also from vs have "- y + (y + u) = u" by simp
  43.391 -    finally show "x + z = u" .
  43.392 -  next
  43.393 -    assume "x + z = u"
  43.394 -    with vs show "x + (y + z) = y + u"
  43.395 -      by (simp only: add_left_commute [of x])
  43.396 -  }
  43.397 -qed
  43.398 -
  43.399 -lemma (in vectorspace) add_cancel_end:
  43.400 -  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + (y + z) = y) = (x = - z)"
  43.401 -proof
  43.402 -  assume vs: "x \<in> V"  "y \<in> V"  "z \<in> V"
  43.403 -  {
  43.404 -    assume "x + (y + z) = y"
  43.405 -    with vs have "(x + z) + y = 0 + y"
  43.406 -      by (simp add: add_ac)
  43.407 -    with vs have "x + z = 0"
  43.408 -      by (simp only: add_right_cancel add_closed zero)
  43.409 -    with vs show "x = - z" by (simp add: add_minus_eq_minus)
  43.410 -  next
  43.411 -    assume eq: "x = - z"
  43.412 -    then have "x + (y + z) = - z + (y + z)" by simp
  43.413 -    also have "\<dots> = y + (- z + z)"
  43.414 -      by (rule add_left_commute) (simp_all add: vs)
  43.415 -    also from vs have "\<dots> = y"  by simp
  43.416 -    finally show "x + (y + z) = y" .
  43.417 -  }
  43.418 -qed
  43.419 -
  43.420 -end
    44.1 --- a/src/HOL/Real/HahnBanach/ZornLemma.thy	Mon Dec 29 13:23:53 2008 +0100
    44.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    44.3 @@ -1,58 +0,0 @@
    44.4 -(*  Title:      HOL/Real/HahnBanach/ZornLemma.thy
    44.5 -    ID:         $Id$
    44.6 -    Author:     Gertrud Bauer, TU Munich
    44.7 -*)
    44.8 -
    44.9 -header {* Zorn's Lemma *}
   44.10 -
   44.11 -theory ZornLemma
   44.12 -imports Zorn
   44.13 -begin
   44.14 -
   44.15 -text {*
   44.16 -  Zorn's Lemmas states: if every linear ordered subset of an ordered
   44.17 -  set @{text S} has an upper bound in @{text S}, then there exists a
   44.18 -  maximal element in @{text S}.  In our application, @{text S} is a
   44.19 -  set of sets ordered by set inclusion. Since the union of a chain of
   44.20 -  sets is an upper bound for all elements of the chain, the conditions
   44.21 -  of Zorn's lemma can be modified: if @{text S} is non-empty, it
   44.22 -  suffices to show that for every non-empty chain @{text c} in @{text
   44.23 -  S} the union of @{text c} also lies in @{text S}.
   44.24 -*}
   44.25 -
   44.26 -theorem Zorn's_Lemma:
   44.27 -  assumes r: "\<And>c. c \<in> chain S \<Longrightarrow> \<exists>x. x \<in> c \<Longrightarrow> \<Union>c \<in> S"
   44.28 -    and aS: "a \<in> S"
   44.29 -  shows "\<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z \<longrightarrow> y = z"
   44.30 -proof (rule Zorn_Lemma2)
   44.31 -  show "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
   44.32 -  proof
   44.33 -    fix c assume "c \<in> chain S"
   44.34 -    show "\<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
   44.35 -    proof cases
   44.36 -
   44.37 -      txt {* If @{text c} is an empty chain, then every element in
   44.38 -	@{text S} is an upper bound of @{text c}. *}
   44.39 -
   44.40 -      assume "c = {}"
   44.41 -      with aS show ?thesis by fast
   44.42 -
   44.43 -      txt {* If @{text c} is non-empty, then @{text "\<Union>c"} is an upper
   44.44 -	bound of @{text c}, lying in @{text S}. *}
   44.45 -
   44.46 -    next
   44.47 -      assume "c \<noteq> {}"
   44.48 -      show ?thesis
   44.49 -      proof
   44.50 -        show "\<forall>z \<in> c. z \<subseteq> \<Union>c" by fast
   44.51 -        show "\<Union>c \<in> S"
   44.52 -        proof (rule r)
   44.53 -          from `c \<noteq> {}` show "\<exists>x. x \<in> c" by fast
   44.54 -	  show "c \<in> chain S" by fact
   44.55 -        qed
   44.56 -      qed
   44.57 -    qed
   44.58 -  qed
   44.59 -qed
   44.60 -
   44.61 -end
    45.1 --- a/src/HOL/Real/HahnBanach/document/root.bib	Mon Dec 29 13:23:53 2008 +0100
    45.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    45.3 @@ -1,27 +0,0 @@
    45.4 -
    45.5 -@Book{Heuser:1986,
    45.6 -  author = 	 {H. Heuser},
    45.7 -  title = 	 {Funktionalanalysis: Theorie und Anwendung},
    45.8 -  publisher = 	 {Teubner},
    45.9 -  year = 	 1986
   45.10 -}
   45.11 -
   45.12 -@InCollection{Narici:1996,
   45.13 -  author = 	 {L. Narici and E. Beckenstein},
   45.14 -  title = 	 {The {Hahn-Banach Theorem}: The Life and Times},
   45.15 -  booktitle = 	 {Topology Atlas},
   45.16 -  publisher =	 {York University, Toronto, Ontario, Canada},
   45.17 -  year =	 1996,
   45.18 -  note =	 {\url{http://at.yorku.ca/topology/preprint.htm} and
   45.19 -                  \url{http://at.yorku.ca/p/a/a/a/16.htm}}
   45.20 -}
   45.21 -
   45.22 -@Article{Nowak:1993,
   45.23 -  author =       {B. Nowak and A. Trybulec},
   45.24 -  title =	 {{Hahn-Banach} Theorem},
   45.25 -  journal =      {Journal of Formalized Mathematics},
   45.26 -  year =         {1993},
   45.27 -  volume =       {5},
   45.28 -  institution =  {University of Bialystok},
   45.29 -  note =         {\url{http://mizar.uwb.edu.pl/JFM/Vol5/hahnban.html}}
   45.30 -}
    46.1 --- a/src/HOL/Real/HahnBanach/document/root.tex	Mon Dec 29 13:23:53 2008 +0100
    46.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    46.3 @@ -1,83 +0,0 @@
    46.4 -\documentclass[10pt,a4paper,twoside]{article}
    46.5 -\usepackage{graphicx}
    46.6 -\usepackage{latexsym,theorem}
    46.7 -\usepackage{isabelle,isabellesym}
    46.8 -\usepackage{pdfsetup} %last one!
    46.9 -
   46.10 -\isabellestyle{it}
   46.11 -\urlstyle{rm}
   46.12 -
   46.13 -\newcommand{\isasymsup}{\isamath{\sup\,}}
   46.14 -\newcommand{\skp}{\smallskip}
   46.15 -
   46.16 -
   46.17 -\begin{document}
   46.18 -
   46.19 -\pagestyle{headings}
   46.20 -\pagenumbering{arabic}
   46.21 -
   46.22 -\title{The Hahn-Banach Theorem \\ for Real Vector Spaces}
   46.23 -\author{Gertrud Bauer \\ \url{http://www.in.tum.de/~bauerg/}}
   46.24 -\maketitle
   46.25 -
   46.26 -\begin{abstract}
   46.27 -  The Hahn-Banach Theorem is one of the most fundamental results in functional
   46.28 -  analysis. We present a fully formal proof of two versions of the theorem,
   46.29 -  one for general linear spaces and another for normed spaces.  This
   46.30 -  development is based on simply-typed classical set-theory, as provided by
   46.31 -  Isabelle/HOL.
   46.32 -\end{abstract}
   46.33 -
   46.34 -
   46.35 -\tableofcontents
   46.36 -\parindent 0pt \parskip 0.5ex
   46.37 -
   46.38 -\clearpage
   46.39 -\section{Preface}
   46.40 -
   46.41 -This is a fully formal proof of the Hahn-Banach Theorem. It closely follows
   46.42 -the informal presentation given in Heuser's textbook \cite[{\S} 36]{Heuser:1986}.
   46.43 -Another formal proof of the same theorem has been done in Mizar
   46.44 -\cite{Nowak:1993}.  A general overview of the relevance and history of the
   46.45 -Hahn-Banach Theorem is given by Narici and Beckenstein \cite{Narici:1996}.
   46.46 -
   46.47 -\medskip The document is structured as follows.  The first part contains
   46.48 -definitions of basic notions of linear algebra: vector spaces, subspaces,
   46.49 -normed spaces, continuous linear-forms, norm of functions and an order on
   46.50 -functions by domain extension.  The second part contains some lemmas about the
   46.51 -supremum (w.r.t.\ the function order) and extension of non-maximal functions.
   46.52 -With these preliminaries, the main proof of the theorem (in its two versions)
   46.53 -is conducted in the third part.  The dependencies of individual theories are
   46.54 -as follows.
   46.55 -
   46.56 -\begin{center}
   46.57 -  \includegraphics[scale=0.5]{session_graph}  
   46.58 -\end{center}
   46.59 -
   46.60 -\clearpage
   46.61 -\part {Basic Notions}
   46.62 -
   46.63 -\input{Bounds}
   46.64 -\input{VectorSpace}
   46.65 -\input{Subspace}
   46.66 -\input{NormedSpace}
   46.67 -\input{Linearform}
   46.68 -\input{FunctionOrder}
   46.69 -\input{FunctionNorm}
   46.70 -\input{ZornLemma}
   46.71 -
   46.72 -\clearpage
   46.73 -\part {Lemmas for the Proof}
   46.74 -
   46.75 -\input{HahnBanachSupLemmas}
   46.76 -\input{HahnBanachExtLemmas}
   46.77 -\input{HahnBanachLemmas}
   46.78 -
   46.79 -\clearpage
   46.80 -\part {The Main Proof}
   46.81 -
   46.82 -\input{HahnBanach}
   46.83 -\bibliographystyle{abbrv}
   46.84 -\bibliography{root}
   46.85 -
   46.86 -\end{document}
    47.1 --- a/src/HOL/Real/RealVector.thy	Mon Dec 29 13:23:53 2008 +0100
    47.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    47.3 @@ -1,841 +0,0 @@
    47.4 -(*  Title:      HOL/Real/RealVector.thy
    47.5 -    Author:     Brian Huffman
    47.6 -*)
    47.7 -
    47.8 -header {* Vector Spaces and Algebras over the Reals *}
    47.9 -
   47.10 -theory RealVector
   47.11 -imports "~~/src/HOL/RealPow"
   47.12 -begin
   47.13 -
   47.14 -subsection {* Locale for additive functions *}
   47.15 -
   47.16 -locale additive =
   47.17 -  fixes f :: "'a::ab_group_add \<Rightarrow> 'b::ab_group_add"
   47.18 -  assumes add: "f (x + y) = f x + f y"
   47.19 -begin
   47.20 -
   47.21 -lemma zero: "f 0 = 0"
   47.22 -proof -
   47.23 -  have "f 0 = f (0 + 0)" by simp
   47.24 -  also have "\<dots> = f 0 + f 0" by (rule add)
   47.25 -  finally show "f 0 = 0" by simp
   47.26 -qed
   47.27 -
   47.28 -lemma minus: "f (- x) = - f x"
   47.29 -proof -
   47.30 -  have "f (- x) + f x = f (- x + x)" by (rule add [symmetric])
   47.31 -  also have "\<dots> = - f x + f x" by (simp add: zero)
   47.32 -  finally show "f (- x) = - f x" by (rule add_right_imp_eq)
   47.33 -qed
   47.34 -
   47.35 -lemma diff: "f (x - y) = f x - f y"
   47.36 -by (simp add: diff_def add minus)
   47.37 -
   47.38 -lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
   47.39 -apply (cases "finite A")
   47.40 -apply (induct set: finite)
   47.41 -apply (simp add: zero)
   47.42 -apply (simp add: add)
   47.43 -apply (simp add: zero)
   47.44 -done
   47.45 -
   47.46 -end
   47.47 -
   47.48 -subsection {* Vector spaces *}
   47.49 -
   47.50 -locale vector_space =
   47.51 -  fixes scale :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b"
   47.52 -  assumes scale_right_distrib: "scale a (x + y) = scale a x + scale a y"
   47.53 -  and scale_left_distrib: "scale (a + b) x = scale a x + scale b x"
   47.54 -  and scale_scale [simp]: "scale a (scale b x) = scale (a * b) x"
   47.55 -  and scale_one [simp]: "scale 1 x = x"
   47.56 -begin
   47.57 -
   47.58 -lemma scale_left_commute:
   47.59 -  "scale a (scale b x) = scale b (scale a x)"
   47.60 -by (simp add: mult_commute)
   47.61 -
   47.62 -lemma scale_zero_left [simp]: "scale 0 x = 0"
   47.63 -  and scale_minus_left [simp]: "scale (- a) x = - (scale a x)"
   47.64 -  and scale_left_diff_distrib: "scale (a - b) x = scale a x - scale b x"
   47.65 -proof -
   47.66 -  interpret s: additive ["\<lambda>a. scale a x"]
   47.67 -    proof qed (rule scale_left_distrib)
   47.68 -  show "scale 0 x = 0" by (rule s.zero)
   47.69 -  show "scale (- a) x = - (scale a x)" by (rule s.minus)
   47.70 -  show "scale (a - b) x = scale a x - scale b x" by (rule s.diff)
   47.71 -qed
   47.72 -
   47.73 -lemma scale_zero_right [simp]: "scale a 0 = 0"
   47.74 -  and scale_minus_right [simp]: "scale a (- x) = - (scale a x)"
   47.75 -  and scale_right_diff_distrib: "scale a (x - y) = scale a x - scale a y"
   47.76 -proof -
   47.77 -  interpret s: additive ["\<lambda>x. scale a x"]
   47.78 -    proof qed (rule scale_right_distrib)
   47.79 -  show "scale a 0 = 0" by (rule s.zero)
   47.80 -  show "scale a (- x) = - (scale a x)" by (rule s.minus)
   47.81 -  show "scale a (x - y) = scale a x - scale a y" by (rule s.diff)
   47.82 -qed
   47.83 -
   47.84 -lemma scale_eq_0_iff [simp]:
   47.85 -  "scale a x = 0 \<longleftrightarrow> a = 0 \<or> x = 0"
   47.86 -proof cases
   47.87 -  assume "a = 0" thus ?thesis by simp
   47.88 -next
   47.89 -  assume anz [simp]: "a \<noteq> 0"
   47.90 -  { assume "scale a x = 0"
   47.91 -    hence "scale (inverse a) (scale a x) = 0" by simp
   47.92 -    hence "x = 0" by simp }
   47.93 -  thus ?thesis by force
   47.94 -qed
   47.95 -
   47.96 -lemma scale_left_imp_eq:
   47.97 -  "\<lbrakk>a \<noteq> 0; scale a x = scale a y\<rbrakk> \<Longrightarrow> x = y"
   47.98 -proof -
   47.99 -  assume nonzero: "a \<noteq> 0"
  47.100 -  assume "scale a x = scale a y"
  47.101 -  hence "scale a (x - y) = 0"
  47.102 -     by (simp add: scale_right_diff_distrib)
  47.103 -  hence "x - y = 0" by (simp add: nonzero)
  47.104 -  thus "x = y" by (simp only: right_minus_eq)
  47.105 -qed
  47.106 -
  47.107 -lemma scale_right_imp_eq:
  47.108 -  "\<lbrakk>x \<noteq> 0; scale a x = scale b x\<rbrakk> \<Longrightarrow> a = b"
  47.109 -proof -
  47.110 -  assume nonzero: "x \<noteq> 0"
  47.111 -  assume "scale a x = scale b x"
  47.112 -  hence "scale (a - b) x = 0"
  47.113 -     by (simp add: scale_left_diff_distrib)
  47.114 -  hence "a - b = 0" by (simp add: nonzero)
  47.115 -  thus "a = b" by (simp only: right_minus_eq)
  47.116 -qed
  47.117 -
  47.118 -lemma scale_cancel_left:
  47.119 -  "scale a x = scale a y \<longleftrightarrow> x = y \<or> a = 0"
  47.120 -by (auto intro: scale_left_imp_eq)
  47.121 -
  47.122 -lemma scale_cancel_right:
  47.123 -  "scale a x = scale b x \<longleftrightarrow> a = b \<or> x = 0"
  47.124 -by (auto intro: scale_right_imp_eq)
  47.125 -
  47.126 -end
  47.127 -
  47.128 -subsection {* Real vector spaces *}
  47.129 -
  47.130 -class scaleR = type +
  47.131 -  fixes scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "*\<^sub>R" 75)
  47.132 -begin
  47.133 -
  47.134 -abbreviation
  47.135 -  divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a" (infixl "'/\<^sub>R" 70)
  47.136 -where
  47.137 -  "x /\<^sub>R r == scaleR (inverse r) x"
  47.138 -
  47.139 -end
  47.140 -
  47.141 -instantiation real :: scaleR
  47.142 -begin
  47.143 -
  47.144 -definition
  47.145 -  real_scaleR_def [simp]: "scaleR a x = a * x"
  47.146 -
  47.147 -instance ..
  47.148 -
  47.149 -end
  47.150 -
  47.151 -class real_vector = scaleR + ab_group_add +
  47.152 -  assumes scaleR_right_distrib: "scaleR a (x + y) = scaleR a x + scaleR a y"
  47.153 -  and scaleR_left_distrib: "scaleR (a + b) x = scaleR a x + scaleR b x"
  47.154 -  and scaleR_scaleR [simp]: "scaleR a (scaleR b x) = scaleR (a * b) x"
  47.155 -  and scaleR_one [simp]: "scaleR 1 x = x"
  47.156 -
  47.157 -interpretation real_vector:
  47.158 -  vector_space ["scaleR :: real \<Rightarrow> 'a \<Rightarrow> 'a::real_vector"]
  47.159 -apply unfold_locales
  47.160 -apply (rule scaleR_right_distrib)
  47.161 -apply (rule scaleR_left_distrib)
  47.162 -apply (rule scaleR_scaleR)
  47.163 -apply (rule scaleR_one)
  47.164 -done
  47.165 -
  47.166 -text {* Recover original theorem names *}
  47.167 -
  47.168 -lemmas scaleR_left_commute = real_vector.scale_left_commute
  47.169 -lemmas scaleR_zero_left = real_vector.scale_zero_left
  47.170 -lemmas scaleR_minus_left = real_vector.scale_minus_left
  47.171 -lemmas scaleR_left_diff_distrib = real_vector.scale_left_diff_distrib
  47.172 -lemmas scaleR_zero_right = real_vector.scale_zero_right
  47.173 -lemmas scaleR_minus_right = real_vector.scale_minus_right
  47.174 -lemmas scaleR_right_diff_distrib = real_vector.scale_right_diff_distrib
  47.175 -lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff
  47.176 -lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq
  47.177 -lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq
  47.178 -lemmas scaleR_cancel_left = real_vector.scale_cancel_left
  47.179 -lemmas scaleR_cancel_right = real_vector.scale_cancel_right
  47.180 -
  47.181 -class real_algebra = real_vector + ring +
  47.182 -  assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)"
  47.183 -  and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"
  47.184 -
  47.185 -class real_algebra_1 = real_algebra + ring_1
  47.186 -
  47.187 -class real_div_algebra = real_algebra_1 + division_ring
  47.188 -
  47.189 -class real_field = real_div_algebra + field
  47.190 -
  47.191 -instance real :: real_field
  47.192 -apply (intro_classes, unfold real_scaleR_def)
  47.193 -apply (rule right_distrib)
  47.194 -apply (rule left_distrib)
  47.195 -apply (rule mult_assoc [symmetric])
  47.196 -apply (rule mult_1_left)
  47.197 -apply (rule mult_assoc)
  47.198 -apply (rule mult_left_commute)
  47.199 -done
  47.200 -
  47.201 -interpretation scaleR_left: additive ["(\<lambda>a. scaleR a x::'a::real_vector)"]
  47.202 -proof qed (rule scaleR_left_distrib)
  47.203 -
  47.204 -interpretation scaleR_right: additive ["(\<lambda>x. scaleR a x::'a::real_vector)"]
  47.205 -proof qed (rule scaleR_right_distrib)
  47.206 -
  47.207 -lemma nonzero_inverse_scaleR_distrib:
  47.208 -  fixes x :: "'a::real_div_algebra" shows
  47.209 -  "\<lbrakk>a \<noteq> 0; x \<noteq> 0\<rbrakk> \<Longrightarrow> inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
  47.210 -by (rule inverse_unique, simp)
  47.211 -
  47.212 -lemma inverse_scaleR_distrib:
  47.213 -  fixes x :: "'a::{real_div_algebra,division_by_zero}"
  47.214 -  shows "inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
  47.215 -apply (case_tac "a = 0", simp)
  47.216 -apply (case_tac "x = 0", simp)
  47.217 -apply (erule (1) nonzero_inverse_scaleR_distrib)
  47.218 -done
  47.219 -
  47.220 -
  47.221 -subsection {* Embedding of the Reals into any @{text real_algebra_1}:
  47.222 -@{term of_real} *}
  47.223 -
  47.224 -definition
  47.225 -  of_real :: "real \<Rightarrow> 'a::real_algebra_1" where
  47.226 -  "of_real r = scaleR r 1"
  47.227 -
  47.228 -lemma scaleR_conv_of_real: "scaleR r x = of_real r * x"
  47.229 -by (simp add: of_real_def)
  47.230 -
  47.231 -lemma of_real_0 [simp]: "of_real 0 = 0"
  47.232 -by (simp add: of_real_def)
  47.233 -
  47.234 -lemma of_real_1 [simp]: "of_real 1 = 1"
  47.235 -by (simp add: of_real_def)
  47.236 -
  47.237 -lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y"
  47.238 -by (simp add: of_real_def scaleR_left_distrib)
  47.239 -
  47.240 -lemma of_real_minus [simp]: "of_real (- x) = - of_real x"
  47.241 -by (simp add: of_real_def)
  47.242 -
  47.243 -lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y"
  47.244 -by (simp add: of_real_def scaleR_left_diff_distrib)
  47.245 -
  47.246 -lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y"
  47.247 -by (simp add: of_real_def mult_commute)
  47.248 -
  47.249 -lemma nonzero_of_real_inverse:
  47.250 -  "x \<noteq> 0 \<Longrightarrow> of_real (inverse x) =
  47.251 -   inverse (of_real x :: 'a::real_div_algebra)"
  47.252 -by (simp add: of_real_def nonzero_inverse_scaleR_distrib)
  47.253 -
  47.254 -lemma of_real_inverse [simp]:
  47.255 -  "of_real (inverse x) =
  47.256 -   inverse (of_real x :: 'a::{real_div_algebra,division_by_zero})"
  47.257 -by (simp add: of_real_def inverse_scaleR_distrib)
  47.258 -
  47.259 -lemma nonzero_of_real_divide:
  47.260 -  "y \<noteq> 0 \<Longrightarrow> of_real (x / y) =
  47.261 -   (of_real x / of_real y :: 'a::real_field)"
  47.262 -by (simp add: divide_inverse nonzero_of_real_inverse)
  47.263 -
  47.264 -lemma of_real_divide [simp]:
  47.265 -  "of_real (x / y) =
  47.266 -   (of_real x / of_real y :: 'a::{real_field,division_by_zero})"
  47.267 -by (simp add: divide_inverse)
  47.268 -
  47.269 -lemma of_real_power [simp]:
  47.270 -  "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1,recpower}) ^ n"
  47.271 -by (induct n) (simp_all add: power_Suc)
  47.272 -
  47.273 -lemma of_real_eq_iff [simp]: "(of_real x = of_real y) = (x = y)"
  47.274 -by (simp add: of_real_def scaleR_cancel_right)
  47.275 -
  47.276 -lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified]
  47.277 -
  47.278 -lemma of_real_eq_id [simp]: "of_real = (id :: real \<Rightarrow> real)"
  47.279 -proof
  47.280 -  fix r
  47.281 -  show "of_real r = id r"
  47.282 -    by (simp add: of_real_def)
  47.283 -qed
  47.284 -
  47.285 -text{*Collapse nested embeddings*}
  47.286 -lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n"
  47.287 -by (induct n) auto
  47.288 -
  47.289 -lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z"
  47.290 -by (cases z rule: int_diff_cases, simp)
  47.291 -
  47.292 -lemma of_real_number_of_eq:
  47.293 -  "of_real (number_of w) = (number_of w :: 'a::{number_ring,real_algebra_1})"
  47.294 -by (simp add: number_of_eq)
  47.295 -
  47.296 -text{*Every real algebra has characteristic zero*}
  47.297 -instance real_algebra_1 < ring_char_0
  47.298 -proof
  47.299 -  fix m n :: nat
  47.300 -  have "(of_real (of_nat m) = (of_real (of_nat n)::'a)) = (m = n)"
  47.301 -    by (simp only: of_real_eq_iff of_nat_eq_iff)
  47.302 -  thus "(of_nat m = (of_nat n::'a)) = (m = n)"
  47.303 -    by (simp only: of_real_of_nat_eq)
  47.304 -qed
  47.305 -
  47.306 -instance real_field < field_char_0 ..
  47.307 -
  47.308 -
  47.309 -subsection {* The Set of Real Numbers *}
  47.310 -
  47.311 -definition
  47.312 -  Reals :: "'a::real_algebra_1 set" where
  47.313 -  [code del]: "Reals \<equiv> range of_real"
  47.314 -
  47.315 -notation (xsymbols)
  47.316 -  Reals  ("\<real>")
  47.317 -
  47.318 -lemma Reals_of_real [simp]: "of_real r \<in> Reals"
  47.319 -by (simp add: Reals_def)
  47.320 -
  47.321 -lemma Reals_of_int [simp]: "of_int z \<in> Reals"
  47.322 -by (subst of_real_of_int_eq [symmetric], rule Reals_of_real)
  47.323 -
  47.324 -lemma Reals_of_nat [simp]: "of_nat n \<in> Reals"
  47.325 -by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real)
  47.326 -
  47.327 -lemma Reals_number_of [simp]:
  47.328 -  "(number_of w::'a::{number_ring,real_algebra_1}) \<in> Reals"
  47.329 -by (subst of_real_number_of_eq [symmetric], rule Reals_of_real)
  47.330 -
  47.331 -lemma Reals_0 [simp]: "0 \<in> Reals"
  47.332 -apply (unfold Reals_def)
  47.333 -apply (rule range_eqI)
  47.334 -apply (rule of_real_0 [symmetric])
  47.335 -done
  47.336 -
  47.337 -lemma Reals_1 [simp]: "1 \<in> Reals"
  47.338 -apply (unfold Reals_def)
  47.339 -apply (rule range_eqI)
  47.340 -apply (rule of_real_1 [symmetric])
  47.341 -done
  47.342 -
  47.343 -lemma Reals_add [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a + b \<in> Reals"
  47.344 -apply (auto simp add: Reals_def)
  47.345 -apply (rule range_eqI)
  47.346 -apply (rule of_real_add [symmetric])
  47.347 -done
  47.348 -
  47.349 -lemma Reals_minus [simp]: "a \<in> Reals \<Longrightarrow> - a \<in> Reals"
  47.350 -apply (auto simp add: Reals_def)
  47.351 -apply (rule range_eqI)
  47.352 -apply (rule of_real_minus [symmetric])
  47.353 -done
  47.354 -
  47.355 -lemma Reals_diff [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a - b \<in> Reals"
  47.356 -apply (auto simp add: Reals_def)
  47.357 -apply (rule range_eqI)
  47.358 -apply (rule of_real_diff [symmetric])
  47.359 -done
  47.360 -
  47.361 -lemma Reals_mult [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a * b \<in> Reals"
  47.362 -apply (auto simp add: Reals_def)
  47.363 -apply (rule range_eqI)
  47.364 -apply (rule of_real_mult [symmetric])
  47.365 -done
  47.366 -
  47.367 -lemma nonzero_Reals_inverse:
  47.368 -  fixes a :: "'a::real_div_algebra"
  47.369 -  shows "\<lbrakk>a \<in> Reals; a \<noteq> 0\<rbrakk> \<Longrightarrow> inverse a \<in> Reals"
  47.370 -apply (auto simp add: Reals_def)
  47.371 -apply (rule range_eqI)
  47.372 -apply (erule nonzero_of_real_inverse [symmetric])
  47.373 -done
  47.374 -
  47.375 -lemma Reals_inverse [simp]:
  47.376 -  fixes a :: "'a::{real_div_algebra,division_by_zero}"
  47.377 -  shows "a \<in> Reals \<Longrightarrow> inverse a \<in> Reals"
  47.378 -apply (auto simp add: Reals_def)
  47.379 -apply (rule range_eqI)
  47.380 -apply (rule of_real_inverse [symmetric])
  47.381 -done
  47.382 -
  47.383 -lemma nonzero_Reals_divide:
  47.384 -  fixes a b :: "'a::real_field"
  47.385 -  shows "\<lbrakk>a \<in> Reals; b \<in> Reals; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
  47.386 -apply (auto simp add: Reals_def)
  47.387 -apply (rule range_eqI)
  47.388 -apply (erule nonzero_of_real_divide [symmetric])
  47.389 -done
  47.390 -
  47.391 -lemma Reals_divide [simp]:
  47.392 -  fixes a b :: "'a::{real_field,division_by_zero}"
  47.393 -  shows "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
  47.394 -apply (auto simp add: Reals_def)
  47.395 -apply (rule range_eqI)
  47.396 -apply (rule of_real_divide [symmetric])
  47.397 -done
  47.398 -
  47.399 -lemma Reals_power [simp]:
  47.400 -  fixes a :: "'a::{real_algebra_1,recpower}"
  47.401 -  shows "a \<in> Reals \<Longrightarrow> a ^ n \<in> Reals"
  47.402 -apply (auto simp add: Reals_def)
  47.403 -apply (rule range_eqI)
  47.404 -apply (rule of_real_power [symmetric])
  47.405 -done
  47.406 -
  47.407 -lemma Reals_cases [cases set: Reals]:
  47.408 -  assumes "q \<in> \<real>"
  47.409 -  obtains (of_real) r where "q = of_real r"
  47.410 -  unfolding Reals_def
  47.411 -proof -
  47.412 -  from `q \<in> \<real>` have "q \<in> range of_real" unfolding Reals_def .
  47.413 -  then obtain r where "q = of_real r" ..
  47.414 -  then show thesis ..
  47.415 -qed
  47.416 -
  47.417 -lemma Reals_induct [case_names of_real, induct set: Reals]:
  47.418 -  "q \<in> \<real> \<Longrightarrow> (\<And>r. P (of_real r)) \<Longrightarrow> P q"
  47.419 -  by (rule Reals_cases) auto
  47.420 -
  47.421 -
  47.422 -subsection {* Real normed vector spaces *}
  47.423 -
  47.424 -class norm = type +
  47.425 -  fixes norm :: "'a \<Rightarrow> real"
  47.426 -
  47.427 -instantiation real :: norm
  47.428 -begin
  47.429 -
  47.430 -definition
  47.431 -  real_norm_def [simp]: "norm r \<equiv> \<bar>r\<bar>"
  47.432 -
  47.433 -instance ..
  47.434 -
  47.435 -end
  47.436 -
  47.437 -class sgn_div_norm = scaleR + norm + sgn +
  47.438 -  assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x"
  47.439 -
  47.440 -class real_normed_vector = real_vector + sgn_div_norm +
  47.441 -  assumes norm_ge_zero [simp]: "0 \<le> norm x"
  47.442 -  and norm_eq_zero [simp]: "norm x = 0 \<longleftrightarrow> x = 0"
  47.443 -  and norm_triangle_ineq: "norm (x + y) \<le> norm x + norm y"
  47.444 -  and norm_scaleR: "norm (scaleR a x) = \<bar>a\<bar> * norm x"
  47.445 -
  47.446 -class real_normed_algebra = real_algebra + real_normed_vector +
  47.447 -  assumes norm_mult_ineq: "norm (x * y) \<le> norm x * norm y"
  47.448 -
  47.449 -class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra +
  47.450 -  assumes norm_one [simp]: "norm 1 = 1"
  47.451 -
  47.452 -class real_normed_div_algebra = real_div_algebra + real_normed_vector +
  47.453 -  assumes norm_mult: "norm (x * y) = norm x * norm y"
  47.454 -
  47.455 -class real_normed_field = real_field + real_normed_div_algebra
  47.456 -
  47.457 -instance real_normed_div_algebra < real_normed_algebra_1
  47.458 -proof
  47.459 -  fix x y :: 'a
  47.460 -  show "norm (x * y) \<le> norm x * norm y"
  47.461 -    by (simp add: norm_mult)
  47.462 -next
  47.463 -  have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)"
  47.464 -    by (rule norm_mult)
  47.465 -  thus "norm (1::'a) = 1" by simp
  47.466 -qed
  47.467 -
  47.468 -instance real :: real_normed_field
  47.469 -apply (intro_classes, unfold real_norm_def real_scaleR_def)
  47.470 -apply (simp add: real_sgn_def)
  47.471 -apply (rule abs_ge_zero)
  47.472 -apply (rule abs_eq_0)
  47.473 -apply (rule abs_triangle_ineq)
  47.474 -apply (rule abs_mult)
  47.475 -apply (rule abs_mult)
  47.476 -done
  47.477 -
  47.478 -lemma norm_zero [simp]: "norm (0::'a::real_normed_vector) = 0"
  47.479 -by simp
  47.480 -
  47.481 -lemma zero_less_norm_iff [simp]:
  47.482 -  fixes x :: "'a::real_normed_vector"
  47.483 -  shows "(0 < norm x) = (x \<noteq> 0)"
  47.484 -by (simp add: order_less_le)
  47.485 -
  47.486 -lemma norm_not_less_zero [simp]:
  47.487 -  fixes x :: "'a::real_normed_vector"
  47.488 -  shows "\<not> norm x < 0"
  47.489 -by (simp add: linorder_not_less)
  47.490 -
  47.491 -lemma norm_le_zero_iff [simp]:
  47.492 -  fixes x :: "'a::real_normed_vector"
  47.493 -  shows "(norm x \<le> 0) = (x = 0)"
  47.494 -by (simp add: order_le_less)
  47.495 -
  47.496 -lemma norm_minus_cancel [simp]:
  47.497 -  fixes x :: "'a::real_normed_vector"
  47.498 -  shows "norm (- x) = norm x"
  47.499 -proof -
  47.500 -  have "norm (- x) = norm (scaleR (- 1) x)"
  47.501 -    by (simp only: scaleR_minus_left scaleR_one)
  47.502 -  also have "\<dots> = \<bar>- 1\<bar> * norm x"
  47.503 -    by (rule norm_scaleR)
  47.504 -  finally show ?thesis by simp
  47.505 -qed
  47.506 -
  47.507 -lemma norm_minus_commute:
  47.508 -  fixes a b :: "'a::real_normed_vector"
  47.509 -  shows "norm (a - b) = norm (b - a)"
  47.510 -proof -
  47.511 -  have "norm (- (b - a)) = norm (b - a)"
  47.512 -    by (rule norm_minus_cancel)
  47.513 -  thus ?thesis by simp
  47.514 -qed
  47.515 -
  47.516 -lemma norm_triangle_ineq2:
  47.517 -  fixes a b :: "'a::real_normed_vector"
  47.518 -  shows "norm a - norm b \<le> norm (a - b)"
  47.519 -proof -
  47.520 -  have "norm (a - b + b) \<le> norm (a - b) + norm b"
  47.521 -    by (rule norm_triangle_ineq)
  47.522 -  thus ?thesis by simp
  47.523 -qed
  47.524 -
  47.525 -lemma norm_triangle_ineq3:
  47.526 -  fixes a b :: "'a::real_normed_vector"
  47.527 -  shows "\<bar>norm a - norm b\<bar> \<le> norm (a - b)"
  47.528 -apply (subst abs_le_iff)
  47.529 -apply auto
  47.530 -apply (rule norm_triangle_ineq2)
  47.531 -apply (subst norm_minus_commute)
  47.532 -apply (rule norm_triangle_ineq2)
  47.533 -done
  47.534 -
  47.535 -lemma norm_triangle_ineq4:
  47.536 -  fixes a b :: "'a::real_normed_vector"
  47.537 -  shows "norm (a - b) \<le> norm a + norm b"
  47.538 -proof -
  47.539 -  have "norm (a + - b) \<le> norm a + norm (- b)"
  47.540 -    by (rule norm_triangle_ineq)
  47.541 -  thus ?thesis
  47.542 -    by (simp only: diff_minus norm_minus_cancel)
  47.543 -qed
  47.544 -
  47.545 -lemma norm_diff_ineq:
  47.546 -  fixes a b :: "'a::real_normed_vector"
  47.547 -  shows "norm a - norm b \<le> norm (a + b)"
  47.548 -proof -
  47.549 -  have "norm a - norm (- b) \<le> norm (a - - b)"
  47.550 -    by (rule norm_triangle_ineq2)
  47.551 -  thus ?thesis by simp
  47.552 -qed
  47.553 -
  47.554 -lemma norm_diff_triangle_ineq:
  47.555 -  fixes a b c d :: "'a::real_normed_vector"
  47.556 -  shows "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
  47.557 -proof -
  47.558 -  have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))"
  47.559 -    by (simp add: diff_minus add_ac)
  47.560 -  also have "\<dots> \<le> norm (a - c) + norm (b - d)"
  47.561 -    by (rule norm_triangle_ineq)
  47.562 -  finally show ?thesis .
  47.563 -qed
  47.564 -
  47.565 -lemma abs_norm_cancel [simp]:
  47.566 -  fixes a :: "'a::real_normed_vector"
  47.567 -  shows "\<bar>norm a\<bar> = norm a"
  47.568 -by (rule abs_of_nonneg [OF norm_ge_zero])
  47.569 -
  47.570 -lemma norm_add_less:
  47.571 -  fixes x y :: "'a::real_normed_vector"
  47.572 -  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x + y) < r + s"
  47.573 -by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono])
  47.574 -
  47.575 -lemma norm_mult_less:
  47.576 -  fixes x y :: "'a::real_normed_algebra"
  47.577 -  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x * y) < r * s"
  47.578 -apply (rule order_le_less_trans [OF norm_mult_ineq])
  47.579 -apply (simp add: mult_strict_mono')
  47.580 -done
  47.581 -
  47.582 -lemma norm_of_real [simp]:
  47.583 -  "norm (of_real r :: 'a::real_normed_algebra_1) = \<bar>r\<bar>"
  47.584 -unfolding of_real_def by (simp add: norm_scaleR)
  47.585 -
  47.586 -lemma norm_number_of [simp]:
  47.587 -  "norm (number_of w::'a::{number_ring,real_normed_algebra_1})
  47.588 -    = \<bar>number_of w\<bar>"
  47.589 -by (subst of_real_number_of_eq [symmetric], rule norm_of_real)
  47.590 -
  47.591 -lemma norm_of_int [simp]:
  47.592 -  "norm (of_int z::'a::real_normed_algebra_1) = \<bar>of_int z\<bar>"
  47.593 -by (subst of_real_of_int_eq [symmetric], rule norm_of_real)
  47.594 -
  47.595 -lemma norm_of_nat [simp]:
  47.596 -  "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n"
  47.597 -apply (subst of_real_of_nat_eq [symmetric])
  47.598 -apply (subst norm_of_real, simp)
  47.599 -done
  47.600 -
  47.601 -lemma nonzero_norm_inverse:
  47.602 -  fixes a :: "'a::real_normed_div_algebra"
  47.603 -  shows "a \<noteq> 0 \<Longrightarrow> norm (inverse a) = inverse (norm a)"
  47.604 -apply (rule inverse_unique [symmetric])
  47.605 -apply (simp add: norm_mult [symmetric])
  47.606 -done
  47.607 -
  47.608 -lemma norm_inverse:
  47.609 -  fixes a :: "'a::{real_normed_div_algebra,division_by_zero}"
  47.610 -  shows "norm (inverse a) = inverse (norm a)"
  47.611 -apply (case_tac "a = 0", simp)
  47.612 -apply (erule nonzero_norm_inverse)
  47.613 -done
  47.614 -
  47.615 -lemma nonzero_norm_divide:
  47.616 -  fixes a b :: "'a::real_normed_field"
  47.617 -  shows "b \<noteq> 0 \<Longrightarrow> norm (a / b) = norm a / norm b"
  47.618 -by (simp add: divide_inverse norm_mult nonzero_norm_inverse)
  47.619 -
  47.620 -lemma norm_divide:
  47.621 -  fixes a b :: "'a::{real_normed_field,division_by_zero}"
  47.622 -  shows "norm (a / b) = norm a / norm b"
  47.623 -by (simp add: divide_inverse norm_mult norm_inverse)
  47.624 -
  47.625 -lemma norm_power_ineq:
  47.626 -  fixes x :: "'a::{real_normed_algebra_1,recpower}"
  47.627 -  shows "norm (x ^ n) \<le> norm x ^ n"
  47.628 -proof (induct n)
  47.629 -  case 0 show "norm (x ^ 0) \<le> norm x ^ 0" by simp
  47.630 -next
  47.631 -  case (Suc n)
  47.632 -  have "norm (x * x ^ n) \<le> norm x * norm (x ^ n)"
  47.633 -    by (rule norm_mult_ineq)
  47.634 -  also from Suc have "\<dots> \<le> norm x * norm x ^ n"
  47.635 -    using norm_ge_zero by (rule mult_left_mono)
  47.636 -  finally show "norm (x ^ Suc n) \<le> norm x ^ Suc n"
  47.637 -    by (simp add: power_Suc)
  47.638 -qed
  47.639 -
  47.640 -lemma norm_power:
  47.641 -  fixes x :: "'a::{real_normed_div_algebra,recpower}"
  47.642 -  shows "norm (x ^ n) = norm x ^ n"
  47.643 -by (induct n) (simp_all add: power_Suc norm_mult)
  47.644 -
  47.645 -
  47.646 -subsection {* Sign function *}
  47.647 -
  47.648 -lemma norm_sgn:
  47.649 -  "norm (sgn(x::'a::real_normed_vector)) = (if x = 0 then 0 else 1)"
  47.650 -by (simp add: sgn_div_norm norm_scaleR)
  47.651 -
  47.652 -lemma sgn_zero [simp]: "sgn(0::'a::real_normed_vector) = 0"
  47.653 -by (simp add: sgn_div_norm)
  47.654 -
  47.655 -lemma sgn_zero_iff: "(sgn(x::'a::real_normed_vector) = 0) = (x = 0)"
  47.656 -by (simp add: sgn_div_norm)
  47.657 -
  47.658 -lemma sgn_minus: "sgn (- x) = - sgn(x::'a::real_normed_vector)"
  47.659 -by (simp add: sgn_div_norm)
  47.660 -
  47.661 -lemma sgn_scaleR:
  47.662 -  "sgn (scaleR r x) = scaleR (sgn r) (sgn(x::'a::real_normed_vector))"
  47.663 -by (simp add: sgn_div_norm norm_scaleR mult_ac)
  47.664 -
  47.665 -lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1"
  47.666 -by (simp add: sgn_div_norm)
  47.667 -
  47.668 -lemma sgn_of_real:
  47.669 -  "sgn (of_real r::'a::real_normed_algebra_1) = of_real (sgn r)"
  47.670 -unfolding of_real_def by (simp only: sgn_scaleR sgn_one)
  47.671 -
  47.672 -lemma sgn_mult:
  47.673 -  fixes x y :: "'a::real_normed_div_algebra"
  47.674 -  shows "sgn (x * y) = sgn x * sgn y"
  47.675 -by (simp add: sgn_div_norm norm_mult mult_commute)
  47.676 -
  47.677 -lemma real_sgn_eq: "sgn (x::real) = x / \<bar>x\<bar>"
  47.678 -by (simp add: sgn_div_norm divide_inverse)
  47.679 -
  47.680 -lemma real_sgn_pos: "0 < (x::real) \<Longrightarrow> sgn x = 1"
  47.681 -unfolding real_sgn_eq by simp
  47.682 -
  47.683 -lemma real_sgn_neg: "(x::real) < 0 \<Longrightarrow> sgn x = -1"
  47.684 -unfolding real_sgn_eq by simp
  47.685 -
  47.686 -
  47.687 -subsection {* Bounded Linear and Bilinear Operators *}
  47.688 -
  47.689 -locale bounded_linear = additive +
  47.690 -  constrains f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
  47.691 -  assumes scaleR: "f (scaleR r x) = scaleR r (f x)"
  47.692 -  assumes bounded: "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
  47.693 -begin
  47.694 -
  47.695 -lemma pos_bounded:
  47.696 -  "\<exists>K>0. \<forall>x. norm (f x) \<le> norm x * K"
  47.697 -proof -
  47.698 -  obtain K where K: "\<And>x. norm (f x) \<le> norm x * K"
  47.699 -    using bounded by fast
  47.700 -  show ?thesis
  47.701 -  proof (intro exI impI conjI allI)
  47.702 -    show "0 < max 1 K"
  47.703 -      by (rule order_less_le_trans [OF zero_less_one le_maxI1])
  47.704 -  next
  47.705 -    fix x
  47.706 -    have "norm (f x) \<le> norm x * K" using K .
  47.707 -    also have "\<dots> \<le> norm x * max 1 K"
  47.708 -      by (rule mult_left_mono [OF le_maxI2 norm_ge_zero])
  47.709 -    finally show "norm (f x) \<le> norm x * max 1 K" .
  47.710 -  qed
  47.711 -qed
  47.712 -
  47.713 -lemma nonneg_bounded:
  47.714 -  "\<exists>K\<ge>0. \<forall>x. norm (f x) \<le> norm x * K"
  47.715 -proof -
  47.716 -  from pos_bounded
  47.717 -  show ?thesis by (auto intro: order_less_imp_le)
  47.718 -qed
  47.719 -
  47.720 -end
  47.721 -
  47.722 -locale bounded_bilinear =
  47.723 -  fixes prod :: "['a::real_normed_vector, 'b::real_normed_vector]
  47.724 -                 \<Rightarrow> 'c::real_normed_vector"
  47.725 -    (infixl "**" 70)
  47.726 -  assumes add_left: "prod (a + a') b = prod a b + prod a' b"
  47.727 -  assumes add_right: "prod a (b + b') = prod a b + prod a b'"
  47.728 -  assumes scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)"
  47.729 -  assumes scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)"
  47.730 -  assumes bounded: "\<exists>K. \<forall>a b. norm (prod a b) \<le> norm a * norm b * K"
  47.731 -begin
  47.732 -
  47.733 -lemma pos_bounded:
  47.734 -  "\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
  47.735 -apply (cut_tac bounded, erule exE)
  47.736 -apply (rule_tac x="max 1 K" in exI, safe)
  47.737 -apply (rule order_less_le_trans [OF zero_less_one le_maxI1])
  47.738 -apply (drule spec, drule spec, erule order_trans)
  47.739 -apply (rule mult_left_mono [OF le_maxI2])
  47.740 -apply (intro mult_nonneg_nonneg norm_ge_zero)
  47.741 -done
  47.742 -
  47.743 -lemma nonneg_bounded:
  47.744 -  "\<exists>K\<ge>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
  47.745 -proof -
  47.746 -  from pos_bounded
  47.747 -  show ?thesis by (auto intro: order_less_imp_le)
  47.748 -qed
  47.749 -
  47.750 -lemma additive_right: "additive (\<lambda>b. prod a b)"
  47.751 -by (rule additive.intro, rule add_right)
  47.752 -
  47.753 -lemma additive_left: "additive (\<lambda>a. prod a b)"
  47.754 -by (rule additive.intro, rule add_left)
  47.755 -
  47.756 -lemma zero_left: "prod 0 b = 0"
  47.757 -by (rule additive.zero [OF additive_left])
  47.758 -
  47.759 -lemma zero_right: "prod a 0 = 0"
  47.760 -by (rule additive.zero [OF additive_right])
  47.761 -
  47.762 -lemma minus_left: "prod (- a) b = - prod a b"
  47.763 -by (rule additive.minus [OF additive_left])
  47.764 -
  47.765 -lemma minus_right: "prod a (- b) = - prod a b"
  47.766 -by (rule additive.minus [OF additive_right])
  47.767 -
  47.768 -lemma diff_left:
  47.769 -  "prod (a - a') b = prod a b - prod a' b"
  47.770 -by (rule additive.diff [OF additive_left])
  47.771 -
  47.772 -lemma diff_right:
  47.773 -  "prod a (b - b') = prod a b - prod a b'"
  47.774 -by (rule additive.diff [OF additive_right])
  47.775 -
  47.776 -lemma bounded_linear_left:
  47.777 -  "bounded_linear (\<lambda>a. a ** b)"
  47.778 -apply (unfold_locales)
  47.779 -apply (rule add_left)
  47.780 -apply (rule scaleR_left)
  47.781 -apply (cut_tac bounded, safe)
  47.782 -apply (rule_tac x="norm b * K" in exI)
  47.783 -apply (simp add: mult_ac)
  47.784 -done
  47.785 -
  47.786 -lemma bounded_linear_right:
  47.787 -  "bounded_linear (\<lambda>b. a ** b)"
  47.788 -apply (unfold_locales)
  47.789 -apply (rule add_right)
  47.790 -apply (rule scaleR_right)
  47.791 -apply (cut_tac bounded, safe)
  47.792 -apply (rule_tac x="norm a * K" in exI)
  47.793 -apply (simp add: mult_ac)
  47.794 -done
  47.795 -
  47.796 -lemma prod_diff_prod:
  47.797 -  "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)"
  47.798 -by (simp add: diff_left diff_right)
  47.799 -
  47.800 -end
  47.801 -
  47.802 -interpretation mult:
  47.803 -  bounded_bilinear ["op * :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra"]
  47.804 -apply (rule bounded_bilinear.intro)
  47.805 -apply (rule left_distrib)
  47.806 -apply (rule right_distrib)
  47.807 -apply (rule mult_scaleR_left)
  47.808 -apply (rule mult_scaleR_right)
  47.809 -apply (rule_tac x="1" in exI)
  47.810 -apply (simp add: norm_mult_ineq)
  47.811 -done
  47.812 -
  47.813 -interpretation mult_left:
  47.814 -  bounded_linear ["(\<lambda>x::'a::real_normed_algebra. x * y)"]
  47.815 -by (rule mult.bounded_linear_left)
  47.816 -
  47.817 -interpretation mult_right:
  47.818 -  bounded_linear ["(\<lambda>y::'a::real_normed_algebra. x * y)"]
  47.819 -by (rule mult.bounded_linear_right)
  47.820 -
  47.821 -interpretation divide:
  47.822 -  bounded_linear ["(\<lambda>x::'a::real_normed_field. x / y)"]
  47.823 -unfolding divide_inverse by (rule mult.bounded_linear_left)
  47.824 -
  47.825 -interpretation scaleR: bounded_bilinear ["scaleR"]
  47.826 -apply (rule bounded_bilinear.intro)
  47.827 -apply (rule scaleR_left_distrib)
  47.828 -apply (rule scaleR_right_distrib)
  47.829 -apply simp
  47.830 -apply (rule scaleR_left_commute)
  47.831 -apply (rule_tac x="1" in exI)
  47.832 -apply (simp add: norm_scaleR)
  47.833 -done
  47.834 -
  47.835 -interpretation scaleR_left: bounded_linear ["\<lambda>r. scaleR r x"]
  47.836 -by (rule scaleR.bounded_linear_left)
  47.837 -
  47.838 -interpretation scaleR_right: bounded_linear ["\<lambda>x. scaleR r x"]
  47.839 -by (rule scaleR.bounded_linear_right)
  47.840 -
  47.841 -interpretation of_real: bounded_linear ["\<lambda>r. of_real r"]
  47.842 -unfolding of_real_def by (rule scaleR.bounded_linear_left)
  47.843 -
  47.844 -end
    48.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    48.2 +++ b/src/HOL/RealVector.thy	Mon Dec 29 14:08:08 2008 +0100
    48.3 @@ -0,0 +1,841 @@
    48.4 +(*  Title:      HOL/RealVector.thy
    48.5 +    Author:     Brian Huffman
    48.6 +*)
    48.7 +
    48.8 +header {* Vector Spaces and Algebras over the Reals *}
    48.9 +
   48.10 +theory RealVector
   48.11 +imports RealPow
   48.12 +begin
   48.13 +
   48.14 +subsection {* Locale for additive functions *}
   48.15 +
   48.16 +locale additive =
   48.17 +  fixes f :: "'a::ab_group_add \<Rightarrow> 'b::ab_group_add"
   48.18 +  assumes add: "f (x + y) = f x + f y"
   48.19 +begin
   48.20 +
   48.21 +lemma zero: "f 0 = 0"
   48.22 +proof -
   48.23 +  have "f 0 = f (0 + 0)" by simp
   48.24 +  also have "\<dots> = f 0 + f 0" by (rule add)
   48.25 +  finally show "f 0 = 0" by simp
   48.26 +qed
   48.27 +
   48.28 +lemma minus: "f (- x) = - f x"
   48.29 +proof -
   48.30 +  have "f (- x) + f x = f (- x + x)" by (rule add [symmetric])
   48.31 +  also have "\<dots> = - f x + f x" by (simp add: zero)
   48.32 +  finally show "f (- x) = - f x" by (rule add_right_imp_eq)
   48.33 +qed
   48.34 +
   48.35 +lemma diff: "f (x - y) = f x - f y"
   48.36 +by (simp add: diff_def add minus)
   48.37 +
   48.38 +lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
   48.39 +apply (cases "finite A")
   48.40 +apply (induct set: finite)
   48.41 +apply (simp add: zero)
   48.42 +apply (simp add: add)
   48.43 +apply (simp add: zero)
   48.44 +done
   48.45 +
   48.46 +end
   48.47 +
   48.48 +subsection {* Vector spaces *}
   48.49 +
   48.50 +locale vector_space =
   48.51 +  fixes scale :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b"
   48.52 +  assumes scale_right_distrib: "scale a (x + y) = scale a x + scale a y"
   48.53 +  and scale_left_distrib: "scale (a + b) x = scale a x + scale b x"
   48.54 +  and scale_scale [simp]: "scale a (scale b x) = scale (a * b) x"
   48.55 +  and scale_one [simp]: "scale 1 x = x"
   48.56 +begin
   48.57 +
   48.58 +lemma scale_left_commute:
   48.59 +  "scale a (scale b x) = scale b (scale a x)"
   48.60 +by (simp add: mult_commute)
   48.61 +
   48.62 +lemma scale_zero_left [simp]: "scale 0 x = 0"
   48.63 +  and scale_minus_left [simp]: "scale (- a) x = - (scale a x)"
   48.64 +  and scale_left_diff_distrib: "scale (a - b) x = scale a x - scale b x"
   48.65 +proof -
   48.66 +  interpret s: additive ["\<lambda>a. scale a x"]
   48.67 +    proof qed (rule scale_left_distrib)
   48.68 +  show "scale 0 x = 0" by (rule s.zero)
   48.69 +  show "scale (- a) x = - (scale a x)" by (rule s.minus)
   48.70 +  show "scale (a - b) x = scale a x - scale b x" by (rule s.diff)
   48.71 +qed
   48.72 +
   48.73 +lemma scale_zero_right [simp]: "scale a 0 = 0"
   48.74 +  and scale_minus_right [simp]: "scale a (- x) = - (scale a x)"
   48.75 +  and scale_right_diff_distrib: "scale a (x - y) = scale a x - scale a y"
   48.76 +proof -
   48.77 +  interpret s: additive ["\<lambda>x. scale a x"]
   48.78 +    proof qed (rule scale_right_distrib)
   48.79 +  show "scale a 0 = 0" by (rule s.zero)
   48.80 +  show "scale a (- x) = - (scale a x)" by (rule s.minus)
   48.81 +  show "scale a (x - y) = scale a x - scale a y" by (rule s.diff)
   48.82 +qed
   48.83 +
   48.84 +lemma scale_eq_0_iff [simp]:
   48.85 +  "scale a x = 0 \<longleftrightarrow> a = 0 \<or> x = 0"
   48.86 +proof cases
   48.87 +  assume "a = 0" thus ?thesis by simp
   48.88 +next
   48.89 +  assume anz [simp]: "a \<noteq> 0"
   48.90 +  { assume "scale a x = 0"
   48.91 +    hence "scale (inverse a) (scale a x) = 0" by simp
   48.92 +    hence "x = 0" by simp }
   48.93 +  thus ?thesis by force
   48.94 +qed
   48.95 +
   48.96 +lemma scale_left_imp_eq:
   48.97 +  "\<lbrakk>a \<noteq> 0; scale a x = scale a y\<rbrakk> \<Longrightarrow> x = y"
   48.98 +proof -
   48.99 +  assume nonzero: "a \<noteq> 0"
  48.100 +  assume "scale a x = scale a y"
  48.101 +  hence "scale a (x - y) = 0"
  48.102 +     by (simp add: scale_right_diff_distrib)
  48.103 +  hence "x - y = 0" by (simp add: nonzero)
  48.104 +  thus "x = y" by (simp only: right_minus_eq)
  48.105 +qed
  48.106 +
  48.107 +lemma scale_right_imp_eq:
  48.108 +  "\<lbrakk>x \<noteq> 0; scale a x = scale b x\<rbrakk> \<Longrightarrow> a = b"
  48.109 +proof -
  48.110 +  assume nonzero: "x \<noteq> 0"
  48.111 +  assume "scale a x = scale b x"
  48.112 +  hence "scale (a - b) x = 0"
  48.113 +     by (simp add: scale_left_diff_distrib)
  48.114 +  hence "a - b = 0" by (simp add: nonzero)
  48.115 +  thus "a = b" by (simp only: right_minus_eq)
  48.116 +qed
  48.117 +
  48.118 +lemma scale_cancel_left:
  48.119 +  "scale a x = scale a y \<longleftrightarrow> x = y \<or> a = 0"
  48.120 +by (auto intro: scale_left_imp_eq)
  48.121 +
  48.122 +lemma scale_cancel_right:
  48.123 +  "scale a x = scale b x \<longleftrightarrow> a = b \<or> x = 0"
  48.124 +by (auto intro: scale_right_imp_eq)
  48.125 +
  48.126 +end
  48.127 +
  48.128 +subsection {* Real vector spaces *}
  48.129 +
  48.130 +class scaleR = type +
  48.131 +  fixes scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "*\<^sub>R" 75)
  48.132 +begin
  48.133 +
  48.134 +abbreviation
  48.135 +  divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a" (infixl "'/\<^sub>R" 70)
  48.136 +where
  48.137 +  "x /\<^sub>R r == scaleR (inverse r) x"
  48.138 +
  48.139 +end
  48.140 +
  48.141 +instantiation real :: scaleR
  48.142 +begin
  48.143 +
  48.144 +definition
  48.145 +  real_scaleR_def [simp]: "scaleR a x = a * x"
  48.146 +
  48.147 +instance ..
  48.148 +
  48.149 +end
  48.150 +
  48.151 +class real_vector = scaleR + ab_group_add +
  48.152 +  assumes scaleR_right_distrib: "scaleR a (x + y) = scaleR a x + scaleR a y"
  48.153 +  and scaleR_left_distrib: "scaleR (a + b) x = scaleR a x + scaleR b x"
  48.154 +  and scaleR_scaleR [simp]: "scaleR a (scaleR b x) = scaleR (a * b) x"
  48.155 +  and scaleR_one [simp]: "scaleR 1 x = x"
  48.156 +
  48.157 +interpretation real_vector:
  48.158 +  vector_space ["scaleR :: real \<Rightarrow> 'a \<Rightarrow> 'a::real_vector"]
  48.159 +apply unfold_locales
  48.160 +apply (rule scaleR_right_distrib)
  48.161 +apply (rule scaleR_left_distrib)
  48.162 +apply (rule scaleR_scaleR)
  48.163 +apply (rule scaleR_one)
  48.164 +done
  48.165 +
  48.166 +text {* Recover original theorem names *}
  48.167 +
  48.168 +lemmas scaleR_left_commute = real_vector.scale_left_commute
  48.169 +lemmas scaleR_zero_left = real_vector.scale_zero_left
  48.170 +lemmas scaleR_minus_left = real_vector.scale_minus_left
  48.171 +lemmas scaleR_left_diff_distrib = real_vector.scale_left_diff_distrib
  48.172 +lemmas scaleR_zero_right = real_vector.scale_zero_right
  48.173 +lemmas scaleR_minus_right = real_vector.scale_minus_right
  48.174 +lemmas scaleR_right_diff_distrib = real_vector.scale_right_diff_distrib
  48.175 +lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff
  48.176 +lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq
  48.177 +lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq
  48.178 +lemmas scaleR_cancel_left = real_vector.scale_cancel_left
  48.179 +lemmas scaleR_cancel_right = real_vector.scale_cancel_right
  48.180 +
  48.181 +class real_algebra = real_vector + ring +
  48.182 +  assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)"
  48.183 +  and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"
  48.184 +
  48.185 +class real_algebra_1 = real_algebra + ring_1
  48.186 +
  48.187 +class real_div_algebra = real_algebra_1 + division_ring
  48.188 +
  48.189 +class real_field = real_div_algebra + field
  48.190 +
  48.191 +instance real :: real_field
  48.192 +apply (intro_classes, unfold real_scaleR_def)
  48.193 +apply (rule right_distrib)
  48.194 +apply (rule left_distrib)
  48.195 +apply (rule mult_assoc [symmetric])
  48.196 +apply (rule mult_1_left)
  48.197 +apply (rule mult_assoc)
  48.198 +apply (rule mult_left_commute)
  48.199 +done
  48.200 +
  48.201 +interpretation scaleR_left: additive ["(\<lambda>a. scaleR a x::'a::real_vector)"]
  48.202 +proof qed (rule scaleR_left_distrib)
  48.203 +
  48.204 +interpretation scaleR_right: additive ["(\<lambda>x. scaleR a x::'a::real_vector)"]
  48.205 +proof qed (rule scaleR_right_distrib)
  48.206 +
  48.207 +lemma nonzero_inverse_scaleR_distrib:
  48.208 +  fixes x :: "'a::real_div_algebra" shows
  48.209 +  "\<lbrakk>a \<noteq> 0; x \<noteq> 0\<rbrakk> \<Longrightarrow> inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
  48.210 +by (rule inverse_unique, simp)
  48.211 +
  48.212 +lemma inverse_scaleR_distrib:
  48.213 +  fixes x :: "'a::{real_div_algebra,division_by_zero}"
  48.214 +  shows "inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
  48.215 +apply (case_tac "a = 0", simp)
  48.216 +apply (case_tac "x = 0", simp)
  48.217 +apply (erule (1) nonzero_inverse_scaleR_distrib)
  48.218 +done
  48.219 +
  48.220 +
  48.221 +subsection {* Embedding of the Reals into any @{text real_algebra_1}:
  48.222 +@{term of_real} *}
  48.223 +
  48.224 +definition
  48.225 +  of_real :: "real \<Rightarrow> 'a::real_algebra_1" where
  48.226 +  "of_real r = scaleR r 1"
  48.227 +
  48.228 +lemma scaleR_conv_of_real: "scaleR r x = of_real r * x"
  48.229 +by (simp add: of_real_def)
  48.230 +
  48.231 +lemma of_real_0 [simp]: "of_real 0 = 0"
  48.232 +by (simp add: of_real_def)
  48.233 +
  48.234 +lemma of_real_1 [simp]: "of_real 1 = 1"
  48.235 +by (simp add: of_real_def)
  48.236 +
  48.237 +lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y"
  48.238 +by (simp add: of_real_def scaleR_left_distrib)
  48.239 +
  48.240 +lemma of_real_minus [simp]: "of_real (- x) = - of_real x"
  48.241 +by (simp add: of_real_def)
  48.242 +
  48.243 +lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y"
  48.244 +by (simp add: of_real_def scaleR_left_diff_distrib)
  48.245 +
  48.246 +lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y"
  48.247 +by (simp add: of_real_def mult_commute)
  48.248 +
  48.249 +lemma nonzero_of_real_inverse:
  48.250 +  "x \<noteq> 0 \<Longrightarrow> of_real (inverse x) =
  48.251 +   inverse (of_real x :: 'a::real_div_algebra)"
  48.252 +by (simp add: of_real_def nonzero_inverse_scaleR_distrib)
  48.253 +
  48.254 +lemma of_real_inverse [simp]:
  48.255 +  "of_real (inverse x) =
  48.256 +   inverse (of_real x :: 'a::{real_div_algebra,division_by_zero})"
  48.257 +by (simp add: of_real_def inverse_scaleR_distrib)
  48.258 +
  48.259 +lemma nonzero_of_real_divide:
  48.260 +  "y \<noteq> 0 \<Longrightarrow> of_real (x / y) =
  48.261 +   (of_real x / of_real y :: 'a::real_field)"
  48.262 +by (simp add: divide_inverse nonzero_of_real_inverse)
  48.263 +
  48.264 +lemma of_real_divide [simp]:
  48.265 +  "of_real (x / y) =
  48.266 +   (of_real x / of_real y :: 'a::{real_field,division_by_zero})"
  48.267 +by (simp add: divide_inverse)
  48.268 +
  48.269 +lemma of_real_power [simp]:
  48.270 +  "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1,recpower}) ^ n"
  48.271 +by (induct n) (simp_all add: power_Suc)
  48.272 +
  48.273 +lemma of_real_eq_iff [simp]: "(of_real x = of_real y) = (x = y)"
  48.274 +by (simp add: of_real_def scaleR_cancel_right)
  48.275 +
  48.276 +lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified]
  48.277 +
  48.278 +lemma of_real_eq_id [simp]: "of_real = (id :: real \<Rightarrow> real)"
  48.279 +proof
  48.280 +  fix r
  48.281 +  show "of_real r = id r"
  48.282 +    by (simp add: of_real_def)
  48.283 +qed
  48.284 +
  48.285 +text{*Collapse nested embeddings*}
  48.286 +lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n"
  48.287 +by (induct n) auto
  48.288 +
  48.289 +lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z"
  48.290 +by (cases z rule: int_diff_cases, simp)
  48.291 +
  48.292 +lemma of_real_number_of_eq:
  48.293 +  "of_real (number_of w) = (number_of w :: 'a::{number_ring,real_algebra_1})"
  48.294 +by (simp add: number_of_eq)
  48.295 +
  48.296 +text{*Every real algebra has characteristic zero*}
  48.297 +instance real_algebra_1 < ring_char_0
  48.298 +proof
  48.299 +  fix m n :: nat
  48.300 +  have "(of_real (of_nat m) = (of_real (of_nat n)::'a)) = (m = n)"
  48.301 +    by (simp only: of_real_eq_iff of_nat_eq_iff)
  48.302 +  thus "(of_nat m = (of_nat n::'a)) = (m = n)"
  48.303 +    by (simp only: of_real_of_nat_eq)
  48.304 +qed
  48.305 +
  48.306 +instance real_field < field_char_0 ..
  48.307 +
  48.308 +
  48.309 +subsection {* The Set of Real Numbers *}
  48.310 +
  48.311 +definition
  48.312 +  Reals :: "'a::real_algebra_1 set" where
  48.313 +  [code del]: "Reals \<equiv> range of_real"
  48.314 +
  48.315 +notation (xsymbols)
  48.316 +  Reals  ("\<real>")
  48.317 +
  48.318 +lemma Reals_of_real [simp]: "of_real r \<in> Reals"
  48.319 +by (simp add: Reals_def)
  48.320 +
  48.321 +lemma Reals_of_int [simp]: "of_int z \<in> Reals"
  48.322 +by (subst of_real_of_int_eq [symmetric], rule Reals_of_real)
  48.323 +
  48.324 +lemma Reals_of_nat [simp]: "of_nat n \<in> Reals"
  48.325 +by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real)
  48.326 +
  48.327 +lemma Reals_number_of [simp]:
  48.328 +  "(number_of w::'a::{number_ring,real_algebra_1}) \<in> Reals"
  48.329 +by (subst of_real_number_of_eq [symmetric], rule Reals_of_real)
  48.330 +
  48.331 +lemma Reals_0 [simp]: "0 \<in> Reals"
  48.332 +apply (unfold Reals_def)
  48.333 +apply (rule range_eqI)
  48.334 +apply (rule of_real_0 [symmetric])
  48.335 +done
  48.336 +
  48.337 +lemma Reals_1 [simp]: "1 \<in> Reals"
  48.338 +apply (unfold Reals_def)
  48.339 +apply (rule range_eqI)
  48.340 +apply (rule of_real_1 [symmetric])
  48.341 +done
  48.342 +
  48.343 +lemma Reals_add [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a + b \<in> Reals"
  48.344 +apply (auto simp add: Reals_def)
  48.345 +apply (rule range_eqI)
  48.346 +apply (rule of_real_add [symmetric])
  48.347 +done
  48.348 +
  48.349 +lemma Reals_minus [simp]: "a \<in> Reals \<Longrightarrow> - a \<in> Reals"
  48.350 +apply (auto simp add: Reals_def)
  48.351 +apply (rule range_eqI)
  48.352 +apply (rule of_real_minus [symmetric])
  48.353 +done
  48.354 +
  48.355 +lemma Reals_diff [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a - b \<in> Reals"
  48.356 +apply (auto simp add: Reals_def)
  48.357 +apply (rule range_eqI)
  48.358 +apply (rule of_real_diff [symmetric])
  48.359 +done
  48.360 +
  48.361 +lemma Reals_mult [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a * b \<in> Reals"
  48.362 +apply (auto simp add: Reals_def)
  48.363 +apply (rule range_eqI)
  48.364 +apply (rule of_real_mult [symmetric])
  48.365 +done
  48.366 +
  48.367 +lemma nonzero_Reals_inverse:
  48.368 +  fixes a :: "'a::real_div_algebra"
  48.369 +  shows "\<lbrakk>a \<in> Reals; a \<noteq> 0\<rbrakk> \<Longrightarrow> inverse a \<in> Reals"
  48.370 +apply (auto simp add: Reals_def)
  48.371 +apply (rule range_eqI)
  48.372 +apply (erule nonzero_of_real_inverse [symmetric])
  48.373 +done
  48.374 +
  48.375 +lemma Reals_inverse [simp]:
  48.376 +  fixes a :: "'a::{real_div_algebra,division_by_zero}"
  48.377 +  shows "a \<in> Reals \<Longrightarrow> inverse a \<in> Reals"
  48.378 +apply (auto simp add: Reals_def)
  48.379 +apply (rule range_eqI)
  48.380 +apply (rule of_real_inverse [symmetric])
  48.381 +done
  48.382 +
  48.383 +lemma nonzero_Reals_divide:
  48.384 +  fixes a b :: "'a::real_field"
  48.385 +  shows "\<lbrakk>a \<in> Reals; b \<in> Reals; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
  48.386 +apply (auto simp add: Reals_def)
  48.387 +apply (rule range_eqI)
  48.388 +apply (erule nonzero_of_real_divide [symmetric])
  48.389 +done
  48.390 +
  48.391 +lemma Reals_divide [simp]:
  48.392 +  fixes a b :: "'a::{real_field,division_by_zero}"
  48.393 +  shows "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
  48.394 +apply (auto simp add: Reals_def)
  48.395 +apply (rule range_eqI)
  48.396 +apply (rule of_real_divide [symmetric])
  48.397 +done
  48.398 +
  48.399 +lemma Reals_power [simp]:
  48.400 +  fixes a :: "'a::{real_algebra_1,recpower}"
  48.401 +  shows "a \<in> Reals \<Longrightarrow> a ^ n \<in> Reals"
  48.402 +apply (auto simp add: Reals_def)
  48.403 +apply (rule range_eqI)
  48.404 +apply (rule of_real_power [symmetric])
  48.405 +done
  48.406 +
  48.407 +lemma Reals_cases [cases set: Reals]:
  48.408 +  assumes "q \<in> \<real>"
  48.409 +  obtains (of_real) r where "q = of_real r"
  48.410 +  unfolding Reals_def
  48.411 +proof -
  48.412 +  from `q \<in> \<real>` have "q \<in> range of_real" unfolding Reals_def .
  48.413 +  then obtain r where "q = of_real r" ..
  48.414 +  then show thesis ..
  48.415 +qed
  48.416 +
  48.417 +lemma Reals_induct [case_names of_real, induct set: Reals]:
  48.418 +  "q \<in> \<real> \<Longrightarrow> (\<And>r. P (of_real r)) \<Longrightarrow> P q"
  48.419 +  by (rule Reals_cases) auto
  48.420 +
  48.421 +
  48.422 +subsection {* Real normed vector spaces *}
  48.423 +
  48.424 +class norm = type +
  48.425 +  fixes norm :: "'a \<Rightarrow> real"
  48.426 +
  48.427 +instantiation real :: norm
  48.428 +begin
  48.429 +
  48.430 +definition
  48.431 +  real_norm_def [simp]: "norm r \<equiv> \<bar>r\<bar>"
  48.432 +
  48.433 +instance ..
  48.434 +
  48.435 +end
  48.436 +
  48.437 +class sgn_div_norm = scaleR + norm + sgn +
  48.438 +  assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x"
  48.439 +
  48.440 +class real_normed_vector = real_vector + sgn_div_norm +
  48.441 +  assumes norm_ge_zero [simp]: "0 \<le> norm x"
  48.442 +  and norm_eq_zero [simp]: "norm x = 0 \<longleftrightarrow> x = 0"
  48.443 +  and norm_triangle_ineq: "norm (x + y) \<le> norm x + norm y"
  48.444 +  and norm_scaleR: "norm (scaleR a x) = \<bar>a\<bar> * norm x"
  48.445 +
  48.446 +class real_normed_algebra = real_algebra + real_normed_vector +
  48.447 +  assumes norm_mult_ineq: "norm (x * y) \<le> norm x * norm y"
  48.448 +
  48.449 +class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra +
  48.450 +  assumes norm_one [simp]: "norm 1 = 1"
  48.451 +
  48.452 +class real_normed_div_algebra = real_div_algebra + real_normed_vector +
  48.453 +  assumes norm_mult: "norm (x * y) = norm x * norm y"
  48.454 +
  48.455 +class real_normed_field = real_field + real_normed_div_algebra
  48.456 +
  48.457 +instance real_normed_div_algebra < real_normed_algebra_1
  48.458 +proof
  48.459 +  fix x y :: 'a
  48.460 +  show "norm (x * y) \<le> norm x * norm y"
  48.461 +    by (simp add: norm_mult)
  48.462 +next
  48.463 +  have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)"
  48.464 +    by (rule norm_mult)
  48.465 +  thus "norm (1::'a) = 1" by simp
  48.466 +qed
  48.467 +
  48.468 +instance real :: real_normed_field
  48.469 +apply (intro_classes, unfold real_norm_def real_scaleR_def)
  48.470 +apply (simp add: real_sgn_def)
  48.471 +apply (rule abs_ge_zero)
  48.472 +apply (rule abs_eq_0)
  48.473 +apply (rule abs_triangle_ineq)
  48.474 +apply (rule abs_mult)
  48.475 +apply (rule abs_mult)
  48.476 +done
  48.477 +
  48.478 +lemma norm_zero [simp]: "norm (0::'a::real_normed_vector) = 0"
  48.479 +by simp
  48.480 +
  48.481 +lemma zero_less_norm_iff [simp]:
  48.482 +  fixes x :: "'a::real_normed_vector"
  48.483 +  shows "(0 < norm x) = (x \<noteq> 0)"
  48.484 +by (simp add: order_less_le)
  48.485 +
  48.486 +lemma norm_not_less_zero [simp]:
  48.487 +  fixes x :: "'a::real_normed_vector"
  48.488 +  shows "\<not> norm x < 0"
  48.489 +by (simp add: linorder_not_less)
  48.490 +
  48.491 +lemma norm_le_zero_iff [simp]:
  48.492 +  fixes x :: "'a::real_normed_vector"
  48.493 +  shows "(norm x \<le> 0) = (x = 0)"
  48.494 +by (simp add: order_le_less)
  48.495 +
  48.496 +lemma norm_minus_cancel [simp]:
  48.497 +  fixes x :: "'a::real_normed_vector"
  48.498 +  shows "norm (- x) = norm x"
  48.499 +proof -
  48.500 +  have "norm (- x) = norm (scaleR (- 1) x)"
  48.501 +    by (simp only: scaleR_minus_left scaleR_one)
  48.502 +  also have "\<dots> = \<bar>- 1\<bar> * norm x"
  48.503 +    by (rule norm_scaleR)
  48.504 +  finally show ?thesis by simp
  48.505 +qed
  48.506 +
  48.507 +lemma norm_minus_commute:
  48.508 +  fixes a b :: "'a::real_normed_vector"
  48.509 +  shows "norm (a - b) = norm (b - a)"
  48.510 +proof -
  48.511 +  have "norm (- (b - a)) = norm (b - a)"
  48.512 +    by (rule norm_minus_cancel)
  48.513 +  thus ?thesis by simp
  48.514 +qed
  48.515 +
  48.516 +lemma norm_triangle_ineq2:
  48.517 +  fixes a b :: "'a::real_normed_vector"
  48.518 +  shows "norm a - norm b \<le> norm (a - b)"
  48.519 +proof -
  48.520 +  have "norm (a - b + b) \<le> norm (a - b) + norm b"
  48.521 +    by (rule norm_triangle_ineq)
  48.522 +  thus ?thesis by simp
  48.523 +qed
  48.524 +
  48.525 +lemma norm_triangle_ineq3:
  48.526 +  fixes a b :: "'a::real_normed_vector"
  48.527 +  shows "\<bar>norm a - norm b\<bar> \<le> norm (a - b)"
  48.528 +apply (subst abs_le_iff)
  48.529 +apply auto
  48.530 +apply (rule norm_triangle_ineq2)
  48.531 +apply (subst norm_minus_commute)
  48.532 +apply (rule norm_triangle_ineq2)
  48.533 +done
  48.534 +
  48.535 +lemma norm_triangle_ineq4:
  48.536 +  fixes a b :: "'a::real_normed_vector"
  48.537 +  shows "norm (a - b) \<le> norm a + norm b"
  48.538 +proof -
  48.539 +  have "norm (a + - b) \<le> norm a + norm (- b)"
  48.540 +    by (rule norm_triangle_ineq)
  48.541 +  thus ?thesis
  48.542 +    by (simp only: diff_minus norm_minus_cancel)
  48.543 +qed
  48.544 +
  48.545 +lemma norm_diff_ineq:
  48.546 +  fixes a b :: "'a::real_normed_vector"
  48.547 +  shows "norm a - norm b \<le> norm (a + b)"
  48.548 +proof -
  48.549 +  have "norm a - norm (- b) \<le> norm (a - - b)"
  48.550 +    by (rule norm_triangle_ineq2)
  48.551 +  thus ?thesis by simp
  48.552 +qed
  48.553 +
  48.554 +lemma norm_diff_triangle_ineq:
  48.555 +  fixes a b c d :: "'a::real_normed_vector"
  48.556 +  shows "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
  48.557 +proof -
  48.558 +  have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))"
  48.559 +    by (simp add: diff_minus add_ac)
  48.560 +  also have "\<dots> \<le> norm (a - c) + norm (b - d)"
  48.561 +    by (rule norm_triangle_ineq)
  48.562 +  finally show ?thesis .
  48.563 +qed
  48.564 +
  48.565 +lemma abs_norm_cancel [simp]:
  48.566 +  fixes a :: "'a::real_normed_vector"
  48.567 +  shows "\<bar>norm a\<bar> = norm a"
  48.568 +by (rule abs_of_nonneg [OF norm_ge_zero])
  48.569 +
  48.570 +lemma norm_add_less:
  48.571 +  fixes x y :: "'a::real_normed_vector"
  48.572 +  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x + y) < r + s"
  48.573 +by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono])
  48.574 +
  48.575 +lemma norm_mult_less:
  48.576 +  fixes x y :: "'a::real_normed_algebra"
  48.577 +  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x * y) < r * s"
  48.578 +apply (rule order_le_less_trans [OF norm_mult_ineq])
  48.579 +apply (simp add: mult_strict_mono')
  48.580 +done
  48.581 +
  48.582 +lemma norm_of_real [simp]:
  48.583 +  "norm (of_real r :: 'a::real_normed_algebra_1) = \<bar>r\<bar>"
  48.584 +unfolding of_real_def by (simp add: norm_scaleR)
  48.585 +
  48.586 +lemma norm_number_of [simp]:
  48.587 +  "norm (number_of w::'a::{number_ring,real_normed_algebra_1})
  48.588 +    = \<bar>number_of w\<bar>"
  48.589 +by (subst of_real_number_of_eq [symmetric], rule norm_of_real)
  48.590 +
  48.591 +lemma norm_of_int [simp]:
  48.592 +  "norm (of_int z::'a::real_normed_algebra_1) = \<bar>of_int z\<bar>"
  48.593 +by (subst of_real_of_int_eq [symmetric], rule norm_of_real)
  48.594 +
  48.595 +lemma norm_of_nat [simp]:
  48.596 +  "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n"
  48.597 +apply (subst of_real_of_nat_eq [symmetric])
  48.598 +apply (subst norm_of_real, simp)
  48.599 +done
  48.600 +
  48.601 +lemma nonzero_norm_inverse:
  48.602 +  fixes a :: "'a::real_normed_div_algebra"
  48.603 +  shows "a \<noteq> 0 \<Longrightarrow> norm (inverse a) = inverse (norm a)"
  48.604 +apply (rule inverse_unique [symmetric])
  48.605 +apply (simp add: norm_mult [symmetric])
  48.606 +done
  48.607 +
  48.608 +lemma norm_inverse:
  48.609 +  fixes a :: "'a::{real_normed_div_algebra,division_by_zero}"
  48.610 +  shows "norm (inverse a) = inverse (norm a)"
  48.611 +apply (case_tac "a = 0", simp)
  48.612 +apply (erule nonzero_norm_inverse)
  48.613 +done
  48.614 +
  48.615 +lemma nonzero_norm_divide:
  48.616 +  fixes a b :: "'a::real_normed_field"
  48.617 +  shows "b \<noteq> 0 \<Longrightarrow> norm (a / b) = norm a / norm b"
  48.618 +by (simp add: divide_inverse norm_mult nonzero_norm_inverse)
  48.619 +
  48.620 +lemma norm_divide:
  48.621 +  fixes a b :: "'a::{real_normed_field,division_by_zero}"
  48.622 +  shows "norm (a / b) = norm a / norm b"
  48.623 +by (simp add: divide_inverse norm_mult norm_inverse)
  48.624 +
  48.625 +lemma norm_power_ineq:
  48.626 +  fixes x :: "'a::{real_normed_algebra_1,recpower}"
  48.627 +  shows "norm (x ^ n) \<le> norm x ^ n"
  48.628 +proof (induct n)
  48.629 +  case 0 show "norm (x ^ 0) \<le> norm x ^ 0" by simp
  48.630 +next
  48.631 +  case (Suc n)
  48.632 +  have "norm (x * x ^ n) \<le> norm x * norm (x ^ n)"
  48.633 +    by (rule norm_mult_ineq)
  48.634 +  also from Suc have "\<dots> \<le> norm x * norm x ^ n"
  48.635 +    using norm_ge_zero by (rule mult_left_mono)
  48.636 +  finally show "norm (x ^ Suc n) \<le> norm x ^ Suc n"
  48.637 +    by (simp add: power_Suc)
  48.638 +qed
  48.639 +
  48.640 +lemma norm_power:
  48.641 +  fixes x :: "'a::{real_normed_div_algebra,recpower}"
  48.642 +  shows "norm (x ^ n) = norm x ^ n"
  48.643 +by (induct n) (simp_all add: power_Suc norm_mult)
  48.644 +
  48.645 +
  48.646 +subsection {* Sign function *}
  48.647 +
  48.648 +lemma norm_sgn:
  48.649 +  "norm (sgn(x::'a::real_normed_vector)) = (if x = 0 then 0 else 1)"
  48.650 +by (simp add: sgn_div_norm norm_scaleR)
  48.651 +
  48.652 +lemma sgn_zero [simp]: "sgn(0::'a::real_normed_vector) = 0"
  48.653 +by (simp add: sgn_div_norm)
  48.654 +
  48.655 +lemma sgn_zero_iff: "(sgn(x::'a::real_normed_vector) = 0) = (x = 0)"
  48.656 +by (simp add: sgn_div_norm)
  48.657 +
  48.658 +lemma sgn_minus: "sgn (- x) = - sgn(x::'a::real_normed_vector)"
  48.659 +by (simp add: sgn_div_norm)
  48.660 +
  48.661 +lemma sgn_scaleR:
  48.662 +  "sgn (scaleR r x) = scaleR (sgn r) (sgn(x::'a::real_normed_vector))"
  48.663 +by (simp add: sgn_div_norm norm_scaleR mult_ac)
  48.664 +
  48.665 +lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1"
  48.666 +by (simp add: sgn_div_norm)
  48.667 +
  48.668 +lemma sgn_of_real:
  48.669 +  "sgn (of_real r::'a::real_normed_algebra_1) = of_real (sgn r)"
  48.670 +unfolding of_real_def by (simp only: sgn_scaleR sgn_one)
  48.671 +
  48.672 +lemma sgn_mult:
  48.673 +  fixes x y :: "'a::real_normed_div_algebra"
  48.674 +  shows "sgn (x * y) = sgn x * sgn y"
  48.675 +by (simp add: sgn_div_norm norm_mult mult_commute)
  48.676 +
  48.677 +lemma real_sgn_eq: "sgn (x::real) = x / \<bar>x\<bar>"
  48.678 +by (simp add: sgn_div_norm divide_inverse)
  48.679 +
  48.680 +lemma real_sgn_pos: "0 < (x::real) \<Longrightarrow> sgn x = 1"
  48.681 +unfolding real_sgn_eq by simp
  48.682 +
  48.683 +lemma real_sgn_neg: "(x::real) < 0 \<Longrightarrow> sgn x = -1"
  48.684 +unfolding real_sgn_eq by simp
  48.685 +
  48.686 +
  48.687 +subsection {* Bounded Linear and Bilinear Operators *}
  48.688 +
  48.689 +locale bounded_linear = additive +
  48.690 +  constrains f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
  48.691 +  assumes scaleR: "f (scaleR r x) = scaleR r (f x)"
  48.692 +  assumes bounded: "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
  48.693 +begin
  48.694 +
  48.695 +lemma pos_bounded:
  48.696 +  "\<exists>K>0. \<forall>x. norm (f x) \<le> norm x * K"
  48.697 +proof -
  48.698 +  obtain K where K: "\<And>x. norm (f x) \<le> norm x * K"
  48.699 +    using bounded by fast
  48.700 +  show ?thesis
  48.701 +  proof (intro exI impI conjI allI)
  48.702 +    show "0 < max 1 K"
  48.703 +      by (rule order_less_le_trans [OF zero_less_one le_maxI1])
  48.704 +  next
  48.705 +    fix x
  48.706 +    have "norm (f x) \<le> norm x * K" using K .
  48.707 +    also have "\<dots> \<le> norm x * max 1 K"
  48.708 +      by (rule mult_left_mono [OF le_maxI2 norm_ge_zero])
  48.709 +    finally show "norm (f x) \<le> norm x * max 1 K" .
  48.710 +  qed
  48.711 +qed
  48.712 +
  48.713 +lemma nonneg_bounded:
  48.714 +  "\<exists>K\<ge>0. \<forall>x. norm (f x) \<le> norm x * K"
  48.715 +proof -
  48.716 +  from pos_bounded
  48.717 +  show ?thesis by (auto intro: order_less_imp_le)
  48.718 +qed
  48.719 +
  48.720 +end
  48.721 +
  48.722 +locale bounded_bilinear =
  48.723 +  fixes prod :: "['a::real_normed_vector, 'b::real_normed_vector]
  48.724 +                 \<Rightarrow> 'c::real_normed_vector"
  48.725 +    (infixl "**" 70)
  48.726 +  assumes add_left: "prod (a + a') b = prod a b + prod a' b"
  48.727 +  assumes add_right: "prod a (b + b') = prod a b + prod a b'"
  48.728 +  assumes scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)"
  48.729 +  assumes scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)"
  48.730 +  assumes bounded: "\<exists>K. \<forall>a b. norm (prod a b) \<le> norm a * norm b * K"
  48.731 +begin
  48.732 +
  48.733 +lemma pos_bounded:
  48.734 +  "\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
  48.735 +apply (cut_tac bounded, erule exE)
  48.736 +apply (rule_tac x="max 1 K" in exI, safe)
  48.737 +apply (rule order_less_le_trans [OF zero_less_one le_maxI1])
  48.738 +apply (drule spec, drule spec, erule order_trans)
  48.739 +apply (rule mult_left_mono [OF le_maxI2])
  48.740 +apply (intro mult_nonneg_nonneg norm_ge_zero)
  48.741 +done
  48.742 +
  48.743 +lemma nonneg_bounded:
  48.744 +  "\<exists>K\<ge>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
  48.745 +proof -
  48.746 +  from pos_bounded
  48.747 +  show ?thesis by (auto intro: order_less_imp_le)
  48.748 +qed
  48.749 +
  48.750 +lemma additive_right: "additive (\<lambda>b. prod a b)"
  48.751 +by (rule additive.intro, rule add_right)
  48.752 +
  48.753 +lemma additive_left: "additive (\<lambda>a. prod a b)"
  48.754 +by (rule additive.intro, rule add_left)
  48.755 +
  48.756 +lemma zero_left: "prod 0 b = 0"
  48.757 +by (rule additive.zero [OF additive_left])
  48.758 +
  48.759 +lemma zero_right: "prod a 0 = 0"
  48.760 +by (rule additive.zero [OF additive_right])
  48.761 +
  48.762 +lemma minus_left: "prod (- a) b = - prod a b"
  48.763 +by (rule additive.minus [OF additive_left])
  48.764 +
  48.765 +lemma minus_right: "prod a (- b) = - prod a b"
  48.766 +by (rule additive.minus [OF additive_right])
  48.767 +
  48.768 +lemma diff_left:
  48.769 +  "prod (a - a') b = prod a b - prod a' b"
  48.770 +by (rule additive.diff [OF additive_left])
  48.771 +
  48.772 +lemma diff_right:
  48.773 +  "prod a (b - b') = prod a b - prod a b'"
  48.774 +by (rule additive.diff [OF additive_right])
  48.775 +
  48.776 +lemma bounded_linear_left:
  48.777 +  "bounded_linear (\<lambda>a. a ** b)"
  48.778 +apply (unfold_locales)
  48.779 +apply (rule add_left)
  48.780 +apply (rule scaleR_left)
  48.781 +apply (cut_tac bounded, safe)
  48.782 +apply (rule_tac x="norm b * K" in exI)
  48.783 +apply (simp add: mult_ac)
  48.784 +done
  48.785 +
  48.786 +lemma bounded_linear_right:
  48.787 +  "bounded_linear (\<lambda>b. a ** b)"
  48.788 +apply (unfold_locales)
  48.789 +apply (rule add_right)
  48.790 +apply (rule scaleR_right)
  48.791 +apply (cut_tac bounded, safe)
  48.792 +apply (rule_tac x="norm a * K" in exI)
  48.793 +apply (simp add: mult_ac)
  48.794 +done
  48.795 +
  48.796 +lemma prod_diff_prod:
  48.797 +  "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)"
  48.798 +by (simp add: diff_left diff_right)
  48.799 +
  48.800 +end
  48.801 +
  48.802 +interpretation mult:
  48.803 +  bounded_bilinear ["op * :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra"]
  48.804 +apply (rule bounded_bilinear.intro)
  48.805 +apply (rule left_distrib)
  48.806 +apply (rule right_distrib)
  48.807 +apply (rule mult_scaleR_left)
  48.808 +apply (rule mult_scaleR_right)
  48.809 +apply (rule_tac x="1" in exI)
  48.810 +apply (simp add: norm_mult_ineq)
  48.811 +done
  48.812 +
  48.813 +interpretation mult_left:
  48.814 +  bounded_linear ["(\<lambda>x::'a::real_normed_algebra. x * y)"]
  48.815 +by (rule mult.bounded_linear_left)
  48.816 +
  48.817 +interpretation mult_right:
  48.818 +  bounded_linear ["(\<lambda>y::'a::real_normed_algebra. x * y)"]
  48.819 +by (rule mult.bounded_linear_right)
  48.820 +
  48.821 +interpretation divide:
  48.822 +  bounded_linear ["(\<lambda>x::'a::real_normed_field. x / y)"]
  48.823 +unfolding divide_inverse by (rule mult.bounded_linear_left)
  48.824 +
  48.825 +interpretation scaleR: bounded_bilinear ["scaleR"]
  48.826 +apply (rule bounded_bilinear.intro)
  48.827 +apply (rule scaleR_left_distrib)
  48.828 +apply (rule scaleR_right_distrib)
  48.829 +apply simp
  48.830 +apply (rule scaleR_left_commute)
  48.831 +apply (rule_tac x="1" in exI)
  48.832 +apply (simp add: norm_scaleR)
  48.833 +done
  48.834 +
  48.835 +interpretation scaleR_left: bounded_linear ["\<lambda>r. scaleR r x"]
  48.836 +by (rule scaleR.bounded_linear_left)
  48.837 +
  48.838 +interpretation scaleR_right: bounded_linear ["\<lambda>x. scaleR r x"]
  48.839 +by (rule scaleR.bounded_linear_right)
  48.840 +
  48.841 +interpretation of_real: bounded_linear ["\<lambda>r. of_real r"]
  48.842 +unfolding of_real_def by (rule scaleR.bounded_linear_left)
  48.843 +
  48.844 +end
    49.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    49.2 +++ b/src/HOL/SEQ.thy	Mon Dec 29 14:08:08 2008 +0100
    49.3 @@ -0,0 +1,1136 @@
    49.4 +(*  Title       : SEQ.thy
    49.5 +    Author      : Jacques D. Fleuriot
    49.6 +    Copyright   : 1998  University of Cambridge
    49.7 +    Description : Convergence of sequences and series
    49.8 +    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
    49.9 +    Additional contributions by Jeremy Avigad and Brian Huffman
   49.10 +*)
   49.11 +
   49.12 +header {* Sequences and Convergence *}
   49.13 +
   49.14 +theory SEQ
   49.15 +imports RealVector RComplete
   49.16 +begin
   49.17 +
   49.18 +definition
   49.19 +  Zseq :: "[nat \<Rightarrow> 'a::real_normed_vector] \<Rightarrow> bool" where
   49.20 +    --{*Standard definition of sequence converging to zero*}
   49.21 +  [code del]: "Zseq X = (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. norm (X n) < r)"
   49.22 +
   49.23 +definition
   49.24 +  LIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool"
   49.25 +    ("((_)/ ----> (_))" [60, 60] 60) where
   49.26 +    --{*Standard definition of convergence of sequence*}
   49.27 +  [code del]: "X ----> L = (\<forall>r. 0 < r --> (\<exists>no. \<forall>n. no \<le> n --> norm (X n - L) < r))"
   49.28 +
   49.29 +definition
   49.30 +  lim :: "(nat => 'a::real_normed_vector) => 'a" where
   49.31 +    --{*Standard definition of limit using choice operator*}
   49.32 +  "lim X = (THE L. X ----> L)"
   49.33 +
   49.34 +definition
   49.35 +  convergent :: "(nat => 'a::real_normed_vector) => bool" where
   49.36 +    --{*Standard definition of convergence*}
   49.37 +  "convergent X = (\<exists>L. X ----> L)"
   49.38 +
   49.39 +definition
   49.40 +  Bseq :: "(nat => 'a::real_normed_vector) => bool" where
   49.41 +    --{*Standard definition for bounded sequence*}
   49.42 +  [code del]: "Bseq X = (\<exists>K>0.\<forall>n. norm (X n) \<le> K)"
   49.43 +
   49.44 +definition
   49.45 +  monoseq :: "(nat=>real)=>bool" where
   49.46 +    --{*Definition for monotonicity*}
   49.47 +  [code del]: "monoseq X = ((\<forall>m. \<forall>n\<ge>m. X m \<le> X n) | (\<forall>m. \<forall>n\<ge>m. X n \<le> X m))"
   49.48 +
   49.49 +definition
   49.50 +  subseq :: "(nat => nat) => bool" where
   49.51 +    --{*Definition of subsequence*}
   49.52 +  [code del]:   "subseq f = (\<forall>m. \<forall>n>m. (f m) < (f n))"
   49.53 +
   49.54 +definition
   49.55 +  Cauchy :: "(nat => 'a::real_normed_vector) => bool" where
   49.56 +    --{*Standard definition of the Cauchy condition*}
   49.57 +  [code del]: "Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. norm (X m - X n) < e)"
   49.58 +
   49.59 +
   49.60 +subsection {* Bounded Sequences *}
   49.61 +
   49.62 +lemma BseqI': assumes K: "\<And>n. norm (X n) \<le> K" shows "Bseq X"
   49.63 +unfolding Bseq_def
   49.64 +proof (intro exI conjI allI)
   49.65 +  show "0 < max K 1" by simp
   49.66 +next
   49.67 +  fix n::nat
   49.68 +  have "norm (X n) \<le> K" by (rule K)
   49.69 +  thus "norm (X n) \<le> max K 1" by simp
   49.70 +qed
   49.71 +
   49.72 +lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
   49.73 +unfolding Bseq_def by auto
   49.74 +
   49.75 +lemma BseqI2': assumes K: "\<forall>n\<ge>N. norm (X n) \<le> K" shows "Bseq X"
   49.76 +proof (rule BseqI')
   49.77 +  let ?A = "norm ` X ` {..N}"
   49.78 +  have 1: "finite ?A" by simp
   49.79 +  fix n::nat
   49.80 +  show "norm (X n) \<le> max K (Max ?A)"
   49.81 +  proof (cases rule: linorder_le_cases)
   49.82 +    assume "n \<ge> N"
   49.83 +    hence "norm (X n) \<le> K" using K by simp
   49.84 +    thus "norm (X n) \<le> max K (Max ?A)" by simp
   49.85 +  next
   49.86 +    assume "n \<le> N"
   49.87 +    hence "norm (X n) \<in> ?A" by simp
   49.88 +    with 1 have "norm (X n) \<le> Max ?A" by (rule Max_ge)
   49.89 +    thus "norm (X n) \<le> max K (Max ?A)" by simp
   49.90 +  qed
   49.91 +qed
   49.92 +
   49.93 +lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))"
   49.94 +unfolding Bseq_def by auto
   49.95 +
   49.96 +lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X"
   49.97 +apply (erule BseqE)
   49.98 +apply (rule_tac N="k" and K="K" in BseqI2')
   49.99 +apply clarify
  49.100 +apply (drule_tac x="n - k" in spec, simp)
  49.101 +done
  49.102 +
  49.103 +
  49.104 +subsection {* Sequences That Converge to Zero *}
  49.105 +
  49.106 +lemma ZseqI:
  49.107 +  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r) \<Longrightarrow> Zseq X"
  49.108 +unfolding Zseq_def by simp
  49.109 +
  49.110 +lemma ZseqD:
  49.111 +  "\<lbrakk>Zseq X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r"
  49.112 +unfolding Zseq_def by simp
  49.113 +
  49.114 +lemma Zseq_zero: "Zseq (\<lambda>n. 0)"
  49.115 +unfolding Zseq_def by simp
  49.116 +
  49.117 +lemma Zseq_const_iff: "Zseq (\<lambda>n. k) = (k = 0)"
  49.118 +unfolding Zseq_def by force
  49.119 +
  49.120 +lemma Zseq_norm_iff: "Zseq (\<lambda>n. norm (X n)) = Zseq (\<lambda>n. X n)"
  49.121 +unfolding Zseq_def by simp
  49.122 +
  49.123 +lemma Zseq_imp_Zseq:
  49.124 +  assumes X: "Zseq X"
  49.125 +  assumes Y: "\<And>n. norm (Y n) \<le> norm (X n) * K"
  49.126 +  shows "Zseq (\<lambda>n. Y n)"
  49.127 +proof (cases)
  49.128 +  assume K: "0 < K"
  49.129 +  show ?thesis
  49.130 +  proof (rule ZseqI)
  49.131 +    fix r::real assume "0 < r"
  49.132 +    hence "0 < r / K"
  49.133 +      using K by (rule divide_pos_pos)
  49.134 +    then obtain N where "\<forall>n\<ge>N. norm (X n) < r / K"
  49.135 +      using ZseqD [OF X] by fast
  49.136 +    hence "\<forall>n\<ge>N. norm (X n) * K < r"
  49.137 +      by (simp add: pos_less_divide_eq K)
  49.138 +    hence "\<forall>n\<ge>N. norm (Y n) < r"
  49.139 +      by (simp add: order_le_less_trans [OF Y])
  49.140 +    thus "\<exists>N. \<forall>n\<ge>N. norm (Y n) < r" ..
  49.141 +  qed
  49.142 +next
  49.143 +  assume "\<not> 0 < K"
  49.144 +  hence K: "K \<le> 0" by (simp only: linorder_not_less)
  49.145 +  {
  49.146 +    fix n::nat
  49.147 +    have "norm (Y n) \<le> norm (X n) * K" by (rule Y)
  49.148 +    also have "\<dots> \<le> norm (X n) * 0"
  49.149 +      using K norm_ge_zero by (rule mult_left_mono)
  49.150 +    finally have "norm (Y n) = 0" by simp
  49.151 +  }
  49.152 +  thus ?thesis by (simp add: Zseq_zero)
  49.153 +qed
  49.154 +
  49.155 +lemma Zseq_le: "\<lbrakk>Zseq Y; \<forall>n. norm (X n) \<le> norm (Y n)\<rbrakk> \<Longrightarrow> Zseq X"
  49.156 +by (erule_tac K="1" in Zseq_imp_Zseq, simp)
  49.157 +
  49.158 +lemma Zseq_add:
  49.159 +  assumes X: "Zseq X"
  49.160 +  assumes Y: "Zseq Y"
  49.161 +  shows "Zseq (\<lambda>n. X n + Y n)"
  49.162 +proof (rule ZseqI)
  49.163 +  fix r::real assume "0 < r"
  49.164 +  hence r: "0 < r / 2" by simp
  49.165 +  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r/2"
  49.166 +    using ZseqD [OF X r] by fast
  49.167 +  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < r/2"
  49.168 +    using ZseqD [OF Y r] by fast
  49.169 +  show "\<exists>N. \<forall>n\<ge>N. norm (X n + Y n) < r"
  49.170 +  proof (intro exI allI impI)
  49.171 +    fix n assume n: "max M N \<le> n"
  49.172 +    have "norm (X n + Y n) \<le> norm (X n) + norm (Y n)"
  49.173 +      by (rule norm_triangle_ineq)
  49.174 +    also have "\<dots> < r/2 + r/2"
  49.175 +    proof (rule add_strict_mono)
  49.176 +      from M n show "norm (X n) < r/2" by simp
  49.177 +      from N n show "norm (Y n) < r/2" by simp
  49.178 +    qed
  49.179 +    finally show "norm (X n + Y n) < r" by simp
  49.180 +  qed
  49.181 +qed
  49.182 +
  49.183 +lemma Zseq_minus: "Zseq X \<Longrightarrow> Zseq (\<lambda>n. - X n)"
  49.184 +unfolding Zseq_def by simp
  49.185 +
  49.186 +lemma Zseq_diff: "\<lbrakk>Zseq X; Zseq Y\<rbrakk> \<Longrightarrow> Zseq (\<lambda>n. X n - Y n)"
  49.187 +by (simp only: diff_minus Zseq_add Zseq_minus)
  49.188 +
  49.189 +lemma (in bounded_linear) Zseq:
  49.190 +  assumes X: "Zseq X"
  49.191 +  shows "Zseq (\<lambda>n. f (X n))"
  49.192 +proof -
  49.193 +  obtain K where "\<And>x. norm (f x) \<le> norm x * K"
  49.194 +    using bounded by fast
  49.195 +  with X show ?thesis
  49.196 +    by (rule Zseq_imp_Zseq)
  49.197 +qed
  49.198 +
  49.199 +lemma (in bounded_bilinear) Zseq:
  49.200 +  assumes X: "Zseq X"
  49.201 +  assumes Y: "Zseq Y"
  49.202 +  shows "Zseq (\<lambda>n. X n ** Y n)"
  49.203 +proof (rule ZseqI)
  49.204 +  fix r::real assume r: "0 < r"
  49.205 +  obtain K where K: "0 < K"
  49.206 +    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
  49.207 +    using pos_bounded by fast
  49.208 +  from K have K': "0 < inverse K"
  49.209 +    by (rule positive_imp_inverse_positive)
  49.210 +  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r"
  49.211 +    using ZseqD [OF X r] by fast
  49.212 +  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < inverse K"
  49.213 +    using ZseqD [OF Y K'] by fast
  49.214 +  show "\<exists>N. \<forall>n\<ge>N. norm (X n ** Y n) < r"
  49.215 +  proof (intro exI allI impI)
  49.216 +    fix n assume n: "max M N \<le> n"
  49.217 +    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
  49.218 +      by (rule norm_le)
  49.219 +    also have "norm (X n) * norm (Y n) * K < r * inverse K * K"
  49.220 +    proof (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero K)
  49.221 +      from M n show Xn: "norm (X n) < r" by simp
  49.222 +      from N n show Yn: "norm (Y n) < inverse K" by simp
  49.223 +    qed
  49.224 +    also from K have "r * inverse K * K = r" by simp
  49.225 +    finally show "norm (X n ** Y n) < r" .
  49.226 +  qed
  49.227 +qed
  49.228 +
  49.229 +lemma (in bounded_bilinear) Zseq_prod_Bseq:
  49.230 +  assumes X: "Zseq X"
  49.231 +  assumes Y: "Bseq Y"
  49.232 +  shows "Zseq (\<lambda>n. X n ** Y n)"
  49.233 +proof -
  49.234 +  obtain K where K: "0 \<le> K"
  49.235 +    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
  49.236 +    using nonneg_bounded by fast
  49.237 +  obtain B where B: "0 < B"
  49.238 +    and norm_Y: "\<And>n. norm (Y n) \<le> B"
  49.239 +    using Y [unfolded Bseq_def] by fast
  49.240 +  from X show ?thesis
  49.241 +  proof (rule Zseq_imp_Zseq)
  49.242 +    fix n::nat
  49.243 +    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
  49.244 +      by (rule norm_le)
  49.245 +    also have "\<dots> \<le> norm (X n) * B * K"
  49.246 +      by (intro mult_mono' order_refl norm_Y norm_ge_zero
  49.247 +                mult_nonneg_nonneg K)
  49.248 +    also have "\<dots> = norm (X n) * (B * K)"
  49.249 +      by (rule mult_assoc)
  49.250 +    finally show "norm (X n ** Y n) \<le> norm (X n) * (B * K)" .
  49.251 +  qed
  49.252 +qed
  49.253 +
  49.254 +lemma (in bounded_bilinear) Bseq_prod_Zseq:
  49.255 +  assumes X: "Bseq X"
  49.256 +  assumes Y: "Zseq Y"
  49.257 +  shows "Zseq (\<lambda>n. X n ** Y n)"
  49.258 +proof -
  49.259 +  obtain K where K: "0 \<le> K"
  49.260 +    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
  49.261 +    using nonneg_bounded by fast
  49.262 +  obtain B where B: "0 < B"
  49.263 +    and norm_X: "\<And>n. norm (X n) \<le> B"
  49.264 +    using X [unfolded Bseq_def] by fast
  49.265 +  from Y show ?thesis
  49.266 +  proof (rule Zseq_imp_Zseq)
  49.267 +    fix n::nat
  49.268 +    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
  49.269 +      by (rule norm_le)
  49.270 +    also have "\<dots> \<le> B * norm (Y n) * K"
  49.271 +      by (intro mult_mono' order_refl norm_X norm_ge_zero
  49.272 +                mult_nonneg_nonneg K)
  49.273 +    also have "\<dots> = norm (Y n) * (B * K)"
  49.274 +      by (simp only: mult_ac)
  49.275 +    finally show "norm (X n ** Y n) \<le> norm (Y n) * (B * K)" .
  49.276 +  qed
  49.277 +qed
  49.278 +
  49.279 +lemma (in bounded_bilinear) Zseq_left:
  49.280 +  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. X n ** a)"
  49.281 +by (rule bounded_linear_left [THEN bounded_linear.Zseq])
  49.282 +
  49.283 +lemma (in bounded_bilinear) Zseq_right:
  49.284 +  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. a ** X n)"
  49.285 +by (rule bounded_linear_right [THEN bounded_linear.Zseq])
  49.286 +
  49.287 +lemmas Zseq_mult = mult.Zseq
  49.288 +lemmas Zseq_mult_right = mult.Zseq_right
  49.289 +lemmas Zseq_mult_left = mult.Zseq_left
  49.290 +
  49.291 +
  49.292 +subsection {* Limits of Sequences *}
  49.293 +
  49.294 +lemma LIMSEQ_iff:
  49.295 +      "(X ----> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)"
  49.296 +by (rule LIMSEQ_def)
  49.297 +
  49.298 +lemma LIMSEQ_Zseq_iff: "((\<lambda>n. X n) ----> L) = Zseq (\<lambda>n. X n - L)"
  49.299 +by (simp only: LIMSEQ_def Zseq_def)
  49.300 +
  49.301 +lemma LIMSEQ_I:
  49.302 +  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X ----> L"
  49.303 +by (simp add: LIMSEQ_def)
  49.304 +
  49.305 +lemma LIMSEQ_D:
  49.306 +  "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
  49.307 +by (simp add: LIMSEQ_def)
  49.308 +
  49.309 +lemma LIMSEQ_const: "(\<lambda>n. k) ----> k"
  49.310 +by (simp add: LIMSEQ_def)
  49.311 +
  49.312 +lemma LIMSEQ_const_iff: "(\<lambda>n. k) ----> l = (k = l)"
  49.313 +by (simp add: LIMSEQ_Zseq_iff Zseq_const_iff)
  49.314 +
  49.315 +lemma LIMSEQ_norm: "X ----> a \<Longrightarrow> (\<lambda>n. norm (X n)) ----> norm a"
  49.316 +apply (simp add: LIMSEQ_def, safe)
  49.317 +apply (drule_tac x="r" in spec, safe)
  49.318 +apply (rule_tac x="no" in exI, safe)
  49.319 +apply (drule_tac x="n" in spec, safe)
  49.320 +apply (erule order_le_less_trans [OF norm_triangle_ineq3])
  49.321 +done
  49.322 +
  49.323 +lemma LIMSEQ_ignore_initial_segment:
  49.324 +  "f ----> a \<Longrightarrow> (\<lambda>n. f (n + k)) ----> a"
  49.325 +apply (rule LIMSEQ_I)
  49.326 +apply (drule (1) LIMSEQ_D)
  49.327 +apply (erule exE, rename_tac N)
  49.328 +apply (rule_tac x=N in exI)
  49.329 +apply simp
  49.330 +done
  49.331 +
  49.332 +lemma LIMSEQ_offset:
  49.333 +  "(\<lambda>n. f (n + k)) ----> a \<Longrightarrow> f ----> a"
  49.334 +apply (rule LIMSEQ_I)
  49.335 +apply (drule (1) LIMSEQ_D)
  49.336 +apply (erule exE, rename_tac N)
  49.337 +apply (rule_tac x="N + k" in exI)
  49.338 +apply clarify
  49.339 +apply (drule_tac x="n - k" in spec)
  49.340 +apply (simp add: le_diff_conv2)
  49.341 +done
  49.342 +
  49.343 +lemma LIMSEQ_Suc: "f ----> l \<Longrightarrow> (\<lambda>n. f (Suc n)) ----> l"
  49.344 +by (drule_tac k="1" in LIMSEQ_ignore_initial_segment, simp)
  49.345 +
  49.346 +lemma LIMSEQ_imp_Suc: "(\<lambda>n. f (Suc n)) ----> l \<Longrightarrow> f ----> l"
  49.347 +by (rule_tac k="1" in LIMSEQ_offset, simp)
  49.348 +
  49.349 +lemma LIMSEQ_Suc_iff: "(\<lambda>n. f (Suc n)) ----> l = f ----> l"
  49.350 +by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc)
  49.351 +
  49.352 +lemma add_diff_add:
  49.353 +  fixes a b c d :: "'a::ab_group_add"
  49.354 +  shows "(a + c) - (b + d) = (a - b) + (c - d)"
  49.355 +by simp
  49.356 +
  49.357 +lemma minus_diff_minus:
  49.358 +  fixes a b :: "'a::ab_group_add"
  49.359 +  shows "(- a) - (- b) = - (a - b)"
  49.360 +by simp
  49.361 +
  49.362 +lemma LIMSEQ_add: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) ----> a + b"
  49.363 +by (simp only: LIMSEQ_Zseq_iff add_diff_add Zseq_add)
  49.364 +
  49.365 +lemma LIMSEQ_minus: "X ----> a \<Longrightarrow> (\<lambda>n. - X n) ----> - a"
  49.366 +by (simp only: LIMSEQ_Zseq_iff minus_diff_minus Zseq_minus)
  49.367 +
  49.368 +lemma LIMSEQ_minus_cancel: "(\<lambda>n. - X n) ----> - a \<Longrightarrow> X ----> a"
  49.369 +by (drule LIMSEQ_minus, simp)
  49.370 +
  49.371 +lemma LIMSEQ_diff: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) ----> a - b"
  49.372 +by (simp add: diff_minus LIMSEQ_add LIMSEQ_minus)
  49.373 +
  49.374 +lemma LIMSEQ_unique: "\<lbrakk>X ----> a; X ----> b\<rbrakk> \<Longrightarrow> a = b"
  49.375 +by (drule (1) LIMSEQ_diff, simp add: LIMSEQ_const_iff)
  49.376 +
  49.377 +lemma (in bounded_linear) LIMSEQ:
  49.378 +  "X ----> a \<Longrightarrow> (\<lambda>n. f (X n)) ----> f a"
  49.379 +by (simp only: LIMSEQ_Zseq_iff diff [symmetric] Zseq)
  49.380 +
  49.381 +lemma (in bounded_bilinear) LIMSEQ:
  49.382 +  "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n ** Y n) ----> a ** b"
  49.383 +by (simp only: LIMSEQ_Zseq_iff prod_diff_prod
  49.384 +               Zseq_add Zseq Zseq_left Zseq_right)
  49.385 +
  49.386 +lemma LIMSEQ_mult:
  49.387 +  fixes a b :: "'a::real_normed_algebra"
  49.388 +  shows "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b"
  49.389 +by (rule mult.LIMSEQ)
  49.390 +
  49.391 +lemma inverse_diff_inverse:
  49.392 +  "\<lbrakk>(a::'a::division_ring) \<noteq> 0; b \<noteq> 0\<rbrakk>
  49.393 +   \<Longrightarrow> inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
  49.394 +by (simp add: ring_simps)
  49.395 +
  49.396 +lemma Bseq_inverse_lemma:
  49.397 +  fixes x :: "'a::real_normed_div_algebra"
  49.398 +  shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
  49.399 +apply (subst nonzero_norm_inverse, clarsimp)
  49.400 +apply (erule (1) le_imp_inverse_le)
  49.401 +done
  49.402 +
  49.403 +lemma Bseq_inverse:
  49.404 +  fixes a :: "'a::real_normed_div_algebra"
  49.405 +  assumes X: "X ----> a"
  49.406 +  assumes a: "a \<noteq> 0"
  49.407 +  shows "Bseq (\<lambda>n. inverse (X n))"
  49.408 +proof -
  49.409 +  from a have "0 < norm a" by simp
  49.410 +  hence "\<exists>r>0. r < norm a" by (rule dense)
  49.411 +  then obtain r where r1: "0 < r" and r2: "r < norm a" by fast
  49.412 +  obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> norm (X n - a) < r"
  49.413 +    using LIMSEQ_D [OF X r1] by fast
  49.414 +  show ?thesis
  49.415 +  proof (rule BseqI2' [rule_format])
  49.416 +    fix n assume n: "N \<le> n"
  49.417 +    hence 1: "norm (X n - a) < r" by (rule N)
  49.418 +    hence 2: "X n \<noteq> 0" using r2 by auto
  49.419 +    hence "norm (inverse (X n)) = inverse (norm (X n))"
  49.420 +      by (rule nonzero_norm_inverse)
  49.421 +    also have "\<dots> \<le> inverse (norm a - r)"
  49.422 +    proof (rule le_imp_inverse_le)
  49.423 +      show "0 < norm a - r" using r2 by simp
  49.424 +    next
  49.425 +      have "norm a - norm (X n) \<le> norm (a - X n)"
  49.426 +        by (rule norm_triangle_ineq2)
  49.427 +      also have "\<dots> = norm (X n - a)"
  49.428 +        by (rule norm_minus_commute)
  49.429 +      also have "\<dots> < r" using 1 .
  49.430 +      finally show "norm a - r \<le> norm (X n)" by simp
  49.431 +    qed
  49.432 +    finally show "norm (inverse (X n)) \<le> inverse (norm a - r)" .
  49.433 +  qed
  49.434 +qed
  49.435 +
  49.436 +lemma LIMSEQ_inverse_lemma:
  49.437 +  fixes a :: "'a::real_normed_div_algebra"
  49.438 +  shows "\<lbrakk>X ----> a; a \<noteq> 0; \<forall>n. X n \<noteq> 0\<rbrakk>
  49.439 +         \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> inverse a"
  49.440 +apply (subst LIMSEQ_Zseq_iff)
  49.441 +apply (simp add: inverse_diff_inverse nonzero_imp_inverse_nonzero)
  49.442 +apply (rule Zseq_minus)
  49.443 +apply (rule Zseq_mult_left)
  49.444 +apply (rule mult.Bseq_prod_Zseq)
  49.445 +apply (erule (1) Bseq_inverse)
  49.446 +apply (simp add: LIMSEQ_Zseq_iff)
  49.447 +done
  49.448 +
  49.449 +lemma LIMSEQ_inverse:
  49.450 +  fixes a :: "'a::real_normed_div_algebra"
  49.451 +  assumes X: "X ----> a"
  49.452 +  assumes a: "a \<noteq> 0"
  49.453 +  shows "(\<lambda>n. inverse (X n)) ----> inverse a"
  49.454 +proof -
  49.455 +  from a have "0 < norm a" by simp
  49.456 +  then obtain k where "\<forall>n\<ge>k. norm (X n - a) < norm a"
  49.457 +    using LIMSEQ_D [OF X] by fast
  49.458 +  hence "\<forall>n\<ge>k. X n \<noteq> 0" by auto
  49.459 +  hence k: "\<forall>n. X (n + k) \<noteq> 0" by simp
  49.460 +
  49.461 +  from X have "(\<lambda>n. X (n + k)) ----> a"
  49.462 +    by (rule LIMSEQ_ignore_initial_segment)
  49.463 +  hence "(\<lambda>n. inverse (X (n + k))) ----> inverse a"
  49.464 +    using a k by (rule LIMSEQ_inverse_lemma)
  49.465 +  thus "(\<lambda>n. inverse (X n)) ----> inverse a"
  49.466 +    by (rule LIMSEQ_offset)
  49.467 +qed
  49.468 +
  49.469 +lemma LIMSEQ_divide:
  49.470 +  fixes a b :: "'a::real_normed_field"
  49.471 +  shows "\<lbrakk>X ----> a; Y ----> b; b \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>n. X n / Y n) ----> a / b"
  49.472 +by (simp add: LIMSEQ_mult LIMSEQ_inverse divide_inverse)
  49.473 +
  49.474 +lemma LIMSEQ_pow:
  49.475 +  fixes a :: "'a::{real_normed_algebra,recpower}"
  49.476 +  shows "X ----> a \<Longrightarrow> (\<lambda>n. (X n) ^ m) ----> a ^ m"
  49.477 +by (induct m) (simp_all add: power_Suc LIMSEQ_const LIMSEQ_mult)
  49.478 +
  49.479 +lemma LIMSEQ_setsum:
  49.480 +  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
  49.481 +  shows "(\<lambda>m. \<Sum>n\<in>S. X n m) ----> (\<Sum>n\<in>S. L n)"
  49.482 +proof (cases "finite S")
  49.483 +  case True
  49.484 +  thus ?thesis using n
  49.485 +  proof (induct)
  49.486 +    case empty
  49.487 +    show ?case
  49.488 +      by (simp add: LIMSEQ_const)
  49.489 +  next
  49.490 +    case insert
  49.491 +    thus ?case
  49.492 +      by (simp add: LIMSEQ_add)
  49.493 +  qed
  49.494 +next
  49.495 +  case False
  49.496 +  thus ?thesis
  49.497 +    by (simp add: LIMSEQ_const)
  49.498 +qed
  49.499 +
  49.500 +lemma LIMSEQ_setprod:
  49.501 +  fixes L :: "'a \<Rightarrow> 'b::{real_normed_algebra,comm_ring_1}"
  49.502 +  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
  49.503 +  shows "(\<lambda>m. \<Prod>n\<in>S. X n m) ----> (\<Prod>n\<in>S. L n)"
  49.504 +proof (cases "finite S")
  49.505 +  case True
  49.506 +  thus ?thesis using n
  49.507 +  proof (induct)
  49.508 +    case empty
  49.509 +    show ?case
  49.510 +      by (simp add: LIMSEQ_const)
  49.511 +  next
  49.512 +    case insert
  49.513 +    thus ?case
  49.514 +      by (simp add: LIMSEQ_mult)
  49.515 +  qed
  49.516 +next
  49.517 +  case False
  49.518 +  thus ?thesis
  49.519 +    by (simp add: setprod_def LIMSEQ_const)
  49.520 +qed
  49.521 +
  49.522 +lemma LIMSEQ_add_const: "f ----> a ==> (%n.(f n + b)) ----> a + b"
  49.523 +by (simp add: LIMSEQ_add LIMSEQ_const)
  49.524 +
  49.525 +(* FIXME: delete *)
  49.526 +lemma LIMSEQ_add_minus:
  49.527 +     "[| X ----> a; Y ----> b |] ==> (%n. X n + -Y n) ----> a + -b"
  49.528 +by (simp only: LIMSEQ_add LIMSEQ_minus)
  49.529 +
  49.530 +lemma LIMSEQ_diff_const: "f ----> a ==> (%n.(f n  - b)) ----> a - b"
  49.531 +by (simp add: LIMSEQ_diff LIMSEQ_const)
  49.532 +
  49.533 +lemma LIMSEQ_diff_approach_zero: 
  49.534 +  "g ----> L ==> (%x. f x - g x) ----> 0  ==>
  49.535 +     f ----> L"
  49.536 +  apply (drule LIMSEQ_add)
  49.537 +  apply assumption
  49.538 +  apply simp
  49.539 +done
  49.540 +
  49.541 +lemma LIMSEQ_diff_approach_zero2: 
  49.542 +  "f ----> L ==> (%x. f x - g x) ----> 0  ==>
  49.543 +     g ----> L";
  49.544 +  apply (drule LIMSEQ_diff)
  49.545 +  apply assumption
  49.546 +  apply simp
  49.547 +done
  49.548 +
  49.549 +text{*A sequence tends to zero iff its abs does*}
  49.550 +lemma LIMSEQ_norm_zero: "((\<lambda>n. norm (X n)) ----> 0) = (X ----> 0)"
  49.551 +by (simp add: LIMSEQ_def)
  49.552 +
  49.553 +lemma LIMSEQ_rabs_zero: "((%n. \<bar>f n\<bar>) ----> 0) = (f ----> (0::real))"
  49.554 +by (simp add: LIMSEQ_def)
  49.555 +
  49.556 +lemma LIMSEQ_imp_rabs: "f ----> (l::real) ==> (%n. \<bar>f n\<bar>) ----> \<bar>l\<bar>"
  49.557 +by (drule LIMSEQ_norm, simp)
  49.558 +
  49.559 +text{*An unbounded sequence's inverse tends to 0*}
  49.560 +
  49.561 +lemma LIMSEQ_inverse_zero:
  49.562 +  "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
  49.563 +apply (rule LIMSEQ_I)
  49.564 +apply (drule_tac x="inverse r" in spec, safe)
  49.565 +apply (rule_tac x="N" in exI, safe)
  49.566 +apply (drule_tac x="n" in spec, safe)
  49.567 +apply (frule positive_imp_inverse_positive)
  49.568 +apply (frule (1) less_imp_inverse_less)
  49.569 +apply (subgoal_tac "0 < X n", simp)
  49.570 +apply (erule (1) order_less_trans)
  49.571 +done
  49.572 +
  49.573 +text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}
  49.574 +
  49.575 +lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
  49.576 +apply (rule LIMSEQ_inverse_zero, safe)
  49.577 +apply (cut_tac x = r in reals_Archimedean2)
  49.578 +apply (safe, rule_tac x = n in exI)
  49.579 +apply (auto simp add: real_of_nat_Suc)
  49.580 +done
  49.581 +
  49.582 +text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
  49.583 +infinity is now easily proved*}
  49.584 +
  49.585 +lemma LIMSEQ_inverse_real_of_nat_add:
  49.586 +     "(%n. r + inverse(real(Suc n))) ----> r"
  49.587 +by (cut_tac LIMSEQ_add [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
  49.588 +
  49.589 +lemma LIMSEQ_inverse_real_of_nat_add_minus:
  49.590 +     "(%n. r + -inverse(real(Suc n))) ----> r"
  49.591 +by (cut_tac LIMSEQ_add_minus [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
  49.592 +
  49.593 +lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
  49.594 +     "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
  49.595 +by (cut_tac b=1 in
  49.596 +        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat_add_minus], auto)
  49.597 +
  49.598 +lemma LIMSEQ_le_const:
  49.599 +  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. a \<le> X n\<rbrakk> \<Longrightarrow> a \<le> x"
  49.600 +apply (rule ccontr, simp only: linorder_not_le)
  49.601 +apply (drule_tac r="a - x" in LIMSEQ_D, simp)
  49.602 +apply clarsimp
  49.603 +apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI1)
  49.604 +apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI2)
  49.605 +apply simp
  49.606 +done
  49.607 +
  49.608 +lemma LIMSEQ_le_const2:
  49.609 +  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. X n \<le> a\<rbrakk> \<Longrightarrow> x \<le> a"
  49.610 +apply (subgoal_tac "- a \<le> - x", simp)
  49.611 +apply (rule LIMSEQ_le_const)
  49.612 +apply (erule LIMSEQ_minus)
  49.613 +apply simp
  49.614 +done
  49.615 +
  49.616 +lemma LIMSEQ_le:
  49.617 +  "\<lbrakk>X ----> x; Y ----> y; \<exists>N. \<forall>n\<ge>N. X n \<le> Y n\<rbrakk> \<Longrightarrow> x \<le> (y::real)"
  49.618 +apply (subgoal_tac "0 \<le> y - x", simp)
  49.619 +apply (rule LIMSEQ_le_const)
  49.620 +apply (erule (1) LIMSEQ_diff)
  49.621 +apply (simp add: le_diff_eq)
  49.622 +done
  49.623 +
  49.624 +
  49.625 +subsection {* Convergence *}
  49.626 +
  49.627 +lemma limI: "X ----> L ==> lim X = L"
  49.628 +apply (simp add: lim_def)
  49.629 +apply (blast intro: LIMSEQ_unique)
  49.630 +done
  49.631 +
  49.632 +lemma convergentD: "convergent X ==> \<exists>L. (X ----> L)"
  49.633 +by (simp add: convergent_def)
  49.634 +
  49.635 +lemma convergentI: "(X ----> L) ==> convergent X"
  49.636 +by (auto simp add: convergent_def)
  49.637 +
  49.638 +lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)"
  49.639 +by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def)
  49.640 +
  49.641 +lemma convergent_minus_iff: "(convergent X) = (convergent (%n. -(X n)))"
  49.642 +apply (simp add: convergent_def)
  49.643 +apply (auto dest: LIMSEQ_minus)
  49.644 +apply (drule LIMSEQ_minus, auto)
  49.645 +done
  49.646 +
  49.647 +
  49.648 +subsection {* Bounded Monotonic Sequences *}
  49.649 +
  49.650 +text{*Subsequence (alternative definition, (e.g. Hoskins)*}
  49.651 +
  49.652 +lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))"
  49.653 +apply (simp add: subseq_def)
  49.654 +apply (auto dest!: less_imp_Suc_add)
  49.655 +apply (induct_tac k)
  49.656 +apply (auto intro: less_trans)
  49.657 +done
  49.658 +
  49.659 +lemma monoseq_Suc:
  49.660 +   "monoseq X = ((\<forall>n. X n \<le> X (Suc n))
  49.661 +                 | (\<forall>n. X (Suc n) \<le> X n))"
  49.662 +apply (simp add: monoseq_def)
  49.663 +apply (auto dest!: le_imp_less_or_eq)
  49.664 +apply (auto intro!: lessI [THEN less_imp_le] dest!: less_imp_Suc_add)
  49.665 +apply (induct_tac "ka")
  49.666 +apply (auto intro: order_trans)
  49.667 +apply (erule contrapos_np)
  49.668 +apply (induct_tac "k")
  49.669 +apply (auto intro: order_trans)
  49.670 +done
  49.671 +
  49.672 +lemma monoI1: "\<forall>m. \<forall> n \<ge> m. X m \<le> X n ==> monoseq X"
  49.673 +by (simp add: monoseq_def)
  49.674 +
  49.675 +lemma monoI2: "\<forall>m. \<forall> n \<ge> m. X n \<le> X m ==> monoseq X"
  49.676 +by (simp add: monoseq_def)
  49.677 +
  49.678 +lemma mono_SucI1: "\<forall>n. X n \<le> X (Suc n) ==> monoseq X"
  49.679 +by (simp add: monoseq_Suc)
  49.680 +
  49.681 +lemma mono_SucI2: "\<forall>n. X (Suc n) \<le> X n ==> monoseq X"
  49.682 +by (simp add: monoseq_Suc)
  49.683 +
  49.684 +text{*Bounded Sequence*}
  49.685 +
  49.686 +lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)"
  49.687 +by (simp add: Bseq_def)
  49.688 +
  49.689 +lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
  49.690 +by (auto simp add: Bseq_def)
  49.691 +
  49.692 +lemma lemma_NBseq_def:
  49.693 +     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) =
  49.694 +      (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
  49.695 +apply auto
  49.696 + prefer 2 apply force
  49.697 +apply (cut_tac x = K in reals_Archimedean2, clarify)
  49.698 +apply (rule_tac x = n in exI, clarify)
  49.699 +apply (drule_tac x = na in spec)
  49.700 +apply (auto simp add: real_of_nat_Suc)
  49.701 +done
  49.702 +
  49.703 +text{* alternative definition for Bseq *}
  49.704 +lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
  49.705 +apply (simp add: Bseq_def)
  49.706 +apply (simp (no_asm) add: lemma_NBseq_def)
  49.707 +done
  49.708 +
  49.709 +lemma lemma_NBseq_def2:
  49.710 +     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
  49.711 +apply (subst lemma_NBseq_def, auto)
  49.712 +apply (rule_tac x = "Suc N" in exI)
  49.713 +apply (rule_tac [2] x = N in exI)
  49.714 +apply (auto simp add: real_of_nat_Suc)
  49.715 + prefer 2 apply (blast intro: order_less_imp_le)
  49.716 +apply (drule_tac x = n in spec, simp)
  49.717 +done
  49.718 +
  49.719 +(* yet another definition for Bseq *)
  49.720 +lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
  49.721 +by (simp add: Bseq_def lemma_NBseq_def2)
  49.722 +
  49.723 +subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
  49.724 +
  49.725 +lemma Bseq_isUb:
  49.726 +  "!!(X::nat=>real). Bseq X ==> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
  49.727 +by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
  49.728 +
  49.729 +
  49.730 +text{* Use completeness of reals (supremum property)
  49.731 +   to show that any bounded sequence has a least upper bound*}
  49.732 +
  49.733 +lemma Bseq_isLub:
  49.734 +  "!!(X::nat=>real). Bseq X ==>
  49.735 +   \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
  49.736 +by (blast intro: reals_complete Bseq_isUb)
  49.737 +
  49.738 +subsubsection{*A Bounded and Monotonic Sequence Converges*}
  49.739 +
  49.740 +lemma lemma_converg1:
  49.741 +     "!!(X::nat=>real). [| \<forall>m. \<forall> n \<ge> m. X m \<le> X n;
  49.742 +                  isLub (UNIV::real set) {x. \<exists>n. X n = x} (X ma)
  49.743 +               |] ==> \<forall>n \<ge> ma. X n = X ma"
  49.744 +apply safe
  49.745 +apply (drule_tac y = "X n" in isLubD2)
  49.746 +apply (blast dest: order_antisym)+
  49.747 +done
  49.748 +
  49.749 +text{* The best of both worlds: Easier to prove this result as a standard
  49.750 +   theorem and then use equivalence to "transfer" it into the
  49.751 +   equivalent nonstandard form if needed!*}
  49.752 +
  49.753 +lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)"
  49.754 +apply (simp add: LIMSEQ_def)
  49.755 +apply (rule_tac x = "X m" in exI, safe)
  49.756 +apply (rule_tac x = m in exI, safe)
  49.757 +apply (drule spec, erule impE, auto)
  49.758 +done
  49.759 +
  49.760 +lemma lemma_converg2:
  49.761 +   "!!(X::nat=>real).
  49.762 +    [| \<forall>m. X m ~= U;  isLub UNIV {x. \<exists>n. X n = x} U |] ==> \<forall>m. X m < U"
  49.763 +apply safe
  49.764 +apply (drule_tac y = "X m" in isLubD2)
  49.765 +apply (auto dest!: order_le_imp_less_or_eq)
  49.766 +done
  49.767 +
  49.768 +lemma lemma_converg3: "!!(X ::nat=>real). \<forall>m. X m \<le> U ==> isUb UNIV {x. \<exists>n. X n = x} U"
  49.769 +by (rule setleI [THEN isUbI], auto)
  49.770 +
  49.771 +text{* FIXME: @{term "U - T < U"} is redundant *}
  49.772 +lemma lemma_converg4: "!!(X::nat=> real).
  49.773 +               [| \<forall>m. X m ~= U;
  49.774 +                  isLub UNIV {x. \<exists>n. X n = x} U;
  49.775 +                  0 < T;
  49.776 +                  U + - T < U
  49.777 +               |] ==> \<exists>m. U + -T < X m & X m < U"
  49.778 +apply (drule lemma_converg2, assumption)
  49.779 +apply (rule ccontr, simp)
  49.780 +apply (simp add: linorder_not_less)
  49.781 +apply (drule lemma_converg3)
  49.782 +apply (drule isLub_le_isUb, assumption)
  49.783 +apply (auto dest: order_less_le_trans)
  49.784 +done
  49.785 +
  49.786 +text{*A standard proof of the theorem for monotone increasing sequence*}
  49.787 +
  49.788 +lemma Bseq_mono_convergent:
  49.789 +     "[| Bseq X; \<forall>m. \<forall>n \<ge> m. X m \<le> X n |] ==> convergent (X::nat=>real)"
  49.790 +apply (simp add: convergent_def)
  49.791 +apply (frule Bseq_isLub, safe)
  49.792 +apply (case_tac "\<exists>m. X m = U", auto)
  49.793 +apply (blast dest: lemma_converg1 Bmonoseq_LIMSEQ)
  49.794 +(* second case *)
  49.795 +apply (rule_tac x = U in exI)
  49.796 +apply (subst LIMSEQ_iff, safe)
  49.797 +apply (frule lemma_converg2, assumption)
  49.798 +apply (drule lemma_converg4, auto)
  49.799 +apply (rule_tac x = m in exI, safe)
  49.800 +apply (subgoal_tac "X m \<le> X n")
  49.801 + prefer 2 apply blast
  49.802 +apply (drule_tac x=n and P="%m. X m < U" in spec, arith)
  49.803 +done
  49.804 +
  49.805 +lemma Bseq_minus_iff: "Bseq (%n. -(X n)) = Bseq X"
  49.806 +by (simp add: Bseq_def)
  49.807 +
  49.808 +text{*Main monotonicity theorem*}
  49.809 +lemma Bseq_monoseq_convergent: "[| Bseq X; monoseq X |] ==> convergent X"
  49.810 +apply (simp add: monoseq_def, safe)
  49.811 +apply (rule_tac [2] convergent_minus_iff [THEN ssubst])
  49.812 +apply (drule_tac [2] Bseq_minus_iff [THEN ssubst])
  49.813 +apply (auto intro!: Bseq_mono_convergent)
  49.814 +done
  49.815 +
  49.816 +subsubsection{*A Few More Equivalence Theorems for Boundedness*}
  49.817 +
  49.818 +text{*alternative formulation for boundedness*}
  49.819 +lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
  49.820 +apply (unfold Bseq_def, safe)
  49.821 +apply (rule_tac [2] x = "k + norm x" in exI)
  49.822 +apply (rule_tac x = K in exI, simp)
  49.823 +apply (rule exI [where x = 0], auto)
  49.824 +apply (erule order_less_le_trans, simp)
  49.825 +apply (drule_tac x=n in spec, fold diff_def)
  49.826 +apply (drule order_trans [OF norm_triangle_ineq2])
  49.827 +apply simp
  49.828 +done
  49.829 +
  49.830 +text{*alternative formulation for boundedness*}
  49.831 +lemma Bseq_iff3: "Bseq X = (\<exists>k > 0. \<exists>N. \<forall>n. norm(X(n) + -X(N)) \<le> k)"
  49.832 +apply safe
  49.833 +apply (simp add: Bseq_def, safe)
  49.834 +apply (rule_tac x = "K + norm (X N)" in exI)
  49.835 +apply auto
  49.836 +apply (erule order_less_le_trans, simp)
  49.837 +apply (rule_tac x = N in exI, safe)
  49.838 +apply (drule_tac x = n in spec)
  49.839 +apply (rule order_trans [OF norm_triangle_ineq], simp)
  49.840 +apply (auto simp add: Bseq_iff2)
  49.841 +done
  49.842 +
  49.843 +lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
  49.844 +apply (simp add: Bseq_def)
  49.845 +apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
  49.846 +apply (drule_tac x = n in spec, arith)
  49.847 +done
  49.848 +
  49.849 +
  49.850 +subsection {* Cauchy Sequences *}
  49.851 +
  49.852 +lemma CauchyI:
  49.853 +  "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
  49.854 +by (simp add: Cauchy_def)
  49.855 +
  49.856 +lemma CauchyD:
  49.857 +  "\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
  49.858 +by (simp add: Cauchy_def)
  49.859 +
  49.860 +subsubsection {* Cauchy Sequences are Bounded *}
  49.861 +
  49.862 +text{*A Cauchy sequence is bounded -- this is the standard
  49.863 +  proof mechanization rather than the nonstandard proof*}
  49.864 +
  49.865 +lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
  49.866 +          ==>  \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
  49.867 +apply (clarify, drule spec, drule (1) mp)
  49.868 +apply (simp only: norm_minus_commute)
  49.869 +apply (drule order_le_less_trans [OF norm_triangle_ineq2])
  49.870 +apply simp
  49.871 +done
  49.872 +
  49.873 +lemma Cauchy_Bseq: "Cauchy X ==> Bseq X"
  49.874 +apply (simp add: Cauchy_def)
  49.875 +apply (drule spec, drule mp, rule zero_less_one, safe)
  49.876 +apply (drule_tac x="M" in spec, simp)
  49.877 +apply (drule lemmaCauchy)
  49.878 +apply (rule_tac k="M" in Bseq_offset)
  49.879 +apply (simp add: Bseq_def)
  49.880 +apply (rule_tac x="1 + norm (X M)" in exI)
  49.881 +apply (rule conjI, rule order_less_le_trans [OF zero_less_one], simp)
  49.882 +apply (simp add: order_less_imp_le)
  49.883 +done
  49.884 +
  49.885 +subsubsection {* Cauchy Sequences are Convergent *}
  49.886 +
  49.887 +axclass banach \<subseteq> real_normed_vector
  49.888 +  Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X"
  49.889 +
  49.890 +theorem LIMSEQ_imp_Cauchy:
  49.891 +  assumes X: "X ----> a" shows "Cauchy X"
  49.892 +proof (rule CauchyI)
  49.893 +  fix e::real assume "0 < e"
  49.894 +  hence "0 < e/2" by simp
  49.895 +  with X have "\<exists>N. \<forall>n\<ge>N. norm (X n - a) < e/2" by (rule LIMSEQ_D)
  49.896 +  then obtain N where N: "\<forall>n\<ge>N. norm (X n - a) < e/2" ..
  49.897 +  show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < e"
  49.898 +  proof (intro exI allI impI)
  49.899 +    fix m assume "N \<le> m"
  49.900 +    hence m: "norm (X m - a) < e/2" using N by fast
  49.901 +    fix n assume "N \<le> n"
  49.902 +    hence n: "norm (X n - a) < e/2" using N by fast
  49.903 +    have "norm (X m - X n) = norm ((X m - a) - (X n - a))" by simp
  49.904 +    also have "\<dots> \<le> norm (X m - a) + norm (X n - a)"
  49.905 +      by (rule norm_triangle_ineq4)
  49.906 +    also from m n have "\<dots> < e" by(simp add:field_simps)
  49.907 +    finally show "norm (X m - X n) < e" .
  49.908 +  qed
  49.909 +qed
  49.910 +
  49.911 +lemma convergent_Cauchy: "convergent X \<Longrightarrow> Cauchy X"
  49.912 +unfolding convergent_def
  49.913 +by (erule exE, erule LIMSEQ_imp_Cauchy)
  49.914 +
  49.915 +text {*
  49.916 +Proof that Cauchy sequences converge based on the one from
  49.917 +http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html
  49.918 +*}
  49.919 +
  49.920 +text {*
  49.921 +  If sequence @{term "X"} is Cauchy, then its limit is the lub of
  49.922 +  @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
  49.923 +*}
  49.924 +
  49.925 +lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u"
  49.926 +by (simp add: isUbI setleI)
  49.927 +
  49.928 +lemma real_abs_diff_less_iff:
  49.929 +  "(\<bar>x - a\<bar> < (r::real)) = (a - r < x \<and> x < a + r)"
  49.930 +by auto
  49.931 +
  49.932 +locale real_Cauchy =
  49.933 +  fixes X :: "nat \<Rightarrow> real"
  49.934 +  assumes X: "Cauchy X"
  49.935 +  fixes S :: "real set"
  49.936 +  defines S_def: "S \<equiv> {x::real. \<exists>N. \<forall>n\<ge>N. x < X n}"
  49.937 +
  49.938 +lemma real_CauchyI:
  49.939 +  assumes "Cauchy X"
  49.940 +  shows "real_Cauchy X"
  49.941 +  proof qed (fact assms)
  49.942 +
  49.943 +lemma (in real_Cauchy) mem_S: "\<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S"
  49.944 +by (unfold S_def, auto)
  49.945 +
  49.946 +lemma (in real_Cauchy) bound_isUb:
  49.947 +  assumes N: "\<forall>n\<ge>N. X n < x"
  49.948 +  shows "isUb UNIV S x"
  49.949 +proof (rule isUb_UNIV_I)
  49.950 +  fix y::real assume "y \<in> S"
  49.951 +  hence "\<exists>M. \<forall>n\<ge>M. y < X n"
  49.952 +    by (simp add: S_def)
  49.953 +  then obtain M where "\<forall>n\<ge>M. y < X n" ..
  49.954 +  hence "y < X (max M N)" by simp
  49.955 +  also have "\<dots> < x" using N by simp
  49.956 +  finally show "y \<le> x"
  49.957 +    by (rule order_less_imp_le)
  49.958 +qed
  49.959 +
  49.960 +lemma (in real_Cauchy) isLub_ex: "\<exists>u. isLub UNIV S u"
  49.961 +proof (rule reals_complete)
  49.962 +  obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < 1"
  49.963 +    using CauchyD [OF X zero_less_one] by fast
  49.964 +  hence N: "\<forall>n\<ge>N. norm (X n - X N) < 1" by simp
  49.965 +  show "\<exists>x. x \<in> S"
  49.966 +  proof
  49.967 +    from N have "\<forall>n\<ge>N. X N - 1 < X n"
  49.968 +      by (simp add: real_abs_diff_less_iff)
  49.969 +    thus "X N - 1 \<in> S" by (rule mem_S)
  49.970 +  qed
  49.971 +  show "\<exists>u. isUb UNIV S u"
  49.972 +  proof
  49.973 +    from N have "\<forall>n\<ge>N. X n < X N + 1"
  49.974 +      by (simp add: real_abs_diff_less_iff)
  49.975 +    thus "isUb UNIV S (X N + 1)"
  49.976 +      by (rule bound_isUb)
  49.977 +  qed
  49.978 +qed
  49.979 +
  49.980 +lemma (in real_Cauchy) isLub_imp_LIMSEQ:
  49.981 +  assumes x: "isLub UNIV S x"
  49.982 +  shows "X ----> x"
  49.983 +proof (rule LIMSEQ_I)
  49.984 +  fix r::real assume "0 < r"
  49.985 +  hence r: "0 < r/2" by simp
  49.986 +  obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. norm (X n - X m) < r/2"
  49.987 +    using CauchyD [OF X r] by fast
  49.988 +  hence "\<forall>n\<ge>N. norm (X n - X N) < r/2" by simp
  49.989 +  hence N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2"
  49.990 +    by (simp only: real_norm_def real_abs_diff_less_iff)
  49.991 +
  49.992 +  from N have "\<forall>n\<ge>N. X N - r/2 < X n" by fast
  49.993 +  hence "X N - r/2 \<in> S" by (rule mem_S)
  49.994 +  hence 1: "X N - r/2 \<le> x" using x isLub_isUb isUbD by fast
  49.995 +
  49.996 +  from N have "\<forall>n\<ge>N. X n < X N + r/2" by fast
  49.997 +  hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb)
  49.998 +  hence 2: "x \<le> X N + r/2" using x isLub_le_isUb by fast
  49.999 +
 49.1000 +  show "\<exists>N. \<forall>n\<ge>N. norm (X n - x) < r"
 49.1001 +  proof (intro exI allI impI)
 49.1002 +    fix n assume n: "N \<le> n"
 49.1003 +    from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
 49.1004 +    thus "norm (X n - x) < r" using 1 2
 49.1005 +      by (simp add: real_abs_diff_less_iff)
 49.1006 +  qed
 49.1007 +qed
 49.1008 +
 49.1009 +lemma (in real_Cauchy) LIMSEQ_ex: "\<exists>x. X ----> x"
 49.1010 +proof -
 49.1011 +  obtain x where "isLub UNIV S x"
 49.1012 +    using isLub_ex by fast
 49.1013 +  hence "X ----> x"
 49.1014 +    by (rule isLub_imp_LIMSEQ)
 49.1015 +  thus ?thesis ..
 49.1016 +qed
 49.1017 +
 49.1018 +lemma real_Cauchy_convergent:
 49.1019 +  fixes X :: "nat \<Rightarrow> real"
 49.1020 +  shows "Cauchy X \<Longrightarrow> convergent X"
 49.1021 +unfolding convergent_def
 49.1022 +by (rule real_Cauchy.LIMSEQ_ex)
 49.1023 + (rule real_CauchyI)
 49.1024 +
 49.1025 +instance real :: banach
 49.1026 +by intro_classes (rule real_Cauchy_convergent)
 49.1027 +
 49.1028 +lemma Cauchy_convergent_iff:
 49.1029 +  fixes X :: "nat \<Rightarrow> 'a::banach"
 49.1030 +  shows "Cauchy X = convergent X"
 49.1031 +by (fast intro: Cauchy_convergent convergent_Cauchy)
 49.1032 +
 49.1033 +
 49.1034 +subsection {* Power Sequences *}
 49.1035 +
 49.1036 +text{*The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
 49.1037 +"x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
 49.1038 +  also fact that bounded and monotonic sequence converges.*}
 49.1039 +
 49.1040 +lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
 49.1041 +apply (simp add: Bseq_def)
 49.1042 +apply (rule_tac x = 1 in exI)
 49.1043 +apply (simp add: power_abs)
 49.1044 +apply (auto dest: power_mono)
 49.1045 +done
 49.1046 +
 49.1047 +lemma monoseq_realpow: "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
 49.1048 +apply (clarify intro!: mono_SucI2)
 49.1049 +apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
 49.1050 +done
 49.1051 +
 49.1052 +lemma convergent_realpow:
 49.1053 +  "[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)"
 49.1054 +by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
 49.1055 +
 49.1056 +lemma LIMSEQ_inverse_realpow_zero_lemma:
 49.1057 +  fixes x :: real
 49.1058 +  assumes x: "0 \<le> x"
 49.1059 +  shows "real n * x + 1 \<le> (x + 1) ^ n"
 49.1060 +apply (induct n)
 49.1061 +apply simp
 49.1062 +apply simp
 49.1063 +apply (rule order_trans)
 49.1064 +prefer 2
 49.1065 +apply (erule mult_left_mono)
 49.1066 +apply (rule add_increasing [OF x], simp)
 49.1067 +apply (simp add: real_of_nat_Suc)
 49.1068 +apply (simp add: ring_distribs)
 49.1069 +apply (simp add: mult_nonneg_nonneg x)
 49.1070 +done
 49.1071 +
 49.1072 +lemma LIMSEQ_inverse_realpow_zero:
 49.1073 +  "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) ----> 0"
 49.1074 +proof (rule LIMSEQ_inverse_zero [rule_format])
 49.1075 +  fix y :: real
 49.1076 +  assume x: "1 < x"
 49.1077 +  hence "0 < x - 1" by simp
 49.1078 +  hence "\<forall>y. \<exists>N::nat. y < real N * (x - 1)"
 49.1079 +    by (rule reals_Archimedean3)
 49.1080 +  hence "\<exists>N::nat. y < real N * (x - 1)" ..
 49.1081 +  then obtain N::nat where "y < real N * (x - 1)" ..
 49.1082 +  also have "\<dots> \<le> real N * (x - 1) + 1" by simp
 49.1083 +  also have "\<dots> \<le> (x - 1 + 1) ^ N"
 49.1084 +    by (rule LIMSEQ_inverse_realpow_zero_lemma, cut_tac x, simp)
 49.1085 +  also have "\<dots> = x ^ N" by simp
 49.1086 +  finally have "y < x ^ N" .
 49.1087 +  hence "\<forall>n\<ge>N. y < x ^ n"
 49.1088 +    apply clarify
 49.1089 +    apply (erule order_less_le_trans)
 49.1090 +    apply (erule power_increasing)
 49.1091 +    apply (rule order_less_imp_le [OF x])
 49.1092 +    done
 49.1093 +  thus "\<exists>N. \<forall>n\<ge>N. y < x ^ n" ..
 49.1094 +qed
 49.1095 +
 49.1096 +lemma LIMSEQ_realpow_zero:
 49.1097 +  "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
 49.1098 +proof (cases)
 49.1099 +  assume "x = 0"
 49.1100 +  hence "(\<lambda>n. x ^ Suc n) ----> 0" by (simp add: LIMSEQ_const)
 49.1101 +  thus ?thesis by (rule LIMSEQ_imp_Suc)
 49.1102 +next
 49.1103 +  assume "0 \<le> x" and "x \<noteq> 0"
 49.1104 +  hence x0: "0 < x" by simp
 49.1105 +  assume x1: "x < 1"
 49.1106 +  from x0 x1 have "1 < inverse x"
 49.1107 +    by (rule real_inverse_gt_one)
 49.1108 +  hence "(\<lambda>n. inverse (inverse x ^ n)) ----> 0"
 49.1109 +    by (rule LIMSEQ_inverse_realpow_zero)
 49.1110 +  thus ?thesis by (simp add: power_inverse)
 49.1111 +qed
 49.1112 +
 49.1113 +lemma LIMSEQ_power_zero:
 49.1114 +  fixes x :: "'a::{real_normed_algebra_1,recpower}"
 49.1115 +  shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
 49.1116 +apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
 49.1117 +apply (simp only: LIMSEQ_Zseq_iff, erule Zseq_le)
 49.1118 +apply (simp add: power_abs norm_power_ineq)
 49.1119 +done
 49.1120 +
 49.1121 +lemma LIMSEQ_divide_realpow_zero:
 49.1122 +  "1 < (x::real) ==> (%n. a / (x ^ n)) ----> 0"
 49.1123 +apply (cut_tac a = a and x1 = "inverse x" in
 49.1124 +        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_realpow_zero])
 49.1125 +apply (auto simp add: divide_inverse power_inverse)
 49.1126 +apply (simp add: inverse_eq_divide pos_divide_less_eq)
 49.1127 +done
 49.1128 +
 49.1129 +text{*Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}*}
 49.1130 +
 49.1131 +lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < (1::real) ==> (%n. \<bar>c\<bar> ^ n) ----> 0"
 49.1132 +by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
 49.1133 +
 49.1134 +lemma LIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < (1::real) ==> (%n. c ^ n) ----> 0"
 49.1135 +apply (rule LIMSEQ_rabs_zero [THEN iffD1])
 49.1136 +apply (auto intro: LIMSEQ_rabs_realpow_zero simp add: power_abs)
 49.1137 +done
 49.1138 +
 49.1139 +end
    50.1 --- a/src/HOL/Series.thy	Mon Dec 29 13:23:53 2008 +0100
    50.2 +++ b/src/HOL/Series.thy	Mon Dec 29 14:08:08 2008 +0100
    50.3 @@ -10,7 +10,7 @@
    50.4  header{*Finite Summation and Infinite Series*}
    50.5  
    50.6  theory Series
    50.7 -imports "~~/src/HOL/Hyperreal/SEQ"
    50.8 +imports SEQ
    50.9  begin
   50.10  
   50.11  definition