adapted HOL source structure to distribution layout
authorhaftmann
Mon, 29 Dec 2008 14:08:08 +0100
changeset 29197 6d4cb27ed19c
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
--- a/NEWS	Mon Dec 29 13:23:53 2008 +0100
+++ b/NEWS	Mon Dec 29 14:08:08 2008 +0100
@@ -157,11 +157,12 @@
 
 *** HOL ***
 
-* Made repository layout more coherent with logical
-distribution structure:
+* Made source layout more coherent with logical distribution
+structure:
 
     src/HOL/Library/RType.thy ~> src/HOL/Typerep.thy
     src/HOL/Library/Code_Message.thy ~> src/HOL/
+    src/HOL/Library/Dense_Linear_Order.thy ~> src/HOL/
     src/HOL/Library/GCD.thy ~> src/HOL/
     src/HOL/Library/Order_Relation.thy ~> src/HOL/
     src/HOL/Library/Parity.thy ~> src/HOL/
@@ -177,6 +178,7 @@
     src/HOL/Complex/Complex_Main.thy ~> src/HOL/
     src/HOL/Complex/Complex.thy ~> src/HOL/
     src/HOL/Complex/FrechetDeriv.thy ~> src/HOL/
+    src/HOL/Complex/Fundamental_Theorem_Algebra.thy ~> src/HOL/
     src/HOL/Hyperreal/Deriv.thy ~> src/HOL/
     src/HOL/Hyperreal/Fact.thy ~> src/HOL/
     src/HOL/Hyperreal/Integration.thy ~> src/HOL/
@@ -186,9 +188,12 @@
     src/HOL/Hyperreal/MacLaurin.thy ~> src/HOL/
     src/HOL/Hyperreal/NthRoot.thy ~> src/HOL/
     src/HOL/Hyperreal/Series.thy ~> src/HOL/
+    src/HOL/Hyperreal/SEQ.thy ~> src/HOL/
     src/HOL/Hyperreal/Taylor.thy ~> src/HOL/
     src/HOL/Hyperreal/Transcendental.thy ~> src/HOL/
     src/HOL/Real/Float ~> src/HOL/Library/
+    src/HOL/Real/HahnBanach ~> src/HOL/HahnBanach
+    src/HOL/Real/RealVector.thy ~> src/HOL/
 
     src/HOL/arith_data.ML ~> src/HOL/Tools
     src/HOL/hologic.ML ~> src/HOL/Tools
--- a/src/HOL/Complex/Fundamental_Theorem_Algebra.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1329 +0,0 @@
-(*  Title:       Fundamental_Theorem_Algebra.thy
-    Author:      Amine Chaieb
-*)
-
-header{*Fundamental Theorem of Algebra*}
-
-theory Fundamental_Theorem_Algebra
-imports "~~/src/HOL/Univ_Poly" "~~/src/HOL/Library/Dense_Linear_Order" "~~/src/HOL/Complex"
-begin
-
-subsection {* Square root of complex numbers *}
-definition csqrt :: "complex \<Rightarrow> complex" where
-"csqrt z = (if Im z = 0 then
-            if 0 \<le> Re z then Complex (sqrt(Re z)) 0
-            else Complex 0 (sqrt(- Re z))
-           else Complex (sqrt((cmod z + Re z) /2))
-                        ((Im z / abs(Im z)) * sqrt((cmod z - Re z) /2)))"
-
-lemma csqrt[algebra]: "csqrt z ^ 2 = z"
-proof-
-  obtain x y where xy: "z = Complex x y" by (cases z, simp_all)
-  {assume y0: "y = 0"
-    {assume x0: "x \<ge> 0" 
-      then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
-	by (simp add: csqrt_def power2_eq_square)}
-    moreover
-    {assume "\<not> x \<ge> 0" hence x0: "- x \<ge> 0" by arith
-      then have ?thesis using y0 xy real_sqrt_pow2[OF x0] 
-	by (simp add: csqrt_def power2_eq_square) }
-    ultimately have ?thesis by blast}
-  moreover
-  {assume y0: "y\<noteq>0"
-    {fix x y
-      let ?z = "Complex x y"
-      from abs_Re_le_cmod[of ?z] have tha: "abs x \<le> cmod ?z" by auto
-      hence "cmod ?z - x \<ge> 0" "cmod ?z + x \<ge> 0" by arith+ 
-      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) }
-    note th = this
-    have sq4: "\<And>x::real. x^2 / 4 = (x / 2) ^ 2" 
-      by (simp add: power2_eq_square) 
-    from th[of x y]
-    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
-    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"
-      unfolding power2_eq_square by simp 
-    have "sqrt 4 = sqrt (2^2)" by simp 
-    hence sqrt4: "sqrt 4 = 2" by (simp only: real_sqrt_abs)
-    have th2: "2 *(y * sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) + x) / 4)) / \<bar>y\<bar> = y"
-      using iffD2[OF real_sqrt_pow2_iff sum_power2_ge_zero[of x y]] y0
-      unfolding power2_eq_square 
-      by (simp add: ring_simps real_sqrt_divide sqrt4)
-     from y0 xy have ?thesis  apply (simp add: csqrt_def power2_eq_square)
-       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])
-      using th1 th2  ..}
-  ultimately show ?thesis by blast
-qed
-
-
-subsection{* More lemmas about module of complex numbers *}
-
-lemma complex_of_real_power: "complex_of_real x ^ n = complex_of_real (x^n)"
-  by (rule of_real_power [symmetric])
-
-lemma real_down2: "(0::real) < d1 \<Longrightarrow> 0 < d2 ==> EX e. 0 < e & e < d1 & e < d2"
-  apply ferrack apply arith done
-
-text{* The triangle inequality for cmod *}
-lemma complex_mod_triangle_sub: "cmod w \<le> cmod (w + z) + norm z"
-  using complex_mod_triangle_ineq2[of "w + z" "-z"] by auto
-
-subsection{* Basic lemmas about complex polynomials *}
-
-lemma poly_bound_exists:
-  shows "\<exists>m. m > 0 \<and> (\<forall>z. cmod z <= r \<longrightarrow> cmod (poly p z) \<le> m)"
-proof(induct p)
-  case Nil thus ?case by (rule exI[where x=1], simp) 
-next
-  case (Cons c cs)
-  from Cons.hyps obtain m where m: "\<forall>z. cmod z \<le> r \<longrightarrow> cmod (poly cs z) \<le> m"
-    by blast
-  let ?k = " 1 + cmod c + \<bar>r * m\<bar>"
-  have kp: "?k > 0" using abs_ge_zero[of "r*m"] norm_ge_zero[of c] by arith
-  {fix z
-    assume H: "cmod z \<le> r"
-    from m H have th: "cmod (poly cs z) \<le> m" by blast
-    from H have rp: "r \<ge> 0" using norm_ge_zero[of z] by arith
-    have "cmod (poly (c # cs) z) \<le> cmod c + cmod (z* poly cs z)"
-      using norm_triangle_ineq[of c "z* poly cs z"] by simp
-    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)
-    also have "\<dots> \<le> ?k" by simp
-    finally have "cmod (poly (c # cs) z) \<le> ?k" .}
-  with kp show ?case by blast
-qed
-
-
-text{* Offsetting the variable in a polynomial gives another of same degree *}
-  (* FIXME : Lemma holds also in locale --- fix it later *)
-lemma  poly_offset_lemma:
-  shows "\<exists>b q. (length q = length p) \<and> (\<forall>x. poly (b#q) (x::complex) = (a + x) * poly p x)"
-proof(induct p)
-  case Nil thus ?case by simp
-next
-  case (Cons c cs)
-  from Cons.hyps obtain b q where 
-    bq: "length q = length cs" "\<forall>x. poly (b # q) x = (a + x) * poly cs x"
-    by blast
-  let ?b = "a*c"
-  let ?q = "(b+c)#q"
-  have lg: "length ?q = length (c#cs)" using bq(1) by simp
-  {fix x
-    from bq(2)[rule_format, of x]
-    have "x*poly (b # q) x = x*((a + x) * poly cs x)" by simp
-    hence "poly (?b# ?q) x = (a + x) * poly (c # cs) x"
-      by (simp add: ring_simps)}
-  with lg  show ?case by blast 
-qed
-
-    (* FIXME : This one too*)
-lemma poly_offset: "\<exists> q. length q = length p \<and> (\<forall>x. poly q (x::complex) = poly p (a + x))"
-proof (induct p)
-  case Nil thus ?case by simp
-next
-  case (Cons c cs)
-  from Cons.hyps obtain q where q: "length q = length cs" "\<forall>x. poly q x = poly cs (a + x)" by blast
-  from poly_offset_lemma[of q a] obtain b p where 
-    bp: "length p = length q" "\<forall>x. poly (b # p) x = (a + x) * poly q x"
-    by blast
-  thus ?case using q bp by - (rule exI[where x="(c + b)#p"], simp)
-qed
-
-text{* An alternative useful formulation of completeness of the reals *}
-lemma real_sup_exists: assumes ex: "\<exists>x. P x" and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z"
-  shows "\<exists>(s::real). \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s"
-proof-
-  from ex bz obtain x Y where x: "P x" and Y: "\<And>x. P x \<Longrightarrow> x < Y"  by blast
-  from ex have thx:"\<exists>x. x \<in> Collect P" by blast
-  from bz have thY: "\<exists>Y. isUb UNIV (Collect P) Y" 
-    by(auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def order_le_less)
-  from reals_complete[OF thx thY] obtain L where L: "isLub UNIV (Collect P) L"
-    by blast
-  from Y[OF x] have xY: "x < Y" .
-  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)  
-  from Y have Y': "\<forall>x. P x \<longrightarrow> x \<le> Y" 
-    apply (clarsimp, atomize (full)) by auto 
-  from L Y' have "L \<le> Y" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
-  {fix y
-    {fix z assume z: "P z" "y < z"
-      from L' z have "y < L" by auto }
-    moreover
-    {assume yL: "y < L" "\<forall>z. P z \<longrightarrow> \<not> y < z"
-      hence nox: "\<forall>z. P z \<longrightarrow> y \<ge> z" by auto
-      from nox L have "y \<ge> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def) 
-      with yL(1) have False  by arith}
-    ultimately have "(\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < L" by blast}
-  thus ?thesis by blast
-qed
-
-
-subsection{* Some theorems about Sequences*}
-text{* Given a binary function @{text "f:: nat \<Rightarrow> 'a \<Rightarrow> 'a"}, its values are uniquely determined by a function g *}
-
-lemma num_Axiom: "EX! g. g 0 = e \<and> (\<forall>n. g (Suc n) = f n (g n))"
-  unfolding Ex1_def
-  apply (rule_tac x="nat_rec e f" in exI)
-  apply (rule conjI)+
-apply (rule def_nat_rec_0, simp)
-apply (rule allI, rule def_nat_rec_Suc, simp)
-apply (rule allI, rule impI, rule ext)
-apply (erule conjE)
-apply (induct_tac x)
-apply (simp add: nat_rec_0)
-apply (erule_tac x="n" in allE)
-apply (simp)
-done
-
- text{* An equivalent formulation of monotony -- Not used here, but might be useful *}
-lemma mono_Suc: "mono f = (\<forall>n. (f n :: 'a :: order) \<le> f (Suc n))"
-unfolding mono_def
-proof auto
-  fix A B :: nat
-  assume H: "\<forall>n. f n \<le> f (Suc n)" "A \<le> B"
-  hence "\<exists>k. B = A + k" apply -  apply (thin_tac "\<forall>n. f n \<le> f (Suc n)") 
-    by presburger
-  then obtain k where k: "B = A + k" by blast
-  {fix a k
-    have "f a \<le> f (a + k)"
-    proof (induct k)
-      case 0 thus ?case by simp
-    next
-      case (Suc k)
-      from Suc.hyps H(1)[rule_format, of "a + k"] show ?case by simp
-    qed}
-  with k show "f A \<le> f B" by blast
-qed
-
-text{* for any sequence, there is a mootonic subsequence *}
-lemma seq_monosub: "\<exists>f. subseq f \<and> monoseq (\<lambda> n. (s (f n)))"
-proof-
-  {assume H: "\<forall>n. \<exists>p >n. \<forall> m\<ge>p. s m \<le> s p"
-    let ?P = "\<lambda> p n. p > n \<and> (\<forall>m \<ge> p. s m \<le> s p)"
-    from num_Axiom[of "SOME p. ?P p 0" "\<lambda>p n. SOME p. ?P p n"]
-    obtain f where f: "f 0 = (SOME p. ?P p 0)" "\<forall>n. f (Suc n) = (SOME p. ?P p (f n))" by blast
-    have "?P (f 0) 0"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p 0"]
-      using H apply - 
-      apply (erule allE[where x=0], erule exE, rule_tac x="p" in exI) 
-      unfolding order_le_less by blast 
-    hence f0: "f 0 > 0" "\<forall>m \<ge> f 0. s m \<le> s (f 0)" by blast+
-    {fix n
-      have "?P (f (Suc n)) (f n)" 
-	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
-	using H apply - 
-      apply (erule allE[where x="f n"], erule exE, rule_tac x="p" in exI) 
-      unfolding order_le_less by blast 
-    hence "f (Suc n) > f n" "\<forall>m \<ge> f (Suc n). s m \<le> s (f (Suc n))" by blast+}
-  note fSuc = this
-    {fix p q assume pq: "p \<ge> f q"
-      have "s p \<le> s(f(q))"  using f0(2)[rule_format, of p] pq fSuc
-	by (cases q, simp_all) }
-    note pqth = this
-    {fix q
-      have "f (Suc q) > f q" apply (induct q) 
-	using f0(1) fSuc(1)[of 0] apply simp by (rule fSuc(1))}
-    note fss = this
-    from fss have th1: "subseq f" unfolding subseq_Suc_iff ..
-    {fix a b 
-      have "f a \<le> f (a + b)"
-      proof(induct b)
-	case 0 thus ?case by simp
-      next
-	case (Suc b)
-	from fSuc(1)[of "a + b"] Suc.hyps show ?case by simp
-      qed}
-    note fmon0 = this
-    have "monoseq (\<lambda>n. s (f n))" 
-    proof-
-      {fix n
-	have "s (f n) \<ge> s (f (Suc n))" 
-	proof(cases n)
-	  case 0
-	  assume n0: "n = 0"
-	  from fSuc(1)[of 0] have th0: "f 0 \<le> f (Suc 0)" by simp
-	  from f0(2)[rule_format, OF th0] show ?thesis  using n0 by simp
-	next
-	  case (Suc m)
-	  assume m: "n = Suc m"
-	  from fSuc(1)[of n] m have th0: "f (Suc m) \<le> f (Suc (Suc m))" by simp
-	  from m fSuc(2)[rule_format, OF th0] show ?thesis by simp 
-	qed}
-      thus "monoseq (\<lambda>n. s (f n))" unfolding monoseq_Suc by blast 
-    qed
-    with th1 have ?thesis by blast}
-  moreover
-  {fix N assume N: "\<forall>p >N. \<exists> m\<ge>p. s m > s p"
-    {fix p assume p: "p \<ge> Suc N" 
-      hence pN: "p > N" by arith with N obtain m where m: "m \<ge> p" "s m > s p" by blast
-      have "m \<noteq> p" using m(2) by auto 
-      with m have "\<exists>m>p. s p < s m" by - (rule exI[where x=m], auto)}
-    note th0 = this
-    let ?P = "\<lambda>m x. m > x \<and> s x < s m"
-    from num_Axiom[of "SOME x. ?P x (Suc N)" "\<lambda>m x. SOME y. ?P y x"]
-    obtain f where f: "f 0 = (SOME x. ?P x (Suc N))" 
-      "\<forall>n. f (Suc n) = (SOME m. ?P m (f n))" by blast
-    have "?P (f 0) (Suc N)"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p (Suc N)"]
-      using N apply - 
-      apply (erule allE[where x="Suc N"], clarsimp)
-      apply (rule_tac x="m" in exI)
-      apply auto
-      apply (subgoal_tac "Suc N \<noteq> m")
-      apply simp
-      apply (rule ccontr, simp)
-      done
-    hence f0: "f 0 > Suc N" "s (Suc N) < s (f 0)" by blast+
-    {fix n
-      have "f n > N \<and> ?P (f (Suc n)) (f n)"
-	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
-      proof (induct n)
-	case 0 thus ?case
-	  using f0 N apply auto 
-	  apply (erule allE[where x="f 0"], clarsimp) 
-	  apply (rule_tac x="m" in exI, simp)
-	  by (subgoal_tac "f 0 \<noteq> m", auto)
-      next
-	case (Suc n)
-	from Suc.hyps have Nfn: "N < f n" by blast
-	from Suc.hyps obtain m where m: "m > f n" "s (f n) < s m" by blast
-	with Nfn have mN: "m > N" by arith
-	note key = Suc.hyps[unfolded some_eq_ex[of "\<lambda>p. ?P p (f n)", symmetric] f(2)[rule_format, of n, symmetric]]
-	
-	from key have th0: "f (Suc n) > N" by simp
-	from N[rule_format, OF th0]
-	obtain m' where m': "m' \<ge> f (Suc n)" "s (f (Suc n)) < s m'" by blast
-	have "m' \<noteq> f (Suc (n))" apply (rule ccontr) using m'(2) by auto
-	hence "m' > f (Suc n)" using m'(1) by simp
-	with key m'(2) show ?case by auto
-      qed}
-    note fSuc = this
-    {fix n
-      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 
-      hence "f n \<ge> Suc N" "f(Suc n) > f n" "s(f n) < s(f(Suc n))" by blast+}
-    note thf = this
-    have sqf: "subseq f" unfolding subseq_Suc_iff using thf by simp
-    have "monoseq (\<lambda>n. s (f n))"  unfolding monoseq_Suc using thf
-      apply -
-      apply (rule disjI1)
-      apply auto
-      apply (rule order_less_imp_le)
-      apply blast
-      done
-    then have ?thesis  using sqf by blast}
-  ultimately show ?thesis unfolding linorder_not_less[symmetric] by blast
-qed
-
-lemma seq_suble: assumes sf: "subseq f" shows "n \<le> f n"
-proof(induct n)
-  case 0 thus ?case by simp
-next
-  case (Suc n)
-  from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps
-  have "n < f (Suc n)" by arith 
-  thus ?case by arith
-qed
-
-subsection {* Fundamental theorem of algebra *}
-lemma  unimodular_reduce_norm:
-  assumes md: "cmod z = 1"
-  shows "cmod (z + 1) < 1 \<or> cmod (z - 1) < 1 \<or> cmod (z + ii) < 1 \<or> cmod (z - ii) < 1"
-proof-
-  obtain x y where z: "z = Complex x y " by (cases z, auto)
-  from md z have xy: "x^2 + y^2 = 1" by (simp add: cmod_def)
-  {assume C: "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + ii) \<ge> 1" "cmod (z - ii) \<ge> 1"
-    from C z xy have "2*x \<le> 1" "2*x \<ge> -1" "2*y \<le> 1" "2*y \<ge> -1"
-      by (simp_all add: cmod_def power2_eq_square ring_simps)
-    hence "abs (2*x) \<le> 1" "abs (2*y) \<le> 1" by simp_all
-    hence "(abs (2 * x))^2 <= 1^2" "(abs (2 * y)) ^2 <= 1^2"
-      by - (rule power_mono, simp, simp)+
-    hence th0: "4*x^2 \<le> 1" "4*y^2 \<le> 1" 
-      by (simp_all  add: power2_abs power_mult_distrib)
-    from add_mono[OF th0] xy have False by simp }
-  thus ?thesis unfolding linorder_not_le[symmetric] by blast
-qed
-
-text{* Hence we can always reduce modulus of @{text "1 + b z^n"} if nonzero *}
-lemma reduce_poly_simple:
- assumes b: "b \<noteq> 0" and n: "n\<noteq>0"
-  shows "\<exists>z. cmod (1 + b * z^n) < 1"
-using n
-proof(induct n rule: nat_less_induct)
-  fix n
-  assume IH: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>z. cmod (1 + b * z ^ m) < 1)" and n: "n \<noteq> 0"
-  let ?P = "\<lambda>z n. cmod (1 + b * z ^ n) < 1"
-  {assume e: "even n"
-    hence "\<exists>m. n = 2*m" by presburger
-    then obtain m where m: "n = 2*m" by blast
-    from n m have "m\<noteq>0" "m < n" by presburger+
-    with IH[rule_format, of m] obtain z where z: "?P z m" by blast
-    from z have "?P (csqrt z) n" by (simp add: m power_mult csqrt)
-    hence "\<exists>z. ?P z n" ..}
-  moreover
-  {assume o: "odd n"
-    from b have b': "b^2 \<noteq> 0" unfolding power2_eq_square by simp
-    have "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
-    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) = 
-    ((Re (inverse b))^2 + (Im (inverse b))^2) * \<bar>Im b * Im b + Re b * Re b\<bar>" by algebra
-    also have "\<dots> = cmod (inverse b) ^2 * cmod b ^ 2" 
-      apply (simp add: cmod_def) using realpow_two_le_add_order[of "Re b" "Im b"]
-      by (simp add: power2_eq_square)
-    finally 
-    have th0: "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
-    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) =
-    1" 
-      apply (simp add: power2_eq_square norm_mult[symmetric] norm_inverse[symmetric])
-      using right_inverse[OF b']
-      by (simp add: power2_eq_square[symmetric] power_inverse[symmetric] ring_simps)
-    have th0: "cmod (complex_of_real (cmod b) / b) = 1"
-      apply (simp add: complex_Re_mult cmod_def power2_eq_square Re_complex_of_real Im_complex_of_real divide_inverse ring_simps )
-      by (simp add: real_sqrt_mult[symmetric] th0)        
-    from o have "\<exists>m. n = Suc (2*m)" by presburger+
-    then obtain m where m: "n = Suc (2*m)" by blast
-    from unimodular_reduce_norm[OF th0] o
-    have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
-      apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1", rule_tac x="1" in exI, simp)
-      apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp add: diff_def)
-      apply (cases "cmod (complex_of_real (cmod b) / b + ii) < 1")
-      apply (cases "even m", rule_tac x="ii" in exI, simp add: m power_mult)
-      apply (rule_tac x="- ii" in exI, simp add: m power_mult)
-      apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult diff_def)
-      apply (rule_tac x="ii" in exI, simp add: m power_mult diff_def)
-      done
-    then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" by blast
-    let ?w = "v / complex_of_real (root n (cmod b))"
-    from odd_real_root_pow[OF o, of "cmod b"]
-    have th1: "?w ^ n = v^n / complex_of_real (cmod b)" 
-      by (simp add: power_divide complex_of_real_power)
-    have th2:"cmod (complex_of_real (cmod b) / b) = 1" using b by (simp add: norm_divide)
-    hence th3: "cmod (complex_of_real (cmod b) / b) \<ge> 0" by simp
-    have th4: "cmod (complex_of_real (cmod b) / b) *
-   cmod (1 + b * (v ^ n / complex_of_real (cmod b)))
-   < cmod (complex_of_real (cmod b) / b) * 1"
-      apply (simp only: norm_mult[symmetric] right_distrib)
-      using b v by (simp add: th2)
-
-    from mult_less_imp_less_left[OF th4 th3]
-    have "?P ?w n" unfolding th1 . 
-    hence "\<exists>z. ?P z n" .. }
-  ultimately show "\<exists>z. ?P z n" by blast
-qed
-
-
-text{* Bolzano-Weierstrass type property for closed disc in complex plane. *}
-
-lemma metric_bound_lemma: "cmod (x - y) <= \<bar>Re x - Re y\<bar> + \<bar>Im x - Im y\<bar>"
-  using real_sqrt_sum_squares_triangle_ineq[of "Re x - Re y" 0 0 "Im x - Im y" ]
-  unfolding cmod_def by simp
-
-lemma bolzano_weierstrass_complex_disc:
-  assumes r: "\<forall>n. cmod (s n) \<le> r"
-  shows "\<exists>f z. subseq f \<and> (\<forall>e >0. \<exists>N. \<forall>n \<ge> N. cmod (s (f n) - z) < e)"
-proof-
-  from seq_monosub[of "Re o s"] 
-  obtain f g where f: "subseq f" "monoseq (\<lambda>n. Re (s (f n)))" 
-    unfolding o_def by blast
-  from seq_monosub[of "Im o s o f"] 
-  obtain g where g: "subseq g" "monoseq (\<lambda>n. Im (s(f(g n))))" unfolding o_def by blast  
-  let ?h = "f o g"
-  from r[rule_format, of 0] have rp: "r \<ge> 0" using norm_ge_zero[of "s 0"] by arith 
-  have th:"\<forall>n. r + 1 \<ge> \<bar> Re (s n)\<bar>" 
-  proof
-    fix n
-    from abs_Re_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Re (s n)\<bar> \<le> r + 1" by arith
-  qed
-  have conv1: "convergent (\<lambda>n. Re (s ( f n)))"
-    apply (rule Bseq_monoseq_convergent)
-    apply (simp add: Bseq_def)
-    apply (rule exI[where x= "r + 1"])
-    using th rp apply simp
-    using f(2) .
-  have th:"\<forall>n. r + 1 \<ge> \<bar> Im (s n)\<bar>" 
-  proof
-    fix n
-    from abs_Im_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Im (s n)\<bar> \<le> r + 1" by arith
-  qed
-
-  have conv2: "convergent (\<lambda>n. Im (s (f (g n))))"
-    apply (rule Bseq_monoseq_convergent)
-    apply (simp add: Bseq_def)
-    apply (rule exI[where x= "r + 1"])
-    using th rp apply simp
-    using g(2) .
-
-  from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x" 
-    by blast 
-  hence  x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Re (s (f n)) - x \<bar> < r" 
-    unfolding LIMSEQ_def real_norm_def .
-
-  from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y" 
-    by blast 
-  hence  y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Im (s (f (g n))) - y \<bar> < r" 
-    unfolding LIMSEQ_def real_norm_def .
-  let ?w = "Complex x y"
-  from f(1) g(1) have hs: "subseq ?h" unfolding subseq_def by auto 
-  {fix e assume ep: "e > (0::real)"
-    hence e2: "e/2 > 0" by simp
-    from x[rule_format, OF e2] y[rule_format, OF e2]
-    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
-    {fix n assume nN12: "n \<ge> N1 + N2"
-      hence nN1: "g n \<ge> N1" and nN2: "n \<ge> N2" using seq_suble[OF g(1), of n] by arith+
-      from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
-      have "cmod (s (?h n) - ?w) < e" 
-	using metric_bound_lemma[of "s (f (g n))" ?w] by simp }
-    hence "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" by blast }
-  with hs show ?thesis  by blast  
-qed
-
-text{* Polynomial is continuous. *}
-
-lemma poly_cont:
-  assumes ep: "e > 0" 
-  shows "\<exists>d >0. \<forall>w. 0 < cmod (w - z) \<and> cmod (w - z) < d \<longrightarrow> cmod (poly p w - poly p z) < e"
-proof-
-  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
-  {fix w
-    note q(2)[of "w - z", simplified]}
-  note th = this
-  show ?thesis unfolding th[symmetric]
-  proof(induct q)
-    case Nil thus ?case  using ep by auto
-  next
-    case (Cons c cs)
-    from poly_bound_exists[of 1 "cs"] 
-    obtain m where m: "m > 0" "\<And>z. cmod z \<le> 1 \<Longrightarrow> cmod (poly cs z) \<le> m" by blast
-    from ep m(1) have em0: "e/m > 0" by (simp add: field_simps)
-    have one0: "1 > (0::real)"  by arith
-    from real_lbound_gt_zero[OF one0 em0] 
-    obtain d where d: "d >0" "d < 1" "d < e / m" by blast
-    from d(1,3) m(1) have dm: "d*m > 0" "d*m < e" 
-      by (simp_all add: field_simps real_mult_order)
-    show ?case 
-      proof(rule ex_forward[OF real_lbound_gt_zero[OF one0 em0]], clarsimp simp add: norm_mult)
-	fix d w
-	assume H: "d > 0" "d < 1" "d < e/m" "w\<noteq>z" "cmod (w-z) < d"
-	hence d1: "cmod (w-z) \<le> 1" "d \<ge> 0" by simp_all
-	from H(3) m(1) have dme: "d*m < e" by (simp add: field_simps)
-	from H have th: "cmod (w-z) \<le> d" by simp 
-	from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
-	show "cmod (w - z) * cmod (poly cs (w - z)) < e" by simp
-      qed  
-    qed
-qed
-
-text{* Hence a polynomial attains minimum on a closed disc 
-  in the complex plane. *}
-lemma  poly_minimum_modulus_disc:
-  "\<exists>z. \<forall>w. cmod w \<le> r \<longrightarrow> cmod (poly p z) \<le> cmod (poly p w)"
-proof-
-  {assume "\<not> r \<ge> 0" hence ?thesis unfolding linorder_not_le
-      apply -
-      apply (rule exI[where x=0]) 
-      apply auto
-      apply (subgoal_tac "cmod w < 0")
-      apply simp
-      apply arith
-      done }
-  moreover
-  {assume rp: "r \<ge> 0"
-    from rp have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))" by simp 
-    hence mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x"  by blast
-    {fix x z
-      assume H: "cmod z \<le> r" "cmod (poly p z) = - x" "\<not>x < 1"
-      hence "- x < 0 " by arith
-      with H(2) norm_ge_zero[of "poly p z"]  have False by simp }
-    then have mth2: "\<exists>z. \<forall>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<longrightarrow> x < z" by blast
-    from real_sup_exists[OF mth1 mth2] obtain s where 
-      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
-    let ?m = "-s"
-    {fix y
-      from s[rule_format, of "-y"] have 
-    "(\<exists>z x. cmod z \<le> r \<and> -(- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y" 
-	unfolding minus_less_iff[of y ] equation_minus_iff by blast }
-    note s1 = this[unfolded minus_minus]
-    from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m" 
-      by auto
-    {fix n::nat
-      from s1[rule_format, of "?m + 1/real (Suc n)"] 
-      have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)"
-	by simp}
-    hence th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
-    from choice[OF th] obtain g where 
-      g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m+1 /real(Suc n)" 
-      by blast
-    from bolzano_weierstrass_complex_disc[OF g(1)] 
-    obtain f z where fz: "subseq f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
-      by blast    
-    {fix w 
-      assume wr: "cmod w \<le> r"
-      let ?e = "\<bar>cmod (poly p z) - ?m\<bar>"
-      {assume e: "?e > 0"
-	hence e2: "?e/2 > 0" by simp
-	from poly_cont[OF e2, of z p] obtain d where
-	  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
-	{fix w assume w: "cmod (w - z) < d"
-	  have "cmod(poly p w - poly p z) < ?e / 2"
-	    using d(2)[rule_format, of w] w e by (cases "w=z", simp_all)}
-	note th1 = this
-	
-	from fz(2)[rule_format, OF d(1)] obtain N1 where 
-	  N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d" by blast
-	from reals_Archimedean2[of "2/?e"] obtain N2::nat where
-	  N2: "2/?e < real N2" by blast
-	have th2: "cmod(poly p (g(f(N1 + N2))) - poly p z) < ?e/2"
-	  using N1[rule_format, of "N1 + N2"] th1 by simp
-	{fix a b e2 m :: real
-	have "a < e2 \<Longrightarrow> abs(b - m) < e2 \<Longrightarrow> 2 * e2 <= abs(b - m) + a
-          ==> False" by arith}
-      note th0 = this
-      have ath: 
-	"\<And>m x e. m <= x \<Longrightarrow>  x < m + e ==> abs(x - m::real) < e" by arith
-      from s1m[OF g(1)[rule_format]]
-      have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
-      from seq_suble[OF fz(1), of "N1+N2"]
-      have th00: "real (Suc (N1+N2)) \<le> real (Suc (f (N1+N2)))" by simp
-      have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1+N2)) > 0"  
-	using N2 by auto
-      from frac_le[OF th000 th00] have th00: "?m +1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))" by simp
-      from g(2)[rule_format, of "f (N1 + N2)"]
-      have th01:"cmod (poly p (g (f (N1 + N2)))) < - s + 1 / real (Suc (f (N1 + N2)))" .
-      from order_less_le_trans[OF th01 th00]
-      have th32: "cmod(poly p (g (f (N1 + N2)))) < ?m + (1/ real(Suc (N1 + N2)))" .
-      from N2 have "2/?e < real (Suc (N1 + N2))" by arith
-      with e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
-      have "?e/2 > 1/ real (Suc (N1 + N2))" by (simp add: inverse_eq_divide)
-      with ath[OF th31 th32]
-      have thc1:"\<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar>< ?e/2" by arith  
-      have ath2: "\<And>(a::real) b c m. \<bar>a - b\<bar> <= c ==> \<bar>b - m\<bar> <= \<bar>a - m\<bar> + c" 
-	by arith
-      have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar>
-\<le> cmod (poly p (g (f (N1 + N2))) - poly p z)" 
-	by (simp add: norm_triangle_ineq3)
-      from ath2[OF th22, of ?m]
-      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
-      from th0[OF th2 thc1 thc2] have False .}
-      hence "?e = 0" by auto
-      then have "cmod (poly p z) = ?m" by simp  
-      with s1m[OF wr]
-      have "cmod (poly p z) \<le> cmod (poly p w)" by simp }
-    hence ?thesis by blast}
-  ultimately show ?thesis by blast
-qed
-
-lemma "(rcis (sqrt (abs r)) (a/2)) ^ 2 = rcis (abs r) a"
-  unfolding power2_eq_square
-  apply (simp add: rcis_mult)
-  apply (simp add: power2_eq_square[symmetric])
-  done
-
-lemma cispi: "cis pi = -1" 
-  unfolding cis_def
-  by simp
-
-lemma "(rcis (sqrt (abs r)) ((pi + a)/2)) ^ 2 = rcis (- abs r) a"
-  unfolding power2_eq_square
-  apply (simp add: rcis_mult add_divide_distrib)
-  apply (simp add: power2_eq_square[symmetric] rcis_def cispi cis_mult[symmetric])
-  done
-
-text {* Nonzero polynomial in z goes to infinity as z does. *}
-
-instance complex::idom_char_0 by (intro_classes)
-instance complex :: recpower_idom_char_0 by intro_classes
-
-lemma poly_infinity:
-  assumes ex: "list_ex (\<lambda>c. c \<noteq> 0) p"
-  shows "\<exists>r. \<forall>z. r \<le> cmod z \<longrightarrow> d \<le> cmod (poly (a#p) z)"
-using ex
-proof(induct p arbitrary: a d)
-  case (Cons c cs a d) 
-  {assume H: "list_ex (\<lambda>c. c\<noteq>0) cs"
-    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
-    let ?r = "1 + \<bar>r\<bar>"
-    {fix z assume h: "1 + \<bar>r\<bar> \<le> cmod z"
-      have r0: "r \<le> cmod z" using h by arith
-      from r[rule_format, OF r0]
-      have th0: "d + cmod a \<le> 1 * cmod(poly (c#cs) z)" by arith
-      from h have z1: "cmod z \<ge> 1" by arith
-      from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (c#cs) z"]]]
-      have th1: "d \<le> cmod(z * poly (c#cs) z) - cmod a"
-	unfolding norm_mult by (simp add: ring_simps)
-      from complex_mod_triangle_sub[of "z * poly (c#cs) z" a]
-      have th2: "cmod(z * poly (c#cs) z) - cmod a \<le> cmod (poly (a#c#cs) z)" 
-	by (simp add: diff_le_eq ring_simps) 
-      from th1 th2 have "d \<le> cmod (poly (a#c#cs) z)"  by arith}
-    hence ?case by blast}
-  moreover
-  {assume cs0: "\<not> (list_ex (\<lambda>c. c \<noteq> 0) cs)"
-    with Cons.prems have c0: "c \<noteq> 0" by simp
-    from cs0 have cs0': "list_all (\<lambda>c. c = 0) cs" 
-      by (auto simp add: list_all_iff list_ex_iff)
-    {fix z
-      assume h: "(\<bar>d\<bar> + cmod a) / cmod c \<le> cmod z"
-      from c0 have "cmod c > 0" by simp
-      from h c0 have th0: "\<bar>d\<bar> + cmod a \<le> cmod (z*c)" 
-	by (simp add: field_simps norm_mult)
-      have ath: "\<And>mzh mazh ma. mzh <= mazh + ma ==> abs(d) + ma <= mzh ==> d <= mazh" by arith
-      from complex_mod_triangle_sub[of "z*c" a ]
-      have th1: "cmod (z * c) \<le> cmod (a + z * c) + cmod a"
-	by (simp add: ring_simps)
-      from ath[OF th1 th0] have "d \<le> cmod (poly (a # c # cs) z)" 
-	using poly_0[OF cs0'] by simp}
-    then have ?case  by blast}
-  ultimately show ?case by blast
-qed simp
-
-text {* Hence polynomial's modulus attains its minimum somewhere. *}
-lemma poly_minimum_modulus:
-  "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)"
-proof(induct p)
-  case (Cons c cs) 
-  {assume cs0: "list_ex (\<lambda>c. c \<noteq> 0) cs"
-    from poly_infinity[OF cs0, of "cmod (poly (c#cs) 0)" c]
-    obtain r where r: "\<And>z. r \<le> cmod z \<Longrightarrow> cmod (poly (c # cs) 0) \<le> cmod (poly (c # cs) z)" by blast
-    have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>" by arith
-    from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "c#cs"] 
-    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
-    {fix z assume z: "r \<le> cmod z"
-      from v[of 0] r[OF z] 
-      have "cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) z)"
-	by simp }
-    note v0 = this
-    from v0 v ath[of r] have ?case by blast}
-  moreover
-  {assume cs0: "\<not> (list_ex (\<lambda>c. c\<noteq>0) cs)"
-    hence th:"list_all (\<lambda>c. c = 0) cs" by (simp add: list_all_iff list_ex_iff)
-    from poly_0[OF th] Cons.hyps have ?case by simp}
-  ultimately show ?case by blast
-qed simp
-
-text{* Constant function (non-syntactic characterization). *}
-definition "constant f = (\<forall>x y. f x = f y)"
-
-lemma nonconstant_length: "\<not> (constant (poly p)) \<Longrightarrow> length p \<ge> 2"
-  unfolding constant_def
-  apply (induct p, auto)
-  apply (unfold not_less[symmetric])
-  apply simp
-  apply (rule ccontr)
-  apply auto
-  done
- 
-lemma poly_replicate_append:
-  "poly ((replicate n 0)@p) (x::'a::{recpower, comm_ring}) = x^n * poly p x"
-  by(induct n, auto simp add: power_Suc ring_simps)
-
-text {* Decomposition of polynomial, skipping zero coefficients 
-  after the first.  *}
-
-lemma poly_decompose_lemma:
- assumes nz: "\<not>(\<forall>z. z\<noteq>0 \<longrightarrow> poly p z = (0::'a::{recpower,idom}))"
-  shows "\<exists>k a q. a\<noteq>0 \<and> Suc (length q + k) = length p \<and> 
-                 (\<forall>z. poly p z = z^k * poly (a#q) z)"
-using nz
-proof(induct p)
-  case Nil thus ?case by simp
-next
-  case (Cons c cs)
-  {assume c0: "c = 0"
-    
-    from Cons.hyps Cons.prems c0 have ?case apply auto
-      apply (rule_tac x="k+1" in exI)
-      apply (rule_tac x="a" in exI, clarsimp)
-      apply (rule_tac x="q" in exI)
-      by (auto simp add: power_Suc)}
-  moreover
-  {assume c0: "c\<noteq>0"
-    hence ?case apply-
-      apply (rule exI[where x=0])
-      apply (rule exI[where x=c], clarsimp)
-      apply (rule exI[where x=cs])
-      apply auto
-      done}
-  ultimately show ?case by blast
-qed
-
-lemma poly_decompose:
-  assumes nc: "~constant(poly p)"
-  shows "\<exists>k a q. a\<noteq>(0::'a::{recpower,idom}) \<and> k\<noteq>0 \<and>
-               length q + k + 1 = length p \<and> 
-              (\<forall>z. poly p z = poly p 0 + z^k * poly (a#q) z)"
-using nc 
-proof(induct p)
-  case Nil thus ?case by (simp add: constant_def)
-next
-  case (Cons c cs)
-  {assume C:"\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0"
-    {fix x y
-      from C have "poly (c#cs) x = poly (c#cs) y" by (cases "x=0", auto)}
-    with Cons.prems have False by (auto simp add: constant_def)}
-  hence th: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)" ..
-  from poly_decompose_lemma[OF th] 
-  show ?case 
-    apply clarsimp    
-    apply (rule_tac x="k+1" in exI)
-    apply (rule_tac x="a" in exI)
-    apply simp
-    apply (rule_tac x="q" in exI)
-    apply (auto simp add: power_Suc)
-    done
-qed
-
-text{* Fundamental theorem of algebral *}
-
-lemma fundamental_theorem_of_algebra:
-  assumes nc: "~constant(poly p)"
-  shows "\<exists>z::complex. poly p z = 0"
-using nc
-proof(induct n\<equiv> "length p" arbitrary: p rule: nat_less_induct)
-  fix n fix p :: "complex list"
-  let ?p = "poly p"
-  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"
-  let ?ths = "\<exists>z. ?p z = 0"
-
-  from nonconstant_length[OF nc] have n2: "n\<ge> 2" by (simp add: n)
-  from poly_minimum_modulus obtain c where 
-    c: "\<forall>w. cmod (?p c) \<le> cmod (?p w)" by blast
-  {assume pc: "?p c = 0" hence ?ths by blast}
-  moreover
-  {assume pc0: "?p c \<noteq> 0"
-    from poly_offset[of p c] obtain q where
-      q: "length q = length p" "\<forall>x. poly q x = ?p (c+x)" by blast
-    {assume h: "constant (poly q)"
-      from q(2) have th: "\<forall>x. poly q (x - c) = ?p x" by auto
-      {fix x y
-	from th have "?p x = poly q (x - c)" by auto 
-	also have "\<dots> = poly q (y - c)" 
-	  using h unfolding constant_def by blast
-	also have "\<dots> = ?p y" using th by auto
-	finally have "?p x = ?p y" .}
-      with nc have False unfolding constant_def by blast }
-    hence qnc: "\<not> constant (poly q)" by blast
-    from q(2) have pqc0: "?p c = poly q 0" by simp
-    from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)" by simp 
-    let ?a0 = "poly q 0"
-    from pc0 pqc0 have a00: "?a0 \<noteq> 0" by simp 
-    from a00 
-    have qr: "\<forall>z. poly q z = poly (map (op * (inverse ?a0)) q) z * ?a0"
-      by (simp add: poly_cmult_map)
-    let ?r = "map (op * (inverse ?a0)) q"
-    have lgqr: "length q = length ?r" by simp 
-    {assume h: "\<And>x y. poly ?r x = poly ?r y"
-      {fix x y
-	from qr[rule_format, of x] 
-	have "poly q x = poly ?r x * ?a0" by auto
-	also have "\<dots> = poly ?r y * ?a0" using h by simp
-	also have "\<dots> = poly q y" using qr[rule_format, of y] by simp
-	finally have "poly q x = poly q y" .} 
-      with qnc have False unfolding constant_def by blast}
-    hence rnc: "\<not> constant (poly ?r)" unfolding constant_def by blast
-    from qr[rule_format, of 0] a00  have r01: "poly ?r 0 = 1" by auto
-    {fix w 
-      have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
-	using qr[rule_format, of w] a00 by simp
-      also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
-	using a00 unfolding norm_divide by (simp add: field_simps)
-      finally have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" .}
-    note mrmq_eq = this
-    from poly_decompose[OF rnc] obtain k a s where 
-      kas: "a\<noteq>0" "k\<noteq>0" "length s + k + 1 = length ?r" 
-      "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (a#s) z" by blast
-    {assume "k + 1 = n"
-      with kas(3) lgqr[symmetric] q(1) n[symmetric] have s0:"s=[]" by auto
-      {fix w
-	have "cmod (poly ?r w) = cmod (1 + a * w ^ k)" 
-	  using kas(4)[rule_format, of w] s0 r01 by (simp add: ring_simps)}
-      note hth = this [symmetric]
-	from reduce_poly_simple[OF kas(1,2)] 
-      have "\<exists>w. cmod (poly ?r w) < 1" unfolding hth by blast}
-    moreover
-    {assume kn: "k+1 \<noteq> n"
-      from kn kas(3) q(1) n[symmetric] have k1n: "k + 1 < n" by simp
-      have th01: "\<not> constant (poly (1#((replicate (k - 1) 0)@[a])))" 
-	unfolding constant_def poly_Nil poly_Cons poly_replicate_append
-	using kas(1) apply simp 
-	by (rule exI[where x=0], rule exI[where x=1], simp)
-      from kas(2) have th02: "k+1 = length (1#((replicate (k - 1) 0)@[a]))" 
-	by simp
-      from H[rule_format, OF k1n th01 th02]
-      obtain w where w: "1 + w^k * a = 0"
-	unfolding poly_Nil poly_Cons poly_replicate_append
-	using kas(2) by (auto simp add: power_Suc[symmetric, of _ "k - Suc 0"] 
-	  mult_assoc[of _ _ a, symmetric])
-      from poly_bound_exists[of "cmod w" s] obtain m where 
-	m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
-      have w0: "w\<noteq>0" using kas(2) w by (auto simp add: power_0_left)
-      from w have "(1 + w ^ k * a) - 1 = 0 - 1" by simp
-      then have wm1: "w^k * a = - 1" by simp
-      have inv0: "0 < inverse (cmod w ^ (k + 1) * m)" 
-	using norm_ge_zero[of w] w0 m(1)
-	  by (simp add: inverse_eq_divide zero_less_mult_iff)
-      with real_down2[OF zero_less_one] obtain t where
-	t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
-      let ?ct = "complex_of_real t"
-      let ?w = "?ct * w"
-      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)
-      also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
-	unfolding wm1 by (simp)
-      finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) = cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)" 
-	apply -
-	apply (rule cong[OF refl[of cmod]])
-	apply assumption
-	done
-      with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"] 
-      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 
-      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
-      have "t *cmod w \<le> 1 * cmod w" apply (rule mult_mono) using t(1,2) by auto
-      then have tw: "cmod ?w \<le> cmod w" using t(1) by (simp add: norm_mult) 
-      from t inv0 have "t* (cmod w ^ (k + 1) * m) < 1"
-	by (simp add: inverse_eq_divide field_simps)
-      with zero_less_power[OF t(1), of k] 
-      have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1" 
-	apply - apply (rule mult_strict_left_mono) by simp_all
-      have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k+1) * cmod (poly s ?w)))"  using w0 t(1)
-	by (simp add: ring_simps power_mult_distrib norm_of_real norm_power norm_mult)
-      then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
-	using t(1,2) m(2)[rule_format, OF tw] w0
-	apply (simp only: )
-	apply auto
-	apply (rule mult_mono, simp_all add: norm_ge_zero)+
-	apply (simp add: zero_le_mult_iff zero_le_power)
-	done
-      with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k" by simp 
-      from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1" 
-	by auto
-      from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] th120 th121]
-      have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" . 
-      from th11 th12
-      have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"  by arith 
-      then have "cmod (poly ?r ?w) < 1" 
-	unfolding kas(4)[rule_format, of ?w] r01 by simp 
-      then have "\<exists>w. cmod (poly ?r w) < 1" by blast}
-    ultimately have cr0_contr: "\<exists>w. cmod (poly ?r w) < 1" by blast
-    from cr0_contr cq0 q(2)
-    have ?ths unfolding mrmq_eq not_less[symmetric] by auto}
-  ultimately show ?ths by blast
-qed
-
-text {* Alternative version with a syntactic notion of constant polynomial. *}
-
-lemma fundamental_theorem_of_algebra_alt:
-  assumes nc: "~(\<exists>a l. a\<noteq> 0 \<and> list_all(\<lambda>b. b = 0) l \<and> p = a#l)"
-  shows "\<exists>z. poly p z = (0::complex)"
-using nc
-proof(induct p)
-  case (Cons c cs)
-  {assume "c=0" hence ?case by auto}
-  moreover
-  {assume c0: "c\<noteq>0"
-    {assume nc: "constant (poly (c#cs))"
-      from nc[unfolded constant_def, rule_format, of 0] 
-      have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto 
-      hence "list_all (\<lambda>c. c=0) cs"
-	proof(induct cs)
-	  case (Cons d ds)
-	  {assume "d=0" hence ?case using Cons.prems Cons.hyps by simp}
-	  moreover
-	  {assume d0: "d\<noteq>0"
-	    from poly_bound_exists[of 1 ds] obtain m where 
-	      m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
-	    have dm: "cmod d / m > 0" using d0 m(1) by (simp add: field_simps)
-	    from real_down2[OF dm zero_less_one] obtain x where 
-	      x: "x > 0" "x < cmod d / m" "x < 1" by blast
-	    let ?x = "complex_of_real x"
-	    from x have cx: "?x \<noteq> 0"  "cmod ?x \<le> 1" by simp_all
-	    from Cons.prems[rule_format, OF cx(1)]
-	    have cth: "cmod (?x*poly ds ?x) = cmod d" by (simp add: eq_diff_eq[symmetric])
-	    from m(2)[rule_format, OF cx(2)] x(1)
-	    have th0: "cmod (?x*poly ds ?x) \<le> x*m"
-	      by (simp add: norm_mult)
-	    from x(2) m(1) have "x*m < cmod d" by (simp add: field_simps)
-	    with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d" by auto
-	    with cth  have ?case by blast}
-	  ultimately show ?case by blast 
-	qed simp}
-      then have nc: "\<not> constant (poly (c#cs))" using Cons.prems c0 
-	by blast
-      from fundamental_theorem_of_algebra[OF nc] have ?case .}
-  ultimately show ?case by blast  
-qed simp
-
-subsection{* Nullstellenstatz, degrees and divisibility of polynomials *}
-
-lemma nullstellensatz_lemma:
-  fixes p :: "complex list"
-  assumes "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
-  and "degree p = n" and "n \<noteq> 0"
-  shows "p divides (pexp q n)"
-using prems
-proof(induct n arbitrary: p q rule: nat_less_induct)
-  fix n::nat fix p q :: "complex list"
-  assume IH: "\<forall>m<n. \<forall>p q.
-                 (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow>
-                 degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p divides (q %^ m)"
-    and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0" 
-    and dpn: "degree p = n" and n0: "n \<noteq> 0"
-  let ?ths = "p divides (q %^ n)"
-  {fix a assume a: "poly p a = 0"
-    {assume p0: "poly p = poly []" 
-      hence ?ths unfolding divides_def  using pq0 n0
-	apply - apply (rule exI[where x="[]"], rule ext)
-	by (auto simp add: poly_mult poly_exp)}
-    moreover
-    {assume p0: "poly p \<noteq> poly []" 
-      and oa: "order  a p \<noteq> 0"
-      from p0 have pne: "p \<noteq> []" by auto
-      let ?op = "order a p"
-      from p0 have ap: "([- a, 1] %^ ?op) divides p" 
-	"\<not> pexp [- a, 1] (Suc ?op) divides p" using order by blast+ 
-      note oop = order_degree[OF p0, unfolded dpn]
-      {assume q0: "q = []"
-	hence ?ths using n0 unfolding divides_def 
-	  apply simp
-	  apply (rule exI[where x="[]"], rule ext)
-	  by (simp add: divides_def poly_exp poly_mult)}
-      moreover
-      {assume q0: "q\<noteq>[]"
-	from pq0[rule_format, OF a, unfolded poly_linear_divides] q0
-	obtain r where r: "q = pmult [- a, 1] r" by blast
-	from ap[unfolded divides_def] obtain s where
-	  s: "poly p = poly (pmult (pexp [- a, 1] ?op) s)" by blast
-	have s0: "poly s \<noteq> poly []"
-	  using s p0 by (simp add: poly_entire)
-	hence pns0: "poly (pnormalize s) \<noteq> poly []" and sne: "s\<noteq>[]" by auto
-	{assume ds0: "degree s = 0"
-	  from ds0 pns0 have "\<exists>k. pnormalize s = [k]" unfolding degree_def 
-	    by (cases "pnormalize s", auto)
-	  then obtain k where kpn: "pnormalize s = [k]" by blast
-	  from pns0[unfolded poly_zero] kpn have k: "k \<noteq>0" "poly s = poly [k]"
-	    using poly_normalize[of s] by simp_all
-	  let ?w = "pmult (pmult [1/k] (pexp [-a,1] (n - ?op))) (pexp r n)"
-	  from k r s oop have "poly (pexp q n) = poly (pmult p ?w)"
-	    by - (rule ext, simp add: poly_mult poly_exp poly_cmult poly_add power_add[symmetric] ring_simps power_mult_distrib[symmetric])
-	  hence ?ths unfolding divides_def by blast}
-	moreover
-	{assume ds0: "degree s \<noteq> 0"
-	  from ds0 s0 dpn degree_unique[OF s, unfolded linear_pow_mul_degree] oa
-	    have dsn: "degree s < n" by auto 
-	    {fix x assume h: "poly s x = 0"
-	      {assume xa: "x = a"
-		from h[unfolded xa poly_linear_divides] sne obtain u where
-		  u: "s = pmult [- a, 1] u" by blast
-		have "poly p = poly (pmult (pexp [- a, 1] (Suc ?op)) u)"
-		  unfolding s u
-		  apply (rule ext)
-		  by (simp add: ring_simps power_mult_distrib[symmetric] poly_mult poly_cmult poly_add poly_exp)
-		with ap(2)[unfolded divides_def] have False by blast}
-	      note xa = this
-	      from h s have "poly p x = 0" by (simp add: poly_mult)
-	      with pq0 have "poly q x = 0" by blast
-	      with r xa have "poly r x = 0"
-		by (auto simp add: poly_mult poly_add poly_cmult eq_diff_eq[symmetric])}
-	    note impth = this
-	    from IH[rule_format, OF dsn, of s r] impth ds0
-	    have "s divides (pexp r (degree s))" by blast
-	    then obtain u where u: "poly (pexp r (degree s)) = poly (pmult s u)"
-	      unfolding divides_def by blast
-	    hence u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
-	      by (simp add: poly_mult[symmetric] poly_exp[symmetric])
-	    let ?w = "pmult (pmult u (pexp [-a,1] (n - ?op))) (pexp r (n - degree s))"
-	    from u' s r oop[of a] dsn have "poly (pexp q n) = poly (pmult p ?w)"
-	      apply - apply (rule ext)
-	      apply (simp only:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult ring_simps)
-	      
-	      apply (simp add:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult mult_assoc[symmetric])
-	      done
-	    hence ?ths unfolding divides_def by blast}
-      ultimately have ?ths by blast }
-      ultimately have ?ths by blast}
-    ultimately have ?ths using a order_root by blast}
-  moreover
-  {assume exa: "\<not> (\<exists>a. poly p a = 0)"
-    from fundamental_theorem_of_algebra_alt[of p] exa obtain c cs where
-      ccs: "c\<noteq>0" "list_all (\<lambda>c. c = 0) cs" "p = c#cs" by blast
-    
-    from poly_0[OF ccs(2)] ccs(3) 
-    have pp: "\<And>x. poly p x =  c" by simp
-    let ?w = "pmult [1/c] (pexp q n)"
-    from pp ccs(1) 
-    have "poly (pexp q n) = poly (pmult p ?w) "
-      apply - apply (rule ext)
-      unfolding poly_mult_assoc[symmetric] by (simp add: poly_mult)
-    hence ?ths unfolding divides_def by blast}
-  ultimately show ?ths by blast
-qed
-
-lemma nullstellensatz_univariate:
-  "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> 
-    p divides (q %^ (degree p)) \<or> (poly p = poly [] \<and> poly q = poly [])"
-proof-
-  {assume pe: "poly p = poly []"
-    hence eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> poly q = poly []"
-      apply auto
-      by (rule ext, simp)
-    {assume "p divides (pexp q (degree p))"
-      then obtain r where r: "poly (pexp q (degree p)) = poly (pmult p r)" 
-	unfolding divides_def by blast
-      from cong[OF r refl] pe degree_unique[OF pe]
-      have False by (simp add: poly_mult degree_def)}
-    with eq pe have ?thesis by blast}
-  moreover
-  {assume pe: "poly p \<noteq> poly []"
-    have p0: "poly [0] = poly []" by (rule ext, simp)
-    {assume dp: "degree p = 0"
-      then obtain k where "pnormalize p = [k]" using pe poly_normalize[of p]
-	unfolding degree_def by (cases "pnormalize p", auto)
-      hence k: "pnormalize p = [k]" "poly p = poly [k]" "k\<noteq>0"
-	using pe poly_normalize[of p] by (auto simp add: p0)
-      hence th1: "\<forall>x. poly p x \<noteq> 0" by simp
-      from k(2,3) dp have "poly (pexp q (degree p)) = poly (pmult p [1/k]) "
-	by - (rule ext, simp add: poly_mult poly_exp)
-      hence th2: "p divides (pexp q (degree p))" unfolding divides_def by blast
-      from th1 th2 pe have ?thesis by blast}
-    moreover
-    {assume dp: "degree p \<noteq> 0"
-      then obtain n where n: "degree p = Suc n " by (cases "degree p", auto)
-      {assume "p divides (pexp q (Suc n))"
-	then obtain u where u: "poly (pexp q (Suc n)) = poly (pmult p u)"
-	  unfolding divides_def by blast
-	hence u' :"\<And>x. poly (pexp q (Suc n)) x = poly (pmult p u) x" by simp_all
-	{fix x assume h: "poly p x = 0" "poly q x \<noteq> 0"
-	  hence "poly (pexp q (Suc n)) x \<noteq> 0" by (simp only: poly_exp) simp	  
-	  hence False using u' h(1) by (simp only: poly_mult poly_exp) simp}}
-	with n nullstellensatz_lemma[of p q "degree p"] dp 
-	have ?thesis by auto}
-    ultimately have ?thesis by blast}
-  ultimately show ?thesis by blast
-qed
-
-text{* Useful lemma *}
-
-lemma (in idom_char_0) constant_degree: "constant (poly p) \<longleftrightarrow> degree p = 0" (is "?lhs = ?rhs")
-proof
-  assume l: ?lhs
-  from l[unfolded constant_def, rule_format, of _ "zero"]
-  have th: "poly p = poly [poly p 0]" apply - by (rule ext, simp)
-  from degree_unique[OF th] show ?rhs by (simp add: degree_def)
-next
-  assume r: ?rhs
-  from r have "pnormalize p = [] \<or> (\<exists>k. pnormalize p = [k])"
-    unfolding degree_def by (cases "pnormalize p", auto)
-  then show ?lhs unfolding constant_def poly_normalize[of p, symmetric]
-    by (auto simp del: poly_normalize)
-qed
-
-(* It would be nicer to prove this without using algebraic closure...        *)
-
-lemma divides_degree_lemma: assumes dpn: "degree (p::complex list) = n"
-  shows "n \<le> degree (p *** q) \<or> poly (p *** q) = poly []"
-  using dpn
-proof(induct n arbitrary: p q)
-  case 0 thus ?case by simp
-next
-  case (Suc n p q)
-  from Suc.prems fundamental_theorem_of_algebra[of p] constant_degree[of p]
-  obtain a where a: "poly p a = 0" by auto
-  then obtain r where r: "p = pmult [-a, 1] r" unfolding poly_linear_divides
-    using Suc.prems by (auto simp add: degree_def)
-  {assume h: "poly (pmult r q) = poly []"
-    hence "poly (pmult p q) = poly []" using r
-      apply - apply (rule ext)  by (auto simp add: poly_entire poly_mult poly_add poly_cmult) hence ?case by blast}
-  moreover
-  {assume h: "poly (pmult r q) \<noteq> poly []" 
-    hence r0: "poly r \<noteq> poly []" and q0: "poly q \<noteq> poly []"
-      by (auto simp add: poly_entire)
-    have eq: "poly (pmult p q) = poly (pmult [-a, 1] (pmult r q))"
-      apply - apply (rule ext)
-      by (simp add: r poly_mult poly_add poly_cmult ring_simps)
-    from linear_mul_degree[OF h, of "- a"]
-    have dqe: "degree (pmult p q) = degree (pmult r q) + 1"
-      unfolding degree_unique[OF eq] .
-    from linear_mul_degree[OF r0, of "- a", unfolded r[symmetric]] r Suc.prems 
-    have dr: "degree r = n" by auto
-    from  Suc.hyps[OF dr, of q] have "Suc n \<le> degree (pmult p q)"
-      unfolding dqe using h by (auto simp del: poly.simps) 
-    hence ?case by blast}
-  ultimately show ?case by blast
-qed
-
-lemma divides_degree: assumes pq: "p divides (q:: complex list)"
-  shows "degree p \<le> degree q \<or> poly q = poly []"
-using pq  divides_degree_lemma[OF refl, of p]
-apply (auto simp add: divides_def poly_entire)
-apply atomize
-apply (erule_tac x="qa" in allE, auto)
-apply (subgoal_tac "degree q = degree (p *** qa)", simp)
-apply (rule degree_unique, simp)
-done
-
-(* Arithmetic operations on multivariate polynomials.                        *)
-
-lemma mpoly_base_conv: 
-  "(0::complex) \<equiv> poly [] x" "c \<equiv> poly [c] x" "x \<equiv> poly [0,1] x" by simp_all
-
-lemma mpoly_norm_conv: 
-  "poly [0] (x::complex) \<equiv> poly [] x" "poly [poly [] y] x \<equiv> poly [] x" by simp_all
-
-lemma mpoly_sub_conv: 
-  "poly p (x::complex) - poly q x \<equiv> poly p x + -1 * poly q x"
-  by (simp add: diff_def)
-
-lemma poly_pad_rule: "poly p x = 0 ==> poly (0#p) x = (0::complex)" by simp
-
-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
-
-lemma resolve_eq_raw:  "poly [] x \<equiv> 0" "poly [c] x \<equiv> (c::complex)" by auto
-lemma  resolve_eq_then: "(P \<Longrightarrow> (Q \<equiv> Q1)) \<Longrightarrow> (\<not>P \<Longrightarrow> (Q \<equiv> Q2))
-  \<Longrightarrow> Q \<equiv> P \<and> Q1 \<or> \<not>P\<and> Q2" apply (atomize (full)) by blast 
-lemma expand_ex_beta_conv: "list_ex P [c] \<equiv> P c" by simp
-
-lemma poly_divides_pad_rule: 
-  fixes p q :: "complex list"
-  assumes pq: "p divides q"
-  shows "p divides ((0::complex)#q)"
-proof-
-  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
-  hence "poly (0#q) = poly (p *** ([0,1] *** r))" 
-    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
-  thus ?thesis unfolding divides_def by blast
-qed
-
-lemma poly_divides_pad_const_rule: 
-  fixes p q :: "complex list"
-  assumes pq: "p divides q"
-  shows "p divides (a %* q)"
-proof-
-  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
-  hence "poly (a %* q) = poly (p *** (a %* r))" 
-    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
-  thus ?thesis unfolding divides_def by blast
-qed
-
-
-lemma poly_divides_conv0:  
-  fixes p :: "complex list"
-  assumes lgpq: "length q < length p" and lq:"last p \<noteq> 0"
-  shows "p divides q \<equiv> (\<not> (list_ex (\<lambda>c. c \<noteq> 0) q))" (is "?lhs \<equiv> ?rhs")
-proof-
-  {assume r: ?rhs 
-    hence eq: "poly q = poly []" unfolding poly_zero 
-      by (simp add: list_all_iff list_ex_iff)
-    hence "poly q = poly (p *** [])" by - (rule ext, simp add: poly_mult)
-    hence ?lhs unfolding divides_def  by blast}
-  moreover
-  {assume l: ?lhs
-    have ath: "\<And>lq lp dq::nat. lq < lp ==> lq \<noteq> 0 \<Longrightarrow> dq <= lq - 1 ==> dq < lp - 1"
-      by arith
-    {assume q0: "length q = 0"
-      hence "q = []" by simp
-      hence ?rhs by simp}
-    moreover
-    {assume lgq0: "length q \<noteq> 0"
-      from pnormalize_length[of q] have dql: "degree q \<le> length q - 1" 
-	unfolding degree_def by simp
-      from ath[OF lgpq lgq0 dql, unfolded pnormal_degree[OF lq, symmetric]] divides_degree[OF l] have "poly q = poly []" by auto
-      hence ?rhs unfolding poly_zero by (simp add: list_all_iff list_ex_iff)}
-    ultimately have ?rhs by blast }
-  ultimately show "?lhs \<equiv> ?rhs" by - (atomize (full), blast) 
-qed
-
-lemma poly_divides_conv1: 
-  assumes a0: "a\<noteq> (0::complex)" and pp': "(p::complex list) divides p'"
-  and qrp': "\<And>x. a * poly q x - poly p' x \<equiv> poly r x"
-  shows "p divides q \<equiv> p divides (r::complex list)" (is "?lhs \<equiv> ?rhs")
-proof-
-  {
-  from pp' obtain t where t: "poly p' = poly (p *** t)" 
-    unfolding divides_def by blast
-  {assume l: ?lhs
-    then obtain u where u: "poly q = poly (p *** u)" unfolding divides_def by blast
-     have "poly r = poly (p *** ((a %* u) +++ (-- t)))"
-       using u qrp' t
-       by - (rule ext, 
-	 simp add: poly_add poly_mult poly_cmult poly_minus ring_simps)
-     then have ?rhs unfolding divides_def by blast}
-  moreover
-  {assume r: ?rhs
-    then obtain u where u: "poly r = poly (p *** u)" unfolding divides_def by blast
-    from u t qrp' a0 have "poly q = poly (p *** ((1/a) %* (u +++ t)))"
-      by - (rule ext, atomize (full), simp add: poly_mult poly_add poly_cmult field_simps)
-    hence ?lhs  unfolding divides_def by blast}
-  ultimately have "?lhs = ?rhs" by blast }
-thus "?lhs \<equiv> ?rhs"  by - (atomize(full), blast) 
-qed
-
-lemma basic_cqe_conv1:
-  "(\<exists>x. poly p x = 0 \<and> poly [] x \<noteq> 0) \<equiv> False"
-  "(\<exists>x. poly [] x \<noteq> 0) \<equiv> False"
-  "(\<exists>x. poly [c] x \<noteq> 0) \<equiv> c\<noteq>0"
-  "(\<exists>x. poly [] x = 0) \<equiv> True"
-  "(\<exists>x. poly [c] x = 0) \<equiv> c = 0" by simp_all
-
-lemma basic_cqe_conv2: 
-  assumes l:"last (a#b#p) \<noteq> 0" 
-  shows "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True"
-proof-
-  {fix h t
-    assume h: "h\<noteq>0" "list_all (\<lambda>c. c=(0::complex)) t"  "a#b#p = h#t"
-    hence "list_all (\<lambda>c. c= 0) (b#p)" by simp
-    moreover have "last (b#p) \<in> set (b#p)" by simp
-    ultimately have "last (b#p) = 0" by (simp add: list_all_iff)
-    with l have False by simp}
-  hence th: "\<not> (\<exists> h t. h\<noteq>0 \<and> list_all (\<lambda>c. c=0) t \<and> a#b#p = h#t)"
-    by blast
-  from fundamental_theorem_of_algebra_alt[OF th] 
-  show "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True" by auto
-qed
-
-lemma  basic_cqe_conv_2b: "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
-proof-
-  have "\<not> (list_ex (\<lambda>c. c \<noteq> 0) p) \<longleftrightarrow> poly p = poly []" 
-    by (simp add: poly_zero list_all_iff list_ex_iff)
-  also have "\<dots> \<longleftrightarrow> (\<not> (\<exists>x. poly p x \<noteq> 0))" by (auto intro: ext)
-  finally show "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
-    by - (atomize (full), blast)
-qed
-
-lemma basic_cqe_conv3:
-  fixes p q :: "complex list"
-  assumes l: "last (a#p) \<noteq> 0" 
-  shows "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
-proof-
-  note np = pnormalize_eq[OF l]
-  {assume "poly (a#p) = poly []" hence False using l
-      unfolding poly_zero apply (auto simp add: list_all_iff del: last.simps)
-      apply (cases p, simp_all) done}
-  then have p0: "poly (a#p) \<noteq> poly []"  by blast
-  from np have dp:"degree (a#p) = length p" by (simp add: degree_def)
-  from nullstellensatz_univariate[of "a#p" q] p0 dp
-  show "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
-    by - (atomize (full), auto)
-qed
-
-lemma basic_cqe_conv4:
-  fixes p q :: "complex list"
-  assumes h: "\<And>x. poly (q %^ n) x \<equiv> poly r x"
-  shows "p divides (q %^ n) \<equiv> p divides r"
-proof-
-  from h have "poly (q %^ n) = poly r" by (auto intro: ext)  
-  thus "p divides (q %^ n) \<equiv> p divides r" unfolding divides_def by simp
-qed
-
-lemma pmult_Cons_Cons: "((a::complex)#b#p) *** q = (a %*q) +++ (0#((b#p) *** q))"
-  by simp
-
-lemma elim_neg_conv: "- z \<equiv> (-1) * (z::complex)" by simp
-lemma eqT_intr: "PROP P \<Longrightarrow> (True \<Longrightarrow> PROP P )" "PROP P \<Longrightarrow> True" by blast+
-lemma negate_negate_rule: "Trueprop P \<equiv> \<not> P \<equiv> False" by (atomize (full), auto)
-lemma last_simps: "last [x] = x" "last (x#y#ys) = last (y#ys)" by simp_all
-lemma length_simps: "length [] = 0" "length (x#y#xs) = length xs + 2" "length [x] = 1" by simp_all
-
-lemma complex_entire: "(z::complex) \<noteq> 0 \<and> w \<noteq> 0 \<equiv> z*w \<noteq> 0" by simp
-lemma resolve_eq_ne: "(P \<equiv> True) \<equiv> (\<not>P \<equiv> False)" "(P \<equiv> False) \<equiv> (\<not>P \<equiv> True)" 
-  by (atomize (full)) simp_all
-lemma cqe_conv1: "poly [] x = 0 \<longleftrightarrow> True"  by simp
-lemma cqe_conv2: "(p \<Longrightarrow> (q \<equiv> r)) \<equiv> ((p \<and> q) \<equiv> (p \<and> r))"  (is "?l \<equiv> ?r")
-proof
-  assume "p \<Longrightarrow> q \<equiv> r" thus "p \<and> q \<equiv> p \<and> r" apply - apply (atomize (full)) by blast
-next
-  assume "p \<and> q \<equiv> p \<and> r" "p"
-  thus "q \<equiv> r" apply - apply (atomize (full)) apply blast done
-qed
-lemma poly_const_conv: "poly [c] (x::complex) = y \<longleftrightarrow> c = y" by simp
-
-end
\ No newline at end of file
--- a/src/HOL/Complex/README.html	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,67 +0,0 @@
-<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
-
-<!-- $Id$ -->
-
-<HTML>
-
-<HEAD>
-  <meta http-equiv="content-type" content="text/html;charset=iso-8859-1">
-  <TITLE>HOL/Complex/README</TITLE>
-</HEAD>
-
-<BODY>
-
-<H1>Complex: The Complex Numbers</H1>
-		<P>This directory defines the type <KBD>complex</KBD> of the complex numbers,
-with numeric constants and some complex analysis.  The development includes
-nonstandard analysis for the complex numbers.  Note that the image
-<KBD>HOL-Complex</KBD> includes theories from the directories 
-<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).
-
-<ul>
-<li><a href="CLim.html">CLim</a> Limits, continuous functions, and derivatives for the complex numbers
-<li><a href="CSeries.html">CSeries</a> Finite summation and infinite series for the complex numbers
-<li><a href="CStar.html">CStar</a> Star-transforms for the complex numbers, to form non-standard extensions of sets and functions
-<li><a href="Complex.html">Complex</a> The complex numbers
-<li><a href="NSCA.html">NSCA</a> Nonstandard complex analysis
-<li><a href="NSComplex.html">NSComplex</a> Ultrapower construction of the nonstandard complex numbers
-</ul>
-
-<h2><a name="Anchor-Real" id="Anchor-Real"></a>Real: Dedekind Cut Construction of the Real Line</h2>
-
-<ul>
-<li><a href="Lubs.html">Lubs</a> Definition of upper bounds, lubs and so on, to support completeness proofs.
-<li><a href="PReal.html">PReal</a> The positive reals constructed using Dedekind cuts
-<li><a href="Rational.html">Rational</a> The rational numbers constructed as equivalence classes of integers
-<li><a href="RComplete.html">RComplete</a> The reals are complete: they satisfy the supremum property. They also have the Archimedean property.
-<li><a href="RealDef.html">RealDef</a> The real numbers, their ordering properties, and embedding of the integers and the natural numbers
-<li><a href="RealPow.html">RealPow</a> Real numbers raised to natural number powers
-</ul>
-<h2><a name="Anchor-Hyperreal" id="Anchor-Hyperreal"></a>Hyperreal: Ultrafilter Construction of the Non-Standard Reals</h2>
-See J. D. Fleuriot and L. C. Paulson. Mechanizing Nonstandard Real Analysis. LMS J. Computation and Mathematics 3 (2000), 140-190.
-<ul>
-<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.
-<li><a href="HLog.html">HLog</a> Non-standard logarithms
-<li><a href="HSeries.html">HSeries</a> Non-standard theory of finite summation and infinite series
-<li><a href="HTranscendental.html">HTranscendental</a> Non-standard extensions of transcendental functions
-<li><a href="HyperDef.html">HyperDef</a> Ultrapower construction of the hyperreals
-<li><a href="HyperNat.html">HyperNat</a> Ultrapower construction of the hypernaturals
-<li><a href="HyperPow.html">HyperPow</a> Powers theory for the hyperreals
-<!-- <li><a href="IntFloor.html">IntFloor</a> Floor and Ceiling functions relating the reals and integers -->
-<li><a href="Integration.html">Integration</a> Gage integrals
-<li><a href="Lim.html">Lim</a> Theory of limits, continuous functions, and derivatives
-<li><a href="Log.html">Log</a> Logarithms for the reals
-<li><a href="MacLaurin.html">MacLaurin</a> MacLaurin series
-<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
-<li><a href="NthRoot.html">NthRoot</a> Existence of n-th roots of real numbers
-<li><a href="NSA.html">NSA</a> Theory defining sets of infinite numbers, infinitesimals, the infinitely close relation, and their various algebraic properties.
-<li><a href="Poly.html">Poly</a> Univariate real polynomials
-<li><a href="SEQ.html">SEQ</a> Convergence of sequences and series using standard and nonstandard analysis
-<li><a href="Series.html">Series</a> Finite summation and infinite series for the reals
-<li><a href="Star.html">Star</a> Nonstandard extensions of real sets and real functions
-<li><a href="Transcendental.html">Transcendental</a> Power series and transcendental functions
-</ul>
-<HR>
-<P>Last modified $Date$
-</BODY>
-</HTML>
--- a/src/HOL/Complex/document/root.tex	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,32 +0,0 @@
-
-% $Id$
-
-\documentclass[11pt,a4paper]{article}
-\usepackage{graphicx,isabelle,isabellesym,latexsym}
-\usepackage[latin1]{inputenc}
-\usepackage{pdfsetup}
-
-\urlstyle{rm}
-\isabellestyle{it}
-\pagestyle{myheadings}
-
-\begin{document}
-
-\title{Isabelle/HOL-Complex --- Higher-Order Logic with Complex Numbers}
-\maketitle
-
-\tableofcontents
-
-\begin{center}
-  \includegraphics[width=\textwidth,height=\textheight,keepaspectratio]{session_graph}
-\end{center}
-
-\newpage
-
-\renewcommand{\isamarkupheader}[1]%
-{\section{\isabellecontext: #1}\markright{THEORY~``\isabellecontext''}}
-
-\parindent 0pt\parskip 0.5ex
-\input{session}
-
-\end{document}
--- a/src/HOL/Complex_Main.thy	Mon Dec 29 13:23:53 2008 +0100
+++ b/src/HOL/Complex_Main.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -9,7 +9,7 @@
 imports
   Main
   Real
-  "~~/src/HOL/Complex/Fundamental_Theorem_Algebra"
+  Fundamental_Theorem_Algebra
   Log
   Ln
   Taylor
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Dense_Linear_Order.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,877 @@
+(* Author: Amine Chaieb, TU Muenchen *)
+
+header {* Dense linear order without endpoints
+  and a quantifier elimination procedure in Ferrante and Rackoff style *}
+
+theory Dense_Linear_Order
+imports Plain Groebner_Basis
+uses
+  "~~/src/HOL/Tools/Qelim/langford_data.ML"
+  "~~/src/HOL/Tools/Qelim/ferrante_rackoff_data.ML"
+  ("~~/src/HOL/Tools/Qelim/langford.ML")
+  ("~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML")
+begin
+
+setup {* Langford_Data.setup #> Ferrante_Rackoff_Data.setup *}
+
+context linorder
+begin
+
+lemma less_not_permute: "\<not> (x < y \<and> y < x)" by (simp add: not_less linear)
+
+lemma gather_simps: 
+  shows 
+  "(\<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)"
+  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)"
+  "(\<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))"
+  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
+
+lemma 
+  gather_start: "(\<exists>x. P x) \<equiv> (\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y\<in> {}. x < y) \<and> P x)" 
+  by simp
+
+text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"}*}
+lemma minf_lt:  "\<exists>z . \<forall>x. x < z \<longrightarrow> (x < t \<longleftrightarrow> True)" by auto
+lemma minf_gt: "\<exists>z . \<forall>x. x < z \<longrightarrow>  (t < x \<longleftrightarrow>  False)"
+  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
+
+lemma minf_le: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<le> t \<longleftrightarrow> True)" by (auto simp add: less_le)
+lemma minf_ge: "\<exists>z. \<forall>x. x < z \<longrightarrow> (t \<le> x \<longleftrightarrow> False)"
+  by (auto simp add: less_le not_less not_le)
+lemma minf_eq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
+lemma minf_neq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
+lemma minf_P: "\<exists>z. \<forall>x. x < z \<longrightarrow> (P \<longleftrightarrow> P)" by blast
+
+text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"}*}
+lemma pinf_gt:  "\<exists>z . \<forall>x. z < x \<longrightarrow> (t < x \<longleftrightarrow> True)" by auto
+lemma pinf_lt: "\<exists>z . \<forall>x. z < x \<longrightarrow>  (x < t \<longleftrightarrow>  False)"
+  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
+
+lemma pinf_ge: "\<exists>z. \<forall>x. z < x \<longrightarrow> (t \<le> x \<longleftrightarrow> True)" by (auto simp add: less_le)
+lemma pinf_le: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<le> t \<longleftrightarrow> False)"
+  by (auto simp add: less_le not_less not_le)
+lemma pinf_eq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
+lemma pinf_neq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
+lemma pinf_P: "\<exists>z. \<forall>x. z < x \<longrightarrow> (P \<longleftrightarrow> P)" by blast
+
+lemma nmi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x < t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
+lemma nmi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t < x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)"
+  by (auto simp add: le_less)
+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
+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
+lemma  nmi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x = t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
+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
+lemma  nmi_P: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
+lemma  nmi_conj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x) ;
+  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
+  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
+lemma  nmi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x) ;
+  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
+  \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
+
+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)
+lemma  npi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t < x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
+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
+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
+lemma  npi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x = t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
+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
+lemma  npi_P: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
+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>
+  \<Longrightarrow>  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
+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>
+  \<Longrightarrow> \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
+
+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)"
+proof(clarsimp)
+  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"
+    and xu: "x<u"  and px: "x < t" and ly: "l<y" and yu:"y < u"
+  from tU noU ly yu have tny: "t\<noteq>y" by auto
+  {assume H: "t < y"
+    from less_trans[OF lx px] less_trans[OF H yu]
+    have "l < t \<and> t < u"  by simp
+    with tU noU have "False" by auto}
+  hence "\<not> t < y"  by auto hence "y \<le> t" by (simp add: not_less)
+  thus "y < t" using tny by (simp add: less_le)
+qed
+
+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)"
+proof(clarsimp)
+  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" and xu: "x<u"
+  and px: "t < x" and ly: "l<y" and yu:"y < u"
+  from tU noU ly yu have tny: "t\<noteq>y" by auto
+  {assume H: "y< t"
+    from less_trans[OF ly H] less_trans[OF px xu] have "l < t \<and> t < u" by simp
+    with tU noU have "False" by auto}
+  hence "\<not> y<t"  by auto hence "t \<le> y" by (auto simp add: not_less)
+  thus "t < y" using tny by (simp add:less_le)
+qed
+
+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)"
+proof(clarsimp)
+  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" and xu: "x<u"
+  and px: "x \<le> t" and ly: "l<y" and yu:"y < u"
+  from tU noU ly yu have tny: "t\<noteq>y" by auto
+  {assume H: "t < y"
+    from less_le_trans[OF lx px] less_trans[OF H yu]
+    have "l < t \<and> t < u" by simp
+    with tU noU have "False" by auto}
+  hence "\<not> t < y"  by auto thus "y \<le> t" by (simp add: not_less)
+qed
+
+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)"
+proof(clarsimp)
+  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" and xu: "x<u"
+  and px: "t \<le> x" and ly: "l<y" and yu:"y < u"
+  from tU noU ly yu have tny: "t\<noteq>y" by auto
+  {assume H: "y< t"
+    from less_trans[OF ly H] le_less_trans[OF px xu]
+    have "l < t \<and> t < u" by simp
+    with tU noU have "False" by auto}
+  hence "\<not> y<t"  by auto thus "t \<le> y" by (simp add: not_less)
+qed
+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
+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
+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
+
+lemma lin_dense_conj:
+  "\<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
+  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
+  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
+  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
+  \<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)
+  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<and> P2 y))"
+  by blast
+lemma lin_dense_disj:
+  "\<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
+  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
+  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
+  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
+  \<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)
+  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<or> P2 y))"
+  by blast
+
+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>
+  \<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')"
+by auto
+
+lemma finite_set_intervals:
+  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
+  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"
+  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"
+proof-
+  let ?Mx = "{y. y\<in> S \<and> y \<le> x}"
+  let ?xM = "{y. y\<in> S \<and> x \<le> y}"
+  let ?a = "Max ?Mx"
+  let ?b = "Min ?xM"
+  have MxS: "?Mx \<subseteq> S" by blast
+  hence fMx: "finite ?Mx" using fS finite_subset by auto
+  from lx linS have linMx: "l \<in> ?Mx" by blast
+  hence Mxne: "?Mx \<noteq> {}" by blast
+  have xMS: "?xM \<subseteq> S" by blast
+  hence fxM: "finite ?xM" using fS finite_subset by auto
+  from xu uinS have linxM: "u \<in> ?xM" by blast
+  hence xMne: "?xM \<noteq> {}" by blast
+  have ax:"?a \<le> x" using Mxne fMx by auto
+  have xb:"x \<le> ?b" using xMne fxM by auto
+  have "?a \<in> ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \<in> S" using MxS by blast
+  have "?b \<in> ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \<in> S" using xMS by blast
+  have noy:"\<forall> y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S"
+  proof(clarsimp)
+    fix y   assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S"
+    from yS have "y\<in> ?Mx \<or> y\<in> ?xM" by (auto simp add: linear)
+    moreover {assume "y \<in> ?Mx" hence "y \<le> ?a" using Mxne fMx by auto with ay have "False" by (simp add: not_le[symmetric])}
+    moreover {assume "y \<in> ?xM" hence "?b \<le> y" using xMne fxM by auto with yb have "False" by (simp add: not_le[symmetric])}
+    ultimately show "False" by blast
+  qed
+  from ainS binS noy ax xb px show ?thesis by blast
+qed
+
+lemma finite_set_intervals2:
+  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
+  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"
+  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)"
+proof-
+  from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su]
+  obtain a and b where
+    as: "a\<in> S" and bs: "b\<in> S" and noS:"\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S"
+    and axb: "a \<le> x \<and> x \<le> b \<and> P x"  by auto
+  from axb have "x= a \<or> x= b \<or> (a < x \<and> x < b)" by (auto simp add: le_less)
+  thus ?thesis using px as bs noS by blast
+qed
+
+end
+
+section {* The classical QE after Langford for dense linear orders *}
+
+context dense_linear_order
+begin
+
+lemma interval_empty_iff:
+  "{y. x < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
+  by (auto dest: dense)
+
+lemma dlo_qe_bnds: 
+  assumes ne: "L \<noteq> {}" and neU: "U \<noteq> {}" and fL: "finite L" and fU: "finite U"
+  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)"
+proof (simp only: atomize_eq, rule iffI)
+  assume H: "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)"
+  then obtain x where xL: "\<forall>y\<in>L. y < x" and xU: "\<forall>y\<in>U. x < y" by blast
+  {fix l u assume l: "l \<in> L" and u: "u \<in> U"
+    have "l < x" using xL l by blast
+    also have "x < u" using xU u by blast
+    finally (less_trans) have "l < u" .}
+  thus "\<forall>l\<in>L. \<forall>u\<in>U. l < u" by blast
+next
+  assume H: "\<forall>l\<in>L. \<forall>u\<in>U. l < u"
+  let ?ML = "Max L"
+  let ?MU = "Min U"  
+  from fL ne have th1: "?ML \<in> L" and th1': "\<forall>l\<in>L. l \<le> ?ML" by auto
+  from fU neU have th2: "?MU \<in> U" and th2': "\<forall>u\<in>U. ?MU \<le> u" by auto
+  from th1 th2 H have "?ML < ?MU" by auto
+  with dense obtain w where th3: "?ML < w" and th4: "w < ?MU" by blast
+  from th3 th1' have "\<forall>l \<in> L. l < w" by auto
+  moreover from th4 th2' have "\<forall>u \<in> U. w < u" by auto
+  ultimately show "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)" by auto
+qed
+
+lemma dlo_qe_noub: 
+  assumes ne: "L \<noteq> {}" and fL: "finite L"
+  shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> {}. x < y)) \<equiv> True"
+proof(simp add: atomize_eq)
+  from gt_ex[of "Max L"] obtain M where M: "Max L < M" by blast
+  from ne fL have "\<forall>x \<in> L. x \<le> Max L" by simp
+  with M have "\<forall>x\<in>L. x < M" by (auto intro: le_less_trans)
+  thus "\<exists>x. \<forall>y\<in>L. y < x" by blast
+qed
+
+lemma dlo_qe_nolb: 
+  assumes ne: "U \<noteq> {}" and fU: "finite U"
+  shows "(\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> True"
+proof(simp add: atomize_eq)
+  from lt_ex[of "Min U"] obtain M where M: "M < Min U" by blast
+  from ne fU have "\<forall>x \<in> U. Min U \<le> x" by simp
+  with M have "\<forall>x\<in>U. M < x" by (auto intro: less_le_trans)
+  thus "\<exists>x. \<forall>y\<in>U. x < y" by blast
+qed
+
+lemma exists_neq: "\<exists>(x::'a). x \<noteq> t" "\<exists>(x::'a). t \<noteq> x" 
+  using gt_ex[of t] by auto
+
+lemmas dlo_simps = order_refl less_irrefl not_less not_le exists_neq 
+  le_less neq_iff linear less_not_permute
+
+lemma axiom: "dense_linear_order (op \<le>) (op <)" by (rule dense_linear_order_axioms)
+lemma atoms:
+  shows "TERM (less :: 'a \<Rightarrow> _)"
+    and "TERM (less_eq :: 'a \<Rightarrow> _)"
+    and "TERM (op = :: 'a \<Rightarrow> _)" .
+
+declare axiom[langford qe: dlo_qe_bnds dlo_qe_nolb dlo_qe_noub gather: gather_start gather_simps atoms: atoms]
+declare dlo_simps[langfordsimp]
+
+end
+
+(* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *)
+lemma dnf:
+  "(P & (Q | R)) = ((P&Q) | (P&R))" 
+  "((Q | R) & P) = ((Q&P) | (R&P))"
+  by blast+
+
+lemmas weak_dnf_simps = simp_thms dnf
+
+lemma nnf_simps:
+    "(\<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)"
+    "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
+  by blast+
+
+lemma ex_distrib: "(\<exists>x. P x \<or> Q x) \<longleftrightarrow> ((\<exists>x. P x) \<or> (\<exists>x. Q x))" by blast
+
+lemmas dnf_simps = weak_dnf_simps nnf_simps ex_distrib
+
+use "~~/src/HOL/Tools/Qelim/langford.ML"
+method_setup dlo = {*
+  Method.ctxt_args (Method.SIMPLE_METHOD' o LangfordQE.dlo_tac)
+*} "Langford's algorithm for quantifier elimination in dense linear orders"
+
+
+section {* Contructive dense linear orders yield QE for linear arithmetic over ordered Fields -- see @{text "Arith_Tools.thy"} *}
+
+text {* Linear order without upper bounds *}
+
+locale linorder_stupid_syntax = linorder
+begin
+notation
+  less_eq  ("op \<sqsubseteq>") and
+  less_eq  ("(_/ \<sqsubseteq> _)" [51, 51] 50) and
+  less  ("op \<sqsubset>") and
+  less  ("(_/ \<sqsubset> _)"  [51, 51] 50)
+
+end
+
+locale linorder_no_ub = linorder_stupid_syntax +
+  assumes gt_ex: "\<exists>y. less x y"
+begin
+lemma ge_ex: "\<exists>y. x \<sqsubseteq> y" using gt_ex by auto
+
+text {* Theorems for @{text "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"} *}
+lemma pinf_conj:
+  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
+  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
+  shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
+proof-
+  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
+     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
+  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
+  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
+  {fix x assume H: "z \<sqsubset> x"
+    from less_trans[OF zz1 H] less_trans[OF zz2 H]
+    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
+  }
+  thus ?thesis by blast
+qed
+
+lemma pinf_disj:
+  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
+  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
+  shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
+proof-
+  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
+     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
+  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
+  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
+  {fix x assume H: "z \<sqsubset> x"
+    from less_trans[OF zz1 H] less_trans[OF zz2 H]
+    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
+  }
+  thus ?thesis by blast
+qed
+
+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"
+proof-
+  from ex obtain z where z: "\<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
+  from gt_ex obtain x where x: "z \<sqsubset> x" by blast
+  from z x p1 show ?thesis by blast
+qed
+
+end
+
+text {* Linear order without upper bounds *}
+
+locale linorder_no_lb = linorder_stupid_syntax +
+  assumes lt_ex: "\<exists>y. less y x"
+begin
+lemma le_ex: "\<exists>y. y \<sqsubseteq> x" using lt_ex by auto
+
+
+text {* Theorems for @{text "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"} *}
+lemma minf_conj:
+  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
+  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
+  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
+proof-
+  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
+  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
+  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
+  {fix x assume H: "x \<sqsubset> z"
+    from less_trans[OF H zz1] less_trans[OF H zz2]
+    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
+  }
+  thus ?thesis by blast
+qed
+
+lemma minf_disj:
+  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
+  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
+  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
+proof-
+  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
+  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
+  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
+  {fix x assume H: "x \<sqsubset> z"
+    from less_trans[OF H zz1] less_trans[OF H zz2]
+    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
+  }
+  thus ?thesis by blast
+qed
+
+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"
+proof-
+  from ex obtain z where z: "\<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
+  from lt_ex obtain x where x: "x \<sqsubset> z" by blast
+  from z x p1 show ?thesis by blast
+qed
+
+end
+
+
+locale constr_dense_linear_order = linorder_no_lb + linorder_no_ub +
+  fixes between
+  assumes between_less: "less x y \<Longrightarrow> less x (between x y) \<and> less (between x y) y"
+     and  between_same: "between x x = x"
+
+interpretation  constr_dense_linear_order < dense_linear_order 
+  apply unfold_locales
+  using gt_ex lt_ex between_less
+    by (auto, rule_tac x="between x y" in exI, simp)
+
+context  constr_dense_linear_order
+begin
+
+lemma rinf_U:
+  assumes fU: "finite U"
+  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
+  \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
+  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')"
+  and nmi: "\<not> MP"  and npi: "\<not> PP"  and ex: "\<exists> x.  P x"
+  shows "\<exists> u\<in> U. \<exists> u' \<in> U. P (between u u')"
+proof-
+  from ex obtain x where px: "P x" by blast
+  from px nmi npi nmpiU have "\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u'" by auto
+  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
+  from uU have Une: "U \<noteq> {}" by auto
+  term "linorder.Min less_eq"
+  let ?l = "linorder.Min less_eq U"
+  let ?u = "linorder.Max less_eq U"
+  have linM: "?l \<in> U" using fU Une by simp
+  have uinM: "?u \<in> U" using fU Une by simp
+  have lM: "\<forall> t\<in> U. ?l \<sqsubseteq> t" using Une fU by auto
+  have Mu: "\<forall> t\<in> U. t \<sqsubseteq> ?u" using Une fU by auto
+  have th:"?l \<sqsubseteq> u" using uU Une lM by auto
+  from order_trans[OF th ux] have lx: "?l \<sqsubseteq> x" .
+  have th: "u' \<sqsubseteq> ?u" using uU' Une Mu by simp
+  from order_trans[OF xu' th] have xu: "x \<sqsubseteq> ?u" .
+  from finite_set_intervals2[where P="P",OF px lx xu linM uinM fU lM Mu]
+  have "(\<exists> s\<in> U. P s) \<or>
+      (\<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)" .
+  moreover { fix u assume um: "u\<in>U" and pu: "P u"
+    have "between u u = u" by (simp add: between_same)
+    with um pu have "P (between u u)" by simp
+    with um have ?thesis by blast}
+  moreover{
+    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"
+      then obtain t1 and t2 where t1M: "t1 \<in> U" and t2M: "t2\<in> U"
+        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"
+        by blast
+      from less_trans[OF t1x xt2] have t1t2: "t1 \<sqsubset> t2" .
+      let ?u = "between t1 t2"
+      from between_less t1t2 have t1lu: "t1 \<sqsubset> ?u" and ut2: "?u \<sqsubset> t2" by auto
+      from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast
+      with t1M t2M have ?thesis by blast}
+    ultimately show ?thesis by blast
+  qed
+
+theorem fr_eq:
+  assumes fU: "finite U"
+  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
+   \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
+  and nmibnd: "\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<sqsubseteq> x)"
+  and npibnd: "\<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<sqsubseteq> u)"
+  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)"
+  shows "(\<exists> x. P x) \<equiv> (MP \<or> PP \<or> (\<exists> u \<in> U. \<exists> u'\<in> U. P (between u u')))"
+  (is "_ \<equiv> (_ \<or> _ \<or> ?F)" is "?E \<equiv> ?D")
+proof-
+ {
+   assume px: "\<exists> x. P x"
+   have "MP \<or> PP \<or> (\<not> MP \<and> \<not> PP)" by blast
+   moreover {assume "MP \<or> PP" hence "?D" by blast}
+   moreover {assume nmi: "\<not> MP" and npi: "\<not> PP"
+     from npmibnd[OF nmibnd npibnd]
+     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')" .
+     from rinf_U[OF fU lin_dense nmpiU nmi npi px] have "?D" by blast}
+   ultimately have "?D" by blast}
+ moreover
+ { assume "?D"
+   moreover {assume m:"MP" from minf_ex[OF mi m] have "?E" .}
+   moreover {assume p: "PP" from pinf_ex[OF pi p] have "?E" . }
+   moreover {assume f:"?F" hence "?E" by blast}
+   ultimately have "?E" by blast}
+ ultimately have "?E = ?D" by blast thus "?E \<equiv> ?D" by simp
+qed
+
+lemmas minf_thms = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P
+lemmas pinf_thms = pinf_conj pinf_disj pinf_eq pinf_neq pinf_lt pinf_le pinf_gt pinf_ge pinf_P
+
+lemmas nmi_thms = nmi_conj nmi_disj nmi_eq nmi_neq nmi_lt nmi_le nmi_gt nmi_ge nmi_P
+lemmas npi_thms = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P
+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
+
+lemma ferrack_axiom: "constr_dense_linear_order less_eq less between"
+  by (rule constr_dense_linear_order_axioms)
+lemma atoms:
+  shows "TERM (less :: 'a \<Rightarrow> _)"
+    and "TERM (less_eq :: 'a \<Rightarrow> _)"
+    and "TERM (op = :: 'a \<Rightarrow> _)" .
+
+declare ferrack_axiom [ferrack minf: minf_thms pinf: pinf_thms
+    nmi: nmi_thms npi: npi_thms lindense:
+    lin_dense_thms qe: fr_eq atoms: atoms]
+
+declaration {*
+let
+fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
+fun generic_whatis phi =
+ let
+  val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
+  fun h x t =
+   case term_of t of
+     Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
+                            else Ferrante_Rackoff_Data.Nox
+   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
+                            else Ferrante_Rackoff_Data.Nox
+   | b$y$z => if Term.could_unify (b, lt) then
+                 if term_of x aconv y then Ferrante_Rackoff_Data.Lt
+                 else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
+                 else Ferrante_Rackoff_Data.Nox
+             else if Term.could_unify (b, le) then
+                 if term_of x aconv y then Ferrante_Rackoff_Data.Le
+                 else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
+                 else Ferrante_Rackoff_Data.Nox
+             else Ferrante_Rackoff_Data.Nox
+   | _ => Ferrante_Rackoff_Data.Nox
+ in h end
+ fun ss phi = HOL_ss addsimps (simps phi)
+in
+ Ferrante_Rackoff_Data.funs  @{thm "ferrack_axiom"}
+  {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss}
+end
+*}
+
+end
+
+use "~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML"
+
+method_setup ferrack = {*
+  Method.ctxt_args (Method.SIMPLE_METHOD' o FerranteRackoff.dlo_tac)
+*} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders"
+
+subsection {* Ferrante and Rackoff algorithm over ordered fields *}
+
+lemma neg_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x < 0) == (x > 0))"
+proof-
+  assume H: "c < 0"
+  have "c*x < 0 = (0/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
+  also have "\<dots> = (0 < x)" by simp
+  finally show  "(c*x < 0) == (x > 0)" by simp
+qed
+
+lemma pos_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x < 0) == (x < 0))"
+proof-
+  assume H: "c > 0"
+  hence "c*x < 0 = (0/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
+  also have "\<dots> = (0 > x)" by simp
+  finally show  "(c*x < 0) == (x < 0)" by simp
+qed
+
+lemma neg_prod_sum_lt: "(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x + t< 0) == (x > (- 1/c)*t))"
+proof-
+  assume H: "c < 0"
+  have "c*x + t< 0 = (c*x < -t)" by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (-t/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
+  also have "\<dots> = ((- 1/c)*t < x)" by simp
+  finally show  "(c*x + t < 0) == (x > (- 1/c)*t)" by simp
+qed
+
+lemma pos_prod_sum_lt:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x + t < 0) == (x < (- 1/c)*t))"
+proof-
+  assume H: "c > 0"
+  have "c*x + t< 0 = (c*x < -t)"  by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (-t/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
+  also have "\<dots> = ((- 1/c)*t > x)" by simp
+  finally show  "(c*x + t < 0) == (x < (- 1/c)*t)" by simp
+qed
+
+lemma sum_lt:"((x::'a::pordered_ab_group_add) + t < 0) == (x < - t)"
+  using less_diff_eq[where a= x and b=t and c=0] by simp
+
+lemma neg_prod_le:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x <= 0) == (x >= 0))"
+proof-
+  assume H: "c < 0"
+  have "c*x <= 0 = (0/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
+  also have "\<dots> = (0 <= x)" by simp
+  finally show  "(c*x <= 0) == (x >= 0)" by simp
+qed
+
+lemma pos_prod_le:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x <= 0) == (x <= 0))"
+proof-
+  assume H: "c > 0"
+  hence "c*x <= 0 = (0/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
+  also have "\<dots> = (0 >= x)" by simp
+  finally show  "(c*x <= 0) == (x <= 0)" by simp
+qed
+
+lemma neg_prod_sum_le: "(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x + t <= 0) == (x >= (- 1/c)*t))"
+proof-
+  assume H: "c < 0"
+  have "c*x + t <= 0 = (c*x <= -t)"  by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (-t/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
+  also have "\<dots> = ((- 1/c)*t <= x)" by simp
+  finally show  "(c*x + t <= 0) == (x >= (- 1/c)*t)" by simp
+qed
+
+lemma pos_prod_sum_le:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x + t <= 0) == (x <= (- 1/c)*t))"
+proof-
+  assume H: "c > 0"
+  have "c*x + t <= 0 = (c*x <= -t)" by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (-t/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
+  also have "\<dots> = ((- 1/c)*t >= x)" by simp
+  finally show  "(c*x + t <= 0) == (x <= (- 1/c)*t)" by simp
+qed
+
+lemma sum_le:"((x::'a::pordered_ab_group_add) + t <= 0) == (x <= - t)"
+  using le_diff_eq[where a= x and b=t and c=0] by simp
+
+lemma nz_prod_eq:"(c\<Colon>'a\<Colon>ordered_field) \<noteq> 0 \<Longrightarrow> ((c*x = 0) == (x = 0))" by simp
+lemma nz_prod_sum_eq: "(c\<Colon>'a\<Colon>ordered_field) \<noteq> 0 \<Longrightarrow> ((c*x + t = 0) == (x = (- 1/c)*t))"
+proof-
+  assume H: "c \<noteq> 0"
+  have "c*x + t = 0 = (c*x = -t)" by (subst eq_iff_diff_eq_0 [of "c*x" "-t"], simp)
+  also have "\<dots> = (x = -t/c)" by (simp only: nonzero_eq_divide_eq[OF H] ring_simps)
+  finally show  "(c*x + t = 0) == (x = (- 1/c)*t)" by simp
+qed
+lemma sum_eq:"((x::'a::pordered_ab_group_add) + t = 0) == (x = - t)"
+  using eq_diff_eq[where a= x and b=t and c=0] by simp
+
+
+interpretation class_ordered_field_dense_linear_order: constr_dense_linear_order
+ ["op <=" "op <"
+   "\<lambda> x y. 1/2 * ((x::'a::{ordered_field,recpower,number_ring}) + y)"]
+proof (unfold_locales, dlo, dlo, auto)
+  fix x y::'a assume lt: "x < y"
+  from  less_half_sum[OF lt] show "x < (x + y) /2" by simp
+next
+  fix x y::'a assume lt: "x < y"
+  from  gt_half_sum[OF lt] show "(x + y) /2 < y" by simp
+qed
+
+declaration{*
+let
+fun earlier [] x y = false
+        | earlier (h::t) x y =
+    if h aconvc y then false else if h aconvc x then true else earlier t x y;
+
+fun dest_frac ct = case term_of ct of
+   Const (@{const_name "HOL.divide"},_) $ a $ b=>
+    Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b))
+ | t => Rat.rat_of_int (snd (HOLogic.dest_number t))
+
+fun mk_frac phi cT x =
+ let val (a, b) = Rat.quotient_of_rat x
+ in if b = 1 then Numeral.mk_cnumber cT a
+    else Thm.capply
+         (Thm.capply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"})
+                     (Numeral.mk_cnumber cT a))
+         (Numeral.mk_cnumber cT b)
+ end
+
+fun whatis x ct = case term_of ct of
+  Const(@{const_name "HOL.plus"}, _)$(Const(@{const_name "HOL.times"},_)$_$y)$_ =>
+     if y aconv term_of x then ("c*x+t",[(funpow 2 Thm.dest_arg1) ct, Thm.dest_arg ct])
+     else ("Nox",[])
+| Const(@{const_name "HOL.plus"}, _)$y$_ =>
+     if y aconv term_of x then ("x+t",[Thm.dest_arg ct])
+     else ("Nox",[])
+| Const(@{const_name "HOL.times"}, _)$_$y =>
+     if y aconv term_of x then ("c*x",[Thm.dest_arg1 ct])
+     else ("Nox",[])
+| t => if t aconv term_of x then ("x",[]) else ("Nox",[]);
+
+fun xnormalize_conv ctxt [] ct = reflexive ct
+| xnormalize_conv ctxt (vs as (x::_)) ct =
+   case term_of ct of
+   Const(@{const_name HOL.less},_)$_$Const(@{const_name "HOL.zero"},_) =>
+    (case whatis x (Thm.dest_arg1 ct) of
+    ("c*x+t",[c,t]) =>
+       let
+        val cr = dest_frac c
+        val clt = Thm.dest_fun2 ct
+        val cz = Thm.dest_arg ct
+        val neg = cr </ Rat.zero
+        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
+               (Thm.capply @{cterm "Trueprop"}
+                  (if neg then Thm.capply (Thm.capply clt c) cz
+                    else Thm.capply (Thm.capply clt cz) c))
+        val cth = equal_elim (symmetric cthp) TrueI
+        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x,t])
+             (if neg then @{thm neg_prod_sum_lt} else @{thm pos_prod_sum_lt})) cth
+        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
+                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
+      in rth end
+    | ("x+t",[t]) =>
+       let
+        val T = ctyp_of_term x
+        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_lt"}
+        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
+              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
+       in  rth end
+    | ("c*x",[c]) =>
+       let
+        val cr = dest_frac c
+        val clt = Thm.dest_fun2 ct
+        val cz = Thm.dest_arg ct
+        val neg = cr </ Rat.zero
+        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
+               (Thm.capply @{cterm "Trueprop"}
+                  (if neg then Thm.capply (Thm.capply clt c) cz
+                    else Thm.capply (Thm.capply clt cz) c))
+        val cth = equal_elim (symmetric cthp) TrueI
+        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
+             (if neg then @{thm neg_prod_lt} else @{thm pos_prod_lt})) cth
+        val rth = th
+      in rth end
+    | _ => reflexive ct)
+
+
+|  Const(@{const_name HOL.less_eq},_)$_$Const(@{const_name "HOL.zero"},_) =>
+   (case whatis x (Thm.dest_arg1 ct) of
+    ("c*x+t",[c,t]) =>
+       let
+        val T = ctyp_of_term x
+        val cr = dest_frac c
+        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
+        val cz = Thm.dest_arg ct
+        val neg = cr </ Rat.zero
+        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
+               (Thm.capply @{cterm "Trueprop"}
+                  (if neg then Thm.capply (Thm.capply clt c) cz
+                    else Thm.capply (Thm.capply clt cz) c))
+        val cth = equal_elim (symmetric cthp) TrueI
+        val th = implies_elim (instantiate' [SOME T] (map SOME [c,x,t])
+             (if neg then @{thm neg_prod_sum_le} else @{thm pos_prod_sum_le})) cth
+        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
+                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
+      in rth end
+    | ("x+t",[t]) =>
+       let
+        val T = ctyp_of_term x
+        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_le"}
+        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
+              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
+       in  rth end
+    | ("c*x",[c]) =>
+       let
+        val T = ctyp_of_term x
+        val cr = dest_frac c
+        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
+        val cz = Thm.dest_arg ct
+        val neg = cr </ Rat.zero
+        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
+               (Thm.capply @{cterm "Trueprop"}
+                  (if neg then Thm.capply (Thm.capply clt c) cz
+                    else Thm.capply (Thm.capply clt cz) c))
+        val cth = equal_elim (symmetric cthp) TrueI
+        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
+             (if neg then @{thm neg_prod_le} else @{thm pos_prod_le})) cth
+        val rth = th
+      in rth end
+    | _ => reflexive ct)
+
+|  Const("op =",_)$_$Const(@{const_name "HOL.zero"},_) =>
+   (case whatis x (Thm.dest_arg1 ct) of
+    ("c*x+t",[c,t]) =>
+       let
+        val T = ctyp_of_term x
+        val cr = dest_frac c
+        val ceq = Thm.dest_fun2 ct
+        val cz = Thm.dest_arg ct
+        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
+            (Thm.capply @{cterm "Trueprop"}
+             (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
+        val cth = equal_elim (symmetric cthp) TrueI
+        val th = implies_elim
+                 (instantiate' [SOME T] (map SOME [c,x,t]) @{thm nz_prod_sum_eq}) cth
+        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
+                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
+      in rth end
+    | ("x+t",[t]) =>
+       let
+        val T = ctyp_of_term x
+        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_eq"}
+        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
+              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
+       in  rth end
+    | ("c*x",[c]) =>
+       let
+        val T = ctyp_of_term x
+        val cr = dest_frac c
+        val ceq = Thm.dest_fun2 ct
+        val cz = Thm.dest_arg ct
+        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
+            (Thm.capply @{cterm "Trueprop"}
+             (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
+        val cth = equal_elim (symmetric cthp) TrueI
+        val rth = implies_elim
+                 (instantiate' [SOME T] (map SOME [c,x]) @{thm nz_prod_eq}) cth
+      in rth end
+    | _ => reflexive ct);
+
+local
+  val less_iff_diff_less_0 = mk_meta_eq @{thm "less_iff_diff_less_0"}
+  val le_iff_diff_le_0 = mk_meta_eq @{thm "le_iff_diff_le_0"}
+  val eq_iff_diff_eq_0 = mk_meta_eq @{thm "eq_iff_diff_eq_0"}
+in
+fun field_isolate_conv phi ctxt vs ct = case term_of ct of
+  Const(@{const_name HOL.less},_)$a$b =>
+   let val (ca,cb) = Thm.dest_binop ct
+       val T = ctyp_of_term ca
+       val th = instantiate' [SOME T] [SOME ca, SOME cb] less_iff_diff_less_0
+       val nth = Conv.fconv_rule
+         (Conv.arg_conv (Conv.arg1_conv
+              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
+       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
+   in rth end
+| Const(@{const_name HOL.less_eq},_)$a$b =>
+   let val (ca,cb) = Thm.dest_binop ct
+       val T = ctyp_of_term ca
+       val th = instantiate' [SOME T] [SOME ca, SOME cb] le_iff_diff_le_0
+       val nth = Conv.fconv_rule
+         (Conv.arg_conv (Conv.arg1_conv
+              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
+       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
+   in rth end
+
+| Const("op =",_)$a$b =>
+   let val (ca,cb) = Thm.dest_binop ct
+       val T = ctyp_of_term ca
+       val th = instantiate' [SOME T] [SOME ca, SOME cb] eq_iff_diff_eq_0
+       val nth = Conv.fconv_rule
+         (Conv.arg_conv (Conv.arg1_conv
+              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
+       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
+   in rth end
+| @{term "Not"} $(Const("op =",_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
+| _ => reflexive ct
+end;
+
+fun classfield_whatis phi =
+ let
+  fun h x t =
+   case term_of t of
+     Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
+                            else Ferrante_Rackoff_Data.Nox
+   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
+                            else Ferrante_Rackoff_Data.Nox
+   | Const(@{const_name HOL.less},_)$y$z =>
+       if term_of x aconv y then Ferrante_Rackoff_Data.Lt
+        else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
+        else Ferrante_Rackoff_Data.Nox
+   | Const (@{const_name HOL.less_eq},_)$y$z =>
+         if term_of x aconv y then Ferrante_Rackoff_Data.Le
+         else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
+         else Ferrante_Rackoff_Data.Nox
+   | _ => Ferrante_Rackoff_Data.Nox
+ in h end;
+fun class_field_ss phi =
+   HOL_basic_ss addsimps ([@{thm "linorder_not_less"}, @{thm "linorder_not_le"}])
+   addsplits [@{thm "abs_split"},@{thm "split_max"}, @{thm "split_min"}]
+
+in
+Ferrante_Rackoff_Data.funs @{thm "class_ordered_field_dense_linear_order.ferrack_axiom"}
+  {isolate_conv = field_isolate_conv, whatis = classfield_whatis, simpset = class_field_ss}
+end
+*}
+
+
+end 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Fundamental_Theorem_Algebra.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,1327 @@
+(* Author: Amine Chaieb, TU Muenchen *)
+
+header{*Fundamental Theorem of Algebra*}
+
+theory Fundamental_Theorem_Algebra
+imports Univ_Poly Dense_Linear_Order Complex
+begin
+
+subsection {* Square root of complex numbers *}
+definition csqrt :: "complex \<Rightarrow> complex" where
+"csqrt z = (if Im z = 0 then
+            if 0 \<le> Re z then Complex (sqrt(Re z)) 0
+            else Complex 0 (sqrt(- Re z))
+           else Complex (sqrt((cmod z + Re z) /2))
+                        ((Im z / abs(Im z)) * sqrt((cmod z - Re z) /2)))"
+
+lemma csqrt[algebra]: "csqrt z ^ 2 = z"
+proof-
+  obtain x y where xy: "z = Complex x y" by (cases z, simp_all)
+  {assume y0: "y = 0"
+    {assume x0: "x \<ge> 0" 
+      then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
+	by (simp add: csqrt_def power2_eq_square)}
+    moreover
+    {assume "\<not> x \<ge> 0" hence x0: "- x \<ge> 0" by arith
+      then have ?thesis using y0 xy real_sqrt_pow2[OF x0] 
+	by (simp add: csqrt_def power2_eq_square) }
+    ultimately have ?thesis by blast}
+  moreover
+  {assume y0: "y\<noteq>0"
+    {fix x y
+      let ?z = "Complex x y"
+      from abs_Re_le_cmod[of ?z] have tha: "abs x \<le> cmod ?z" by auto
+      hence "cmod ?z - x \<ge> 0" "cmod ?z + x \<ge> 0" by arith+ 
+      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) }
+    note th = this
+    have sq4: "\<And>x::real. x^2 / 4 = (x / 2) ^ 2" 
+      by (simp add: power2_eq_square) 
+    from th[of x y]
+    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
+    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"
+      unfolding power2_eq_square by simp 
+    have "sqrt 4 = sqrt (2^2)" by simp 
+    hence sqrt4: "sqrt 4 = 2" by (simp only: real_sqrt_abs)
+    have th2: "2 *(y * sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) + x) / 4)) / \<bar>y\<bar> = y"
+      using iffD2[OF real_sqrt_pow2_iff sum_power2_ge_zero[of x y]] y0
+      unfolding power2_eq_square 
+      by (simp add: ring_simps real_sqrt_divide sqrt4)
+     from y0 xy have ?thesis  apply (simp add: csqrt_def power2_eq_square)
+       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])
+      using th1 th2  ..}
+  ultimately show ?thesis by blast
+qed
+
+
+subsection{* More lemmas about module of complex numbers *}
+
+lemma complex_of_real_power: "complex_of_real x ^ n = complex_of_real (x^n)"
+  by (rule of_real_power [symmetric])
+
+lemma real_down2: "(0::real) < d1 \<Longrightarrow> 0 < d2 ==> EX e. 0 < e & e < d1 & e < d2"
+  apply ferrack apply arith done
+
+text{* The triangle inequality for cmod *}
+lemma complex_mod_triangle_sub: "cmod w \<le> cmod (w + z) + norm z"
+  using complex_mod_triangle_ineq2[of "w + z" "-z"] by auto
+
+subsection{* Basic lemmas about complex polynomials *}
+
+lemma poly_bound_exists:
+  shows "\<exists>m. m > 0 \<and> (\<forall>z. cmod z <= r \<longrightarrow> cmod (poly p z) \<le> m)"
+proof(induct p)
+  case Nil thus ?case by (rule exI[where x=1], simp) 
+next
+  case (Cons c cs)
+  from Cons.hyps obtain m where m: "\<forall>z. cmod z \<le> r \<longrightarrow> cmod (poly cs z) \<le> m"
+    by blast
+  let ?k = " 1 + cmod c + \<bar>r * m\<bar>"
+  have kp: "?k > 0" using abs_ge_zero[of "r*m"] norm_ge_zero[of c] by arith
+  {fix z
+    assume H: "cmod z \<le> r"
+    from m H have th: "cmod (poly cs z) \<le> m" by blast
+    from H have rp: "r \<ge> 0" using norm_ge_zero[of z] by arith
+    have "cmod (poly (c # cs) z) \<le> cmod c + cmod (z* poly cs z)"
+      using norm_triangle_ineq[of c "z* poly cs z"] by simp
+    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)
+    also have "\<dots> \<le> ?k" by simp
+    finally have "cmod (poly (c # cs) z) \<le> ?k" .}
+  with kp show ?case by blast
+qed
+
+
+text{* Offsetting the variable in a polynomial gives another of same degree *}
+  (* FIXME : Lemma holds also in locale --- fix it later *)
+lemma  poly_offset_lemma:
+  shows "\<exists>b q. (length q = length p) \<and> (\<forall>x. poly (b#q) (x::complex) = (a + x) * poly p x)"
+proof(induct p)
+  case Nil thus ?case by simp
+next
+  case (Cons c cs)
+  from Cons.hyps obtain b q where 
+    bq: "length q = length cs" "\<forall>x. poly (b # q) x = (a + x) * poly cs x"
+    by blast
+  let ?b = "a*c"
+  let ?q = "(b+c)#q"
+  have lg: "length ?q = length (c#cs)" using bq(1) by simp
+  {fix x
+    from bq(2)[rule_format, of x]
+    have "x*poly (b # q) x = x*((a + x) * poly cs x)" by simp
+    hence "poly (?b# ?q) x = (a + x) * poly (c # cs) x"
+      by (simp add: ring_simps)}
+  with lg  show ?case by blast 
+qed
+
+    (* FIXME : This one too*)
+lemma poly_offset: "\<exists> q. length q = length p \<and> (\<forall>x. poly q (x::complex) = poly p (a + x))"
+proof (induct p)
+  case Nil thus ?case by simp
+next
+  case (Cons c cs)
+  from Cons.hyps obtain q where q: "length q = length cs" "\<forall>x. poly q x = poly cs (a + x)" by blast
+  from poly_offset_lemma[of q a] obtain b p where 
+    bp: "length p = length q" "\<forall>x. poly (b # p) x = (a + x) * poly q x"
+    by blast
+  thus ?case using q bp by - (rule exI[where x="(c + b)#p"], simp)
+qed
+
+text{* An alternative useful formulation of completeness of the reals *}
+lemma real_sup_exists: assumes ex: "\<exists>x. P x" and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z"
+  shows "\<exists>(s::real). \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s"
+proof-
+  from ex bz obtain x Y where x: "P x" and Y: "\<And>x. P x \<Longrightarrow> x < Y"  by blast
+  from ex have thx:"\<exists>x. x \<in> Collect P" by blast
+  from bz have thY: "\<exists>Y. isUb UNIV (Collect P) Y" 
+    by(auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def order_le_less)
+  from reals_complete[OF thx thY] obtain L where L: "isLub UNIV (Collect P) L"
+    by blast
+  from Y[OF x] have xY: "x < Y" .
+  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)  
+  from Y have Y': "\<forall>x. P x \<longrightarrow> x \<le> Y" 
+    apply (clarsimp, atomize (full)) by auto 
+  from L Y' have "L \<le> Y" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
+  {fix y
+    {fix z assume z: "P z" "y < z"
+      from L' z have "y < L" by auto }
+    moreover
+    {assume yL: "y < L" "\<forall>z. P z \<longrightarrow> \<not> y < z"
+      hence nox: "\<forall>z. P z \<longrightarrow> y \<ge> z" by auto
+      from nox L have "y \<ge> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def) 
+      with yL(1) have False  by arith}
+    ultimately have "(\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < L" by blast}
+  thus ?thesis by blast
+qed
+
+
+subsection{* Some theorems about Sequences*}
+text{* Given a binary function @{text "f:: nat \<Rightarrow> 'a \<Rightarrow> 'a"}, its values are uniquely determined by a function g *}
+
+lemma num_Axiom: "EX! g. g 0 = e \<and> (\<forall>n. g (Suc n) = f n (g n))"
+  unfolding Ex1_def
+  apply (rule_tac x="nat_rec e f" in exI)
+  apply (rule conjI)+
+apply (rule def_nat_rec_0, simp)
+apply (rule allI, rule def_nat_rec_Suc, simp)
+apply (rule allI, rule impI, rule ext)
+apply (erule conjE)
+apply (induct_tac x)
+apply (simp add: nat_rec_0)
+apply (erule_tac x="n" in allE)
+apply (simp)
+done
+
+ text{* An equivalent formulation of monotony -- Not used here, but might be useful *}
+lemma mono_Suc: "mono f = (\<forall>n. (f n :: 'a :: order) \<le> f (Suc n))"
+unfolding mono_def
+proof auto
+  fix A B :: nat
+  assume H: "\<forall>n. f n \<le> f (Suc n)" "A \<le> B"
+  hence "\<exists>k. B = A + k" apply -  apply (thin_tac "\<forall>n. f n \<le> f (Suc n)") 
+    by presburger
+  then obtain k where k: "B = A + k" by blast
+  {fix a k
+    have "f a \<le> f (a + k)"
+    proof (induct k)
+      case 0 thus ?case by simp
+    next
+      case (Suc k)
+      from Suc.hyps H(1)[rule_format, of "a + k"] show ?case by simp
+    qed}
+  with k show "f A \<le> f B" by blast
+qed
+
+text{* for any sequence, there is a mootonic subsequence *}
+lemma seq_monosub: "\<exists>f. subseq f \<and> monoseq (\<lambda> n. (s (f n)))"
+proof-
+  {assume H: "\<forall>n. \<exists>p >n. \<forall> m\<ge>p. s m \<le> s p"
+    let ?P = "\<lambda> p n. p > n \<and> (\<forall>m \<ge> p. s m \<le> s p)"
+    from num_Axiom[of "SOME p. ?P p 0" "\<lambda>p n. SOME p. ?P p n"]
+    obtain f where f: "f 0 = (SOME p. ?P p 0)" "\<forall>n. f (Suc n) = (SOME p. ?P p (f n))" by blast
+    have "?P (f 0) 0"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p 0"]
+      using H apply - 
+      apply (erule allE[where x=0], erule exE, rule_tac x="p" in exI) 
+      unfolding order_le_less by blast 
+    hence f0: "f 0 > 0" "\<forall>m \<ge> f 0. s m \<le> s (f 0)" by blast+
+    {fix n
+      have "?P (f (Suc n)) (f n)" 
+	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
+	using H apply - 
+      apply (erule allE[where x="f n"], erule exE, rule_tac x="p" in exI) 
+      unfolding order_le_less by blast 
+    hence "f (Suc n) > f n" "\<forall>m \<ge> f (Suc n). s m \<le> s (f (Suc n))" by blast+}
+  note fSuc = this
+    {fix p q assume pq: "p \<ge> f q"
+      have "s p \<le> s(f(q))"  using f0(2)[rule_format, of p] pq fSuc
+	by (cases q, simp_all) }
+    note pqth = this
+    {fix q
+      have "f (Suc q) > f q" apply (induct q) 
+	using f0(1) fSuc(1)[of 0] apply simp by (rule fSuc(1))}
+    note fss = this
+    from fss have th1: "subseq f" unfolding subseq_Suc_iff ..
+    {fix a b 
+      have "f a \<le> f (a + b)"
+      proof(induct b)
+	case 0 thus ?case by simp
+      next
+	case (Suc b)
+	from fSuc(1)[of "a + b"] Suc.hyps show ?case by simp
+      qed}
+    note fmon0 = this
+    have "monoseq (\<lambda>n. s (f n))" 
+    proof-
+      {fix n
+	have "s (f n) \<ge> s (f (Suc n))" 
+	proof(cases n)
+	  case 0
+	  assume n0: "n = 0"
+	  from fSuc(1)[of 0] have th0: "f 0 \<le> f (Suc 0)" by simp
+	  from f0(2)[rule_format, OF th0] show ?thesis  using n0 by simp
+	next
+	  case (Suc m)
+	  assume m: "n = Suc m"
+	  from fSuc(1)[of n] m have th0: "f (Suc m) \<le> f (Suc (Suc m))" by simp
+	  from m fSuc(2)[rule_format, OF th0] show ?thesis by simp 
+	qed}
+      thus "monoseq (\<lambda>n. s (f n))" unfolding monoseq_Suc by blast 
+    qed
+    with th1 have ?thesis by blast}
+  moreover
+  {fix N assume N: "\<forall>p >N. \<exists> m\<ge>p. s m > s p"
+    {fix p assume p: "p \<ge> Suc N" 
+      hence pN: "p > N" by arith with N obtain m where m: "m \<ge> p" "s m > s p" by blast
+      have "m \<noteq> p" using m(2) by auto 
+      with m have "\<exists>m>p. s p < s m" by - (rule exI[where x=m], auto)}
+    note th0 = this
+    let ?P = "\<lambda>m x. m > x \<and> s x < s m"
+    from num_Axiom[of "SOME x. ?P x (Suc N)" "\<lambda>m x. SOME y. ?P y x"]
+    obtain f where f: "f 0 = (SOME x. ?P x (Suc N))" 
+      "\<forall>n. f (Suc n) = (SOME m. ?P m (f n))" by blast
+    have "?P (f 0) (Suc N)"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p (Suc N)"]
+      using N apply - 
+      apply (erule allE[where x="Suc N"], clarsimp)
+      apply (rule_tac x="m" in exI)
+      apply auto
+      apply (subgoal_tac "Suc N \<noteq> m")
+      apply simp
+      apply (rule ccontr, simp)
+      done
+    hence f0: "f 0 > Suc N" "s (Suc N) < s (f 0)" by blast+
+    {fix n
+      have "f n > N \<and> ?P (f (Suc n)) (f n)"
+	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
+      proof (induct n)
+	case 0 thus ?case
+	  using f0 N apply auto 
+	  apply (erule allE[where x="f 0"], clarsimp) 
+	  apply (rule_tac x="m" in exI, simp)
+	  by (subgoal_tac "f 0 \<noteq> m", auto)
+      next
+	case (Suc n)
+	from Suc.hyps have Nfn: "N < f n" by blast
+	from Suc.hyps obtain m where m: "m > f n" "s (f n) < s m" by blast
+	with Nfn have mN: "m > N" by arith
+	note key = Suc.hyps[unfolded some_eq_ex[of "\<lambda>p. ?P p (f n)", symmetric] f(2)[rule_format, of n, symmetric]]
+	
+	from key have th0: "f (Suc n) > N" by simp
+	from N[rule_format, OF th0]
+	obtain m' where m': "m' \<ge> f (Suc n)" "s (f (Suc n)) < s m'" by blast
+	have "m' \<noteq> f (Suc (n))" apply (rule ccontr) using m'(2) by auto
+	hence "m' > f (Suc n)" using m'(1) by simp
+	with key m'(2) show ?case by auto
+      qed}
+    note fSuc = this
+    {fix n
+      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 
+      hence "f n \<ge> Suc N" "f(Suc n) > f n" "s(f n) < s(f(Suc n))" by blast+}
+    note thf = this
+    have sqf: "subseq f" unfolding subseq_Suc_iff using thf by simp
+    have "monoseq (\<lambda>n. s (f n))"  unfolding monoseq_Suc using thf
+      apply -
+      apply (rule disjI1)
+      apply auto
+      apply (rule order_less_imp_le)
+      apply blast
+      done
+    then have ?thesis  using sqf by blast}
+  ultimately show ?thesis unfolding linorder_not_less[symmetric] by blast
+qed
+
+lemma seq_suble: assumes sf: "subseq f" shows "n \<le> f n"
+proof(induct n)
+  case 0 thus ?case by simp
+next
+  case (Suc n)
+  from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps
+  have "n < f (Suc n)" by arith 
+  thus ?case by arith
+qed
+
+subsection {* Fundamental theorem of algebra *}
+lemma  unimodular_reduce_norm:
+  assumes md: "cmod z = 1"
+  shows "cmod (z + 1) < 1 \<or> cmod (z - 1) < 1 \<or> cmod (z + ii) < 1 \<or> cmod (z - ii) < 1"
+proof-
+  obtain x y where z: "z = Complex x y " by (cases z, auto)
+  from md z have xy: "x^2 + y^2 = 1" by (simp add: cmod_def)
+  {assume C: "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + ii) \<ge> 1" "cmod (z - ii) \<ge> 1"
+    from C z xy have "2*x \<le> 1" "2*x \<ge> -1" "2*y \<le> 1" "2*y \<ge> -1"
+      by (simp_all add: cmod_def power2_eq_square ring_simps)
+    hence "abs (2*x) \<le> 1" "abs (2*y) \<le> 1" by simp_all
+    hence "(abs (2 * x))^2 <= 1^2" "(abs (2 * y)) ^2 <= 1^2"
+      by - (rule power_mono, simp, simp)+
+    hence th0: "4*x^2 \<le> 1" "4*y^2 \<le> 1" 
+      by (simp_all  add: power2_abs power_mult_distrib)
+    from add_mono[OF th0] xy have False by simp }
+  thus ?thesis unfolding linorder_not_le[symmetric] by blast
+qed
+
+text{* Hence we can always reduce modulus of @{text "1 + b z^n"} if nonzero *}
+lemma reduce_poly_simple:
+ assumes b: "b \<noteq> 0" and n: "n\<noteq>0"
+  shows "\<exists>z. cmod (1 + b * z^n) < 1"
+using n
+proof(induct n rule: nat_less_induct)
+  fix n
+  assume IH: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>z. cmod (1 + b * z ^ m) < 1)" and n: "n \<noteq> 0"
+  let ?P = "\<lambda>z n. cmod (1 + b * z ^ n) < 1"
+  {assume e: "even n"
+    hence "\<exists>m. n = 2*m" by presburger
+    then obtain m where m: "n = 2*m" by blast
+    from n m have "m\<noteq>0" "m < n" by presburger+
+    with IH[rule_format, of m] obtain z where z: "?P z m" by blast
+    from z have "?P (csqrt z) n" by (simp add: m power_mult csqrt)
+    hence "\<exists>z. ?P z n" ..}
+  moreover
+  {assume o: "odd n"
+    from b have b': "b^2 \<noteq> 0" unfolding power2_eq_square by simp
+    have "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
+    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) = 
+    ((Re (inverse b))^2 + (Im (inverse b))^2) * \<bar>Im b * Im b + Re b * Re b\<bar>" by algebra
+    also have "\<dots> = cmod (inverse b) ^2 * cmod b ^ 2" 
+      apply (simp add: cmod_def) using realpow_two_le_add_order[of "Re b" "Im b"]
+      by (simp add: power2_eq_square)
+    finally 
+    have th0: "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
+    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) =
+    1" 
+      apply (simp add: power2_eq_square norm_mult[symmetric] norm_inverse[symmetric])
+      using right_inverse[OF b']
+      by (simp add: power2_eq_square[symmetric] power_inverse[symmetric] ring_simps)
+    have th0: "cmod (complex_of_real (cmod b) / b) = 1"
+      apply (simp add: complex_Re_mult cmod_def power2_eq_square Re_complex_of_real Im_complex_of_real divide_inverse ring_simps )
+      by (simp add: real_sqrt_mult[symmetric] th0)        
+    from o have "\<exists>m. n = Suc (2*m)" by presburger+
+    then obtain m where m: "n = Suc (2*m)" by blast
+    from unimodular_reduce_norm[OF th0] o
+    have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
+      apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1", rule_tac x="1" in exI, simp)
+      apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp add: diff_def)
+      apply (cases "cmod (complex_of_real (cmod b) / b + ii) < 1")
+      apply (cases "even m", rule_tac x="ii" in exI, simp add: m power_mult)
+      apply (rule_tac x="- ii" in exI, simp add: m power_mult)
+      apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult diff_def)
+      apply (rule_tac x="ii" in exI, simp add: m power_mult diff_def)
+      done
+    then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" by blast
+    let ?w = "v / complex_of_real (root n (cmod b))"
+    from odd_real_root_pow[OF o, of "cmod b"]
+    have th1: "?w ^ n = v^n / complex_of_real (cmod b)" 
+      by (simp add: power_divide complex_of_real_power)
+    have th2:"cmod (complex_of_real (cmod b) / b) = 1" using b by (simp add: norm_divide)
+    hence th3: "cmod (complex_of_real (cmod b) / b) \<ge> 0" by simp
+    have th4: "cmod (complex_of_real (cmod b) / b) *
+   cmod (1 + b * (v ^ n / complex_of_real (cmod b)))
+   < cmod (complex_of_real (cmod b) / b) * 1"
+      apply (simp only: norm_mult[symmetric] right_distrib)
+      using b v by (simp add: th2)
+
+    from mult_less_imp_less_left[OF th4 th3]
+    have "?P ?w n" unfolding th1 . 
+    hence "\<exists>z. ?P z n" .. }
+  ultimately show "\<exists>z. ?P z n" by blast
+qed
+
+
+text{* Bolzano-Weierstrass type property for closed disc in complex plane. *}
+
+lemma metric_bound_lemma: "cmod (x - y) <= \<bar>Re x - Re y\<bar> + \<bar>Im x - Im y\<bar>"
+  using real_sqrt_sum_squares_triangle_ineq[of "Re x - Re y" 0 0 "Im x - Im y" ]
+  unfolding cmod_def by simp
+
+lemma bolzano_weierstrass_complex_disc:
+  assumes r: "\<forall>n. cmod (s n) \<le> r"
+  shows "\<exists>f z. subseq f \<and> (\<forall>e >0. \<exists>N. \<forall>n \<ge> N. cmod (s (f n) - z) < e)"
+proof-
+  from seq_monosub[of "Re o s"] 
+  obtain f g where f: "subseq f" "monoseq (\<lambda>n. Re (s (f n)))" 
+    unfolding o_def by blast
+  from seq_monosub[of "Im o s o f"] 
+  obtain g where g: "subseq g" "monoseq (\<lambda>n. Im (s(f(g n))))" unfolding o_def by blast  
+  let ?h = "f o g"
+  from r[rule_format, of 0] have rp: "r \<ge> 0" using norm_ge_zero[of "s 0"] by arith 
+  have th:"\<forall>n. r + 1 \<ge> \<bar> Re (s n)\<bar>" 
+  proof
+    fix n
+    from abs_Re_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Re (s n)\<bar> \<le> r + 1" by arith
+  qed
+  have conv1: "convergent (\<lambda>n. Re (s ( f n)))"
+    apply (rule Bseq_monoseq_convergent)
+    apply (simp add: Bseq_def)
+    apply (rule exI[where x= "r + 1"])
+    using th rp apply simp
+    using f(2) .
+  have th:"\<forall>n. r + 1 \<ge> \<bar> Im (s n)\<bar>" 
+  proof
+    fix n
+    from abs_Im_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Im (s n)\<bar> \<le> r + 1" by arith
+  qed
+
+  have conv2: "convergent (\<lambda>n. Im (s (f (g n))))"
+    apply (rule Bseq_monoseq_convergent)
+    apply (simp add: Bseq_def)
+    apply (rule exI[where x= "r + 1"])
+    using th rp apply simp
+    using g(2) .
+
+  from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x" 
+    by blast 
+  hence  x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Re (s (f n)) - x \<bar> < r" 
+    unfolding LIMSEQ_def real_norm_def .
+
+  from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y" 
+    by blast 
+  hence  y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Im (s (f (g n))) - y \<bar> < r" 
+    unfolding LIMSEQ_def real_norm_def .
+  let ?w = "Complex x y"
+  from f(1) g(1) have hs: "subseq ?h" unfolding subseq_def by auto 
+  {fix e assume ep: "e > (0::real)"
+    hence e2: "e/2 > 0" by simp
+    from x[rule_format, OF e2] y[rule_format, OF e2]
+    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
+    {fix n assume nN12: "n \<ge> N1 + N2"
+      hence nN1: "g n \<ge> N1" and nN2: "n \<ge> N2" using seq_suble[OF g(1), of n] by arith+
+      from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
+      have "cmod (s (?h n) - ?w) < e" 
+	using metric_bound_lemma[of "s (f (g n))" ?w] by simp }
+    hence "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" by blast }
+  with hs show ?thesis  by blast  
+qed
+
+text{* Polynomial is continuous. *}
+
+lemma poly_cont:
+  assumes ep: "e > 0" 
+  shows "\<exists>d >0. \<forall>w. 0 < cmod (w - z) \<and> cmod (w - z) < d \<longrightarrow> cmod (poly p w - poly p z) < e"
+proof-
+  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
+  {fix w
+    note q(2)[of "w - z", simplified]}
+  note th = this
+  show ?thesis unfolding th[symmetric]
+  proof(induct q)
+    case Nil thus ?case  using ep by auto
+  next
+    case (Cons c cs)
+    from poly_bound_exists[of 1 "cs"] 
+    obtain m where m: "m > 0" "\<And>z. cmod z \<le> 1 \<Longrightarrow> cmod (poly cs z) \<le> m" by blast
+    from ep m(1) have em0: "e/m > 0" by (simp add: field_simps)
+    have one0: "1 > (0::real)"  by arith
+    from real_lbound_gt_zero[OF one0 em0] 
+    obtain d where d: "d >0" "d < 1" "d < e / m" by blast
+    from d(1,3) m(1) have dm: "d*m > 0" "d*m < e" 
+      by (simp_all add: field_simps real_mult_order)
+    show ?case 
+      proof(rule ex_forward[OF real_lbound_gt_zero[OF one0 em0]], clarsimp simp add: norm_mult)
+	fix d w
+	assume H: "d > 0" "d < 1" "d < e/m" "w\<noteq>z" "cmod (w-z) < d"
+	hence d1: "cmod (w-z) \<le> 1" "d \<ge> 0" by simp_all
+	from H(3) m(1) have dme: "d*m < e" by (simp add: field_simps)
+	from H have th: "cmod (w-z) \<le> d" by simp 
+	from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
+	show "cmod (w - z) * cmod (poly cs (w - z)) < e" by simp
+      qed  
+    qed
+qed
+
+text{* Hence a polynomial attains minimum on a closed disc 
+  in the complex plane. *}
+lemma  poly_minimum_modulus_disc:
+  "\<exists>z. \<forall>w. cmod w \<le> r \<longrightarrow> cmod (poly p z) \<le> cmod (poly p w)"
+proof-
+  {assume "\<not> r \<ge> 0" hence ?thesis unfolding linorder_not_le
+      apply -
+      apply (rule exI[where x=0]) 
+      apply auto
+      apply (subgoal_tac "cmod w < 0")
+      apply simp
+      apply arith
+      done }
+  moreover
+  {assume rp: "r \<ge> 0"
+    from rp have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))" by simp 
+    hence mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x"  by blast
+    {fix x z
+      assume H: "cmod z \<le> r" "cmod (poly p z) = - x" "\<not>x < 1"
+      hence "- x < 0 " by arith
+      with H(2) norm_ge_zero[of "poly p z"]  have False by simp }
+    then have mth2: "\<exists>z. \<forall>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<longrightarrow> x < z" by blast
+    from real_sup_exists[OF mth1 mth2] obtain s where 
+      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
+    let ?m = "-s"
+    {fix y
+      from s[rule_format, of "-y"] have 
+    "(\<exists>z x. cmod z \<le> r \<and> -(- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y" 
+	unfolding minus_less_iff[of y ] equation_minus_iff by blast }
+    note s1 = this[unfolded minus_minus]
+    from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m" 
+      by auto
+    {fix n::nat
+      from s1[rule_format, of "?m + 1/real (Suc n)"] 
+      have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)"
+	by simp}
+    hence th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
+    from choice[OF th] obtain g where 
+      g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m+1 /real(Suc n)" 
+      by blast
+    from bolzano_weierstrass_complex_disc[OF g(1)] 
+    obtain f z where fz: "subseq f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
+      by blast    
+    {fix w 
+      assume wr: "cmod w \<le> r"
+      let ?e = "\<bar>cmod (poly p z) - ?m\<bar>"
+      {assume e: "?e > 0"
+	hence e2: "?e/2 > 0" by simp
+	from poly_cont[OF e2, of z p] obtain d where
+	  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
+	{fix w assume w: "cmod (w - z) < d"
+	  have "cmod(poly p w - poly p z) < ?e / 2"
+	    using d(2)[rule_format, of w] w e by (cases "w=z", simp_all)}
+	note th1 = this
+	
+	from fz(2)[rule_format, OF d(1)] obtain N1 where 
+	  N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d" by blast
+	from reals_Archimedean2[of "2/?e"] obtain N2::nat where
+	  N2: "2/?e < real N2" by blast
+	have th2: "cmod(poly p (g(f(N1 + N2))) - poly p z) < ?e/2"
+	  using N1[rule_format, of "N1 + N2"] th1 by simp
+	{fix a b e2 m :: real
+	have "a < e2 \<Longrightarrow> abs(b - m) < e2 \<Longrightarrow> 2 * e2 <= abs(b - m) + a
+          ==> False" by arith}
+      note th0 = this
+      have ath: 
+	"\<And>m x e. m <= x \<Longrightarrow>  x < m + e ==> abs(x - m::real) < e" by arith
+      from s1m[OF g(1)[rule_format]]
+      have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
+      from seq_suble[OF fz(1), of "N1+N2"]
+      have th00: "real (Suc (N1+N2)) \<le> real (Suc (f (N1+N2)))" by simp
+      have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1+N2)) > 0"  
+	using N2 by auto
+      from frac_le[OF th000 th00] have th00: "?m +1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))" by simp
+      from g(2)[rule_format, of "f (N1 + N2)"]
+      have th01:"cmod (poly p (g (f (N1 + N2)))) < - s + 1 / real (Suc (f (N1 + N2)))" .
+      from order_less_le_trans[OF th01 th00]
+      have th32: "cmod(poly p (g (f (N1 + N2)))) < ?m + (1/ real(Suc (N1 + N2)))" .
+      from N2 have "2/?e < real (Suc (N1 + N2))" by arith
+      with e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
+      have "?e/2 > 1/ real (Suc (N1 + N2))" by (simp add: inverse_eq_divide)
+      with ath[OF th31 th32]
+      have thc1:"\<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar>< ?e/2" by arith  
+      have ath2: "\<And>(a::real) b c m. \<bar>a - b\<bar> <= c ==> \<bar>b - m\<bar> <= \<bar>a - m\<bar> + c" 
+	by arith
+      have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar>
+\<le> cmod (poly p (g (f (N1 + N2))) - poly p z)" 
+	by (simp add: norm_triangle_ineq3)
+      from ath2[OF th22, of ?m]
+      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
+      from th0[OF th2 thc1 thc2] have False .}
+      hence "?e = 0" by auto
+      then have "cmod (poly p z) = ?m" by simp  
+      with s1m[OF wr]
+      have "cmod (poly p z) \<le> cmod (poly p w)" by simp }
+    hence ?thesis by blast}
+  ultimately show ?thesis by blast
+qed
+
+lemma "(rcis (sqrt (abs r)) (a/2)) ^ 2 = rcis (abs r) a"
+  unfolding power2_eq_square
+  apply (simp add: rcis_mult)
+  apply (simp add: power2_eq_square[symmetric])
+  done
+
+lemma cispi: "cis pi = -1" 
+  unfolding cis_def
+  by simp
+
+lemma "(rcis (sqrt (abs r)) ((pi + a)/2)) ^ 2 = rcis (- abs r) a"
+  unfolding power2_eq_square
+  apply (simp add: rcis_mult add_divide_distrib)
+  apply (simp add: power2_eq_square[symmetric] rcis_def cispi cis_mult[symmetric])
+  done
+
+text {* Nonzero polynomial in z goes to infinity as z does. *}
+
+instance complex::idom_char_0 by (intro_classes)
+instance complex :: recpower_idom_char_0 by intro_classes
+
+lemma poly_infinity:
+  assumes ex: "list_ex (\<lambda>c. c \<noteq> 0) p"
+  shows "\<exists>r. \<forall>z. r \<le> cmod z \<longrightarrow> d \<le> cmod (poly (a#p) z)"
+using ex
+proof(induct p arbitrary: a d)
+  case (Cons c cs a d) 
+  {assume H: "list_ex (\<lambda>c. c\<noteq>0) cs"
+    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
+    let ?r = "1 + \<bar>r\<bar>"
+    {fix z assume h: "1 + \<bar>r\<bar> \<le> cmod z"
+      have r0: "r \<le> cmod z" using h by arith
+      from r[rule_format, OF r0]
+      have th0: "d + cmod a \<le> 1 * cmod(poly (c#cs) z)" by arith
+      from h have z1: "cmod z \<ge> 1" by arith
+      from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (c#cs) z"]]]
+      have th1: "d \<le> cmod(z * poly (c#cs) z) - cmod a"
+	unfolding norm_mult by (simp add: ring_simps)
+      from complex_mod_triangle_sub[of "z * poly (c#cs) z" a]
+      have th2: "cmod(z * poly (c#cs) z) - cmod a \<le> cmod (poly (a#c#cs) z)" 
+	by (simp add: diff_le_eq ring_simps) 
+      from th1 th2 have "d \<le> cmod (poly (a#c#cs) z)"  by arith}
+    hence ?case by blast}
+  moreover
+  {assume cs0: "\<not> (list_ex (\<lambda>c. c \<noteq> 0) cs)"
+    with Cons.prems have c0: "c \<noteq> 0" by simp
+    from cs0 have cs0': "list_all (\<lambda>c. c = 0) cs" 
+      by (auto simp add: list_all_iff list_ex_iff)
+    {fix z
+      assume h: "(\<bar>d\<bar> + cmod a) / cmod c \<le> cmod z"
+      from c0 have "cmod c > 0" by simp
+      from h c0 have th0: "\<bar>d\<bar> + cmod a \<le> cmod (z*c)" 
+	by (simp add: field_simps norm_mult)
+      have ath: "\<And>mzh mazh ma. mzh <= mazh + ma ==> abs(d) + ma <= mzh ==> d <= mazh" by arith
+      from complex_mod_triangle_sub[of "z*c" a ]
+      have th1: "cmod (z * c) \<le> cmod (a + z * c) + cmod a"
+	by (simp add: ring_simps)
+      from ath[OF th1 th0] have "d \<le> cmod (poly (a # c # cs) z)" 
+	using poly_0[OF cs0'] by simp}
+    then have ?case  by blast}
+  ultimately show ?case by blast
+qed simp
+
+text {* Hence polynomial's modulus attains its minimum somewhere. *}
+lemma poly_minimum_modulus:
+  "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)"
+proof(induct p)
+  case (Cons c cs) 
+  {assume cs0: "list_ex (\<lambda>c. c \<noteq> 0) cs"
+    from poly_infinity[OF cs0, of "cmod (poly (c#cs) 0)" c]
+    obtain r where r: "\<And>z. r \<le> cmod z \<Longrightarrow> cmod (poly (c # cs) 0) \<le> cmod (poly (c # cs) z)" by blast
+    have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>" by arith
+    from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "c#cs"] 
+    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
+    {fix z assume z: "r \<le> cmod z"
+      from v[of 0] r[OF z] 
+      have "cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) z)"
+	by simp }
+    note v0 = this
+    from v0 v ath[of r] have ?case by blast}
+  moreover
+  {assume cs0: "\<not> (list_ex (\<lambda>c. c\<noteq>0) cs)"
+    hence th:"list_all (\<lambda>c. c = 0) cs" by (simp add: list_all_iff list_ex_iff)
+    from poly_0[OF th] Cons.hyps have ?case by simp}
+  ultimately show ?case by blast
+qed simp
+
+text{* Constant function (non-syntactic characterization). *}
+definition "constant f = (\<forall>x y. f x = f y)"
+
+lemma nonconstant_length: "\<not> (constant (poly p)) \<Longrightarrow> length p \<ge> 2"
+  unfolding constant_def
+  apply (induct p, auto)
+  apply (unfold not_less[symmetric])
+  apply simp
+  apply (rule ccontr)
+  apply auto
+  done
+ 
+lemma poly_replicate_append:
+  "poly ((replicate n 0)@p) (x::'a::{recpower, comm_ring}) = x^n * poly p x"
+  by(induct n, auto simp add: power_Suc ring_simps)
+
+text {* Decomposition of polynomial, skipping zero coefficients 
+  after the first.  *}
+
+lemma poly_decompose_lemma:
+ assumes nz: "\<not>(\<forall>z. z\<noteq>0 \<longrightarrow> poly p z = (0::'a::{recpower,idom}))"
+  shows "\<exists>k a q. a\<noteq>0 \<and> Suc (length q + k) = length p \<and> 
+                 (\<forall>z. poly p z = z^k * poly (a#q) z)"
+using nz
+proof(induct p)
+  case Nil thus ?case by simp
+next
+  case (Cons c cs)
+  {assume c0: "c = 0"
+    
+    from Cons.hyps Cons.prems c0 have ?case apply auto
+      apply (rule_tac x="k+1" in exI)
+      apply (rule_tac x="a" in exI, clarsimp)
+      apply (rule_tac x="q" in exI)
+      by (auto simp add: power_Suc)}
+  moreover
+  {assume c0: "c\<noteq>0"
+    hence ?case apply-
+      apply (rule exI[where x=0])
+      apply (rule exI[where x=c], clarsimp)
+      apply (rule exI[where x=cs])
+      apply auto
+      done}
+  ultimately show ?case by blast
+qed
+
+lemma poly_decompose:
+  assumes nc: "~constant(poly p)"
+  shows "\<exists>k a q. a\<noteq>(0::'a::{recpower,idom}) \<and> k\<noteq>0 \<and>
+               length q + k + 1 = length p \<and> 
+              (\<forall>z. poly p z = poly p 0 + z^k * poly (a#q) z)"
+using nc 
+proof(induct p)
+  case Nil thus ?case by (simp add: constant_def)
+next
+  case (Cons c cs)
+  {assume C:"\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0"
+    {fix x y
+      from C have "poly (c#cs) x = poly (c#cs) y" by (cases "x=0", auto)}
+    with Cons.prems have False by (auto simp add: constant_def)}
+  hence th: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)" ..
+  from poly_decompose_lemma[OF th] 
+  show ?case 
+    apply clarsimp    
+    apply (rule_tac x="k+1" in exI)
+    apply (rule_tac x="a" in exI)
+    apply simp
+    apply (rule_tac x="q" in exI)
+    apply (auto simp add: power_Suc)
+    done
+qed
+
+text{* Fundamental theorem of algebral *}
+
+lemma fundamental_theorem_of_algebra:
+  assumes nc: "~constant(poly p)"
+  shows "\<exists>z::complex. poly p z = 0"
+using nc
+proof(induct n\<equiv> "length p" arbitrary: p rule: nat_less_induct)
+  fix n fix p :: "complex list"
+  let ?p = "poly p"
+  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"
+  let ?ths = "\<exists>z. ?p z = 0"
+
+  from nonconstant_length[OF nc] have n2: "n\<ge> 2" by (simp add: n)
+  from poly_minimum_modulus obtain c where 
+    c: "\<forall>w. cmod (?p c) \<le> cmod (?p w)" by blast
+  {assume pc: "?p c = 0" hence ?ths by blast}
+  moreover
+  {assume pc0: "?p c \<noteq> 0"
+    from poly_offset[of p c] obtain q where
+      q: "length q = length p" "\<forall>x. poly q x = ?p (c+x)" by blast
+    {assume h: "constant (poly q)"
+      from q(2) have th: "\<forall>x. poly q (x - c) = ?p x" by auto
+      {fix x y
+	from th have "?p x = poly q (x - c)" by auto 
+	also have "\<dots> = poly q (y - c)" 
+	  using h unfolding constant_def by blast
+	also have "\<dots> = ?p y" using th by auto
+	finally have "?p x = ?p y" .}
+      with nc have False unfolding constant_def by blast }
+    hence qnc: "\<not> constant (poly q)" by blast
+    from q(2) have pqc0: "?p c = poly q 0" by simp
+    from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)" by simp 
+    let ?a0 = "poly q 0"
+    from pc0 pqc0 have a00: "?a0 \<noteq> 0" by simp 
+    from a00 
+    have qr: "\<forall>z. poly q z = poly (map (op * (inverse ?a0)) q) z * ?a0"
+      by (simp add: poly_cmult_map)
+    let ?r = "map (op * (inverse ?a0)) q"
+    have lgqr: "length q = length ?r" by simp 
+    {assume h: "\<And>x y. poly ?r x = poly ?r y"
+      {fix x y
+	from qr[rule_format, of x] 
+	have "poly q x = poly ?r x * ?a0" by auto
+	also have "\<dots> = poly ?r y * ?a0" using h by simp
+	also have "\<dots> = poly q y" using qr[rule_format, of y] by simp
+	finally have "poly q x = poly q y" .} 
+      with qnc have False unfolding constant_def by blast}
+    hence rnc: "\<not> constant (poly ?r)" unfolding constant_def by blast
+    from qr[rule_format, of 0] a00  have r01: "poly ?r 0 = 1" by auto
+    {fix w 
+      have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
+	using qr[rule_format, of w] a00 by simp
+      also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
+	using a00 unfolding norm_divide by (simp add: field_simps)
+      finally have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" .}
+    note mrmq_eq = this
+    from poly_decompose[OF rnc] obtain k a s where 
+      kas: "a\<noteq>0" "k\<noteq>0" "length s + k + 1 = length ?r" 
+      "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (a#s) z" by blast
+    {assume "k + 1 = n"
+      with kas(3) lgqr[symmetric] q(1) n[symmetric] have s0:"s=[]" by auto
+      {fix w
+	have "cmod (poly ?r w) = cmod (1 + a * w ^ k)" 
+	  using kas(4)[rule_format, of w] s0 r01 by (simp add: ring_simps)}
+      note hth = this [symmetric]
+	from reduce_poly_simple[OF kas(1,2)] 
+      have "\<exists>w. cmod (poly ?r w) < 1" unfolding hth by blast}
+    moreover
+    {assume kn: "k+1 \<noteq> n"
+      from kn kas(3) q(1) n[symmetric] have k1n: "k + 1 < n" by simp
+      have th01: "\<not> constant (poly (1#((replicate (k - 1) 0)@[a])))" 
+	unfolding constant_def poly_Nil poly_Cons poly_replicate_append
+	using kas(1) apply simp 
+	by (rule exI[where x=0], rule exI[where x=1], simp)
+      from kas(2) have th02: "k+1 = length (1#((replicate (k - 1) 0)@[a]))" 
+	by simp
+      from H[rule_format, OF k1n th01 th02]
+      obtain w where w: "1 + w^k * a = 0"
+	unfolding poly_Nil poly_Cons poly_replicate_append
+	using kas(2) by (auto simp add: power_Suc[symmetric, of _ "k - Suc 0"] 
+	  mult_assoc[of _ _ a, symmetric])
+      from poly_bound_exists[of "cmod w" s] obtain m where 
+	m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
+      have w0: "w\<noteq>0" using kas(2) w by (auto simp add: power_0_left)
+      from w have "(1 + w ^ k * a) - 1 = 0 - 1" by simp
+      then have wm1: "w^k * a = - 1" by simp
+      have inv0: "0 < inverse (cmod w ^ (k + 1) * m)" 
+	using norm_ge_zero[of w] w0 m(1)
+	  by (simp add: inverse_eq_divide zero_less_mult_iff)
+      with real_down2[OF zero_less_one] obtain t where
+	t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
+      let ?ct = "complex_of_real t"
+      let ?w = "?ct * w"
+      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)
+      also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
+	unfolding wm1 by (simp)
+      finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) = cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)" 
+	apply -
+	apply (rule cong[OF refl[of cmod]])
+	apply assumption
+	done
+      with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"] 
+      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 
+      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
+      have "t *cmod w \<le> 1 * cmod w" apply (rule mult_mono) using t(1,2) by auto
+      then have tw: "cmod ?w \<le> cmod w" using t(1) by (simp add: norm_mult) 
+      from t inv0 have "t* (cmod w ^ (k + 1) * m) < 1"
+	by (simp add: inverse_eq_divide field_simps)
+      with zero_less_power[OF t(1), of k] 
+      have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1" 
+	apply - apply (rule mult_strict_left_mono) by simp_all
+      have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k+1) * cmod (poly s ?w)))"  using w0 t(1)
+	by (simp add: ring_simps power_mult_distrib norm_of_real norm_power norm_mult)
+      then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
+	using t(1,2) m(2)[rule_format, OF tw] w0
+	apply (simp only: )
+	apply auto
+	apply (rule mult_mono, simp_all add: norm_ge_zero)+
+	apply (simp add: zero_le_mult_iff zero_le_power)
+	done
+      with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k" by simp 
+      from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1" 
+	by auto
+      from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] th120 th121]
+      have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" . 
+      from th11 th12
+      have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"  by arith 
+      then have "cmod (poly ?r ?w) < 1" 
+	unfolding kas(4)[rule_format, of ?w] r01 by simp 
+      then have "\<exists>w. cmod (poly ?r w) < 1" by blast}
+    ultimately have cr0_contr: "\<exists>w. cmod (poly ?r w) < 1" by blast
+    from cr0_contr cq0 q(2)
+    have ?ths unfolding mrmq_eq not_less[symmetric] by auto}
+  ultimately show ?ths by blast
+qed
+
+text {* Alternative version with a syntactic notion of constant polynomial. *}
+
+lemma fundamental_theorem_of_algebra_alt:
+  assumes nc: "~(\<exists>a l. a\<noteq> 0 \<and> list_all(\<lambda>b. b = 0) l \<and> p = a#l)"
+  shows "\<exists>z. poly p z = (0::complex)"
+using nc
+proof(induct p)
+  case (Cons c cs)
+  {assume "c=0" hence ?case by auto}
+  moreover
+  {assume c0: "c\<noteq>0"
+    {assume nc: "constant (poly (c#cs))"
+      from nc[unfolded constant_def, rule_format, of 0] 
+      have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto 
+      hence "list_all (\<lambda>c. c=0) cs"
+	proof(induct cs)
+	  case (Cons d ds)
+	  {assume "d=0" hence ?case using Cons.prems Cons.hyps by simp}
+	  moreover
+	  {assume d0: "d\<noteq>0"
+	    from poly_bound_exists[of 1 ds] obtain m where 
+	      m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
+	    have dm: "cmod d / m > 0" using d0 m(1) by (simp add: field_simps)
+	    from real_down2[OF dm zero_less_one] obtain x where 
+	      x: "x > 0" "x < cmod d / m" "x < 1" by blast
+	    let ?x = "complex_of_real x"
+	    from x have cx: "?x \<noteq> 0"  "cmod ?x \<le> 1" by simp_all
+	    from Cons.prems[rule_format, OF cx(1)]
+	    have cth: "cmod (?x*poly ds ?x) = cmod d" by (simp add: eq_diff_eq[symmetric])
+	    from m(2)[rule_format, OF cx(2)] x(1)
+	    have th0: "cmod (?x*poly ds ?x) \<le> x*m"
+	      by (simp add: norm_mult)
+	    from x(2) m(1) have "x*m < cmod d" by (simp add: field_simps)
+	    with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d" by auto
+	    with cth  have ?case by blast}
+	  ultimately show ?case by blast 
+	qed simp}
+      then have nc: "\<not> constant (poly (c#cs))" using Cons.prems c0 
+	by blast
+      from fundamental_theorem_of_algebra[OF nc] have ?case .}
+  ultimately show ?case by blast  
+qed simp
+
+subsection{* Nullstellenstatz, degrees and divisibility of polynomials *}
+
+lemma nullstellensatz_lemma:
+  fixes p :: "complex list"
+  assumes "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
+  and "degree p = n" and "n \<noteq> 0"
+  shows "p divides (pexp q n)"
+using prems
+proof(induct n arbitrary: p q rule: nat_less_induct)
+  fix n::nat fix p q :: "complex list"
+  assume IH: "\<forall>m<n. \<forall>p q.
+                 (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow>
+                 degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p divides (q %^ m)"
+    and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0" 
+    and dpn: "degree p = n" and n0: "n \<noteq> 0"
+  let ?ths = "p divides (q %^ n)"
+  {fix a assume a: "poly p a = 0"
+    {assume p0: "poly p = poly []" 
+      hence ?ths unfolding divides_def  using pq0 n0
+	apply - apply (rule exI[where x="[]"], rule ext)
+	by (auto simp add: poly_mult poly_exp)}
+    moreover
+    {assume p0: "poly p \<noteq> poly []" 
+      and oa: "order  a p \<noteq> 0"
+      from p0 have pne: "p \<noteq> []" by auto
+      let ?op = "order a p"
+      from p0 have ap: "([- a, 1] %^ ?op) divides p" 
+	"\<not> pexp [- a, 1] (Suc ?op) divides p" using order by blast+ 
+      note oop = order_degree[OF p0, unfolded dpn]
+      {assume q0: "q = []"
+	hence ?ths using n0 unfolding divides_def 
+	  apply simp
+	  apply (rule exI[where x="[]"], rule ext)
+	  by (simp add: divides_def poly_exp poly_mult)}
+      moreover
+      {assume q0: "q\<noteq>[]"
+	from pq0[rule_format, OF a, unfolded poly_linear_divides] q0
+	obtain r where r: "q = pmult [- a, 1] r" by blast
+	from ap[unfolded divides_def] obtain s where
+	  s: "poly p = poly (pmult (pexp [- a, 1] ?op) s)" by blast
+	have s0: "poly s \<noteq> poly []"
+	  using s p0 by (simp add: poly_entire)
+	hence pns0: "poly (pnormalize s) \<noteq> poly []" and sne: "s\<noteq>[]" by auto
+	{assume ds0: "degree s = 0"
+	  from ds0 pns0 have "\<exists>k. pnormalize s = [k]" unfolding degree_def 
+	    by (cases "pnormalize s", auto)
+	  then obtain k where kpn: "pnormalize s = [k]" by blast
+	  from pns0[unfolded poly_zero] kpn have k: "k \<noteq>0" "poly s = poly [k]"
+	    using poly_normalize[of s] by simp_all
+	  let ?w = "pmult (pmult [1/k] (pexp [-a,1] (n - ?op))) (pexp r n)"
+	  from k r s oop have "poly (pexp q n) = poly (pmult p ?w)"
+	    by - (rule ext, simp add: poly_mult poly_exp poly_cmult poly_add power_add[symmetric] ring_simps power_mult_distrib[symmetric])
+	  hence ?ths unfolding divides_def by blast}
+	moreover
+	{assume ds0: "degree s \<noteq> 0"
+	  from ds0 s0 dpn degree_unique[OF s, unfolded linear_pow_mul_degree] oa
+	    have dsn: "degree s < n" by auto 
+	    {fix x assume h: "poly s x = 0"
+	      {assume xa: "x = a"
+		from h[unfolded xa poly_linear_divides] sne obtain u where
+		  u: "s = pmult [- a, 1] u" by blast
+		have "poly p = poly (pmult (pexp [- a, 1] (Suc ?op)) u)"
+		  unfolding s u
+		  apply (rule ext)
+		  by (simp add: ring_simps power_mult_distrib[symmetric] poly_mult poly_cmult poly_add poly_exp)
+		with ap(2)[unfolded divides_def] have False by blast}
+	      note xa = this
+	      from h s have "poly p x = 0" by (simp add: poly_mult)
+	      with pq0 have "poly q x = 0" by blast
+	      with r xa have "poly r x = 0"
+		by (auto simp add: poly_mult poly_add poly_cmult eq_diff_eq[symmetric])}
+	    note impth = this
+	    from IH[rule_format, OF dsn, of s r] impth ds0
+	    have "s divides (pexp r (degree s))" by blast
+	    then obtain u where u: "poly (pexp r (degree s)) = poly (pmult s u)"
+	      unfolding divides_def by blast
+	    hence u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
+	      by (simp add: poly_mult[symmetric] poly_exp[symmetric])
+	    let ?w = "pmult (pmult u (pexp [-a,1] (n - ?op))) (pexp r (n - degree s))"
+	    from u' s r oop[of a] dsn have "poly (pexp q n) = poly (pmult p ?w)"
+	      apply - apply (rule ext)
+	      apply (simp only:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult ring_simps)
+	      
+	      apply (simp add:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult mult_assoc[symmetric])
+	      done
+	    hence ?ths unfolding divides_def by blast}
+      ultimately have ?ths by blast }
+      ultimately have ?ths by blast}
+    ultimately have ?ths using a order_root by blast}
+  moreover
+  {assume exa: "\<not> (\<exists>a. poly p a = 0)"
+    from fundamental_theorem_of_algebra_alt[of p] exa obtain c cs where
+      ccs: "c\<noteq>0" "list_all (\<lambda>c. c = 0) cs" "p = c#cs" by blast
+    
+    from poly_0[OF ccs(2)] ccs(3) 
+    have pp: "\<And>x. poly p x =  c" by simp
+    let ?w = "pmult [1/c] (pexp q n)"
+    from pp ccs(1) 
+    have "poly (pexp q n) = poly (pmult p ?w) "
+      apply - apply (rule ext)
+      unfolding poly_mult_assoc[symmetric] by (simp add: poly_mult)
+    hence ?ths unfolding divides_def by blast}
+  ultimately show ?ths by blast
+qed
+
+lemma nullstellensatz_univariate:
+  "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> 
+    p divides (q %^ (degree p)) \<or> (poly p = poly [] \<and> poly q = poly [])"
+proof-
+  {assume pe: "poly p = poly []"
+    hence eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> poly q = poly []"
+      apply auto
+      by (rule ext, simp)
+    {assume "p divides (pexp q (degree p))"
+      then obtain r where r: "poly (pexp q (degree p)) = poly (pmult p r)" 
+	unfolding divides_def by blast
+      from cong[OF r refl] pe degree_unique[OF pe]
+      have False by (simp add: poly_mult degree_def)}
+    with eq pe have ?thesis by blast}
+  moreover
+  {assume pe: "poly p \<noteq> poly []"
+    have p0: "poly [0] = poly []" by (rule ext, simp)
+    {assume dp: "degree p = 0"
+      then obtain k where "pnormalize p = [k]" using pe poly_normalize[of p]
+	unfolding degree_def by (cases "pnormalize p", auto)
+      hence k: "pnormalize p = [k]" "poly p = poly [k]" "k\<noteq>0"
+	using pe poly_normalize[of p] by (auto simp add: p0)
+      hence th1: "\<forall>x. poly p x \<noteq> 0" by simp
+      from k(2,3) dp have "poly (pexp q (degree p)) = poly (pmult p [1/k]) "
+	by - (rule ext, simp add: poly_mult poly_exp)
+      hence th2: "p divides (pexp q (degree p))" unfolding divides_def by blast
+      from th1 th2 pe have ?thesis by blast}
+    moreover
+    {assume dp: "degree p \<noteq> 0"
+      then obtain n where n: "degree p = Suc n " by (cases "degree p", auto)
+      {assume "p divides (pexp q (Suc n))"
+	then obtain u where u: "poly (pexp q (Suc n)) = poly (pmult p u)"
+	  unfolding divides_def by blast
+	hence u' :"\<And>x. poly (pexp q (Suc n)) x = poly (pmult p u) x" by simp_all
+	{fix x assume h: "poly p x = 0" "poly q x \<noteq> 0"
+	  hence "poly (pexp q (Suc n)) x \<noteq> 0" by (simp only: poly_exp) simp	  
+	  hence False using u' h(1) by (simp only: poly_mult poly_exp) simp}}
+	with n nullstellensatz_lemma[of p q "degree p"] dp 
+	have ?thesis by auto}
+    ultimately have ?thesis by blast}
+  ultimately show ?thesis by blast
+qed
+
+text{* Useful lemma *}
+
+lemma (in idom_char_0) constant_degree: "constant (poly p) \<longleftrightarrow> degree p = 0" (is "?lhs = ?rhs")
+proof
+  assume l: ?lhs
+  from l[unfolded constant_def, rule_format, of _ "zero"]
+  have th: "poly p = poly [poly p 0]" apply - by (rule ext, simp)
+  from degree_unique[OF th] show ?rhs by (simp add: degree_def)
+next
+  assume r: ?rhs
+  from r have "pnormalize p = [] \<or> (\<exists>k. pnormalize p = [k])"
+    unfolding degree_def by (cases "pnormalize p", auto)
+  then show ?lhs unfolding constant_def poly_normalize[of p, symmetric]
+    by (auto simp del: poly_normalize)
+qed
+
+(* It would be nicer to prove this without using algebraic closure...        *)
+
+lemma divides_degree_lemma: assumes dpn: "degree (p::complex list) = n"
+  shows "n \<le> degree (p *** q) \<or> poly (p *** q) = poly []"
+  using dpn
+proof(induct n arbitrary: p q)
+  case 0 thus ?case by simp
+next
+  case (Suc n p q)
+  from Suc.prems fundamental_theorem_of_algebra[of p] constant_degree[of p]
+  obtain a where a: "poly p a = 0" by auto
+  then obtain r where r: "p = pmult [-a, 1] r" unfolding poly_linear_divides
+    using Suc.prems by (auto simp add: degree_def)
+  {assume h: "poly (pmult r q) = poly []"
+    hence "poly (pmult p q) = poly []" using r
+      apply - apply (rule ext)  by (auto simp add: poly_entire poly_mult poly_add poly_cmult) hence ?case by blast}
+  moreover
+  {assume h: "poly (pmult r q) \<noteq> poly []" 
+    hence r0: "poly r \<noteq> poly []" and q0: "poly q \<noteq> poly []"
+      by (auto simp add: poly_entire)
+    have eq: "poly (pmult p q) = poly (pmult [-a, 1] (pmult r q))"
+      apply - apply (rule ext)
+      by (simp add: r poly_mult poly_add poly_cmult ring_simps)
+    from linear_mul_degree[OF h, of "- a"]
+    have dqe: "degree (pmult p q) = degree (pmult r q) + 1"
+      unfolding degree_unique[OF eq] .
+    from linear_mul_degree[OF r0, of "- a", unfolded r[symmetric]] r Suc.prems 
+    have dr: "degree r = n" by auto
+    from  Suc.hyps[OF dr, of q] have "Suc n \<le> degree (pmult p q)"
+      unfolding dqe using h by (auto simp del: poly.simps) 
+    hence ?case by blast}
+  ultimately show ?case by blast
+qed
+
+lemma divides_degree: assumes pq: "p divides (q:: complex list)"
+  shows "degree p \<le> degree q \<or> poly q = poly []"
+using pq  divides_degree_lemma[OF refl, of p]
+apply (auto simp add: divides_def poly_entire)
+apply atomize
+apply (erule_tac x="qa" in allE, auto)
+apply (subgoal_tac "degree q = degree (p *** qa)", simp)
+apply (rule degree_unique, simp)
+done
+
+(* Arithmetic operations on multivariate polynomials.                        *)
+
+lemma mpoly_base_conv: 
+  "(0::complex) \<equiv> poly [] x" "c \<equiv> poly [c] x" "x \<equiv> poly [0,1] x" by simp_all
+
+lemma mpoly_norm_conv: 
+  "poly [0] (x::complex) \<equiv> poly [] x" "poly [poly [] y] x \<equiv> poly [] x" by simp_all
+
+lemma mpoly_sub_conv: 
+  "poly p (x::complex) - poly q x \<equiv> poly p x + -1 * poly q x"
+  by (simp add: diff_def)
+
+lemma poly_pad_rule: "poly p x = 0 ==> poly (0#p) x = (0::complex)" by simp
+
+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
+
+lemma resolve_eq_raw:  "poly [] x \<equiv> 0" "poly [c] x \<equiv> (c::complex)" by auto
+lemma  resolve_eq_then: "(P \<Longrightarrow> (Q \<equiv> Q1)) \<Longrightarrow> (\<not>P \<Longrightarrow> (Q \<equiv> Q2))
+  \<Longrightarrow> Q \<equiv> P \<and> Q1 \<or> \<not>P\<and> Q2" apply (atomize (full)) by blast 
+lemma expand_ex_beta_conv: "list_ex P [c] \<equiv> P c" by simp
+
+lemma poly_divides_pad_rule: 
+  fixes p q :: "complex list"
+  assumes pq: "p divides q"
+  shows "p divides ((0::complex)#q)"
+proof-
+  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
+  hence "poly (0#q) = poly (p *** ([0,1] *** r))" 
+    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
+  thus ?thesis unfolding divides_def by blast
+qed
+
+lemma poly_divides_pad_const_rule: 
+  fixes p q :: "complex list"
+  assumes pq: "p divides q"
+  shows "p divides (a %* q)"
+proof-
+  from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
+  hence "poly (a %* q) = poly (p *** (a %* r))" 
+    by - (rule ext, simp add: poly_mult poly_cmult poly_add)
+  thus ?thesis unfolding divides_def by blast
+qed
+
+
+lemma poly_divides_conv0:  
+  fixes p :: "complex list"
+  assumes lgpq: "length q < length p" and lq:"last p \<noteq> 0"
+  shows "p divides q \<equiv> (\<not> (list_ex (\<lambda>c. c \<noteq> 0) q))" (is "?lhs \<equiv> ?rhs")
+proof-
+  {assume r: ?rhs 
+    hence eq: "poly q = poly []" unfolding poly_zero 
+      by (simp add: list_all_iff list_ex_iff)
+    hence "poly q = poly (p *** [])" by - (rule ext, simp add: poly_mult)
+    hence ?lhs unfolding divides_def  by blast}
+  moreover
+  {assume l: ?lhs
+    have ath: "\<And>lq lp dq::nat. lq < lp ==> lq \<noteq> 0 \<Longrightarrow> dq <= lq - 1 ==> dq < lp - 1"
+      by arith
+    {assume q0: "length q = 0"
+      hence "q = []" by simp
+      hence ?rhs by simp}
+    moreover
+    {assume lgq0: "length q \<noteq> 0"
+      from pnormalize_length[of q] have dql: "degree q \<le> length q - 1" 
+	unfolding degree_def by simp
+      from ath[OF lgpq lgq0 dql, unfolded pnormal_degree[OF lq, symmetric]] divides_degree[OF l] have "poly q = poly []" by auto
+      hence ?rhs unfolding poly_zero by (simp add: list_all_iff list_ex_iff)}
+    ultimately have ?rhs by blast }
+  ultimately show "?lhs \<equiv> ?rhs" by - (atomize (full), blast) 
+qed
+
+lemma poly_divides_conv1: 
+  assumes a0: "a\<noteq> (0::complex)" and pp': "(p::complex list) divides p'"
+  and qrp': "\<And>x. a * poly q x - poly p' x \<equiv> poly r x"
+  shows "p divides q \<equiv> p divides (r::complex list)" (is "?lhs \<equiv> ?rhs")
+proof-
+  {
+  from pp' obtain t where t: "poly p' = poly (p *** t)" 
+    unfolding divides_def by blast
+  {assume l: ?lhs
+    then obtain u where u: "poly q = poly (p *** u)" unfolding divides_def by blast
+     have "poly r = poly (p *** ((a %* u) +++ (-- t)))"
+       using u qrp' t
+       by - (rule ext, 
+	 simp add: poly_add poly_mult poly_cmult poly_minus ring_simps)
+     then have ?rhs unfolding divides_def by blast}
+  moreover
+  {assume r: ?rhs
+    then obtain u where u: "poly r = poly (p *** u)" unfolding divides_def by blast
+    from u t qrp' a0 have "poly q = poly (p *** ((1/a) %* (u +++ t)))"
+      by - (rule ext, atomize (full), simp add: poly_mult poly_add poly_cmult field_simps)
+    hence ?lhs  unfolding divides_def by blast}
+  ultimately have "?lhs = ?rhs" by blast }
+thus "?lhs \<equiv> ?rhs"  by - (atomize(full), blast) 
+qed
+
+lemma basic_cqe_conv1:
+  "(\<exists>x. poly p x = 0 \<and> poly [] x \<noteq> 0) \<equiv> False"
+  "(\<exists>x. poly [] x \<noteq> 0) \<equiv> False"
+  "(\<exists>x. poly [c] x \<noteq> 0) \<equiv> c\<noteq>0"
+  "(\<exists>x. poly [] x = 0) \<equiv> True"
+  "(\<exists>x. poly [c] x = 0) \<equiv> c = 0" by simp_all
+
+lemma basic_cqe_conv2: 
+  assumes l:"last (a#b#p) \<noteq> 0" 
+  shows "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True"
+proof-
+  {fix h t
+    assume h: "h\<noteq>0" "list_all (\<lambda>c. c=(0::complex)) t"  "a#b#p = h#t"
+    hence "list_all (\<lambda>c. c= 0) (b#p)" by simp
+    moreover have "last (b#p) \<in> set (b#p)" by simp
+    ultimately have "last (b#p) = 0" by (simp add: list_all_iff)
+    with l have False by simp}
+  hence th: "\<not> (\<exists> h t. h\<noteq>0 \<and> list_all (\<lambda>c. c=0) t \<and> a#b#p = h#t)"
+    by blast
+  from fundamental_theorem_of_algebra_alt[OF th] 
+  show "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True" by auto
+qed
+
+lemma  basic_cqe_conv_2b: "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
+proof-
+  have "\<not> (list_ex (\<lambda>c. c \<noteq> 0) p) \<longleftrightarrow> poly p = poly []" 
+    by (simp add: poly_zero list_all_iff list_ex_iff)
+  also have "\<dots> \<longleftrightarrow> (\<not> (\<exists>x. poly p x \<noteq> 0))" by (auto intro: ext)
+  finally show "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
+    by - (atomize (full), blast)
+qed
+
+lemma basic_cqe_conv3:
+  fixes p q :: "complex list"
+  assumes l: "last (a#p) \<noteq> 0" 
+  shows "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
+proof-
+  note np = pnormalize_eq[OF l]
+  {assume "poly (a#p) = poly []" hence False using l
+      unfolding poly_zero apply (auto simp add: list_all_iff del: last.simps)
+      apply (cases p, simp_all) done}
+  then have p0: "poly (a#p) \<noteq> poly []"  by blast
+  from np have dp:"degree (a#p) = length p" by (simp add: degree_def)
+  from nullstellensatz_univariate[of "a#p" q] p0 dp
+  show "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
+    by - (atomize (full), auto)
+qed
+
+lemma basic_cqe_conv4:
+  fixes p q :: "complex list"
+  assumes h: "\<And>x. poly (q %^ n) x \<equiv> poly r x"
+  shows "p divides (q %^ n) \<equiv> p divides r"
+proof-
+  from h have "poly (q %^ n) = poly r" by (auto intro: ext)  
+  thus "p divides (q %^ n) \<equiv> p divides r" unfolding divides_def by simp
+qed
+
+lemma pmult_Cons_Cons: "((a::complex)#b#p) *** q = (a %*q) +++ (0#((b#p) *** q))"
+  by simp
+
+lemma elim_neg_conv: "- z \<equiv> (-1) * (z::complex)" by simp
+lemma eqT_intr: "PROP P \<Longrightarrow> (True \<Longrightarrow> PROP P )" "PROP P \<Longrightarrow> True" by blast+
+lemma negate_negate_rule: "Trueprop P \<equiv> \<not> P \<equiv> False" by (atomize (full), auto)
+lemma last_simps: "last [x] = x" "last (x#y#ys) = last (y#ys)" by simp_all
+lemma length_simps: "length [] = 0" "length (x#y#xs) = length xs + 2" "length [x] = 1" by simp_all
+
+lemma complex_entire: "(z::complex) \<noteq> 0 \<and> w \<noteq> 0 \<equiv> z*w \<noteq> 0" by simp
+lemma resolve_eq_ne: "(P \<equiv> True) \<equiv> (\<not>P \<equiv> False)" "(P \<equiv> False) \<equiv> (\<not>P \<equiv> True)" 
+  by (atomize (full)) simp_all
+lemma cqe_conv1: "poly [] x = 0 \<longleftrightarrow> True"  by simp
+lemma cqe_conv2: "(p \<Longrightarrow> (q \<equiv> r)) \<equiv> ((p \<and> q) \<equiv> (p \<and> r))"  (is "?l \<equiv> ?r")
+proof
+  assume "p \<Longrightarrow> q \<equiv> r" thus "p \<and> q \<equiv> p \<and> r" apply - apply (atomize (full)) by blast
+next
+  assume "p \<and> q \<equiv> p \<and> r" "p"
+  thus "q \<equiv> r" apply - apply (atomize (full)) apply blast done
+qed
+lemma poly_const_conv: "poly [c] (x::complex) = y \<longleftrightarrow> c = y" by simp
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/Bounds.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,83 @@
+(*  Title:      HOL/Real/HahnBanach/Bounds.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Bounds *}
+
+theory Bounds
+imports Main ContNotDenum
+begin
+
+locale lub =
+  fixes A and x
+  assumes least [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> a \<le> b) \<Longrightarrow> x \<le> b"
+    and upper [intro?]: "a \<in> A \<Longrightarrow> a \<le> x"
+
+lemmas [elim?] = lub.least lub.upper
+
+definition
+  the_lub :: "'a::order set \<Rightarrow> 'a" where
+  "the_lub A = The (lub A)"
+
+notation (xsymbols)
+  the_lub  ("\<Squnion>_" [90] 90)
+
+lemma the_lub_equality [elim?]:
+  assumes "lub A x"
+  shows "\<Squnion>A = (x::'a::order)"
+proof -
+  interpret lub [A x] by fact
+  show ?thesis
+  proof (unfold the_lub_def)
+    from `lub A x` show "The (lub A) = x"
+    proof
+      fix x' assume lub': "lub A x'"
+      show "x' = x"
+      proof (rule order_antisym)
+	from lub' show "x' \<le> x"
+	proof
+          fix a assume "a \<in> A"
+          then show "a \<le> x" ..
+	qed
+	show "x \<le> x'"
+	proof
+          fix a assume "a \<in> A"
+          with lub' show "a \<le> x'" ..
+	qed
+      qed
+    qed
+  qed
+qed
+
+lemma the_lubI_ex:
+  assumes ex: "\<exists>x. lub A x"
+  shows "lub A (\<Squnion>A)"
+proof -
+  from ex obtain x where x: "lub A x" ..
+  also from x have [symmetric]: "\<Squnion>A = x" ..
+  finally show ?thesis .
+qed
+
+lemma lub_compat: "lub A x = isLub UNIV A x"
+proof -
+  have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)"
+    by (rule ext) (simp only: isUb_def)
+  then show ?thesis
+    by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast
+qed
+
+lemma real_complete:
+  fixes A :: "real set"
+  assumes nonempty: "\<exists>a. a \<in> A"
+    and ex_upper: "\<exists>y. \<forall>a \<in> A. a \<le> y"
+  shows "\<exists>x. lub A x"
+proof -
+  from ex_upper have "\<exists>y. isUb UNIV A y"
+    unfolding isUb_def setle_def by blast
+  with nonempty have "\<exists>x. isLub UNIV A x"
+    by (rule reals_complete)
+  then show ?thesis by (simp only: lub_compat)
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/FunctionNorm.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,279 @@
+(*  Title:      HOL/Real/HahnBanach/FunctionNorm.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* The norm of a function *}
+
+theory FunctionNorm
+imports NormedSpace FunctionOrder
+begin
+
+subsection {* Continuous linear forms*}
+
+text {*
+  A linear form @{text f} on a normed vector space @{text "(V, \<parallel>\<cdot>\<parallel>)"}
+  is \emph{continuous}, iff it is bounded, i.e.
+  \begin{center}
+  @{text "\<exists>c \<in> R. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
+  \end{center}
+  In our application no other functions than linear forms are
+  considered, so we can define continuous linear forms as bounded
+  linear forms:
+*}
+
+locale continuous = var V + norm_syntax + linearform +
+  assumes bounded: "\<exists>c. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
+
+declare continuous.intro [intro?] continuous_axioms.intro [intro?]
+
+lemma continuousI [intro]:
+  fixes norm :: "_ \<Rightarrow> real"  ("\<parallel>_\<parallel>")
+  assumes "linearform V f"
+  assumes r: "\<And>x. x \<in> V \<Longrightarrow> \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
+  shows "continuous V norm f"
+proof
+  show "linearform V f" by fact
+  from r have "\<exists>c. \<forall>x\<in>V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" by blast
+  then show "continuous_axioms V norm f" ..
+qed
+
+
+subsection {* The norm of a linear form *}
+
+text {*
+  The least real number @{text c} for which holds
+  \begin{center}
+  @{text "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
+  \end{center}
+  is called the \emph{norm} of @{text f}.
+
+  For non-trivial vector spaces @{text "V \<noteq> {0}"} the norm can be
+  defined as
+  \begin{center}
+  @{text "\<parallel>f\<parallel> = \<sup>x \<noteq> 0. \<bar>f x\<bar> / \<parallel>x\<parallel>"}
+  \end{center}
+
+  For the case @{text "V = {0}"} the supremum would be taken from an
+  empty set. Since @{text \<real>} is unbounded, there would be no supremum.
+  To avoid this situation it must be guaranteed that there is an
+  element in this set. This element must be @{text "{} \<ge> 0"} so that
+  @{text fn_norm} has the norm properties. Furthermore it does not
+  have to change the norm in all other cases, so it must be @{text 0},
+  as all other elements are @{text "{} \<ge> 0"}.
+
+  Thus we define the set @{text B} where the supremum is taken from as
+  follows:
+  \begin{center}
+  @{text "{0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel>. x \<noteq> 0 \<and> x \<in> F}"}
+  \end{center}
+
+  @{text fn_norm} is equal to the supremum of @{text B}, if the
+  supremum exists (otherwise it is undefined).
+*}
+
+locale fn_norm = norm_syntax +
+  fixes B defines "B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
+  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
+  defines "\<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
+
+locale normed_vectorspace_with_fn_norm = normed_vectorspace + fn_norm
+
+lemma (in fn_norm) B_not_empty [intro]: "0 \<in> B V f"
+  by (simp add: B_def)
+
+text {*
+  The following lemma states that every continuous linear form on a
+  normed space @{text "(V, \<parallel>\<cdot>\<parallel>)"} has a function norm.
+*}
+
+lemma (in normed_vectorspace_with_fn_norm) fn_norm_works:
+  assumes "continuous V norm f"
+  shows "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
+proof -
+  interpret continuous [V norm f] by fact
+  txt {* The existence of the supremum is shown using the
+    completeness of the reals. Completeness means, that every
+    non-empty bounded set of reals has a supremum. *}
+  have "\<exists>a. lub (B V f) a"
+  proof (rule real_complete)
+    txt {* First we have to show that @{text B} is non-empty: *}
+    have "0 \<in> B V f" ..
+    then show "\<exists>x. x \<in> B V f" ..
+
+    txt {* Then we have to show that @{text B} is bounded: *}
+    show "\<exists>c. \<forall>y \<in> B V f. y \<le> c"
+    proof -
+      txt {* We know that @{text f} is bounded by some value @{text c}. *}
+      from bounded obtain c where c: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
+
+      txt {* To prove the thesis, we have to show that there is some
+        @{text b}, such that @{text "y \<le> b"} for all @{text "y \<in>
+        B"}. Due to the definition of @{text B} there are two cases. *}
+
+      def b \<equiv> "max c 0"
+      have "\<forall>y \<in> B V f. y \<le> b"
+      proof
+        fix y assume y: "y \<in> B V f"
+        show "y \<le> b"
+        proof cases
+          assume "y = 0"
+          then show ?thesis unfolding b_def by arith
+        next
+          txt {* The second case is @{text "y = \<bar>f x\<bar> / \<parallel>x\<parallel>"} for some
+            @{text "x \<in> V"} with @{text "x \<noteq> 0"}. *}
+          assume "y \<noteq> 0"
+          with y obtain x where y_rep: "y = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
+              and x: "x \<in> V" and neq: "x \<noteq> 0"
+            by (auto simp add: B_def real_divide_def)
+          from x neq have gt: "0 < \<parallel>x\<parallel>" ..
+
+          txt {* The thesis follows by a short calculation using the
+            fact that @{text f} is bounded. *}
+
+          note y_rep
+          also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
+          proof (rule mult_right_mono)
+            from c x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
+            from gt have "0 < inverse \<parallel>x\<parallel>" 
+              by (rule positive_imp_inverse_positive)
+            then show "0 \<le> inverse \<parallel>x\<parallel>" by (rule order_less_imp_le)
+          qed
+          also have "\<dots> = c * (\<parallel>x\<parallel> * inverse \<parallel>x\<parallel>)"
+            by (rule real_mult_assoc)
+          also
+          from gt have "\<parallel>x\<parallel> \<noteq> 0" by simp
+          then have "\<parallel>x\<parallel> * inverse \<parallel>x\<parallel> = 1" by simp 
+          also have "c * 1 \<le> b" by (simp add: b_def le_maxI1)
+          finally show "y \<le> b" .
+        qed
+      qed
+      then show ?thesis ..
+    qed
+  qed
+  then show ?thesis unfolding fn_norm_def by (rule the_lubI_ex)
+qed
+
+lemma (in normed_vectorspace_with_fn_norm) fn_norm_ub [iff?]:
+  assumes "continuous V norm f"
+  assumes b: "b \<in> B V f"
+  shows "b \<le> \<parallel>f\<parallel>\<hyphen>V"
+proof -
+  interpret continuous [V norm f] by fact
+  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
+    using `continuous V norm f` by (rule fn_norm_works)
+  from this and b show ?thesis ..
+qed
+
+lemma (in normed_vectorspace_with_fn_norm) fn_norm_leastB:
+  assumes "continuous V norm f"
+  assumes b: "\<And>b. b \<in> B V f \<Longrightarrow> b \<le> y"
+  shows "\<parallel>f\<parallel>\<hyphen>V \<le> y"
+proof -
+  interpret continuous [V norm f] by fact
+  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
+    using `continuous V norm f` by (rule fn_norm_works)
+  from this and b show ?thesis ..
+qed
+
+text {* The norm of a continuous function is always @{text "\<ge> 0"}. *}
+
+lemma (in normed_vectorspace_with_fn_norm) fn_norm_ge_zero [iff]:
+  assumes "continuous V norm f"
+  shows "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
+proof -
+  interpret continuous [V norm f] by fact
+  txt {* The function norm is defined as the supremum of @{text B}.
+    So it is @{text "\<ge> 0"} if all elements in @{text B} are @{text "\<ge>
+    0"}, provided the supremum exists and @{text B} is not empty. *}
+  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
+    using `continuous V norm f` by (rule fn_norm_works)
+  moreover have "0 \<in> B V f" ..
+  ultimately show ?thesis ..
+qed
+
+text {*
+  \medskip The fundamental property of function norms is:
+  \begin{center}
+  @{text "\<bar>f x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
+  \end{center}
+*}
+
+lemma (in normed_vectorspace_with_fn_norm) fn_norm_le_cong:
+  assumes "continuous V norm f" "linearform V f"
+  assumes x: "x \<in> V"
+  shows "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
+proof -
+  interpret continuous [V norm f] by fact
+  interpret linearform [V f] .
+  show ?thesis
+  proof cases
+    assume "x = 0"
+    then have "\<bar>f x\<bar> = \<bar>f 0\<bar>" by simp
+    also have "f 0 = 0" by rule unfold_locales
+    also have "\<bar>\<dots>\<bar> = 0" by simp
+    also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
+      using `continuous V norm f` by (rule fn_norm_ge_zero)
+    from x have "0 \<le> norm x" ..
+    with a have "0 \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" by (simp add: zero_le_mult_iff)
+    finally show "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" .
+  next
+    assume "x \<noteq> 0"
+    with x have neq: "\<parallel>x\<parallel> \<noteq> 0" by simp
+    then have "\<bar>f x\<bar> = (\<bar>f x\<bar> * inverse \<parallel>x\<parallel>) * \<parallel>x\<parallel>" by simp
+    also have "\<dots> \<le>  \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
+    proof (rule mult_right_mono)
+      from x show "0 \<le> \<parallel>x\<parallel>" ..
+      from x and neq have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<in> B V f"
+	by (auto simp add: B_def real_divide_def)
+      with `continuous V norm f` show "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>V"
+	by (rule fn_norm_ub)
+    qed
+    finally show ?thesis .
+  qed
+qed
+
+text {*
+  \medskip The function norm is the least positive real number for
+  which the following inequation holds:
+  \begin{center}
+    @{text "\<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
+  \end{center}
+*}
+
+lemma (in normed_vectorspace_with_fn_norm) fn_norm_least [intro?]:
+  assumes "continuous V norm f"
+  assumes ineq: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" and ge: "0 \<le> c"
+  shows "\<parallel>f\<parallel>\<hyphen>V \<le> c"
+proof -
+  interpret continuous [V norm f] by fact
+  show ?thesis
+  proof (rule fn_norm_leastB [folded B_def fn_norm_def])
+    fix b assume b: "b \<in> B V f"
+    show "b \<le> c"
+    proof cases
+      assume "b = 0"
+      with ge show ?thesis by simp
+    next
+      assume "b \<noteq> 0"
+      with b obtain x where b_rep: "b = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
+        and x_neq: "x \<noteq> 0" and x: "x \<in> V"
+	by (auto simp add: B_def real_divide_def)
+      note b_rep
+      also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
+      proof (rule mult_right_mono)
+	have "0 < \<parallel>x\<parallel>" using x x_neq ..
+	then show "0 \<le> inverse \<parallel>x\<parallel>" by simp
+	from ineq and x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
+      qed
+      also have "\<dots> = c"
+      proof -
+	from x_neq and x have "\<parallel>x\<parallel> \<noteq> 0" by simp
+	then show ?thesis by simp
+      qed
+      finally show ?thesis .
+    qed
+  qed (insert `continuous V norm f`, simp_all add: continuous_def)
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/FunctionOrder.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,142 @@
+(*  Title:      HOL/Real/HahnBanach/FunctionOrder.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* An order on functions *}
+
+theory FunctionOrder
+imports Subspace Linearform
+begin
+
+subsection {* The graph of a function *}
+
+text {*
+  We define the \emph{graph} of a (real) function @{text f} with
+  domain @{text F} as the set
+  \begin{center}
+  @{text "{(x, f x). x \<in> F}"}
+  \end{center}
+  So we are modeling partial functions by specifying the domain and
+  the mapping function. We use the term ``function'' also for its
+  graph.
+*}
+
+types 'a graph = "('a \<times> real) set"
+
+definition
+  graph :: "'a set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a graph" where
+  "graph F f = {(x, f x) | x. x \<in> F}"
+
+lemma graphI [intro]: "x \<in> F \<Longrightarrow> (x, f x) \<in> graph F f"
+  unfolding graph_def by blast
+
+lemma graphI2 [intro?]: "x \<in> F \<Longrightarrow> \<exists>t \<in> graph F f. t = (x, f x)"
+  unfolding graph_def by blast
+
+lemma graphE [elim?]:
+    "(x, y) \<in> graph F f \<Longrightarrow> (x \<in> F \<Longrightarrow> y = f x \<Longrightarrow> C) \<Longrightarrow> C"
+  unfolding graph_def by blast
+
+
+subsection {* Functions ordered by domain extension *}
+
+text {*
+  A function @{text h'} is an extension of @{text h}, iff the graph of
+  @{text h} is a subset of the graph of @{text h'}.
+*}
+
+lemma graph_extI:
+  "(\<And>x. x \<in> H \<Longrightarrow> h x = h' x) \<Longrightarrow> H \<subseteq> H'
+    \<Longrightarrow> graph H h \<subseteq> graph H' h'"
+  unfolding graph_def by blast
+
+lemma graph_extD1 [dest?]:
+  "graph H h \<subseteq> graph H' h' \<Longrightarrow> x \<in> H \<Longrightarrow> h x = h' x"
+  unfolding graph_def by blast
+
+lemma graph_extD2 [dest?]:
+  "graph H h \<subseteq> graph H' h' \<Longrightarrow> H \<subseteq> H'"
+  unfolding graph_def by blast
+
+
+subsection {* Domain and function of a graph *}
+
+text {*
+  The inverse functions to @{text graph} are @{text domain} and @{text
+  funct}.
+*}
+
+definition
+  "domain" :: "'a graph \<Rightarrow> 'a set" where
+  "domain g = {x. \<exists>y. (x, y) \<in> g}"
+
+definition
+  funct :: "'a graph \<Rightarrow> ('a \<Rightarrow> real)" where
+  "funct g = (\<lambda>x. (SOME y. (x, y) \<in> g))"
+
+text {*
+  The following lemma states that @{text g} is the graph of a function
+  if the relation induced by @{text g} is unique.
+*}
+
+lemma graph_domain_funct:
+  assumes uniq: "\<And>x y z. (x, y) \<in> g \<Longrightarrow> (x, z) \<in> g \<Longrightarrow> z = y"
+  shows "graph (domain g) (funct g) = g"
+  unfolding domain_def funct_def graph_def
+proof auto  (* FIXME !? *)
+  fix a b assume g: "(a, b) \<in> g"
+  from g show "(a, SOME y. (a, y) \<in> g) \<in> g" by (rule someI2)
+  from g show "\<exists>y. (a, y) \<in> g" ..
+  from g show "b = (SOME y. (a, y) \<in> g)"
+  proof (rule some_equality [symmetric])
+    fix y assume "(a, y) \<in> g"
+    with g show "y = b" by (rule uniq)
+  qed
+qed
+
+
+subsection {* Norm-preserving extensions of a function *}
+
+text {*
+  Given a linear form @{text f} on the space @{text F} and a seminorm
+  @{text p} on @{text E}. The set of all linear extensions of @{text
+  f}, to superspaces @{text H} of @{text F}, which are bounded by
+  @{text p}, is defined as follows.
+*}
+
+definition
+  norm_pres_extensions ::
+    "'a::{plus, minus, uminus, zero} set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> real)
+      \<Rightarrow> 'a graph set" where
+    "norm_pres_extensions E p F f
+      = {g. \<exists>H h. g = graph H h
+          \<and> linearform H h
+          \<and> H \<unlhd> E
+          \<and> F \<unlhd> H
+          \<and> graph F f \<subseteq> graph H h
+          \<and> (\<forall>x \<in> H. h x \<le> p x)}"
+
+lemma norm_pres_extensionE [elim]:
+  "g \<in> norm_pres_extensions E p F f
+  \<Longrightarrow> (\<And>H h. g = graph H h \<Longrightarrow> linearform H h
+        \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H \<Longrightarrow> graph F f \<subseteq> graph H h
+        \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x \<Longrightarrow> C) \<Longrightarrow> C"
+  unfolding norm_pres_extensions_def by blast
+
+lemma norm_pres_extensionI2 [intro]:
+  "linearform H h \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H
+    \<Longrightarrow> graph F f \<subseteq> graph H h \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x
+    \<Longrightarrow> graph H h \<in> norm_pres_extensions E p F f"
+  unfolding norm_pres_extensions_def by blast
+
+lemma norm_pres_extensionI:  (* FIXME ? *)
+  "\<exists>H h. g = graph H h
+    \<and> linearform H h
+    \<and> H \<unlhd> E
+    \<and> F \<unlhd> H
+    \<and> graph F f \<subseteq> graph H h
+    \<and> (\<forall>x \<in> H. h x \<le> p x) \<Longrightarrow> g \<in> norm_pres_extensions E p F f"
+  unfolding norm_pres_extensions_def by blast
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/HahnBanach.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,510 @@
+(*  Title:      HOL/Real/HahnBanach/HahnBanach.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* The Hahn-Banach Theorem *}
+
+theory HahnBanach
+imports HahnBanachLemmas
+begin
+
+text {*
+  We present the proof of two different versions of the Hahn-Banach
+  Theorem, closely following \cite[\S36]{Heuser:1986}.
+*}
+
+subsection {* The Hahn-Banach Theorem for vector spaces *}
+
+text {*
+  \textbf{Hahn-Banach Theorem.} Let @{text F} be a subspace of a real
+  vector space @{text E}, let @{text p} be a semi-norm on @{text E},
+  and @{text f} be a linear form defined on @{text F} such that @{text
+  f} is bounded by @{text p}, i.e.  @{text "\<forall>x \<in> F. f x \<le> p x"}.  Then
+  @{text f} can be extended to a linear form @{text h} on @{text E}
+  such that @{text h} is norm-preserving, i.e. @{text h} is also
+  bounded by @{text p}.
+
+  \bigskip
+  \textbf{Proof Sketch.}
+  \begin{enumerate}
+
+  \item Define @{text M} as the set of norm-preserving extensions of
+  @{text f} to subspaces of @{text E}. The linear forms in @{text M}
+  are ordered by domain extension.
+
+  \item We show that every non-empty chain in @{text M} has an upper
+  bound in @{text M}.
+
+  \item With Zorn's Lemma we conclude that there is a maximal function
+  @{text g} in @{text M}.
+
+  \item The domain @{text H} of @{text g} is the whole space @{text
+  E}, as shown by classical contradiction:
+
+  \begin{itemize}
+
+  \item Assuming @{text g} is not defined on whole @{text E}, it can
+  still be extended in a norm-preserving way to a super-space @{text
+  H'} of @{text H}.
+
+  \item Thus @{text g} can not be maximal. Contradiction!
+
+  \end{itemize}
+  \end{enumerate}
+*}
+
+theorem HahnBanach:
+  assumes E: "vectorspace E" and "subspace F E"
+    and "seminorm E p" and "linearform F f"
+  assumes fp: "\<forall>x \<in> F. f x \<le> p x"
+  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)"
+    -- {* Let @{text E} be a vector space, @{text F} a subspace of @{text E}, @{text p} a seminorm on @{text E}, *}
+    -- {* and @{text f} a linear form on @{text F} such that @{text f} is bounded by @{text p}, *}
+    -- {* then @{text f} can be extended to a linear form @{text h} on @{text E} in a norm-preserving way. \skp *}
+proof -
+  interpret vectorspace [E] by fact
+  interpret subspace [F E] by fact
+  interpret seminorm [E p] by fact
+  interpret linearform [F f] by fact
+  def M \<equiv> "norm_pres_extensions E p F f"
+  then have M: "M = \<dots>" by (simp only:)
+  from E have F: "vectorspace F" ..
+  note FE = `F \<unlhd> E`
+  {
+    fix c assume cM: "c \<in> chain M" and ex: "\<exists>x. x \<in> c"
+    have "\<Union>c \<in> M"
+      -- {* Show that every non-empty chain @{text c} of @{text M} has an upper bound in @{text M}: *}
+      -- {* @{text "\<Union>c"} is greater than any element of the chain @{text c}, so it suffices to show @{text "\<Union>c \<in> M"}. *}
+      unfolding M_def
+    proof (rule norm_pres_extensionI)
+      let ?H = "domain (\<Union>c)"
+      let ?h = "funct (\<Union>c)"
+
+      have a: "graph ?H ?h = \<Union>c"
+      proof (rule graph_domain_funct)
+        fix x y z assume "(x, y) \<in> \<Union>c" and "(x, z) \<in> \<Union>c"
+        with M_def cM show "z = y" by (rule sup_definite)
+      qed
+      moreover from M cM a have "linearform ?H ?h"
+        by (rule sup_lf)
+      moreover from a M cM ex FE E have "?H \<unlhd> E"
+        by (rule sup_subE)
+      moreover from a M cM ex FE have "F \<unlhd> ?H"
+        by (rule sup_supF)
+      moreover from a M cM ex have "graph F f \<subseteq> graph ?H ?h"
+        by (rule sup_ext)
+      moreover from a M cM have "\<forall>x \<in> ?H. ?h x \<le> p x"
+        by (rule sup_norm_pres)
+      ultimately show "\<exists>H h. \<Union>c = graph H h
+          \<and> linearform H h
+          \<and> H \<unlhd> E
+          \<and> F \<unlhd> H
+          \<and> graph F f \<subseteq> graph H h
+          \<and> (\<forall>x \<in> H. h x \<le> p x)" by blast
+    qed
+  }
+  then have "\<exists>g \<in> M. \<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
+  -- {* With Zorn's Lemma we can conclude that there is a maximal element in @{text M}. \skp *}
+  proof (rule Zorn's_Lemma)
+      -- {* We show that @{text M} is non-empty: *}
+    show "graph F f \<in> M"
+      unfolding M_def
+    proof (rule norm_pres_extensionI2)
+      show "linearform F f" by fact
+      show "F \<unlhd> E" by fact
+      from F show "F \<unlhd> F" by (rule vectorspace.subspace_refl)
+      show "graph F f \<subseteq> graph F f" ..
+      show "\<forall>x\<in>F. f x \<le> p x" by fact
+    qed
+  qed
+  then obtain g where gM: "g \<in> M" and gx: "\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
+    by blast
+  from gM obtain H h where
+      g_rep: "g = graph H h"
+    and linearform: "linearform H h"
+    and HE: "H \<unlhd> E" and FH: "F \<unlhd> H"
+    and graphs: "graph F f \<subseteq> graph H h"
+    and hp: "\<forall>x \<in> H. h x \<le> p x" unfolding M_def ..
+      -- {* @{text g} is a norm-preserving extension of @{text f}, in other words: *}
+      -- {* @{text g} is the graph of some linear form @{text h} defined on a subspace @{text H} of @{text E}, *}
+      -- {* and @{text h} is an extension of @{text f} that is again bounded by @{text p}. \skp *}
+  from HE E have H: "vectorspace H"
+    by (rule subspace.vectorspace)
+
+  have HE_eq: "H = E"
+    -- {* We show that @{text h} is defined on whole @{text E} by classical contradiction. \skp *}
+  proof (rule classical)
+    assume neq: "H \<noteq> E"
+      -- {* Assume @{text h} is not defined on whole @{text E}. Then show that @{text h} can be extended *}
+      -- {* in a norm-preserving way to a function @{text h'} with the graph @{text g'}. \skp *}
+    have "\<exists>g' \<in> M. g \<subseteq> g' \<and> g \<noteq> g'"
+    proof -
+      from HE have "H \<subseteq> E" ..
+      with neq obtain x' where x'E: "x' \<in> E" and "x' \<notin> H" by blast
+      obtain x': "x' \<noteq> 0"
+      proof
+        show "x' \<noteq> 0"
+        proof
+          assume "x' = 0"
+          with H have "x' \<in> H" by (simp only: vectorspace.zero)
+          with `x' \<notin> H` show False by contradiction
+        qed
+      qed
+
+      def H' \<equiv> "H + lin x'"
+        -- {* Define @{text H'} as the direct sum of @{text H} and the linear closure of @{text x'}. \skp *}
+      have HH': "H \<unlhd> H'"
+      proof (unfold H'_def)
+        from x'E have "vectorspace (lin x')" ..
+        with H show "H \<unlhd> H + lin x'" ..
+      qed
+
+      obtain xi where
+        xi: "\<forall>y \<in> H. - p (y + x') - h y \<le> xi
+          \<and> xi \<le> p (y + x') - h y"
+        -- {* Pick a real number @{text \<xi>} that fulfills certain inequations; this will *}
+        -- {* be used to establish that @{text h'} is a norm-preserving extension of @{text h}.
+           \label{ex-xi-use}\skp *}
+      proof -
+        from H have "\<exists>xi. \<forall>y \<in> H. - p (y + x') - h y \<le> xi
+            \<and> xi \<le> p (y + x') - h y"
+        proof (rule ex_xi)
+          fix u v assume u: "u \<in> H" and v: "v \<in> H"
+          with HE have uE: "u \<in> E" and vE: "v \<in> E" by auto
+          from H u v linearform have "h v - h u = h (v - u)"
+            by (simp add: linearform.diff)
+          also from hp and H u v have "\<dots> \<le> p (v - u)"
+            by (simp only: vectorspace.diff_closed)
+          also from x'E uE vE have "v - u = x' + - x' + v + - u"
+            by (simp add: diff_eq1)
+          also from x'E uE vE have "\<dots> = v + x' + - (u + x')"
+            by (simp add: add_ac)
+          also from x'E uE vE have "\<dots> = (v + x') - (u + x')"
+            by (simp add: diff_eq1)
+          also from x'E uE vE E have "p \<dots> \<le> p (v + x') + p (u + x')"
+            by (simp add: diff_subadditive)
+          finally have "h v - h u \<le> p (v + x') + p (u + x')" .
+          then show "- p (u + x') - h u \<le> p (v + x') - h v" by simp
+        qed
+        then show thesis by (blast intro: that)
+      qed
+
+      def h' \<equiv> "\<lambda>x. let (y, a) =
+          SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H in h y + a * xi"
+        -- {* Define the extension @{text h'} of @{text h} to @{text H'} using @{text \<xi>}. \skp *}
+
+      have "g \<subseteq> graph H' h' \<and> g \<noteq> graph H' h'"
+        -- {* @{text h'} is an extension of @{text h} \dots \skp *}
+      proof
+        show "g \<subseteq> graph H' h'"
+        proof -
+          have  "graph H h \<subseteq> graph H' h'"
+          proof (rule graph_extI)
+            fix t assume t: "t \<in> H"
+            from E HE t have "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
+	      using `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` by (rule decomp_H'_H)
+            with h'_def show "h t = h' t" by (simp add: Let_def)
+          next
+            from HH' show "H \<subseteq> H'" ..
+          qed
+          with g_rep show ?thesis by (simp only:)
+        qed
+
+        show "g \<noteq> graph H' h'"
+        proof -
+          have "graph H h \<noteq> graph H' h'"
+          proof
+            assume eq: "graph H h = graph H' h'"
+            have "x' \<in> H'"
+	      unfolding H'_def
+            proof
+              from H show "0 \<in> H" by (rule vectorspace.zero)
+              from x'E show "x' \<in> lin x'" by (rule x_lin_x)
+              from x'E show "x' = 0 + x'" by simp
+            qed
+            then have "(x', h' x') \<in> graph H' h'" ..
+            with eq have "(x', h' x') \<in> graph H h" by (simp only:)
+            then have "x' \<in> H" ..
+            with `x' \<notin> H` show False by contradiction
+          qed
+          with g_rep show ?thesis by simp
+        qed
+      qed
+      moreover have "graph H' h' \<in> M"
+        -- {* and @{text h'} is norm-preserving. \skp *}
+      proof (unfold M_def)
+        show "graph H' h' \<in> norm_pres_extensions E p F f"
+        proof (rule norm_pres_extensionI2)
+          show "linearform H' h'"
+	    using h'_def H'_def HE linearform `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E
+	    by (rule h'_lf)
+          show "H' \<unlhd> E"
+	  unfolding H'_def
+          proof
+            show "H \<unlhd> E" by fact
+            show "vectorspace E" by fact
+            from x'E show "lin x' \<unlhd> E" ..
+          qed
+          from H `F \<unlhd> H` HH' show FH': "F \<unlhd> H'"
+            by (rule vectorspace.subspace_trans)
+          show "graph F f \<subseteq> graph H' h'"
+          proof (rule graph_extI)
+            fix x assume x: "x \<in> F"
+            with graphs have "f x = h x" ..
+            also have "\<dots> = h x + 0 * xi" by simp
+            also have "\<dots> = (let (y, a) = (x, 0) in h y + a * xi)"
+              by (simp add: Let_def)
+            also have "(x, 0) =
+                (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)"
+	      using E HE
+            proof (rule decomp_H'_H [symmetric])
+              from FH x show "x \<in> H" ..
+              from x' show "x' \<noteq> 0" .
+	      show "x' \<notin> H" by fact
+	      show "x' \<in> E" by fact
+            qed
+            also have
+              "(let (y, a) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)
+              in h y + a * xi) = h' x" by (simp only: h'_def)
+            finally show "f x = h' x" .
+          next
+            from FH' show "F \<subseteq> H'" ..
+          qed
+          show "\<forall>x \<in> H'. h' x \<le> p x"
+	    using h'_def H'_def `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E HE
+	      `seminorm E p` linearform and hp xi
+	    by (rule h'_norm_pres)
+        qed
+      qed
+      ultimately show ?thesis ..
+    qed
+    then have "\<not> (\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x)" by simp
+      -- {* So the graph @{text g} of @{text h} cannot be maximal. Contradiction! \skp *}
+    with gx show "H = E" by contradiction
+  qed
+
+  from HE_eq and linearform have "linearform E h"
+    by (simp only:)
+  moreover have "\<forall>x \<in> F. h x = f x"
+  proof
+    fix x assume "x \<in> F"
+    with graphs have "f x = h x" ..
+    then show "h x = f x" ..
+  qed
+  moreover from HE_eq and hp have "\<forall>x \<in> E. h x \<le> p x"
+    by (simp only:)
+  ultimately show ?thesis by blast
+qed
+
+
+subsection  {* Alternative formulation *}
+
+text {*
+  The following alternative formulation of the Hahn-Banach
+  Theorem\label{abs-HahnBanach} uses the fact that for a real linear
+  form @{text f} and a seminorm @{text p} the following inequations
+  are equivalent:\footnote{This was shown in lemma @{thm [source]
+  abs_ineq_iff} (see page \pageref{abs-ineq-iff}).}
+  \begin{center}
+  \begin{tabular}{lll}
+  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
+  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
+  \end{tabular}
+  \end{center}
+*}
+
+theorem abs_HahnBanach:
+  assumes E: "vectorspace E" and FE: "subspace F E"
+    and lf: "linearform F f" and sn: "seminorm E p"
+  assumes fp: "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
+  shows "\<exists>g. linearform E g
+    \<and> (\<forall>x \<in> F. g x = f x)
+    \<and> (\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x)"
+proof -
+  interpret vectorspace [E] by fact
+  interpret subspace [F E] by fact
+  interpret linearform [F f] by fact
+  interpret seminorm [E p] by fact
+  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)"
+    using E FE sn lf
+  proof (rule HahnBanach)
+    show "\<forall>x \<in> F. f x \<le> p x"
+      using FE E sn lf and fp by (rule abs_ineq_iff [THEN iffD1])
+  qed
+  then obtain g where lg: "linearform E g" and *: "\<forall>x \<in> F. g x = f x"
+      and **: "\<forall>x \<in> E. g x \<le> p x" by blast
+  have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
+    using _ E sn lg **
+  proof (rule abs_ineq_iff [THEN iffD2])
+    show "E \<unlhd> E" ..
+  qed
+  with lg * show ?thesis by blast
+qed
+
+
+subsection {* The Hahn-Banach Theorem for normed spaces *}
+
+text {*
+  Every continuous linear form @{text f} on a subspace @{text F} of a
+  norm space @{text E}, can be extended to a continuous linear form
+  @{text g} on @{text E} such that @{text "\<parallel>f\<parallel> = \<parallel>g\<parallel>"}.
+*}
+
+theorem norm_HahnBanach:
+  fixes V and norm ("\<parallel>_\<parallel>")
+  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}"
+  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
+  defines "\<And>V f. \<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
+  assumes E_norm: "normed_vectorspace E norm" and FE: "subspace F E"
+    and linearform: "linearform F f" and "continuous F norm f"
+  shows "\<exists>g. linearform E g
+     \<and> continuous E norm g
+     \<and> (\<forall>x \<in> F. g x = f x)
+     \<and> \<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
+proof -
+  interpret normed_vectorspace [E norm] by fact
+  interpret normed_vectorspace_with_fn_norm [E norm B fn_norm]
+    by (auto simp: B_def fn_norm_def) intro_locales
+  interpret subspace [F E] by fact
+  interpret linearform [F f] by fact
+  interpret continuous [F norm f] by fact
+  have E: "vectorspace E" by intro_locales
+  have F: "vectorspace F" by rule intro_locales
+  have F_norm: "normed_vectorspace F norm"
+    using FE E_norm by (rule subspace_normed_vs)
+  have ge_zero: "0 \<le> \<parallel>f\<parallel>\<hyphen>F"
+    by (rule normed_vectorspace_with_fn_norm.fn_norm_ge_zero
+      [OF normed_vectorspace_with_fn_norm.intro,
+       OF F_norm `continuous F norm f` , folded B_def fn_norm_def])
+  txt {* We define a function @{text p} on @{text E} as follows:
+    @{text "p x = \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"} *}
+  def p \<equiv> "\<lambda>x. \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
+
+  txt {* @{text p} is a seminorm on @{text E}: *}
+  have q: "seminorm E p"
+  proof
+    fix x y a assume x: "x \<in> E" and y: "y \<in> E"
+    
+    txt {* @{text p} is positive definite: *}
+    have "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
+    moreover from x have "0 \<le> \<parallel>x\<parallel>" ..
+    ultimately show "0 \<le> p x"  
+      by (simp add: p_def zero_le_mult_iff)
+
+    txt {* @{text p} is absolutely homogenous: *}
+
+    show "p (a \<cdot> x) = \<bar>a\<bar> * p x"
+    proof -
+      have "p (a \<cdot> x) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>a \<cdot> x\<parallel>" by (simp only: p_def)
+      also from x have "\<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>" by (rule abs_homogenous)
+      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
+      also have "\<dots> = \<bar>a\<bar> * p x" by (simp only: p_def)
+      finally show ?thesis .
+    qed
+
+    txt {* Furthermore, @{text p} is subadditive: *}
+
+    show "p (x + y) \<le> p x + p y"
+    proof -
+      have "p (x + y) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel>" by (simp only: p_def)
+      also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
+      from x y have "\<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>" ..
+      with a have " \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>F * (\<parallel>x\<parallel> + \<parallel>y\<parallel>)"
+        by (simp add: mult_left_mono)
+      also have "\<dots> = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel> + \<parallel>f\<parallel>\<hyphen>F * \<parallel>y\<parallel>" by (simp only: right_distrib)
+      also have "\<dots> = p x + p y" by (simp only: p_def)
+      finally show ?thesis .
+    qed
+  qed
+
+  txt {* @{text f} is bounded by @{text p}. *}
+
+  have "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
+  proof
+    fix x assume "x \<in> F"
+    with `continuous F norm f` and linearform
+    show "\<bar>f x\<bar> \<le> p x"
+      unfolding p_def by (rule normed_vectorspace_with_fn_norm.fn_norm_le_cong
+        [OF normed_vectorspace_with_fn_norm.intro,
+         OF F_norm, folded B_def fn_norm_def])
+  qed
+
+  txt {* Using the fact that @{text p} is a seminorm and @{text f} is bounded
+    by @{text p} we can apply the Hahn-Banach Theorem for real vector
+    spaces. So @{text f} can be extended in a norm-preserving way to
+    some function @{text g} on the whole vector space @{text E}. *}
+
+  with E FE linearform q obtain g where
+      linearformE: "linearform E g"
+    and a: "\<forall>x \<in> F. g x = f x"
+    and b: "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
+    by (rule abs_HahnBanach [elim_format]) iprover
+
+  txt {* We furthermore have to show that @{text g} is also continuous: *}
+
+  have g_cont: "continuous E norm g" using linearformE
+  proof
+    fix x assume "x \<in> E"
+    with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
+      by (simp only: p_def)
+  qed
+
+  txt {* To complete the proof, we show that @{text "\<parallel>g\<parallel> = \<parallel>f\<parallel>"}. *}
+
+  have "\<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
+  proof (rule order_antisym)
+    txt {*
+      First we show @{text "\<parallel>g\<parallel> \<le> \<parallel>f\<parallel>"}.  The function norm @{text
+      "\<parallel>g\<parallel>"} is defined as the smallest @{text "c \<in> \<real>"} such that
+      \begin{center}
+      \begin{tabular}{l}
+      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
+      \end{tabular}
+      \end{center}
+      \noindent Furthermore holds
+      \begin{center}
+      \begin{tabular}{l}
+      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
+      \end{tabular}
+      \end{center}
+    *}
+
+    have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
+    proof
+      fix x assume "x \<in> E"
+      with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
+        by (simp only: p_def)
+    qed
+    from g_cont this ge_zero
+    show "\<parallel>g\<parallel>\<hyphen>E \<le> \<parallel>f\<parallel>\<hyphen>F"
+      by (rule fn_norm_least [of g, folded B_def fn_norm_def])
+
+    txt {* The other direction is achieved by a similar argument. *}
+
+    show "\<parallel>f\<parallel>\<hyphen>F \<le> \<parallel>g\<parallel>\<hyphen>E"
+    proof (rule normed_vectorspace_with_fn_norm.fn_norm_least
+	[OF normed_vectorspace_with_fn_norm.intro,
+	 OF F_norm, folded B_def fn_norm_def])
+      show "\<forall>x \<in> F. \<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
+      proof
+	fix x assume x: "x \<in> F"
+	from a x have "g x = f x" ..
+	then have "\<bar>f x\<bar> = \<bar>g x\<bar>" by (simp only:)
+	also from g_cont
+	have "\<dots> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
+	proof (rule fn_norm_le_cong [of g, folded B_def fn_norm_def])
+	  from FE x show "x \<in> E" ..
+	qed
+	finally show "\<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>" .
+      qed
+      show "0 \<le> \<parallel>g\<parallel>\<hyphen>E"
+	using g_cont
+	by (rule fn_norm_ge_zero [of g, folded B_def fn_norm_def])
+      show "continuous F norm f" by fact
+    qed
+  qed
+  with linearformE a g_cont show ?thesis by blast
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/HahnBanachExtLemmas.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,281 @@
+(*  Title:      HOL/Real/HahnBanach/HahnBanachExtLemmas.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Extending non-maximal functions *}
+
+theory HahnBanachExtLemmas
+imports FunctionNorm
+begin
+
+text {*
+  In this section the following context is presumed.  Let @{text E} be
+  a real vector space with a seminorm @{text q} on @{text E}. @{text
+  F} is a subspace of @{text E} and @{text f} a linear function on
+  @{text F}. We consider a subspace @{text H} of @{text E} that is a
+  superspace of @{text F} and a linear form @{text h} on @{text
+  H}. @{text H} is a not equal to @{text E} and @{text "x\<^sub>0"} is
+  an element in @{text "E - H"}.  @{text H} is extended to the direct
+  sum @{text "H' = H + lin x\<^sub>0"}, so for any @{text "x \<in> H'"}
+  the decomposition of @{text "x = y + a \<cdot> x"} with @{text "y \<in> H"} is
+  unique. @{text h'} is defined on @{text H'} by @{text "h' x = h y +
+  a \<cdot> \<xi>"} for a certain @{text \<xi>}.
+
+  Subsequently we show some properties of this extension @{text h'} of
+  @{text h}.
+
+  \medskip This lemma will be used to show the existence of a linear
+  extension of @{text f} (see page \pageref{ex-xi-use}). It is a
+  consequence of the completeness of @{text \<real>}. To show
+  \begin{center}
+  \begin{tabular}{l}
+  @{text "\<exists>\<xi>. \<forall>y \<in> F. a y \<le> \<xi> \<and> \<xi> \<le> b y"}
+  \end{tabular}
+  \end{center}
+  \noindent it suffices to show that
+  \begin{center}
+  \begin{tabular}{l}
+  @{text "\<forall>u \<in> F. \<forall>v \<in> F. a u \<le> b v"}
+  \end{tabular}
+  \end{center}
+*}
+
+lemma ex_xi:
+  assumes "vectorspace F"
+  assumes r: "\<And>u v. u \<in> F \<Longrightarrow> v \<in> F \<Longrightarrow> a u \<le> b v"
+  shows "\<exists>xi::real. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y"
+proof -
+  interpret vectorspace [F] by fact
+  txt {* From the completeness of the reals follows:
+    The set @{text "S = {a u. u \<in> F}"} has a supremum, if it is
+    non-empty and has an upper bound. *}
+
+  let ?S = "{a u | u. u \<in> F}"
+  have "\<exists>xi. lub ?S xi"
+  proof (rule real_complete)
+    have "a 0 \<in> ?S" by blast
+    then show "\<exists>X. X \<in> ?S" ..
+    have "\<forall>y \<in> ?S. y \<le> b 0"
+    proof
+      fix y assume y: "y \<in> ?S"
+      then obtain u where u: "u \<in> F" and y: "y = a u" by blast
+      from u and zero have "a u \<le> b 0" by (rule r)
+      with y show "y \<le> b 0" by (simp only:)
+    qed
+    then show "\<exists>u. \<forall>y \<in> ?S. y \<le> u" ..
+  qed
+  then obtain xi where xi: "lub ?S xi" ..
+  {
+    fix y assume "y \<in> F"
+    then have "a y \<in> ?S" by blast
+    with xi have "a y \<le> xi" by (rule lub.upper)
+  } moreover {
+    fix y assume y: "y \<in> F"
+    from xi have "xi \<le> b y"
+    proof (rule lub.least)
+      fix au assume "au \<in> ?S"
+      then obtain u where u: "u \<in> F" and au: "au = a u" by blast
+      from u y have "a u \<le> b y" by (rule r)
+      with au show "au \<le> b y" by (simp only:)
+    qed
+  } ultimately show "\<exists>xi. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y" by blast
+qed
+
+text {*
+  \medskip The function @{text h'} is defined as a @{text "h' x = h y
+  + a \<cdot> \<xi>"} where @{text "x = y + a \<cdot> \<xi>"} is a linear extension of
+  @{text h} to @{text H'}.
+*}
+
+lemma h'_lf:
+  assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
+      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
+    and H'_def: "H' \<equiv> H + lin x0"
+    and HE: "H \<unlhd> E"
+  assumes "linearform H h"
+  assumes x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
+  assumes E: "vectorspace E"
+  shows "linearform H' h'"
+proof -
+  interpret linearform [H h] by fact
+  interpret vectorspace [E] by fact
+  show ?thesis
+  proof
+    note E = `vectorspace E`
+    have H': "vectorspace H'"
+    proof (unfold H'_def)
+      from `x0 \<in> E`
+      have "lin x0 \<unlhd> E" ..
+      with HE show "vectorspace (H + lin x0)" using E ..
+    qed
+    {
+      fix x1 x2 assume x1: "x1 \<in> H'" and x2: "x2 \<in> H'"
+      show "h' (x1 + x2) = h' x1 + h' x2"
+      proof -
+	from H' x1 x2 have "x1 + x2 \<in> H'"
+          by (rule vectorspace.add_closed)
+	with x1 x2 obtain y y1 y2 a a1 a2 where
+          x1x2: "x1 + x2 = y + a \<cdot> x0" and y: "y \<in> H"
+          and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
+          and x2_rep: "x2 = y2 + a2 \<cdot> x0" and y2: "y2 \<in> H"
+          unfolding H'_def sum_def lin_def by blast
+	
+	have ya: "y1 + y2 = y \<and> a1 + a2 = a" using E HE _ y x0
+	proof (rule decomp_H') txt_raw {* \label{decomp-H-use} *}
+          from HE y1 y2 show "y1 + y2 \<in> H"
+            by (rule subspace.add_closed)
+          from x0 and HE y y1 y2
+          have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E"  "y2 \<in> E" by auto
+          with x1_rep x2_rep have "(y1 + y2) + (a1 + a2) \<cdot> x0 = x1 + x2"
+            by (simp add: add_ac add_mult_distrib2)
+          also note x1x2
+          finally show "(y1 + y2) + (a1 + a2) \<cdot> x0 = y + a \<cdot> x0" .
+	qed
+	
+	from h'_def x1x2 E HE y x0
+	have "h' (x1 + x2) = h y + a * xi"
+          by (rule h'_definite)
+	also have "\<dots> = h (y1 + y2) + (a1 + a2) * xi"
+          by (simp only: ya)
+	also from y1 y2 have "h (y1 + y2) = h y1 + h y2"
+          by simp
+	also have "\<dots> + (a1 + a2) * xi = (h y1 + a1 * xi) + (h y2 + a2 * xi)"
+          by (simp add: left_distrib)
+	also from h'_def x1_rep E HE y1 x0
+	have "h y1 + a1 * xi = h' x1"
+          by (rule h'_definite [symmetric])
+	also from h'_def x2_rep E HE y2 x0
+	have "h y2 + a2 * xi = h' x2"
+          by (rule h'_definite [symmetric])
+	finally show ?thesis .
+      qed
+    next
+      fix x1 c assume x1: "x1 \<in> H'"
+      show "h' (c \<cdot> x1) = c * (h' x1)"
+      proof -
+	from H' x1 have ax1: "c \<cdot> x1 \<in> H'"
+          by (rule vectorspace.mult_closed)
+	with x1 obtain y a y1 a1 where
+            cx1_rep: "c \<cdot> x1 = y + a \<cdot> x0" and y: "y \<in> H"
+          and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
+          unfolding H'_def sum_def lin_def by blast
+	
+	have ya: "c \<cdot> y1 = y \<and> c * a1 = a" using E HE _ y x0
+	proof (rule decomp_H')
+          from HE y1 show "c \<cdot> y1 \<in> H"
+            by (rule subspace.mult_closed)
+          from x0 and HE y y1
+          have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E" by auto
+          with x1_rep have "c \<cdot> y1 + (c * a1) \<cdot> x0 = c \<cdot> x1"
+            by (simp add: mult_assoc add_mult_distrib1)
+          also note cx1_rep
+          finally show "c \<cdot> y1 + (c * a1) \<cdot> x0 = y + a \<cdot> x0" .
+	qed
+	
+	from h'_def cx1_rep E HE y x0 have "h' (c \<cdot> x1) = h y + a * xi"
+          by (rule h'_definite)
+	also have "\<dots> = h (c \<cdot> y1) + (c * a1) * xi"
+          by (simp only: ya)
+	also from y1 have "h (c \<cdot> y1) = c * h y1"
+          by simp
+	also have "\<dots> + (c * a1) * xi = c * (h y1 + a1 * xi)"
+          by (simp only: right_distrib)
+	also from h'_def x1_rep E HE y1 x0 have "h y1 + a1 * xi = h' x1"
+          by (rule h'_definite [symmetric])
+	finally show ?thesis .
+      qed
+    }
+  qed
+qed
+
+text {* \medskip The linear extension @{text h'} of @{text h}
+  is bounded by the seminorm @{text p}. *}
+
+lemma h'_norm_pres:
+  assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
+      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
+    and H'_def: "H' \<equiv> H + lin x0"
+    and x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
+  assumes E: "vectorspace E" and HE: "subspace H E"
+    and "seminorm E p" and "linearform H h"
+  assumes a: "\<forall>y \<in> H. h y \<le> p y"
+    and a': "\<forall>y \<in> H. - p (y + x0) - h y \<le> xi \<and> xi \<le> p (y + x0) - h y"
+  shows "\<forall>x \<in> H'. h' x \<le> p x"
+proof -
+  interpret vectorspace [E] by fact
+  interpret subspace [H E] by fact
+  interpret seminorm [E p] by fact
+  interpret linearform [H h] by fact
+  show ?thesis
+  proof
+    fix x assume x': "x \<in> H'"
+    show "h' x \<le> p x"
+    proof -
+      from a' have a1: "\<forall>ya \<in> H. - p (ya + x0) - h ya \<le> xi"
+	and a2: "\<forall>ya \<in> H. xi \<le> p (ya + x0) - h ya" by auto
+      from x' obtain y a where
+          x_rep: "x = y + a \<cdot> x0" and y: "y \<in> H"
+	unfolding H'_def sum_def lin_def by blast
+      from y have y': "y \<in> E" ..
+      from y have ay: "inverse a \<cdot> y \<in> H" by simp
+      
+      from h'_def x_rep E HE y x0 have "h' x = h y + a * xi"
+	by (rule h'_definite)
+      also have "\<dots> \<le> p (y + a \<cdot> x0)"
+      proof (rule linorder_cases)
+	assume z: "a = 0"
+	then have "h y + a * xi = h y" by simp
+	also from a y have "\<dots> \<le> p y" ..
+	also from x0 y' z have "p y = p (y + a \<cdot> x0)" by simp
+	finally show ?thesis .
+      next
+	txt {* In the case @{text "a < 0"}, we use @{text "a\<^sub>1"}
+          with @{text ya} taken as @{text "y / a"}: *}
+	assume lz: "a < 0" then have nz: "a \<noteq> 0" by simp
+	from a1 ay
+	have "- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y) \<le> xi" ..
+	with lz have "a * xi \<le>
+          a * (- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
+          by (simp add: mult_left_mono_neg order_less_imp_le)
+	
+	also have "\<dots> =
+          - a * (p (inverse a \<cdot> y + x0)) - a * (h (inverse a \<cdot> y))"
+	  by (simp add: right_diff_distrib)
+	also from lz x0 y' have "- a * (p (inverse a \<cdot> y + x0)) =
+          p (a \<cdot> (inverse a \<cdot> y + x0))"
+          by (simp add: abs_homogenous)
+	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
+          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
+	also from nz y have "a * (h (inverse a \<cdot> y)) =  h y"
+          by simp
+	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
+	then show ?thesis by simp
+      next
+	txt {* In the case @{text "a > 0"}, we use @{text "a\<^sub>2"}
+          with @{text ya} taken as @{text "y / a"}: *}
+	assume gz: "0 < a" then have nz: "a \<noteq> 0" by simp
+	from a2 ay
+	have "xi \<le> p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y)" ..
+	with gz have "a * xi \<le>
+          a * (p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
+          by simp
+	also have "\<dots> = a * p (inverse a \<cdot> y + x0) - a * h (inverse a \<cdot> y)"
+	  by (simp add: right_diff_distrib)
+	also from gz x0 y'
+	have "a * p (inverse a \<cdot> y + x0) = p (a \<cdot> (inverse a \<cdot> y + x0))"
+          by (simp add: abs_homogenous)
+	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
+          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
+	also from nz y have "a * h (inverse a \<cdot> y) = h y"
+          by simp
+	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
+	then show ?thesis by simp
+      qed
+      also from x_rep have "\<dots> = p x" by (simp only:)
+      finally show ?thesis .
+    qed
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/HahnBanachLemmas.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,4 @@
+(*<*)
+theory HahnBanachLemmas imports HahnBanachSupLemmas HahnBanachExtLemmas begin
+end
+(*>*)
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/HahnBanachSupLemmas.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,446 @@
+(*  Title:      HOL/Real/HahnBanach/HahnBanachSupLemmas.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* The supremum w.r.t.~the function order *}
+
+theory HahnBanachSupLemmas
+imports FunctionNorm ZornLemma
+begin
+
+text {*
+  This section contains some lemmas that will be used in the proof of
+  the Hahn-Banach Theorem.  In this section the following context is
+  presumed.  Let @{text E} be a real vector space with a seminorm
+  @{text p} on @{text E}.  @{text F} is a subspace of @{text E} and
+  @{text f} a linear form on @{text F}. We consider a chain @{text c}
+  of norm-preserving extensions of @{text f}, such that @{text "\<Union>c =
+  graph H h"}.  We will show some properties about the limit function
+  @{text h}, i.e.\ the supremum of the chain @{text c}.
+
+  \medskip Let @{text c} be a chain of norm-preserving extensions of
+  the function @{text f} and let @{text "graph H h"} be the supremum
+  of @{text c}.  Every element in @{text H} is member of one of the
+  elements of the chain.
+*}
+lemmas [dest?] = chainD
+lemmas chainE2 [elim?] = chainD2 [elim_format, standard]
+
+lemma some_H'h't:
+  assumes M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and u: "graph H h = \<Union>c"
+    and x: "x \<in> H"
+  shows "\<exists>H' h'. graph H' h' \<in> c
+    \<and> (x, h x) \<in> graph H' h'
+    \<and> linearform H' h' \<and> H' \<unlhd> E
+    \<and> F \<unlhd> H' \<and> graph F f \<subseteq> graph H' h'
+    \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
+proof -
+  from x have "(x, h x) \<in> graph H h" ..
+  also from u have "\<dots> = \<Union>c" .
+  finally obtain g where gc: "g \<in> c" and gh: "(x, h x) \<in> g" by blast
+
+  from cM have "c \<subseteq> M" ..
+  with gc have "g \<in> M" ..
+  also from M have "\<dots> = norm_pres_extensions E p F f" .
+  finally obtain H' and h' where g: "g = graph H' h'"
+    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
+      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x" ..
+
+  from gc and g have "graph H' h' \<in> c" by (simp only:)
+  moreover from gh and g have "(x, h x) \<in> graph H' h'" by (simp only:)
+  ultimately show ?thesis using * by blast
+qed
+
+text {*
+  \medskip Let @{text c} be a chain of norm-preserving extensions of
+  the function @{text f} and let @{text "graph H h"} be the supremum
+  of @{text c}.  Every element in the domain @{text H} of the supremum
+  function is member of the domain @{text H'} of some function @{text
+  h'}, such that @{text h} extends @{text h'}.
+*}
+
+lemma some_H'h':
+  assumes M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and u: "graph H h = \<Union>c"
+    and x: "x \<in> H"
+  shows "\<exists>H' h'. x \<in> H' \<and> graph H' h' \<subseteq> graph H h
+    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
+    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
+proof -
+  from M cM u x obtain H' h' where
+      x_hx: "(x, h x) \<in> graph H' h'"
+    and c: "graph H' h' \<in> c"
+    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
+      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
+    by (rule some_H'h't [elim_format]) blast
+  from x_hx have "x \<in> H'" ..
+  moreover from cM u c have "graph H' h' \<subseteq> graph H h"
+    by (simp only: chain_ball_Union_upper)
+  ultimately show ?thesis using * by blast
+qed
+
+text {*
+  \medskip Any two elements @{text x} and @{text y} in the domain
+  @{text H} of the supremum function @{text h} are both in the domain
+  @{text H'} of some function @{text h'}, such that @{text h} extends
+  @{text h'}.
+*}
+
+lemma some_H'h'2:
+  assumes M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and u: "graph H h = \<Union>c"
+    and x: "x \<in> H"
+    and y: "y \<in> H"
+  shows "\<exists>H' h'. x \<in> H' \<and> y \<in> H'
+    \<and> graph H' h' \<subseteq> graph H h
+    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
+    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
+proof -
+  txt {* @{text y} is in the domain @{text H''} of some function @{text h''},
+  such that @{text h} extends @{text h''}. *}
+
+  from M cM u and y obtain H' h' where
+      y_hy: "(y, h y) \<in> graph H' h'"
+    and c': "graph H' h' \<in> c"
+    and * :
+      "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
+      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
+    by (rule some_H'h't [elim_format]) blast
+
+  txt {* @{text x} is in the domain @{text H'} of some function @{text h'},
+    such that @{text h} extends @{text h'}. *}
+
+  from M cM u and x obtain H'' h'' where
+      x_hx: "(x, h x) \<in> graph H'' h''"
+    and c'': "graph H'' h'' \<in> c"
+    and ** :
+      "linearform H'' h''"  "H'' \<unlhd> E"  "F \<unlhd> H''"
+      "graph F f \<subseteq> graph H'' h''"  "\<forall>x \<in> H''. h'' x \<le> p x"
+    by (rule some_H'h't [elim_format]) blast
+
+  txt {* Since both @{text h'} and @{text h''} are elements of the chain,
+    @{text h''} is an extension of @{text h'} or vice versa. Thus both
+    @{text x} and @{text y} are contained in the greater
+    one. \label{cases1}*}
+
+  from cM c'' c' have "graph H'' h'' \<subseteq> graph H' h' \<or> graph H' h' \<subseteq> graph H'' h''"
+    (is "?case1 \<or> ?case2") ..
+  then show ?thesis
+  proof
+    assume ?case1
+    have "(x, h x) \<in> graph H'' h''" by fact
+    also have "\<dots> \<subseteq> graph H' h'" by fact
+    finally have xh:"(x, h x) \<in> graph H' h'" .
+    then have "x \<in> H'" ..
+    moreover from y_hy have "y \<in> H'" ..
+    moreover from cM u and c' have "graph H' h' \<subseteq> graph H h"
+      by (simp only: chain_ball_Union_upper)
+    ultimately show ?thesis using * by blast
+  next
+    assume ?case2
+    from x_hx have "x \<in> H''" ..
+    moreover {
+      have "(y, h y) \<in> graph H' h'" by (rule y_hy)
+      also have "\<dots> \<subseteq> graph H'' h''" by fact
+      finally have "(y, h y) \<in> graph H'' h''" .
+    } then have "y \<in> H''" ..
+    moreover from cM u and c'' have "graph H'' h'' \<subseteq> graph H h"
+      by (simp only: chain_ball_Union_upper)
+    ultimately show ?thesis using ** by blast
+  qed
+qed
+
+text {*
+  \medskip The relation induced by the graph of the supremum of a
+  chain @{text c} is definite, i.~e.~t is the graph of a function.
+*}
+
+lemma sup_definite:
+  assumes M_def: "M \<equiv> norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and xy: "(x, y) \<in> \<Union>c"
+    and xz: "(x, z) \<in> \<Union>c"
+  shows "z = y"
+proof -
+  from cM have c: "c \<subseteq> M" ..
+  from xy obtain G1 where xy': "(x, y) \<in> G1" and G1: "G1 \<in> c" ..
+  from xz obtain G2 where xz': "(x, z) \<in> G2" and G2: "G2 \<in> c" ..
+
+  from G1 c have "G1 \<in> M" ..
+  then obtain H1 h1 where G1_rep: "G1 = graph H1 h1"
+    unfolding M_def by blast
+
+  from G2 c have "G2 \<in> M" ..
+  then obtain H2 h2 where G2_rep: "G2 = graph H2 h2"
+    unfolding M_def by blast
+
+  txt {* @{text "G\<^sub>1"} is contained in @{text "G\<^sub>2"}
+    or vice versa, since both @{text "G\<^sub>1"} and @{text
+    "G\<^sub>2"} are members of @{text c}. \label{cases2}*}
+
+  from cM G1 G2 have "G1 \<subseteq> G2 \<or> G2 \<subseteq> G1" (is "?case1 \<or> ?case2") ..
+  then show ?thesis
+  proof
+    assume ?case1
+    with xy' G2_rep have "(x, y) \<in> graph H2 h2" by blast
+    then have "y = h2 x" ..
+    also
+    from xz' G2_rep have "(x, z) \<in> graph H2 h2" by (simp only:)
+    then have "z = h2 x" ..
+    finally show ?thesis .
+  next
+    assume ?case2
+    with xz' G1_rep have "(x, z) \<in> graph H1 h1" by blast
+    then have "z = h1 x" ..
+    also
+    from xy' G1_rep have "(x, y) \<in> graph H1 h1" by (simp only:)
+    then have "y = h1 x" ..
+    finally show ?thesis ..
+  qed
+qed
+
+text {*
+  \medskip The limit function @{text h} is linear. Every element
+  @{text x} in the domain of @{text h} is in the domain of a function
+  @{text h'} in the chain of norm preserving extensions.  Furthermore,
+  @{text h} is an extension of @{text h'} so the function values of
+  @{text x} are identical for @{text h'} and @{text h}.  Finally, the
+  function @{text h'} is linear by construction of @{text M}.
+*}
+
+lemma sup_lf:
+  assumes M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and u: "graph H h = \<Union>c"
+  shows "linearform H h"
+proof
+  fix x y assume x: "x \<in> H" and y: "y \<in> H"
+  with M cM u obtain H' h' where
+        x': "x \<in> H'" and y': "y \<in> H'"
+      and b: "graph H' h' \<subseteq> graph H h"
+      and linearform: "linearform H' h'"
+      and subspace: "H' \<unlhd> E"
+    by (rule some_H'h'2 [elim_format]) blast
+
+  show "h (x + y) = h x + h y"
+  proof -
+    from linearform x' y' have "h' (x + y) = h' x + h' y"
+      by (rule linearform.add)
+    also from b x' have "h' x = h x" ..
+    also from b y' have "h' y = h y" ..
+    also from subspace x' y' have "x + y \<in> H'"
+      by (rule subspace.add_closed)
+    with b have "h' (x + y) = h (x + y)" ..
+    finally show ?thesis .
+  qed
+next
+  fix x a assume x: "x \<in> H"
+  with M cM u obtain H' h' where
+        x': "x \<in> H'"
+      and b: "graph H' h' \<subseteq> graph H h"
+      and linearform: "linearform H' h'"
+      and subspace: "H' \<unlhd> E"
+    by (rule some_H'h' [elim_format]) blast
+
+  show "h (a \<cdot> x) = a * h x"
+  proof -
+    from linearform x' have "h' (a \<cdot> x) = a * h' x"
+      by (rule linearform.mult)
+    also from b x' have "h' x = h x" ..
+    also from subspace x' have "a \<cdot> x \<in> H'"
+      by (rule subspace.mult_closed)
+    with b have "h' (a \<cdot> x) = h (a \<cdot> x)" ..
+    finally show ?thesis .
+  qed
+qed
+
+text {*
+  \medskip The limit of a non-empty chain of norm preserving
+  extensions of @{text f} is an extension of @{text f}, since every
+  element of the chain is an extension of @{text f} and the supremum
+  is an extension for every element of the chain.
+*}
+
+lemma sup_ext:
+  assumes graph: "graph H h = \<Union>c"
+    and M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and ex: "\<exists>x. x \<in> c"
+  shows "graph F f \<subseteq> graph H h"
+proof -
+  from ex obtain x where xc: "x \<in> c" ..
+  from cM have "c \<subseteq> M" ..
+  with xc have "x \<in> M" ..
+  with M have "x \<in> norm_pres_extensions E p F f"
+    by (simp only:)
+  then obtain G g where "x = graph G g" and "graph F f \<subseteq> graph G g" ..
+  then have "graph F f \<subseteq> x" by (simp only:)
+  also from xc have "\<dots> \<subseteq> \<Union>c" by blast
+  also from graph have "\<dots> = graph H h" ..
+  finally show ?thesis .
+qed
+
+text {*
+  \medskip The domain @{text H} of the limit function is a superspace
+  of @{text F}, since @{text F} is a subset of @{text H}. The
+  existence of the @{text 0} element in @{text F} and the closure
+  properties follow from the fact that @{text F} is a vector space.
+*}
+
+lemma sup_supF:
+  assumes graph: "graph H h = \<Union>c"
+    and M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and ex: "\<exists>x. x \<in> c"
+    and FE: "F \<unlhd> E"
+  shows "F \<unlhd> H"
+proof
+  from FE show "F \<noteq> {}" by (rule subspace.non_empty)
+  from graph M cM ex have "graph F f \<subseteq> graph H h" by (rule sup_ext)
+  then show "F \<subseteq> H" ..
+  fix x y assume "x \<in> F" and "y \<in> F"
+  with FE show "x + y \<in> F" by (rule subspace.add_closed)
+next
+  fix x a assume "x \<in> F"
+  with FE show "a \<cdot> x \<in> F" by (rule subspace.mult_closed)
+qed
+
+text {*
+  \medskip The domain @{text H} of the limit function is a subspace of
+  @{text E}.
+*}
+
+lemma sup_subE:
+  assumes graph: "graph H h = \<Union>c"
+    and M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+    and ex: "\<exists>x. x \<in> c"
+    and FE: "F \<unlhd> E"
+    and E: "vectorspace E"
+  shows "H \<unlhd> E"
+proof
+  show "H \<noteq> {}"
+  proof -
+    from FE E have "0 \<in> F" by (rule subspace.zero)
+    also from graph M cM ex FE have "F \<unlhd> H" by (rule sup_supF)
+    then have "F \<subseteq> H" ..
+    finally show ?thesis by blast
+  qed
+  show "H \<subseteq> E"
+  proof
+    fix x assume "x \<in> H"
+    with M cM graph
+    obtain H' h' where x: "x \<in> H'" and H'E: "H' \<unlhd> E"
+      by (rule some_H'h' [elim_format]) blast
+    from H'E have "H' \<subseteq> E" ..
+    with x show "x \<in> E" ..
+  qed
+  fix x y assume x: "x \<in> H" and y: "y \<in> H"
+  show "x + y \<in> H"
+  proof -
+    from M cM graph x y obtain H' h' where
+          x': "x \<in> H'" and y': "y \<in> H'" and H'E: "H' \<unlhd> E"
+        and graphs: "graph H' h' \<subseteq> graph H h"
+      by (rule some_H'h'2 [elim_format]) blast
+    from H'E x' y' have "x + y \<in> H'"
+      by (rule subspace.add_closed)
+    also from graphs have "H' \<subseteq> H" ..
+    finally show ?thesis .
+  qed
+next
+  fix x a assume x: "x \<in> H"
+  show "a \<cdot> x \<in> H"
+  proof -
+    from M cM graph x
+    obtain H' h' where x': "x \<in> H'" and H'E: "H' \<unlhd> E"
+        and graphs: "graph H' h' \<subseteq> graph H h"
+      by (rule some_H'h' [elim_format]) blast
+    from H'E x' have "a \<cdot> x \<in> H'" by (rule subspace.mult_closed)
+    also from graphs have "H' \<subseteq> H" ..
+    finally show ?thesis .
+  qed
+qed
+
+text {*
+  \medskip The limit function is bounded by the norm @{text p} as
+  well, since all elements in the chain are bounded by @{text p}.
+*}
+
+lemma sup_norm_pres:
+  assumes graph: "graph H h = \<Union>c"
+    and M: "M = norm_pres_extensions E p F f"
+    and cM: "c \<in> chain M"
+  shows "\<forall>x \<in> H. h x \<le> p x"
+proof
+  fix x assume "x \<in> H"
+  with M cM graph obtain H' h' where x': "x \<in> H'"
+      and graphs: "graph H' h' \<subseteq> graph H h"
+      and a: "\<forall>x \<in> H'. h' x \<le> p x"
+    by (rule some_H'h' [elim_format]) blast
+  from graphs x' have [symmetric]: "h' x = h x" ..
+  also from a x' have "h' x \<le> p x " ..
+  finally show "h x \<le> p x" .
+qed
+
+text {*
+  \medskip The following lemma is a property of linear forms on real
+  vector spaces. It will be used for the lemma @{text abs_HahnBanach}
+  (see page \pageref{abs-HahnBanach}). \label{abs-ineq-iff} For real
+  vector spaces the following inequations are equivalent:
+  \begin{center}
+  \begin{tabular}{lll}
+  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
+  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
+  \end{tabular}
+  \end{center}
+*}
+
+lemma abs_ineq_iff:
+  assumes "subspace H E" and "vectorspace E" and "seminorm E p"
+    and "linearform H h"
+  shows "(\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x) = (\<forall>x \<in> H. h x \<le> p x)" (is "?L = ?R")
+proof
+  interpret subspace [H E] by fact
+  interpret vectorspace [E] by fact
+  interpret seminorm [E p] by fact
+  interpret linearform [H h] by fact
+  have H: "vectorspace H" using `vectorspace E` ..
+  {
+    assume l: ?L
+    show ?R
+    proof
+      fix x assume x: "x \<in> H"
+      have "h x \<le> \<bar>h x\<bar>" by arith
+      also from l x have "\<dots> \<le> p x" ..
+      finally show "h x \<le> p x" .
+    qed
+  next
+    assume r: ?R
+    show ?L
+    proof
+      fix x assume x: "x \<in> H"
+      show "\<And>a b :: real. - a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> \<bar>b\<bar> \<le> a"
+        by arith
+      from `linearform H h` and H x
+      have "- h x = h (- x)" by (rule linearform.neg [symmetric])
+      also
+      from H x have "- x \<in> H" by (rule vectorspace.neg_closed)
+      with r have "h (- x) \<le> p (- x)" ..
+      also have "\<dots> = p x"
+	using `seminorm E p` `vectorspace E`
+      proof (rule seminorm.minus)
+        from x show "x \<in> E" ..
+      qed
+      finally have "- h x \<le> p x" .
+      then show "- p x \<le> h x" by simp
+      from r x show "h x \<le> p x" ..
+    qed
+  }
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/Linearform.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,61 @@
+(*  Title:      HOL/Real/HahnBanach/Linearform.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Linearforms *}
+
+theory Linearform
+imports VectorSpace
+begin
+
+text {*
+  A \emph{linear form} is a function on a vector space into the reals
+  that is additive and multiplicative.
+*}
+
+locale linearform = var V + var f +
+  constrains V :: "'a\<Colon>{minus, plus, zero, uminus} set"
+  assumes add [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x + y) = f x + f y"
+    and mult [iff]: "x \<in> V \<Longrightarrow> f (a \<cdot> x) = a * f x"
+
+declare linearform.intro [intro?]
+
+lemma (in linearform) neg [iff]:
+  assumes "vectorspace V"
+  shows "x \<in> V \<Longrightarrow> f (- x) = - f x"
+proof -
+  interpret vectorspace [V] by fact
+  assume x: "x \<in> V"
+  then have "f (- x) = f ((- 1) \<cdot> x)" by (simp add: negate_eq1)
+  also from x have "\<dots> = (- 1) * (f x)" by (rule mult)
+  also from x have "\<dots> = - (f x)" by simp
+  finally show ?thesis .
+qed
+
+lemma (in linearform) diff [iff]:
+  assumes "vectorspace V"
+  shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x - y) = f x - f y"
+proof -
+  interpret vectorspace [V] by fact
+  assume x: "x \<in> V" and y: "y \<in> V"
+  then have "x - y = x + - y" by (rule diff_eq1)
+  also have "f \<dots> = f x + f (- y)" by (rule add) (simp_all add: x y)
+  also have "f (- y) = - f y" using `vectorspace V` y by (rule neg)
+  finally show ?thesis by simp
+qed
+
+text {* Every linear form yields @{text 0} for the @{text 0} vector. *}
+
+lemma (in linearform) zero [iff]:
+  assumes "vectorspace V"
+  shows "f 0 = 0"
+proof -
+  interpret vectorspace [V] by fact
+  have "f 0 = f (0 - 0)" by simp
+  also have "\<dots> = f 0 - f 0" using `vectorspace V` by (rule diff) simp_all
+  also have "\<dots> = 0" by simp
+  finally show ?thesis .
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/NormedSpace.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,118 @@
+(*  Title:      HOL/Real/HahnBanach/NormedSpace.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Normed vector spaces *}
+
+theory NormedSpace
+imports Subspace
+begin
+
+subsection {* Quasinorms *}
+
+text {*
+  A \emph{seminorm} @{text "\<parallel>\<cdot>\<parallel>"} is a function on a real vector space
+  into the reals that has the following properties: it is positive
+  definite, absolute homogenous and subadditive.
+*}
+
+locale norm_syntax =
+  fixes norm :: "'a \<Rightarrow> real"    ("\<parallel>_\<parallel>")
+
+locale seminorm = var V + norm_syntax +
+  constrains V :: "'a\<Colon>{minus, plus, zero, uminus} set"
+  assumes ge_zero [iff?]: "x \<in> V \<Longrightarrow> 0 \<le> \<parallel>x\<parallel>"
+    and abs_homogenous [iff?]: "x \<in> V \<Longrightarrow> \<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>"
+    and subadditive [iff?]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> \<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>"
+
+declare seminorm.intro [intro?]
+
+lemma (in seminorm) diff_subadditive:
+  assumes "vectorspace V"
+  shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> \<parallel>x - y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>"
+proof -
+  interpret vectorspace [V] by fact
+  assume x: "x \<in> V" and y: "y \<in> V"
+  then have "x - y = x + - 1 \<cdot> y"
+    by (simp add: diff_eq2 negate_eq2a)
+  also from x y have "\<parallel>\<dots>\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>- 1 \<cdot> y\<parallel>"
+    by (simp add: subadditive)
+  also from y have "\<parallel>- 1 \<cdot> y\<parallel> = \<bar>- 1\<bar> * \<parallel>y\<parallel>"
+    by (rule abs_homogenous)
+  also have "\<dots> = \<parallel>y\<parallel>" by simp
+  finally show ?thesis .
+qed
+
+lemma (in seminorm) minus:
+  assumes "vectorspace V"
+  shows "x \<in> V \<Longrightarrow> \<parallel>- x\<parallel> = \<parallel>x\<parallel>"
+proof -
+  interpret vectorspace [V] by fact
+  assume x: "x \<in> V"
+  then have "- x = - 1 \<cdot> x" by (simp only: negate_eq1)
+  also from x have "\<parallel>\<dots>\<parallel> = \<bar>- 1\<bar> * \<parallel>x\<parallel>"
+    by (rule abs_homogenous)
+  also have "\<dots> = \<parallel>x\<parallel>" by simp
+  finally show ?thesis .
+qed
+
+
+subsection {* Norms *}
+
+text {*
+  A \emph{norm} @{text "\<parallel>\<cdot>\<parallel>"} is a seminorm that maps only the
+  @{text 0} vector to @{text 0}.
+*}
+
+locale norm = seminorm +
+  assumes zero_iff [iff]: "x \<in> V \<Longrightarrow> (\<parallel>x\<parallel> = 0) = (x = 0)"
+
+
+subsection {* Normed vector spaces *}
+
+text {*
+  A vector space together with a norm is called a \emph{normed
+  space}.
+*}
+
+locale normed_vectorspace = vectorspace + norm
+
+declare normed_vectorspace.intro [intro?]
+
+lemma (in normed_vectorspace) gt_zero [intro?]:
+  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> 0 < \<parallel>x\<parallel>"
+proof -
+  assume x: "x \<in> V" and neq: "x \<noteq> 0"
+  from x have "0 \<le> \<parallel>x\<parallel>" ..
+  also have [symmetric]: "\<dots> \<noteq> 0"
+  proof
+    assume "\<parallel>x\<parallel> = 0"
+    with x have "x = 0" by simp
+    with neq show False by contradiction
+  qed
+  finally show ?thesis .
+qed
+
+text {*
+  Any subspace of a normed vector space is again a normed vectorspace.
+*}
+
+lemma subspace_normed_vs [intro?]:
+  fixes F E norm
+  assumes "subspace F E" "normed_vectorspace E norm"
+  shows "normed_vectorspace F norm"
+proof -
+  interpret subspace [F E] by fact
+  interpret normed_vectorspace [E norm] by fact
+  show ?thesis
+  proof
+    show "vectorspace F" by (rule vectorspace) unfold_locales
+  next
+    have "NormedSpace.norm E norm" ..
+    with subset show "NormedSpace.norm F norm"
+      by (simp add: norm_def seminorm_def norm_axioms_def)
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/README.html	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,38 @@
+<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
+
+<!-- $Id$ -->
+
+<HTML>
+
+<HEAD>
+  <meta http-equiv="Content-Type" content="text/html; charset=iso-8859-1">
+  <TITLE>HOL/Real/HahnBanach/README</TITLE>
+</HEAD>
+
+<BODY>
+
+<H3>The Hahn-Banach Theorem for Real Vector Spaces (Isabelle/Isar)</H3>
+
+Author: Gertrud Bauer, Technische Universit&auml;t M&uuml;nchen<P>
+
+This directory contains the proof of the Hahn-Banach theorem for real vectorspaces,
+following H. Heuser, Funktionalanalysis, p. 228 -232.
+The Hahn-Banach theorem is one of the fundamental theorems of functioal analysis.
+It is a conclusion of Zorn's lemma.<P>
+
+Two different formaulations of the theorem are presented, one for general real vectorspaces
+and its application to normed vectorspaces. <P>
+
+The theorem says, that every continous linearform, defined on arbitrary subspaces
+(not only one-dimensional subspaces), can be extended to a continous linearform on
+the whole vectorspace.
+
+
+<HR>
+
+<ADDRESS>
+<A NAME="bauerg@in.tum.de" HREF="mailto:bauerg@in.tum.de">bauerg@in.tum.de</A>
+</ADDRESS>
+
+</BODY>
+</HTML>
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/ROOT.ML	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,8 @@
+(*  Title:      HOL/Real/HahnBanach/ROOT.ML
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+
+The Hahn-Banach theorem for real vector spaces (Isabelle/Isar).
+*)
+
+time_use_thy "HahnBanach";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/Subspace.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,514 @@
+(*  Title:      HOL/Real/HahnBanach/Subspace.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Subspaces *}
+
+theory Subspace
+imports VectorSpace
+begin
+
+subsection {* Definition *}
+
+text {*
+  A non-empty subset @{text U} of a vector space @{text V} is a
+  \emph{subspace} of @{text V}, iff @{text U} is closed under addition
+  and scalar multiplication.
+*}
+
+locale subspace = var U + var V +
+  constrains U :: "'a\<Colon>{minus, plus, zero, uminus} set"
+  assumes non_empty [iff, intro]: "U \<noteq> {}"
+    and subset [iff]: "U \<subseteq> V"
+    and add_closed [iff]: "x \<in> U \<Longrightarrow> y \<in> U \<Longrightarrow> x + y \<in> U"
+    and mult_closed [iff]: "x \<in> U \<Longrightarrow> a \<cdot> x \<in> U"
+
+notation (symbols)
+  subspace  (infix "\<unlhd>" 50)
+
+declare vectorspace.intro [intro?] subspace.intro [intro?]
+
+lemma subspace_subset [elim]: "U \<unlhd> V \<Longrightarrow> U \<subseteq> V"
+  by (rule subspace.subset)
+
+lemma (in subspace) subsetD [iff]: "x \<in> U \<Longrightarrow> x \<in> V"
+  using subset by blast
+
+lemma subspaceD [elim]: "U \<unlhd> V \<Longrightarrow> x \<in> U \<Longrightarrow> x \<in> V"
+  by (rule subspace.subsetD)
+
+lemma rev_subspaceD [elim?]: "x \<in> U \<Longrightarrow> U \<unlhd> V \<Longrightarrow> x \<in> V"
+  by (rule subspace.subsetD)
+
+lemma (in subspace) diff_closed [iff]:
+  assumes "vectorspace V"
+  assumes x: "x \<in> U" and y: "y \<in> U"
+  shows "x - y \<in> U"
+proof -
+  interpret vectorspace [V] by fact
+  from x y show ?thesis by (simp add: diff_eq1 negate_eq1)
+qed
+
+text {*
+  \medskip Similar as for linear spaces, the existence of the zero
+  element in every subspace follows from the non-emptiness of the
+  carrier set and by vector space laws.
+*}
+
+lemma (in subspace) zero [intro]:
+  assumes "vectorspace V"
+  shows "0 \<in> U"
+proof -
+  interpret vectorspace [V] by fact
+  have "U \<noteq> {}" by (rule U_V.non_empty)
+  then obtain x where x: "x \<in> U" by blast
+  then have "x \<in> V" .. then have "0 = x - x" by simp
+  also from `vectorspace V` x x have "\<dots> \<in> U" by (rule U_V.diff_closed)
+  finally show ?thesis .
+qed
+
+lemma (in subspace) neg_closed [iff]:
+  assumes "vectorspace V"
+  assumes x: "x \<in> U"
+  shows "- x \<in> U"
+proof -
+  interpret vectorspace [V] by fact
+  from x show ?thesis by (simp add: negate_eq1)
+qed
+
+text {* \medskip Further derived laws: every subspace is a vector space. *}
+
+lemma (in subspace) vectorspace [iff]:
+  assumes "vectorspace V"
+  shows "vectorspace U"
+proof -
+  interpret vectorspace [V] by fact
+  show ?thesis
+  proof
+    show "U \<noteq> {}" ..
+    fix x y z assume x: "x \<in> U" and y: "y \<in> U" and z: "z \<in> U"
+    fix a b :: real
+    from x y show "x + y \<in> U" by simp
+    from x show "a \<cdot> x \<in> U" by simp
+    from x y z show "(x + y) + z = x + (y + z)" by (simp add: add_ac)
+    from x y show "x + y = y + x" by (simp add: add_ac)
+    from x show "x - x = 0" by simp
+    from x show "0 + x = x" by simp
+    from x y show "a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y" by (simp add: distrib)
+    from x show "(a + b) \<cdot> x = a \<cdot> x + b \<cdot> x" by (simp add: distrib)
+    from x show "(a * b) \<cdot> x = a \<cdot> b \<cdot> x" by (simp add: mult_assoc)
+    from x show "1 \<cdot> x = x" by simp
+    from x show "- x = - 1 \<cdot> x" by (simp add: negate_eq1)
+    from x y show "x - y = x + - y" by (simp add: diff_eq1)
+  qed
+qed
+
+
+text {* The subspace relation is reflexive. *}
+
+lemma (in vectorspace) subspace_refl [intro]: "V \<unlhd> V"
+proof
+  show "V \<noteq> {}" ..
+  show "V \<subseteq> V" ..
+  fix x y assume x: "x \<in> V" and y: "y \<in> V"
+  fix a :: real
+  from x y show "x + y \<in> V" by simp
+  from x show "a \<cdot> x \<in> V" by simp
+qed
+
+text {* The subspace relation is transitive. *}
+
+lemma (in vectorspace) subspace_trans [trans]:
+  "U \<unlhd> V \<Longrightarrow> V \<unlhd> W \<Longrightarrow> U \<unlhd> W"
+proof
+  assume uv: "U \<unlhd> V" and vw: "V \<unlhd> W"
+  from uv show "U \<noteq> {}" by (rule subspace.non_empty)
+  show "U \<subseteq> W"
+  proof -
+    from uv have "U \<subseteq> V" by (rule subspace.subset)
+    also from vw have "V \<subseteq> W" by (rule subspace.subset)
+    finally show ?thesis .
+  qed
+  fix x y assume x: "x \<in> U" and y: "y \<in> U"
+  from uv and x y show "x + y \<in> U" by (rule subspace.add_closed)
+  from uv and x show "\<And>a. a \<cdot> x \<in> U" by (rule subspace.mult_closed)
+qed
+
+
+subsection {* Linear closure *}
+
+text {*
+  The \emph{linear closure} of a vector @{text x} is the set of all
+  scalar multiples of @{text x}.
+*}
+
+definition
+  lin :: "('a::{minus, plus, zero}) \<Rightarrow> 'a set" where
+  "lin x = {a \<cdot> x | a. True}"
+
+lemma linI [intro]: "y = a \<cdot> x \<Longrightarrow> y \<in> lin x"
+  unfolding lin_def by blast
+
+lemma linI' [iff]: "a \<cdot> x \<in> lin x"
+  unfolding lin_def by blast
+
+lemma linE [elim]: "x \<in> lin v \<Longrightarrow> (\<And>a::real. x = a \<cdot> v \<Longrightarrow> C) \<Longrightarrow> C"
+  unfolding lin_def by blast
+
+
+text {* Every vector is contained in its linear closure. *}
+
+lemma (in vectorspace) x_lin_x [iff]: "x \<in> V \<Longrightarrow> x \<in> lin x"
+proof -
+  assume "x \<in> V"
+  then have "x = 1 \<cdot> x" by simp
+  also have "\<dots> \<in> lin x" ..
+  finally show ?thesis .
+qed
+
+lemma (in vectorspace) "0_lin_x" [iff]: "x \<in> V \<Longrightarrow> 0 \<in> lin x"
+proof
+  assume "x \<in> V"
+  then show "0 = 0 \<cdot> x" by simp
+qed
+
+text {* Any linear closure is a subspace. *}
+
+lemma (in vectorspace) lin_subspace [intro]:
+  "x \<in> V \<Longrightarrow> lin x \<unlhd> V"
+proof
+  assume x: "x \<in> V"
+  then show "lin x \<noteq> {}" by (auto simp add: x_lin_x)
+  show "lin x \<subseteq> V"
+  proof
+    fix x' assume "x' \<in> lin x"
+    then obtain a where "x' = a \<cdot> x" ..
+    with x show "x' \<in> V" by simp
+  qed
+  fix x' x'' assume x': "x' \<in> lin x" and x'': "x'' \<in> lin x"
+  show "x' + x'' \<in> lin x"
+  proof -
+    from x' obtain a' where "x' = a' \<cdot> x" ..
+    moreover from x'' obtain a'' where "x'' = a'' \<cdot> x" ..
+    ultimately have "x' + x'' = (a' + a'') \<cdot> x"
+      using x by (simp add: distrib)
+    also have "\<dots> \<in> lin x" ..
+    finally show ?thesis .
+  qed
+  fix a :: real
+  show "a \<cdot> x' \<in> lin x"
+  proof -
+    from x' obtain a' where "x' = a' \<cdot> x" ..
+    with x have "a \<cdot> x' = (a * a') \<cdot> x" by (simp add: mult_assoc)
+    also have "\<dots> \<in> lin x" ..
+    finally show ?thesis .
+  qed
+qed
+
+
+text {* Any linear closure is a vector space. *}
+
+lemma (in vectorspace) lin_vectorspace [intro]:
+  assumes "x \<in> V"
+  shows "vectorspace (lin x)"
+proof -
+  from `x \<in> V` have "subspace (lin x) V"
+    by (rule lin_subspace)
+  from this and vectorspace_axioms show ?thesis
+    by (rule subspace.vectorspace)
+qed
+
+
+subsection {* Sum of two vectorspaces *}
+
+text {*
+  The \emph{sum} of two vectorspaces @{text U} and @{text V} is the
+  set of all sums of elements from @{text U} and @{text V}.
+*}
+
+instantiation "fun" :: (type, type) plus
+begin
+
+definition
+  sum_def: "plus_fun U V = {u + v | u v. u \<in> U \<and> v \<in> V}"  (* FIXME not fully general!? *)
+
+instance ..
+
+end
+
+lemma sumE [elim]:
+    "x \<in> U + V \<Longrightarrow> (\<And>u v. x = u + v \<Longrightarrow> u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> C) \<Longrightarrow> C"
+  unfolding sum_def by blast
+
+lemma sumI [intro]:
+    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> x = u + v \<Longrightarrow> x \<in> U + V"
+  unfolding sum_def by blast
+
+lemma sumI' [intro]:
+    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> u + v \<in> U + V"
+  unfolding sum_def by blast
+
+text {* @{text U} is a subspace of @{text "U + V"}. *}
+
+lemma subspace_sum1 [iff]:
+  assumes "vectorspace U" "vectorspace V"
+  shows "U \<unlhd> U + V"
+proof -
+  interpret vectorspace [U] by fact
+  interpret vectorspace [V] by fact
+  show ?thesis
+  proof
+    show "U \<noteq> {}" ..
+    show "U \<subseteq> U + V"
+    proof
+      fix x assume x: "x \<in> U"
+      moreover have "0 \<in> V" ..
+      ultimately have "x + 0 \<in> U + V" ..
+      with x show "x \<in> U + V" by simp
+    qed
+    fix x y assume x: "x \<in> U" and "y \<in> U"
+    then show "x + y \<in> U" by simp
+    from x show "\<And>a. a \<cdot> x \<in> U" by simp
+  qed
+qed
+
+text {* The sum of two subspaces is again a subspace. *}
+
+lemma sum_subspace [intro?]:
+  assumes "subspace U E" "vectorspace E" "subspace V E"
+  shows "U + V \<unlhd> E"
+proof -
+  interpret subspace [U E] by fact
+  interpret vectorspace [E] by fact
+  interpret subspace [V E] by fact
+  show ?thesis
+  proof
+    have "0 \<in> U + V"
+    proof
+      show "0 \<in> U" using `vectorspace E` ..
+      show "0 \<in> V" using `vectorspace E` ..
+      show "(0::'a) = 0 + 0" by simp
+    qed
+    then show "U + V \<noteq> {}" by blast
+    show "U + V \<subseteq> E"
+    proof
+      fix x assume "x \<in> U + V"
+      then obtain u v where "x = u + v" and
+	"u \<in> U" and "v \<in> V" ..
+      then show "x \<in> E" by simp
+    qed
+    fix x y assume x: "x \<in> U + V" and y: "y \<in> U + V"
+    show "x + y \<in> U + V"
+    proof -
+      from x obtain ux vx where "x = ux + vx" and "ux \<in> U" and "vx \<in> V" ..
+      moreover
+      from y obtain uy vy where "y = uy + vy" and "uy \<in> U" and "vy \<in> V" ..
+      ultimately
+      have "ux + uy \<in> U"
+	and "vx + vy \<in> V"
+	and "x + y = (ux + uy) + (vx + vy)"
+	using x y by (simp_all add: add_ac)
+      then show ?thesis ..
+    qed
+    fix a show "a \<cdot> x \<in> U + V"
+    proof -
+      from x obtain u v where "x = u + v" and "u \<in> U" and "v \<in> V" ..
+      then have "a \<cdot> u \<in> U" and "a \<cdot> v \<in> V"
+	and "a \<cdot> x = (a \<cdot> u) + (a \<cdot> v)" by (simp_all add: distrib)
+      then show ?thesis ..
+    qed
+  qed
+qed
+
+text{* The sum of two subspaces is a vectorspace. *}
+
+lemma sum_vs [intro?]:
+    "U \<unlhd> E \<Longrightarrow> V \<unlhd> E \<Longrightarrow> vectorspace E \<Longrightarrow> vectorspace (U + V)"
+  by (rule subspace.vectorspace) (rule sum_subspace)
+
+
+subsection {* Direct sums *}
+
+text {*
+  The sum of @{text U} and @{text V} is called \emph{direct}, iff the
+  zero element is the only common element of @{text U} and @{text
+  V}. For every element @{text x} of the direct sum of @{text U} and
+  @{text V} the decomposition in @{text "x = u + v"} with
+  @{text "u \<in> U"} and @{text "v \<in> V"} is unique.
+*}
+
+lemma decomp:
+  assumes "vectorspace E" "subspace U E" "subspace V E"
+  assumes direct: "U \<inter> V = {0}"
+    and u1: "u1 \<in> U" and u2: "u2 \<in> U"
+    and v1: "v1 \<in> V" and v2: "v2 \<in> V"
+    and sum: "u1 + v1 = u2 + v2"
+  shows "u1 = u2 \<and> v1 = v2"
+proof -
+  interpret vectorspace [E] by fact
+  interpret subspace [U E] by fact
+  interpret subspace [V E] by fact
+  show ?thesis
+  proof
+    have U: "vectorspace U"  (* FIXME: use interpret *)
+      using `subspace U E` `vectorspace E` by (rule subspace.vectorspace)
+    have V: "vectorspace V"
+      using `subspace V E` `vectorspace E` by (rule subspace.vectorspace)
+    from u1 u2 v1 v2 and sum have eq: "u1 - u2 = v2 - v1"
+      by (simp add: add_diff_swap)
+    from u1 u2 have u: "u1 - u2 \<in> U"
+      by (rule vectorspace.diff_closed [OF U])
+    with eq have v': "v2 - v1 \<in> U" by (simp only:)
+    from v2 v1 have v: "v2 - v1 \<in> V"
+      by (rule vectorspace.diff_closed [OF V])
+    with eq have u': " u1 - u2 \<in> V" by (simp only:)
+    
+    show "u1 = u2"
+    proof (rule add_minus_eq)
+      from u1 show "u1 \<in> E" ..
+      from u2 show "u2 \<in> E" ..
+      from u u' and direct show "u1 - u2 = 0" by blast
+    qed
+    show "v1 = v2"
+    proof (rule add_minus_eq [symmetric])
+      from v1 show "v1 \<in> E" ..
+      from v2 show "v2 \<in> E" ..
+      from v v' and direct show "v2 - v1 = 0" by blast
+    qed
+  qed
+qed
+
+text {*
+  An application of the previous lemma will be used in the proof of
+  the Hahn-Banach Theorem (see page \pageref{decomp-H-use}): for any
+  element @{text "y + a \<cdot> x\<^sub>0"} of the direct sum of a
+  vectorspace @{text H} and the linear closure of @{text "x\<^sub>0"}
+  the components @{text "y \<in> H"} and @{text a} are uniquely
+  determined.
+*}
+
+lemma decomp_H':
+  assumes "vectorspace E" "subspace H E"
+  assumes y1: "y1 \<in> H" and y2: "y2 \<in> H"
+    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
+    and eq: "y1 + a1 \<cdot> x' = y2 + a2 \<cdot> x'"
+  shows "y1 = y2 \<and> a1 = a2"
+proof -
+  interpret vectorspace [E] by fact
+  interpret subspace [H E] by fact
+  show ?thesis
+  proof
+    have c: "y1 = y2 \<and> a1 \<cdot> x' = a2 \<cdot> x'"
+    proof (rule decomp)
+      show "a1 \<cdot> x' \<in> lin x'" ..
+      show "a2 \<cdot> x' \<in> lin x'" ..
+      show "H \<inter> lin x' = {0}"
+      proof
+	show "H \<inter> lin x' \<subseteq> {0}"
+	proof
+          fix x assume x: "x \<in> H \<inter> lin x'"
+          then obtain a where xx': "x = a \<cdot> x'"
+            by blast
+          have "x = 0"
+          proof cases
+            assume "a = 0"
+            with xx' and x' show ?thesis by simp
+          next
+            assume a: "a \<noteq> 0"
+            from x have "x \<in> H" ..
+            with xx' have "inverse a \<cdot> a \<cdot> x' \<in> H" by simp
+            with a and x' have "x' \<in> H" by (simp add: mult_assoc2)
+            with `x' \<notin> H` show ?thesis by contradiction
+          qed
+          then show "x \<in> {0}" ..
+	qed
+	show "{0} \<subseteq> H \<inter> lin x'"
+	proof -
+          have "0 \<in> H" using `vectorspace E` ..
+          moreover have "0 \<in> lin x'" using `x' \<in> E` ..
+          ultimately show ?thesis by blast
+	qed
+      qed
+      show "lin x' \<unlhd> E" using `x' \<in> E` ..
+    qed (rule `vectorspace E`, rule `subspace H E`, rule y1, rule y2, rule eq)
+    then show "y1 = y2" ..
+    from c have "a1 \<cdot> x' = a2 \<cdot> x'" ..
+    with x' show "a1 = a2" by (simp add: mult_right_cancel)
+  qed
+qed
+
+text {*
+  Since for any element @{text "y + a \<cdot> x'"} of the direct sum of a
+  vectorspace @{text H} and the linear closure of @{text x'} the
+  components @{text "y \<in> H"} and @{text a} are unique, it follows from
+  @{text "y \<in> H"} that @{text "a = 0"}.
+*}
+
+lemma decomp_H'_H:
+  assumes "vectorspace E" "subspace H E"
+  assumes t: "t \<in> H"
+    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
+  shows "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
+proof -
+  interpret vectorspace [E] by fact
+  interpret subspace [H E] by fact
+  show ?thesis
+  proof (rule, simp_all only: split_paired_all split_conv)
+    from t x' show "t = t + 0 \<cdot> x' \<and> t \<in> H" by simp
+    fix y and a assume ya: "t = y + a \<cdot> x' \<and> y \<in> H"
+    have "y = t \<and> a = 0"
+    proof (rule decomp_H')
+      from ya x' show "y + a \<cdot> x' = t + 0 \<cdot> x'" by simp
+      from ya show "y \<in> H" ..
+    qed (rule `vectorspace E`, rule `subspace H E`, rule t, (rule x')+)
+    with t x' show "(y, a) = (y + a \<cdot> x', 0)" by simp
+  qed
+qed
+
+text {*
+  The components @{text "y \<in> H"} and @{text a} in @{text "y + a \<cdot> x'"}
+  are unique, so the function @{text h'} defined by
+  @{text "h' (y + a \<cdot> x') = h y + a \<cdot> \<xi>"} is definite.
+*}
+
+lemma h'_definite:
+  fixes H
+  assumes h'_def:
+    "h' \<equiv> (\<lambda>x. let (y, a) = SOME (y, a). (x = y + a \<cdot> x' \<and> y \<in> H)
+                in (h y) + a * xi)"
+    and x: "x = y + a \<cdot> x'"
+  assumes "vectorspace E" "subspace H E"
+  assumes y: "y \<in> H"
+    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
+  shows "h' x = h y + a * xi"
+proof -
+  interpret vectorspace [E] by fact
+  interpret subspace [H E] by fact
+  from x y x' have "x \<in> H + lin x'" by auto
+  have "\<exists>!p. (\<lambda>(y, a). x = y + a \<cdot> x' \<and> y \<in> H) p" (is "\<exists>!p. ?P p")
+  proof (rule ex_ex1I)
+    from x y show "\<exists>p. ?P p" by blast
+    fix p q assume p: "?P p" and q: "?P q"
+    show "p = q"
+    proof -
+      from p have xp: "x = fst p + snd p \<cdot> x' \<and> fst p \<in> H"
+        by (cases p) simp
+      from q have xq: "x = fst q + snd q \<cdot> x' \<and> fst q \<in> H"
+        by (cases q) simp
+      have "fst p = fst q \<and> snd p = snd q"
+      proof (rule decomp_H')
+        from xp show "fst p \<in> H" ..
+        from xq show "fst q \<in> H" ..
+        from xp and xq show "fst p + snd p \<cdot> x' = fst q + snd q \<cdot> x'"
+          by simp
+      qed (rule `vectorspace E`, rule `subspace H E`, (rule x')+)
+      then show ?thesis by (cases p, cases q) simp
+    qed
+  qed
+  then have eq: "(SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H) = (y, a)"
+    by (rule some1_equality) (simp add: x y)
+  with h'_def show "h' x = h y + a * xi" by (simp add: Let_def)
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/VectorSpace.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,417 @@
+(*  Title:      HOL/Real/HahnBanach/VectorSpace.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Vector spaces *}
+
+theory VectorSpace
+imports Real Bounds Zorn
+begin
+
+subsection {* Signature *}
+
+text {*
+  For the definition of real vector spaces a type @{typ 'a} of the
+  sort @{text "{plus, minus, zero}"} is considered, on which a real
+  scalar multiplication @{text \<cdot>} is declared.
+*}
+
+consts
+  prod  :: "real \<Rightarrow> 'a::{plus, minus, zero} \<Rightarrow> 'a"     (infixr "'(*')" 70)
+
+notation (xsymbols)
+  prod  (infixr "\<cdot>" 70)
+notation (HTML output)
+  prod  (infixr "\<cdot>" 70)
+
+
+subsection {* Vector space laws *}
+
+text {*
+  A \emph{vector space} is a non-empty set @{text V} of elements from
+  @{typ 'a} with the following vector space laws: The set @{text V} is
+  closed under addition and scalar multiplication, addition is
+  associative and commutative; @{text "- x"} is the inverse of @{text
+  x} w.~r.~t.~addition and @{text 0} is the neutral element of
+  addition.  Addition and multiplication are distributive; scalar
+  multiplication is associative and the real number @{text "1"} is
+  the neutral element of scalar multiplication.
+*}
+
+locale vectorspace = var V +
+  assumes non_empty [iff, intro?]: "V \<noteq> {}"
+    and add_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y \<in> V"
+    and mult_closed [iff]: "x \<in> V \<Longrightarrow> a \<cdot> x \<in> V"
+    and add_assoc: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y) + z = x + (y + z)"
+    and add_commute: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = y + x"
+    and diff_self [simp]: "x \<in> V \<Longrightarrow> x - x = 0"
+    and add_zero_left [simp]: "x \<in> V \<Longrightarrow> 0 + x = x"
+    and add_mult_distrib1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y"
+    and add_mult_distrib2: "x \<in> V \<Longrightarrow> (a + b) \<cdot> x = a \<cdot> x + b \<cdot> x"
+    and mult_assoc: "x \<in> V \<Longrightarrow> (a * b) \<cdot> x = a \<cdot> (b \<cdot> x)"
+    and mult_1 [simp]: "x \<in> V \<Longrightarrow> 1 \<cdot> x = x"
+    and negate_eq1: "x \<in> V \<Longrightarrow> - x = (- 1) \<cdot> x"
+    and diff_eq1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = x + - y"
+
+lemma (in vectorspace) negate_eq2: "x \<in> V \<Longrightarrow> (- 1) \<cdot> x = - x"
+  by (rule negate_eq1 [symmetric])
+
+lemma (in vectorspace) negate_eq2a: "x \<in> V \<Longrightarrow> -1 \<cdot> x = - x"
+  by (simp add: negate_eq1)
+
+lemma (in vectorspace) diff_eq2: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + - y = x - y"
+  by (rule diff_eq1 [symmetric])
+
+lemma (in vectorspace) diff_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y \<in> V"
+  by (simp add: diff_eq1 negate_eq1)
+
+lemma (in vectorspace) neg_closed [iff]: "x \<in> V \<Longrightarrow> - x \<in> V"
+  by (simp add: negate_eq1)
+
+lemma (in vectorspace) add_left_commute:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> x + (y + z) = y + (x + z)"
+proof -
+  assume xyz: "x \<in> V"  "y \<in> V"  "z \<in> V"
+  then have "x + (y + z) = (x + y) + z"
+    by (simp only: add_assoc)
+  also from xyz have "\<dots> = (y + x) + z" by (simp only: add_commute)
+  also from xyz have "\<dots> = y + (x + z)" by (simp only: add_assoc)
+  finally show ?thesis .
+qed
+
+theorems (in vectorspace) add_ac =
+  add_assoc add_commute add_left_commute
+
+
+text {* The existence of the zero element of a vector space
+  follows from the non-emptiness of carrier set. *}
+
+lemma (in vectorspace) zero [iff]: "0 \<in> V"
+proof -
+  from non_empty obtain x where x: "x \<in> V" by blast
+  then have "0 = x - x" by (rule diff_self [symmetric])
+  also from x x have "\<dots> \<in> V" by (rule diff_closed)
+  finally show ?thesis .
+qed
+
+lemma (in vectorspace) add_zero_right [simp]:
+  "x \<in> V \<Longrightarrow>  x + 0 = x"
+proof -
+  assume x: "x \<in> V"
+  from this and zero have "x + 0 = 0 + x" by (rule add_commute)
+  also from x have "\<dots> = x" by (rule add_zero_left)
+  finally show ?thesis .
+qed
+
+lemma (in vectorspace) mult_assoc2:
+    "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = (a * b) \<cdot> x"
+  by (simp only: mult_assoc)
+
+lemma (in vectorspace) diff_mult_distrib1:
+    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x - y) = a \<cdot> x - a \<cdot> y"
+  by (simp add: diff_eq1 negate_eq1 add_mult_distrib1 mult_assoc2)
+
+lemma (in vectorspace) diff_mult_distrib2:
+  "x \<in> V \<Longrightarrow> (a - b) \<cdot> x = a \<cdot> x - (b \<cdot> x)"
+proof -
+  assume x: "x \<in> V"
+  have " (a - b) \<cdot> x = (a + - b) \<cdot> x"
+    by (simp add: real_diff_def)
+  also from x have "\<dots> = a \<cdot> x + (- b) \<cdot> x"
+    by (rule add_mult_distrib2)
+  also from x have "\<dots> = a \<cdot> x + - (b \<cdot> x)"
+    by (simp add: negate_eq1 mult_assoc2)
+  also from x have "\<dots> = a \<cdot> x - (b \<cdot> x)"
+    by (simp add: diff_eq1)
+  finally show ?thesis .
+qed
+
+lemmas (in vectorspace) distrib =
+  add_mult_distrib1 add_mult_distrib2
+  diff_mult_distrib1 diff_mult_distrib2
+
+
+text {* \medskip Further derived laws: *}
+
+lemma (in vectorspace) mult_zero_left [simp]:
+  "x \<in> V \<Longrightarrow> 0 \<cdot> x = 0"
+proof -
+  assume x: "x \<in> V"
+  have "0 \<cdot> x = (1 - 1) \<cdot> x" by simp
+  also have "\<dots> = (1 + - 1) \<cdot> x" by simp
+  also from x have "\<dots> =  1 \<cdot> x + (- 1) \<cdot> x"
+    by (rule add_mult_distrib2)
+  also from x have "\<dots> = x + (- 1) \<cdot> x" by simp
+  also from x have "\<dots> = x + - x" by (simp add: negate_eq2a)
+  also from x have "\<dots> = x - x" by (simp add: diff_eq2)
+  also from x have "\<dots> = 0" by simp
+  finally show ?thesis .
+qed
+
+lemma (in vectorspace) mult_zero_right [simp]:
+  "a \<cdot> 0 = (0::'a)"
+proof -
+  have "a \<cdot> 0 = a \<cdot> (0 - (0::'a))" by simp
+  also have "\<dots> =  a \<cdot> 0 - a \<cdot> 0"
+    by (rule diff_mult_distrib1) simp_all
+  also have "\<dots> = 0" by simp
+  finally show ?thesis .
+qed
+
+lemma (in vectorspace) minus_mult_cancel [simp]:
+    "x \<in> V \<Longrightarrow> (- a) \<cdot> - x = a \<cdot> x"
+  by (simp add: negate_eq1 mult_assoc2)
+
+lemma (in vectorspace) add_minus_left_eq_diff:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + y = y - x"
+proof -
+  assume xy: "x \<in> V"  "y \<in> V"
+  then have "- x + y = y + - x" by (simp add: add_commute)
+  also from xy have "\<dots> = y - x" by (simp add: diff_eq1)
+  finally show ?thesis .
+qed
+
+lemma (in vectorspace) add_minus [simp]:
+    "x \<in> V \<Longrightarrow> x + - x = 0"
+  by (simp add: diff_eq2)
+
+lemma (in vectorspace) add_minus_left [simp]:
+    "x \<in> V \<Longrightarrow> - x + x = 0"
+  by (simp add: diff_eq2 add_commute)
+
+lemma (in vectorspace) minus_minus [simp]:
+    "x \<in> V \<Longrightarrow> - (- x) = x"
+  by (simp add: negate_eq1 mult_assoc2)
+
+lemma (in vectorspace) minus_zero [simp]:
+    "- (0::'a) = 0"
+  by (simp add: negate_eq1)
+
+lemma (in vectorspace) minus_zero_iff [simp]:
+  "x \<in> V \<Longrightarrow> (- x = 0) = (x = 0)"
+proof
+  assume x: "x \<in> V"
+  {
+    from x have "x = - (- x)" by (simp add: minus_minus)
+    also assume "- x = 0"
+    also have "- \<dots> = 0" by (rule minus_zero)
+    finally show "x = 0" .
+  next
+    assume "x = 0"
+    then show "- x = 0" by simp
+  }
+qed
+
+lemma (in vectorspace) add_minus_cancel [simp]:
+    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + (- x + y) = y"
+  by (simp add: add_assoc [symmetric] del: add_commute)
+
+lemma (in vectorspace) minus_add_cancel [simp]:
+    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + (x + y) = y"
+  by (simp add: add_assoc [symmetric] del: add_commute)
+
+lemma (in vectorspace) minus_add_distrib [simp]:
+    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - (x + y) = - x + - y"
+  by (simp add: negate_eq1 add_mult_distrib1)
+
+lemma (in vectorspace) diff_zero [simp]:
+    "x \<in> V \<Longrightarrow> x - 0 = x"
+  by (simp add: diff_eq1)
+
+lemma (in vectorspace) diff_zero_right [simp]:
+    "x \<in> V \<Longrightarrow> 0 - x = - x"
+  by (simp add: diff_eq1)
+
+lemma (in vectorspace) add_left_cancel:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y = x + z) = (y = z)"
+proof
+  assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V"
+  {
+    from y have "y = 0 + y" by simp
+    also from x y have "\<dots> = (- x + x) + y" by simp
+    also from x y have "\<dots> = - x + (x + y)"
+      by (simp add: add_assoc neg_closed)
+    also assume "x + y = x + z"
+    also from x z have "- x + (x + z) = - x + x + z"
+      by (simp add: add_assoc [symmetric] neg_closed)
+    also from x z have "\<dots> = z" by simp
+    finally show "y = z" .
+  next
+    assume "y = z"
+    then show "x + y = x + z" by (simp only:)
+  }
+qed
+
+lemma (in vectorspace) add_right_cancel:
+    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (y + x = z + x) = (y = z)"
+  by (simp only: add_commute add_left_cancel)
+
+lemma (in vectorspace) add_assoc_cong:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x' \<in> V \<Longrightarrow> y' \<in> V \<Longrightarrow> z \<in> V
+    \<Longrightarrow> x + y = x' + y' \<Longrightarrow> x + (y + z) = x' + (y' + z)"
+  by (simp only: add_assoc [symmetric])
+
+lemma (in vectorspace) mult_left_commute:
+    "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = b \<cdot> a \<cdot> x"
+  by (simp add: real_mult_commute mult_assoc2)
+
+lemma (in vectorspace) mult_zero_uniq:
+  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> a \<cdot> x = 0 \<Longrightarrow> a = 0"
+proof (rule classical)
+  assume a: "a \<noteq> 0"
+  assume x: "x \<in> V"  "x \<noteq> 0" and ax: "a \<cdot> x = 0"
+  from x a have "x = (inverse a * a) \<cdot> x" by simp
+  also from `x \<in> V` have "\<dots> = inverse a \<cdot> (a \<cdot> x)" by (rule mult_assoc)
+  also from ax have "\<dots> = inverse a \<cdot> 0" by simp
+  also have "\<dots> = 0" by simp
+  finally have "x = 0" .
+  with `x \<noteq> 0` show "a = 0" by contradiction
+qed
+
+lemma (in vectorspace) mult_left_cancel:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (a \<cdot> x = a \<cdot> y) = (x = y)"
+proof
+  assume x: "x \<in> V" and y: "y \<in> V" and a: "a \<noteq> 0"
+  from x have "x = 1 \<cdot> x" by simp
+  also from a have "\<dots> = (inverse a * a) \<cdot> x" by simp
+  also from x have "\<dots> = inverse a \<cdot> (a \<cdot> x)"
+    by (simp only: mult_assoc)
+  also assume "a \<cdot> x = a \<cdot> y"
+  also from a y have "inverse a \<cdot> \<dots> = y"
+    by (simp add: mult_assoc2)
+  finally show "x = y" .
+next
+  assume "x = y"
+  then show "a \<cdot> x = a \<cdot> y" by (simp only:)
+qed
+
+lemma (in vectorspace) mult_right_cancel:
+  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> (a \<cdot> x = b \<cdot> x) = (a = b)"
+proof
+  assume x: "x \<in> V" and neq: "x \<noteq> 0"
+  {
+    from x have "(a - b) \<cdot> x = a \<cdot> x - b \<cdot> x"
+      by (simp add: diff_mult_distrib2)
+    also assume "a \<cdot> x = b \<cdot> x"
+    with x have "a \<cdot> x - b \<cdot> x = 0" by simp
+    finally have "(a - b) \<cdot> x = 0" .
+    with x neq have "a - b = 0" by (rule mult_zero_uniq)
+    then show "a = b" by simp
+  next
+    assume "a = b"
+    then show "a \<cdot> x = b \<cdot> x" by (simp only:)
+  }
+qed
+
+lemma (in vectorspace) eq_diff_eq:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x = z - y) = (x + y = z)"
+proof
+  assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V"
+  {
+    assume "x = z - y"
+    then have "x + y = z - y + y" by simp
+    also from y z have "\<dots> = z + - y + y"
+      by (simp add: diff_eq1)
+    also have "\<dots> = z + (- y + y)"
+      by (rule add_assoc) (simp_all add: y z)
+    also from y z have "\<dots> = z + 0"
+      by (simp only: add_minus_left)
+    also from z have "\<dots> = z"
+      by (simp only: add_zero_right)
+    finally show "x + y = z" .
+  next
+    assume "x + y = z"
+    then have "z - y = (x + y) - y" by simp
+    also from x y have "\<dots> = x + y + - y"
+      by (simp add: diff_eq1)
+    also have "\<dots> = x + (y + - y)"
+      by (rule add_assoc) (simp_all add: x y)
+    also from x y have "\<dots> = x" by simp
+    finally show "x = z - y" ..
+  }
+qed
+
+lemma (in vectorspace) add_minus_eq_minus:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = 0 \<Longrightarrow> x = - y"
+proof -
+  assume x: "x \<in> V" and y: "y \<in> V"
+  from x y have "x = (- y + y) + x" by simp
+  also from x y have "\<dots> = - y + (x + y)" by (simp add: add_ac)
+  also assume "x + y = 0"
+  also from y have "- y + 0 = - y" by simp
+  finally show "x = - y" .
+qed
+
+lemma (in vectorspace) add_minus_eq:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = 0 \<Longrightarrow> x = y"
+proof -
+  assume x: "x \<in> V" and y: "y \<in> V"
+  assume "x - y = 0"
+  with x y have eq: "x + - y = 0" by (simp add: diff_eq1)
+  with _ _ have "x = - (- y)"
+    by (rule add_minus_eq_minus) (simp_all add: x y)
+  with x y show "x = y" by simp
+qed
+
+lemma (in vectorspace) add_diff_swap:
+  "a \<in> V \<Longrightarrow> b \<in> V \<Longrightarrow> c \<in> V \<Longrightarrow> d \<in> V \<Longrightarrow> a + b = c + d
+    \<Longrightarrow> a - c = d - b"
+proof -
+  assume vs: "a \<in> V"  "b \<in> V"  "c \<in> V"  "d \<in> V"
+    and eq: "a + b = c + d"
+  then have "- c + (a + b) = - c + (c + d)"
+    by (simp add: add_left_cancel)
+  also have "\<dots> = d" using `c \<in> V` `d \<in> V` by (rule minus_add_cancel)
+  finally have eq: "- c + (a + b) = d" .
+  from vs have "a - c = (- c + (a + b)) + - b"
+    by (simp add: add_ac diff_eq1)
+  also from vs eq have "\<dots>  = d + - b"
+    by (simp add: add_right_cancel)
+  also from vs have "\<dots> = d - b" by (simp add: diff_eq2)
+  finally show "a - c = d - b" .
+qed
+
+lemma (in vectorspace) vs_add_cancel_21:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> u \<in> V
+    \<Longrightarrow> (x + (y + z) = y + u) = (x + z = u)"
+proof
+  assume vs: "x \<in> V"  "y \<in> V"  "z \<in> V"  "u \<in> V"
+  {
+    from vs have "x + z = - y + y + (x + z)" by simp
+    also have "\<dots> = - y + (y + (x + z))"
+      by (rule add_assoc) (simp_all add: vs)
+    also from vs have "y + (x + z) = x + (y + z)"
+      by (simp add: add_ac)
+    also assume "x + (y + z) = y + u"
+    also from vs have "- y + (y + u) = u" by simp
+    finally show "x + z = u" .
+  next
+    assume "x + z = u"
+    with vs show "x + (y + z) = y + u"
+      by (simp only: add_left_commute [of x])
+  }
+qed
+
+lemma (in vectorspace) add_cancel_end:
+  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + (y + z) = y) = (x = - z)"
+proof
+  assume vs: "x \<in> V"  "y \<in> V"  "z \<in> V"
+  {
+    assume "x + (y + z) = y"
+    with vs have "(x + z) + y = 0 + y"
+      by (simp add: add_ac)
+    with vs have "x + z = 0"
+      by (simp only: add_right_cancel add_closed zero)
+    with vs show "x = - z" by (simp add: add_minus_eq_minus)
+  next
+    assume eq: "x = - z"
+    then have "x + (y + z) = - z + (y + z)" by simp
+    also have "\<dots> = y + (- z + z)"
+      by (rule add_left_commute) (simp_all add: vs)
+    also from vs have "\<dots> = y"  by simp
+    finally show "x + (y + z) = y" .
+  }
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/ZornLemma.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,58 @@
+(*  Title:      HOL/Real/HahnBanach/ZornLemma.thy
+    ID:         $Id$
+    Author:     Gertrud Bauer, TU Munich
+*)
+
+header {* Zorn's Lemma *}
+
+theory ZornLemma
+imports Zorn
+begin
+
+text {*
+  Zorn's Lemmas states: if every linear ordered subset of an ordered
+  set @{text S} has an upper bound in @{text S}, then there exists a
+  maximal element in @{text S}.  In our application, @{text S} is a
+  set of sets ordered by set inclusion. Since the union of a chain of
+  sets is an upper bound for all elements of the chain, the conditions
+  of Zorn's lemma can be modified: if @{text S} is non-empty, it
+  suffices to show that for every non-empty chain @{text c} in @{text
+  S} the union of @{text c} also lies in @{text S}.
+*}
+
+theorem Zorn's_Lemma:
+  assumes r: "\<And>c. c \<in> chain S \<Longrightarrow> \<exists>x. x \<in> c \<Longrightarrow> \<Union>c \<in> S"
+    and aS: "a \<in> S"
+  shows "\<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z \<longrightarrow> y = z"
+proof (rule Zorn_Lemma2)
+  show "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
+  proof
+    fix c assume "c \<in> chain S"
+    show "\<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
+    proof cases
+
+      txt {* If @{text c} is an empty chain, then every element in
+	@{text S} is an upper bound of @{text c}. *}
+
+      assume "c = {}"
+      with aS show ?thesis by fast
+
+      txt {* If @{text c} is non-empty, then @{text "\<Union>c"} is an upper
+	bound of @{text c}, lying in @{text S}. *}
+
+    next
+      assume "c \<noteq> {}"
+      show ?thesis
+      proof
+        show "\<forall>z \<in> c. z \<subseteq> \<Union>c" by fast
+        show "\<Union>c \<in> S"
+        proof (rule r)
+          from `c \<noteq> {}` show "\<exists>x. x \<in> c" by fast
+	  show "c \<in> chain S" by fact
+        qed
+      qed
+    qed
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/document/root.bib	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,27 @@
+
+@Book{Heuser:1986,
+  author = 	 {H. Heuser},
+  title = 	 {Funktionalanalysis: Theorie und Anwendung},
+  publisher = 	 {Teubner},
+  year = 	 1986
+}
+
+@InCollection{Narici:1996,
+  author = 	 {L. Narici and E. Beckenstein},
+  title = 	 {The {Hahn-Banach Theorem}: The Life and Times},
+  booktitle = 	 {Topology Atlas},
+  publisher =	 {York University, Toronto, Ontario, Canada},
+  year =	 1996,
+  note =	 {\url{http://at.yorku.ca/topology/preprint.htm} and
+                  \url{http://at.yorku.ca/p/a/a/a/16.htm}}
+}
+
+@Article{Nowak:1993,
+  author =       {B. Nowak and A. Trybulec},
+  title =	 {{Hahn-Banach} Theorem},
+  journal =      {Journal of Formalized Mathematics},
+  year =         {1993},
+  volume =       {5},
+  institution =  {University of Bialystok},
+  note =         {\url{http://mizar.uwb.edu.pl/JFM/Vol5/hahnban.html}}
+}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HahnBanach/document/root.tex	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,83 @@
+\documentclass[10pt,a4paper,twoside]{article}
+\usepackage{graphicx}
+\usepackage{latexsym,theorem}
+\usepackage{isabelle,isabellesym}
+\usepackage{pdfsetup} %last one!
+
+\isabellestyle{it}
+\urlstyle{rm}
+
+\newcommand{\isasymsup}{\isamath{\sup\,}}
+\newcommand{\skp}{\smallskip}
+
+
+\begin{document}
+
+\pagestyle{headings}
+\pagenumbering{arabic}
+
+\title{The Hahn-Banach Theorem \\ for Real Vector Spaces}
+\author{Gertrud Bauer \\ \url{http://www.in.tum.de/~bauerg/}}
+\maketitle
+
+\begin{abstract}
+  The Hahn-Banach Theorem is one of the most fundamental results in functional
+  analysis. We present a fully formal proof of two versions of the theorem,
+  one for general linear spaces and another for normed spaces.  This
+  development is based on simply-typed classical set-theory, as provided by
+  Isabelle/HOL.
+\end{abstract}
+
+
+\tableofcontents
+\parindent 0pt \parskip 0.5ex
+
+\clearpage
+\section{Preface}
+
+This is a fully formal proof of the Hahn-Banach Theorem. It closely follows
+the informal presentation given in Heuser's textbook \cite[{\S} 36]{Heuser:1986}.
+Another formal proof of the same theorem has been done in Mizar
+\cite{Nowak:1993}.  A general overview of the relevance and history of the
+Hahn-Banach Theorem is given by Narici and Beckenstein \cite{Narici:1996}.
+
+\medskip The document is structured as follows.  The first part contains
+definitions of basic notions of linear algebra: vector spaces, subspaces,
+normed spaces, continuous linear-forms, norm of functions and an order on
+functions by domain extension.  The second part contains some lemmas about the
+supremum (w.r.t.\ the function order) and extension of non-maximal functions.
+With these preliminaries, the main proof of the theorem (in its two versions)
+is conducted in the third part.  The dependencies of individual theories are
+as follows.
+
+\begin{center}
+  \includegraphics[scale=0.5]{session_graph}  
+\end{center}
+
+\clearpage
+\part {Basic Notions}
+
+\input{Bounds}
+\input{VectorSpace}
+\input{Subspace}
+\input{NormedSpace}
+\input{Linearform}
+\input{FunctionOrder}
+\input{FunctionNorm}
+\input{ZornLemma}
+
+\clearpage
+\part {Lemmas for the Proof}
+
+\input{HahnBanachSupLemmas}
+\input{HahnBanachExtLemmas}
+\input{HahnBanachLemmas}
+
+\clearpage
+\part {The Main Proof}
+
+\input{HahnBanach}
+\bibliographystyle{abbrv}
+\bibliography{root}
+
+\end{document}
--- a/src/HOL/Hyperreal/SEQ.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1136 +0,0 @@
-(*  Title       : SEQ.thy
-    Author      : Jacques D. Fleuriot
-    Copyright   : 1998  University of Cambridge
-    Description : Convergence of sequences and series
-    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
-    Additional contributions by Jeremy Avigad and Brian Huffman
-*)
-
-header {* Sequences and Convergence *}
-
-theory SEQ
-imports "../Real/RealVector" "../RComplete"
-begin
-
-definition
-  Zseq :: "[nat \<Rightarrow> 'a::real_normed_vector] \<Rightarrow> bool" where
-    --{*Standard definition of sequence converging to zero*}
-  [code del]: "Zseq X = (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. norm (X n) < r)"
-
-definition
-  LIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool"
-    ("((_)/ ----> (_))" [60, 60] 60) where
-    --{*Standard definition of convergence of sequence*}
-  [code del]: "X ----> L = (\<forall>r. 0 < r --> (\<exists>no. \<forall>n. no \<le> n --> norm (X n - L) < r))"
-
-definition
-  lim :: "(nat => 'a::real_normed_vector) => 'a" where
-    --{*Standard definition of limit using choice operator*}
-  "lim X = (THE L. X ----> L)"
-
-definition
-  convergent :: "(nat => 'a::real_normed_vector) => bool" where
-    --{*Standard definition of convergence*}
-  "convergent X = (\<exists>L. X ----> L)"
-
-definition
-  Bseq :: "(nat => 'a::real_normed_vector) => bool" where
-    --{*Standard definition for bounded sequence*}
-  [code del]: "Bseq X = (\<exists>K>0.\<forall>n. norm (X n) \<le> K)"
-
-definition
-  monoseq :: "(nat=>real)=>bool" where
-    --{*Definition for monotonicity*}
-  [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))"
-
-definition
-  subseq :: "(nat => nat) => bool" where
-    --{*Definition of subsequence*}
-  [code del]:   "subseq f = (\<forall>m. \<forall>n>m. (f m) < (f n))"
-
-definition
-  Cauchy :: "(nat => 'a::real_normed_vector) => bool" where
-    --{*Standard definition of the Cauchy condition*}
-  [code del]: "Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. norm (X m - X n) < e)"
-
-
-subsection {* Bounded Sequences *}
-
-lemma BseqI': assumes K: "\<And>n. norm (X n) \<le> K" shows "Bseq X"
-unfolding Bseq_def
-proof (intro exI conjI allI)
-  show "0 < max K 1" by simp
-next
-  fix n::nat
-  have "norm (X n) \<le> K" by (rule K)
-  thus "norm (X n) \<le> max K 1" by simp
-qed
-
-lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
-unfolding Bseq_def by auto
-
-lemma BseqI2': assumes K: "\<forall>n\<ge>N. norm (X n) \<le> K" shows "Bseq X"
-proof (rule BseqI')
-  let ?A = "norm ` X ` {..N}"
-  have 1: "finite ?A" by simp
-  fix n::nat
-  show "norm (X n) \<le> max K (Max ?A)"
-  proof (cases rule: linorder_le_cases)
-    assume "n \<ge> N"
-    hence "norm (X n) \<le> K" using K by simp
-    thus "norm (X n) \<le> max K (Max ?A)" by simp
-  next
-    assume "n \<le> N"
-    hence "norm (X n) \<in> ?A" by simp
-    with 1 have "norm (X n) \<le> Max ?A" by (rule Max_ge)
-    thus "norm (X n) \<le> max K (Max ?A)" by simp
-  qed
-qed
-
-lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))"
-unfolding Bseq_def by auto
-
-lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X"
-apply (erule BseqE)
-apply (rule_tac N="k" and K="K" in BseqI2')
-apply clarify
-apply (drule_tac x="n - k" in spec, simp)
-done
-
-
-subsection {* Sequences That Converge to Zero *}
-
-lemma ZseqI:
-  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r) \<Longrightarrow> Zseq X"
-unfolding Zseq_def by simp
-
-lemma ZseqD:
-  "\<lbrakk>Zseq X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r"
-unfolding Zseq_def by simp
-
-lemma Zseq_zero: "Zseq (\<lambda>n. 0)"
-unfolding Zseq_def by simp
-
-lemma Zseq_const_iff: "Zseq (\<lambda>n. k) = (k = 0)"
-unfolding Zseq_def by force
-
-lemma Zseq_norm_iff: "Zseq (\<lambda>n. norm (X n)) = Zseq (\<lambda>n. X n)"
-unfolding Zseq_def by simp
-
-lemma Zseq_imp_Zseq:
-  assumes X: "Zseq X"
-  assumes Y: "\<And>n. norm (Y n) \<le> norm (X n) * K"
-  shows "Zseq (\<lambda>n. Y n)"
-proof (cases)
-  assume K: "0 < K"
-  show ?thesis
-  proof (rule ZseqI)
-    fix r::real assume "0 < r"
-    hence "0 < r / K"
-      using K by (rule divide_pos_pos)
-    then obtain N where "\<forall>n\<ge>N. norm (X n) < r / K"
-      using ZseqD [OF X] by fast
-    hence "\<forall>n\<ge>N. norm (X n) * K < r"
-      by (simp add: pos_less_divide_eq K)
-    hence "\<forall>n\<ge>N. norm (Y n) < r"
-      by (simp add: order_le_less_trans [OF Y])
-    thus "\<exists>N. \<forall>n\<ge>N. norm (Y n) < r" ..
-  qed
-next
-  assume "\<not> 0 < K"
-  hence K: "K \<le> 0" by (simp only: linorder_not_less)
-  {
-    fix n::nat
-    have "norm (Y n) \<le> norm (X n) * K" by (rule Y)
-    also have "\<dots> \<le> norm (X n) * 0"
-      using K norm_ge_zero by (rule mult_left_mono)
-    finally have "norm (Y n) = 0" by simp
-  }
-  thus ?thesis by (simp add: Zseq_zero)
-qed
-
-lemma Zseq_le: "\<lbrakk>Zseq Y; \<forall>n. norm (X n) \<le> norm (Y n)\<rbrakk> \<Longrightarrow> Zseq X"
-by (erule_tac K="1" in Zseq_imp_Zseq, simp)
-
-lemma Zseq_add:
-  assumes X: "Zseq X"
-  assumes Y: "Zseq Y"
-  shows "Zseq (\<lambda>n. X n + Y n)"
-proof (rule ZseqI)
-  fix r::real assume "0 < r"
-  hence r: "0 < r / 2" by simp
-  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r/2"
-    using ZseqD [OF X r] by fast
-  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < r/2"
-    using ZseqD [OF Y r] by fast
-  show "\<exists>N. \<forall>n\<ge>N. norm (X n + Y n) < r"
-  proof (intro exI allI impI)
-    fix n assume n: "max M N \<le> n"
-    have "norm (X n + Y n) \<le> norm (X n) + norm (Y n)"
-      by (rule norm_triangle_ineq)
-    also have "\<dots> < r/2 + r/2"
-    proof (rule add_strict_mono)
-      from M n show "norm (X n) < r/2" by simp
-      from N n show "norm (Y n) < r/2" by simp
-    qed
-    finally show "norm (X n + Y n) < r" by simp
-  qed
-qed
-
-lemma Zseq_minus: "Zseq X \<Longrightarrow> Zseq (\<lambda>n. - X n)"
-unfolding Zseq_def by simp
-
-lemma Zseq_diff: "\<lbrakk>Zseq X; Zseq Y\<rbrakk> \<Longrightarrow> Zseq (\<lambda>n. X n - Y n)"
-by (simp only: diff_minus Zseq_add Zseq_minus)
-
-lemma (in bounded_linear) Zseq:
-  assumes X: "Zseq X"
-  shows "Zseq (\<lambda>n. f (X n))"
-proof -
-  obtain K where "\<And>x. norm (f x) \<le> norm x * K"
-    using bounded by fast
-  with X show ?thesis
-    by (rule Zseq_imp_Zseq)
-qed
-
-lemma (in bounded_bilinear) Zseq:
-  assumes X: "Zseq X"
-  assumes Y: "Zseq Y"
-  shows "Zseq (\<lambda>n. X n ** Y n)"
-proof (rule ZseqI)
-  fix r::real assume r: "0 < r"
-  obtain K where K: "0 < K"
-    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
-    using pos_bounded by fast
-  from K have K': "0 < inverse K"
-    by (rule positive_imp_inverse_positive)
-  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r"
-    using ZseqD [OF X r] by fast
-  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < inverse K"
-    using ZseqD [OF Y K'] by fast
-  show "\<exists>N. \<forall>n\<ge>N. norm (X n ** Y n) < r"
-  proof (intro exI allI impI)
-    fix n assume n: "max M N \<le> n"
-    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
-      by (rule norm_le)
-    also have "norm (X n) * norm (Y n) * K < r * inverse K * K"
-    proof (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero K)
-      from M n show Xn: "norm (X n) < r" by simp
-      from N n show Yn: "norm (Y n) < inverse K" by simp
-    qed
-    also from K have "r * inverse K * K = r" by simp
-    finally show "norm (X n ** Y n) < r" .
-  qed
-qed
-
-lemma (in bounded_bilinear) Zseq_prod_Bseq:
-  assumes X: "Zseq X"
-  assumes Y: "Bseq Y"
-  shows "Zseq (\<lambda>n. X n ** Y n)"
-proof -
-  obtain K where K: "0 \<le> K"
-    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
-    using nonneg_bounded by fast
-  obtain B where B: "0 < B"
-    and norm_Y: "\<And>n. norm (Y n) \<le> B"
-    using Y [unfolded Bseq_def] by fast
-  from X show ?thesis
-  proof (rule Zseq_imp_Zseq)
-    fix n::nat
-    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
-      by (rule norm_le)
-    also have "\<dots> \<le> norm (X n) * B * K"
-      by (intro mult_mono' order_refl norm_Y norm_ge_zero
-                mult_nonneg_nonneg K)
-    also have "\<dots> = norm (X n) * (B * K)"
-      by (rule mult_assoc)
-    finally show "norm (X n ** Y n) \<le> norm (X n) * (B * K)" .
-  qed
-qed
-
-lemma (in bounded_bilinear) Bseq_prod_Zseq:
-  assumes X: "Bseq X"
-  assumes Y: "Zseq Y"
-  shows "Zseq (\<lambda>n. X n ** Y n)"
-proof -
-  obtain K where K: "0 \<le> K"
-    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
-    using nonneg_bounded by fast
-  obtain B where B: "0 < B"
-    and norm_X: "\<And>n. norm (X n) \<le> B"
-    using X [unfolded Bseq_def] by fast
-  from Y show ?thesis
-  proof (rule Zseq_imp_Zseq)
-    fix n::nat
-    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
-      by (rule norm_le)
-    also have "\<dots> \<le> B * norm (Y n) * K"
-      by (intro mult_mono' order_refl norm_X norm_ge_zero
-                mult_nonneg_nonneg K)
-    also have "\<dots> = norm (Y n) * (B * K)"
-      by (simp only: mult_ac)
-    finally show "norm (X n ** Y n) \<le> norm (Y n) * (B * K)" .
-  qed
-qed
-
-lemma (in bounded_bilinear) Zseq_left:
-  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. X n ** a)"
-by (rule bounded_linear_left [THEN bounded_linear.Zseq])
-
-lemma (in bounded_bilinear) Zseq_right:
-  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. a ** X n)"
-by (rule bounded_linear_right [THEN bounded_linear.Zseq])
-
-lemmas Zseq_mult = mult.Zseq
-lemmas Zseq_mult_right = mult.Zseq_right
-lemmas Zseq_mult_left = mult.Zseq_left
-
-
-subsection {* Limits of Sequences *}
-
-lemma LIMSEQ_iff:
-      "(X ----> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)"
-by (rule LIMSEQ_def)
-
-lemma LIMSEQ_Zseq_iff: "((\<lambda>n. X n) ----> L) = Zseq (\<lambda>n. X n - L)"
-by (simp only: LIMSEQ_def Zseq_def)
-
-lemma LIMSEQ_I:
-  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X ----> L"
-by (simp add: LIMSEQ_def)
-
-lemma LIMSEQ_D:
-  "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
-by (simp add: LIMSEQ_def)
-
-lemma LIMSEQ_const: "(\<lambda>n. k) ----> k"
-by (simp add: LIMSEQ_def)
-
-lemma LIMSEQ_const_iff: "(\<lambda>n. k) ----> l = (k = l)"
-by (simp add: LIMSEQ_Zseq_iff Zseq_const_iff)
-
-lemma LIMSEQ_norm: "X ----> a \<Longrightarrow> (\<lambda>n. norm (X n)) ----> norm a"
-apply (simp add: LIMSEQ_def, safe)
-apply (drule_tac x="r" in spec, safe)
-apply (rule_tac x="no" in exI, safe)
-apply (drule_tac x="n" in spec, safe)
-apply (erule order_le_less_trans [OF norm_triangle_ineq3])
-done
-
-lemma LIMSEQ_ignore_initial_segment:
-  "f ----> a \<Longrightarrow> (\<lambda>n. f (n + k)) ----> a"
-apply (rule LIMSEQ_I)
-apply (drule (1) LIMSEQ_D)
-apply (erule exE, rename_tac N)
-apply (rule_tac x=N in exI)
-apply simp
-done
-
-lemma LIMSEQ_offset:
-  "(\<lambda>n. f (n + k)) ----> a \<Longrightarrow> f ----> a"
-apply (rule LIMSEQ_I)
-apply (drule (1) LIMSEQ_D)
-apply (erule exE, rename_tac N)
-apply (rule_tac x="N + k" in exI)
-apply clarify
-apply (drule_tac x="n - k" in spec)
-apply (simp add: le_diff_conv2)
-done
-
-lemma LIMSEQ_Suc: "f ----> l \<Longrightarrow> (\<lambda>n. f (Suc n)) ----> l"
-by (drule_tac k="1" in LIMSEQ_ignore_initial_segment, simp)
-
-lemma LIMSEQ_imp_Suc: "(\<lambda>n. f (Suc n)) ----> l \<Longrightarrow> f ----> l"
-by (rule_tac k="1" in LIMSEQ_offset, simp)
-
-lemma LIMSEQ_Suc_iff: "(\<lambda>n. f (Suc n)) ----> l = f ----> l"
-by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc)
-
-lemma add_diff_add:
-  fixes a b c d :: "'a::ab_group_add"
-  shows "(a + c) - (b + d) = (a - b) + (c - d)"
-by simp
-
-lemma minus_diff_minus:
-  fixes a b :: "'a::ab_group_add"
-  shows "(- a) - (- b) = - (a - b)"
-by simp
-
-lemma LIMSEQ_add: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) ----> a + b"
-by (simp only: LIMSEQ_Zseq_iff add_diff_add Zseq_add)
-
-lemma LIMSEQ_minus: "X ----> a \<Longrightarrow> (\<lambda>n. - X n) ----> - a"
-by (simp only: LIMSEQ_Zseq_iff minus_diff_minus Zseq_minus)
-
-lemma LIMSEQ_minus_cancel: "(\<lambda>n. - X n) ----> - a \<Longrightarrow> X ----> a"
-by (drule LIMSEQ_minus, simp)
-
-lemma LIMSEQ_diff: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) ----> a - b"
-by (simp add: diff_minus LIMSEQ_add LIMSEQ_minus)
-
-lemma LIMSEQ_unique: "\<lbrakk>X ----> a; X ----> b\<rbrakk> \<Longrightarrow> a = b"
-by (drule (1) LIMSEQ_diff, simp add: LIMSEQ_const_iff)
-
-lemma (in bounded_linear) LIMSEQ:
-  "X ----> a \<Longrightarrow> (\<lambda>n. f (X n)) ----> f a"
-by (simp only: LIMSEQ_Zseq_iff diff [symmetric] Zseq)
-
-lemma (in bounded_bilinear) LIMSEQ:
-  "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n ** Y n) ----> a ** b"
-by (simp only: LIMSEQ_Zseq_iff prod_diff_prod
-               Zseq_add Zseq Zseq_left Zseq_right)
-
-lemma LIMSEQ_mult:
-  fixes a b :: "'a::real_normed_algebra"
-  shows "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b"
-by (rule mult.LIMSEQ)
-
-lemma inverse_diff_inverse:
-  "\<lbrakk>(a::'a::division_ring) \<noteq> 0; b \<noteq> 0\<rbrakk>
-   \<Longrightarrow> inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
-by (simp add: ring_simps)
-
-lemma Bseq_inverse_lemma:
-  fixes x :: "'a::real_normed_div_algebra"
-  shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
-apply (subst nonzero_norm_inverse, clarsimp)
-apply (erule (1) le_imp_inverse_le)
-done
-
-lemma Bseq_inverse:
-  fixes a :: "'a::real_normed_div_algebra"
-  assumes X: "X ----> a"
-  assumes a: "a \<noteq> 0"
-  shows "Bseq (\<lambda>n. inverse (X n))"
-proof -
-  from a have "0 < norm a" by simp
-  hence "\<exists>r>0. r < norm a" by (rule dense)
-  then obtain r where r1: "0 < r" and r2: "r < norm a" by fast
-  obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> norm (X n - a) < r"
-    using LIMSEQ_D [OF X r1] by fast
-  show ?thesis
-  proof (rule BseqI2' [rule_format])
-    fix n assume n: "N \<le> n"
-    hence 1: "norm (X n - a) < r" by (rule N)
-    hence 2: "X n \<noteq> 0" using r2 by auto
-    hence "norm (inverse (X n)) = inverse (norm (X n))"
-      by (rule nonzero_norm_inverse)
-    also have "\<dots> \<le> inverse (norm a - r)"
-    proof (rule le_imp_inverse_le)
-      show "0 < norm a - r" using r2 by simp
-    next
-      have "norm a - norm (X n) \<le> norm (a - X n)"
-        by (rule norm_triangle_ineq2)
-      also have "\<dots> = norm (X n - a)"
-        by (rule norm_minus_commute)
-      also have "\<dots> < r" using 1 .
-      finally show "norm a - r \<le> norm (X n)" by simp
-    qed
-    finally show "norm (inverse (X n)) \<le> inverse (norm a - r)" .
-  qed
-qed
-
-lemma LIMSEQ_inverse_lemma:
-  fixes a :: "'a::real_normed_div_algebra"
-  shows "\<lbrakk>X ----> a; a \<noteq> 0; \<forall>n. X n \<noteq> 0\<rbrakk>
-         \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> inverse a"
-apply (subst LIMSEQ_Zseq_iff)
-apply (simp add: inverse_diff_inverse nonzero_imp_inverse_nonzero)
-apply (rule Zseq_minus)
-apply (rule Zseq_mult_left)
-apply (rule mult.Bseq_prod_Zseq)
-apply (erule (1) Bseq_inverse)
-apply (simp add: LIMSEQ_Zseq_iff)
-done
-
-lemma LIMSEQ_inverse:
-  fixes a :: "'a::real_normed_div_algebra"
-  assumes X: "X ----> a"
-  assumes a: "a \<noteq> 0"
-  shows "(\<lambda>n. inverse (X n)) ----> inverse a"
-proof -
-  from a have "0 < norm a" by simp
-  then obtain k where "\<forall>n\<ge>k. norm (X n - a) < norm a"
-    using LIMSEQ_D [OF X] by fast
-  hence "\<forall>n\<ge>k. X n \<noteq> 0" by auto
-  hence k: "\<forall>n. X (n + k) \<noteq> 0" by simp
-
-  from X have "(\<lambda>n. X (n + k)) ----> a"
-    by (rule LIMSEQ_ignore_initial_segment)
-  hence "(\<lambda>n. inverse (X (n + k))) ----> inverse a"
-    using a k by (rule LIMSEQ_inverse_lemma)
-  thus "(\<lambda>n. inverse (X n)) ----> inverse a"
-    by (rule LIMSEQ_offset)
-qed
-
-lemma LIMSEQ_divide:
-  fixes a b :: "'a::real_normed_field"
-  shows "\<lbrakk>X ----> a; Y ----> b; b \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>n. X n / Y n) ----> a / b"
-by (simp add: LIMSEQ_mult LIMSEQ_inverse divide_inverse)
-
-lemma LIMSEQ_pow:
-  fixes a :: "'a::{real_normed_algebra,recpower}"
-  shows "X ----> a \<Longrightarrow> (\<lambda>n. (X n) ^ m) ----> a ^ m"
-by (induct m) (simp_all add: power_Suc LIMSEQ_const LIMSEQ_mult)
-
-lemma LIMSEQ_setsum:
-  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
-  shows "(\<lambda>m. \<Sum>n\<in>S. X n m) ----> (\<Sum>n\<in>S. L n)"
-proof (cases "finite S")
-  case True
-  thus ?thesis using n
-  proof (induct)
-    case empty
-    show ?case
-      by (simp add: LIMSEQ_const)
-  next
-    case insert
-    thus ?case
-      by (simp add: LIMSEQ_add)
-  qed
-next
-  case False
-  thus ?thesis
-    by (simp add: LIMSEQ_const)
-qed
-
-lemma LIMSEQ_setprod:
-  fixes L :: "'a \<Rightarrow> 'b::{real_normed_algebra,comm_ring_1}"
-  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
-  shows "(\<lambda>m. \<Prod>n\<in>S. X n m) ----> (\<Prod>n\<in>S. L n)"
-proof (cases "finite S")
-  case True
-  thus ?thesis using n
-  proof (induct)
-    case empty
-    show ?case
-      by (simp add: LIMSEQ_const)
-  next
-    case insert
-    thus ?case
-      by (simp add: LIMSEQ_mult)
-  qed
-next
-  case False
-  thus ?thesis
-    by (simp add: setprod_def LIMSEQ_const)
-qed
-
-lemma LIMSEQ_add_const: "f ----> a ==> (%n.(f n + b)) ----> a + b"
-by (simp add: LIMSEQ_add LIMSEQ_const)
-
-(* FIXME: delete *)
-lemma LIMSEQ_add_minus:
-     "[| X ----> a; Y ----> b |] ==> (%n. X n + -Y n) ----> a + -b"
-by (simp only: LIMSEQ_add LIMSEQ_minus)
-
-lemma LIMSEQ_diff_const: "f ----> a ==> (%n.(f n  - b)) ----> a - b"
-by (simp add: LIMSEQ_diff LIMSEQ_const)
-
-lemma LIMSEQ_diff_approach_zero: 
-  "g ----> L ==> (%x. f x - g x) ----> 0  ==>
-     f ----> L"
-  apply (drule LIMSEQ_add)
-  apply assumption
-  apply simp
-done
-
-lemma LIMSEQ_diff_approach_zero2: 
-  "f ----> L ==> (%x. f x - g x) ----> 0  ==>
-     g ----> L";
-  apply (drule LIMSEQ_diff)
-  apply assumption
-  apply simp
-done
-
-text{*A sequence tends to zero iff its abs does*}
-lemma LIMSEQ_norm_zero: "((\<lambda>n. norm (X n)) ----> 0) = (X ----> 0)"
-by (simp add: LIMSEQ_def)
-
-lemma LIMSEQ_rabs_zero: "((%n. \<bar>f n\<bar>) ----> 0) = (f ----> (0::real))"
-by (simp add: LIMSEQ_def)
-
-lemma LIMSEQ_imp_rabs: "f ----> (l::real) ==> (%n. \<bar>f n\<bar>) ----> \<bar>l\<bar>"
-by (drule LIMSEQ_norm, simp)
-
-text{*An unbounded sequence's inverse tends to 0*}
-
-lemma LIMSEQ_inverse_zero:
-  "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
-apply (rule LIMSEQ_I)
-apply (drule_tac x="inverse r" in spec, safe)
-apply (rule_tac x="N" in exI, safe)
-apply (drule_tac x="n" in spec, safe)
-apply (frule positive_imp_inverse_positive)
-apply (frule (1) less_imp_inverse_less)
-apply (subgoal_tac "0 < X n", simp)
-apply (erule (1) order_less_trans)
-done
-
-text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}
-
-lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
-apply (rule LIMSEQ_inverse_zero, safe)
-apply (cut_tac x = r in reals_Archimedean2)
-apply (safe, rule_tac x = n in exI)
-apply (auto simp add: real_of_nat_Suc)
-done
-
-text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
-infinity is now easily proved*}
-
-lemma LIMSEQ_inverse_real_of_nat_add:
-     "(%n. r + inverse(real(Suc n))) ----> r"
-by (cut_tac LIMSEQ_add [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
-
-lemma LIMSEQ_inverse_real_of_nat_add_minus:
-     "(%n. r + -inverse(real(Suc n))) ----> r"
-by (cut_tac LIMSEQ_add_minus [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
-
-lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
-     "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
-by (cut_tac b=1 in
-        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat_add_minus], auto)
-
-lemma LIMSEQ_le_const:
-  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. a \<le> X n\<rbrakk> \<Longrightarrow> a \<le> x"
-apply (rule ccontr, simp only: linorder_not_le)
-apply (drule_tac r="a - x" in LIMSEQ_D, simp)
-apply clarsimp
-apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI1)
-apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI2)
-apply simp
-done
-
-lemma LIMSEQ_le_const2:
-  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. X n \<le> a\<rbrakk> \<Longrightarrow> x \<le> a"
-apply (subgoal_tac "- a \<le> - x", simp)
-apply (rule LIMSEQ_le_const)
-apply (erule LIMSEQ_minus)
-apply simp
-done
-
-lemma LIMSEQ_le:
-  "\<lbrakk>X ----> x; Y ----> y; \<exists>N. \<forall>n\<ge>N. X n \<le> Y n\<rbrakk> \<Longrightarrow> x \<le> (y::real)"
-apply (subgoal_tac "0 \<le> y - x", simp)
-apply (rule LIMSEQ_le_const)
-apply (erule (1) LIMSEQ_diff)
-apply (simp add: le_diff_eq)
-done
-
-
-subsection {* Convergence *}
-
-lemma limI: "X ----> L ==> lim X = L"
-apply (simp add: lim_def)
-apply (blast intro: LIMSEQ_unique)
-done
-
-lemma convergentD: "convergent X ==> \<exists>L. (X ----> L)"
-by (simp add: convergent_def)
-
-lemma convergentI: "(X ----> L) ==> convergent X"
-by (auto simp add: convergent_def)
-
-lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)"
-by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def)
-
-lemma convergent_minus_iff: "(convergent X) = (convergent (%n. -(X n)))"
-apply (simp add: convergent_def)
-apply (auto dest: LIMSEQ_minus)
-apply (drule LIMSEQ_minus, auto)
-done
-
-
-subsection {* Bounded Monotonic Sequences *}
-
-text{*Subsequence (alternative definition, (e.g. Hoskins)*}
-
-lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))"
-apply (simp add: subseq_def)
-apply (auto dest!: less_imp_Suc_add)
-apply (induct_tac k)
-apply (auto intro: less_trans)
-done
-
-lemma monoseq_Suc:
-   "monoseq X = ((\<forall>n. X n \<le> X (Suc n))
-                 | (\<forall>n. X (Suc n) \<le> X n))"
-apply (simp add: monoseq_def)
-apply (auto dest!: le_imp_less_or_eq)
-apply (auto intro!: lessI [THEN less_imp_le] dest!: less_imp_Suc_add)
-apply (induct_tac "ka")
-apply (auto intro: order_trans)
-apply (erule contrapos_np)
-apply (induct_tac "k")
-apply (auto intro: order_trans)
-done
-
-lemma monoI1: "\<forall>m. \<forall> n \<ge> m. X m \<le> X n ==> monoseq X"
-by (simp add: monoseq_def)
-
-lemma monoI2: "\<forall>m. \<forall> n \<ge> m. X n \<le> X m ==> monoseq X"
-by (simp add: monoseq_def)
-
-lemma mono_SucI1: "\<forall>n. X n \<le> X (Suc n) ==> monoseq X"
-by (simp add: monoseq_Suc)
-
-lemma mono_SucI2: "\<forall>n. X (Suc n) \<le> X n ==> monoseq X"
-by (simp add: monoseq_Suc)
-
-text{*Bounded Sequence*}
-
-lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)"
-by (simp add: Bseq_def)
-
-lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
-by (auto simp add: Bseq_def)
-
-lemma lemma_NBseq_def:
-     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) =
-      (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
-apply auto
- prefer 2 apply force
-apply (cut_tac x = K in reals_Archimedean2, clarify)
-apply (rule_tac x = n in exI, clarify)
-apply (drule_tac x = na in spec)
-apply (auto simp add: real_of_nat_Suc)
-done
-
-text{* alternative definition for Bseq *}
-lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
-apply (simp add: Bseq_def)
-apply (simp (no_asm) add: lemma_NBseq_def)
-done
-
-lemma lemma_NBseq_def2:
-     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
-apply (subst lemma_NBseq_def, auto)
-apply (rule_tac x = "Suc N" in exI)
-apply (rule_tac [2] x = N in exI)
-apply (auto simp add: real_of_nat_Suc)
- prefer 2 apply (blast intro: order_less_imp_le)
-apply (drule_tac x = n in spec, simp)
-done
-
-(* yet another definition for Bseq *)
-lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
-by (simp add: Bseq_def lemma_NBseq_def2)
-
-subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
-
-lemma Bseq_isUb:
-  "!!(X::nat=>real). Bseq X ==> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
-by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
-
-
-text{* Use completeness of reals (supremum property)
-   to show that any bounded sequence has a least upper bound*}
-
-lemma Bseq_isLub:
-  "!!(X::nat=>real). Bseq X ==>
-   \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
-by (blast intro: reals_complete Bseq_isUb)
-
-subsubsection{*A Bounded and Monotonic Sequence Converges*}
-
-lemma lemma_converg1:
-     "!!(X::nat=>real). [| \<forall>m. \<forall> n \<ge> m. X m \<le> X n;
-                  isLub (UNIV::real set) {x. \<exists>n. X n = x} (X ma)
-               |] ==> \<forall>n \<ge> ma. X n = X ma"
-apply safe
-apply (drule_tac y = "X n" in isLubD2)
-apply (blast dest: order_antisym)+
-done
-
-text{* The best of both worlds: Easier to prove this result as a standard
-   theorem and then use equivalence to "transfer" it into the
-   equivalent nonstandard form if needed!*}
-
-lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)"
-apply (simp add: LIMSEQ_def)
-apply (rule_tac x = "X m" in exI, safe)
-apply (rule_tac x = m in exI, safe)
-apply (drule spec, erule impE, auto)
-done
-
-lemma lemma_converg2:
-   "!!(X::nat=>real).
-    [| \<forall>m. X m ~= U;  isLub UNIV {x. \<exists>n. X n = x} U |] ==> \<forall>m. X m < U"
-apply safe
-apply (drule_tac y = "X m" in isLubD2)
-apply (auto dest!: order_le_imp_less_or_eq)
-done
-
-lemma lemma_converg3: "!!(X ::nat=>real). \<forall>m. X m \<le> U ==> isUb UNIV {x. \<exists>n. X n = x} U"
-by (rule setleI [THEN isUbI], auto)
-
-text{* FIXME: @{term "U - T < U"} is redundant *}
-lemma lemma_converg4: "!!(X::nat=> real).
-               [| \<forall>m. X m ~= U;
-                  isLub UNIV {x. \<exists>n. X n = x} U;
-                  0 < T;
-                  U + - T < U
-               |] ==> \<exists>m. U + -T < X m & X m < U"
-apply (drule lemma_converg2, assumption)
-apply (rule ccontr, simp)
-apply (simp add: linorder_not_less)
-apply (drule lemma_converg3)
-apply (drule isLub_le_isUb, assumption)
-apply (auto dest: order_less_le_trans)
-done
-
-text{*A standard proof of the theorem for monotone increasing sequence*}
-
-lemma Bseq_mono_convergent:
-     "[| Bseq X; \<forall>m. \<forall>n \<ge> m. X m \<le> X n |] ==> convergent (X::nat=>real)"
-apply (simp add: convergent_def)
-apply (frule Bseq_isLub, safe)
-apply (case_tac "\<exists>m. X m = U", auto)
-apply (blast dest: lemma_converg1 Bmonoseq_LIMSEQ)
-(* second case *)
-apply (rule_tac x = U in exI)
-apply (subst LIMSEQ_iff, safe)
-apply (frule lemma_converg2, assumption)
-apply (drule lemma_converg4, auto)
-apply (rule_tac x = m in exI, safe)
-apply (subgoal_tac "X m \<le> X n")
- prefer 2 apply blast
-apply (drule_tac x=n and P="%m. X m < U" in spec, arith)
-done
-
-lemma Bseq_minus_iff: "Bseq (%n. -(X n)) = Bseq X"
-by (simp add: Bseq_def)
-
-text{*Main monotonicity theorem*}
-lemma Bseq_monoseq_convergent: "[| Bseq X; monoseq X |] ==> convergent X"
-apply (simp add: monoseq_def, safe)
-apply (rule_tac [2] convergent_minus_iff [THEN ssubst])
-apply (drule_tac [2] Bseq_minus_iff [THEN ssubst])
-apply (auto intro!: Bseq_mono_convergent)
-done
-
-subsubsection{*A Few More Equivalence Theorems for Boundedness*}
-
-text{*alternative formulation for boundedness*}
-lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
-apply (unfold Bseq_def, safe)
-apply (rule_tac [2] x = "k + norm x" in exI)
-apply (rule_tac x = K in exI, simp)
-apply (rule exI [where x = 0], auto)
-apply (erule order_less_le_trans, simp)
-apply (drule_tac x=n in spec, fold diff_def)
-apply (drule order_trans [OF norm_triangle_ineq2])
-apply simp
-done
-
-text{*alternative formulation for boundedness*}
-lemma Bseq_iff3: "Bseq X = (\<exists>k > 0. \<exists>N. \<forall>n. norm(X(n) + -X(N)) \<le> k)"
-apply safe
-apply (simp add: Bseq_def, safe)
-apply (rule_tac x = "K + norm (X N)" in exI)
-apply auto
-apply (erule order_less_le_trans, simp)
-apply (rule_tac x = N in exI, safe)
-apply (drule_tac x = n in spec)
-apply (rule order_trans [OF norm_triangle_ineq], simp)
-apply (auto simp add: Bseq_iff2)
-done
-
-lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
-apply (simp add: Bseq_def)
-apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
-apply (drule_tac x = n in spec, arith)
-done
-
-
-subsection {* Cauchy Sequences *}
-
-lemma CauchyI:
-  "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
-by (simp add: Cauchy_def)
-
-lemma CauchyD:
-  "\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
-by (simp add: Cauchy_def)
-
-subsubsection {* Cauchy Sequences are Bounded *}
-
-text{*A Cauchy sequence is bounded -- this is the standard
-  proof mechanization rather than the nonstandard proof*}
-
-lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
-          ==>  \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
-apply (clarify, drule spec, drule (1) mp)
-apply (simp only: norm_minus_commute)
-apply (drule order_le_less_trans [OF norm_triangle_ineq2])
-apply simp
-done
-
-lemma Cauchy_Bseq: "Cauchy X ==> Bseq X"
-apply (simp add: Cauchy_def)
-apply (drule spec, drule mp, rule zero_less_one, safe)
-apply (drule_tac x="M" in spec, simp)
-apply (drule lemmaCauchy)
-apply (rule_tac k="M" in Bseq_offset)
-apply (simp add: Bseq_def)
-apply (rule_tac x="1 + norm (X M)" in exI)
-apply (rule conjI, rule order_less_le_trans [OF zero_less_one], simp)
-apply (simp add: order_less_imp_le)
-done
-
-subsubsection {* Cauchy Sequences are Convergent *}
-
-axclass banach \<subseteq> real_normed_vector
-  Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X"
-
-theorem LIMSEQ_imp_Cauchy:
-  assumes X: "X ----> a" shows "Cauchy X"
-proof (rule CauchyI)
-  fix e::real assume "0 < e"
-  hence "0 < e/2" by simp
-  with X have "\<exists>N. \<forall>n\<ge>N. norm (X n - a) < e/2" by (rule LIMSEQ_D)
-  then obtain N where N: "\<forall>n\<ge>N. norm (X n - a) < e/2" ..
-  show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < e"
-  proof (intro exI allI impI)
-    fix m assume "N \<le> m"
-    hence m: "norm (X m - a) < e/2" using N by fast
-    fix n assume "N \<le> n"
-    hence n: "norm (X n - a) < e/2" using N by fast
-    have "norm (X m - X n) = norm ((X m - a) - (X n - a))" by simp
-    also have "\<dots> \<le> norm (X m - a) + norm (X n - a)"
-      by (rule norm_triangle_ineq4)
-    also from m n have "\<dots> < e" by(simp add:field_simps)
-    finally show "norm (X m - X n) < e" .
-  qed
-qed
-
-lemma convergent_Cauchy: "convergent X \<Longrightarrow> Cauchy X"
-unfolding convergent_def
-by (erule exE, erule LIMSEQ_imp_Cauchy)
-
-text {*
-Proof that Cauchy sequences converge based on the one from
-http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html
-*}
-
-text {*
-  If sequence @{term "X"} is Cauchy, then its limit is the lub of
-  @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
-*}
-
-lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u"
-by (simp add: isUbI setleI)
-
-lemma real_abs_diff_less_iff:
-  "(\<bar>x - a\<bar> < (r::real)) = (a - r < x \<and> x < a + r)"
-by auto
-
-locale real_Cauchy =
-  fixes X :: "nat \<Rightarrow> real"
-  assumes X: "Cauchy X"
-  fixes S :: "real set"
-  defines S_def: "S \<equiv> {x::real. \<exists>N. \<forall>n\<ge>N. x < X n}"
-
-lemma real_CauchyI:
-  assumes "Cauchy X"
-  shows "real_Cauchy X"
-  proof qed (fact assms)
-
-lemma (in real_Cauchy) mem_S: "\<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S"
-by (unfold S_def, auto)
-
-lemma (in real_Cauchy) bound_isUb:
-  assumes N: "\<forall>n\<ge>N. X n < x"
-  shows "isUb UNIV S x"
-proof (rule isUb_UNIV_I)
-  fix y::real assume "y \<in> S"
-  hence "\<exists>M. \<forall>n\<ge>M. y < X n"
-    by (simp add: S_def)
-  then obtain M where "\<forall>n\<ge>M. y < X n" ..
-  hence "y < X (max M N)" by simp
-  also have "\<dots> < x" using N by simp
-  finally show "y \<le> x"
-    by (rule order_less_imp_le)
-qed
-
-lemma (in real_Cauchy) isLub_ex: "\<exists>u. isLub UNIV S u"
-proof (rule reals_complete)
-  obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < 1"
-    using CauchyD [OF X zero_less_one] by fast
-  hence N: "\<forall>n\<ge>N. norm (X n - X N) < 1" by simp
-  show "\<exists>x. x \<in> S"
-  proof
-    from N have "\<forall>n\<ge>N. X N - 1 < X n"
-      by (simp add: real_abs_diff_less_iff)
-    thus "X N - 1 \<in> S" by (rule mem_S)
-  qed
-  show "\<exists>u. isUb UNIV S u"
-  proof
-    from N have "\<forall>n\<ge>N. X n < X N + 1"
-      by (simp add: real_abs_diff_less_iff)
-    thus "isUb UNIV S (X N + 1)"
-      by (rule bound_isUb)
-  qed
-qed
-
-lemma (in real_Cauchy) isLub_imp_LIMSEQ:
-  assumes x: "isLub UNIV S x"
-  shows "X ----> x"
-proof (rule LIMSEQ_I)
-  fix r::real assume "0 < r"
-  hence r: "0 < r/2" by simp
-  obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. norm (X n - X m) < r/2"
-    using CauchyD [OF X r] by fast
-  hence "\<forall>n\<ge>N. norm (X n - X N) < r/2" by simp
-  hence N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2"
-    by (simp only: real_norm_def real_abs_diff_less_iff)
-
-  from N have "\<forall>n\<ge>N. X N - r/2 < X n" by fast
-  hence "X N - r/2 \<in> S" by (rule mem_S)
-  hence 1: "X N - r/2 \<le> x" using x isLub_isUb isUbD by fast
-
-  from N have "\<forall>n\<ge>N. X n < X N + r/2" by fast
-  hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb)
-  hence 2: "x \<le> X N + r/2" using x isLub_le_isUb by fast
-
-  show "\<exists>N. \<forall>n\<ge>N. norm (X n - x) < r"
-  proof (intro exI allI impI)
-    fix n assume n: "N \<le> n"
-    from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
-    thus "norm (X n - x) < r" using 1 2
-      by (simp add: real_abs_diff_less_iff)
-  qed
-qed
-
-lemma (in real_Cauchy) LIMSEQ_ex: "\<exists>x. X ----> x"
-proof -
-  obtain x where "isLub UNIV S x"
-    using isLub_ex by fast
-  hence "X ----> x"
-    by (rule isLub_imp_LIMSEQ)
-  thus ?thesis ..
-qed
-
-lemma real_Cauchy_convergent:
-  fixes X :: "nat \<Rightarrow> real"
-  shows "Cauchy X \<Longrightarrow> convergent X"
-unfolding convergent_def
-by (rule real_Cauchy.LIMSEQ_ex)
- (rule real_CauchyI)
-
-instance real :: banach
-by intro_classes (rule real_Cauchy_convergent)
-
-lemma Cauchy_convergent_iff:
-  fixes X :: "nat \<Rightarrow> 'a::banach"
-  shows "Cauchy X = convergent X"
-by (fast intro: Cauchy_convergent convergent_Cauchy)
-
-
-subsection {* Power Sequences *}
-
-text{*The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
-"x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
-  also fact that bounded and monotonic sequence converges.*}
-
-lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
-apply (simp add: Bseq_def)
-apply (rule_tac x = 1 in exI)
-apply (simp add: power_abs)
-apply (auto dest: power_mono)
-done
-
-lemma monoseq_realpow: "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
-apply (clarify intro!: mono_SucI2)
-apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
-done
-
-lemma convergent_realpow:
-  "[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)"
-by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
-
-lemma LIMSEQ_inverse_realpow_zero_lemma:
-  fixes x :: real
-  assumes x: "0 \<le> x"
-  shows "real n * x + 1 \<le> (x + 1) ^ n"
-apply (induct n)
-apply simp
-apply simp
-apply (rule order_trans)
-prefer 2
-apply (erule mult_left_mono)
-apply (rule add_increasing [OF x], simp)
-apply (simp add: real_of_nat_Suc)
-apply (simp add: ring_distribs)
-apply (simp add: mult_nonneg_nonneg x)
-done
-
-lemma LIMSEQ_inverse_realpow_zero:
-  "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) ----> 0"
-proof (rule LIMSEQ_inverse_zero [rule_format])
-  fix y :: real
-  assume x: "1 < x"
-  hence "0 < x - 1" by simp
-  hence "\<forall>y. \<exists>N::nat. y < real N * (x - 1)"
-    by (rule reals_Archimedean3)
-  hence "\<exists>N::nat. y < real N * (x - 1)" ..
-  then obtain N::nat where "y < real N * (x - 1)" ..
-  also have "\<dots> \<le> real N * (x - 1) + 1" by simp
-  also have "\<dots> \<le> (x - 1 + 1) ^ N"
-    by (rule LIMSEQ_inverse_realpow_zero_lemma, cut_tac x, simp)
-  also have "\<dots> = x ^ N" by simp
-  finally have "y < x ^ N" .
-  hence "\<forall>n\<ge>N. y < x ^ n"
-    apply clarify
-    apply (erule order_less_le_trans)
-    apply (erule power_increasing)
-    apply (rule order_less_imp_le [OF x])
-    done
-  thus "\<exists>N. \<forall>n\<ge>N. y < x ^ n" ..
-qed
-
-lemma LIMSEQ_realpow_zero:
-  "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
-proof (cases)
-  assume "x = 0"
-  hence "(\<lambda>n. x ^ Suc n) ----> 0" by (simp add: LIMSEQ_const)
-  thus ?thesis by (rule LIMSEQ_imp_Suc)
-next
-  assume "0 \<le> x" and "x \<noteq> 0"
-  hence x0: "0 < x" by simp
-  assume x1: "x < 1"
-  from x0 x1 have "1 < inverse x"
-    by (rule real_inverse_gt_one)
-  hence "(\<lambda>n. inverse (inverse x ^ n)) ----> 0"
-    by (rule LIMSEQ_inverse_realpow_zero)
-  thus ?thesis by (simp add: power_inverse)
-qed
-
-lemma LIMSEQ_power_zero:
-  fixes x :: "'a::{real_normed_algebra_1,recpower}"
-  shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
-apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
-apply (simp only: LIMSEQ_Zseq_iff, erule Zseq_le)
-apply (simp add: power_abs norm_power_ineq)
-done
-
-lemma LIMSEQ_divide_realpow_zero:
-  "1 < (x::real) ==> (%n. a / (x ^ n)) ----> 0"
-apply (cut_tac a = a and x1 = "inverse x" in
-        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_realpow_zero])
-apply (auto simp add: divide_inverse power_inverse)
-apply (simp add: inverse_eq_divide pos_divide_less_eq)
-done
-
-text{*Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}*}
-
-lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < (1::real) ==> (%n. \<bar>c\<bar> ^ n) ----> 0"
-by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
-
-lemma LIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < (1::real) ==> (%n. c ^ n) ----> 0"
-apply (rule LIMSEQ_rabs_zero [THEN iffD1])
-apply (auto intro: LIMSEQ_rabs_realpow_zero simp add: power_abs)
-done
-
-end
--- a/src/HOL/IsaMakefile	Mon Dec 29 13:23:53 2008 +0100
+++ b/src/HOL/IsaMakefile	Mon Dec 29 14:08:08 2008 +0100
@@ -261,7 +261,7 @@
 $(OUT)/HOL: ROOT.ML $(MAIN_DEPENDENCIES) \
   Complex_Main.thy \
   Complex.thy \
-  Complex/Fundamental_Theorem_Algebra.thy \
+  Fundamental_Theorem_Algebra.thy \
   Deriv.thy \
   Fact.thy \
   FrechetDeriv.thy \
@@ -271,11 +271,11 @@
   Log.thy \
   MacLaurin.thy \
   NthRoot.thy \
-  Hyperreal/SEQ.thy \
+  SEQ.thy \
   Series.thy \
   Taylor.thy \
   Transcendental.thy \
-  Library/Dense_Linear_Order.thy \
+  Dense_Linear_Order.thy \
   GCD.thy \
   Order_Relation.thy \
   Parity.thy \
@@ -287,7 +287,7 @@
   RealDef.thy \
   RealPow.thy \
   Real.thy \
-  Real/RealVector.thy \
+  RealVector.thy \
   Tools/float_syntax.ML \
   Tools/rat_arith.ML \
   Tools/real_arith.ML \
@@ -337,16 +337,16 @@
 HOL-HahnBanach: HOL $(LOG)/HOL-HahnBanach.gz
 
 $(LOG)/HOL-HahnBanach.gz: $(OUT)/HOL			\
-  Real/HahnBanach/Bounds.thy Real/HahnBanach/FunctionNorm.thy		\
-  Real/HahnBanach/FunctionOrder.thy Real/HahnBanach/HahnBanach.thy	\
-  Real/HahnBanach/HahnBanachExtLemmas.thy				\
-  Real/HahnBanach/HahnBanachSupLemmas.thy				\
-  Real/HahnBanach/Linearform.thy Real/HahnBanach/NormedSpace.thy	\
-  Real/HahnBanach/README.html Real/HahnBanach/ROOT.ML			\
-  Real/HahnBanach/Subspace.thy Real/HahnBanach/VectorSpace.thy		\
-  Real/HahnBanach/ZornLemma.thy Real/HahnBanach/document/root.bib	\
-  Real/HahnBanach/document/root.tex
-	@cd Real; $(ISABELLE_TOOL) usedir -g true $(OUT)/HOL HahnBanach
+  HahnBanach/Bounds.thy HahnBanach/FunctionNorm.thy		\
+  HahnBanach/FunctionOrder.thy HahnBanach/HahnBanach.thy	\
+  HahnBanach/HahnBanachExtLemmas.thy				\
+  HahnBanach/HahnBanachSupLemmas.thy				\
+  HahnBanach/Linearform.thy HahnBanach/NormedSpace.thy	\
+  HahnBanach/README.html HahnBanach/ROOT.ML			\
+  HahnBanach/Subspace.thy HahnBanach/VectorSpace.thy		\
+  HahnBanach/ZornLemma.thy HahnBanach/document/root.bib	\
+  HahnBanach/document/root.tex
+	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL HahnBanach
 
 
 ## HOL-Subst
--- a/src/HOL/Library/Dense_Linear_Order.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,880 +0,0 @@
-(*
-    ID:         $Id$
-    Author:     Amine Chaieb, TU Muenchen
-*)
-
-header {* Dense linear order without endpoints
-  and a quantifier elimination procedure in Ferrante and Rackoff style *}
-
-theory Dense_Linear_Order
-imports Plain "~~/src/HOL/Groebner_Basis"
-uses
-  "~~/src/HOL/Tools/Qelim/langford_data.ML"
-  "~~/src/HOL/Tools/Qelim/ferrante_rackoff_data.ML"
-  ("~~/src/HOL/Tools/Qelim/langford.ML")
-  ("~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML")
-begin
-
-setup {* Langford_Data.setup #> Ferrante_Rackoff_Data.setup *}
-
-context linorder
-begin
-
-lemma less_not_permute: "\<not> (x < y \<and> y < x)" by (simp add: not_less linear)
-
-lemma gather_simps: 
-  shows 
-  "(\<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)"
-  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)"
-  "(\<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))"
-  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
-
-lemma 
-  gather_start: "(\<exists>x. P x) \<equiv> (\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y\<in> {}. x < y) \<and> P x)" 
-  by simp
-
-text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"}*}
-lemma minf_lt:  "\<exists>z . \<forall>x. x < z \<longrightarrow> (x < t \<longleftrightarrow> True)" by auto
-lemma minf_gt: "\<exists>z . \<forall>x. x < z \<longrightarrow>  (t < x \<longleftrightarrow>  False)"
-  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
-
-lemma minf_le: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<le> t \<longleftrightarrow> True)" by (auto simp add: less_le)
-lemma minf_ge: "\<exists>z. \<forall>x. x < z \<longrightarrow> (t \<le> x \<longleftrightarrow> False)"
-  by (auto simp add: less_le not_less not_le)
-lemma minf_eq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
-lemma minf_neq: "\<exists>z. \<forall>x. x < z \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
-lemma minf_P: "\<exists>z. \<forall>x. x < z \<longrightarrow> (P \<longleftrightarrow> P)" by blast
-
-text{* Theorems for @{text "\<exists>z. \<forall>x. x < z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"}*}
-lemma pinf_gt:  "\<exists>z . \<forall>x. z < x \<longrightarrow> (t < x \<longleftrightarrow> True)" by auto
-lemma pinf_lt: "\<exists>z . \<forall>x. z < x \<longrightarrow>  (x < t \<longleftrightarrow>  False)"
-  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)
-
-lemma pinf_ge: "\<exists>z. \<forall>x. z < x \<longrightarrow> (t \<le> x \<longleftrightarrow> True)" by (auto simp add: less_le)
-lemma pinf_le: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<le> t \<longleftrightarrow> False)"
-  by (auto simp add: less_le not_less not_le)
-lemma pinf_eq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x = t \<longleftrightarrow> False)" by auto
-lemma pinf_neq: "\<exists>z. \<forall>x. z < x \<longrightarrow> (x \<noteq> t \<longleftrightarrow> True)" by auto
-lemma pinf_P: "\<exists>z. \<forall>x. z < x \<longrightarrow> (P \<longleftrightarrow> P)" by blast
-
-lemma nmi_lt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> x < t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
-lemma nmi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and> t < x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)"
-  by (auto simp add: le_less)
-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
-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
-lemma  nmi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x = t \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
-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
-lemma  nmi_P: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
-lemma  nmi_conj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x) ;
-  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
-  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
-lemma  nmi_disj: "\<lbrakk>\<forall>x. \<not>P1' \<and> P1 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x) ;
-  \<forall>x. \<not>P2' \<and> P2 x \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)\<rbrakk> \<Longrightarrow>
-  \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. u \<le> x)" by auto
-
-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)
-lemma  npi_gt: "t \<in> U \<Longrightarrow> \<forall>x. \<not>True \<and> t < x \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
-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
-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
-lemma  npi_eq: "t \<in> U \<Longrightarrow> \<forall>x. \<not>False \<and>  x = t \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
-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
-lemma  npi_P: "\<forall> x. ~P \<and> P \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
-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>
-  \<Longrightarrow>  \<forall>x. \<not>(P1' \<and> P2') \<and> (P1 x \<and> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
-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>
-  \<Longrightarrow> \<forall>x. \<not>(P1' \<or> P2') \<and> (P1 x \<or> P2 x) \<longrightarrow>  (\<exists> u\<in> U. x \<le> u)" by auto
-
-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)"
-proof(clarsimp)
-  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"
-    and xu: "x<u"  and px: "x < t" and ly: "l<y" and yu:"y < u"
-  from tU noU ly yu have tny: "t\<noteq>y" by auto
-  {assume H: "t < y"
-    from less_trans[OF lx px] less_trans[OF H yu]
-    have "l < t \<and> t < u"  by simp
-    with tU noU have "False" by auto}
-  hence "\<not> t < y"  by auto hence "y \<le> t" by (simp add: not_less)
-  thus "y < t" using tny by (simp add: less_le)
-qed
-
-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)"
-proof(clarsimp)
-  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" and xu: "x<u"
-  and px: "t < x" and ly: "l<y" and yu:"y < u"
-  from tU noU ly yu have tny: "t\<noteq>y" by auto
-  {assume H: "y< t"
-    from less_trans[OF ly H] less_trans[OF px xu] have "l < t \<and> t < u" by simp
-    with tU noU have "False" by auto}
-  hence "\<not> y<t"  by auto hence "t \<le> y" by (auto simp add: not_less)
-  thus "t < y" using tny by (simp add:less_le)
-qed
-
-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)"
-proof(clarsimp)
-  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" and xu: "x<u"
-  and px: "x \<le> t" and ly: "l<y" and yu:"y < u"
-  from tU noU ly yu have tny: "t\<noteq>y" by auto
-  {assume H: "t < y"
-    from less_le_trans[OF lx px] less_trans[OF H yu]
-    have "l < t \<and> t < u" by simp
-    with tU noU have "False" by auto}
-  hence "\<not> t < y"  by auto thus "y \<le> t" by (simp add: not_less)
-qed
-
-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)"
-proof(clarsimp)
-  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" and xu: "x<u"
-  and px: "t \<le> x" and ly: "l<y" and yu:"y < u"
-  from tU noU ly yu have tny: "t\<noteq>y" by auto
-  {assume H: "y< t"
-    from less_trans[OF ly H] le_less_trans[OF px xu]
-    have "l < t \<and> t < u" by simp
-    with tU noU have "False" by auto}
-  hence "\<not> y<t"  by auto thus "t \<le> y" by (simp add: not_less)
-qed
-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
-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
-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
-
-lemma lin_dense_conj:
-  "\<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
-  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
-  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
-  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
-  \<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)
-  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<and> P2 y))"
-  by blast
-lemma lin_dense_disj:
-  "\<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
-  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P1 y) ;
-  \<forall>x l u. (\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> U) \<and> l< x \<and> x < u \<and> P2 x
-  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> P2 y)\<rbrakk> \<Longrightarrow>
-  \<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)
-  \<longrightarrow> (\<forall> y. l < y \<and> y < u \<longrightarrow> (P1 y \<or> P2 y))"
-  by blast
-
-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>
-  \<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')"
-by auto
-
-lemma finite_set_intervals:
-  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
-  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"
-  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"
-proof-
-  let ?Mx = "{y. y\<in> S \<and> y \<le> x}"
-  let ?xM = "{y. y\<in> S \<and> x \<le> y}"
-  let ?a = "Max ?Mx"
-  let ?b = "Min ?xM"
-  have MxS: "?Mx \<subseteq> S" by blast
-  hence fMx: "finite ?Mx" using fS finite_subset by auto
-  from lx linS have linMx: "l \<in> ?Mx" by blast
-  hence Mxne: "?Mx \<noteq> {}" by blast
-  have xMS: "?xM \<subseteq> S" by blast
-  hence fxM: "finite ?xM" using fS finite_subset by auto
-  from xu uinS have linxM: "u \<in> ?xM" by blast
-  hence xMne: "?xM \<noteq> {}" by blast
-  have ax:"?a \<le> x" using Mxne fMx by auto
-  have xb:"x \<le> ?b" using xMne fxM by auto
-  have "?a \<in> ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \<in> S" using MxS by blast
-  have "?b \<in> ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \<in> S" using xMS by blast
-  have noy:"\<forall> y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S"
-  proof(clarsimp)
-    fix y   assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S"
-    from yS have "y\<in> ?Mx \<or> y\<in> ?xM" by (auto simp add: linear)
-    moreover {assume "y \<in> ?Mx" hence "y \<le> ?a" using Mxne fMx by auto with ay have "False" by (simp add: not_le[symmetric])}
-    moreover {assume "y \<in> ?xM" hence "?b \<le> y" using xMne fxM by auto with yb have "False" by (simp add: not_le[symmetric])}
-    ultimately show "False" by blast
-  qed
-  from ainS binS noy ax xb px show ?thesis by blast
-qed
-
-lemma finite_set_intervals2:
-  assumes px: "P x" and lx: "l \<le> x" and xu: "x \<le> u" and linS: "l\<in> S"
-  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"
-  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)"
-proof-
-  from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su]
-  obtain a and b where
-    as: "a\<in> S" and bs: "b\<in> S" and noS:"\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S"
-    and axb: "a \<le> x \<and> x \<le> b \<and> P x"  by auto
-  from axb have "x= a \<or> x= b \<or> (a < x \<and> x < b)" by (auto simp add: le_less)
-  thus ?thesis using px as bs noS by blast
-qed
-
-end
-
-section {* The classical QE after Langford for dense linear orders *}
-
-context dense_linear_order
-begin
-
-lemma interval_empty_iff:
-  "{y. x < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
-  by (auto dest: dense)
-
-lemma dlo_qe_bnds: 
-  assumes ne: "L \<noteq> {}" and neU: "U \<noteq> {}" and fL: "finite L" and fU: "finite U"
-  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)"
-proof (simp only: atomize_eq, rule iffI)
-  assume H: "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)"
-  then obtain x where xL: "\<forall>y\<in>L. y < x" and xU: "\<forall>y\<in>U. x < y" by blast
-  {fix l u assume l: "l \<in> L" and u: "u \<in> U"
-    have "l < x" using xL l by blast
-    also have "x < u" using xU u by blast
-    finally (less_trans) have "l < u" .}
-  thus "\<forall>l\<in>L. \<forall>u\<in>U. l < u" by blast
-next
-  assume H: "\<forall>l\<in>L. \<forall>u\<in>U. l < u"
-  let ?ML = "Max L"
-  let ?MU = "Min U"  
-  from fL ne have th1: "?ML \<in> L" and th1': "\<forall>l\<in>L. l \<le> ?ML" by auto
-  from fU neU have th2: "?MU \<in> U" and th2': "\<forall>u\<in>U. ?MU \<le> u" by auto
-  from th1 th2 H have "?ML < ?MU" by auto
-  with dense obtain w where th3: "?ML < w" and th4: "w < ?MU" by blast
-  from th3 th1' have "\<forall>l \<in> L. l < w" by auto
-  moreover from th4 th2' have "\<forall>u \<in> U. w < u" by auto
-  ultimately show "\<exists>x. (\<forall>y\<in>L. y < x) \<and> (\<forall>y\<in>U. x < y)" by auto
-qed
-
-lemma dlo_qe_noub: 
-  assumes ne: "L \<noteq> {}" and fL: "finite L"
-  shows "(\<exists>x. (\<forall>y \<in> L. y < x) \<and> (\<forall>y \<in> {}. x < y)) \<equiv> True"
-proof(simp add: atomize_eq)
-  from gt_ex[of "Max L"] obtain M where M: "Max L < M" by blast
-  from ne fL have "\<forall>x \<in> L. x \<le> Max L" by simp
-  with M have "\<forall>x\<in>L. x < M" by (auto intro: le_less_trans)
-  thus "\<exists>x. \<forall>y\<in>L. y < x" by blast
-qed
-
-lemma dlo_qe_nolb: 
-  assumes ne: "U \<noteq> {}" and fU: "finite U"
-  shows "(\<exists>x. (\<forall>y \<in> {}. y < x) \<and> (\<forall>y \<in> U. x < y)) \<equiv> True"
-proof(simp add: atomize_eq)
-  from lt_ex[of "Min U"] obtain M where M: "M < Min U" by blast
-  from ne fU have "\<forall>x \<in> U. Min U \<le> x" by simp
-  with M have "\<forall>x\<in>U. M < x" by (auto intro: less_le_trans)
-  thus "\<exists>x. \<forall>y\<in>U. x < y" by blast
-qed
-
-lemma exists_neq: "\<exists>(x::'a). x \<noteq> t" "\<exists>(x::'a). t \<noteq> x" 
-  using gt_ex[of t] by auto
-
-lemmas dlo_simps = order_refl less_irrefl not_less not_le exists_neq 
-  le_less neq_iff linear less_not_permute
-
-lemma axiom: "dense_linear_order (op \<le>) (op <)" by (rule dense_linear_order_axioms)
-lemma atoms:
-  shows "TERM (less :: 'a \<Rightarrow> _)"
-    and "TERM (less_eq :: 'a \<Rightarrow> _)"
-    and "TERM (op = :: 'a \<Rightarrow> _)" .
-
-declare axiom[langford qe: dlo_qe_bnds dlo_qe_nolb dlo_qe_noub gather: gather_start gather_simps atoms: atoms]
-declare dlo_simps[langfordsimp]
-
-end
-
-(* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *)
-lemma dnf:
-  "(P & (Q | R)) = ((P&Q) | (P&R))" 
-  "((Q | R) & P) = ((Q&P) | (R&P))"
-  by blast+
-
-lemmas weak_dnf_simps = simp_thms dnf
-
-lemma nnf_simps:
-    "(\<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)"
-    "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
-  by blast+
-
-lemma ex_distrib: "(\<exists>x. P x \<or> Q x) \<longleftrightarrow> ((\<exists>x. P x) \<or> (\<exists>x. Q x))" by blast
-
-lemmas dnf_simps = weak_dnf_simps nnf_simps ex_distrib
-
-use "~~/src/HOL/Tools/Qelim/langford.ML"
-method_setup dlo = {*
-  Method.ctxt_args (Method.SIMPLE_METHOD' o LangfordQE.dlo_tac)
-*} "Langford's algorithm for quantifier elimination in dense linear orders"
-
-
-section {* Contructive dense linear orders yield QE for linear arithmetic over ordered Fields -- see @{text "Arith_Tools.thy"} *}
-
-text {* Linear order without upper bounds *}
-
-locale linorder_stupid_syntax = linorder
-begin
-notation
-  less_eq  ("op \<sqsubseteq>") and
-  less_eq  ("(_/ \<sqsubseteq> _)" [51, 51] 50) and
-  less  ("op \<sqsubset>") and
-  less  ("(_/ \<sqsubset> _)"  [51, 51] 50)
-
-end
-
-locale linorder_no_ub = linorder_stupid_syntax +
-  assumes gt_ex: "\<exists>y. less x y"
-begin
-lemma ge_ex: "\<exists>y. x \<sqsubseteq> y" using gt_ex by auto
-
-text {* Theorems for @{text "\<exists>z. \<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>+\<infinity>\<^esub>)"} *}
-lemma pinf_conj:
-  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
-  shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
-proof-
-  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
-  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
-  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
-  {fix x assume H: "z \<sqsubset> x"
-    from less_trans[OF zz1 H] less_trans[OF zz2 H]
-    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
-  }
-  thus ?thesis by blast
-qed
-
-lemma pinf_disj:
-  assumes ex1: "\<exists>z1. \<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-  and ex2: "\<exists>z2. \<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
-  shows "\<exists>z. \<forall>x. z \<sqsubset>  x \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
-proof-
-  from ex1 ex2 obtain z1 and z2 where z1: "\<forall>x. z1 \<sqsubset> x \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-     and z2: "\<forall>x. z2 \<sqsubset> x \<longrightarrow> (P2 x \<longleftrightarrow> P2')" by blast
-  from gt_ex obtain z where z:"ord.max less_eq z1 z2 \<sqsubset> z" by blast
-  from z have zz1: "z1 \<sqsubset> z" and zz2: "z2 \<sqsubset> z" by simp_all
-  {fix x assume H: "z \<sqsubset> x"
-    from less_trans[OF zz1 H] less_trans[OF zz2 H]
-    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
-  }
-  thus ?thesis by blast
-qed
-
-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"
-proof-
-  from ex obtain z where z: "\<forall>x. z \<sqsubset> x \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
-  from gt_ex obtain x where x: "z \<sqsubset> x" by blast
-  from z x p1 show ?thesis by blast
-qed
-
-end
-
-text {* Linear order without upper bounds *}
-
-locale linorder_no_lb = linorder_stupid_syntax +
-  assumes lt_ex: "\<exists>y. less y x"
-begin
-lemma le_ex: "\<exists>y. y \<sqsubseteq> x" using lt_ex by auto
-
-
-text {* Theorems for @{text "\<exists>z. \<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P\<^bsub>-\<infinity>\<^esub>)"} *}
-lemma minf_conj:
-  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
-  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2'))"
-proof-
-  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
-  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
-  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
-  {fix x assume H: "x \<sqsubset> z"
-    from less_trans[OF H zz1] less_trans[OF H zz2]
-    have "(P1 x \<and> P2 x) \<longleftrightarrow> (P1' \<and> P2')"  using z1 zz1 z2 zz2 by auto
-  }
-  thus ?thesis by blast
-qed
-
-lemma minf_disj:
-  assumes ex1: "\<exists>z1. \<forall>x. x \<sqsubset> z1 \<longrightarrow> (P1 x \<longleftrightarrow> P1')"
-  and ex2: "\<exists>z2. \<forall>x. x \<sqsubset> z2 \<longrightarrow> (P2 x \<longleftrightarrow> P2')"
-  shows "\<exists>z. \<forall>x. x \<sqsubset>  z \<longrightarrow> ((P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2'))"
-proof-
-  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
-  from lt_ex obtain z where z:"z \<sqsubset> ord.min less_eq z1 z2" by blast
-  from z have zz1: "z \<sqsubset> z1" and zz2: "z \<sqsubset> z2" by simp_all
-  {fix x assume H: "x \<sqsubset> z"
-    from less_trans[OF H zz1] less_trans[OF H zz2]
-    have "(P1 x \<or> P2 x) \<longleftrightarrow> (P1' \<or> P2')"  using z1 zz1 z2 zz2 by auto
-  }
-  thus ?thesis by blast
-qed
-
-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"
-proof-
-  from ex obtain z where z: "\<forall>x. x \<sqsubset> z \<longrightarrow> (P x \<longleftrightarrow> P1)" by blast
-  from lt_ex obtain x where x: "x \<sqsubset> z" by blast
-  from z x p1 show ?thesis by blast
-qed
-
-end
-
-
-locale constr_dense_linear_order = linorder_no_lb + linorder_no_ub +
-  fixes between
-  assumes between_less: "less x y \<Longrightarrow> less x (between x y) \<and> less (between x y) y"
-     and  between_same: "between x x = x"
-
-interpretation  constr_dense_linear_order < dense_linear_order 
-  apply unfold_locales
-  using gt_ex lt_ex between_less
-    by (auto, rule_tac x="between x y" in exI, simp)
-
-context  constr_dense_linear_order
-begin
-
-lemma rinf_U:
-  assumes fU: "finite U"
-  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
-  \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
-  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')"
-  and nmi: "\<not> MP"  and npi: "\<not> PP"  and ex: "\<exists> x.  P x"
-  shows "\<exists> u\<in> U. \<exists> u' \<in> U. P (between u u')"
-proof-
-  from ex obtain x where px: "P x" by blast
-  from px nmi npi nmpiU have "\<exists> u\<in> U. \<exists> u' \<in> U. u \<sqsubseteq> x \<and> x \<sqsubseteq> u'" by auto
-  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
-  from uU have Une: "U \<noteq> {}" by auto
-  term "linorder.Min less_eq"
-  let ?l = "linorder.Min less_eq U"
-  let ?u = "linorder.Max less_eq U"
-  have linM: "?l \<in> U" using fU Une by simp
-  have uinM: "?u \<in> U" using fU Une by simp
-  have lM: "\<forall> t\<in> U. ?l \<sqsubseteq> t" using Une fU by auto
-  have Mu: "\<forall> t\<in> U. t \<sqsubseteq> ?u" using Une fU by auto
-  have th:"?l \<sqsubseteq> u" using uU Une lM by auto
-  from order_trans[OF th ux] have lx: "?l \<sqsubseteq> x" .
-  have th: "u' \<sqsubseteq> ?u" using uU' Une Mu by simp
-  from order_trans[OF xu' th] have xu: "x \<sqsubseteq> ?u" .
-  from finite_set_intervals2[where P="P",OF px lx xu linM uinM fU lM Mu]
-  have "(\<exists> s\<in> U. P s) \<or>
-      (\<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)" .
-  moreover { fix u assume um: "u\<in>U" and pu: "P u"
-    have "between u u = u" by (simp add: between_same)
-    with um pu have "P (between u u)" by simp
-    with um have ?thesis by blast}
-  moreover{
-    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"
-      then obtain t1 and t2 where t1M: "t1 \<in> U" and t2M: "t2\<in> U"
-        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"
-        by blast
-      from less_trans[OF t1x xt2] have t1t2: "t1 \<sqsubset> t2" .
-      let ?u = "between t1 t2"
-      from between_less t1t2 have t1lu: "t1 \<sqsubset> ?u" and ut2: "?u \<sqsubset> t2" by auto
-      from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast
-      with t1M t2M have ?thesis by blast}
-    ultimately show ?thesis by blast
-  qed
-
-theorem fr_eq:
-  assumes fU: "finite U"
-  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
-   \<longrightarrow> (\<forall> y. l \<sqsubset> y \<and> y \<sqsubset> u \<longrightarrow> P y )"
-  and nmibnd: "\<forall>x. \<not> MP \<and> P x \<longrightarrow> (\<exists> u\<in> U. u \<sqsubseteq> x)"
-  and npibnd: "\<forall>x. \<not>PP \<and> P x \<longrightarrow> (\<exists> u\<in> U. x \<sqsubseteq> u)"
-  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)"
-  shows "(\<exists> x. P x) \<equiv> (MP \<or> PP \<or> (\<exists> u \<in> U. \<exists> u'\<in> U. P (between u u')))"
-  (is "_ \<equiv> (_ \<or> _ \<or> ?F)" is "?E \<equiv> ?D")
-proof-
- {
-   assume px: "\<exists> x. P x"
-   have "MP \<or> PP \<or> (\<not> MP \<and> \<not> PP)" by blast
-   moreover {assume "MP \<or> PP" hence "?D" by blast}
-   moreover {assume nmi: "\<not> MP" and npi: "\<not> PP"
-     from npmibnd[OF nmibnd npibnd]
-     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')" .
-     from rinf_U[OF fU lin_dense nmpiU nmi npi px] have "?D" by blast}
-   ultimately have "?D" by blast}
- moreover
- { assume "?D"
-   moreover {assume m:"MP" from minf_ex[OF mi m] have "?E" .}
-   moreover {assume p: "PP" from pinf_ex[OF pi p] have "?E" . }
-   moreover {assume f:"?F" hence "?E" by blast}
-   ultimately have "?E" by blast}
- ultimately have "?E = ?D" by blast thus "?E \<equiv> ?D" by simp
-qed
-
-lemmas minf_thms = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P
-lemmas pinf_thms = pinf_conj pinf_disj pinf_eq pinf_neq pinf_lt pinf_le pinf_gt pinf_ge pinf_P
-
-lemmas nmi_thms = nmi_conj nmi_disj nmi_eq nmi_neq nmi_lt nmi_le nmi_gt nmi_ge nmi_P
-lemmas npi_thms = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P
-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
-
-lemma ferrack_axiom: "constr_dense_linear_order less_eq less between"
-  by (rule constr_dense_linear_order_axioms)
-lemma atoms:
-  shows "TERM (less :: 'a \<Rightarrow> _)"
-    and "TERM (less_eq :: 'a \<Rightarrow> _)"
-    and "TERM (op = :: 'a \<Rightarrow> _)" .
-
-declare ferrack_axiom [ferrack minf: minf_thms pinf: pinf_thms
-    nmi: nmi_thms npi: npi_thms lindense:
-    lin_dense_thms qe: fr_eq atoms: atoms]
-
-declaration {*
-let
-fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
-fun generic_whatis phi =
- let
-  val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
-  fun h x t =
-   case term_of t of
-     Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
-                            else Ferrante_Rackoff_Data.Nox
-   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
-                            else Ferrante_Rackoff_Data.Nox
-   | b$y$z => if Term.could_unify (b, lt) then
-                 if term_of x aconv y then Ferrante_Rackoff_Data.Lt
-                 else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
-                 else Ferrante_Rackoff_Data.Nox
-             else if Term.could_unify (b, le) then
-                 if term_of x aconv y then Ferrante_Rackoff_Data.Le
-                 else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
-                 else Ferrante_Rackoff_Data.Nox
-             else Ferrante_Rackoff_Data.Nox
-   | _ => Ferrante_Rackoff_Data.Nox
- in h end
- fun ss phi = HOL_ss addsimps (simps phi)
-in
- Ferrante_Rackoff_Data.funs  @{thm "ferrack_axiom"}
-  {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss}
-end
-*}
-
-end
-
-use "~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML"
-
-method_setup ferrack = {*
-  Method.ctxt_args (Method.SIMPLE_METHOD' o FerranteRackoff.dlo_tac)
-*} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders"
-
-subsection {* Ferrante and Rackoff algorithm over ordered fields *}
-
-lemma neg_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x < 0) == (x > 0))"
-proof-
-  assume H: "c < 0"
-  have "c*x < 0 = (0/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
-  also have "\<dots> = (0 < x)" by simp
-  finally show  "(c*x < 0) == (x > 0)" by simp
-qed
-
-lemma pos_prod_lt:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x < 0) == (x < 0))"
-proof-
-  assume H: "c > 0"
-  hence "c*x < 0 = (0/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
-  also have "\<dots> = (0 > x)" by simp
-  finally show  "(c*x < 0) == (x < 0)" by simp
-qed
-
-lemma neg_prod_sum_lt: "(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x + t< 0) == (x > (- 1/c)*t))"
-proof-
-  assume H: "c < 0"
-  have "c*x + t< 0 = (c*x < -t)" by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c < x)" by (simp only: neg_divide_less_eq[OF H] ring_simps)
-  also have "\<dots> = ((- 1/c)*t < x)" by simp
-  finally show  "(c*x + t < 0) == (x > (- 1/c)*t)" by simp
-qed
-
-lemma pos_prod_sum_lt:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x + t < 0) == (x < (- 1/c)*t))"
-proof-
-  assume H: "c > 0"
-  have "c*x + t< 0 = (c*x < -t)"  by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c > x)" by (simp only: pos_less_divide_eq[OF H] ring_simps)
-  also have "\<dots> = ((- 1/c)*t > x)" by simp
-  finally show  "(c*x + t < 0) == (x < (- 1/c)*t)" by simp
-qed
-
-lemma sum_lt:"((x::'a::pordered_ab_group_add) + t < 0) == (x < - t)"
-  using less_diff_eq[where a= x and b=t and c=0] by simp
-
-lemma neg_prod_le:"(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x <= 0) == (x >= 0))"
-proof-
-  assume H: "c < 0"
-  have "c*x <= 0 = (0/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
-  also have "\<dots> = (0 <= x)" by simp
-  finally show  "(c*x <= 0) == (x >= 0)" by simp
-qed
-
-lemma pos_prod_le:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x <= 0) == (x <= 0))"
-proof-
-  assume H: "c > 0"
-  hence "c*x <= 0 = (0/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
-  also have "\<dots> = (0 >= x)" by simp
-  finally show  "(c*x <= 0) == (x <= 0)" by simp
-qed
-
-lemma neg_prod_sum_le: "(c\<Colon>'a\<Colon>ordered_field) < 0 \<Longrightarrow> ((c*x + t <= 0) == (x >= (- 1/c)*t))"
-proof-
-  assume H: "c < 0"
-  have "c*x + t <= 0 = (c*x <= -t)"  by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c <= x)" by (simp only: neg_divide_le_eq[OF H] ring_simps)
-  also have "\<dots> = ((- 1/c)*t <= x)" by simp
-  finally show  "(c*x + t <= 0) == (x >= (- 1/c)*t)" by simp
-qed
-
-lemma pos_prod_sum_le:"(c\<Colon>'a\<Colon>ordered_field) > 0 \<Longrightarrow> ((c*x + t <= 0) == (x <= (- 1/c)*t))"
-proof-
-  assume H: "c > 0"
-  have "c*x + t <= 0 = (c*x <= -t)" by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (-t/c >= x)" by (simp only: pos_le_divide_eq[OF H] ring_simps)
-  also have "\<dots> = ((- 1/c)*t >= x)" by simp
-  finally show  "(c*x + t <= 0) == (x <= (- 1/c)*t)" by simp
-qed
-
-lemma sum_le:"((x::'a::pordered_ab_group_add) + t <= 0) == (x <= - t)"
-  using le_diff_eq[where a= x and b=t and c=0] by simp
-
-lemma nz_prod_eq:"(c\<Colon>'a\<Colon>ordered_field) \<noteq> 0 \<Longrightarrow> ((c*x = 0) == (x = 0))" by simp
-lemma nz_prod_sum_eq: "(c\<Colon>'a\<Colon>ordered_field) \<noteq> 0 \<Longrightarrow> ((c*x + t = 0) == (x = (- 1/c)*t))"
-proof-
-  assume H: "c \<noteq> 0"
-  have "c*x + t = 0 = (c*x = -t)" by (subst eq_iff_diff_eq_0 [of "c*x" "-t"], simp)
-  also have "\<dots> = (x = -t/c)" by (simp only: nonzero_eq_divide_eq[OF H] ring_simps)
-  finally show  "(c*x + t = 0) == (x = (- 1/c)*t)" by simp
-qed
-lemma sum_eq:"((x::'a::pordered_ab_group_add) + t = 0) == (x = - t)"
-  using eq_diff_eq[where a= x and b=t and c=0] by simp
-
-
-interpretation class_ordered_field_dense_linear_order: constr_dense_linear_order
- ["op <=" "op <"
-   "\<lambda> x y. 1/2 * ((x::'a::{ordered_field,recpower,number_ring}) + y)"]
-proof (unfold_locales, dlo, dlo, auto)
-  fix x y::'a assume lt: "x < y"
-  from  less_half_sum[OF lt] show "x < (x + y) /2" by simp
-next
-  fix x y::'a assume lt: "x < y"
-  from  gt_half_sum[OF lt] show "(x + y) /2 < y" by simp
-qed
-
-declaration{*
-let
-fun earlier [] x y = false
-        | earlier (h::t) x y =
-    if h aconvc y then false else if h aconvc x then true else earlier t x y;
-
-fun dest_frac ct = case term_of ct of
-   Const (@{const_name "HOL.divide"},_) $ a $ b=>
-    Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b))
- | t => Rat.rat_of_int (snd (HOLogic.dest_number t))
-
-fun mk_frac phi cT x =
- let val (a, b) = Rat.quotient_of_rat x
- in if b = 1 then Numeral.mk_cnumber cT a
-    else Thm.capply
-         (Thm.capply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"})
-                     (Numeral.mk_cnumber cT a))
-         (Numeral.mk_cnumber cT b)
- end
-
-fun whatis x ct = case term_of ct of
-  Const(@{const_name "HOL.plus"}, _)$(Const(@{const_name "HOL.times"},_)$_$y)$_ =>
-     if y aconv term_of x then ("c*x+t",[(funpow 2 Thm.dest_arg1) ct, Thm.dest_arg ct])
-     else ("Nox",[])
-| Const(@{const_name "HOL.plus"}, _)$y$_ =>
-     if y aconv term_of x then ("x+t",[Thm.dest_arg ct])
-     else ("Nox",[])
-| Const(@{const_name "HOL.times"}, _)$_$y =>
-     if y aconv term_of x then ("c*x",[Thm.dest_arg1 ct])
-     else ("Nox",[])
-| t => if t aconv term_of x then ("x",[]) else ("Nox",[]);
-
-fun xnormalize_conv ctxt [] ct = reflexive ct
-| xnormalize_conv ctxt (vs as (x::_)) ct =
-   case term_of ct of
-   Const(@{const_name HOL.less},_)$_$Const(@{const_name "HOL.zero"},_) =>
-    (case whatis x (Thm.dest_arg1 ct) of
-    ("c*x+t",[c,t]) =>
-       let
-        val cr = dest_frac c
-        val clt = Thm.dest_fun2 ct
-        val cz = Thm.dest_arg ct
-        val neg = cr </ Rat.zero
-        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
-               (Thm.capply @{cterm "Trueprop"}
-                  (if neg then Thm.capply (Thm.capply clt c) cz
-                    else Thm.capply (Thm.capply clt cz) c))
-        val cth = equal_elim (symmetric cthp) TrueI
-        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x,t])
-             (if neg then @{thm neg_prod_sum_lt} else @{thm pos_prod_sum_lt})) cth
-        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
-                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
-      in rth end
-    | ("x+t",[t]) =>
-       let
-        val T = ctyp_of_term x
-        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_lt"}
-        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
-              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
-       in  rth end
-    | ("c*x",[c]) =>
-       let
-        val cr = dest_frac c
-        val clt = Thm.dest_fun2 ct
-        val cz = Thm.dest_arg ct
-        val neg = cr </ Rat.zero
-        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
-               (Thm.capply @{cterm "Trueprop"}
-                  (if neg then Thm.capply (Thm.capply clt c) cz
-                    else Thm.capply (Thm.capply clt cz) c))
-        val cth = equal_elim (symmetric cthp) TrueI
-        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
-             (if neg then @{thm neg_prod_lt} else @{thm pos_prod_lt})) cth
-        val rth = th
-      in rth end
-    | _ => reflexive ct)
-
-
-|  Const(@{const_name HOL.less_eq},_)$_$Const(@{const_name "HOL.zero"},_) =>
-   (case whatis x (Thm.dest_arg1 ct) of
-    ("c*x+t",[c,t]) =>
-       let
-        val T = ctyp_of_term x
-        val cr = dest_frac c
-        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
-        val cz = Thm.dest_arg ct
-        val neg = cr </ Rat.zero
-        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
-               (Thm.capply @{cterm "Trueprop"}
-                  (if neg then Thm.capply (Thm.capply clt c) cz
-                    else Thm.capply (Thm.capply clt cz) c))
-        val cth = equal_elim (symmetric cthp) TrueI
-        val th = implies_elim (instantiate' [SOME T] (map SOME [c,x,t])
-             (if neg then @{thm neg_prod_sum_le} else @{thm pos_prod_sum_le})) cth
-        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
-                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
-      in rth end
-    | ("x+t",[t]) =>
-       let
-        val T = ctyp_of_term x
-        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_le"}
-        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
-              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
-       in  rth end
-    | ("c*x",[c]) =>
-       let
-        val T = ctyp_of_term x
-        val cr = dest_frac c
-        val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
-        val cz = Thm.dest_arg ct
-        val neg = cr </ Rat.zero
-        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
-               (Thm.capply @{cterm "Trueprop"}
-                  (if neg then Thm.capply (Thm.capply clt c) cz
-                    else Thm.capply (Thm.capply clt cz) c))
-        val cth = equal_elim (symmetric cthp) TrueI
-        val th = implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
-             (if neg then @{thm neg_prod_le} else @{thm pos_prod_le})) cth
-        val rth = th
-      in rth end
-    | _ => reflexive ct)
-
-|  Const("op =",_)$_$Const(@{const_name "HOL.zero"},_) =>
-   (case whatis x (Thm.dest_arg1 ct) of
-    ("c*x+t",[c,t]) =>
-       let
-        val T = ctyp_of_term x
-        val cr = dest_frac c
-        val ceq = Thm.dest_fun2 ct
-        val cz = Thm.dest_arg ct
-        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
-            (Thm.capply @{cterm "Trueprop"}
-             (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
-        val cth = equal_elim (symmetric cthp) TrueI
-        val th = implies_elim
-                 (instantiate' [SOME T] (map SOME [c,x,t]) @{thm nz_prod_sum_eq}) cth
-        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
-                   (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
-      in rth end
-    | ("x+t",[t]) =>
-       let
-        val T = ctyp_of_term x
-        val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_eq"}
-        val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
-              (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
-       in  rth end
-    | ("c*x",[c]) =>
-       let
-        val T = ctyp_of_term x
-        val cr = dest_frac c
-        val ceq = Thm.dest_fun2 ct
-        val cz = Thm.dest_arg ct
-        val cthp = Simplifier.rewrite (local_simpset_of ctxt)
-            (Thm.capply @{cterm "Trueprop"}
-             (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz)))
-        val cth = equal_elim (symmetric cthp) TrueI
-        val rth = implies_elim
-                 (instantiate' [SOME T] (map SOME [c,x]) @{thm nz_prod_eq}) cth
-      in rth end
-    | _ => reflexive ct);
-
-local
-  val less_iff_diff_less_0 = mk_meta_eq @{thm "less_iff_diff_less_0"}
-  val le_iff_diff_le_0 = mk_meta_eq @{thm "le_iff_diff_le_0"}
-  val eq_iff_diff_eq_0 = mk_meta_eq @{thm "eq_iff_diff_eq_0"}
-in
-fun field_isolate_conv phi ctxt vs ct = case term_of ct of
-  Const(@{const_name HOL.less},_)$a$b =>
-   let val (ca,cb) = Thm.dest_binop ct
-       val T = ctyp_of_term ca
-       val th = instantiate' [SOME T] [SOME ca, SOME cb] less_iff_diff_less_0
-       val nth = Conv.fconv_rule
-         (Conv.arg_conv (Conv.arg1_conv
-              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
-       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
-   in rth end
-| Const(@{const_name HOL.less_eq},_)$a$b =>
-   let val (ca,cb) = Thm.dest_binop ct
-       val T = ctyp_of_term ca
-       val th = instantiate' [SOME T] [SOME ca, SOME cb] le_iff_diff_le_0
-       val nth = Conv.fconv_rule
-         (Conv.arg_conv (Conv.arg1_conv
-              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
-       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
-   in rth end
-
-| Const("op =",_)$a$b =>
-   let val (ca,cb) = Thm.dest_binop ct
-       val T = ctyp_of_term ca
-       val th = instantiate' [SOME T] [SOME ca, SOME cb] eq_iff_diff_eq_0
-       val nth = Conv.fconv_rule
-         (Conv.arg_conv (Conv.arg1_conv
-              (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
-       val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
-   in rth end
-| @{term "Not"} $(Const("op =",_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
-| _ => reflexive ct
-end;
-
-fun classfield_whatis phi =
- let
-  fun h x t =
-   case term_of t of
-     Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
-                            else Ferrante_Rackoff_Data.Nox
-   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
-                            else Ferrante_Rackoff_Data.Nox
-   | Const(@{const_name HOL.less},_)$y$z =>
-       if term_of x aconv y then Ferrante_Rackoff_Data.Lt
-        else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
-        else Ferrante_Rackoff_Data.Nox
-   | Const (@{const_name HOL.less_eq},_)$y$z =>
-         if term_of x aconv y then Ferrante_Rackoff_Data.Le
-         else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
-         else Ferrante_Rackoff_Data.Nox
-   | _ => Ferrante_Rackoff_Data.Nox
- in h end;
-fun class_field_ss phi =
-   HOL_basic_ss addsimps ([@{thm "linorder_not_less"}, @{thm "linorder_not_le"}])
-   addsplits [@{thm "abs_split"},@{thm "split_max"}, @{thm "split_min"}]
-
-in
-Ferrante_Rackoff_Data.funs @{thm "class_ordered_field_dense_linear_order.ferrack_axiom"}
-  {isolate_conv = field_isolate_conv, whatis = classfield_whatis, simpset = class_field_ss}
-end
-*}
-
-
-end 
--- a/src/HOL/Library/Library.thy	Mon Dec 29 13:23:53 2008 +0100
+++ b/src/HOL/Library/Library.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -16,7 +16,6 @@
   Continuity
   ContNotDenum
   Countable
-  Dense_Linear_Order
   Efficient_Nat
   Enum
   Eval_Witness
--- a/src/HOL/Lim.thy	Mon Dec 29 13:23:53 2008 +0100
+++ b/src/HOL/Lim.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -7,7 +7,7 @@
 header{* Limits and Continuity *}
 
 theory Lim
-imports "~~/src/HOL/Hyperreal/SEQ"
+imports SEQ
 begin
 
 text{*Standard Definitions*}
--- a/src/HOL/PReal.thy	Mon Dec 29 13:23:53 2008 +0100
+++ b/src/HOL/PReal.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -9,7 +9,7 @@
 header {* Positive real numbers *}
 
 theory PReal
-imports Rational "~~/src/HOL/Library/Dense_Linear_Order"
+imports Rational Dense_Linear_Order
 begin
 
 text{*Could be generalized and moved to @{text Ring_and_Field}*}
--- a/src/HOL/Real.thy	Mon Dec 29 13:23:53 2008 +0100
+++ b/src/HOL/Real.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -1,5 +1,5 @@
 theory Real
-imports RComplete "~~/src/HOL/Real/RealVector"
+imports RComplete RealVector
 begin
 
 end
--- a/src/HOL/Real/HahnBanach/Bounds.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,83 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/Bounds.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Bounds *}
-
-theory Bounds
-imports Main ContNotDenum
-begin
-
-locale lub =
-  fixes A and x
-  assumes least [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> a \<le> b) \<Longrightarrow> x \<le> b"
-    and upper [intro?]: "a \<in> A \<Longrightarrow> a \<le> x"
-
-lemmas [elim?] = lub.least lub.upper
-
-definition
-  the_lub :: "'a::order set \<Rightarrow> 'a" where
-  "the_lub A = The (lub A)"
-
-notation (xsymbols)
-  the_lub  ("\<Squnion>_" [90] 90)
-
-lemma the_lub_equality [elim?]:
-  assumes "lub A x"
-  shows "\<Squnion>A = (x::'a::order)"
-proof -
-  interpret lub [A x] by fact
-  show ?thesis
-  proof (unfold the_lub_def)
-    from `lub A x` show "The (lub A) = x"
-    proof
-      fix x' assume lub': "lub A x'"
-      show "x' = x"
-      proof (rule order_antisym)
-	from lub' show "x' \<le> x"
-	proof
-          fix a assume "a \<in> A"
-          then show "a \<le> x" ..
-	qed
-	show "x \<le> x'"
-	proof
-          fix a assume "a \<in> A"
-          with lub' show "a \<le> x'" ..
-	qed
-      qed
-    qed
-  qed
-qed
-
-lemma the_lubI_ex:
-  assumes ex: "\<exists>x. lub A x"
-  shows "lub A (\<Squnion>A)"
-proof -
-  from ex obtain x where x: "lub A x" ..
-  also from x have [symmetric]: "\<Squnion>A = x" ..
-  finally show ?thesis .
-qed
-
-lemma lub_compat: "lub A x = isLub UNIV A x"
-proof -
-  have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)"
-    by (rule ext) (simp only: isUb_def)
-  then show ?thesis
-    by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast
-qed
-
-lemma real_complete:
-  fixes A :: "real set"
-  assumes nonempty: "\<exists>a. a \<in> A"
-    and ex_upper: "\<exists>y. \<forall>a \<in> A. a \<le> y"
-  shows "\<exists>x. lub A x"
-proof -
-  from ex_upper have "\<exists>y. isUb UNIV A y"
-    unfolding isUb_def setle_def by blast
-  with nonempty have "\<exists>x. isLub UNIV A x"
-    by (rule reals_complete)
-  then show ?thesis by (simp only: lub_compat)
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/FunctionNorm.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,279 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/FunctionNorm.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* The norm of a function *}
-
-theory FunctionNorm
-imports NormedSpace FunctionOrder
-begin
-
-subsection {* Continuous linear forms*}
-
-text {*
-  A linear form @{text f} on a normed vector space @{text "(V, \<parallel>\<cdot>\<parallel>)"}
-  is \emph{continuous}, iff it is bounded, i.e.
-  \begin{center}
-  @{text "\<exists>c \<in> R. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
-  \end{center}
-  In our application no other functions than linear forms are
-  considered, so we can define continuous linear forms as bounded
-  linear forms:
-*}
-
-locale continuous = var V + norm_syntax + linearform +
-  assumes bounded: "\<exists>c. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
-
-declare continuous.intro [intro?] continuous_axioms.intro [intro?]
-
-lemma continuousI [intro]:
-  fixes norm :: "_ \<Rightarrow> real"  ("\<parallel>_\<parallel>")
-  assumes "linearform V f"
-  assumes r: "\<And>x. x \<in> V \<Longrightarrow> \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
-  shows "continuous V norm f"
-proof
-  show "linearform V f" by fact
-  from r have "\<exists>c. \<forall>x\<in>V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" by blast
-  then show "continuous_axioms V norm f" ..
-qed
-
-
-subsection {* The norm of a linear form *}
-
-text {*
-  The least real number @{text c} for which holds
-  \begin{center}
-  @{text "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
-  \end{center}
-  is called the \emph{norm} of @{text f}.
-
-  For non-trivial vector spaces @{text "V \<noteq> {0}"} the norm can be
-  defined as
-  \begin{center}
-  @{text "\<parallel>f\<parallel> = \<sup>x \<noteq> 0. \<bar>f x\<bar> / \<parallel>x\<parallel>"}
-  \end{center}
-
-  For the case @{text "V = {0}"} the supremum would be taken from an
-  empty set. Since @{text \<real>} is unbounded, there would be no supremum.
-  To avoid this situation it must be guaranteed that there is an
-  element in this set. This element must be @{text "{} \<ge> 0"} so that
-  @{text fn_norm} has the norm properties. Furthermore it does not
-  have to change the norm in all other cases, so it must be @{text 0},
-  as all other elements are @{text "{} \<ge> 0"}.
-
-  Thus we define the set @{text B} where the supremum is taken from as
-  follows:
-  \begin{center}
-  @{text "{0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel>. x \<noteq> 0 \<and> x \<in> F}"}
-  \end{center}
-
-  @{text fn_norm} is equal to the supremum of @{text B}, if the
-  supremum exists (otherwise it is undefined).
-*}
-
-locale fn_norm = norm_syntax +
-  fixes B defines "B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
-  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
-  defines "\<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
-
-locale normed_vectorspace_with_fn_norm = normed_vectorspace + fn_norm
-
-lemma (in fn_norm) B_not_empty [intro]: "0 \<in> B V f"
-  by (simp add: B_def)
-
-text {*
-  The following lemma states that every continuous linear form on a
-  normed space @{text "(V, \<parallel>\<cdot>\<parallel>)"} has a function norm.
-*}
-
-lemma (in normed_vectorspace_with_fn_norm) fn_norm_works:
-  assumes "continuous V norm f"
-  shows "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
-proof -
-  interpret continuous [V norm f] by fact
-  txt {* The existence of the supremum is shown using the
-    completeness of the reals. Completeness means, that every
-    non-empty bounded set of reals has a supremum. *}
-  have "\<exists>a. lub (B V f) a"
-  proof (rule real_complete)
-    txt {* First we have to show that @{text B} is non-empty: *}
-    have "0 \<in> B V f" ..
-    then show "\<exists>x. x \<in> B V f" ..
-
-    txt {* Then we have to show that @{text B} is bounded: *}
-    show "\<exists>c. \<forall>y \<in> B V f. y \<le> c"
-    proof -
-      txt {* We know that @{text f} is bounded by some value @{text c}. *}
-      from bounded obtain c where c: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
-
-      txt {* To prove the thesis, we have to show that there is some
-        @{text b}, such that @{text "y \<le> b"} for all @{text "y \<in>
-        B"}. Due to the definition of @{text B} there are two cases. *}
-
-      def b \<equiv> "max c 0"
-      have "\<forall>y \<in> B V f. y \<le> b"
-      proof
-        fix y assume y: "y \<in> B V f"
-        show "y \<le> b"
-        proof cases
-          assume "y = 0"
-          then show ?thesis unfolding b_def by arith
-        next
-          txt {* The second case is @{text "y = \<bar>f x\<bar> / \<parallel>x\<parallel>"} for some
-            @{text "x \<in> V"} with @{text "x \<noteq> 0"}. *}
-          assume "y \<noteq> 0"
-          with y obtain x where y_rep: "y = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
-              and x: "x \<in> V" and neq: "x \<noteq> 0"
-            by (auto simp add: B_def real_divide_def)
-          from x neq have gt: "0 < \<parallel>x\<parallel>" ..
-
-          txt {* The thesis follows by a short calculation using the
-            fact that @{text f} is bounded. *}
-
-          note y_rep
-          also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
-          proof (rule mult_right_mono)
-            from c x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
-            from gt have "0 < inverse \<parallel>x\<parallel>" 
-              by (rule positive_imp_inverse_positive)
-            then show "0 \<le> inverse \<parallel>x\<parallel>" by (rule order_less_imp_le)
-          qed
-          also have "\<dots> = c * (\<parallel>x\<parallel> * inverse \<parallel>x\<parallel>)"
-            by (rule real_mult_assoc)
-          also
-          from gt have "\<parallel>x\<parallel> \<noteq> 0" by simp
-          then have "\<parallel>x\<parallel> * inverse \<parallel>x\<parallel> = 1" by simp 
-          also have "c * 1 \<le> b" by (simp add: b_def le_maxI1)
-          finally show "y \<le> b" .
-        qed
-      qed
-      then show ?thesis ..
-    qed
-  qed
-  then show ?thesis unfolding fn_norm_def by (rule the_lubI_ex)
-qed
-
-lemma (in normed_vectorspace_with_fn_norm) fn_norm_ub [iff?]:
-  assumes "continuous V norm f"
-  assumes b: "b \<in> B V f"
-  shows "b \<le> \<parallel>f\<parallel>\<hyphen>V"
-proof -
-  interpret continuous [V norm f] by fact
-  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
-    using `continuous V norm f` by (rule fn_norm_works)
-  from this and b show ?thesis ..
-qed
-
-lemma (in normed_vectorspace_with_fn_norm) fn_norm_leastB:
-  assumes "continuous V norm f"
-  assumes b: "\<And>b. b \<in> B V f \<Longrightarrow> b \<le> y"
-  shows "\<parallel>f\<parallel>\<hyphen>V \<le> y"
-proof -
-  interpret continuous [V norm f] by fact
-  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
-    using `continuous V norm f` by (rule fn_norm_works)
-  from this and b show ?thesis ..
-qed
-
-text {* The norm of a continuous function is always @{text "\<ge> 0"}. *}
-
-lemma (in normed_vectorspace_with_fn_norm) fn_norm_ge_zero [iff]:
-  assumes "continuous V norm f"
-  shows "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
-proof -
-  interpret continuous [V norm f] by fact
-  txt {* The function norm is defined as the supremum of @{text B}.
-    So it is @{text "\<ge> 0"} if all elements in @{text B} are @{text "\<ge>
-    0"}, provided the supremum exists and @{text B} is not empty. *}
-  have "lub (B V f) (\<parallel>f\<parallel>\<hyphen>V)"
-    using `continuous V norm f` by (rule fn_norm_works)
-  moreover have "0 \<in> B V f" ..
-  ultimately show ?thesis ..
-qed
-
-text {*
-  \medskip The fundamental property of function norms is:
-  \begin{center}
-  @{text "\<bar>f x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
-  \end{center}
-*}
-
-lemma (in normed_vectorspace_with_fn_norm) fn_norm_le_cong:
-  assumes "continuous V norm f" "linearform V f"
-  assumes x: "x \<in> V"
-  shows "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
-proof -
-  interpret continuous [V norm f] by fact
-  interpret linearform [V f] .
-  show ?thesis
-  proof cases
-    assume "x = 0"
-    then have "\<bar>f x\<bar> = \<bar>f 0\<bar>" by simp
-    also have "f 0 = 0" by rule unfold_locales
-    also have "\<bar>\<dots>\<bar> = 0" by simp
-    also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>V"
-      using `continuous V norm f` by (rule fn_norm_ge_zero)
-    from x have "0 \<le> norm x" ..
-    with a have "0 \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" by (simp add: zero_le_mult_iff)
-    finally show "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>" .
-  next
-    assume "x \<noteq> 0"
-    with x have neq: "\<parallel>x\<parallel> \<noteq> 0" by simp
-    then have "\<bar>f x\<bar> = (\<bar>f x\<bar> * inverse \<parallel>x\<parallel>) * \<parallel>x\<parallel>" by simp
-    also have "\<dots> \<le>  \<parallel>f\<parallel>\<hyphen>V * \<parallel>x\<parallel>"
-    proof (rule mult_right_mono)
-      from x show "0 \<le> \<parallel>x\<parallel>" ..
-      from x and neq have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<in> B V f"
-	by (auto simp add: B_def real_divide_def)
-      with `continuous V norm f` show "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>V"
-	by (rule fn_norm_ub)
-    qed
-    finally show ?thesis .
-  qed
-qed
-
-text {*
-  \medskip The function norm is the least positive real number for
-  which the following inequation holds:
-  \begin{center}
-    @{text "\<bar>f x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
-  \end{center}
-*}
-
-lemma (in normed_vectorspace_with_fn_norm) fn_norm_least [intro?]:
-  assumes "continuous V norm f"
-  assumes ineq: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" and ge: "0 \<le> c"
-  shows "\<parallel>f\<parallel>\<hyphen>V \<le> c"
-proof -
-  interpret continuous [V norm f] by fact
-  show ?thesis
-  proof (rule fn_norm_leastB [folded B_def fn_norm_def])
-    fix b assume b: "b \<in> B V f"
-    show "b \<le> c"
-    proof cases
-      assume "b = 0"
-      with ge show ?thesis by simp
-    next
-      assume "b \<noteq> 0"
-      with b obtain x where b_rep: "b = \<bar>f x\<bar> * inverse \<parallel>x\<parallel>"
-        and x_neq: "x \<noteq> 0" and x: "x \<in> V"
-	by (auto simp add: B_def real_divide_def)
-      note b_rep
-      also have "\<bar>f x\<bar> * inverse \<parallel>x\<parallel> \<le> (c * \<parallel>x\<parallel>) * inverse \<parallel>x\<parallel>"
-      proof (rule mult_right_mono)
-	have "0 < \<parallel>x\<parallel>" using x x_neq ..
-	then show "0 \<le> inverse \<parallel>x\<parallel>" by simp
-	from ineq and x show "\<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>" ..
-      qed
-      also have "\<dots> = c"
-      proof -
-	from x_neq and x have "\<parallel>x\<parallel> \<noteq> 0" by simp
-	then show ?thesis by simp
-      qed
-      finally show ?thesis .
-    qed
-  qed (insert `continuous V norm f`, simp_all add: continuous_def)
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/FunctionOrder.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,142 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/FunctionOrder.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* An order on functions *}
-
-theory FunctionOrder
-imports Subspace Linearform
-begin
-
-subsection {* The graph of a function *}
-
-text {*
-  We define the \emph{graph} of a (real) function @{text f} with
-  domain @{text F} as the set
-  \begin{center}
-  @{text "{(x, f x). x \<in> F}"}
-  \end{center}
-  So we are modeling partial functions by specifying the domain and
-  the mapping function. We use the term ``function'' also for its
-  graph.
-*}
-
-types 'a graph = "('a \<times> real) set"
-
-definition
-  graph :: "'a set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a graph" where
-  "graph F f = {(x, f x) | x. x \<in> F}"
-
-lemma graphI [intro]: "x \<in> F \<Longrightarrow> (x, f x) \<in> graph F f"
-  unfolding graph_def by blast
-
-lemma graphI2 [intro?]: "x \<in> F \<Longrightarrow> \<exists>t \<in> graph F f. t = (x, f x)"
-  unfolding graph_def by blast
-
-lemma graphE [elim?]:
-    "(x, y) \<in> graph F f \<Longrightarrow> (x \<in> F \<Longrightarrow> y = f x \<Longrightarrow> C) \<Longrightarrow> C"
-  unfolding graph_def by blast
-
-
-subsection {* Functions ordered by domain extension *}
-
-text {*
-  A function @{text h'} is an extension of @{text h}, iff the graph of
-  @{text h} is a subset of the graph of @{text h'}.
-*}
-
-lemma graph_extI:
-  "(\<And>x. x \<in> H \<Longrightarrow> h x = h' x) \<Longrightarrow> H \<subseteq> H'
-    \<Longrightarrow> graph H h \<subseteq> graph H' h'"
-  unfolding graph_def by blast
-
-lemma graph_extD1 [dest?]:
-  "graph H h \<subseteq> graph H' h' \<Longrightarrow> x \<in> H \<Longrightarrow> h x = h' x"
-  unfolding graph_def by blast
-
-lemma graph_extD2 [dest?]:
-  "graph H h \<subseteq> graph H' h' \<Longrightarrow> H \<subseteq> H'"
-  unfolding graph_def by blast
-
-
-subsection {* Domain and function of a graph *}
-
-text {*
-  The inverse functions to @{text graph} are @{text domain} and @{text
-  funct}.
-*}
-
-definition
-  "domain" :: "'a graph \<Rightarrow> 'a set" where
-  "domain g = {x. \<exists>y. (x, y) \<in> g}"
-
-definition
-  funct :: "'a graph \<Rightarrow> ('a \<Rightarrow> real)" where
-  "funct g = (\<lambda>x. (SOME y. (x, y) \<in> g))"
-
-text {*
-  The following lemma states that @{text g} is the graph of a function
-  if the relation induced by @{text g} is unique.
-*}
-
-lemma graph_domain_funct:
-  assumes uniq: "\<And>x y z. (x, y) \<in> g \<Longrightarrow> (x, z) \<in> g \<Longrightarrow> z = y"
-  shows "graph (domain g) (funct g) = g"
-  unfolding domain_def funct_def graph_def
-proof auto  (* FIXME !? *)
-  fix a b assume g: "(a, b) \<in> g"
-  from g show "(a, SOME y. (a, y) \<in> g) \<in> g" by (rule someI2)
-  from g show "\<exists>y. (a, y) \<in> g" ..
-  from g show "b = (SOME y. (a, y) \<in> g)"
-  proof (rule some_equality [symmetric])
-    fix y assume "(a, y) \<in> g"
-    with g show "y = b" by (rule uniq)
-  qed
-qed
-
-
-subsection {* Norm-preserving extensions of a function *}
-
-text {*
-  Given a linear form @{text f} on the space @{text F} and a seminorm
-  @{text p} on @{text E}. The set of all linear extensions of @{text
-  f}, to superspaces @{text H} of @{text F}, which are bounded by
-  @{text p}, is defined as follows.
-*}
-
-definition
-  norm_pres_extensions ::
-    "'a::{plus, minus, uminus, zero} set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> real)
-      \<Rightarrow> 'a graph set" where
-    "norm_pres_extensions E p F f
-      = {g. \<exists>H h. g = graph H h
-          \<and> linearform H h
-          \<and> H \<unlhd> E
-          \<and> F \<unlhd> H
-          \<and> graph F f \<subseteq> graph H h
-          \<and> (\<forall>x \<in> H. h x \<le> p x)}"
-
-lemma norm_pres_extensionE [elim]:
-  "g \<in> norm_pres_extensions E p F f
-  \<Longrightarrow> (\<And>H h. g = graph H h \<Longrightarrow> linearform H h
-        \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H \<Longrightarrow> graph F f \<subseteq> graph H h
-        \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x \<Longrightarrow> C) \<Longrightarrow> C"
-  unfolding norm_pres_extensions_def by blast
-
-lemma norm_pres_extensionI2 [intro]:
-  "linearform H h \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H
-    \<Longrightarrow> graph F f \<subseteq> graph H h \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x
-    \<Longrightarrow> graph H h \<in> norm_pres_extensions E p F f"
-  unfolding norm_pres_extensions_def by blast
-
-lemma norm_pres_extensionI:  (* FIXME ? *)
-  "\<exists>H h. g = graph H h
-    \<and> linearform H h
-    \<and> H \<unlhd> E
-    \<and> F \<unlhd> H
-    \<and> graph F f \<subseteq> graph H h
-    \<and> (\<forall>x \<in> H. h x \<le> p x) \<Longrightarrow> g \<in> norm_pres_extensions E p F f"
-  unfolding norm_pres_extensions_def by blast
-
-end
--- a/src/HOL/Real/HahnBanach/HahnBanach.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,510 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/HahnBanach.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* The Hahn-Banach Theorem *}
-
-theory HahnBanach
-imports HahnBanachLemmas
-begin
-
-text {*
-  We present the proof of two different versions of the Hahn-Banach
-  Theorem, closely following \cite[\S36]{Heuser:1986}.
-*}
-
-subsection {* The Hahn-Banach Theorem for vector spaces *}
-
-text {*
-  \textbf{Hahn-Banach Theorem.} Let @{text F} be a subspace of a real
-  vector space @{text E}, let @{text p} be a semi-norm on @{text E},
-  and @{text f} be a linear form defined on @{text F} such that @{text
-  f} is bounded by @{text p}, i.e.  @{text "\<forall>x \<in> F. f x \<le> p x"}.  Then
-  @{text f} can be extended to a linear form @{text h} on @{text E}
-  such that @{text h} is norm-preserving, i.e. @{text h} is also
-  bounded by @{text p}.
-
-  \bigskip
-  \textbf{Proof Sketch.}
-  \begin{enumerate}
-
-  \item Define @{text M} as the set of norm-preserving extensions of
-  @{text f} to subspaces of @{text E}. The linear forms in @{text M}
-  are ordered by domain extension.
-
-  \item We show that every non-empty chain in @{text M} has an upper
-  bound in @{text M}.
-
-  \item With Zorn's Lemma we conclude that there is a maximal function
-  @{text g} in @{text M}.
-
-  \item The domain @{text H} of @{text g} is the whole space @{text
-  E}, as shown by classical contradiction:
-
-  \begin{itemize}
-
-  \item Assuming @{text g} is not defined on whole @{text E}, it can
-  still be extended in a norm-preserving way to a super-space @{text
-  H'} of @{text H}.
-
-  \item Thus @{text g} can not be maximal. Contradiction!
-
-  \end{itemize}
-  \end{enumerate}
-*}
-
-theorem HahnBanach:
-  assumes E: "vectorspace E" and "subspace F E"
-    and "seminorm E p" and "linearform F f"
-  assumes fp: "\<forall>x \<in> F. f x \<le> p x"
-  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)"
-    -- {* Let @{text E} be a vector space, @{text F} a subspace of @{text E}, @{text p} a seminorm on @{text E}, *}
-    -- {* and @{text f} a linear form on @{text F} such that @{text f} is bounded by @{text p}, *}
-    -- {* then @{text f} can be extended to a linear form @{text h} on @{text E} in a norm-preserving way. \skp *}
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [F E] by fact
-  interpret seminorm [E p] by fact
-  interpret linearform [F f] by fact
-  def M \<equiv> "norm_pres_extensions E p F f"
-  then have M: "M = \<dots>" by (simp only:)
-  from E have F: "vectorspace F" ..
-  note FE = `F \<unlhd> E`
-  {
-    fix c assume cM: "c \<in> chain M" and ex: "\<exists>x. x \<in> c"
-    have "\<Union>c \<in> M"
-      -- {* Show that every non-empty chain @{text c} of @{text M} has an upper bound in @{text M}: *}
-      -- {* @{text "\<Union>c"} is greater than any element of the chain @{text c}, so it suffices to show @{text "\<Union>c \<in> M"}. *}
-      unfolding M_def
-    proof (rule norm_pres_extensionI)
-      let ?H = "domain (\<Union>c)"
-      let ?h = "funct (\<Union>c)"
-
-      have a: "graph ?H ?h = \<Union>c"
-      proof (rule graph_domain_funct)
-        fix x y z assume "(x, y) \<in> \<Union>c" and "(x, z) \<in> \<Union>c"
-        with M_def cM show "z = y" by (rule sup_definite)
-      qed
-      moreover from M cM a have "linearform ?H ?h"
-        by (rule sup_lf)
-      moreover from a M cM ex FE E have "?H \<unlhd> E"
-        by (rule sup_subE)
-      moreover from a M cM ex FE have "F \<unlhd> ?H"
-        by (rule sup_supF)
-      moreover from a M cM ex have "graph F f \<subseteq> graph ?H ?h"
-        by (rule sup_ext)
-      moreover from a M cM have "\<forall>x \<in> ?H. ?h x \<le> p x"
-        by (rule sup_norm_pres)
-      ultimately show "\<exists>H h. \<Union>c = graph H h
-          \<and> linearform H h
-          \<and> H \<unlhd> E
-          \<and> F \<unlhd> H
-          \<and> graph F f \<subseteq> graph H h
-          \<and> (\<forall>x \<in> H. h x \<le> p x)" by blast
-    qed
-  }
-  then have "\<exists>g \<in> M. \<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
-  -- {* With Zorn's Lemma we can conclude that there is a maximal element in @{text M}. \skp *}
-  proof (rule Zorn's_Lemma)
-      -- {* We show that @{text M} is non-empty: *}
-    show "graph F f \<in> M"
-      unfolding M_def
-    proof (rule norm_pres_extensionI2)
-      show "linearform F f" by fact
-      show "F \<unlhd> E" by fact
-      from F show "F \<unlhd> F" by (rule vectorspace.subspace_refl)
-      show "graph F f \<subseteq> graph F f" ..
-      show "\<forall>x\<in>F. f x \<le> p x" by fact
-    qed
-  qed
-  then obtain g where gM: "g \<in> M" and gx: "\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x"
-    by blast
-  from gM obtain H h where
-      g_rep: "g = graph H h"
-    and linearform: "linearform H h"
-    and HE: "H \<unlhd> E" and FH: "F \<unlhd> H"
-    and graphs: "graph F f \<subseteq> graph H h"
-    and hp: "\<forall>x \<in> H. h x \<le> p x" unfolding M_def ..
-      -- {* @{text g} is a norm-preserving extension of @{text f}, in other words: *}
-      -- {* @{text g} is the graph of some linear form @{text h} defined on a subspace @{text H} of @{text E}, *}
-      -- {* and @{text h} is an extension of @{text f} that is again bounded by @{text p}. \skp *}
-  from HE E have H: "vectorspace H"
-    by (rule subspace.vectorspace)
-
-  have HE_eq: "H = E"
-    -- {* We show that @{text h} is defined on whole @{text E} by classical contradiction. \skp *}
-  proof (rule classical)
-    assume neq: "H \<noteq> E"
-      -- {* Assume @{text h} is not defined on whole @{text E}. Then show that @{text h} can be extended *}
-      -- {* in a norm-preserving way to a function @{text h'} with the graph @{text g'}. \skp *}
-    have "\<exists>g' \<in> M. g \<subseteq> g' \<and> g \<noteq> g'"
-    proof -
-      from HE have "H \<subseteq> E" ..
-      with neq obtain x' where x'E: "x' \<in> E" and "x' \<notin> H" by blast
-      obtain x': "x' \<noteq> 0"
-      proof
-        show "x' \<noteq> 0"
-        proof
-          assume "x' = 0"
-          with H have "x' \<in> H" by (simp only: vectorspace.zero)
-          with `x' \<notin> H` show False by contradiction
-        qed
-      qed
-
-      def H' \<equiv> "H + lin x'"
-        -- {* Define @{text H'} as the direct sum of @{text H} and the linear closure of @{text x'}. \skp *}
-      have HH': "H \<unlhd> H'"
-      proof (unfold H'_def)
-        from x'E have "vectorspace (lin x')" ..
-        with H show "H \<unlhd> H + lin x'" ..
-      qed
-
-      obtain xi where
-        xi: "\<forall>y \<in> H. - p (y + x') - h y \<le> xi
-          \<and> xi \<le> p (y + x') - h y"
-        -- {* Pick a real number @{text \<xi>} that fulfills certain inequations; this will *}
-        -- {* be used to establish that @{text h'} is a norm-preserving extension of @{text h}.
-           \label{ex-xi-use}\skp *}
-      proof -
-        from H have "\<exists>xi. \<forall>y \<in> H. - p (y + x') - h y \<le> xi
-            \<and> xi \<le> p (y + x') - h y"
-        proof (rule ex_xi)
-          fix u v assume u: "u \<in> H" and v: "v \<in> H"
-          with HE have uE: "u \<in> E" and vE: "v \<in> E" by auto
-          from H u v linearform have "h v - h u = h (v - u)"
-            by (simp add: linearform.diff)
-          also from hp and H u v have "\<dots> \<le> p (v - u)"
-            by (simp only: vectorspace.diff_closed)
-          also from x'E uE vE have "v - u = x' + - x' + v + - u"
-            by (simp add: diff_eq1)
-          also from x'E uE vE have "\<dots> = v + x' + - (u + x')"
-            by (simp add: add_ac)
-          also from x'E uE vE have "\<dots> = (v + x') - (u + x')"
-            by (simp add: diff_eq1)
-          also from x'E uE vE E have "p \<dots> \<le> p (v + x') + p (u + x')"
-            by (simp add: diff_subadditive)
-          finally have "h v - h u \<le> p (v + x') + p (u + x')" .
-          then show "- p (u + x') - h u \<le> p (v + x') - h v" by simp
-        qed
-        then show thesis by (blast intro: that)
-      qed
-
-      def h' \<equiv> "\<lambda>x. let (y, a) =
-          SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H in h y + a * xi"
-        -- {* Define the extension @{text h'} of @{text h} to @{text H'} using @{text \<xi>}. \skp *}
-
-      have "g \<subseteq> graph H' h' \<and> g \<noteq> graph H' h'"
-        -- {* @{text h'} is an extension of @{text h} \dots \skp *}
-      proof
-        show "g \<subseteq> graph H' h'"
-        proof -
-          have  "graph H h \<subseteq> graph H' h'"
-          proof (rule graph_extI)
-            fix t assume t: "t \<in> H"
-            from E HE t have "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
-	      using `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` by (rule decomp_H'_H)
-            with h'_def show "h t = h' t" by (simp add: Let_def)
-          next
-            from HH' show "H \<subseteq> H'" ..
-          qed
-          with g_rep show ?thesis by (simp only:)
-        qed
-
-        show "g \<noteq> graph H' h'"
-        proof -
-          have "graph H h \<noteq> graph H' h'"
-          proof
-            assume eq: "graph H h = graph H' h'"
-            have "x' \<in> H'"
-	      unfolding H'_def
-            proof
-              from H show "0 \<in> H" by (rule vectorspace.zero)
-              from x'E show "x' \<in> lin x'" by (rule x_lin_x)
-              from x'E show "x' = 0 + x'" by simp
-            qed
-            then have "(x', h' x') \<in> graph H' h'" ..
-            with eq have "(x', h' x') \<in> graph H h" by (simp only:)
-            then have "x' \<in> H" ..
-            with `x' \<notin> H` show False by contradiction
-          qed
-          with g_rep show ?thesis by simp
-        qed
-      qed
-      moreover have "graph H' h' \<in> M"
-        -- {* and @{text h'} is norm-preserving. \skp *}
-      proof (unfold M_def)
-        show "graph H' h' \<in> norm_pres_extensions E p F f"
-        proof (rule norm_pres_extensionI2)
-          show "linearform H' h'"
-	    using h'_def H'_def HE linearform `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E
-	    by (rule h'_lf)
-          show "H' \<unlhd> E"
-	  unfolding H'_def
-          proof
-            show "H \<unlhd> E" by fact
-            show "vectorspace E" by fact
-            from x'E show "lin x' \<unlhd> E" ..
-          qed
-          from H `F \<unlhd> H` HH' show FH': "F \<unlhd> H'"
-            by (rule vectorspace.subspace_trans)
-          show "graph F f \<subseteq> graph H' h'"
-          proof (rule graph_extI)
-            fix x assume x: "x \<in> F"
-            with graphs have "f x = h x" ..
-            also have "\<dots> = h x + 0 * xi" by simp
-            also have "\<dots> = (let (y, a) = (x, 0) in h y + a * xi)"
-              by (simp add: Let_def)
-            also have "(x, 0) =
-                (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)"
-	      using E HE
-            proof (rule decomp_H'_H [symmetric])
-              from FH x show "x \<in> H" ..
-              from x' show "x' \<noteq> 0" .
-	      show "x' \<notin> H" by fact
-	      show "x' \<in> E" by fact
-            qed
-            also have
-              "(let (y, a) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)
-              in h y + a * xi) = h' x" by (simp only: h'_def)
-            finally show "f x = h' x" .
-          next
-            from FH' show "F \<subseteq> H'" ..
-          qed
-          show "\<forall>x \<in> H'. h' x \<le> p x"
-	    using h'_def H'_def `x' \<notin> H` `x' \<in> E` `x' \<noteq> 0` E HE
-	      `seminorm E p` linearform and hp xi
-	    by (rule h'_norm_pres)
-        qed
-      qed
-      ultimately show ?thesis ..
-    qed
-    then have "\<not> (\<forall>x \<in> M. g \<subseteq> x \<longrightarrow> g = x)" by simp
-      -- {* So the graph @{text g} of @{text h} cannot be maximal. Contradiction! \skp *}
-    with gx show "H = E" by contradiction
-  qed
-
-  from HE_eq and linearform have "linearform E h"
-    by (simp only:)
-  moreover have "\<forall>x \<in> F. h x = f x"
-  proof
-    fix x assume "x \<in> F"
-    with graphs have "f x = h x" ..
-    then show "h x = f x" ..
-  qed
-  moreover from HE_eq and hp have "\<forall>x \<in> E. h x \<le> p x"
-    by (simp only:)
-  ultimately show ?thesis by blast
-qed
-
-
-subsection  {* Alternative formulation *}
-
-text {*
-  The following alternative formulation of the Hahn-Banach
-  Theorem\label{abs-HahnBanach} uses the fact that for a real linear
-  form @{text f} and a seminorm @{text p} the following inequations
-  are equivalent:\footnote{This was shown in lemma @{thm [source]
-  abs_ineq_iff} (see page \pageref{abs-ineq-iff}).}
-  \begin{center}
-  \begin{tabular}{lll}
-  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
-  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
-  \end{tabular}
-  \end{center}
-*}
-
-theorem abs_HahnBanach:
-  assumes E: "vectorspace E" and FE: "subspace F E"
-    and lf: "linearform F f" and sn: "seminorm E p"
-  assumes fp: "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
-  shows "\<exists>g. linearform E g
-    \<and> (\<forall>x \<in> F. g x = f x)
-    \<and> (\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x)"
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [F E] by fact
-  interpret linearform [F f] by fact
-  interpret seminorm [E p] by fact
-  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)"
-    using E FE sn lf
-  proof (rule HahnBanach)
-    show "\<forall>x \<in> F. f x \<le> p x"
-      using FE E sn lf and fp by (rule abs_ineq_iff [THEN iffD1])
-  qed
-  then obtain g where lg: "linearform E g" and *: "\<forall>x \<in> F. g x = f x"
-      and **: "\<forall>x \<in> E. g x \<le> p x" by blast
-  have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
-    using _ E sn lg **
-  proof (rule abs_ineq_iff [THEN iffD2])
-    show "E \<unlhd> E" ..
-  qed
-  with lg * show ?thesis by blast
-qed
-
-
-subsection {* The Hahn-Banach Theorem for normed spaces *}
-
-text {*
-  Every continuous linear form @{text f} on a subspace @{text F} of a
-  norm space @{text E}, can be extended to a continuous linear form
-  @{text g} on @{text E} such that @{text "\<parallel>f\<parallel> = \<parallel>g\<parallel>"}.
-*}
-
-theorem norm_HahnBanach:
-  fixes V and norm ("\<parallel>_\<parallel>")
-  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}"
-  fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
-  defines "\<And>V f. \<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
-  assumes E_norm: "normed_vectorspace E norm" and FE: "subspace F E"
-    and linearform: "linearform F f" and "continuous F norm f"
-  shows "\<exists>g. linearform E g
-     \<and> continuous E norm g
-     \<and> (\<forall>x \<in> F. g x = f x)
-     \<and> \<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
-proof -
-  interpret normed_vectorspace [E norm] by fact
-  interpret normed_vectorspace_with_fn_norm [E norm B fn_norm]
-    by (auto simp: B_def fn_norm_def) intro_locales
-  interpret subspace [F E] by fact
-  interpret linearform [F f] by fact
-  interpret continuous [F norm f] by fact
-  have E: "vectorspace E" by intro_locales
-  have F: "vectorspace F" by rule intro_locales
-  have F_norm: "normed_vectorspace F norm"
-    using FE E_norm by (rule subspace_normed_vs)
-  have ge_zero: "0 \<le> \<parallel>f\<parallel>\<hyphen>F"
-    by (rule normed_vectorspace_with_fn_norm.fn_norm_ge_zero
-      [OF normed_vectorspace_with_fn_norm.intro,
-       OF F_norm `continuous F norm f` , folded B_def fn_norm_def])
-  txt {* We define a function @{text p} on @{text E} as follows:
-    @{text "p x = \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"} *}
-  def p \<equiv> "\<lambda>x. \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
-
-  txt {* @{text p} is a seminorm on @{text E}: *}
-  have q: "seminorm E p"
-  proof
-    fix x y a assume x: "x \<in> E" and y: "y \<in> E"
-    
-    txt {* @{text p} is positive definite: *}
-    have "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
-    moreover from x have "0 \<le> \<parallel>x\<parallel>" ..
-    ultimately show "0 \<le> p x"  
-      by (simp add: p_def zero_le_mult_iff)
-
-    txt {* @{text p} is absolutely homogenous: *}
-
-    show "p (a \<cdot> x) = \<bar>a\<bar> * p x"
-    proof -
-      have "p (a \<cdot> x) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>a \<cdot> x\<parallel>" by (simp only: p_def)
-      also from x have "\<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>" by (rule abs_homogenous)
-      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
-      also have "\<dots> = \<bar>a\<bar> * p x" by (simp only: p_def)
-      finally show ?thesis .
-    qed
-
-    txt {* Furthermore, @{text p} is subadditive: *}
-
-    show "p (x + y) \<le> p x + p y"
-    proof -
-      have "p (x + y) = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel>" by (simp only: p_def)
-      also have a: "0 \<le> \<parallel>f\<parallel>\<hyphen>F" by (rule ge_zero)
-      from x y have "\<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>" ..
-      with a have " \<parallel>f\<parallel>\<hyphen>F * \<parallel>x + y\<parallel> \<le> \<parallel>f\<parallel>\<hyphen>F * (\<parallel>x\<parallel> + \<parallel>y\<parallel>)"
-        by (simp add: mult_left_mono)
-      also have "\<dots> = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel> + \<parallel>f\<parallel>\<hyphen>F * \<parallel>y\<parallel>" by (simp only: right_distrib)
-      also have "\<dots> = p x + p y" by (simp only: p_def)
-      finally show ?thesis .
-    qed
-  qed
-
-  txt {* @{text f} is bounded by @{text p}. *}
-
-  have "\<forall>x \<in> F. \<bar>f x\<bar> \<le> p x"
-  proof
-    fix x assume "x \<in> F"
-    with `continuous F norm f` and linearform
-    show "\<bar>f x\<bar> \<le> p x"
-      unfolding p_def by (rule normed_vectorspace_with_fn_norm.fn_norm_le_cong
-        [OF normed_vectorspace_with_fn_norm.intro,
-         OF F_norm, folded B_def fn_norm_def])
-  qed
-
-  txt {* Using the fact that @{text p} is a seminorm and @{text f} is bounded
-    by @{text p} we can apply the Hahn-Banach Theorem for real vector
-    spaces. So @{text f} can be extended in a norm-preserving way to
-    some function @{text g} on the whole vector space @{text E}. *}
-
-  with E FE linearform q obtain g where
-      linearformE: "linearform E g"
-    and a: "\<forall>x \<in> F. g x = f x"
-    and b: "\<forall>x \<in> E. \<bar>g x\<bar> \<le> p x"
-    by (rule abs_HahnBanach [elim_format]) iprover
-
-  txt {* We furthermore have to show that @{text g} is also continuous: *}
-
-  have g_cont: "continuous E norm g" using linearformE
-  proof
-    fix x assume "x \<in> E"
-    with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
-      by (simp only: p_def)
-  qed
-
-  txt {* To complete the proof, we show that @{text "\<parallel>g\<parallel> = \<parallel>f\<parallel>"}. *}
-
-  have "\<parallel>g\<parallel>\<hyphen>E = \<parallel>f\<parallel>\<hyphen>F"
-  proof (rule order_antisym)
-    txt {*
-      First we show @{text "\<parallel>g\<parallel> \<le> \<parallel>f\<parallel>"}.  The function norm @{text
-      "\<parallel>g\<parallel>"} is defined as the smallest @{text "c \<in> \<real>"} such that
-      \begin{center}
-      \begin{tabular}{l}
-      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> c \<cdot> \<parallel>x\<parallel>"}
-      \end{tabular}
-      \end{center}
-      \noindent Furthermore holds
-      \begin{center}
-      \begin{tabular}{l}
-      @{text "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>"}
-      \end{tabular}
-      \end{center}
-    *}
-
-    have "\<forall>x \<in> E. \<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
-    proof
-      fix x assume "x \<in> E"
-      with b show "\<bar>g x\<bar> \<le> \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
-        by (simp only: p_def)
-    qed
-    from g_cont this ge_zero
-    show "\<parallel>g\<parallel>\<hyphen>E \<le> \<parallel>f\<parallel>\<hyphen>F"
-      by (rule fn_norm_least [of g, folded B_def fn_norm_def])
-
-    txt {* The other direction is achieved by a similar argument. *}
-
-    show "\<parallel>f\<parallel>\<hyphen>F \<le> \<parallel>g\<parallel>\<hyphen>E"
-    proof (rule normed_vectorspace_with_fn_norm.fn_norm_least
-	[OF normed_vectorspace_with_fn_norm.intro,
-	 OF F_norm, folded B_def fn_norm_def])
-      show "\<forall>x \<in> F. \<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
-      proof
-	fix x assume x: "x \<in> F"
-	from a x have "g x = f x" ..
-	then have "\<bar>f x\<bar> = \<bar>g x\<bar>" by (simp only:)
-	also from g_cont
-	have "\<dots> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>"
-	proof (rule fn_norm_le_cong [of g, folded B_def fn_norm_def])
-	  from FE x show "x \<in> E" ..
-	qed
-	finally show "\<bar>f x\<bar> \<le> \<parallel>g\<parallel>\<hyphen>E * \<parallel>x\<parallel>" .
-      qed
-      show "0 \<le> \<parallel>g\<parallel>\<hyphen>E"
-	using g_cont
-	by (rule fn_norm_ge_zero [of g, folded B_def fn_norm_def])
-      show "continuous F norm f" by fact
-    qed
-  qed
-  with linearformE a g_cont show ?thesis by blast
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/HahnBanachExtLemmas.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,281 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/HahnBanachExtLemmas.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Extending non-maximal functions *}
-
-theory HahnBanachExtLemmas
-imports FunctionNorm
-begin
-
-text {*
-  In this section the following context is presumed.  Let @{text E} be
-  a real vector space with a seminorm @{text q} on @{text E}. @{text
-  F} is a subspace of @{text E} and @{text f} a linear function on
-  @{text F}. We consider a subspace @{text H} of @{text E} that is a
-  superspace of @{text F} and a linear form @{text h} on @{text
-  H}. @{text H} is a not equal to @{text E} and @{text "x\<^sub>0"} is
-  an element in @{text "E - H"}.  @{text H} is extended to the direct
-  sum @{text "H' = H + lin x\<^sub>0"}, so for any @{text "x \<in> H'"}
-  the decomposition of @{text "x = y + a \<cdot> x"} with @{text "y \<in> H"} is
-  unique. @{text h'} is defined on @{text H'} by @{text "h' x = h y +
-  a \<cdot> \<xi>"} for a certain @{text \<xi>}.
-
-  Subsequently we show some properties of this extension @{text h'} of
-  @{text h}.
-
-  \medskip This lemma will be used to show the existence of a linear
-  extension of @{text f} (see page \pageref{ex-xi-use}). It is a
-  consequence of the completeness of @{text \<real>}. To show
-  \begin{center}
-  \begin{tabular}{l}
-  @{text "\<exists>\<xi>. \<forall>y \<in> F. a y \<le> \<xi> \<and> \<xi> \<le> b y"}
-  \end{tabular}
-  \end{center}
-  \noindent it suffices to show that
-  \begin{center}
-  \begin{tabular}{l}
-  @{text "\<forall>u \<in> F. \<forall>v \<in> F. a u \<le> b v"}
-  \end{tabular}
-  \end{center}
-*}
-
-lemma ex_xi:
-  assumes "vectorspace F"
-  assumes r: "\<And>u v. u \<in> F \<Longrightarrow> v \<in> F \<Longrightarrow> a u \<le> b v"
-  shows "\<exists>xi::real. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y"
-proof -
-  interpret vectorspace [F] by fact
-  txt {* From the completeness of the reals follows:
-    The set @{text "S = {a u. u \<in> F}"} has a supremum, if it is
-    non-empty and has an upper bound. *}
-
-  let ?S = "{a u | u. u \<in> F}"
-  have "\<exists>xi. lub ?S xi"
-  proof (rule real_complete)
-    have "a 0 \<in> ?S" by blast
-    then show "\<exists>X. X \<in> ?S" ..
-    have "\<forall>y \<in> ?S. y \<le> b 0"
-    proof
-      fix y assume y: "y \<in> ?S"
-      then obtain u where u: "u \<in> F" and y: "y = a u" by blast
-      from u and zero have "a u \<le> b 0" by (rule r)
-      with y show "y \<le> b 0" by (simp only:)
-    qed
-    then show "\<exists>u. \<forall>y \<in> ?S. y \<le> u" ..
-  qed
-  then obtain xi where xi: "lub ?S xi" ..
-  {
-    fix y assume "y \<in> F"
-    then have "a y \<in> ?S" by blast
-    with xi have "a y \<le> xi" by (rule lub.upper)
-  } moreover {
-    fix y assume y: "y \<in> F"
-    from xi have "xi \<le> b y"
-    proof (rule lub.least)
-      fix au assume "au \<in> ?S"
-      then obtain u where u: "u \<in> F" and au: "au = a u" by blast
-      from u y have "a u \<le> b y" by (rule r)
-      with au show "au \<le> b y" by (simp only:)
-    qed
-  } ultimately show "\<exists>xi. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y" by blast
-qed
-
-text {*
-  \medskip The function @{text h'} is defined as a @{text "h' x = h y
-  + a \<cdot> \<xi>"} where @{text "x = y + a \<cdot> \<xi>"} is a linear extension of
-  @{text h} to @{text H'}.
-*}
-
-lemma h'_lf:
-  assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
-      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
-    and H'_def: "H' \<equiv> H + lin x0"
-    and HE: "H \<unlhd> E"
-  assumes "linearform H h"
-  assumes x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
-  assumes E: "vectorspace E"
-  shows "linearform H' h'"
-proof -
-  interpret linearform [H h] by fact
-  interpret vectorspace [E] by fact
-  show ?thesis
-  proof
-    note E = `vectorspace E`
-    have H': "vectorspace H'"
-    proof (unfold H'_def)
-      from `x0 \<in> E`
-      have "lin x0 \<unlhd> E" ..
-      with HE show "vectorspace (H + lin x0)" using E ..
-    qed
-    {
-      fix x1 x2 assume x1: "x1 \<in> H'" and x2: "x2 \<in> H'"
-      show "h' (x1 + x2) = h' x1 + h' x2"
-      proof -
-	from H' x1 x2 have "x1 + x2 \<in> H'"
-          by (rule vectorspace.add_closed)
-	with x1 x2 obtain y y1 y2 a a1 a2 where
-          x1x2: "x1 + x2 = y + a \<cdot> x0" and y: "y \<in> H"
-          and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
-          and x2_rep: "x2 = y2 + a2 \<cdot> x0" and y2: "y2 \<in> H"
-          unfolding H'_def sum_def lin_def by blast
-	
-	have ya: "y1 + y2 = y \<and> a1 + a2 = a" using E HE _ y x0
-	proof (rule decomp_H') txt_raw {* \label{decomp-H-use} *}
-          from HE y1 y2 show "y1 + y2 \<in> H"
-            by (rule subspace.add_closed)
-          from x0 and HE y y1 y2
-          have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E"  "y2 \<in> E" by auto
-          with x1_rep x2_rep have "(y1 + y2) + (a1 + a2) \<cdot> x0 = x1 + x2"
-            by (simp add: add_ac add_mult_distrib2)
-          also note x1x2
-          finally show "(y1 + y2) + (a1 + a2) \<cdot> x0 = y + a \<cdot> x0" .
-	qed
-	
-	from h'_def x1x2 E HE y x0
-	have "h' (x1 + x2) = h y + a * xi"
-          by (rule h'_definite)
-	also have "\<dots> = h (y1 + y2) + (a1 + a2) * xi"
-          by (simp only: ya)
-	also from y1 y2 have "h (y1 + y2) = h y1 + h y2"
-          by simp
-	also have "\<dots> + (a1 + a2) * xi = (h y1 + a1 * xi) + (h y2 + a2 * xi)"
-          by (simp add: left_distrib)
-	also from h'_def x1_rep E HE y1 x0
-	have "h y1 + a1 * xi = h' x1"
-          by (rule h'_definite [symmetric])
-	also from h'_def x2_rep E HE y2 x0
-	have "h y2 + a2 * xi = h' x2"
-          by (rule h'_definite [symmetric])
-	finally show ?thesis .
-      qed
-    next
-      fix x1 c assume x1: "x1 \<in> H'"
-      show "h' (c \<cdot> x1) = c * (h' x1)"
-      proof -
-	from H' x1 have ax1: "c \<cdot> x1 \<in> H'"
-          by (rule vectorspace.mult_closed)
-	with x1 obtain y a y1 a1 where
-            cx1_rep: "c \<cdot> x1 = y + a \<cdot> x0" and y: "y \<in> H"
-          and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
-          unfolding H'_def sum_def lin_def by blast
-	
-	have ya: "c \<cdot> y1 = y \<and> c * a1 = a" using E HE _ y x0
-	proof (rule decomp_H')
-          from HE y1 show "c \<cdot> y1 \<in> H"
-            by (rule subspace.mult_closed)
-          from x0 and HE y y1
-          have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E" by auto
-          with x1_rep have "c \<cdot> y1 + (c * a1) \<cdot> x0 = c \<cdot> x1"
-            by (simp add: mult_assoc add_mult_distrib1)
-          also note cx1_rep
-          finally show "c \<cdot> y1 + (c * a1) \<cdot> x0 = y + a \<cdot> x0" .
-	qed
-	
-	from h'_def cx1_rep E HE y x0 have "h' (c \<cdot> x1) = h y + a * xi"
-          by (rule h'_definite)
-	also have "\<dots> = h (c \<cdot> y1) + (c * a1) * xi"
-          by (simp only: ya)
-	also from y1 have "h (c \<cdot> y1) = c * h y1"
-          by simp
-	also have "\<dots> + (c * a1) * xi = c * (h y1 + a1 * xi)"
-          by (simp only: right_distrib)
-	also from h'_def x1_rep E HE y1 x0 have "h y1 + a1 * xi = h' x1"
-          by (rule h'_definite [symmetric])
-	finally show ?thesis .
-      qed
-    }
-  qed
-qed
-
-text {* \medskip The linear extension @{text h'} of @{text h}
-  is bounded by the seminorm @{text p}. *}
-
-lemma h'_norm_pres:
-  assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
-      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
-    and H'_def: "H' \<equiv> H + lin x0"
-    and x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
-  assumes E: "vectorspace E" and HE: "subspace H E"
-    and "seminorm E p" and "linearform H h"
-  assumes a: "\<forall>y \<in> H. h y \<le> p y"
-    and a': "\<forall>y \<in> H. - p (y + x0) - h y \<le> xi \<and> xi \<le> p (y + x0) - h y"
-  shows "\<forall>x \<in> H'. h' x \<le> p x"
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [H E] by fact
-  interpret seminorm [E p] by fact
-  interpret linearform [H h] by fact
-  show ?thesis
-  proof
-    fix x assume x': "x \<in> H'"
-    show "h' x \<le> p x"
-    proof -
-      from a' have a1: "\<forall>ya \<in> H. - p (ya + x0) - h ya \<le> xi"
-	and a2: "\<forall>ya \<in> H. xi \<le> p (ya + x0) - h ya" by auto
-      from x' obtain y a where
-          x_rep: "x = y + a \<cdot> x0" and y: "y \<in> H"
-	unfolding H'_def sum_def lin_def by blast
-      from y have y': "y \<in> E" ..
-      from y have ay: "inverse a \<cdot> y \<in> H" by simp
-      
-      from h'_def x_rep E HE y x0 have "h' x = h y + a * xi"
-	by (rule h'_definite)
-      also have "\<dots> \<le> p (y + a \<cdot> x0)"
-      proof (rule linorder_cases)
-	assume z: "a = 0"
-	then have "h y + a * xi = h y" by simp
-	also from a y have "\<dots> \<le> p y" ..
-	also from x0 y' z have "p y = p (y + a \<cdot> x0)" by simp
-	finally show ?thesis .
-      next
-	txt {* In the case @{text "a < 0"}, we use @{text "a\<^sub>1"}
-          with @{text ya} taken as @{text "y / a"}: *}
-	assume lz: "a < 0" then have nz: "a \<noteq> 0" by simp
-	from a1 ay
-	have "- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y) \<le> xi" ..
-	with lz have "a * xi \<le>
-          a * (- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
-          by (simp add: mult_left_mono_neg order_less_imp_le)
-	
-	also have "\<dots> =
-          - a * (p (inverse a \<cdot> y + x0)) - a * (h (inverse a \<cdot> y))"
-	  by (simp add: right_diff_distrib)
-	also from lz x0 y' have "- a * (p (inverse a \<cdot> y + x0)) =
-          p (a \<cdot> (inverse a \<cdot> y + x0))"
-          by (simp add: abs_homogenous)
-	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
-          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
-	also from nz y have "a * (h (inverse a \<cdot> y)) =  h y"
-          by simp
-	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
-	then show ?thesis by simp
-      next
-	txt {* In the case @{text "a > 0"}, we use @{text "a\<^sub>2"}
-          with @{text ya} taken as @{text "y / a"}: *}
-	assume gz: "0 < a" then have nz: "a \<noteq> 0" by simp
-	from a2 ay
-	have "xi \<le> p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y)" ..
-	with gz have "a * xi \<le>
-          a * (p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
-          by simp
-	also have "\<dots> = a * p (inverse a \<cdot> y + x0) - a * h (inverse a \<cdot> y)"
-	  by (simp add: right_diff_distrib)
-	also from gz x0 y'
-	have "a * p (inverse a \<cdot> y + x0) = p (a \<cdot> (inverse a \<cdot> y + x0))"
-          by (simp add: abs_homogenous)
-	also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
-          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
-	also from nz y have "a * h (inverse a \<cdot> y) = h y"
-          by simp
-	finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
-	then show ?thesis by simp
-      qed
-      also from x_rep have "\<dots> = p x" by (simp only:)
-      finally show ?thesis .
-    qed
-  qed
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/HahnBanachLemmas.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(*<*)
-theory HahnBanachLemmas imports HahnBanachSupLemmas HahnBanachExtLemmas begin
-end
-(*>*)
\ No newline at end of file
--- a/src/HOL/Real/HahnBanach/HahnBanachSupLemmas.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,446 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/HahnBanachSupLemmas.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* The supremum w.r.t.~the function order *}
-
-theory HahnBanachSupLemmas
-imports FunctionNorm ZornLemma
-begin
-
-text {*
-  This section contains some lemmas that will be used in the proof of
-  the Hahn-Banach Theorem.  In this section the following context is
-  presumed.  Let @{text E} be a real vector space with a seminorm
-  @{text p} on @{text E}.  @{text F} is a subspace of @{text E} and
-  @{text f} a linear form on @{text F}. We consider a chain @{text c}
-  of norm-preserving extensions of @{text f}, such that @{text "\<Union>c =
-  graph H h"}.  We will show some properties about the limit function
-  @{text h}, i.e.\ the supremum of the chain @{text c}.
-
-  \medskip Let @{text c} be a chain of norm-preserving extensions of
-  the function @{text f} and let @{text "graph H h"} be the supremum
-  of @{text c}.  Every element in @{text H} is member of one of the
-  elements of the chain.
-*}
-lemmas [dest?] = chainD
-lemmas chainE2 [elim?] = chainD2 [elim_format, standard]
-
-lemma some_H'h't:
-  assumes M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and u: "graph H h = \<Union>c"
-    and x: "x \<in> H"
-  shows "\<exists>H' h'. graph H' h' \<in> c
-    \<and> (x, h x) \<in> graph H' h'
-    \<and> linearform H' h' \<and> H' \<unlhd> E
-    \<and> F \<unlhd> H' \<and> graph F f \<subseteq> graph H' h'
-    \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
-proof -
-  from x have "(x, h x) \<in> graph H h" ..
-  also from u have "\<dots> = \<Union>c" .
-  finally obtain g where gc: "g \<in> c" and gh: "(x, h x) \<in> g" by blast
-
-  from cM have "c \<subseteq> M" ..
-  with gc have "g \<in> M" ..
-  also from M have "\<dots> = norm_pres_extensions E p F f" .
-  finally obtain H' and h' where g: "g = graph H' h'"
-    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
-      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x" ..
-
-  from gc and g have "graph H' h' \<in> c" by (simp only:)
-  moreover from gh and g have "(x, h x) \<in> graph H' h'" by (simp only:)
-  ultimately show ?thesis using * by blast
-qed
-
-text {*
-  \medskip Let @{text c} be a chain of norm-preserving extensions of
-  the function @{text f} and let @{text "graph H h"} be the supremum
-  of @{text c}.  Every element in the domain @{text H} of the supremum
-  function is member of the domain @{text H'} of some function @{text
-  h'}, such that @{text h} extends @{text h'}.
-*}
-
-lemma some_H'h':
-  assumes M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and u: "graph H h = \<Union>c"
-    and x: "x \<in> H"
-  shows "\<exists>H' h'. x \<in> H' \<and> graph H' h' \<subseteq> graph H h
-    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
-    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
-proof -
-  from M cM u x obtain H' h' where
-      x_hx: "(x, h x) \<in> graph H' h'"
-    and c: "graph H' h' \<in> c"
-    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
-      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
-    by (rule some_H'h't [elim_format]) blast
-  from x_hx have "x \<in> H'" ..
-  moreover from cM u c have "graph H' h' \<subseteq> graph H h"
-    by (simp only: chain_ball_Union_upper)
-  ultimately show ?thesis using * by blast
-qed
-
-text {*
-  \medskip Any two elements @{text x} and @{text y} in the domain
-  @{text H} of the supremum function @{text h} are both in the domain
-  @{text H'} of some function @{text h'}, such that @{text h} extends
-  @{text h'}.
-*}
-
-lemma some_H'h'2:
-  assumes M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and u: "graph H h = \<Union>c"
-    and x: "x \<in> H"
-    and y: "y \<in> H"
-  shows "\<exists>H' h'. x \<in> H' \<and> y \<in> H'
-    \<and> graph H' h' \<subseteq> graph H h
-    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
-    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
-proof -
-  txt {* @{text y} is in the domain @{text H''} of some function @{text h''},
-  such that @{text h} extends @{text h''}. *}
-
-  from M cM u and y obtain H' h' where
-      y_hy: "(y, h y) \<in> graph H' h'"
-    and c': "graph H' h' \<in> c"
-    and * :
-      "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
-      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
-    by (rule some_H'h't [elim_format]) blast
-
-  txt {* @{text x} is in the domain @{text H'} of some function @{text h'},
-    such that @{text h} extends @{text h'}. *}
-
-  from M cM u and x obtain H'' h'' where
-      x_hx: "(x, h x) \<in> graph H'' h''"
-    and c'': "graph H'' h'' \<in> c"
-    and ** :
-      "linearform H'' h''"  "H'' \<unlhd> E"  "F \<unlhd> H''"
-      "graph F f \<subseteq> graph H'' h''"  "\<forall>x \<in> H''. h'' x \<le> p x"
-    by (rule some_H'h't [elim_format]) blast
-
-  txt {* Since both @{text h'} and @{text h''} are elements of the chain,
-    @{text h''} is an extension of @{text h'} or vice versa. Thus both
-    @{text x} and @{text y} are contained in the greater
-    one. \label{cases1}*}
-
-  from cM c'' c' have "graph H'' h'' \<subseteq> graph H' h' \<or> graph H' h' \<subseteq> graph H'' h''"
-    (is "?case1 \<or> ?case2") ..
-  then show ?thesis
-  proof
-    assume ?case1
-    have "(x, h x) \<in> graph H'' h''" by fact
-    also have "\<dots> \<subseteq> graph H' h'" by fact
-    finally have xh:"(x, h x) \<in> graph H' h'" .
-    then have "x \<in> H'" ..
-    moreover from y_hy have "y \<in> H'" ..
-    moreover from cM u and c' have "graph H' h' \<subseteq> graph H h"
-      by (simp only: chain_ball_Union_upper)
-    ultimately show ?thesis using * by blast
-  next
-    assume ?case2
-    from x_hx have "x \<in> H''" ..
-    moreover {
-      have "(y, h y) \<in> graph H' h'" by (rule y_hy)
-      also have "\<dots> \<subseteq> graph H'' h''" by fact
-      finally have "(y, h y) \<in> graph H'' h''" .
-    } then have "y \<in> H''" ..
-    moreover from cM u and c'' have "graph H'' h'' \<subseteq> graph H h"
-      by (simp only: chain_ball_Union_upper)
-    ultimately show ?thesis using ** by blast
-  qed
-qed
-
-text {*
-  \medskip The relation induced by the graph of the supremum of a
-  chain @{text c} is definite, i.~e.~t is the graph of a function.
-*}
-
-lemma sup_definite:
-  assumes M_def: "M \<equiv> norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and xy: "(x, y) \<in> \<Union>c"
-    and xz: "(x, z) \<in> \<Union>c"
-  shows "z = y"
-proof -
-  from cM have c: "c \<subseteq> M" ..
-  from xy obtain G1 where xy': "(x, y) \<in> G1" and G1: "G1 \<in> c" ..
-  from xz obtain G2 where xz': "(x, z) \<in> G2" and G2: "G2 \<in> c" ..
-
-  from G1 c have "G1 \<in> M" ..
-  then obtain H1 h1 where G1_rep: "G1 = graph H1 h1"
-    unfolding M_def by blast
-
-  from G2 c have "G2 \<in> M" ..
-  then obtain H2 h2 where G2_rep: "G2 = graph H2 h2"
-    unfolding M_def by blast
-
-  txt {* @{text "G\<^sub>1"} is contained in @{text "G\<^sub>2"}
-    or vice versa, since both @{text "G\<^sub>1"} and @{text
-    "G\<^sub>2"} are members of @{text c}. \label{cases2}*}
-
-  from cM G1 G2 have "G1 \<subseteq> G2 \<or> G2 \<subseteq> G1" (is "?case1 \<or> ?case2") ..
-  then show ?thesis
-  proof
-    assume ?case1
-    with xy' G2_rep have "(x, y) \<in> graph H2 h2" by blast
-    then have "y = h2 x" ..
-    also
-    from xz' G2_rep have "(x, z) \<in> graph H2 h2" by (simp only:)
-    then have "z = h2 x" ..
-    finally show ?thesis .
-  next
-    assume ?case2
-    with xz' G1_rep have "(x, z) \<in> graph H1 h1" by blast
-    then have "z = h1 x" ..
-    also
-    from xy' G1_rep have "(x, y) \<in> graph H1 h1" by (simp only:)
-    then have "y = h1 x" ..
-    finally show ?thesis ..
-  qed
-qed
-
-text {*
-  \medskip The limit function @{text h} is linear. Every element
-  @{text x} in the domain of @{text h} is in the domain of a function
-  @{text h'} in the chain of norm preserving extensions.  Furthermore,
-  @{text h} is an extension of @{text h'} so the function values of
-  @{text x} are identical for @{text h'} and @{text h}.  Finally, the
-  function @{text h'} is linear by construction of @{text M}.
-*}
-
-lemma sup_lf:
-  assumes M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and u: "graph H h = \<Union>c"
-  shows "linearform H h"
-proof
-  fix x y assume x: "x \<in> H" and y: "y \<in> H"
-  with M cM u obtain H' h' where
-        x': "x \<in> H'" and y': "y \<in> H'"
-      and b: "graph H' h' \<subseteq> graph H h"
-      and linearform: "linearform H' h'"
-      and subspace: "H' \<unlhd> E"
-    by (rule some_H'h'2 [elim_format]) blast
-
-  show "h (x + y) = h x + h y"
-  proof -
-    from linearform x' y' have "h' (x + y) = h' x + h' y"
-      by (rule linearform.add)
-    also from b x' have "h' x = h x" ..
-    also from b y' have "h' y = h y" ..
-    also from subspace x' y' have "x + y \<in> H'"
-      by (rule subspace.add_closed)
-    with b have "h' (x + y) = h (x + y)" ..
-    finally show ?thesis .
-  qed
-next
-  fix x a assume x: "x \<in> H"
-  with M cM u obtain H' h' where
-        x': "x \<in> H'"
-      and b: "graph H' h' \<subseteq> graph H h"
-      and linearform: "linearform H' h'"
-      and subspace: "H' \<unlhd> E"
-    by (rule some_H'h' [elim_format]) blast
-
-  show "h (a \<cdot> x) = a * h x"
-  proof -
-    from linearform x' have "h' (a \<cdot> x) = a * h' x"
-      by (rule linearform.mult)
-    also from b x' have "h' x = h x" ..
-    also from subspace x' have "a \<cdot> x \<in> H'"
-      by (rule subspace.mult_closed)
-    with b have "h' (a \<cdot> x) = h (a \<cdot> x)" ..
-    finally show ?thesis .
-  qed
-qed
-
-text {*
-  \medskip The limit of a non-empty chain of norm preserving
-  extensions of @{text f} is an extension of @{text f}, since every
-  element of the chain is an extension of @{text f} and the supremum
-  is an extension for every element of the chain.
-*}
-
-lemma sup_ext:
-  assumes graph: "graph H h = \<Union>c"
-    and M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and ex: "\<exists>x. x \<in> c"
-  shows "graph F f \<subseteq> graph H h"
-proof -
-  from ex obtain x where xc: "x \<in> c" ..
-  from cM have "c \<subseteq> M" ..
-  with xc have "x \<in> M" ..
-  with M have "x \<in> norm_pres_extensions E p F f"
-    by (simp only:)
-  then obtain G g where "x = graph G g" and "graph F f \<subseteq> graph G g" ..
-  then have "graph F f \<subseteq> x" by (simp only:)
-  also from xc have "\<dots> \<subseteq> \<Union>c" by blast
-  also from graph have "\<dots> = graph H h" ..
-  finally show ?thesis .
-qed
-
-text {*
-  \medskip The domain @{text H} of the limit function is a superspace
-  of @{text F}, since @{text F} is a subset of @{text H}. The
-  existence of the @{text 0} element in @{text F} and the closure
-  properties follow from the fact that @{text F} is a vector space.
-*}
-
-lemma sup_supF:
-  assumes graph: "graph H h = \<Union>c"
-    and M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and ex: "\<exists>x. x \<in> c"
-    and FE: "F \<unlhd> E"
-  shows "F \<unlhd> H"
-proof
-  from FE show "F \<noteq> {}" by (rule subspace.non_empty)
-  from graph M cM ex have "graph F f \<subseteq> graph H h" by (rule sup_ext)
-  then show "F \<subseteq> H" ..
-  fix x y assume "x \<in> F" and "y \<in> F"
-  with FE show "x + y \<in> F" by (rule subspace.add_closed)
-next
-  fix x a assume "x \<in> F"
-  with FE show "a \<cdot> x \<in> F" by (rule subspace.mult_closed)
-qed
-
-text {*
-  \medskip The domain @{text H} of the limit function is a subspace of
-  @{text E}.
-*}
-
-lemma sup_subE:
-  assumes graph: "graph H h = \<Union>c"
-    and M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-    and ex: "\<exists>x. x \<in> c"
-    and FE: "F \<unlhd> E"
-    and E: "vectorspace E"
-  shows "H \<unlhd> E"
-proof
-  show "H \<noteq> {}"
-  proof -
-    from FE E have "0 \<in> F" by (rule subspace.zero)
-    also from graph M cM ex FE have "F \<unlhd> H" by (rule sup_supF)
-    then have "F \<subseteq> H" ..
-    finally show ?thesis by blast
-  qed
-  show "H \<subseteq> E"
-  proof
-    fix x assume "x \<in> H"
-    with M cM graph
-    obtain H' h' where x: "x \<in> H'" and H'E: "H' \<unlhd> E"
-      by (rule some_H'h' [elim_format]) blast
-    from H'E have "H' \<subseteq> E" ..
-    with x show "x \<in> E" ..
-  qed
-  fix x y assume x: "x \<in> H" and y: "y \<in> H"
-  show "x + y \<in> H"
-  proof -
-    from M cM graph x y obtain H' h' where
-          x': "x \<in> H'" and y': "y \<in> H'" and H'E: "H' \<unlhd> E"
-        and graphs: "graph H' h' \<subseteq> graph H h"
-      by (rule some_H'h'2 [elim_format]) blast
-    from H'E x' y' have "x + y \<in> H'"
-      by (rule subspace.add_closed)
-    also from graphs have "H' \<subseteq> H" ..
-    finally show ?thesis .
-  qed
-next
-  fix x a assume x: "x \<in> H"
-  show "a \<cdot> x \<in> H"
-  proof -
-    from M cM graph x
-    obtain H' h' where x': "x \<in> H'" and H'E: "H' \<unlhd> E"
-        and graphs: "graph H' h' \<subseteq> graph H h"
-      by (rule some_H'h' [elim_format]) blast
-    from H'E x' have "a \<cdot> x \<in> H'" by (rule subspace.mult_closed)
-    also from graphs have "H' \<subseteq> H" ..
-    finally show ?thesis .
-  qed
-qed
-
-text {*
-  \medskip The limit function is bounded by the norm @{text p} as
-  well, since all elements in the chain are bounded by @{text p}.
-*}
-
-lemma sup_norm_pres:
-  assumes graph: "graph H h = \<Union>c"
-    and M: "M = norm_pres_extensions E p F f"
-    and cM: "c \<in> chain M"
-  shows "\<forall>x \<in> H. h x \<le> p x"
-proof
-  fix x assume "x \<in> H"
-  with M cM graph obtain H' h' where x': "x \<in> H'"
-      and graphs: "graph H' h' \<subseteq> graph H h"
-      and a: "\<forall>x \<in> H'. h' x \<le> p x"
-    by (rule some_H'h' [elim_format]) blast
-  from graphs x' have [symmetric]: "h' x = h x" ..
-  also from a x' have "h' x \<le> p x " ..
-  finally show "h x \<le> p x" .
-qed
-
-text {*
-  \medskip The following lemma is a property of linear forms on real
-  vector spaces. It will be used for the lemma @{text abs_HahnBanach}
-  (see page \pageref{abs-HahnBanach}). \label{abs-ineq-iff} For real
-  vector spaces the following inequations are equivalent:
-  \begin{center}
-  \begin{tabular}{lll}
-  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
-  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
-  \end{tabular}
-  \end{center}
-*}
-
-lemma abs_ineq_iff:
-  assumes "subspace H E" and "vectorspace E" and "seminorm E p"
-    and "linearform H h"
-  shows "(\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x) = (\<forall>x \<in> H. h x \<le> p x)" (is "?L = ?R")
-proof
-  interpret subspace [H E] by fact
-  interpret vectorspace [E] by fact
-  interpret seminorm [E p] by fact
-  interpret linearform [H h] by fact
-  have H: "vectorspace H" using `vectorspace E` ..
-  {
-    assume l: ?L
-    show ?R
-    proof
-      fix x assume x: "x \<in> H"
-      have "h x \<le> \<bar>h x\<bar>" by arith
-      also from l x have "\<dots> \<le> p x" ..
-      finally show "h x \<le> p x" .
-    qed
-  next
-    assume r: ?R
-    show ?L
-    proof
-      fix x assume x: "x \<in> H"
-      show "\<And>a b :: real. - a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> \<bar>b\<bar> \<le> a"
-        by arith
-      from `linearform H h` and H x
-      have "- h x = h (- x)" by (rule linearform.neg [symmetric])
-      also
-      from H x have "- x \<in> H" by (rule vectorspace.neg_closed)
-      with r have "h (- x) \<le> p (- x)" ..
-      also have "\<dots> = p x"
-	using `seminorm E p` `vectorspace E`
-      proof (rule seminorm.minus)
-        from x show "x \<in> E" ..
-      qed
-      finally have "- h x \<le> p x" .
-      then show "- p x \<le> h x" by simp
-      from r x show "h x \<le> p x" ..
-    qed
-  }
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/Linearform.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,61 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/Linearform.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Linearforms *}
-
-theory Linearform
-imports VectorSpace
-begin
-
-text {*
-  A \emph{linear form} is a function on a vector space into the reals
-  that is additive and multiplicative.
-*}
-
-locale linearform = var V + var f +
-  constrains V :: "'a\<Colon>{minus, plus, zero, uminus} set"
-  assumes add [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x + y) = f x + f y"
-    and mult [iff]: "x \<in> V \<Longrightarrow> f (a \<cdot> x) = a * f x"
-
-declare linearform.intro [intro?]
-
-lemma (in linearform) neg [iff]:
-  assumes "vectorspace V"
-  shows "x \<in> V \<Longrightarrow> f (- x) = - f x"
-proof -
-  interpret vectorspace [V] by fact
-  assume x: "x \<in> V"
-  then have "f (- x) = f ((- 1) \<cdot> x)" by (simp add: negate_eq1)
-  also from x have "\<dots> = (- 1) * (f x)" by (rule mult)
-  also from x have "\<dots> = - (f x)" by simp
-  finally show ?thesis .
-qed
-
-lemma (in linearform) diff [iff]:
-  assumes "vectorspace V"
-  shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x - y) = f x - f y"
-proof -
-  interpret vectorspace [V] by fact
-  assume x: "x \<in> V" and y: "y \<in> V"
-  then have "x - y = x + - y" by (rule diff_eq1)
-  also have "f \<dots> = f x + f (- y)" by (rule add) (simp_all add: x y)
-  also have "f (- y) = - f y" using `vectorspace V` y by (rule neg)
-  finally show ?thesis by simp
-qed
-
-text {* Every linear form yields @{text 0} for the @{text 0} vector. *}
-
-lemma (in linearform) zero [iff]:
-  assumes "vectorspace V"
-  shows "f 0 = 0"
-proof -
-  interpret vectorspace [V] by fact
-  have "f 0 = f (0 - 0)" by simp
-  also have "\<dots> = f 0 - f 0" using `vectorspace V` by (rule diff) simp_all
-  also have "\<dots> = 0" by simp
-  finally show ?thesis .
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/NormedSpace.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,118 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/NormedSpace.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Normed vector spaces *}
-
-theory NormedSpace
-imports Subspace
-begin
-
-subsection {* Quasinorms *}
-
-text {*
-  A \emph{seminorm} @{text "\<parallel>\<cdot>\<parallel>"} is a function on a real vector space
-  into the reals that has the following properties: it is positive
-  definite, absolute homogenous and subadditive.
-*}
-
-locale norm_syntax =
-  fixes norm :: "'a \<Rightarrow> real"    ("\<parallel>_\<parallel>")
-
-locale seminorm = var V + norm_syntax +
-  constrains V :: "'a\<Colon>{minus, plus, zero, uminus} set"
-  assumes ge_zero [iff?]: "x \<in> V \<Longrightarrow> 0 \<le> \<parallel>x\<parallel>"
-    and abs_homogenous [iff?]: "x \<in> V \<Longrightarrow> \<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>"
-    and subadditive [iff?]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> \<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>"
-
-declare seminorm.intro [intro?]
-
-lemma (in seminorm) diff_subadditive:
-  assumes "vectorspace V"
-  shows "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> \<parallel>x - y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>"
-proof -
-  interpret vectorspace [V] by fact
-  assume x: "x \<in> V" and y: "y \<in> V"
-  then have "x - y = x + - 1 \<cdot> y"
-    by (simp add: diff_eq2 negate_eq2a)
-  also from x y have "\<parallel>\<dots>\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>- 1 \<cdot> y\<parallel>"
-    by (simp add: subadditive)
-  also from y have "\<parallel>- 1 \<cdot> y\<parallel> = \<bar>- 1\<bar> * \<parallel>y\<parallel>"
-    by (rule abs_homogenous)
-  also have "\<dots> = \<parallel>y\<parallel>" by simp
-  finally show ?thesis .
-qed
-
-lemma (in seminorm) minus:
-  assumes "vectorspace V"
-  shows "x \<in> V \<Longrightarrow> \<parallel>- x\<parallel> = \<parallel>x\<parallel>"
-proof -
-  interpret vectorspace [V] by fact
-  assume x: "x \<in> V"
-  then have "- x = - 1 \<cdot> x" by (simp only: negate_eq1)
-  also from x have "\<parallel>\<dots>\<parallel> = \<bar>- 1\<bar> * \<parallel>x\<parallel>"
-    by (rule abs_homogenous)
-  also have "\<dots> = \<parallel>x\<parallel>" by simp
-  finally show ?thesis .
-qed
-
-
-subsection {* Norms *}
-
-text {*
-  A \emph{norm} @{text "\<parallel>\<cdot>\<parallel>"} is a seminorm that maps only the
-  @{text 0} vector to @{text 0}.
-*}
-
-locale norm = seminorm +
-  assumes zero_iff [iff]: "x \<in> V \<Longrightarrow> (\<parallel>x\<parallel> = 0) = (x = 0)"
-
-
-subsection {* Normed vector spaces *}
-
-text {*
-  A vector space together with a norm is called a \emph{normed
-  space}.
-*}
-
-locale normed_vectorspace = vectorspace + norm
-
-declare normed_vectorspace.intro [intro?]
-
-lemma (in normed_vectorspace) gt_zero [intro?]:
-  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> 0 < \<parallel>x\<parallel>"
-proof -
-  assume x: "x \<in> V" and neq: "x \<noteq> 0"
-  from x have "0 \<le> \<parallel>x\<parallel>" ..
-  also have [symmetric]: "\<dots> \<noteq> 0"
-  proof
-    assume "\<parallel>x\<parallel> = 0"
-    with x have "x = 0" by simp
-    with neq show False by contradiction
-  qed
-  finally show ?thesis .
-qed
-
-text {*
-  Any subspace of a normed vector space is again a normed vectorspace.
-*}
-
-lemma subspace_normed_vs [intro?]:
-  fixes F E norm
-  assumes "subspace F E" "normed_vectorspace E norm"
-  shows "normed_vectorspace F norm"
-proof -
-  interpret subspace [F E] by fact
-  interpret normed_vectorspace [E norm] by fact
-  show ?thesis
-  proof
-    show "vectorspace F" by (rule vectorspace) unfold_locales
-  next
-    have "NormedSpace.norm E norm" ..
-    with subset show "NormedSpace.norm F norm"
-      by (simp add: norm_def seminorm_def norm_axioms_def)
-  qed
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/README.html	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,38 +0,0 @@
-<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
-
-<!-- $Id$ -->
-
-<HTML>
-
-<HEAD>
-  <meta http-equiv="Content-Type" content="text/html; charset=iso-8859-1">
-  <TITLE>HOL/Real/HahnBanach/README</TITLE>
-</HEAD>
-
-<BODY>
-
-<H3>The Hahn-Banach Theorem for Real Vector Spaces (Isabelle/Isar)</H3>
-
-Author: Gertrud Bauer, Technische Universit&auml;t M&uuml;nchen<P>
-
-This directory contains the proof of the Hahn-Banach theorem for real vectorspaces,
-following H. Heuser, Funktionalanalysis, p. 228 -232.
-The Hahn-Banach theorem is one of the fundamental theorems of functioal analysis.
-It is a conclusion of Zorn's lemma.<P>
-
-Two different formaulations of the theorem are presented, one for general real vectorspaces
-and its application to normed vectorspaces. <P>
-
-The theorem says, that every continous linearform, defined on arbitrary subspaces
-(not only one-dimensional subspaces), can be extended to a continous linearform on
-the whole vectorspace.
-
-
-<HR>
-
-<ADDRESS>
-<A NAME="bauerg@in.tum.de" HREF="mailto:bauerg@in.tum.de">bauerg@in.tum.de</A>
-</ADDRESS>
-
-</BODY>
-</HTML>
--- a/src/HOL/Real/HahnBanach/ROOT.ML	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/ROOT.ML
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-
-The Hahn-Banach theorem for real vector spaces (Isabelle/Isar).
-*)
-
-time_use_thy "HahnBanach";
--- a/src/HOL/Real/HahnBanach/Subspace.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,514 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/Subspace.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Subspaces *}
-
-theory Subspace
-imports VectorSpace
-begin
-
-subsection {* Definition *}
-
-text {*
-  A non-empty subset @{text U} of a vector space @{text V} is a
-  \emph{subspace} of @{text V}, iff @{text U} is closed under addition
-  and scalar multiplication.
-*}
-
-locale subspace = var U + var V +
-  constrains U :: "'a\<Colon>{minus, plus, zero, uminus} set"
-  assumes non_empty [iff, intro]: "U \<noteq> {}"
-    and subset [iff]: "U \<subseteq> V"
-    and add_closed [iff]: "x \<in> U \<Longrightarrow> y \<in> U \<Longrightarrow> x + y \<in> U"
-    and mult_closed [iff]: "x \<in> U \<Longrightarrow> a \<cdot> x \<in> U"
-
-notation (symbols)
-  subspace  (infix "\<unlhd>" 50)
-
-declare vectorspace.intro [intro?] subspace.intro [intro?]
-
-lemma subspace_subset [elim]: "U \<unlhd> V \<Longrightarrow> U \<subseteq> V"
-  by (rule subspace.subset)
-
-lemma (in subspace) subsetD [iff]: "x \<in> U \<Longrightarrow> x \<in> V"
-  using subset by blast
-
-lemma subspaceD [elim]: "U \<unlhd> V \<Longrightarrow> x \<in> U \<Longrightarrow> x \<in> V"
-  by (rule subspace.subsetD)
-
-lemma rev_subspaceD [elim?]: "x \<in> U \<Longrightarrow> U \<unlhd> V \<Longrightarrow> x \<in> V"
-  by (rule subspace.subsetD)
-
-lemma (in subspace) diff_closed [iff]:
-  assumes "vectorspace V"
-  assumes x: "x \<in> U" and y: "y \<in> U"
-  shows "x - y \<in> U"
-proof -
-  interpret vectorspace [V] by fact
-  from x y show ?thesis by (simp add: diff_eq1 negate_eq1)
-qed
-
-text {*
-  \medskip Similar as for linear spaces, the existence of the zero
-  element in every subspace follows from the non-emptiness of the
-  carrier set and by vector space laws.
-*}
-
-lemma (in subspace) zero [intro]:
-  assumes "vectorspace V"
-  shows "0 \<in> U"
-proof -
-  interpret vectorspace [V] by fact
-  have "U \<noteq> {}" by (rule U_V.non_empty)
-  then obtain x where x: "x \<in> U" by blast
-  then have "x \<in> V" .. then have "0 = x - x" by simp
-  also from `vectorspace V` x x have "\<dots> \<in> U" by (rule U_V.diff_closed)
-  finally show ?thesis .
-qed
-
-lemma (in subspace) neg_closed [iff]:
-  assumes "vectorspace V"
-  assumes x: "x \<in> U"
-  shows "- x \<in> U"
-proof -
-  interpret vectorspace [V] by fact
-  from x show ?thesis by (simp add: negate_eq1)
-qed
-
-text {* \medskip Further derived laws: every subspace is a vector space. *}
-
-lemma (in subspace) vectorspace [iff]:
-  assumes "vectorspace V"
-  shows "vectorspace U"
-proof -
-  interpret vectorspace [V] by fact
-  show ?thesis
-  proof
-    show "U \<noteq> {}" ..
-    fix x y z assume x: "x \<in> U" and y: "y \<in> U" and z: "z \<in> U"
-    fix a b :: real
-    from x y show "x + y \<in> U" by simp
-    from x show "a \<cdot> x \<in> U" by simp
-    from x y z show "(x + y) + z = x + (y + z)" by (simp add: add_ac)
-    from x y show "x + y = y + x" by (simp add: add_ac)
-    from x show "x - x = 0" by simp
-    from x show "0 + x = x" by simp
-    from x y show "a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y" by (simp add: distrib)
-    from x show "(a + b) \<cdot> x = a \<cdot> x + b \<cdot> x" by (simp add: distrib)
-    from x show "(a * b) \<cdot> x = a \<cdot> b \<cdot> x" by (simp add: mult_assoc)
-    from x show "1 \<cdot> x = x" by simp
-    from x show "- x = - 1 \<cdot> x" by (simp add: negate_eq1)
-    from x y show "x - y = x + - y" by (simp add: diff_eq1)
-  qed
-qed
-
-
-text {* The subspace relation is reflexive. *}
-
-lemma (in vectorspace) subspace_refl [intro]: "V \<unlhd> V"
-proof
-  show "V \<noteq> {}" ..
-  show "V \<subseteq> V" ..
-  fix x y assume x: "x \<in> V" and y: "y \<in> V"
-  fix a :: real
-  from x y show "x + y \<in> V" by simp
-  from x show "a \<cdot> x \<in> V" by simp
-qed
-
-text {* The subspace relation is transitive. *}
-
-lemma (in vectorspace) subspace_trans [trans]:
-  "U \<unlhd> V \<Longrightarrow> V \<unlhd> W \<Longrightarrow> U \<unlhd> W"
-proof
-  assume uv: "U \<unlhd> V" and vw: "V \<unlhd> W"
-  from uv show "U \<noteq> {}" by (rule subspace.non_empty)
-  show "U \<subseteq> W"
-  proof -
-    from uv have "U \<subseteq> V" by (rule subspace.subset)
-    also from vw have "V \<subseteq> W" by (rule subspace.subset)
-    finally show ?thesis .
-  qed
-  fix x y assume x: "x \<in> U" and y: "y \<in> U"
-  from uv and x y show "x + y \<in> U" by (rule subspace.add_closed)
-  from uv and x show "\<And>a. a \<cdot> x \<in> U" by (rule subspace.mult_closed)
-qed
-
-
-subsection {* Linear closure *}
-
-text {*
-  The \emph{linear closure} of a vector @{text x} is the set of all
-  scalar multiples of @{text x}.
-*}
-
-definition
-  lin :: "('a::{minus, plus, zero}) \<Rightarrow> 'a set" where
-  "lin x = {a \<cdot> x | a. True}"
-
-lemma linI [intro]: "y = a \<cdot> x \<Longrightarrow> y \<in> lin x"
-  unfolding lin_def by blast
-
-lemma linI' [iff]: "a \<cdot> x \<in> lin x"
-  unfolding lin_def by blast
-
-lemma linE [elim]: "x \<in> lin v \<Longrightarrow> (\<And>a::real. x = a \<cdot> v \<Longrightarrow> C) \<Longrightarrow> C"
-  unfolding lin_def by blast
-
-
-text {* Every vector is contained in its linear closure. *}
-
-lemma (in vectorspace) x_lin_x [iff]: "x \<in> V \<Longrightarrow> x \<in> lin x"
-proof -
-  assume "x \<in> V"
-  then have "x = 1 \<cdot> x" by simp
-  also have "\<dots> \<in> lin x" ..
-  finally show ?thesis .
-qed
-
-lemma (in vectorspace) "0_lin_x" [iff]: "x \<in> V \<Longrightarrow> 0 \<in> lin x"
-proof
-  assume "x \<in> V"
-  then show "0 = 0 \<cdot> x" by simp
-qed
-
-text {* Any linear closure is a subspace. *}
-
-lemma (in vectorspace) lin_subspace [intro]:
-  "x \<in> V \<Longrightarrow> lin x \<unlhd> V"
-proof
-  assume x: "x \<in> V"
-  then show "lin x \<noteq> {}" by (auto simp add: x_lin_x)
-  show "lin x \<subseteq> V"
-  proof
-    fix x' assume "x' \<in> lin x"
-    then obtain a where "x' = a \<cdot> x" ..
-    with x show "x' \<in> V" by simp
-  qed
-  fix x' x'' assume x': "x' \<in> lin x" and x'': "x'' \<in> lin x"
-  show "x' + x'' \<in> lin x"
-  proof -
-    from x' obtain a' where "x' = a' \<cdot> x" ..
-    moreover from x'' obtain a'' where "x'' = a'' \<cdot> x" ..
-    ultimately have "x' + x'' = (a' + a'') \<cdot> x"
-      using x by (simp add: distrib)
-    also have "\<dots> \<in> lin x" ..
-    finally show ?thesis .
-  qed
-  fix a :: real
-  show "a \<cdot> x' \<in> lin x"
-  proof -
-    from x' obtain a' where "x' = a' \<cdot> x" ..
-    with x have "a \<cdot> x' = (a * a') \<cdot> x" by (simp add: mult_assoc)
-    also have "\<dots> \<in> lin x" ..
-    finally show ?thesis .
-  qed
-qed
-
-
-text {* Any linear closure is a vector space. *}
-
-lemma (in vectorspace) lin_vectorspace [intro]:
-  assumes "x \<in> V"
-  shows "vectorspace (lin x)"
-proof -
-  from `x \<in> V` have "subspace (lin x) V"
-    by (rule lin_subspace)
-  from this and vectorspace_axioms show ?thesis
-    by (rule subspace.vectorspace)
-qed
-
-
-subsection {* Sum of two vectorspaces *}
-
-text {*
-  The \emph{sum} of two vectorspaces @{text U} and @{text V} is the
-  set of all sums of elements from @{text U} and @{text V}.
-*}
-
-instantiation "fun" :: (type, type) plus
-begin
-
-definition
-  sum_def: "plus_fun U V = {u + v | u v. u \<in> U \<and> v \<in> V}"  (* FIXME not fully general!? *)
-
-instance ..
-
-end
-
-lemma sumE [elim]:
-    "x \<in> U + V \<Longrightarrow> (\<And>u v. x = u + v \<Longrightarrow> u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> C) \<Longrightarrow> C"
-  unfolding sum_def by blast
-
-lemma sumI [intro]:
-    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> x = u + v \<Longrightarrow> x \<in> U + V"
-  unfolding sum_def by blast
-
-lemma sumI' [intro]:
-    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> u + v \<in> U + V"
-  unfolding sum_def by blast
-
-text {* @{text U} is a subspace of @{text "U + V"}. *}
-
-lemma subspace_sum1 [iff]:
-  assumes "vectorspace U" "vectorspace V"
-  shows "U \<unlhd> U + V"
-proof -
-  interpret vectorspace [U] by fact
-  interpret vectorspace [V] by fact
-  show ?thesis
-  proof
-    show "U \<noteq> {}" ..
-    show "U \<subseteq> U + V"
-    proof
-      fix x assume x: "x \<in> U"
-      moreover have "0 \<in> V" ..
-      ultimately have "x + 0 \<in> U + V" ..
-      with x show "x \<in> U + V" by simp
-    qed
-    fix x y assume x: "x \<in> U" and "y \<in> U"
-    then show "x + y \<in> U" by simp
-    from x show "\<And>a. a \<cdot> x \<in> U" by simp
-  qed
-qed
-
-text {* The sum of two subspaces is again a subspace. *}
-
-lemma sum_subspace [intro?]:
-  assumes "subspace U E" "vectorspace E" "subspace V E"
-  shows "U + V \<unlhd> E"
-proof -
-  interpret subspace [U E] by fact
-  interpret vectorspace [E] by fact
-  interpret subspace [V E] by fact
-  show ?thesis
-  proof
-    have "0 \<in> U + V"
-    proof
-      show "0 \<in> U" using `vectorspace E` ..
-      show "0 \<in> V" using `vectorspace E` ..
-      show "(0::'a) = 0 + 0" by simp
-    qed
-    then show "U + V \<noteq> {}" by blast
-    show "U + V \<subseteq> E"
-    proof
-      fix x assume "x \<in> U + V"
-      then obtain u v where "x = u + v" and
-	"u \<in> U" and "v \<in> V" ..
-      then show "x \<in> E" by simp
-    qed
-    fix x y assume x: "x \<in> U + V" and y: "y \<in> U + V"
-    show "x + y \<in> U + V"
-    proof -
-      from x obtain ux vx where "x = ux + vx" and "ux \<in> U" and "vx \<in> V" ..
-      moreover
-      from y obtain uy vy where "y = uy + vy" and "uy \<in> U" and "vy \<in> V" ..
-      ultimately
-      have "ux + uy \<in> U"
-	and "vx + vy \<in> V"
-	and "x + y = (ux + uy) + (vx + vy)"
-	using x y by (simp_all add: add_ac)
-      then show ?thesis ..
-    qed
-    fix a show "a \<cdot> x \<in> U + V"
-    proof -
-      from x obtain u v where "x = u + v" and "u \<in> U" and "v \<in> V" ..
-      then have "a \<cdot> u \<in> U" and "a \<cdot> v \<in> V"
-	and "a \<cdot> x = (a \<cdot> u) + (a \<cdot> v)" by (simp_all add: distrib)
-      then show ?thesis ..
-    qed
-  qed
-qed
-
-text{* The sum of two subspaces is a vectorspace. *}
-
-lemma sum_vs [intro?]:
-    "U \<unlhd> E \<Longrightarrow> V \<unlhd> E \<Longrightarrow> vectorspace E \<Longrightarrow> vectorspace (U + V)"
-  by (rule subspace.vectorspace) (rule sum_subspace)
-
-
-subsection {* Direct sums *}
-
-text {*
-  The sum of @{text U} and @{text V} is called \emph{direct}, iff the
-  zero element is the only common element of @{text U} and @{text
-  V}. For every element @{text x} of the direct sum of @{text U} and
-  @{text V} the decomposition in @{text "x = u + v"} with
-  @{text "u \<in> U"} and @{text "v \<in> V"} is unique.
-*}
-
-lemma decomp:
-  assumes "vectorspace E" "subspace U E" "subspace V E"
-  assumes direct: "U \<inter> V = {0}"
-    and u1: "u1 \<in> U" and u2: "u2 \<in> U"
-    and v1: "v1 \<in> V" and v2: "v2 \<in> V"
-    and sum: "u1 + v1 = u2 + v2"
-  shows "u1 = u2 \<and> v1 = v2"
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [U E] by fact
-  interpret subspace [V E] by fact
-  show ?thesis
-  proof
-    have U: "vectorspace U"  (* FIXME: use interpret *)
-      using `subspace U E` `vectorspace E` by (rule subspace.vectorspace)
-    have V: "vectorspace V"
-      using `subspace V E` `vectorspace E` by (rule subspace.vectorspace)
-    from u1 u2 v1 v2 and sum have eq: "u1 - u2 = v2 - v1"
-      by (simp add: add_diff_swap)
-    from u1 u2 have u: "u1 - u2 \<in> U"
-      by (rule vectorspace.diff_closed [OF U])
-    with eq have v': "v2 - v1 \<in> U" by (simp only:)
-    from v2 v1 have v: "v2 - v1 \<in> V"
-      by (rule vectorspace.diff_closed [OF V])
-    with eq have u': " u1 - u2 \<in> V" by (simp only:)
-    
-    show "u1 = u2"
-    proof (rule add_minus_eq)
-      from u1 show "u1 \<in> E" ..
-      from u2 show "u2 \<in> E" ..
-      from u u' and direct show "u1 - u2 = 0" by blast
-    qed
-    show "v1 = v2"
-    proof (rule add_minus_eq [symmetric])
-      from v1 show "v1 \<in> E" ..
-      from v2 show "v2 \<in> E" ..
-      from v v' and direct show "v2 - v1 = 0" by blast
-    qed
-  qed
-qed
-
-text {*
-  An application of the previous lemma will be used in the proof of
-  the Hahn-Banach Theorem (see page \pageref{decomp-H-use}): for any
-  element @{text "y + a \<cdot> x\<^sub>0"} of the direct sum of a
-  vectorspace @{text H} and the linear closure of @{text "x\<^sub>0"}
-  the components @{text "y \<in> H"} and @{text a} are uniquely
-  determined.
-*}
-
-lemma decomp_H':
-  assumes "vectorspace E" "subspace H E"
-  assumes y1: "y1 \<in> H" and y2: "y2 \<in> H"
-    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
-    and eq: "y1 + a1 \<cdot> x' = y2 + a2 \<cdot> x'"
-  shows "y1 = y2 \<and> a1 = a2"
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [H E] by fact
-  show ?thesis
-  proof
-    have c: "y1 = y2 \<and> a1 \<cdot> x' = a2 \<cdot> x'"
-    proof (rule decomp)
-      show "a1 \<cdot> x' \<in> lin x'" ..
-      show "a2 \<cdot> x' \<in> lin x'" ..
-      show "H \<inter> lin x' = {0}"
-      proof
-	show "H \<inter> lin x' \<subseteq> {0}"
-	proof
-          fix x assume x: "x \<in> H \<inter> lin x'"
-          then obtain a where xx': "x = a \<cdot> x'"
-            by blast
-          have "x = 0"
-          proof cases
-            assume "a = 0"
-            with xx' and x' show ?thesis by simp
-          next
-            assume a: "a \<noteq> 0"
-            from x have "x \<in> H" ..
-            with xx' have "inverse a \<cdot> a \<cdot> x' \<in> H" by simp
-            with a and x' have "x' \<in> H" by (simp add: mult_assoc2)
-            with `x' \<notin> H` show ?thesis by contradiction
-          qed
-          then show "x \<in> {0}" ..
-	qed
-	show "{0} \<subseteq> H \<inter> lin x'"
-	proof -
-          have "0 \<in> H" using `vectorspace E` ..
-          moreover have "0 \<in> lin x'" using `x' \<in> E` ..
-          ultimately show ?thesis by blast
-	qed
-      qed
-      show "lin x' \<unlhd> E" using `x' \<in> E` ..
-    qed (rule `vectorspace E`, rule `subspace H E`, rule y1, rule y2, rule eq)
-    then show "y1 = y2" ..
-    from c have "a1 \<cdot> x' = a2 \<cdot> x'" ..
-    with x' show "a1 = a2" by (simp add: mult_right_cancel)
-  qed
-qed
-
-text {*
-  Since for any element @{text "y + a \<cdot> x'"} of the direct sum of a
-  vectorspace @{text H} and the linear closure of @{text x'} the
-  components @{text "y \<in> H"} and @{text a} are unique, it follows from
-  @{text "y \<in> H"} that @{text "a = 0"}.
-*}
-
-lemma decomp_H'_H:
-  assumes "vectorspace E" "subspace H E"
-  assumes t: "t \<in> H"
-    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
-  shows "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [H E] by fact
-  show ?thesis
-  proof (rule, simp_all only: split_paired_all split_conv)
-    from t x' show "t = t + 0 \<cdot> x' \<and> t \<in> H" by simp
-    fix y and a assume ya: "t = y + a \<cdot> x' \<and> y \<in> H"
-    have "y = t \<and> a = 0"
-    proof (rule decomp_H')
-      from ya x' show "y + a \<cdot> x' = t + 0 \<cdot> x'" by simp
-      from ya show "y \<in> H" ..
-    qed (rule `vectorspace E`, rule `subspace H E`, rule t, (rule x')+)
-    with t x' show "(y, a) = (y + a \<cdot> x', 0)" by simp
-  qed
-qed
-
-text {*
-  The components @{text "y \<in> H"} and @{text a} in @{text "y + a \<cdot> x'"}
-  are unique, so the function @{text h'} defined by
-  @{text "h' (y + a \<cdot> x') = h y + a \<cdot> \<xi>"} is definite.
-*}
-
-lemma h'_definite:
-  fixes H
-  assumes h'_def:
-    "h' \<equiv> (\<lambda>x. let (y, a) = SOME (y, a). (x = y + a \<cdot> x' \<and> y \<in> H)
-                in (h y) + a * xi)"
-    and x: "x = y + a \<cdot> x'"
-  assumes "vectorspace E" "subspace H E"
-  assumes y: "y \<in> H"
-    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
-  shows "h' x = h y + a * xi"
-proof -
-  interpret vectorspace [E] by fact
-  interpret subspace [H E] by fact
-  from x y x' have "x \<in> H + lin x'" by auto
-  have "\<exists>!p. (\<lambda>(y, a). x = y + a \<cdot> x' \<and> y \<in> H) p" (is "\<exists>!p. ?P p")
-  proof (rule ex_ex1I)
-    from x y show "\<exists>p. ?P p" by blast
-    fix p q assume p: "?P p" and q: "?P q"
-    show "p = q"
-    proof -
-      from p have xp: "x = fst p + snd p \<cdot> x' \<and> fst p \<in> H"
-        by (cases p) simp
-      from q have xq: "x = fst q + snd q \<cdot> x' \<and> fst q \<in> H"
-        by (cases q) simp
-      have "fst p = fst q \<and> snd p = snd q"
-      proof (rule decomp_H')
-        from xp show "fst p \<in> H" ..
-        from xq show "fst q \<in> H" ..
-        from xp and xq show "fst p + snd p \<cdot> x' = fst q + snd q \<cdot> x'"
-          by simp
-      qed (rule `vectorspace E`, rule `subspace H E`, (rule x')+)
-      then show ?thesis by (cases p, cases q) simp
-    qed
-  qed
-  then have eq: "(SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H) = (y, a)"
-    by (rule some1_equality) (simp add: x y)
-  with h'_def show "h' x = h y + a * xi" by (simp add: Let_def)
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/VectorSpace.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,417 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/VectorSpace.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Vector spaces *}
-
-theory VectorSpace
-imports Real Bounds Zorn
-begin
-
-subsection {* Signature *}
-
-text {*
-  For the definition of real vector spaces a type @{typ 'a} of the
-  sort @{text "{plus, minus, zero}"} is considered, on which a real
-  scalar multiplication @{text \<cdot>} is declared.
-*}
-
-consts
-  prod  :: "real \<Rightarrow> 'a::{plus, minus, zero} \<Rightarrow> 'a"     (infixr "'(*')" 70)
-
-notation (xsymbols)
-  prod  (infixr "\<cdot>" 70)
-notation (HTML output)
-  prod  (infixr "\<cdot>" 70)
-
-
-subsection {* Vector space laws *}
-
-text {*
-  A \emph{vector space} is a non-empty set @{text V} of elements from
-  @{typ 'a} with the following vector space laws: The set @{text V} is
-  closed under addition and scalar multiplication, addition is
-  associative and commutative; @{text "- x"} is the inverse of @{text
-  x} w.~r.~t.~addition and @{text 0} is the neutral element of
-  addition.  Addition and multiplication are distributive; scalar
-  multiplication is associative and the real number @{text "1"} is
-  the neutral element of scalar multiplication.
-*}
-
-locale vectorspace = var V +
-  assumes non_empty [iff, intro?]: "V \<noteq> {}"
-    and add_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y \<in> V"
-    and mult_closed [iff]: "x \<in> V \<Longrightarrow> a \<cdot> x \<in> V"
-    and add_assoc: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y) + z = x + (y + z)"
-    and add_commute: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = y + x"
-    and diff_self [simp]: "x \<in> V \<Longrightarrow> x - x = 0"
-    and add_zero_left [simp]: "x \<in> V \<Longrightarrow> 0 + x = x"
-    and add_mult_distrib1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y"
-    and add_mult_distrib2: "x \<in> V \<Longrightarrow> (a + b) \<cdot> x = a \<cdot> x + b \<cdot> x"
-    and mult_assoc: "x \<in> V \<Longrightarrow> (a * b) \<cdot> x = a \<cdot> (b \<cdot> x)"
-    and mult_1 [simp]: "x \<in> V \<Longrightarrow> 1 \<cdot> x = x"
-    and negate_eq1: "x \<in> V \<Longrightarrow> - x = (- 1) \<cdot> x"
-    and diff_eq1: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = x + - y"
-
-lemma (in vectorspace) negate_eq2: "x \<in> V \<Longrightarrow> (- 1) \<cdot> x = - x"
-  by (rule negate_eq1 [symmetric])
-
-lemma (in vectorspace) negate_eq2a: "x \<in> V \<Longrightarrow> -1 \<cdot> x = - x"
-  by (simp add: negate_eq1)
-
-lemma (in vectorspace) diff_eq2: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + - y = x - y"
-  by (rule diff_eq1 [symmetric])
-
-lemma (in vectorspace) diff_closed [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y \<in> V"
-  by (simp add: diff_eq1 negate_eq1)
-
-lemma (in vectorspace) neg_closed [iff]: "x \<in> V \<Longrightarrow> - x \<in> V"
-  by (simp add: negate_eq1)
-
-lemma (in vectorspace) add_left_commute:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> x + (y + z) = y + (x + z)"
-proof -
-  assume xyz: "x \<in> V"  "y \<in> V"  "z \<in> V"
-  then have "x + (y + z) = (x + y) + z"
-    by (simp only: add_assoc)
-  also from xyz have "\<dots> = (y + x) + z" by (simp only: add_commute)
-  also from xyz have "\<dots> = y + (x + z)" by (simp only: add_assoc)
-  finally show ?thesis .
-qed
-
-theorems (in vectorspace) add_ac =
-  add_assoc add_commute add_left_commute
-
-
-text {* The existence of the zero element of a vector space
-  follows from the non-emptiness of carrier set. *}
-
-lemma (in vectorspace) zero [iff]: "0 \<in> V"
-proof -
-  from non_empty obtain x where x: "x \<in> V" by blast
-  then have "0 = x - x" by (rule diff_self [symmetric])
-  also from x x have "\<dots> \<in> V" by (rule diff_closed)
-  finally show ?thesis .
-qed
-
-lemma (in vectorspace) add_zero_right [simp]:
-  "x \<in> V \<Longrightarrow>  x + 0 = x"
-proof -
-  assume x: "x \<in> V"
-  from this and zero have "x + 0 = 0 + x" by (rule add_commute)
-  also from x have "\<dots> = x" by (rule add_zero_left)
-  finally show ?thesis .
-qed
-
-lemma (in vectorspace) mult_assoc2:
-    "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = (a * b) \<cdot> x"
-  by (simp only: mult_assoc)
-
-lemma (in vectorspace) diff_mult_distrib1:
-    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<cdot> (x - y) = a \<cdot> x - a \<cdot> y"
-  by (simp add: diff_eq1 negate_eq1 add_mult_distrib1 mult_assoc2)
-
-lemma (in vectorspace) diff_mult_distrib2:
-  "x \<in> V \<Longrightarrow> (a - b) \<cdot> x = a \<cdot> x - (b \<cdot> x)"
-proof -
-  assume x: "x \<in> V"
-  have " (a - b) \<cdot> x = (a + - b) \<cdot> x"
-    by (simp add: real_diff_def)
-  also from x have "\<dots> = a \<cdot> x + (- b) \<cdot> x"
-    by (rule add_mult_distrib2)
-  also from x have "\<dots> = a \<cdot> x + - (b \<cdot> x)"
-    by (simp add: negate_eq1 mult_assoc2)
-  also from x have "\<dots> = a \<cdot> x - (b \<cdot> x)"
-    by (simp add: diff_eq1)
-  finally show ?thesis .
-qed
-
-lemmas (in vectorspace) distrib =
-  add_mult_distrib1 add_mult_distrib2
-  diff_mult_distrib1 diff_mult_distrib2
-
-
-text {* \medskip Further derived laws: *}
-
-lemma (in vectorspace) mult_zero_left [simp]:
-  "x \<in> V \<Longrightarrow> 0 \<cdot> x = 0"
-proof -
-  assume x: "x \<in> V"
-  have "0 \<cdot> x = (1 - 1) \<cdot> x" by simp
-  also have "\<dots> = (1 + - 1) \<cdot> x" by simp
-  also from x have "\<dots> =  1 \<cdot> x + (- 1) \<cdot> x"
-    by (rule add_mult_distrib2)
-  also from x have "\<dots> = x + (- 1) \<cdot> x" by simp
-  also from x have "\<dots> = x + - x" by (simp add: negate_eq2a)
-  also from x have "\<dots> = x - x" by (simp add: diff_eq2)
-  also from x have "\<dots> = 0" by simp
-  finally show ?thesis .
-qed
-
-lemma (in vectorspace) mult_zero_right [simp]:
-  "a \<cdot> 0 = (0::'a)"
-proof -
-  have "a \<cdot> 0 = a \<cdot> (0 - (0::'a))" by simp
-  also have "\<dots> =  a \<cdot> 0 - a \<cdot> 0"
-    by (rule diff_mult_distrib1) simp_all
-  also have "\<dots> = 0" by simp
-  finally show ?thesis .
-qed
-
-lemma (in vectorspace) minus_mult_cancel [simp]:
-    "x \<in> V \<Longrightarrow> (- a) \<cdot> - x = a \<cdot> x"
-  by (simp add: negate_eq1 mult_assoc2)
-
-lemma (in vectorspace) add_minus_left_eq_diff:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + y = y - x"
-proof -
-  assume xy: "x \<in> V"  "y \<in> V"
-  then have "- x + y = y + - x" by (simp add: add_commute)
-  also from xy have "\<dots> = y - x" by (simp add: diff_eq1)
-  finally show ?thesis .
-qed
-
-lemma (in vectorspace) add_minus [simp]:
-    "x \<in> V \<Longrightarrow> x + - x = 0"
-  by (simp add: diff_eq2)
-
-lemma (in vectorspace) add_minus_left [simp]:
-    "x \<in> V \<Longrightarrow> - x + x = 0"
-  by (simp add: diff_eq2 add_commute)
-
-lemma (in vectorspace) minus_minus [simp]:
-    "x \<in> V \<Longrightarrow> - (- x) = x"
-  by (simp add: negate_eq1 mult_assoc2)
-
-lemma (in vectorspace) minus_zero [simp]:
-    "- (0::'a) = 0"
-  by (simp add: negate_eq1)
-
-lemma (in vectorspace) minus_zero_iff [simp]:
-  "x \<in> V \<Longrightarrow> (- x = 0) = (x = 0)"
-proof
-  assume x: "x \<in> V"
-  {
-    from x have "x = - (- x)" by (simp add: minus_minus)
-    also assume "- x = 0"
-    also have "- \<dots> = 0" by (rule minus_zero)
-    finally show "x = 0" .
-  next
-    assume "x = 0"
-    then show "- x = 0" by simp
-  }
-qed
-
-lemma (in vectorspace) add_minus_cancel [simp]:
-    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + (- x + y) = y"
-  by (simp add: add_assoc [symmetric] del: add_commute)
-
-lemma (in vectorspace) minus_add_cancel [simp]:
-    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - x + (x + y) = y"
-  by (simp add: add_assoc [symmetric] del: add_commute)
-
-lemma (in vectorspace) minus_add_distrib [simp]:
-    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> - (x + y) = - x + - y"
-  by (simp add: negate_eq1 add_mult_distrib1)
-
-lemma (in vectorspace) diff_zero [simp]:
-    "x \<in> V \<Longrightarrow> x - 0 = x"
-  by (simp add: diff_eq1)
-
-lemma (in vectorspace) diff_zero_right [simp]:
-    "x \<in> V \<Longrightarrow> 0 - x = - x"
-  by (simp add: diff_eq1)
-
-lemma (in vectorspace) add_left_cancel:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + y = x + z) = (y = z)"
-proof
-  assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V"
-  {
-    from y have "y = 0 + y" by simp
-    also from x y have "\<dots> = (- x + x) + y" by simp
-    also from x y have "\<dots> = - x + (x + y)"
-      by (simp add: add_assoc neg_closed)
-    also assume "x + y = x + z"
-    also from x z have "- x + (x + z) = - x + x + z"
-      by (simp add: add_assoc [symmetric] neg_closed)
-    also from x z have "\<dots> = z" by simp
-    finally show "y = z" .
-  next
-    assume "y = z"
-    then show "x + y = x + z" by (simp only:)
-  }
-qed
-
-lemma (in vectorspace) add_right_cancel:
-    "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (y + x = z + x) = (y = z)"
-  by (simp only: add_commute add_left_cancel)
-
-lemma (in vectorspace) add_assoc_cong:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x' \<in> V \<Longrightarrow> y' \<in> V \<Longrightarrow> z \<in> V
-    \<Longrightarrow> x + y = x' + y' \<Longrightarrow> x + (y + z) = x' + (y' + z)"
-  by (simp only: add_assoc [symmetric])
-
-lemma (in vectorspace) mult_left_commute:
-    "x \<in> V \<Longrightarrow> a \<cdot> b \<cdot> x = b \<cdot> a \<cdot> x"
-  by (simp add: real_mult_commute mult_assoc2)
-
-lemma (in vectorspace) mult_zero_uniq:
-  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> a \<cdot> x = 0 \<Longrightarrow> a = 0"
-proof (rule classical)
-  assume a: "a \<noteq> 0"
-  assume x: "x \<in> V"  "x \<noteq> 0" and ax: "a \<cdot> x = 0"
-  from x a have "x = (inverse a * a) \<cdot> x" by simp
-  also from `x \<in> V` have "\<dots> = inverse a \<cdot> (a \<cdot> x)" by (rule mult_assoc)
-  also from ax have "\<dots> = inverse a \<cdot> 0" by simp
-  also have "\<dots> = 0" by simp
-  finally have "x = 0" .
-  with `x \<noteq> 0` show "a = 0" by contradiction
-qed
-
-lemma (in vectorspace) mult_left_cancel:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (a \<cdot> x = a \<cdot> y) = (x = y)"
-proof
-  assume x: "x \<in> V" and y: "y \<in> V" and a: "a \<noteq> 0"
-  from x have "x = 1 \<cdot> x" by simp
-  also from a have "\<dots> = (inverse a * a) \<cdot> x" by simp
-  also from x have "\<dots> = inverse a \<cdot> (a \<cdot> x)"
-    by (simp only: mult_assoc)
-  also assume "a \<cdot> x = a \<cdot> y"
-  also from a y have "inverse a \<cdot> \<dots> = y"
-    by (simp add: mult_assoc2)
-  finally show "x = y" .
-next
-  assume "x = y"
-  then show "a \<cdot> x = a \<cdot> y" by (simp only:)
-qed
-
-lemma (in vectorspace) mult_right_cancel:
-  "x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> (a \<cdot> x = b \<cdot> x) = (a = b)"
-proof
-  assume x: "x \<in> V" and neq: "x \<noteq> 0"
-  {
-    from x have "(a - b) \<cdot> x = a \<cdot> x - b \<cdot> x"
-      by (simp add: diff_mult_distrib2)
-    also assume "a \<cdot> x = b \<cdot> x"
-    with x have "a \<cdot> x - b \<cdot> x = 0" by simp
-    finally have "(a - b) \<cdot> x = 0" .
-    with x neq have "a - b = 0" by (rule mult_zero_uniq)
-    then show "a = b" by simp
-  next
-    assume "a = b"
-    then show "a \<cdot> x = b \<cdot> x" by (simp only:)
-  }
-qed
-
-lemma (in vectorspace) eq_diff_eq:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x = z - y) = (x + y = z)"
-proof
-  assume x: "x \<in> V" and y: "y \<in> V" and z: "z \<in> V"
-  {
-    assume "x = z - y"
-    then have "x + y = z - y + y" by simp
-    also from y z have "\<dots> = z + - y + y"
-      by (simp add: diff_eq1)
-    also have "\<dots> = z + (- y + y)"
-      by (rule add_assoc) (simp_all add: y z)
-    also from y z have "\<dots> = z + 0"
-      by (simp only: add_minus_left)
-    also from z have "\<dots> = z"
-      by (simp only: add_zero_right)
-    finally show "x + y = z" .
-  next
-    assume "x + y = z"
-    then have "z - y = (x + y) - y" by simp
-    also from x y have "\<dots> = x + y + - y"
-      by (simp add: diff_eq1)
-    also have "\<dots> = x + (y + - y)"
-      by (rule add_assoc) (simp_all add: x y)
-    also from x y have "\<dots> = x" by simp
-    finally show "x = z - y" ..
-  }
-qed
-
-lemma (in vectorspace) add_minus_eq_minus:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x + y = 0 \<Longrightarrow> x = - y"
-proof -
-  assume x: "x \<in> V" and y: "y \<in> V"
-  from x y have "x = (- y + y) + x" by simp
-  also from x y have "\<dots> = - y + (x + y)" by (simp add: add_ac)
-  also assume "x + y = 0"
-  also from y have "- y + 0 = - y" by simp
-  finally show "x = - y" .
-qed
-
-lemma (in vectorspace) add_minus_eq:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> x - y = 0 \<Longrightarrow> x = y"
-proof -
-  assume x: "x \<in> V" and y: "y \<in> V"
-  assume "x - y = 0"
-  with x y have eq: "x + - y = 0" by (simp add: diff_eq1)
-  with _ _ have "x = - (- y)"
-    by (rule add_minus_eq_minus) (simp_all add: x y)
-  with x y show "x = y" by simp
-qed
-
-lemma (in vectorspace) add_diff_swap:
-  "a \<in> V \<Longrightarrow> b \<in> V \<Longrightarrow> c \<in> V \<Longrightarrow> d \<in> V \<Longrightarrow> a + b = c + d
-    \<Longrightarrow> a - c = d - b"
-proof -
-  assume vs: "a \<in> V"  "b \<in> V"  "c \<in> V"  "d \<in> V"
-    and eq: "a + b = c + d"
-  then have "- c + (a + b) = - c + (c + d)"
-    by (simp add: add_left_cancel)
-  also have "\<dots> = d" using `c \<in> V` `d \<in> V` by (rule minus_add_cancel)
-  finally have eq: "- c + (a + b) = d" .
-  from vs have "a - c = (- c + (a + b)) + - b"
-    by (simp add: add_ac diff_eq1)
-  also from vs eq have "\<dots>  = d + - b"
-    by (simp add: add_right_cancel)
-  also from vs have "\<dots> = d - b" by (simp add: diff_eq2)
-  finally show "a - c = d - b" .
-qed
-
-lemma (in vectorspace) vs_add_cancel_21:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> u \<in> V
-    \<Longrightarrow> (x + (y + z) = y + u) = (x + z = u)"
-proof
-  assume vs: "x \<in> V"  "y \<in> V"  "z \<in> V"  "u \<in> V"
-  {
-    from vs have "x + z = - y + y + (x + z)" by simp
-    also have "\<dots> = - y + (y + (x + z))"
-      by (rule add_assoc) (simp_all add: vs)
-    also from vs have "y + (x + z) = x + (y + z)"
-      by (simp add: add_ac)
-    also assume "x + (y + z) = y + u"
-    also from vs have "- y + (y + u) = u" by simp
-    finally show "x + z = u" .
-  next
-    assume "x + z = u"
-    with vs show "x + (y + z) = y + u"
-      by (simp only: add_left_commute [of x])
-  }
-qed
-
-lemma (in vectorspace) add_cancel_end:
-  "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> z \<in> V \<Longrightarrow> (x + (y + z) = y) = (x = - z)"
-proof
-  assume vs: "x \<in> V"  "y \<in> V"  "z \<in> V"
-  {
-    assume "x + (y + z) = y"
-    with vs have "(x + z) + y = 0 + y"
-      by (simp add: add_ac)
-    with vs have "x + z = 0"
-      by (simp only: add_right_cancel add_closed zero)
-    with vs show "x = - z" by (simp add: add_minus_eq_minus)
-  next
-    assume eq: "x = - z"
-    then have "x + (y + z) = - z + (y + z)" by simp
-    also have "\<dots> = y + (- z + z)"
-      by (rule add_left_commute) (simp_all add: vs)
-    also from vs have "\<dots> = y"  by simp
-    finally show "x + (y + z) = y" .
-  }
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/ZornLemma.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,58 +0,0 @@
-(*  Title:      HOL/Real/HahnBanach/ZornLemma.thy
-    ID:         $Id$
-    Author:     Gertrud Bauer, TU Munich
-*)
-
-header {* Zorn's Lemma *}
-
-theory ZornLemma
-imports Zorn
-begin
-
-text {*
-  Zorn's Lemmas states: if every linear ordered subset of an ordered
-  set @{text S} has an upper bound in @{text S}, then there exists a
-  maximal element in @{text S}.  In our application, @{text S} is a
-  set of sets ordered by set inclusion. Since the union of a chain of
-  sets is an upper bound for all elements of the chain, the conditions
-  of Zorn's lemma can be modified: if @{text S} is non-empty, it
-  suffices to show that for every non-empty chain @{text c} in @{text
-  S} the union of @{text c} also lies in @{text S}.
-*}
-
-theorem Zorn's_Lemma:
-  assumes r: "\<And>c. c \<in> chain S \<Longrightarrow> \<exists>x. x \<in> c \<Longrightarrow> \<Union>c \<in> S"
-    and aS: "a \<in> S"
-  shows "\<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z \<longrightarrow> y = z"
-proof (rule Zorn_Lemma2)
-  show "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
-  proof
-    fix c assume "c \<in> chain S"
-    show "\<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
-    proof cases
-
-      txt {* If @{text c} is an empty chain, then every element in
-	@{text S} is an upper bound of @{text c}. *}
-
-      assume "c = {}"
-      with aS show ?thesis by fast
-
-      txt {* If @{text c} is non-empty, then @{text "\<Union>c"} is an upper
-	bound of @{text c}, lying in @{text S}. *}
-
-    next
-      assume "c \<noteq> {}"
-      show ?thesis
-      proof
-        show "\<forall>z \<in> c. z \<subseteq> \<Union>c" by fast
-        show "\<Union>c \<in> S"
-        proof (rule r)
-          from `c \<noteq> {}` show "\<exists>x. x \<in> c" by fast
-	  show "c \<in> chain S" by fact
-        qed
-      qed
-    qed
-  qed
-qed
-
-end
--- a/src/HOL/Real/HahnBanach/document/root.bib	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,27 +0,0 @@
-
-@Book{Heuser:1986,
-  author = 	 {H. Heuser},
-  title = 	 {Funktionalanalysis: Theorie und Anwendung},
-  publisher = 	 {Teubner},
-  year = 	 1986
-}
-
-@InCollection{Narici:1996,
-  author = 	 {L. Narici and E. Beckenstein},
-  title = 	 {The {Hahn-Banach Theorem}: The Life and Times},
-  booktitle = 	 {Topology Atlas},
-  publisher =	 {York University, Toronto, Ontario, Canada},
-  year =	 1996,
-  note =	 {\url{http://at.yorku.ca/topology/preprint.htm} and
-                  \url{http://at.yorku.ca/p/a/a/a/16.htm}}
-}
-
-@Article{Nowak:1993,
-  author =       {B. Nowak and A. Trybulec},
-  title =	 {{Hahn-Banach} Theorem},
-  journal =      {Journal of Formalized Mathematics},
-  year =         {1993},
-  volume =       {5},
-  institution =  {University of Bialystok},
-  note =         {\url{http://mizar.uwb.edu.pl/JFM/Vol5/hahnban.html}}
-}
--- a/src/HOL/Real/HahnBanach/document/root.tex	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,83 +0,0 @@
-\documentclass[10pt,a4paper,twoside]{article}
-\usepackage{graphicx}
-\usepackage{latexsym,theorem}
-\usepackage{isabelle,isabellesym}
-\usepackage{pdfsetup} %last one!
-
-\isabellestyle{it}
-\urlstyle{rm}
-
-\newcommand{\isasymsup}{\isamath{\sup\,}}
-\newcommand{\skp}{\smallskip}
-
-
-\begin{document}
-
-\pagestyle{headings}
-\pagenumbering{arabic}
-
-\title{The Hahn-Banach Theorem \\ for Real Vector Spaces}
-\author{Gertrud Bauer \\ \url{http://www.in.tum.de/~bauerg/}}
-\maketitle
-
-\begin{abstract}
-  The Hahn-Banach Theorem is one of the most fundamental results in functional
-  analysis. We present a fully formal proof of two versions of the theorem,
-  one for general linear spaces and another for normed spaces.  This
-  development is based on simply-typed classical set-theory, as provided by
-  Isabelle/HOL.
-\end{abstract}
-
-
-\tableofcontents
-\parindent 0pt \parskip 0.5ex
-
-\clearpage
-\section{Preface}
-
-This is a fully formal proof of the Hahn-Banach Theorem. It closely follows
-the informal presentation given in Heuser's textbook \cite[{\S} 36]{Heuser:1986}.
-Another formal proof of the same theorem has been done in Mizar
-\cite{Nowak:1993}.  A general overview of the relevance and history of the
-Hahn-Banach Theorem is given by Narici and Beckenstein \cite{Narici:1996}.
-
-\medskip The document is structured as follows.  The first part contains
-definitions of basic notions of linear algebra: vector spaces, subspaces,
-normed spaces, continuous linear-forms, norm of functions and an order on
-functions by domain extension.  The second part contains some lemmas about the
-supremum (w.r.t.\ the function order) and extension of non-maximal functions.
-With these preliminaries, the main proof of the theorem (in its two versions)
-is conducted in the third part.  The dependencies of individual theories are
-as follows.
-
-\begin{center}
-  \includegraphics[scale=0.5]{session_graph}  
-\end{center}
-
-\clearpage
-\part {Basic Notions}
-
-\input{Bounds}
-\input{VectorSpace}
-\input{Subspace}
-\input{NormedSpace}
-\input{Linearform}
-\input{FunctionOrder}
-\input{FunctionNorm}
-\input{ZornLemma}
-
-\clearpage
-\part {Lemmas for the Proof}
-
-\input{HahnBanachSupLemmas}
-\input{HahnBanachExtLemmas}
-\input{HahnBanachLemmas}
-
-\clearpage
-\part {The Main Proof}
-
-\input{HahnBanach}
-\bibliographystyle{abbrv}
-\bibliography{root}
-
-\end{document}
--- a/src/HOL/Real/RealVector.thy	Mon Dec 29 13:23:53 2008 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,841 +0,0 @@
-(*  Title:      HOL/Real/RealVector.thy
-    Author:     Brian Huffman
-*)
-
-header {* Vector Spaces and Algebras over the Reals *}
-
-theory RealVector
-imports "~~/src/HOL/RealPow"
-begin
-
-subsection {* Locale for additive functions *}
-
-locale additive =
-  fixes f :: "'a::ab_group_add \<Rightarrow> 'b::ab_group_add"
-  assumes add: "f (x + y) = f x + f y"
-begin
-
-lemma zero: "f 0 = 0"
-proof -
-  have "f 0 = f (0 + 0)" by simp
-  also have "\<dots> = f 0 + f 0" by (rule add)
-  finally show "f 0 = 0" by simp
-qed
-
-lemma minus: "f (- x) = - f x"
-proof -
-  have "f (- x) + f x = f (- x + x)" by (rule add [symmetric])
-  also have "\<dots> = - f x + f x" by (simp add: zero)
-  finally show "f (- x) = - f x" by (rule add_right_imp_eq)
-qed
-
-lemma diff: "f (x - y) = f x - f y"
-by (simp add: diff_def add minus)
-
-lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
-apply (cases "finite A")
-apply (induct set: finite)
-apply (simp add: zero)
-apply (simp add: add)
-apply (simp add: zero)
-done
-
-end
-
-subsection {* Vector spaces *}
-
-locale vector_space =
-  fixes scale :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b"
-  assumes scale_right_distrib: "scale a (x + y) = scale a x + scale a y"
-  and scale_left_distrib: "scale (a + b) x = scale a x + scale b x"
-  and scale_scale [simp]: "scale a (scale b x) = scale (a * b) x"
-  and scale_one [simp]: "scale 1 x = x"
-begin
-
-lemma scale_left_commute:
-  "scale a (scale b x) = scale b (scale a x)"
-by (simp add: mult_commute)
-
-lemma scale_zero_left [simp]: "scale 0 x = 0"
-  and scale_minus_left [simp]: "scale (- a) x = - (scale a x)"
-  and scale_left_diff_distrib: "scale (a - b) x = scale a x - scale b x"
-proof -
-  interpret s: additive ["\<lambda>a. scale a x"]
-    proof qed (rule scale_left_distrib)
-  show "scale 0 x = 0" by (rule s.zero)
-  show "scale (- a) x = - (scale a x)" by (rule s.minus)
-  show "scale (a - b) x = scale a x - scale b x" by (rule s.diff)
-qed
-
-lemma scale_zero_right [simp]: "scale a 0 = 0"
-  and scale_minus_right [simp]: "scale a (- x) = - (scale a x)"
-  and scale_right_diff_distrib: "scale a (x - y) = scale a x - scale a y"
-proof -
-  interpret s: additive ["\<lambda>x. scale a x"]
-    proof qed (rule scale_right_distrib)
-  show "scale a 0 = 0" by (rule s.zero)
-  show "scale a (- x) = - (scale a x)" by (rule s.minus)
-  show "scale a (x - y) = scale a x - scale a y" by (rule s.diff)
-qed
-
-lemma scale_eq_0_iff [simp]:
-  "scale a x = 0 \<longleftrightarrow> a = 0 \<or> x = 0"
-proof cases
-  assume "a = 0" thus ?thesis by simp
-next
-  assume anz [simp]: "a \<noteq> 0"
-  { assume "scale a x = 0"
-    hence "scale (inverse a) (scale a x) = 0" by simp
-    hence "x = 0" by simp }
-  thus ?thesis by force
-qed
-
-lemma scale_left_imp_eq:
-  "\<lbrakk>a \<noteq> 0; scale a x = scale a y\<rbrakk> \<Longrightarrow> x = y"
-proof -
-  assume nonzero: "a \<noteq> 0"
-  assume "scale a x = scale a y"
-  hence "scale a (x - y) = 0"
-     by (simp add: scale_right_diff_distrib)
-  hence "x - y = 0" by (simp add: nonzero)
-  thus "x = y" by (simp only: right_minus_eq)
-qed
-
-lemma scale_right_imp_eq:
-  "\<lbrakk>x \<noteq> 0; scale a x = scale b x\<rbrakk> \<Longrightarrow> a = b"
-proof -
-  assume nonzero: "x \<noteq> 0"
-  assume "scale a x = scale b x"
-  hence "scale (a - b) x = 0"
-     by (simp add: scale_left_diff_distrib)
-  hence "a - b = 0" by (simp add: nonzero)
-  thus "a = b" by (simp only: right_minus_eq)
-qed
-
-lemma scale_cancel_left:
-  "scale a x = scale a y \<longleftrightarrow> x = y \<or> a = 0"
-by (auto intro: scale_left_imp_eq)
-
-lemma scale_cancel_right:
-  "scale a x = scale b x \<longleftrightarrow> a = b \<or> x = 0"
-by (auto intro: scale_right_imp_eq)
-
-end
-
-subsection {* Real vector spaces *}
-
-class scaleR = type +
-  fixes scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "*\<^sub>R" 75)
-begin
-
-abbreviation
-  divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a" (infixl "'/\<^sub>R" 70)
-where
-  "x /\<^sub>R r == scaleR (inverse r) x"
-
-end
-
-instantiation real :: scaleR
-begin
-
-definition
-  real_scaleR_def [simp]: "scaleR a x = a * x"
-
-instance ..
-
-end
-
-class real_vector = scaleR + ab_group_add +
-  assumes scaleR_right_distrib: "scaleR a (x + y) = scaleR a x + scaleR a y"
-  and scaleR_left_distrib: "scaleR (a + b) x = scaleR a x + scaleR b x"
-  and scaleR_scaleR [simp]: "scaleR a (scaleR b x) = scaleR (a * b) x"
-  and scaleR_one [simp]: "scaleR 1 x = x"
-
-interpretation real_vector:
-  vector_space ["scaleR :: real \<Rightarrow> 'a \<Rightarrow> 'a::real_vector"]
-apply unfold_locales
-apply (rule scaleR_right_distrib)
-apply (rule scaleR_left_distrib)
-apply (rule scaleR_scaleR)
-apply (rule scaleR_one)
-done
-
-text {* Recover original theorem names *}
-
-lemmas scaleR_left_commute = real_vector.scale_left_commute
-lemmas scaleR_zero_left = real_vector.scale_zero_left
-lemmas scaleR_minus_left = real_vector.scale_minus_left
-lemmas scaleR_left_diff_distrib = real_vector.scale_left_diff_distrib
-lemmas scaleR_zero_right = real_vector.scale_zero_right
-lemmas scaleR_minus_right = real_vector.scale_minus_right
-lemmas scaleR_right_diff_distrib = real_vector.scale_right_diff_distrib
-lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff
-lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq
-lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq
-lemmas scaleR_cancel_left = real_vector.scale_cancel_left
-lemmas scaleR_cancel_right = real_vector.scale_cancel_right
-
-class real_algebra = real_vector + ring +
-  assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)"
-  and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"
-
-class real_algebra_1 = real_algebra + ring_1
-
-class real_div_algebra = real_algebra_1 + division_ring
-
-class real_field = real_div_algebra + field
-
-instance real :: real_field
-apply (intro_classes, unfold real_scaleR_def)
-apply (rule right_distrib)
-apply (rule left_distrib)
-apply (rule mult_assoc [symmetric])
-apply (rule mult_1_left)
-apply (rule mult_assoc)
-apply (rule mult_left_commute)
-done
-
-interpretation scaleR_left: additive ["(\<lambda>a. scaleR a x::'a::real_vector)"]
-proof qed (rule scaleR_left_distrib)
-
-interpretation scaleR_right: additive ["(\<lambda>x. scaleR a x::'a::real_vector)"]
-proof qed (rule scaleR_right_distrib)
-
-lemma nonzero_inverse_scaleR_distrib:
-  fixes x :: "'a::real_div_algebra" shows
-  "\<lbrakk>a \<noteq> 0; x \<noteq> 0\<rbrakk> \<Longrightarrow> inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
-by (rule inverse_unique, simp)
-
-lemma inverse_scaleR_distrib:
-  fixes x :: "'a::{real_div_algebra,division_by_zero}"
-  shows "inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
-apply (case_tac "a = 0", simp)
-apply (case_tac "x = 0", simp)
-apply (erule (1) nonzero_inverse_scaleR_distrib)
-done
-
-
-subsection {* Embedding of the Reals into any @{text real_algebra_1}:
-@{term of_real} *}
-
-definition
-  of_real :: "real \<Rightarrow> 'a::real_algebra_1" where
-  "of_real r = scaleR r 1"
-
-lemma scaleR_conv_of_real: "scaleR r x = of_real r * x"
-by (simp add: of_real_def)
-
-lemma of_real_0 [simp]: "of_real 0 = 0"
-by (simp add: of_real_def)
-
-lemma of_real_1 [simp]: "of_real 1 = 1"
-by (simp add: of_real_def)
-
-lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y"
-by (simp add: of_real_def scaleR_left_distrib)
-
-lemma of_real_minus [simp]: "of_real (- x) = - of_real x"
-by (simp add: of_real_def)
-
-lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y"
-by (simp add: of_real_def scaleR_left_diff_distrib)
-
-lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y"
-by (simp add: of_real_def mult_commute)
-
-lemma nonzero_of_real_inverse:
-  "x \<noteq> 0 \<Longrightarrow> of_real (inverse x) =
-   inverse (of_real x :: 'a::real_div_algebra)"
-by (simp add: of_real_def nonzero_inverse_scaleR_distrib)
-
-lemma of_real_inverse [simp]:
-  "of_real (inverse x) =
-   inverse (of_real x :: 'a::{real_div_algebra,division_by_zero})"
-by (simp add: of_real_def inverse_scaleR_distrib)
-
-lemma nonzero_of_real_divide:
-  "y \<noteq> 0 \<Longrightarrow> of_real (x / y) =
-   (of_real x / of_real y :: 'a::real_field)"
-by (simp add: divide_inverse nonzero_of_real_inverse)
-
-lemma of_real_divide [simp]:
-  "of_real (x / y) =
-   (of_real x / of_real y :: 'a::{real_field,division_by_zero})"
-by (simp add: divide_inverse)
-
-lemma of_real_power [simp]:
-  "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1,recpower}) ^ n"
-by (induct n) (simp_all add: power_Suc)
-
-lemma of_real_eq_iff [simp]: "(of_real x = of_real y) = (x = y)"
-by (simp add: of_real_def scaleR_cancel_right)
-
-lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified]
-
-lemma of_real_eq_id [simp]: "of_real = (id :: real \<Rightarrow> real)"
-proof
-  fix r
-  show "of_real r = id r"
-    by (simp add: of_real_def)
-qed
-
-text{*Collapse nested embeddings*}
-lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n"
-by (induct n) auto
-
-lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z"
-by (cases z rule: int_diff_cases, simp)
-
-lemma of_real_number_of_eq:
-  "of_real (number_of w) = (number_of w :: 'a::{number_ring,real_algebra_1})"
-by (simp add: number_of_eq)
-
-text{*Every real algebra has characteristic zero*}
-instance real_algebra_1 < ring_char_0
-proof
-  fix m n :: nat
-  have "(of_real (of_nat m) = (of_real (of_nat n)::'a)) = (m = n)"
-    by (simp only: of_real_eq_iff of_nat_eq_iff)
-  thus "(of_nat m = (of_nat n::'a)) = (m = n)"
-    by (simp only: of_real_of_nat_eq)
-qed
-
-instance real_field < field_char_0 ..
-
-
-subsection {* The Set of Real Numbers *}
-
-definition
-  Reals :: "'a::real_algebra_1 set" where
-  [code del]: "Reals \<equiv> range of_real"
-
-notation (xsymbols)
-  Reals  ("\<real>")
-
-lemma Reals_of_real [simp]: "of_real r \<in> Reals"
-by (simp add: Reals_def)
-
-lemma Reals_of_int [simp]: "of_int z \<in> Reals"
-by (subst of_real_of_int_eq [symmetric], rule Reals_of_real)
-
-lemma Reals_of_nat [simp]: "of_nat n \<in> Reals"
-by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real)
-
-lemma Reals_number_of [simp]:
-  "(number_of w::'a::{number_ring,real_algebra_1}) \<in> Reals"
-by (subst of_real_number_of_eq [symmetric], rule Reals_of_real)
-
-lemma Reals_0 [simp]: "0 \<in> Reals"
-apply (unfold Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_0 [symmetric])
-done
-
-lemma Reals_1 [simp]: "1 \<in> Reals"
-apply (unfold Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_1 [symmetric])
-done
-
-lemma Reals_add [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a + b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_add [symmetric])
-done
-
-lemma Reals_minus [simp]: "a \<in> Reals \<Longrightarrow> - a \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_minus [symmetric])
-done
-
-lemma Reals_diff [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a - b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_diff [symmetric])
-done
-
-lemma Reals_mult [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a * b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_mult [symmetric])
-done
-
-lemma nonzero_Reals_inverse:
-  fixes a :: "'a::real_div_algebra"
-  shows "\<lbrakk>a \<in> Reals; a \<noteq> 0\<rbrakk> \<Longrightarrow> inverse a \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (erule nonzero_of_real_inverse [symmetric])
-done
-
-lemma Reals_inverse [simp]:
-  fixes a :: "'a::{real_div_algebra,division_by_zero}"
-  shows "a \<in> Reals \<Longrightarrow> inverse a \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_inverse [symmetric])
-done
-
-lemma nonzero_Reals_divide:
-  fixes a b :: "'a::real_field"
-  shows "\<lbrakk>a \<in> Reals; b \<in> Reals; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (erule nonzero_of_real_divide [symmetric])
-done
-
-lemma Reals_divide [simp]:
-  fixes a b :: "'a::{real_field,division_by_zero}"
-  shows "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_divide [symmetric])
-done
-
-lemma Reals_power [simp]:
-  fixes a :: "'a::{real_algebra_1,recpower}"
-  shows "a \<in> Reals \<Longrightarrow> a ^ n \<in> Reals"
-apply (auto simp add: Reals_def)
-apply (rule range_eqI)
-apply (rule of_real_power [symmetric])
-done
-
-lemma Reals_cases [cases set: Reals]:
-  assumes "q \<in> \<real>"
-  obtains (of_real) r where "q = of_real r"
-  unfolding Reals_def
-proof -
-  from `q \<in> \<real>` have "q \<in> range of_real" unfolding Reals_def .
-  then obtain r where "q = of_real r" ..
-  then show thesis ..
-qed
-
-lemma Reals_induct [case_names of_real, induct set: Reals]:
-  "q \<in> \<real> \<Longrightarrow> (\<And>r. P (of_real r)) \<Longrightarrow> P q"
-  by (rule Reals_cases) auto
-
-
-subsection {* Real normed vector spaces *}
-
-class norm = type +
-  fixes norm :: "'a \<Rightarrow> real"
-
-instantiation real :: norm
-begin
-
-definition
-  real_norm_def [simp]: "norm r \<equiv> \<bar>r\<bar>"
-
-instance ..
-
-end
-
-class sgn_div_norm = scaleR + norm + sgn +
-  assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x"
-
-class real_normed_vector = real_vector + sgn_div_norm +
-  assumes norm_ge_zero [simp]: "0 \<le> norm x"
-  and norm_eq_zero [simp]: "norm x = 0 \<longleftrightarrow> x = 0"
-  and norm_triangle_ineq: "norm (x + y) \<le> norm x + norm y"
-  and norm_scaleR: "norm (scaleR a x) = \<bar>a\<bar> * norm x"
-
-class real_normed_algebra = real_algebra + real_normed_vector +
-  assumes norm_mult_ineq: "norm (x * y) \<le> norm x * norm y"
-
-class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra +
-  assumes norm_one [simp]: "norm 1 = 1"
-
-class real_normed_div_algebra = real_div_algebra + real_normed_vector +
-  assumes norm_mult: "norm (x * y) = norm x * norm y"
-
-class real_normed_field = real_field + real_normed_div_algebra
-
-instance real_normed_div_algebra < real_normed_algebra_1
-proof
-  fix x y :: 'a
-  show "norm (x * y) \<le> norm x * norm y"
-    by (simp add: norm_mult)
-next
-  have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)"
-    by (rule norm_mult)
-  thus "norm (1::'a) = 1" by simp
-qed
-
-instance real :: real_normed_field
-apply (intro_classes, unfold real_norm_def real_scaleR_def)
-apply (simp add: real_sgn_def)
-apply (rule abs_ge_zero)
-apply (rule abs_eq_0)
-apply (rule abs_triangle_ineq)
-apply (rule abs_mult)
-apply (rule abs_mult)
-done
-
-lemma norm_zero [simp]: "norm (0::'a::real_normed_vector) = 0"
-by simp
-
-lemma zero_less_norm_iff [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "(0 < norm x) = (x \<noteq> 0)"
-by (simp add: order_less_le)
-
-lemma norm_not_less_zero [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "\<not> norm x < 0"
-by (simp add: linorder_not_less)
-
-lemma norm_le_zero_iff [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "(norm x \<le> 0) = (x = 0)"
-by (simp add: order_le_less)
-
-lemma norm_minus_cancel [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "norm (- x) = norm x"
-proof -
-  have "norm (- x) = norm (scaleR (- 1) x)"
-    by (simp only: scaleR_minus_left scaleR_one)
-  also have "\<dots> = \<bar>- 1\<bar> * norm x"
-    by (rule norm_scaleR)
-  finally show ?thesis by simp
-qed
-
-lemma norm_minus_commute:
-  fixes a b :: "'a::real_normed_vector"
-  shows "norm (a - b) = norm (b - a)"
-proof -
-  have "norm (- (b - a)) = norm (b - a)"
-    by (rule norm_minus_cancel)
-  thus ?thesis by simp
-qed
-
-lemma norm_triangle_ineq2:
-  fixes a b :: "'a::real_normed_vector"
-  shows "norm a - norm b \<le> norm (a - b)"
-proof -
-  have "norm (a - b + b) \<le> norm (a - b) + norm b"
-    by (rule norm_triangle_ineq)
-  thus ?thesis by simp
-qed
-
-lemma norm_triangle_ineq3:
-  fixes a b :: "'a::real_normed_vector"
-  shows "\<bar>norm a - norm b\<bar> \<le> norm (a - b)"
-apply (subst abs_le_iff)
-apply auto
-apply (rule norm_triangle_ineq2)
-apply (subst norm_minus_commute)
-apply (rule norm_triangle_ineq2)
-done
-
-lemma norm_triangle_ineq4:
-  fixes a b :: "'a::real_normed_vector"
-  shows "norm (a - b) \<le> norm a + norm b"
-proof -
-  have "norm (a + - b) \<le> norm a + norm (- b)"
-    by (rule norm_triangle_ineq)
-  thus ?thesis
-    by (simp only: diff_minus norm_minus_cancel)
-qed
-
-lemma norm_diff_ineq:
-  fixes a b :: "'a::real_normed_vector"
-  shows "norm a - norm b \<le> norm (a + b)"
-proof -
-  have "norm a - norm (- b) \<le> norm (a - - b)"
-    by (rule norm_triangle_ineq2)
-  thus ?thesis by simp
-qed
-
-lemma norm_diff_triangle_ineq:
-  fixes a b c d :: "'a::real_normed_vector"
-  shows "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
-proof -
-  have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))"
-    by (simp add: diff_minus add_ac)
-  also have "\<dots> \<le> norm (a - c) + norm (b - d)"
-    by (rule norm_triangle_ineq)
-  finally show ?thesis .
-qed
-
-lemma abs_norm_cancel [simp]:
-  fixes a :: "'a::real_normed_vector"
-  shows "\<bar>norm a\<bar> = norm a"
-by (rule abs_of_nonneg [OF norm_ge_zero])
-
-lemma norm_add_less:
-  fixes x y :: "'a::real_normed_vector"
-  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x + y) < r + s"
-by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono])
-
-lemma norm_mult_less:
-  fixes x y :: "'a::real_normed_algebra"
-  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x * y) < r * s"
-apply (rule order_le_less_trans [OF norm_mult_ineq])
-apply (simp add: mult_strict_mono')
-done
-
-lemma norm_of_real [simp]:
-  "norm (of_real r :: 'a::real_normed_algebra_1) = \<bar>r\<bar>"
-unfolding of_real_def by (simp add: norm_scaleR)
-
-lemma norm_number_of [simp]:
-  "norm (number_of w::'a::{number_ring,real_normed_algebra_1})
-    = \<bar>number_of w\<bar>"
-by (subst of_real_number_of_eq [symmetric], rule norm_of_real)
-
-lemma norm_of_int [simp]:
-  "norm (of_int z::'a::real_normed_algebra_1) = \<bar>of_int z\<bar>"
-by (subst of_real_of_int_eq [symmetric], rule norm_of_real)
-
-lemma norm_of_nat [simp]:
-  "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n"
-apply (subst of_real_of_nat_eq [symmetric])
-apply (subst norm_of_real, simp)
-done
-
-lemma nonzero_norm_inverse:
-  fixes a :: "'a::real_normed_div_algebra"
-  shows "a \<noteq> 0 \<Longrightarrow> norm (inverse a) = inverse (norm a)"
-apply (rule inverse_unique [symmetric])
-apply (simp add: norm_mult [symmetric])
-done
-
-lemma norm_inverse:
-  fixes a :: "'a::{real_normed_div_algebra,division_by_zero}"
-  shows "norm (inverse a) = inverse (norm a)"
-apply (case_tac "a = 0", simp)
-apply (erule nonzero_norm_inverse)
-done
-
-lemma nonzero_norm_divide:
-  fixes a b :: "'a::real_normed_field"
-  shows "b \<noteq> 0 \<Longrightarrow> norm (a / b) = norm a / norm b"
-by (simp add: divide_inverse norm_mult nonzero_norm_inverse)
-
-lemma norm_divide:
-  fixes a b :: "'a::{real_normed_field,division_by_zero}"
-  shows "norm (a / b) = norm a / norm b"
-by (simp add: divide_inverse norm_mult norm_inverse)
-
-lemma norm_power_ineq:
-  fixes x :: "'a::{real_normed_algebra_1,recpower}"
-  shows "norm (x ^ n) \<le> norm x ^ n"
-proof (induct n)
-  case 0 show "norm (x ^ 0) \<le> norm x ^ 0" by simp
-next
-  case (Suc n)
-  have "norm (x * x ^ n) \<le> norm x * norm (x ^ n)"
-    by (rule norm_mult_ineq)
-  also from Suc have "\<dots> \<le> norm x * norm x ^ n"
-    using norm_ge_zero by (rule mult_left_mono)
-  finally show "norm (x ^ Suc n) \<le> norm x ^ Suc n"
-    by (simp add: power_Suc)
-qed
-
-lemma norm_power:
-  fixes x :: "'a::{real_normed_div_algebra,recpower}"
-  shows "norm (x ^ n) = norm x ^ n"
-by (induct n) (simp_all add: power_Suc norm_mult)
-
-
-subsection {* Sign function *}
-
-lemma norm_sgn:
-  "norm (sgn(x::'a::real_normed_vector)) = (if x = 0 then 0 else 1)"
-by (simp add: sgn_div_norm norm_scaleR)
-
-lemma sgn_zero [simp]: "sgn(0::'a::real_normed_vector) = 0"
-by (simp add: sgn_div_norm)
-
-lemma sgn_zero_iff: "(sgn(x::'a::real_normed_vector) = 0) = (x = 0)"
-by (simp add: sgn_div_norm)
-
-lemma sgn_minus: "sgn (- x) = - sgn(x::'a::real_normed_vector)"
-by (simp add: sgn_div_norm)
-
-lemma sgn_scaleR:
-  "sgn (scaleR r x) = scaleR (sgn r) (sgn(x::'a::real_normed_vector))"
-by (simp add: sgn_div_norm norm_scaleR mult_ac)
-
-lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1"
-by (simp add: sgn_div_norm)
-
-lemma sgn_of_real:
-  "sgn (of_real r::'a::real_normed_algebra_1) = of_real (sgn r)"
-unfolding of_real_def by (simp only: sgn_scaleR sgn_one)
-
-lemma sgn_mult:
-  fixes x y :: "'a::real_normed_div_algebra"
-  shows "sgn (x * y) = sgn x * sgn y"
-by (simp add: sgn_div_norm norm_mult mult_commute)
-
-lemma real_sgn_eq: "sgn (x::real) = x / \<bar>x\<bar>"
-by (simp add: sgn_div_norm divide_inverse)
-
-lemma real_sgn_pos: "0 < (x::real) \<Longrightarrow> sgn x = 1"
-unfolding real_sgn_eq by simp
-
-lemma real_sgn_neg: "(x::real) < 0 \<Longrightarrow> sgn x = -1"
-unfolding real_sgn_eq by simp
-
-
-subsection {* Bounded Linear and Bilinear Operators *}
-
-locale bounded_linear = additive +
-  constrains f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
-  assumes scaleR: "f (scaleR r x) = scaleR r (f x)"
-  assumes bounded: "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
-begin
-
-lemma pos_bounded:
-  "\<exists>K>0. \<forall>x. norm (f x) \<le> norm x * K"
-proof -
-  obtain K where K: "\<And>x. norm (f x) \<le> norm x * K"
-    using bounded by fast
-  show ?thesis
-  proof (intro exI impI conjI allI)
-    show "0 < max 1 K"
-      by (rule order_less_le_trans [OF zero_less_one le_maxI1])
-  next
-    fix x
-    have "norm (f x) \<le> norm x * K" using K .
-    also have "\<dots> \<le> norm x * max 1 K"
-      by (rule mult_left_mono [OF le_maxI2 norm_ge_zero])
-    finally show "norm (f x) \<le> norm x * max 1 K" .
-  qed
-qed
-
-lemma nonneg_bounded:
-  "\<exists>K\<ge>0. \<forall>x. norm (f x) \<le> norm x * K"
-proof -
-  from pos_bounded
-  show ?thesis by (auto intro: order_less_imp_le)
-qed
-
-end
-
-locale bounded_bilinear =
-  fixes prod :: "['a::real_normed_vector, 'b::real_normed_vector]
-                 \<Rightarrow> 'c::real_normed_vector"
-    (infixl "**" 70)
-  assumes add_left: "prod (a + a') b = prod a b + prod a' b"
-  assumes add_right: "prod a (b + b') = prod a b + prod a b'"
-  assumes scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)"
-  assumes scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)"
-  assumes bounded: "\<exists>K. \<forall>a b. norm (prod a b) \<le> norm a * norm b * K"
-begin
-
-lemma pos_bounded:
-  "\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
-apply (cut_tac bounded, erule exE)
-apply (rule_tac x="max 1 K" in exI, safe)
-apply (rule order_less_le_trans [OF zero_less_one le_maxI1])
-apply (drule spec, drule spec, erule order_trans)
-apply (rule mult_left_mono [OF le_maxI2])
-apply (intro mult_nonneg_nonneg norm_ge_zero)
-done
-
-lemma nonneg_bounded:
-  "\<exists>K\<ge>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
-proof -
-  from pos_bounded
-  show ?thesis by (auto intro: order_less_imp_le)
-qed
-
-lemma additive_right: "additive (\<lambda>b. prod a b)"
-by (rule additive.intro, rule add_right)
-
-lemma additive_left: "additive (\<lambda>a. prod a b)"
-by (rule additive.intro, rule add_left)
-
-lemma zero_left: "prod 0 b = 0"
-by (rule additive.zero [OF additive_left])
-
-lemma zero_right: "prod a 0 = 0"
-by (rule additive.zero [OF additive_right])
-
-lemma minus_left: "prod (- a) b = - prod a b"
-by (rule additive.minus [OF additive_left])
-
-lemma minus_right: "prod a (- b) = - prod a b"
-by (rule additive.minus [OF additive_right])
-
-lemma diff_left:
-  "prod (a - a') b = prod a b - prod a' b"
-by (rule additive.diff [OF additive_left])
-
-lemma diff_right:
-  "prod a (b - b') = prod a b - prod a b'"
-by (rule additive.diff [OF additive_right])
-
-lemma bounded_linear_left:
-  "bounded_linear (\<lambda>a. a ** b)"
-apply (unfold_locales)
-apply (rule add_left)
-apply (rule scaleR_left)
-apply (cut_tac bounded, safe)
-apply (rule_tac x="norm b * K" in exI)
-apply (simp add: mult_ac)
-done
-
-lemma bounded_linear_right:
-  "bounded_linear (\<lambda>b. a ** b)"
-apply (unfold_locales)
-apply (rule add_right)
-apply (rule scaleR_right)
-apply (cut_tac bounded, safe)
-apply (rule_tac x="norm a * K" in exI)
-apply (simp add: mult_ac)
-done
-
-lemma prod_diff_prod:
-  "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)"
-by (simp add: diff_left diff_right)
-
-end
-
-interpretation mult:
-  bounded_bilinear ["op * :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra"]
-apply (rule bounded_bilinear.intro)
-apply (rule left_distrib)
-apply (rule right_distrib)
-apply (rule mult_scaleR_left)
-apply (rule mult_scaleR_right)
-apply (rule_tac x="1" in exI)
-apply (simp add: norm_mult_ineq)
-done
-
-interpretation mult_left:
-  bounded_linear ["(\<lambda>x::'a::real_normed_algebra. x * y)"]
-by (rule mult.bounded_linear_left)
-
-interpretation mult_right:
-  bounded_linear ["(\<lambda>y::'a::real_normed_algebra. x * y)"]
-by (rule mult.bounded_linear_right)
-
-interpretation divide:
-  bounded_linear ["(\<lambda>x::'a::real_normed_field. x / y)"]
-unfolding divide_inverse by (rule mult.bounded_linear_left)
-
-interpretation scaleR: bounded_bilinear ["scaleR"]
-apply (rule bounded_bilinear.intro)
-apply (rule scaleR_left_distrib)
-apply (rule scaleR_right_distrib)
-apply simp
-apply (rule scaleR_left_commute)
-apply (rule_tac x="1" in exI)
-apply (simp add: norm_scaleR)
-done
-
-interpretation scaleR_left: bounded_linear ["\<lambda>r. scaleR r x"]
-by (rule scaleR.bounded_linear_left)
-
-interpretation scaleR_right: bounded_linear ["\<lambda>x. scaleR r x"]
-by (rule scaleR.bounded_linear_right)
-
-interpretation of_real: bounded_linear ["\<lambda>r. of_real r"]
-unfolding of_real_def by (rule scaleR.bounded_linear_left)
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/RealVector.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,841 @@
+(*  Title:      HOL/RealVector.thy
+    Author:     Brian Huffman
+*)
+
+header {* Vector Spaces and Algebras over the Reals *}
+
+theory RealVector
+imports RealPow
+begin
+
+subsection {* Locale for additive functions *}
+
+locale additive =
+  fixes f :: "'a::ab_group_add \<Rightarrow> 'b::ab_group_add"
+  assumes add: "f (x + y) = f x + f y"
+begin
+
+lemma zero: "f 0 = 0"
+proof -
+  have "f 0 = f (0 + 0)" by simp
+  also have "\<dots> = f 0 + f 0" by (rule add)
+  finally show "f 0 = 0" by simp
+qed
+
+lemma minus: "f (- x) = - f x"
+proof -
+  have "f (- x) + f x = f (- x + x)" by (rule add [symmetric])
+  also have "\<dots> = - f x + f x" by (simp add: zero)
+  finally show "f (- x) = - f x" by (rule add_right_imp_eq)
+qed
+
+lemma diff: "f (x - y) = f x - f y"
+by (simp add: diff_def add minus)
+
+lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
+apply (cases "finite A")
+apply (induct set: finite)
+apply (simp add: zero)
+apply (simp add: add)
+apply (simp add: zero)
+done
+
+end
+
+subsection {* Vector spaces *}
+
+locale vector_space =
+  fixes scale :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b"
+  assumes scale_right_distrib: "scale a (x + y) = scale a x + scale a y"
+  and scale_left_distrib: "scale (a + b) x = scale a x + scale b x"
+  and scale_scale [simp]: "scale a (scale b x) = scale (a * b) x"
+  and scale_one [simp]: "scale 1 x = x"
+begin
+
+lemma scale_left_commute:
+  "scale a (scale b x) = scale b (scale a x)"
+by (simp add: mult_commute)
+
+lemma scale_zero_left [simp]: "scale 0 x = 0"
+  and scale_minus_left [simp]: "scale (- a) x = - (scale a x)"
+  and scale_left_diff_distrib: "scale (a - b) x = scale a x - scale b x"
+proof -
+  interpret s: additive ["\<lambda>a. scale a x"]
+    proof qed (rule scale_left_distrib)
+  show "scale 0 x = 0" by (rule s.zero)
+  show "scale (- a) x = - (scale a x)" by (rule s.minus)
+  show "scale (a - b) x = scale a x - scale b x" by (rule s.diff)
+qed
+
+lemma scale_zero_right [simp]: "scale a 0 = 0"
+  and scale_minus_right [simp]: "scale a (- x) = - (scale a x)"
+  and scale_right_diff_distrib: "scale a (x - y) = scale a x - scale a y"
+proof -
+  interpret s: additive ["\<lambda>x. scale a x"]
+    proof qed (rule scale_right_distrib)
+  show "scale a 0 = 0" by (rule s.zero)
+  show "scale a (- x) = - (scale a x)" by (rule s.minus)
+  show "scale a (x - y) = scale a x - scale a y" by (rule s.diff)
+qed
+
+lemma scale_eq_0_iff [simp]:
+  "scale a x = 0 \<longleftrightarrow> a = 0 \<or> x = 0"
+proof cases
+  assume "a = 0" thus ?thesis by simp
+next
+  assume anz [simp]: "a \<noteq> 0"
+  { assume "scale a x = 0"
+    hence "scale (inverse a) (scale a x) = 0" by simp
+    hence "x = 0" by simp }
+  thus ?thesis by force
+qed
+
+lemma scale_left_imp_eq:
+  "\<lbrakk>a \<noteq> 0; scale a x = scale a y\<rbrakk> \<Longrightarrow> x = y"
+proof -
+  assume nonzero: "a \<noteq> 0"
+  assume "scale a x = scale a y"
+  hence "scale a (x - y) = 0"
+     by (simp add: scale_right_diff_distrib)
+  hence "x - y = 0" by (simp add: nonzero)
+  thus "x = y" by (simp only: right_minus_eq)
+qed
+
+lemma scale_right_imp_eq:
+  "\<lbrakk>x \<noteq> 0; scale a x = scale b x\<rbrakk> \<Longrightarrow> a = b"
+proof -
+  assume nonzero: "x \<noteq> 0"
+  assume "scale a x = scale b x"
+  hence "scale (a - b) x = 0"
+     by (simp add: scale_left_diff_distrib)
+  hence "a - b = 0" by (simp add: nonzero)
+  thus "a = b" by (simp only: right_minus_eq)
+qed
+
+lemma scale_cancel_left:
+  "scale a x = scale a y \<longleftrightarrow> x = y \<or> a = 0"
+by (auto intro: scale_left_imp_eq)
+
+lemma scale_cancel_right:
+  "scale a x = scale b x \<longleftrightarrow> a = b \<or> x = 0"
+by (auto intro: scale_right_imp_eq)
+
+end
+
+subsection {* Real vector spaces *}
+
+class scaleR = type +
+  fixes scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "*\<^sub>R" 75)
+begin
+
+abbreviation
+  divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a" (infixl "'/\<^sub>R" 70)
+where
+  "x /\<^sub>R r == scaleR (inverse r) x"
+
+end
+
+instantiation real :: scaleR
+begin
+
+definition
+  real_scaleR_def [simp]: "scaleR a x = a * x"
+
+instance ..
+
+end
+
+class real_vector = scaleR + ab_group_add +
+  assumes scaleR_right_distrib: "scaleR a (x + y) = scaleR a x + scaleR a y"
+  and scaleR_left_distrib: "scaleR (a + b) x = scaleR a x + scaleR b x"
+  and scaleR_scaleR [simp]: "scaleR a (scaleR b x) = scaleR (a * b) x"
+  and scaleR_one [simp]: "scaleR 1 x = x"
+
+interpretation real_vector:
+  vector_space ["scaleR :: real \<Rightarrow> 'a \<Rightarrow> 'a::real_vector"]
+apply unfold_locales
+apply (rule scaleR_right_distrib)
+apply (rule scaleR_left_distrib)
+apply (rule scaleR_scaleR)
+apply (rule scaleR_one)
+done
+
+text {* Recover original theorem names *}
+
+lemmas scaleR_left_commute = real_vector.scale_left_commute
+lemmas scaleR_zero_left = real_vector.scale_zero_left
+lemmas scaleR_minus_left = real_vector.scale_minus_left
+lemmas scaleR_left_diff_distrib = real_vector.scale_left_diff_distrib
+lemmas scaleR_zero_right = real_vector.scale_zero_right
+lemmas scaleR_minus_right = real_vector.scale_minus_right
+lemmas scaleR_right_diff_distrib = real_vector.scale_right_diff_distrib
+lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff
+lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq
+lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq
+lemmas scaleR_cancel_left = real_vector.scale_cancel_left
+lemmas scaleR_cancel_right = real_vector.scale_cancel_right
+
+class real_algebra = real_vector + ring +
+  assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)"
+  and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"
+
+class real_algebra_1 = real_algebra + ring_1
+
+class real_div_algebra = real_algebra_1 + division_ring
+
+class real_field = real_div_algebra + field
+
+instance real :: real_field
+apply (intro_classes, unfold real_scaleR_def)
+apply (rule right_distrib)
+apply (rule left_distrib)
+apply (rule mult_assoc [symmetric])
+apply (rule mult_1_left)
+apply (rule mult_assoc)
+apply (rule mult_left_commute)
+done
+
+interpretation scaleR_left: additive ["(\<lambda>a. scaleR a x::'a::real_vector)"]
+proof qed (rule scaleR_left_distrib)
+
+interpretation scaleR_right: additive ["(\<lambda>x. scaleR a x::'a::real_vector)"]
+proof qed (rule scaleR_right_distrib)
+
+lemma nonzero_inverse_scaleR_distrib:
+  fixes x :: "'a::real_div_algebra" shows
+  "\<lbrakk>a \<noteq> 0; x \<noteq> 0\<rbrakk> \<Longrightarrow> inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
+by (rule inverse_unique, simp)
+
+lemma inverse_scaleR_distrib:
+  fixes x :: "'a::{real_div_algebra,division_by_zero}"
+  shows "inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
+apply (case_tac "a = 0", simp)
+apply (case_tac "x = 0", simp)
+apply (erule (1) nonzero_inverse_scaleR_distrib)
+done
+
+
+subsection {* Embedding of the Reals into any @{text real_algebra_1}:
+@{term of_real} *}
+
+definition
+  of_real :: "real \<Rightarrow> 'a::real_algebra_1" where
+  "of_real r = scaleR r 1"
+
+lemma scaleR_conv_of_real: "scaleR r x = of_real r * x"
+by (simp add: of_real_def)
+
+lemma of_real_0 [simp]: "of_real 0 = 0"
+by (simp add: of_real_def)
+
+lemma of_real_1 [simp]: "of_real 1 = 1"
+by (simp add: of_real_def)
+
+lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y"
+by (simp add: of_real_def scaleR_left_distrib)
+
+lemma of_real_minus [simp]: "of_real (- x) = - of_real x"
+by (simp add: of_real_def)
+
+lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y"
+by (simp add: of_real_def scaleR_left_diff_distrib)
+
+lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y"
+by (simp add: of_real_def mult_commute)
+
+lemma nonzero_of_real_inverse:
+  "x \<noteq> 0 \<Longrightarrow> of_real (inverse x) =
+   inverse (of_real x :: 'a::real_div_algebra)"
+by (simp add: of_real_def nonzero_inverse_scaleR_distrib)
+
+lemma of_real_inverse [simp]:
+  "of_real (inverse x) =
+   inverse (of_real x :: 'a::{real_div_algebra,division_by_zero})"
+by (simp add: of_real_def inverse_scaleR_distrib)
+
+lemma nonzero_of_real_divide:
+  "y \<noteq> 0 \<Longrightarrow> of_real (x / y) =
+   (of_real x / of_real y :: 'a::real_field)"
+by (simp add: divide_inverse nonzero_of_real_inverse)
+
+lemma of_real_divide [simp]:
+  "of_real (x / y) =
+   (of_real x / of_real y :: 'a::{real_field,division_by_zero})"
+by (simp add: divide_inverse)
+
+lemma of_real_power [simp]:
+  "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1,recpower}) ^ n"
+by (induct n) (simp_all add: power_Suc)
+
+lemma of_real_eq_iff [simp]: "(of_real x = of_real y) = (x = y)"
+by (simp add: of_real_def scaleR_cancel_right)
+
+lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified]
+
+lemma of_real_eq_id [simp]: "of_real = (id :: real \<Rightarrow> real)"
+proof
+  fix r
+  show "of_real r = id r"
+    by (simp add: of_real_def)
+qed
+
+text{*Collapse nested embeddings*}
+lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n"
+by (induct n) auto
+
+lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z"
+by (cases z rule: int_diff_cases, simp)
+
+lemma of_real_number_of_eq:
+  "of_real (number_of w) = (number_of w :: 'a::{number_ring,real_algebra_1})"
+by (simp add: number_of_eq)
+
+text{*Every real algebra has characteristic zero*}
+instance real_algebra_1 < ring_char_0
+proof
+  fix m n :: nat
+  have "(of_real (of_nat m) = (of_real (of_nat n)::'a)) = (m = n)"
+    by (simp only: of_real_eq_iff of_nat_eq_iff)
+  thus "(of_nat m = (of_nat n::'a)) = (m = n)"
+    by (simp only: of_real_of_nat_eq)
+qed
+
+instance real_field < field_char_0 ..
+
+
+subsection {* The Set of Real Numbers *}
+
+definition
+  Reals :: "'a::real_algebra_1 set" where
+  [code del]: "Reals \<equiv> range of_real"
+
+notation (xsymbols)
+  Reals  ("\<real>")
+
+lemma Reals_of_real [simp]: "of_real r \<in> Reals"
+by (simp add: Reals_def)
+
+lemma Reals_of_int [simp]: "of_int z \<in> Reals"
+by (subst of_real_of_int_eq [symmetric], rule Reals_of_real)
+
+lemma Reals_of_nat [simp]: "of_nat n \<in> Reals"
+by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real)
+
+lemma Reals_number_of [simp]:
+  "(number_of w::'a::{number_ring,real_algebra_1}) \<in> Reals"
+by (subst of_real_number_of_eq [symmetric], rule Reals_of_real)
+
+lemma Reals_0 [simp]: "0 \<in> Reals"
+apply (unfold Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_0 [symmetric])
+done
+
+lemma Reals_1 [simp]: "1 \<in> Reals"
+apply (unfold Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_1 [symmetric])
+done
+
+lemma Reals_add [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a + b \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_add [symmetric])
+done
+
+lemma Reals_minus [simp]: "a \<in> Reals \<Longrightarrow> - a \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_minus [symmetric])
+done
+
+lemma Reals_diff [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a - b \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_diff [symmetric])
+done
+
+lemma Reals_mult [simp]: "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a * b \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_mult [symmetric])
+done
+
+lemma nonzero_Reals_inverse:
+  fixes a :: "'a::real_div_algebra"
+  shows "\<lbrakk>a \<in> Reals; a \<noteq> 0\<rbrakk> \<Longrightarrow> inverse a \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (erule nonzero_of_real_inverse [symmetric])
+done
+
+lemma Reals_inverse [simp]:
+  fixes a :: "'a::{real_div_algebra,division_by_zero}"
+  shows "a \<in> Reals \<Longrightarrow> inverse a \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_inverse [symmetric])
+done
+
+lemma nonzero_Reals_divide:
+  fixes a b :: "'a::real_field"
+  shows "\<lbrakk>a \<in> Reals; b \<in> Reals; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (erule nonzero_of_real_divide [symmetric])
+done
+
+lemma Reals_divide [simp]:
+  fixes a b :: "'a::{real_field,division_by_zero}"
+  shows "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_divide [symmetric])
+done
+
+lemma Reals_power [simp]:
+  fixes a :: "'a::{real_algebra_1,recpower}"
+  shows "a \<in> Reals \<Longrightarrow> a ^ n \<in> Reals"
+apply (auto simp add: Reals_def)
+apply (rule range_eqI)
+apply (rule of_real_power [symmetric])
+done
+
+lemma Reals_cases [cases set: Reals]:
+  assumes "q \<in> \<real>"
+  obtains (of_real) r where "q = of_real r"
+  unfolding Reals_def
+proof -
+  from `q \<in> \<real>` have "q \<in> range of_real" unfolding Reals_def .
+  then obtain r where "q = of_real r" ..
+  then show thesis ..
+qed
+
+lemma Reals_induct [case_names of_real, induct set: Reals]:
+  "q \<in> \<real> \<Longrightarrow> (\<And>r. P (of_real r)) \<Longrightarrow> P q"
+  by (rule Reals_cases) auto
+
+
+subsection {* Real normed vector spaces *}
+
+class norm = type +
+  fixes norm :: "'a \<Rightarrow> real"
+
+instantiation real :: norm
+begin
+
+definition
+  real_norm_def [simp]: "norm r \<equiv> \<bar>r\<bar>"
+
+instance ..
+
+end
+
+class sgn_div_norm = scaleR + norm + sgn +
+  assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x"
+
+class real_normed_vector = real_vector + sgn_div_norm +
+  assumes norm_ge_zero [simp]: "0 \<le> norm x"
+  and norm_eq_zero [simp]: "norm x = 0 \<longleftrightarrow> x = 0"
+  and norm_triangle_ineq: "norm (x + y) \<le> norm x + norm y"
+  and norm_scaleR: "norm (scaleR a x) = \<bar>a\<bar> * norm x"
+
+class real_normed_algebra = real_algebra + real_normed_vector +
+  assumes norm_mult_ineq: "norm (x * y) \<le> norm x * norm y"
+
+class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra +
+  assumes norm_one [simp]: "norm 1 = 1"
+
+class real_normed_div_algebra = real_div_algebra + real_normed_vector +
+  assumes norm_mult: "norm (x * y) = norm x * norm y"
+
+class real_normed_field = real_field + real_normed_div_algebra
+
+instance real_normed_div_algebra < real_normed_algebra_1
+proof
+  fix x y :: 'a
+  show "norm (x * y) \<le> norm x * norm y"
+    by (simp add: norm_mult)
+next
+  have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)"
+    by (rule norm_mult)
+  thus "norm (1::'a) = 1" by simp
+qed
+
+instance real :: real_normed_field
+apply (intro_classes, unfold real_norm_def real_scaleR_def)
+apply (simp add: real_sgn_def)
+apply (rule abs_ge_zero)
+apply (rule abs_eq_0)
+apply (rule abs_triangle_ineq)
+apply (rule abs_mult)
+apply (rule abs_mult)
+done
+
+lemma norm_zero [simp]: "norm (0::'a::real_normed_vector) = 0"
+by simp
+
+lemma zero_less_norm_iff [simp]:
+  fixes x :: "'a::real_normed_vector"
+  shows "(0 < norm x) = (x \<noteq> 0)"
+by (simp add: order_less_le)
+
+lemma norm_not_less_zero [simp]:
+  fixes x :: "'a::real_normed_vector"
+  shows "\<not> norm x < 0"
+by (simp add: linorder_not_less)
+
+lemma norm_le_zero_iff [simp]:
+  fixes x :: "'a::real_normed_vector"
+  shows "(norm x \<le> 0) = (x = 0)"
+by (simp add: order_le_less)
+
+lemma norm_minus_cancel [simp]:
+  fixes x :: "'a::real_normed_vector"
+  shows "norm (- x) = norm x"
+proof -
+  have "norm (- x) = norm (scaleR (- 1) x)"
+    by (simp only: scaleR_minus_left scaleR_one)
+  also have "\<dots> = \<bar>- 1\<bar> * norm x"
+    by (rule norm_scaleR)
+  finally show ?thesis by simp
+qed
+
+lemma norm_minus_commute:
+  fixes a b :: "'a::real_normed_vector"
+  shows "norm (a - b) = norm (b - a)"
+proof -
+  have "norm (- (b - a)) = norm (b - a)"
+    by (rule norm_minus_cancel)
+  thus ?thesis by simp
+qed
+
+lemma norm_triangle_ineq2:
+  fixes a b :: "'a::real_normed_vector"
+  shows "norm a - norm b \<le> norm (a - b)"
+proof -
+  have "norm (a - b + b) \<le> norm (a - b) + norm b"
+    by (rule norm_triangle_ineq)
+  thus ?thesis by simp
+qed
+
+lemma norm_triangle_ineq3:
+  fixes a b :: "'a::real_normed_vector"
+  shows "\<bar>norm a - norm b\<bar> \<le> norm (a - b)"
+apply (subst abs_le_iff)
+apply auto
+apply (rule norm_triangle_ineq2)
+apply (subst norm_minus_commute)
+apply (rule norm_triangle_ineq2)
+done
+
+lemma norm_triangle_ineq4:
+  fixes a b :: "'a::real_normed_vector"
+  shows "norm (a - b) \<le> norm a + norm b"
+proof -
+  have "norm (a + - b) \<le> norm a + norm (- b)"
+    by (rule norm_triangle_ineq)
+  thus ?thesis
+    by (simp only: diff_minus norm_minus_cancel)
+qed
+
+lemma norm_diff_ineq:
+  fixes a b :: "'a::real_normed_vector"
+  shows "norm a - norm b \<le> norm (a + b)"
+proof -
+  have "norm a - norm (- b) \<le> norm (a - - b)"
+    by (rule norm_triangle_ineq2)
+  thus ?thesis by simp
+qed
+
+lemma norm_diff_triangle_ineq:
+  fixes a b c d :: "'a::real_normed_vector"
+  shows "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
+proof -
+  have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))"
+    by (simp add: diff_minus add_ac)
+  also have "\<dots> \<le> norm (a - c) + norm (b - d)"
+    by (rule norm_triangle_ineq)
+  finally show ?thesis .
+qed
+
+lemma abs_norm_cancel [simp]:
+  fixes a :: "'a::real_normed_vector"
+  shows "\<bar>norm a\<bar> = norm a"
+by (rule abs_of_nonneg [OF norm_ge_zero])
+
+lemma norm_add_less:
+  fixes x y :: "'a::real_normed_vector"
+  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x + y) < r + s"
+by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono])
+
+lemma norm_mult_less:
+  fixes x y :: "'a::real_normed_algebra"
+  shows "\<lbrakk>norm x < r; norm y < s\<rbrakk> \<Longrightarrow> norm (x * y) < r * s"
+apply (rule order_le_less_trans [OF norm_mult_ineq])
+apply (simp add: mult_strict_mono')
+done
+
+lemma norm_of_real [simp]:
+  "norm (of_real r :: 'a::real_normed_algebra_1) = \<bar>r\<bar>"
+unfolding of_real_def by (simp add: norm_scaleR)
+
+lemma norm_number_of [simp]:
+  "norm (number_of w::'a::{number_ring,real_normed_algebra_1})
+    = \<bar>number_of w\<bar>"
+by (subst of_real_number_of_eq [symmetric], rule norm_of_real)
+
+lemma norm_of_int [simp]:
+  "norm (of_int z::'a::real_normed_algebra_1) = \<bar>of_int z\<bar>"
+by (subst of_real_of_int_eq [symmetric], rule norm_of_real)
+
+lemma norm_of_nat [simp]:
+  "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n"
+apply (subst of_real_of_nat_eq [symmetric])
+apply (subst norm_of_real, simp)
+done
+
+lemma nonzero_norm_inverse:
+  fixes a :: "'a::real_normed_div_algebra"
+  shows "a \<noteq> 0 \<Longrightarrow> norm (inverse a) = inverse (norm a)"
+apply (rule inverse_unique [symmetric])
+apply (simp add: norm_mult [symmetric])
+done
+
+lemma norm_inverse:
+  fixes a :: "'a::{real_normed_div_algebra,division_by_zero}"
+  shows "norm (inverse a) = inverse (norm a)"
+apply (case_tac "a = 0", simp)
+apply (erule nonzero_norm_inverse)
+done
+
+lemma nonzero_norm_divide:
+  fixes a b :: "'a::real_normed_field"
+  shows "b \<noteq> 0 \<Longrightarrow> norm (a / b) = norm a / norm b"
+by (simp add: divide_inverse norm_mult nonzero_norm_inverse)
+
+lemma norm_divide:
+  fixes a b :: "'a::{real_normed_field,division_by_zero}"
+  shows "norm (a / b) = norm a / norm b"
+by (simp add: divide_inverse norm_mult norm_inverse)
+
+lemma norm_power_ineq:
+  fixes x :: "'a::{real_normed_algebra_1,recpower}"
+  shows "norm (x ^ n) \<le> norm x ^ n"
+proof (induct n)
+  case 0 show "norm (x ^ 0) \<le> norm x ^ 0" by simp
+next
+  case (Suc n)
+  have "norm (x * x ^ n) \<le> norm x * norm (x ^ n)"
+    by (rule norm_mult_ineq)
+  also from Suc have "\<dots> \<le> norm x * norm x ^ n"
+    using norm_ge_zero by (rule mult_left_mono)
+  finally show "norm (x ^ Suc n) \<le> norm x ^ Suc n"
+    by (simp add: power_Suc)
+qed
+
+lemma norm_power:
+  fixes x :: "'a::{real_normed_div_algebra,recpower}"
+  shows "norm (x ^ n) = norm x ^ n"
+by (induct n) (simp_all add: power_Suc norm_mult)
+
+
+subsection {* Sign function *}
+
+lemma norm_sgn:
+  "norm (sgn(x::'a::real_normed_vector)) = (if x = 0 then 0 else 1)"
+by (simp add: sgn_div_norm norm_scaleR)
+
+lemma sgn_zero [simp]: "sgn(0::'a::real_normed_vector) = 0"
+by (simp add: sgn_div_norm)
+
+lemma sgn_zero_iff: "(sgn(x::'a::real_normed_vector) = 0) = (x = 0)"
+by (simp add: sgn_div_norm)
+
+lemma sgn_minus: "sgn (- x) = - sgn(x::'a::real_normed_vector)"
+by (simp add: sgn_div_norm)
+
+lemma sgn_scaleR:
+  "sgn (scaleR r x) = scaleR (sgn r) (sgn(x::'a::real_normed_vector))"
+by (simp add: sgn_div_norm norm_scaleR mult_ac)
+
+lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1"
+by (simp add: sgn_div_norm)
+
+lemma sgn_of_real:
+  "sgn (of_real r::'a::real_normed_algebra_1) = of_real (sgn r)"
+unfolding of_real_def by (simp only: sgn_scaleR sgn_one)
+
+lemma sgn_mult:
+  fixes x y :: "'a::real_normed_div_algebra"
+  shows "sgn (x * y) = sgn x * sgn y"
+by (simp add: sgn_div_norm norm_mult mult_commute)
+
+lemma real_sgn_eq: "sgn (x::real) = x / \<bar>x\<bar>"
+by (simp add: sgn_div_norm divide_inverse)
+
+lemma real_sgn_pos: "0 < (x::real) \<Longrightarrow> sgn x = 1"
+unfolding real_sgn_eq by simp
+
+lemma real_sgn_neg: "(x::real) < 0 \<Longrightarrow> sgn x = -1"
+unfolding real_sgn_eq by simp
+
+
+subsection {* Bounded Linear and Bilinear Operators *}
+
+locale bounded_linear = additive +
+  constrains f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes scaleR: "f (scaleR r x) = scaleR r (f x)"
+  assumes bounded: "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
+begin
+
+lemma pos_bounded:
+  "\<exists>K>0. \<forall>x. norm (f x) \<le> norm x * K"
+proof -
+  obtain K where K: "\<And>x. norm (f x) \<le> norm x * K"
+    using bounded by fast
+  show ?thesis
+  proof (intro exI impI conjI allI)
+    show "0 < max 1 K"
+      by (rule order_less_le_trans [OF zero_less_one le_maxI1])
+  next
+    fix x
+    have "norm (f x) \<le> norm x * K" using K .
+    also have "\<dots> \<le> norm x * max 1 K"
+      by (rule mult_left_mono [OF le_maxI2 norm_ge_zero])
+    finally show "norm (f x) \<le> norm x * max 1 K" .
+  qed
+qed
+
+lemma nonneg_bounded:
+  "\<exists>K\<ge>0. \<forall>x. norm (f x) \<le> norm x * K"
+proof -
+  from pos_bounded
+  show ?thesis by (auto intro: order_less_imp_le)
+qed
+
+end
+
+locale bounded_bilinear =
+  fixes prod :: "['a::real_normed_vector, 'b::real_normed_vector]
+                 \<Rightarrow> 'c::real_normed_vector"
+    (infixl "**" 70)
+  assumes add_left: "prod (a + a') b = prod a b + prod a' b"
+  assumes add_right: "prod a (b + b') = prod a b + prod a b'"
+  assumes scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)"
+  assumes scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)"
+  assumes bounded: "\<exists>K. \<forall>a b. norm (prod a b) \<le> norm a * norm b * K"
+begin
+
+lemma pos_bounded:
+  "\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
+apply (cut_tac bounded, erule exE)
+apply (rule_tac x="max 1 K" in exI, safe)
+apply (rule order_less_le_trans [OF zero_less_one le_maxI1])
+apply (drule spec, drule spec, erule order_trans)
+apply (rule mult_left_mono [OF le_maxI2])
+apply (intro mult_nonneg_nonneg norm_ge_zero)
+done
+
+lemma nonneg_bounded:
+  "\<exists>K\<ge>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
+proof -
+  from pos_bounded
+  show ?thesis by (auto intro: order_less_imp_le)
+qed
+
+lemma additive_right: "additive (\<lambda>b. prod a b)"
+by (rule additive.intro, rule add_right)
+
+lemma additive_left: "additive (\<lambda>a. prod a b)"
+by (rule additive.intro, rule add_left)
+
+lemma zero_left: "prod 0 b = 0"
+by (rule additive.zero [OF additive_left])
+
+lemma zero_right: "prod a 0 = 0"
+by (rule additive.zero [OF additive_right])
+
+lemma minus_left: "prod (- a) b = - prod a b"
+by (rule additive.minus [OF additive_left])
+
+lemma minus_right: "prod a (- b) = - prod a b"
+by (rule additive.minus [OF additive_right])
+
+lemma diff_left:
+  "prod (a - a') b = prod a b - prod a' b"
+by (rule additive.diff [OF additive_left])
+
+lemma diff_right:
+  "prod a (b - b') = prod a b - prod a b'"
+by (rule additive.diff [OF additive_right])
+
+lemma bounded_linear_left:
+  "bounded_linear (\<lambda>a. a ** b)"
+apply (unfold_locales)
+apply (rule add_left)
+apply (rule scaleR_left)
+apply (cut_tac bounded, safe)
+apply (rule_tac x="norm b * K" in exI)
+apply (simp add: mult_ac)
+done
+
+lemma bounded_linear_right:
+  "bounded_linear (\<lambda>b. a ** b)"
+apply (unfold_locales)
+apply (rule add_right)
+apply (rule scaleR_right)
+apply (cut_tac bounded, safe)
+apply (rule_tac x="norm a * K" in exI)
+apply (simp add: mult_ac)
+done
+
+lemma prod_diff_prod:
+  "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)"
+by (simp add: diff_left diff_right)
+
+end
+
+interpretation mult:
+  bounded_bilinear ["op * :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra"]
+apply (rule bounded_bilinear.intro)
+apply (rule left_distrib)
+apply (rule right_distrib)
+apply (rule mult_scaleR_left)
+apply (rule mult_scaleR_right)
+apply (rule_tac x="1" in exI)
+apply (simp add: norm_mult_ineq)
+done
+
+interpretation mult_left:
+  bounded_linear ["(\<lambda>x::'a::real_normed_algebra. x * y)"]
+by (rule mult.bounded_linear_left)
+
+interpretation mult_right:
+  bounded_linear ["(\<lambda>y::'a::real_normed_algebra. x * y)"]
+by (rule mult.bounded_linear_right)
+
+interpretation divide:
+  bounded_linear ["(\<lambda>x::'a::real_normed_field. x / y)"]
+unfolding divide_inverse by (rule mult.bounded_linear_left)
+
+interpretation scaleR: bounded_bilinear ["scaleR"]
+apply (rule bounded_bilinear.intro)
+apply (rule scaleR_left_distrib)
+apply (rule scaleR_right_distrib)
+apply simp
+apply (rule scaleR_left_commute)
+apply (rule_tac x="1" in exI)
+apply (simp add: norm_scaleR)
+done
+
+interpretation scaleR_left: bounded_linear ["\<lambda>r. scaleR r x"]
+by (rule scaleR.bounded_linear_left)
+
+interpretation scaleR_right: bounded_linear ["\<lambda>x. scaleR r x"]
+by (rule scaleR.bounded_linear_right)
+
+interpretation of_real: bounded_linear ["\<lambda>r. of_real r"]
+unfolding of_real_def by (rule scaleR.bounded_linear_left)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SEQ.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -0,0 +1,1136 @@
+(*  Title       : SEQ.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Convergence of sequences and series
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
+    Additional contributions by Jeremy Avigad and Brian Huffman
+*)
+
+header {* Sequences and Convergence *}
+
+theory SEQ
+imports RealVector RComplete
+begin
+
+definition
+  Zseq :: "[nat \<Rightarrow> 'a::real_normed_vector] \<Rightarrow> bool" where
+    --{*Standard definition of sequence converging to zero*}
+  [code del]: "Zseq X = (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. norm (X n) < r)"
+
+definition
+  LIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool"
+    ("((_)/ ----> (_))" [60, 60] 60) where
+    --{*Standard definition of convergence of sequence*}
+  [code del]: "X ----> L = (\<forall>r. 0 < r --> (\<exists>no. \<forall>n. no \<le> n --> norm (X n - L) < r))"
+
+definition
+  lim :: "(nat => 'a::real_normed_vector) => 'a" where
+    --{*Standard definition of limit using choice operator*}
+  "lim X = (THE L. X ----> L)"
+
+definition
+  convergent :: "(nat => 'a::real_normed_vector) => bool" where
+    --{*Standard definition of convergence*}
+  "convergent X = (\<exists>L. X ----> L)"
+
+definition
+  Bseq :: "(nat => 'a::real_normed_vector) => bool" where
+    --{*Standard definition for bounded sequence*}
+  [code del]: "Bseq X = (\<exists>K>0.\<forall>n. norm (X n) \<le> K)"
+
+definition
+  monoseq :: "(nat=>real)=>bool" where
+    --{*Definition for monotonicity*}
+  [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))"
+
+definition
+  subseq :: "(nat => nat) => bool" where
+    --{*Definition of subsequence*}
+  [code del]:   "subseq f = (\<forall>m. \<forall>n>m. (f m) < (f n))"
+
+definition
+  Cauchy :: "(nat => 'a::real_normed_vector) => bool" where
+    --{*Standard definition of the Cauchy condition*}
+  [code del]: "Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. norm (X m - X n) < e)"
+
+
+subsection {* Bounded Sequences *}
+
+lemma BseqI': assumes K: "\<And>n. norm (X n) \<le> K" shows "Bseq X"
+unfolding Bseq_def
+proof (intro exI conjI allI)
+  show "0 < max K 1" by simp
+next
+  fix n::nat
+  have "norm (X n) \<le> K" by (rule K)
+  thus "norm (X n) \<le> max K 1" by simp
+qed
+
+lemma BseqE: "\<lbrakk>Bseq X; \<And>K. \<lbrakk>0 < K; \<forall>n. norm (X n) \<le> K\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
+unfolding Bseq_def by auto
+
+lemma BseqI2': assumes K: "\<forall>n\<ge>N. norm (X n) \<le> K" shows "Bseq X"
+proof (rule BseqI')
+  let ?A = "norm ` X ` {..N}"
+  have 1: "finite ?A" by simp
+  fix n::nat
+  show "norm (X n) \<le> max K (Max ?A)"
+  proof (cases rule: linorder_le_cases)
+    assume "n \<ge> N"
+    hence "norm (X n) \<le> K" using K by simp
+    thus "norm (X n) \<le> max K (Max ?A)" by simp
+  next
+    assume "n \<le> N"
+    hence "norm (X n) \<in> ?A" by simp
+    with 1 have "norm (X n) \<le> Max ?A" by (rule Max_ge)
+    thus "norm (X n) \<le> max K (Max ?A)" by simp
+  qed
+qed
+
+lemma Bseq_ignore_initial_segment: "Bseq X \<Longrightarrow> Bseq (\<lambda>n. X (n + k))"
+unfolding Bseq_def by auto
+
+lemma Bseq_offset: "Bseq (\<lambda>n. X (n + k)) \<Longrightarrow> Bseq X"
+apply (erule BseqE)
+apply (rule_tac N="k" and K="K" in BseqI2')
+apply clarify
+apply (drule_tac x="n - k" in spec, simp)
+done
+
+
+subsection {* Sequences That Converge to Zero *}
+
+lemma ZseqI:
+  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r) \<Longrightarrow> Zseq X"
+unfolding Zseq_def by simp
+
+lemma ZseqD:
+  "\<lbrakk>Zseq X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n) < r"
+unfolding Zseq_def by simp
+
+lemma Zseq_zero: "Zseq (\<lambda>n. 0)"
+unfolding Zseq_def by simp
+
+lemma Zseq_const_iff: "Zseq (\<lambda>n. k) = (k = 0)"
+unfolding Zseq_def by force
+
+lemma Zseq_norm_iff: "Zseq (\<lambda>n. norm (X n)) = Zseq (\<lambda>n. X n)"
+unfolding Zseq_def by simp
+
+lemma Zseq_imp_Zseq:
+  assumes X: "Zseq X"
+  assumes Y: "\<And>n. norm (Y n) \<le> norm (X n) * K"
+  shows "Zseq (\<lambda>n. Y n)"
+proof (cases)
+  assume K: "0 < K"
+  show ?thesis
+  proof (rule ZseqI)
+    fix r::real assume "0 < r"
+    hence "0 < r / K"
+      using K by (rule divide_pos_pos)
+    then obtain N where "\<forall>n\<ge>N. norm (X n) < r / K"
+      using ZseqD [OF X] by fast
+    hence "\<forall>n\<ge>N. norm (X n) * K < r"
+      by (simp add: pos_less_divide_eq K)
+    hence "\<forall>n\<ge>N. norm (Y n) < r"
+      by (simp add: order_le_less_trans [OF Y])
+    thus "\<exists>N. \<forall>n\<ge>N. norm (Y n) < r" ..
+  qed
+next
+  assume "\<not> 0 < K"
+  hence K: "K \<le> 0" by (simp only: linorder_not_less)
+  {
+    fix n::nat
+    have "norm (Y n) \<le> norm (X n) * K" by (rule Y)
+    also have "\<dots> \<le> norm (X n) * 0"
+      using K norm_ge_zero by (rule mult_left_mono)
+    finally have "norm (Y n) = 0" by simp
+  }
+  thus ?thesis by (simp add: Zseq_zero)
+qed
+
+lemma Zseq_le: "\<lbrakk>Zseq Y; \<forall>n. norm (X n) \<le> norm (Y n)\<rbrakk> \<Longrightarrow> Zseq X"
+by (erule_tac K="1" in Zseq_imp_Zseq, simp)
+
+lemma Zseq_add:
+  assumes X: "Zseq X"
+  assumes Y: "Zseq Y"
+  shows "Zseq (\<lambda>n. X n + Y n)"
+proof (rule ZseqI)
+  fix r::real assume "0 < r"
+  hence r: "0 < r / 2" by simp
+  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r/2"
+    using ZseqD [OF X r] by fast
+  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < r/2"
+    using ZseqD [OF Y r] by fast
+  show "\<exists>N. \<forall>n\<ge>N. norm (X n + Y n) < r"
+  proof (intro exI allI impI)
+    fix n assume n: "max M N \<le> n"
+    have "norm (X n + Y n) \<le> norm (X n) + norm (Y n)"
+      by (rule norm_triangle_ineq)
+    also have "\<dots> < r/2 + r/2"
+    proof (rule add_strict_mono)
+      from M n show "norm (X n) < r/2" by simp
+      from N n show "norm (Y n) < r/2" by simp
+    qed
+    finally show "norm (X n + Y n) < r" by simp
+  qed
+qed
+
+lemma Zseq_minus: "Zseq X \<Longrightarrow> Zseq (\<lambda>n. - X n)"
+unfolding Zseq_def by simp
+
+lemma Zseq_diff: "\<lbrakk>Zseq X; Zseq Y\<rbrakk> \<Longrightarrow> Zseq (\<lambda>n. X n - Y n)"
+by (simp only: diff_minus Zseq_add Zseq_minus)
+
+lemma (in bounded_linear) Zseq:
+  assumes X: "Zseq X"
+  shows "Zseq (\<lambda>n. f (X n))"
+proof -
+  obtain K where "\<And>x. norm (f x) \<le> norm x * K"
+    using bounded by fast
+  with X show ?thesis
+    by (rule Zseq_imp_Zseq)
+qed
+
+lemma (in bounded_bilinear) Zseq:
+  assumes X: "Zseq X"
+  assumes Y: "Zseq Y"
+  shows "Zseq (\<lambda>n. X n ** Y n)"
+proof (rule ZseqI)
+  fix r::real assume r: "0 < r"
+  obtain K where K: "0 < K"
+    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
+    using pos_bounded by fast
+  from K have K': "0 < inverse K"
+    by (rule positive_imp_inverse_positive)
+  obtain M where M: "\<forall>n\<ge>M. norm (X n) < r"
+    using ZseqD [OF X r] by fast
+  obtain N where N: "\<forall>n\<ge>N. norm (Y n) < inverse K"
+    using ZseqD [OF Y K'] by fast
+  show "\<exists>N. \<forall>n\<ge>N. norm (X n ** Y n) < r"
+  proof (intro exI allI impI)
+    fix n assume n: "max M N \<le> n"
+    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
+      by (rule norm_le)
+    also have "norm (X n) * norm (Y n) * K < r * inverse K * K"
+    proof (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero K)
+      from M n show Xn: "norm (X n) < r" by simp
+      from N n show Yn: "norm (Y n) < inverse K" by simp
+    qed
+    also from K have "r * inverse K * K = r" by simp
+    finally show "norm (X n ** Y n) < r" .
+  qed
+qed
+
+lemma (in bounded_bilinear) Zseq_prod_Bseq:
+  assumes X: "Zseq X"
+  assumes Y: "Bseq Y"
+  shows "Zseq (\<lambda>n. X n ** Y n)"
+proof -
+  obtain K where K: "0 \<le> K"
+    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
+    using nonneg_bounded by fast
+  obtain B where B: "0 < B"
+    and norm_Y: "\<And>n. norm (Y n) \<le> B"
+    using Y [unfolded Bseq_def] by fast
+  from X show ?thesis
+  proof (rule Zseq_imp_Zseq)
+    fix n::nat
+    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
+      by (rule norm_le)
+    also have "\<dots> \<le> norm (X n) * B * K"
+      by (intro mult_mono' order_refl norm_Y norm_ge_zero
+                mult_nonneg_nonneg K)
+    also have "\<dots> = norm (X n) * (B * K)"
+      by (rule mult_assoc)
+    finally show "norm (X n ** Y n) \<le> norm (X n) * (B * K)" .
+  qed
+qed
+
+lemma (in bounded_bilinear) Bseq_prod_Zseq:
+  assumes X: "Bseq X"
+  assumes Y: "Zseq Y"
+  shows "Zseq (\<lambda>n. X n ** Y n)"
+proof -
+  obtain K where K: "0 \<le> K"
+    and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
+    using nonneg_bounded by fast
+  obtain B where B: "0 < B"
+    and norm_X: "\<And>n. norm (X n) \<le> B"
+    using X [unfolded Bseq_def] by fast
+  from Y show ?thesis
+  proof (rule Zseq_imp_Zseq)
+    fix n::nat
+    have "norm (X n ** Y n) \<le> norm (X n) * norm (Y n) * K"
+      by (rule norm_le)
+    also have "\<dots> \<le> B * norm (Y n) * K"
+      by (intro mult_mono' order_refl norm_X norm_ge_zero
+                mult_nonneg_nonneg K)
+    also have "\<dots> = norm (Y n) * (B * K)"
+      by (simp only: mult_ac)
+    finally show "norm (X n ** Y n) \<le> norm (Y n) * (B * K)" .
+  qed
+qed
+
+lemma (in bounded_bilinear) Zseq_left:
+  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. X n ** a)"
+by (rule bounded_linear_left [THEN bounded_linear.Zseq])
+
+lemma (in bounded_bilinear) Zseq_right:
+  "Zseq X \<Longrightarrow> Zseq (\<lambda>n. a ** X n)"
+by (rule bounded_linear_right [THEN bounded_linear.Zseq])
+
+lemmas Zseq_mult = mult.Zseq
+lemmas Zseq_mult_right = mult.Zseq_right
+lemmas Zseq_mult_left = mult.Zseq_left
+
+
+subsection {* Limits of Sequences *}
+
+lemma LIMSEQ_iff:
+      "(X ----> L) = (\<forall>r>0. \<exists>no. \<forall>n \<ge> no. norm (X n - L) < r)"
+by (rule LIMSEQ_def)
+
+lemma LIMSEQ_Zseq_iff: "((\<lambda>n. X n) ----> L) = Zseq (\<lambda>n. X n - L)"
+by (simp only: LIMSEQ_def Zseq_def)
+
+lemma LIMSEQ_I:
+  "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r) \<Longrightarrow> X ----> L"
+by (simp add: LIMSEQ_def)
+
+lemma LIMSEQ_D:
+  "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
+by (simp add: LIMSEQ_def)
+
+lemma LIMSEQ_const: "(\<lambda>n. k) ----> k"
+by (simp add: LIMSEQ_def)
+
+lemma LIMSEQ_const_iff: "(\<lambda>n. k) ----> l = (k = l)"
+by (simp add: LIMSEQ_Zseq_iff Zseq_const_iff)
+
+lemma LIMSEQ_norm: "X ----> a \<Longrightarrow> (\<lambda>n. norm (X n)) ----> norm a"
+apply (simp add: LIMSEQ_def, safe)
+apply (drule_tac x="r" in spec, safe)
+apply (rule_tac x="no" in exI, safe)
+apply (drule_tac x="n" in spec, safe)
+apply (erule order_le_less_trans [OF norm_triangle_ineq3])
+done
+
+lemma LIMSEQ_ignore_initial_segment:
+  "f ----> a \<Longrightarrow> (\<lambda>n. f (n + k)) ----> a"
+apply (rule LIMSEQ_I)
+apply (drule (1) LIMSEQ_D)
+apply (erule exE, rename_tac N)
+apply (rule_tac x=N in exI)
+apply simp
+done
+
+lemma LIMSEQ_offset:
+  "(\<lambda>n. f (n + k)) ----> a \<Longrightarrow> f ----> a"
+apply (rule LIMSEQ_I)
+apply (drule (1) LIMSEQ_D)
+apply (erule exE, rename_tac N)
+apply (rule_tac x="N + k" in exI)
+apply clarify
+apply (drule_tac x="n - k" in spec)
+apply (simp add: le_diff_conv2)
+done
+
+lemma LIMSEQ_Suc: "f ----> l \<Longrightarrow> (\<lambda>n. f (Suc n)) ----> l"
+by (drule_tac k="1" in LIMSEQ_ignore_initial_segment, simp)
+
+lemma LIMSEQ_imp_Suc: "(\<lambda>n. f (Suc n)) ----> l \<Longrightarrow> f ----> l"
+by (rule_tac k="1" in LIMSEQ_offset, simp)
+
+lemma LIMSEQ_Suc_iff: "(\<lambda>n. f (Suc n)) ----> l = f ----> l"
+by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc)
+
+lemma add_diff_add:
+  fixes a b c d :: "'a::ab_group_add"
+  shows "(a + c) - (b + d) = (a - b) + (c - d)"
+by simp
+
+lemma minus_diff_minus:
+  fixes a b :: "'a::ab_group_add"
+  shows "(- a) - (- b) = - (a - b)"
+by simp
+
+lemma LIMSEQ_add: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) ----> a + b"
+by (simp only: LIMSEQ_Zseq_iff add_diff_add Zseq_add)
+
+lemma LIMSEQ_minus: "X ----> a \<Longrightarrow> (\<lambda>n. - X n) ----> - a"
+by (simp only: LIMSEQ_Zseq_iff minus_diff_minus Zseq_minus)
+
+lemma LIMSEQ_minus_cancel: "(\<lambda>n. - X n) ----> - a \<Longrightarrow> X ----> a"
+by (drule LIMSEQ_minus, simp)
+
+lemma LIMSEQ_diff: "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) ----> a - b"
+by (simp add: diff_minus LIMSEQ_add LIMSEQ_minus)
+
+lemma LIMSEQ_unique: "\<lbrakk>X ----> a; X ----> b\<rbrakk> \<Longrightarrow> a = b"
+by (drule (1) LIMSEQ_diff, simp add: LIMSEQ_const_iff)
+
+lemma (in bounded_linear) LIMSEQ:
+  "X ----> a \<Longrightarrow> (\<lambda>n. f (X n)) ----> f a"
+by (simp only: LIMSEQ_Zseq_iff diff [symmetric] Zseq)
+
+lemma (in bounded_bilinear) LIMSEQ:
+  "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n ** Y n) ----> a ** b"
+by (simp only: LIMSEQ_Zseq_iff prod_diff_prod
+               Zseq_add Zseq Zseq_left Zseq_right)
+
+lemma LIMSEQ_mult:
+  fixes a b :: "'a::real_normed_algebra"
+  shows "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b"
+by (rule mult.LIMSEQ)
+
+lemma inverse_diff_inverse:
+  "\<lbrakk>(a::'a::division_ring) \<noteq> 0; b \<noteq> 0\<rbrakk>
+   \<Longrightarrow> inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
+by (simp add: ring_simps)
+
+lemma Bseq_inverse_lemma:
+  fixes x :: "'a::real_normed_div_algebra"
+  shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
+apply (subst nonzero_norm_inverse, clarsimp)
+apply (erule (1) le_imp_inverse_le)
+done
+
+lemma Bseq_inverse:
+  fixes a :: "'a::real_normed_div_algebra"
+  assumes X: "X ----> a"
+  assumes a: "a \<noteq> 0"
+  shows "Bseq (\<lambda>n. inverse (X n))"
+proof -
+  from a have "0 < norm a" by simp
+  hence "\<exists>r>0. r < norm a" by (rule dense)
+  then obtain r where r1: "0 < r" and r2: "r < norm a" by fast
+  obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> norm (X n - a) < r"
+    using LIMSEQ_D [OF X r1] by fast
+  show ?thesis
+  proof (rule BseqI2' [rule_format])
+    fix n assume n: "N \<le> n"
+    hence 1: "norm (X n - a) < r" by (rule N)
+    hence 2: "X n \<noteq> 0" using r2 by auto
+    hence "norm (inverse (X n)) = inverse (norm (X n))"
+      by (rule nonzero_norm_inverse)
+    also have "\<dots> \<le> inverse (norm a - r)"
+    proof (rule le_imp_inverse_le)
+      show "0 < norm a - r" using r2 by simp
+    next
+      have "norm a - norm (X n) \<le> norm (a - X n)"
+        by (rule norm_triangle_ineq2)
+      also have "\<dots> = norm (X n - a)"
+        by (rule norm_minus_commute)
+      also have "\<dots> < r" using 1 .
+      finally show "norm a - r \<le> norm (X n)" by simp
+    qed
+    finally show "norm (inverse (X n)) \<le> inverse (norm a - r)" .
+  qed
+qed
+
+lemma LIMSEQ_inverse_lemma:
+  fixes a :: "'a::real_normed_div_algebra"
+  shows "\<lbrakk>X ----> a; a \<noteq> 0; \<forall>n. X n \<noteq> 0\<rbrakk>
+         \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> inverse a"
+apply (subst LIMSEQ_Zseq_iff)
+apply (simp add: inverse_diff_inverse nonzero_imp_inverse_nonzero)
+apply (rule Zseq_minus)
+apply (rule Zseq_mult_left)
+apply (rule mult.Bseq_prod_Zseq)
+apply (erule (1) Bseq_inverse)
+apply (simp add: LIMSEQ_Zseq_iff)
+done
+
+lemma LIMSEQ_inverse:
+  fixes a :: "'a::real_normed_div_algebra"
+  assumes X: "X ----> a"
+  assumes a: "a \<noteq> 0"
+  shows "(\<lambda>n. inverse (X n)) ----> inverse a"
+proof -
+  from a have "0 < norm a" by simp
+  then obtain k where "\<forall>n\<ge>k. norm (X n - a) < norm a"
+    using LIMSEQ_D [OF X] by fast
+  hence "\<forall>n\<ge>k. X n \<noteq> 0" by auto
+  hence k: "\<forall>n. X (n + k) \<noteq> 0" by simp
+
+  from X have "(\<lambda>n. X (n + k)) ----> a"
+    by (rule LIMSEQ_ignore_initial_segment)
+  hence "(\<lambda>n. inverse (X (n + k))) ----> inverse a"
+    using a k by (rule LIMSEQ_inverse_lemma)
+  thus "(\<lambda>n. inverse (X n)) ----> inverse a"
+    by (rule LIMSEQ_offset)
+qed
+
+lemma LIMSEQ_divide:
+  fixes a b :: "'a::real_normed_field"
+  shows "\<lbrakk>X ----> a; Y ----> b; b \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>n. X n / Y n) ----> a / b"
+by (simp add: LIMSEQ_mult LIMSEQ_inverse divide_inverse)
+
+lemma LIMSEQ_pow:
+  fixes a :: "'a::{real_normed_algebra,recpower}"
+  shows "X ----> a \<Longrightarrow> (\<lambda>n. (X n) ^ m) ----> a ^ m"
+by (induct m) (simp_all add: power_Suc LIMSEQ_const LIMSEQ_mult)
+
+lemma LIMSEQ_setsum:
+  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
+  shows "(\<lambda>m. \<Sum>n\<in>S. X n m) ----> (\<Sum>n\<in>S. L n)"
+proof (cases "finite S")
+  case True
+  thus ?thesis using n
+  proof (induct)
+    case empty
+    show ?case
+      by (simp add: LIMSEQ_const)
+  next
+    case insert
+    thus ?case
+      by (simp add: LIMSEQ_add)
+  qed
+next
+  case False
+  thus ?thesis
+    by (simp add: LIMSEQ_const)
+qed
+
+lemma LIMSEQ_setprod:
+  fixes L :: "'a \<Rightarrow> 'b::{real_normed_algebra,comm_ring_1}"
+  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
+  shows "(\<lambda>m. \<Prod>n\<in>S. X n m) ----> (\<Prod>n\<in>S. L n)"
+proof (cases "finite S")
+  case True
+  thus ?thesis using n
+  proof (induct)
+    case empty
+    show ?case
+      by (simp add: LIMSEQ_const)
+  next
+    case insert
+    thus ?case
+      by (simp add: LIMSEQ_mult)
+  qed
+next
+  case False
+  thus ?thesis
+    by (simp add: setprod_def LIMSEQ_const)
+qed
+
+lemma LIMSEQ_add_const: "f ----> a ==> (%n.(f n + b)) ----> a + b"
+by (simp add: LIMSEQ_add LIMSEQ_const)
+
+(* FIXME: delete *)
+lemma LIMSEQ_add_minus:
+     "[| X ----> a; Y ----> b |] ==> (%n. X n + -Y n) ----> a + -b"
+by (simp only: LIMSEQ_add LIMSEQ_minus)
+
+lemma LIMSEQ_diff_const: "f ----> a ==> (%n.(f n  - b)) ----> a - b"
+by (simp add: LIMSEQ_diff LIMSEQ_const)
+
+lemma LIMSEQ_diff_approach_zero: 
+  "g ----> L ==> (%x. f x - g x) ----> 0  ==>
+     f ----> L"
+  apply (drule LIMSEQ_add)
+  apply assumption
+  apply simp
+done
+
+lemma LIMSEQ_diff_approach_zero2: 
+  "f ----> L ==> (%x. f x - g x) ----> 0  ==>
+     g ----> L";
+  apply (drule LIMSEQ_diff)
+  apply assumption
+  apply simp
+done
+
+text{*A sequence tends to zero iff its abs does*}
+lemma LIMSEQ_norm_zero: "((\<lambda>n. norm (X n)) ----> 0) = (X ----> 0)"
+by (simp add: LIMSEQ_def)
+
+lemma LIMSEQ_rabs_zero: "((%n. \<bar>f n\<bar>) ----> 0) = (f ----> (0::real))"
+by (simp add: LIMSEQ_def)
+
+lemma LIMSEQ_imp_rabs: "f ----> (l::real) ==> (%n. \<bar>f n\<bar>) ----> \<bar>l\<bar>"
+by (drule LIMSEQ_norm, simp)
+
+text{*An unbounded sequence's inverse tends to 0*}
+
+lemma LIMSEQ_inverse_zero:
+  "\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
+apply (rule LIMSEQ_I)
+apply (drule_tac x="inverse r" in spec, safe)
+apply (rule_tac x="N" in exI, safe)
+apply (drule_tac x="n" in spec, safe)
+apply (frule positive_imp_inverse_positive)
+apply (frule (1) less_imp_inverse_less)
+apply (subgoal_tac "0 < X n", simp)
+apply (erule (1) order_less_trans)
+done
+
+text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}
+
+lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
+apply (rule LIMSEQ_inverse_zero, safe)
+apply (cut_tac x = r in reals_Archimedean2)
+apply (safe, rule_tac x = n in exI)
+apply (auto simp add: real_of_nat_Suc)
+done
+
+text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
+infinity is now easily proved*}
+
+lemma LIMSEQ_inverse_real_of_nat_add:
+     "(%n. r + inverse(real(Suc n))) ----> r"
+by (cut_tac LIMSEQ_add [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
+
+lemma LIMSEQ_inverse_real_of_nat_add_minus:
+     "(%n. r + -inverse(real(Suc n))) ----> r"
+by (cut_tac LIMSEQ_add_minus [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)
+
+lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
+     "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
+by (cut_tac b=1 in
+        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat_add_minus], auto)
+
+lemma LIMSEQ_le_const:
+  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. a \<le> X n\<rbrakk> \<Longrightarrow> a \<le> x"
+apply (rule ccontr, simp only: linorder_not_le)
+apply (drule_tac r="a - x" in LIMSEQ_D, simp)
+apply clarsimp
+apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI1)
+apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI2)
+apply simp
+done
+
+lemma LIMSEQ_le_const2:
+  "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. X n \<le> a\<rbrakk> \<Longrightarrow> x \<le> a"
+apply (subgoal_tac "- a \<le> - x", simp)
+apply (rule LIMSEQ_le_const)
+apply (erule LIMSEQ_minus)
+apply simp
+done
+
+lemma LIMSEQ_le:
+  "\<lbrakk>X ----> x; Y ----> y; \<exists>N. \<forall>n\<ge>N. X n \<le> Y n\<rbrakk> \<Longrightarrow> x \<le> (y::real)"
+apply (subgoal_tac "0 \<le> y - x", simp)
+apply (rule LIMSEQ_le_const)
+apply (erule (1) LIMSEQ_diff)
+apply (simp add: le_diff_eq)
+done
+
+
+subsection {* Convergence *}
+
+lemma limI: "X ----> L ==> lim X = L"
+apply (simp add: lim_def)
+apply (blast intro: LIMSEQ_unique)
+done
+
+lemma convergentD: "convergent X ==> \<exists>L. (X ----> L)"
+by (simp add: convergent_def)
+
+lemma convergentI: "(X ----> L) ==> convergent X"
+by (auto simp add: convergent_def)
+
+lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)"
+by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def)
+
+lemma convergent_minus_iff: "(convergent X) = (convergent (%n. -(X n)))"
+apply (simp add: convergent_def)
+apply (auto dest: LIMSEQ_minus)
+apply (drule LIMSEQ_minus, auto)
+done
+
+
+subsection {* Bounded Monotonic Sequences *}
+
+text{*Subsequence (alternative definition, (e.g. Hoskins)*}
+
+lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))"
+apply (simp add: subseq_def)
+apply (auto dest!: less_imp_Suc_add)
+apply (induct_tac k)
+apply (auto intro: less_trans)
+done
+
+lemma monoseq_Suc:
+   "monoseq X = ((\<forall>n. X n \<le> X (Suc n))
+                 | (\<forall>n. X (Suc n) \<le> X n))"
+apply (simp add: monoseq_def)
+apply (auto dest!: le_imp_less_or_eq)
+apply (auto intro!: lessI [THEN less_imp_le] dest!: less_imp_Suc_add)
+apply (induct_tac "ka")
+apply (auto intro: order_trans)
+apply (erule contrapos_np)
+apply (induct_tac "k")
+apply (auto intro: order_trans)
+done
+
+lemma monoI1: "\<forall>m. \<forall> n \<ge> m. X m \<le> X n ==> monoseq X"
+by (simp add: monoseq_def)
+
+lemma monoI2: "\<forall>m. \<forall> n \<ge> m. X n \<le> X m ==> monoseq X"
+by (simp add: monoseq_def)
+
+lemma mono_SucI1: "\<forall>n. X n \<le> X (Suc n) ==> monoseq X"
+by (simp add: monoseq_Suc)
+
+lemma mono_SucI2: "\<forall>n. X (Suc n) \<le> X n ==> monoseq X"
+by (simp add: monoseq_Suc)
+
+text{*Bounded Sequence*}
+
+lemma BseqD: "Bseq X ==> \<exists>K. 0 < K & (\<forall>n. norm (X n) \<le> K)"
+by (simp add: Bseq_def)
+
+lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
+by (auto simp add: Bseq_def)
+
+lemma lemma_NBseq_def:
+     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) =
+      (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
+apply auto
+ prefer 2 apply force
+apply (cut_tac x = K in reals_Archimedean2, clarify)
+apply (rule_tac x = n in exI, clarify)
+apply (drule_tac x = na in spec)
+apply (auto simp add: real_of_nat_Suc)
+done
+
+text{* alternative definition for Bseq *}
+lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
+apply (simp add: Bseq_def)
+apply (simp (no_asm) add: lemma_NBseq_def)
+done
+
+lemma lemma_NBseq_def2:
+     "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
+apply (subst lemma_NBseq_def, auto)
+apply (rule_tac x = "Suc N" in exI)
+apply (rule_tac [2] x = N in exI)
+apply (auto simp add: real_of_nat_Suc)
+ prefer 2 apply (blast intro: order_less_imp_le)
+apply (drule_tac x = n in spec, simp)
+done
+
+(* yet another definition for Bseq *)
+lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
+by (simp add: Bseq_def lemma_NBseq_def2)
+
+subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
+
+lemma Bseq_isUb:
+  "!!(X::nat=>real). Bseq X ==> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
+by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
+
+
+text{* Use completeness of reals (supremum property)
+   to show that any bounded sequence has a least upper bound*}
+
+lemma Bseq_isLub:
+  "!!(X::nat=>real). Bseq X ==>
+   \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
+by (blast intro: reals_complete Bseq_isUb)
+
+subsubsection{*A Bounded and Monotonic Sequence Converges*}
+
+lemma lemma_converg1:
+     "!!(X::nat=>real). [| \<forall>m. \<forall> n \<ge> m. X m \<le> X n;
+                  isLub (UNIV::real set) {x. \<exists>n. X n = x} (X ma)
+               |] ==> \<forall>n \<ge> ma. X n = X ma"
+apply safe
+apply (drule_tac y = "X n" in isLubD2)
+apply (blast dest: order_antisym)+
+done
+
+text{* The best of both worlds: Easier to prove this result as a standard
+   theorem and then use equivalence to "transfer" it into the
+   equivalent nonstandard form if needed!*}
+
+lemma Bmonoseq_LIMSEQ: "\<forall>n. m \<le> n --> X n = X m ==> \<exists>L. (X ----> L)"
+apply (simp add: LIMSEQ_def)
+apply (rule_tac x = "X m" in exI, safe)
+apply (rule_tac x = m in exI, safe)
+apply (drule spec, erule impE, auto)
+done
+
+lemma lemma_converg2:
+   "!!(X::nat=>real).
+    [| \<forall>m. X m ~= U;  isLub UNIV {x. \<exists>n. X n = x} U |] ==> \<forall>m. X m < U"
+apply safe
+apply (drule_tac y = "X m" in isLubD2)
+apply (auto dest!: order_le_imp_less_or_eq)
+done
+
+lemma lemma_converg3: "!!(X ::nat=>real). \<forall>m. X m \<le> U ==> isUb UNIV {x. \<exists>n. X n = x} U"
+by (rule setleI [THEN isUbI], auto)
+
+text{* FIXME: @{term "U - T < U"} is redundant *}
+lemma lemma_converg4: "!!(X::nat=> real).
+               [| \<forall>m. X m ~= U;
+                  isLub UNIV {x. \<exists>n. X n = x} U;
+                  0 < T;
+                  U + - T < U
+               |] ==> \<exists>m. U + -T < X m & X m < U"
+apply (drule lemma_converg2, assumption)
+apply (rule ccontr, simp)
+apply (simp add: linorder_not_less)
+apply (drule lemma_converg3)
+apply (drule isLub_le_isUb, assumption)
+apply (auto dest: order_less_le_trans)
+done
+
+text{*A standard proof of the theorem for monotone increasing sequence*}
+
+lemma Bseq_mono_convergent:
+     "[| Bseq X; \<forall>m. \<forall>n \<ge> m. X m \<le> X n |] ==> convergent (X::nat=>real)"
+apply (simp add: convergent_def)
+apply (frule Bseq_isLub, safe)
+apply (case_tac "\<exists>m. X m = U", auto)
+apply (blast dest: lemma_converg1 Bmonoseq_LIMSEQ)
+(* second case *)
+apply (rule_tac x = U in exI)
+apply (subst LIMSEQ_iff, safe)
+apply (frule lemma_converg2, assumption)
+apply (drule lemma_converg4, auto)
+apply (rule_tac x = m in exI, safe)
+apply (subgoal_tac "X m \<le> X n")
+ prefer 2 apply blast
+apply (drule_tac x=n and P="%m. X m < U" in spec, arith)
+done
+
+lemma Bseq_minus_iff: "Bseq (%n. -(X n)) = Bseq X"
+by (simp add: Bseq_def)
+
+text{*Main monotonicity theorem*}
+lemma Bseq_monoseq_convergent: "[| Bseq X; monoseq X |] ==> convergent X"
+apply (simp add: monoseq_def, safe)
+apply (rule_tac [2] convergent_minus_iff [THEN ssubst])
+apply (drule_tac [2] Bseq_minus_iff [THEN ssubst])
+apply (auto intro!: Bseq_mono_convergent)
+done
+
+subsubsection{*A Few More Equivalence Theorems for Boundedness*}
+
+text{*alternative formulation for boundedness*}
+lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
+apply (unfold Bseq_def, safe)
+apply (rule_tac [2] x = "k + norm x" in exI)
+apply (rule_tac x = K in exI, simp)
+apply (rule exI [where x = 0], auto)
+apply (erule order_less_le_trans, simp)
+apply (drule_tac x=n in spec, fold diff_def)
+apply (drule order_trans [OF norm_triangle_ineq2])
+apply simp
+done
+
+text{*alternative formulation for boundedness*}
+lemma Bseq_iff3: "Bseq X = (\<exists>k > 0. \<exists>N. \<forall>n. norm(X(n) + -X(N)) \<le> k)"
+apply safe
+apply (simp add: Bseq_def, safe)
+apply (rule_tac x = "K + norm (X N)" in exI)
+apply auto
+apply (erule order_less_le_trans, simp)
+apply (rule_tac x = N in exI, safe)
+apply (drule_tac x = n in spec)
+apply (rule order_trans [OF norm_triangle_ineq], simp)
+apply (auto simp add: Bseq_iff2)
+done
+
+lemma BseqI2: "(\<forall>n. k \<le> f n & f n \<le> (K::real)) ==> Bseq f"
+apply (simp add: Bseq_def)
+apply (rule_tac x = " (\<bar>k\<bar> + \<bar>K\<bar>) + 1" in exI, auto)
+apply (drule_tac x = n in spec, arith)
+done
+
+
+subsection {* Cauchy Sequences *}
+
+lemma CauchyI:
+  "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e) \<Longrightarrow> Cauchy X"
+by (simp add: Cauchy_def)
+
+lemma CauchyD:
+  "\<lbrakk>Cauchy X; 0 < e\<rbrakk> \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < e"
+by (simp add: Cauchy_def)
+
+subsubsection {* Cauchy Sequences are Bounded *}
+
+text{*A Cauchy sequence is bounded -- this is the standard
+  proof mechanization rather than the nonstandard proof*}
+
+lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
+          ==>  \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
+apply (clarify, drule spec, drule (1) mp)
+apply (simp only: norm_minus_commute)
+apply (drule order_le_less_trans [OF norm_triangle_ineq2])
+apply simp
+done
+
+lemma Cauchy_Bseq: "Cauchy X ==> Bseq X"
+apply (simp add: Cauchy_def)
+apply (drule spec, drule mp, rule zero_less_one, safe)
+apply (drule_tac x="M" in spec, simp)
+apply (drule lemmaCauchy)
+apply (rule_tac k="M" in Bseq_offset)
+apply (simp add: Bseq_def)
+apply (rule_tac x="1 + norm (X M)" in exI)
+apply (rule conjI, rule order_less_le_trans [OF zero_less_one], simp)
+apply (simp add: order_less_imp_le)
+done
+
+subsubsection {* Cauchy Sequences are Convergent *}
+
+axclass banach \<subseteq> real_normed_vector
+  Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X"
+
+theorem LIMSEQ_imp_Cauchy:
+  assumes X: "X ----> a" shows "Cauchy X"
+proof (rule CauchyI)
+  fix e::real assume "0 < e"
+  hence "0 < e/2" by simp
+  with X have "\<exists>N. \<forall>n\<ge>N. norm (X n - a) < e/2" by (rule LIMSEQ_D)
+  then obtain N where N: "\<forall>n\<ge>N. norm (X n - a) < e/2" ..
+  show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < e"
+  proof (intro exI allI impI)
+    fix m assume "N \<le> m"
+    hence m: "norm (X m - a) < e/2" using N by fast
+    fix n assume "N \<le> n"
+    hence n: "norm (X n - a) < e/2" using N by fast
+    have "norm (X m - X n) = norm ((X m - a) - (X n - a))" by simp
+    also have "\<dots> \<le> norm (X m - a) + norm (X n - a)"
+      by (rule norm_triangle_ineq4)
+    also from m n have "\<dots> < e" by(simp add:field_simps)
+    finally show "norm (X m - X n) < e" .
+  qed
+qed
+
+lemma convergent_Cauchy: "convergent X \<Longrightarrow> Cauchy X"
+unfolding convergent_def
+by (erule exE, erule LIMSEQ_imp_Cauchy)
+
+text {*
+Proof that Cauchy sequences converge based on the one from
+http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html
+*}
+
+text {*
+  If sequence @{term "X"} is Cauchy, then its limit is the lub of
+  @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
+*}
+
+lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u"
+by (simp add: isUbI setleI)
+
+lemma real_abs_diff_less_iff:
+  "(\<bar>x - a\<bar> < (r::real)) = (a - r < x \<and> x < a + r)"
+by auto
+
+locale real_Cauchy =
+  fixes X :: "nat \<Rightarrow> real"
+  assumes X: "Cauchy X"
+  fixes S :: "real set"
+  defines S_def: "S \<equiv> {x::real. \<exists>N. \<forall>n\<ge>N. x < X n}"
+
+lemma real_CauchyI:
+  assumes "Cauchy X"
+  shows "real_Cauchy X"
+  proof qed (fact assms)
+
+lemma (in real_Cauchy) mem_S: "\<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S"
+by (unfold S_def, auto)
+
+lemma (in real_Cauchy) bound_isUb:
+  assumes N: "\<forall>n\<ge>N. X n < x"
+  shows "isUb UNIV S x"
+proof (rule isUb_UNIV_I)
+  fix y::real assume "y \<in> S"
+  hence "\<exists>M. \<forall>n\<ge>M. y < X n"
+    by (simp add: S_def)
+  then obtain M where "\<forall>n\<ge>M. y < X n" ..
+  hence "y < X (max M N)" by simp
+  also have "\<dots> < x" using N by simp
+  finally show "y \<le> x"
+    by (rule order_less_imp_le)
+qed
+
+lemma (in real_Cauchy) isLub_ex: "\<exists>u. isLub UNIV S u"
+proof (rule reals_complete)
+  obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (X m - X n) < 1"
+    using CauchyD [OF X zero_less_one] by fast
+  hence N: "\<forall>n\<ge>N. norm (X n - X N) < 1" by simp
+  show "\<exists>x. x \<in> S"
+  proof
+    from N have "\<forall>n\<ge>N. X N - 1 < X n"
+      by (simp add: real_abs_diff_less_iff)
+    thus "X N - 1 \<in> S" by (rule mem_S)
+  qed
+  show "\<exists>u. isUb UNIV S u"
+  proof
+    from N have "\<forall>n\<ge>N. X n < X N + 1"
+      by (simp add: real_abs_diff_less_iff)
+    thus "isUb UNIV S (X N + 1)"
+      by (rule bound_isUb)
+  qed
+qed
+
+lemma (in real_Cauchy) isLub_imp_LIMSEQ:
+  assumes x: "isLub UNIV S x"
+  shows "X ----> x"
+proof (rule LIMSEQ_I)
+  fix r::real assume "0 < r"
+  hence r: "0 < r/2" by simp
+  obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. norm (X n - X m) < r/2"
+    using CauchyD [OF X r] by fast
+  hence "\<forall>n\<ge>N. norm (X n - X N) < r/2" by simp
+  hence N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2"
+    by (simp only: real_norm_def real_abs_diff_less_iff)
+
+  from N have "\<forall>n\<ge>N. X N - r/2 < X n" by fast
+  hence "X N - r/2 \<in> S" by (rule mem_S)
+  hence 1: "X N - r/2 \<le> x" using x isLub_isUb isUbD by fast
+
+  from N have "\<forall>n\<ge>N. X n < X N + r/2" by fast
+  hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb)
+  hence 2: "x \<le> X N + r/2" using x isLub_le_isUb by fast
+
+  show "\<exists>N. \<forall>n\<ge>N. norm (X n - x) < r"
+  proof (intro exI allI impI)
+    fix n assume n: "N \<le> n"
+    from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
+    thus "norm (X n - x) < r" using 1 2
+      by (simp add: real_abs_diff_less_iff)
+  qed
+qed
+
+lemma (in real_Cauchy) LIMSEQ_ex: "\<exists>x. X ----> x"
+proof -
+  obtain x where "isLub UNIV S x"
+    using isLub_ex by fast
+  hence "X ----> x"
+    by (rule isLub_imp_LIMSEQ)
+  thus ?thesis ..
+qed
+
+lemma real_Cauchy_convergent:
+  fixes X :: "nat \<Rightarrow> real"
+  shows "Cauchy X \<Longrightarrow> convergent X"
+unfolding convergent_def
+by (rule real_Cauchy.LIMSEQ_ex)
+ (rule real_CauchyI)
+
+instance real :: banach
+by intro_classes (rule real_Cauchy_convergent)
+
+lemma Cauchy_convergent_iff:
+  fixes X :: "nat \<Rightarrow> 'a::banach"
+  shows "Cauchy X = convergent X"
+by (fast intro: Cauchy_convergent convergent_Cauchy)
+
+
+subsection {* Power Sequences *}
+
+text{*The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
+"x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
+  also fact that bounded and monotonic sequence converges.*}
+
+lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
+apply (simp add: Bseq_def)
+apply (rule_tac x = 1 in exI)
+apply (simp add: power_abs)
+apply (auto dest: power_mono)
+done
+
+lemma monoseq_realpow: "[| 0 \<le> x; x \<le> 1 |] ==> monoseq (%n. x ^ n)"
+apply (clarify intro!: mono_SucI2)
+apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
+done
+
+lemma convergent_realpow:
+  "[| 0 \<le> (x::real); x \<le> 1 |] ==> convergent (%n. x ^ n)"
+by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)
+
+lemma LIMSEQ_inverse_realpow_zero_lemma:
+  fixes x :: real
+  assumes x: "0 \<le> x"
+  shows "real n * x + 1 \<le> (x + 1) ^ n"
+apply (induct n)
+apply simp
+apply simp
+apply (rule order_trans)
+prefer 2
+apply (erule mult_left_mono)
+apply (rule add_increasing [OF x], simp)
+apply (simp add: real_of_nat_Suc)
+apply (simp add: ring_distribs)
+apply (simp add: mult_nonneg_nonneg x)
+done
+
+lemma LIMSEQ_inverse_realpow_zero:
+  "1 < (x::real) \<Longrightarrow> (\<lambda>n. inverse (x ^ n)) ----> 0"
+proof (rule LIMSEQ_inverse_zero [rule_format])
+  fix y :: real
+  assume x: "1 < x"
+  hence "0 < x - 1" by simp
+  hence "\<forall>y. \<exists>N::nat. y < real N * (x - 1)"
+    by (rule reals_Archimedean3)
+  hence "\<exists>N::nat. y < real N * (x - 1)" ..
+  then obtain N::nat where "y < real N * (x - 1)" ..
+  also have "\<dots> \<le> real N * (x - 1) + 1" by simp
+  also have "\<dots> \<le> (x - 1 + 1) ^ N"
+    by (rule LIMSEQ_inverse_realpow_zero_lemma, cut_tac x, simp)
+  also have "\<dots> = x ^ N" by simp
+  finally have "y < x ^ N" .
+  hence "\<forall>n\<ge>N. y < x ^ n"
+    apply clarify
+    apply (erule order_less_le_trans)
+    apply (erule power_increasing)
+    apply (rule order_less_imp_le [OF x])
+    done
+  thus "\<exists>N. \<forall>n\<ge>N. y < x ^ n" ..
+qed
+
+lemma LIMSEQ_realpow_zero:
+  "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
+proof (cases)
+  assume "x = 0"
+  hence "(\<lambda>n. x ^ Suc n) ----> 0" by (simp add: LIMSEQ_const)
+  thus ?thesis by (rule LIMSEQ_imp_Suc)
+next
+  assume "0 \<le> x" and "x \<noteq> 0"
+  hence x0: "0 < x" by simp
+  assume x1: "x < 1"
+  from x0 x1 have "1 < inverse x"
+    by (rule real_inverse_gt_one)
+  hence "(\<lambda>n. inverse (inverse x ^ n)) ----> 0"
+    by (rule LIMSEQ_inverse_realpow_zero)
+  thus ?thesis by (simp add: power_inverse)
+qed
+
+lemma LIMSEQ_power_zero:
+  fixes x :: "'a::{real_normed_algebra_1,recpower}"
+  shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
+apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
+apply (simp only: LIMSEQ_Zseq_iff, erule Zseq_le)
+apply (simp add: power_abs norm_power_ineq)
+done
+
+lemma LIMSEQ_divide_realpow_zero:
+  "1 < (x::real) ==> (%n. a / (x ^ n)) ----> 0"
+apply (cut_tac a = a and x1 = "inverse x" in
+        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_realpow_zero])
+apply (auto simp add: divide_inverse power_inverse)
+apply (simp add: inverse_eq_divide pos_divide_less_eq)
+done
+
+text{*Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}*}
+
+lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < (1::real) ==> (%n. \<bar>c\<bar> ^ n) ----> 0"
+by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
+
+lemma LIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < (1::real) ==> (%n. c ^ n) ----> 0"
+apply (rule LIMSEQ_rabs_zero [THEN iffD1])
+apply (auto intro: LIMSEQ_rabs_realpow_zero simp add: power_abs)
+done
+
+end
--- a/src/HOL/Series.thy	Mon Dec 29 13:23:53 2008 +0100
+++ b/src/HOL/Series.thy	Mon Dec 29 14:08:08 2008 +0100
@@ -10,7 +10,7 @@
 header{*Finite Summation and Infinite Series*}
 
 theory Series
-imports "~~/src/HOL/Hyperreal/SEQ"
+imports SEQ
 begin
 
 definition