simplified "sos" method;
authorwenzelm
Wed, 08 Oct 2014 11:09:17 +0200
changeset 58630 71cdb885b3bb
parent 58629 a6a6cd499d4e
child 58631 41333b45bff9
simplified "sos" method;
NEWS
src/HOL/Library/Sum_of_Squares.thy
src/HOL/Library/Sum_of_Squares/etc/settings
src/HOL/Library/Sum_of_Squares/neos_csdp_client
src/HOL/Library/Sum_of_Squares/sos_wrapper.ML
src/HOL/ROOT
src/HOL/ex/SOS.thy
src/HOL/ex/SOS_Cert.thy
src/HOL/ex/SOS_Remote.thy
--- a/NEWS	Wed Oct 08 10:15:04 2014 +0200
+++ b/NEWS	Wed Oct 08 11:09:17 2014 +0200
@@ -126,6 +126,12 @@
   generated code in target languages in HOL/Library/Code_Test. See
   HOL/Codegenerator_Test/Code_Test* for examples.
 
+* Library/Sum_of_Squares: simplified and improved "sos" method. Always
+use local CSDP executable, which is much faster than the NEOS server.
+The "sos_cert" functionality is invoked as "sos" with additional
+argument. Minor INCOMPATIBILITY.
+
+
 *** ML ***
 
 * Tactical PARALLEL_ALLGOALS is the most common way to refer to
--- a/src/HOL/Library/Sum_of_Squares.thy	Wed Oct 08 10:15:04 2014 +0200
+++ b/src/HOL/Library/Sum_of_Squares.thy	Wed Oct 08 11:09:17 2014 +0200
@@ -3,8 +3,8 @@
     Author:     Philipp Meyer, TU Muenchen
 *)
 
-header {* A decision method for universal multivariate real arithmetic with addition, 
-  multiplication and ordering using semidefinite programming *}
+header {* A decision procedure for universal multivariate real arithmetic
+  with addition, multiplication and ordering using semidefinite programming *}
 
 theory Sum_of_Squares
 imports Complex_Main
@@ -15,27 +15,4 @@
 ML_file "Sum_of_Squares/positivstellensatz_tools.ML"
 ML_file "Sum_of_Squares/sos_wrapper.ML"
 
-text {*
-  Proof method sos invocations:
-  \begin{itemize}
-
-  \item remote solver: @{text "(sos remote_csdp)"}
-
-  \item local solver: @{text "(sos csdp)"}
-
-  The latter requires a local executable from
-  @{url "https://projects.coin-or.org/Csdp"} and the Isabelle settings variable
-  variable @{text ISABELLE_CSDP} pointing to it.
-
-  \end{itemize}
-
-  By default, method sos calls @{text remote_csdp}.  This can take of
-  the order of a minute for one sos call, because sos calls CSDP
-  repeatedly.  If you install CSDP locally, sos calls typically takes
-  only a few seconds.
-
-  The sos method generates a certificate which can be used to repeat
-  the proof without calling an external prover.
-*}
-
 end
--- a/src/HOL/Library/Sum_of_Squares/etc/settings	Wed Oct 08 10:15:04 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-# -*- shell-script -*- :mode=shellscript:
-
-ISABELLE_SUM_OF_SQUARES="$COMPONENT"
-ISABELLE_NEOS_SERVER="http://neos-server.org:3332"
-
-# local SDP Solver, cf. https://projects.coin-or.org/Csdp
-#ISABELLE_CSDP="/usr/local/bin/csdp"
-
--- a/src/HOL/Library/Sum_of_Squares/neos_csdp_client	Wed Oct 08 10:15:04 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,85 +0,0 @@
-#!/usr/bin/env python
-import sys
-import signal
-import xmlrpclib
-import time
-import re
-import os
-
-# Neos server config
-neos = xmlrpclib.Server(os.getenv("ISABELLE_NEOS_SERVER"))
-
-jobNumber = 0
-password = ""
-inputfile = None
-outputfile = None
-# interrupt handler
-def cleanup(signal, frame):
-  sys.stdout.write("Caught interrupt, cleaning up\n")
-  if jobNumber != 0:
-    neos.killJob(jobNumber, password)
-  if inputfile != None:
-    inputfile.close()
-  if outputfile != None:
-    outputfile.close()
-  sys.exit(21)
-
-signal.signal(signal.SIGHUP, cleanup)
-signal.signal(signal.SIGINT, cleanup)
-signal.signal(signal.SIGQUIT, cleanup)
-signal.signal(signal.SIGTERM, cleanup)
-
-if len(sys.argv) <> 3:
-  sys.stderr.write("Usage: neos_csdp_client <input_filename> <output_filename>\n")
-  sys.exit(19)
-
-xml_pre = "<document>\n<category>sdp</category>\n<solver>csdp</solver>\n<inputMethod>SPARSE_SDPA</inputMethod>\n<dat><![CDATA["
-xml_post = "]]></dat>\n</document>\n"
-xml = xml_pre
-inputfile = open(sys.argv[1],"r")
-buffer = 1
-while buffer:
-  buffer = inputfile.read()
-  xml += buffer
-inputfile.close()
-xml += xml_post
-
-(jobNumber,password) = neos.submitJob(xml)
-
-if jobNumber == 0:
-  sys.stdout.write("error submitting job: %s" % password)
-  sys.exit(20)
-else:
-  sys.stdout.write("jobNumber = %d\tpassword = %s\n" % (jobNumber,password))
-
-offset=0
-messages = ""
-status="Waiting"
-while status == "Running" or status=="Waiting":
-  time.sleep(1)
-  (msg,offset) = neos.getIntermediateResults(jobNumber,password,offset)
-  messages += msg.data
-  sys.stdout.write(msg.data)
-  status = neos.getJobStatus(jobNumber, password)
-
-msg = neos.getFinalResults(jobNumber, password).data
-sys.stdout.write("---------- Begin CSDP Output -------------\n");
-sys.stdout.write(msg)
-
-# extract solution
-result = msg.split("Solution:")
-if len(result) > 1:
-  solution = result[1].strip()
-  if solution != "":
-    outputfile = open(sys.argv[2],"w")
-    outputfile.write(solution)
-    outputfile.close()
-
-# extract return code
-p = re.compile(r"^Error: Command exited with non-zero status (\d+)$", re.MULTILINE)
-m = p.search(messages)
-if m:
-  sys.exit(int(m.group(1)))
-else:
-  sys.exit(0)
-
--- a/src/HOL/Library/Sum_of_Squares/sos_wrapper.ML	Wed Oct 08 10:15:04 2014 +0200
+++ b/src/HOL/Library/Sum_of_Squares/sos_wrapper.ML	Wed Oct 08 11:09:17 2014 +0200
@@ -6,9 +6,7 @@
 
 signature SOS_WRAPPER =
 sig
-  datatype prover_result = Success | Failure | Error
-  val dest_dir: string Config.T
-  val prover_name: string Config.T
+  val sos_tac: Proof.context -> string option -> int -> tactic
 end
 
 structure SOS_Wrapper: SOS_WRAPPER =
@@ -21,33 +19,33 @@
   | str_of_result Error = "Error"
 
 
-(*** calling provers ***)
+fun filename name =
+  File.tmp_path (Path.basic (name ^ serial_string ()))
 
-val dest_dir = Attrib.setup_config_string @{binding sos_dest_dir} (K "")
+fun find_failure rv =
+  case rv of
+    0 => (Success, "SDP solved")
+  | 1 => (Failure, "SDP is primal infeasible")
+  | 2 => (Failure, "SDP is dual infeasible")
+  | 3 => (Success, "SDP solved with reduced accuracy")
+  | 4 => (Failure, "Maximum iterations reached")
+  | 5 => (Failure, "Stuck at edge of primal feasibility")
+  | 6 => (Failure, "Stuck at edge of dual infeasibility")
+  | 7 => (Failure, "Lack of progress")
+  | 8 => (Failure, "X, Z, or O was singular")
+  | 9 => (Failure, "Detected NaN or Inf values")
+  | _ => (Error, "return code is " ^ string_of_int rv)
 
-fun filename dir name =
+val exe = Path.explode "$ISABELLE_CSDP"
+
+fun run_solver ctxt input =
   let
-    val probfile = Path.basic (name ^ serial_string ())
-    val dir_path = Path.explode dir
-  in
-    if dir = "" then
-      File.tmp_path probfile
-    else if File.exists dir_path then
-      Path.append dir_path probfile
-    else error ("No such directory: " ^ dir)
-  end
-
-fun run_solver ctxt name exe find_failure input =
-  let
-    val _ = warning ("Calling solver: " ^ name)
-
     (* create input file *)
-    val dir = Config.get ctxt dest_dir
-    val input_file = filename dir "sos_in"
+    val input_file = filename "sos_in"
     val _ = File.write input_file input
 
     (* call solver *)
-    val output_file = filename dir "sos_out"
+    val output_file = filename "sos_out"
     val (output, rv) =
       Isabelle_System.bash_output
        (if File.exists exe then
@@ -59,10 +57,8 @@
     val result = if File.exists output_file then File.read output_file else ""
 
     (* remove temporary files *)
-    val _ =
-      if dir = "" then
-        (File.rm input_file; if File.exists output_file then File.rm output_file else ())
-      else ()
+    val _ = File.rm input_file
+    val _ = if File.exists output_file then File.rm output_file else ()
 
     val _ =
       if Config.get ctxt Sum_of_Squares.trace
@@ -78,78 +74,24 @@
   end
 
 
-(*** various provers ***)
-
-(* local csdp client *)
-
-fun find_csdp_failure rv =
-  case rv of
-    0 => (Success, "SDP solved")
-  | 1 => (Failure, "SDP is primal infeasible")
-  | 2 => (Failure, "SDP is dual infeasible")
-  | 3 => (Success, "SDP solved with reduced accuracy")
-  | 4 => (Failure, "Maximum iterations reached")
-  | 5 => (Failure, "Stuck at edge of primal feasibility")
-  | 6 => (Failure, "Stuck at edge of dual infeasibility")
-  | 7 => (Failure, "Lack of progress")
-  | 8 => (Failure, "X, Z, or O was singular")
-  | 9 => (Failure, "Detected NaN or Inf values")
-  | _ => (Error, "return code is " ^ string_of_int rv)
-
-val csdp = (Path.explode "$ISABELLE_CSDP", find_csdp_failure)
-
-
-(* remote neos server *)
-
-fun find_neos_failure rv =
-  case rv of
-    20 => (Error, "error submitting job")
-  | 21 => (Error, "interrupt")
-  |  _ => find_csdp_failure rv
-
-val neos_csdp = (Path.explode "$ISABELLE_SUM_OF_SQUARES/neos_csdp_client", find_neos_failure)
-
-
-(* named provers *)
-
-fun prover "remote_csdp" = neos_csdp
-  | prover "csdp" = csdp
-  | prover name = error ("Unknown prover: " ^ name)
-
-val prover_name = Attrib.setup_config_string @{binding sos_prover_name} (K "remote_csdp")
-
-fun call_solver ctxt opt_name =
-  let
-    val name = the_default (Config.get ctxt prover_name) opt_name
-    val (exe, find_failure) = prover name
-  in run_solver ctxt name exe find_failure end
-
-
-(* certificate output *)
-
-fun output_line cert =
-  "To repeat this proof with a certificate use this command:\n" ^
-    Active.sendback_markup [] ("apply (sos_cert \"" ^ cert ^ "\")")
-
-val print_cert = warning o output_line o Positivstellensatz_Tools.print_cert
-
-
 (* method setup *)
 
-fun sos_solver print method = SIMPLE_METHOD' o Sum_of_Squares.sos_tac print method
+fun print_cert cert =
+  warning
+    ("To repeat this proof with a certificate use this proof method:\n" ^
+      Active.sendback_markup [] ("(sos \"" ^ Positivstellensatz_Tools.print_cert cert ^ "\")"))
+
+fun sos_tac ctxt NONE =
+      Sum_of_Squares.sos_tac print_cert
+        (Sum_of_Squares.Prover (run_solver ctxt)) ctxt
+  | sos_tac ctxt (SOME cert) =
+      Sum_of_Squares.sos_tac ignore
+        (Sum_of_Squares.Certificate (Positivstellensatz_Tools.read_cert ctxt cert)) ctxt
 
 val _ = Theory.setup
  (Method.setup @{binding sos}
-    (Scan.lift (Scan.option Parse.xname)
-      >> (fn opt_name => fn ctxt =>
-        sos_solver print_cert
-          (Sum_of_Squares.Prover (call_solver ctxt opt_name)) ctxt))
-    "prove universal problems over the reals using sums of squares" #>
-  Method.setup @{binding sos_cert}
-    (Scan.lift Parse.string
-      >> (fn cert => fn ctxt =>
-        sos_solver ignore
-          (Sum_of_Squares.Certificate (Positivstellensatz_Tools.read_cert ctxt cert)) ctxt))
-    "prove universal problems over the reals using sums of squares with certificates")
+    (Scan.lift (Scan.option Parse.string)
+      >> (fn arg => fn ctxt => SIMPLE_METHOD' (sos_tac ctxt arg)))
+    "prove universal problems over the reals using sums of squares")
 
 end
--- a/src/HOL/ROOT	Wed Oct 08 10:15:04 2014 +0200
+++ b/src/HOL/ROOT	Wed Oct 08 11:09:17 2014 +0200
@@ -600,11 +600,8 @@
     ML
     SAT_Examples
     Nominal2_Dummy
+    SOS
     SOS_Cert
-  theories [condition = ISABELLE_CSDP]
-    SOS
-  theories [condition = ISABELLE_FULL_TEST]
-    SOS_Remote
   theories [skip_proofs = false]
     Meson_Test
   theories [condition = SVC_HOME]
--- a/src/HOL/ex/SOS.thy	Wed Oct 08 10:15:04 2014 +0200
+++ b/src/HOL/ex/SOS.thy	Wed Oct 08 11:09:17 2014 +0200
@@ -10,121 +10,121 @@
 begin
 
 lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x \<Longrightarrow> a < 0"
-  by (sos csdp)
+  by sos
 
 lemma "a1 >= 0 & a2 >= 0 \<and> (a1 * a1 + a2 * a2 = b1 * b1 + b2 * b2 + 2) \<and> (a1 * b1 + a2 * b2 = 0) --> a1 * a2 - b1 * b2 >= (0::real)"
-  by (sos csdp)
+  by sos
 
 lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x --> a < 0"
-  by (sos csdp)
+  by sos
 
 lemma "(0::real) <= x & x <= 1 & 0 <= y & y <= 1  --> x^2 + y^2 < 1 |(x - 1)^2 + y^2 < 1 | x^2 + (y - 1)^2 < 1 | (x - 1)^2 + (y - 1)^2 < 1"
-  by (sos csdp)
+  by sos
 
 lemma "(0::real) <= x & 0 <= y & 0 <= z & x + y + z <= 3 --> x * y + x * z + y * z >= 3 * x * y * z"
-  by (sos csdp)
+  by sos
 
 lemma "((x::real)^2 + y^2 + z^2 = 1) --> (x + y + z)^2 <= 3"
-  by (sos csdp)
+  by sos
 
 lemma "(w^2 + x^2 + y^2 + z^2 = 1) --> (w + x + y + z)^2 <= (4::real)"
-  by (sos csdp)
+  by sos
 
 lemma "(x::real) >= 1 & y >= 1 --> x * y >= x + y - 1"
-  by (sos csdp)
+  by sos
 
 lemma "(x::real) > 1 & y > 1 --> x * y > x + y - 1"
-  by (sos csdp)
+  by sos
 
 lemma "abs(x) <= 1 --> abs(64 * x^7 - 112 * x^5 + 56 * x^3 - 7 * x) <= (1::real)"
-  by (sos csdp)
+  by sos
 
 
 text \<open>One component of denominator in dodecahedral example.\<close>
 
 lemma "2 <= x & x <= 125841 / 50000 & 2 <= y & y <= 125841 / 50000 & 2 <= z & z <= 125841 / 50000 --> 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z) >= (0::real)"
-  by (sos csdp)
+  by sos
 
 
 text \<open>Over a larger but simpler interval.\<close>
 
 lemma "(2::real) <= x & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 --> 0 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)"
-  by (sos csdp)
+  by sos
 
 
 text \<open>We can do 12. I think 12 is a sharp bound; see PP's certificate.\<close>
 
 lemma "2 <= (x::real) & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 --> 12 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)"
-  by (sos csdp)
+  by sos
 
 
 text \<open>Inequality from sci.math (see "Leon-Sotelo, por favor").\<close>
 
 lemma "0 <= (x::real) & 0 <= y & (x * y = 1) --> x + y <= x^2 + y^2"
-  by (sos csdp)
+  by sos
 
 lemma "0 <= (x::real) & 0 <= y & (x * y = 1) --> x * y * (x + y) <= x^2 + y^2"
-  by (sos csdp)
+  by sos
 
 lemma "0 <= (x::real) & 0 <= y --> x * y * (x + y)^2 <= (x^2 + y^2)^2"
-  by (sos csdp)
+  by sos
 
 lemma "(0::real) <= a & 0 <= b & 0 <= c & c * (2 * a + b)^3/ 27 <= x \<longrightarrow> c * a^2 * b <= x"
-  by (sos csdp)
+  by sos
 
 lemma "(0::real) < x --> 0 < 1 + x + x^2"
-  by (sos csdp)
+  by sos
 
 lemma "(0::real) <= x --> 0 < 1 + x + x^2"
-  by (sos csdp)
+  by sos
 
 lemma "(0::real) < 1 + x^2"
-  by (sos csdp)
+  by sos
 
 lemma "(0::real) <= 1 + 2 * x + x^2"
-  by (sos csdp)
+  by sos
 
 lemma "(0::real) < 1 + abs x"
-  by (sos csdp)
+  by sos
 
 lemma "(0::real) < 1 + (1 + x)^2 * (abs x)"
-  by (sos csdp)
+  by sos
 
 
 lemma "abs ((1::real) + x^2) = (1::real) + x^2"
-  by (sos csdp)
+  by sos
 lemma "(3::real) * x + 7 * a < 4 \<and> 3 < 2 * x \<longrightarrow> a < 0"
-  by (sos csdp)
+  by sos
 
 lemma "(0::real) < x --> 1 < y --> y * x <= z --> x < z"
-  by (sos csdp)
+  by sos
 lemma "(1::real) < x --> x^2 < y --> 1 < y"
-  by (sos csdp)
+  by sos
 lemma "(b::real)^2 < 4 * a * c --> ~(a * x^2 + b * x + c = 0)"
-  by (sos csdp)
+  by sos
 lemma "(b::real)^2 < 4 * a * c --> ~(a * x^2 + b * x + c = 0)"
-  by (sos csdp)
+  by sos
 lemma "((a::real) * x^2 + b * x + c = 0) --> b^2 >= 4 * a * c"
-  by (sos csdp)
+  by sos
 lemma "(0::real) <= b & 0 <= c & 0 <= x & 0 <= y & (x^2 = c) & (y^2 = a^2 * c + b) --> a * c <= y * x"
-  by (sos csdp)
+  by sos
 lemma "abs(x - z) <= e & abs(y - z) <= e & 0 <= u & 0 <= v & (u + v = 1) --> abs((u * x + v * y) - z) <= (e::real)"
-  by (sos csdp)
+  by sos
 
 
 (* lemma "((x::real) - y - 2 * x^4 = 0) & 0 <= x & x <= 2 & 0 <= y & y <= 3 --> y^2 - 7 * y - 12 * x + 17 >= 0" by sos *) (* Too hard?*)
 
 lemma "(0::real) <= x --> (1 + x + x^2)/(1 + x^2) <= 1 + x"
-  by (sos csdp)
+  by sos
 
 lemma "(0::real) <= x --> 1 - x <= 1 / (1 + x + x^2)"
-  by (sos csdp)
+  by sos
 
 lemma "(x::real) <= 1 / 2 --> - x - 2 * x^2 <= - x / (1 - x)"
-  by (sos csdp)
+  by sos
 
 lemma "4*r^2 = p^2 - 4*q & r >= (0::real) & x^2 + p*x + q = 0 --> 2*(x::real) = - p + 2*r | 2*x = -p - 2*r"
-  by (sos csdp)
+  by sos
 
 end
 
--- a/src/HOL/ex/SOS_Cert.thy	Wed Oct 08 10:15:04 2014 +0200
+++ b/src/HOL/ex/SOS_Cert.thy	Wed Oct 08 11:09:17 2014 +0200
@@ -10,121 +10,121 @@
 begin
 
 lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x \<Longrightarrow> a < 0"
-  by (sos_cert "((R<1 + (((A<1 * R<1) * (R<2 * [1]^2)) + (((A<0 * R<1) * (R<3 * [1]^2)) + ((A<=0 * R<1) * (R<14 * [1]^2))))))")
+  by (sos "((R<1 + (((A<1 * R<1) * (R<2 * [1]^2)) + (((A<0 * R<1) * (R<3 * [1]^2)) + ((A<=0 * R<1) * (R<14 * [1]^2))))))")
 
 lemma "a1 >= 0 & a2 >= 0 \<and> (a1 * a1 + a2 * a2 = b1 * b1 + b2 * b2 + 2) \<and> (a1 * b1 + a2 * b2 = 0) --> a1 * a2 - b1 * b2 >= (0::real)"
-  by (sos_cert "(((A<0 * R<1) + (([~1/2*a1*b2 + ~1/2*a2*b1] * A=0) + (([~1/2*a1*a2 + 1/2*b1*b2] * A=1) + (((A<0 * R<1) * ((R<1/2 * [b2]^2) + (R<1/2 * [b1]^2))) + ((A<=0 * (A<=1 * R<1)) * ((R<1/2 * [b2]^2) + ((R<1/2 * [b1]^2) + ((R<1/2 * [a2]^2) + (R<1/2 * [a1]^2))))))))))")
+  by (sos "(((A<0 * R<1) + (([~1/2*a1*b2 + ~1/2*a2*b1] * A=0) + (([~1/2*a1*a2 + 1/2*b1*b2] * A=1) + (((A<0 * R<1) * ((R<1/2 * [b2]^2) + (R<1/2 * [b1]^2))) + ((A<=0 * (A<=1 * R<1)) * ((R<1/2 * [b2]^2) + ((R<1/2 * [b1]^2) + ((R<1/2 * [a2]^2) + (R<1/2 * [a1]^2))))))))))")
 
 lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x --> a < 0"
-  by (sos_cert "((R<1 + (((A<1 * R<1) * (R<2 * [1]^2)) + (((A<0 * R<1) * (R<3 * [1]^2)) + ((A<=0 * R<1) * (R<14 * [1]^2))))))")
+  by (sos "((R<1 + (((A<1 * R<1) * (R<2 * [1]^2)) + (((A<0 * R<1) * (R<3 * [1]^2)) + ((A<=0 * R<1) * (R<14 * [1]^2))))))")
 
 lemma "(0::real) <= x & x <= 1 & 0 <= y & y <= 1  --> x^2 + y^2 < 1 |(x - 1)^2 + y^2 < 1 | x^2 + (y - 1)^2 < 1 | (x - 1)^2 + (y - 1)^2 < 1"
-  by (sos_cert "((R<1 + (((A<=3 * (A<=4 * R<1)) * (R<1 * [1]^2)) + (((A<=2 * (A<=7 * R<1)) * (R<1 * [1]^2)) + (((A<=1 * (A<=6 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<=5 * R<1)) * (R<1 * [1]^2)))))))")
+  by (sos "((R<1 + (((A<=3 * (A<=4 * R<1)) * (R<1 * [1]^2)) + (((A<=2 * (A<=7 * R<1)) * (R<1 * [1]^2)) + (((A<=1 * (A<=6 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<=5 * R<1)) * (R<1 * [1]^2)))))))")
 
 lemma "(0::real) <= x & 0 <= y & 0 <= z & x + y + z <= 3 --> x * y + x * z + y * z >= 3 * x * y * z"
-  by (sos_cert "(((A<0 * R<1) + (((A<0 * R<1) * (R<1/2 * [1]^2)) + (((A<=2 * R<1) * (R<1/2 * [~1*x + y]^2)) + (((A<=1 * R<1) * (R<1/2 * [~1*x + z]^2)) + (((A<=1 * (A<=2 * (A<=3 * R<1))) * (R<1/2 * [1]^2)) + (((A<=0 * R<1) * (R<1/2 * [~1*y + z]^2)) + (((A<=0 * (A<=2 * (A<=3 * R<1))) * (R<1/2 * [1]^2)) + ((A<=0 * (A<=1 * (A<=3 * R<1))) * (R<1/2 * [1]^2))))))))))")
+  by (sos "(((A<0 * R<1) + (((A<0 * R<1) * (R<1/2 * [1]^2)) + (((A<=2 * R<1) * (R<1/2 * [~1*x + y]^2)) + (((A<=1 * R<1) * (R<1/2 * [~1*x + z]^2)) + (((A<=1 * (A<=2 * (A<=3 * R<1))) * (R<1/2 * [1]^2)) + (((A<=0 * R<1) * (R<1/2 * [~1*y + z]^2)) + (((A<=0 * (A<=2 * (A<=3 * R<1))) * (R<1/2 * [1]^2)) + ((A<=0 * (A<=1 * (A<=3 * R<1))) * (R<1/2 * [1]^2))))))))))")
 
 lemma "((x::real)^2 + y^2 + z^2 = 1) --> (x + y + z)^2 <= 3"
-  by (sos_cert "(((A<0 * R<1) + (([~3] * A=0) + (R<1 * ((R<2 * [~1/2*x + ~1/2*y + z]^2) + (R<3/2 * [~1*x + y]^2))))))")
+  by (sos "(((A<0 * R<1) + (([~3] * A=0) + (R<1 * ((R<2 * [~1/2*x + ~1/2*y + z]^2) + (R<3/2 * [~1*x + y]^2))))))")
 
 lemma "(w^2 + x^2 + y^2 + z^2 = 1) --> (w + x + y + z)^2 <= (4::real)"
-  by (sos_cert "(((A<0 * R<1) + (([~4] * A=0) + (R<1 * ((R<3 * [~1/3*w + ~1/3*x + ~1/3*y + z]^2) + ((R<8/3 * [~1/2*w + ~1/2*x + y]^2) + (R<2 * [~1*w + x]^2)))))))")
+  by (sos "(((A<0 * R<1) + (([~4] * A=0) + (R<1 * ((R<3 * [~1/3*w + ~1/3*x + ~1/3*y + z]^2) + ((R<8/3 * [~1/2*w + ~1/2*x + y]^2) + (R<2 * [~1*w + x]^2)))))))")
 
 lemma "(x::real) >= 1 & y >= 1 --> x * y >= x + y - 1"
-  by (sos_cert "(((A<0 * R<1) + ((A<=0 * (A<=1 * R<1)) * (R<1 * [1]^2))))")
+  by (sos "(((A<0 * R<1) + ((A<=0 * (A<=1 * R<1)) * (R<1 * [1]^2))))")
 
 lemma "(x::real) > 1 & y > 1 --> x * y > x + y - 1"
-  by (sos_cert "((((A<0 * A<1) * R<1) + ((A<=0 * R<1) * (R<1 * [1]^2))))")
+  by (sos "((((A<0 * A<1) * R<1) + ((A<=0 * R<1) * (R<1 * [1]^2))))")
 
 lemma "abs(x) <= 1 --> abs(64 * x^7 - 112 * x^5 + 56 * x^3 - 7 * x) <= (1::real)"
-  by (sos_cert "((((A<0 * R<1) + ((A<=1 * R<1) * (R<1 * [~8*x^3 + ~4*x^2 + 4*x + 1]^2)))) & ((((A<0 * A<1) * R<1) + ((A<=1 * (A<0 * R<1)) * (R<1 * [8*x^3 + ~4*x^2 + ~4*x + 1]^2)))))")
+  by (sos "((((A<0 * R<1) + ((A<=1 * R<1) * (R<1 * [~8*x^3 + ~4*x^2 + 4*x + 1]^2)))) & ((((A<0 * A<1) * R<1) + ((A<=1 * (A<0 * R<1)) * (R<1 * [8*x^3 + ~4*x^2 + ~4*x + 1]^2)))))")
 
 
 text \<open>One component of denominator in dodecahedral example.\<close>
 
 lemma "2 <= x & x <= 125841 / 50000 & 2 <= y & y <= 125841 / 50000 & 2 <= z & z <= 125841 / 50000 --> 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z) >= (0::real)"
-  by (sos_cert "(((A<0 * R<1) + ((R<1 * ((R<5749028157/5000000000 * [~25000/222477*x + ~25000/222477*y + ~25000/222477*z + 1]^2) + ((R<864067/1779816 * [419113/864067*x + 419113/864067*y + z]^2) + ((R<320795/864067 * [419113/1283180*x + y]^2) + (R<1702293/5132720 * [x]^2))))) + (((A<=4 * (A<=5 * R<1)) * (R<3/2 * [1]^2)) + (((A<=3 * (A<=5 * R<1)) * (R<1/2 * [1]^2)) + (((A<=2 * (A<=4 * R<1)) * (R<1 * [1]^2)) + (((A<=2 * (A<=3 * R<1)) * (R<3/2 * [1]^2)) + (((A<=1 * (A<=5 * R<1)) * (R<1/2 * [1]^2)) + (((A<=1 * (A<=3 * R<1)) * (R<1/2 * [1]^2)) + (((A<=0 * (A<=4 * R<1)) * (R<1 * [1]^2)) + (((A<=0 * (A<=2 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<=1 * R<1)) * (R<3/2 * [1]^2)))))))))))))")
+  by (sos "(((A<0 * R<1) + ((R<1 * ((R<5749028157/5000000000 * [~25000/222477*x + ~25000/222477*y + ~25000/222477*z + 1]^2) + ((R<864067/1779816 * [419113/864067*x + 419113/864067*y + z]^2) + ((R<320795/864067 * [419113/1283180*x + y]^2) + (R<1702293/5132720 * [x]^2))))) + (((A<=4 * (A<=5 * R<1)) * (R<3/2 * [1]^2)) + (((A<=3 * (A<=5 * R<1)) * (R<1/2 * [1]^2)) + (((A<=2 * (A<=4 * R<1)) * (R<1 * [1]^2)) + (((A<=2 * (A<=3 * R<1)) * (R<3/2 * [1]^2)) + (((A<=1 * (A<=5 * R<1)) * (R<1/2 * [1]^2)) + (((A<=1 * (A<=3 * R<1)) * (R<1/2 * [1]^2)) + (((A<=0 * (A<=4 * R<1)) * (R<1 * [1]^2)) + (((A<=0 * (A<=2 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<=1 * R<1)) * (R<3/2 * [1]^2)))))))))))))")
 
 
 text \<open>Over a larger but simpler interval.\<close>
 
 lemma "(2::real) <= x & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 --> 0 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)"
-  by (sos_cert "((R<1 + ((R<1 * ((R<1 * [~1/6*x + ~1/6*y + ~1/6*z + 1]^2) + ((R<1/18 * [~1/2*x + ~1/2*y + z]^2) + (R<1/24 * [~1*x + y]^2)))) + (((A<0 * R<1) * (R<1/12 * [1]^2)) + (((A<=4 * (A<=5 * R<1)) * (R<1/6 * [1]^2)) + (((A<=2 * (A<=4 * R<1)) * (R<1/6 * [1]^2)) + (((A<=2 * (A<=3 * R<1)) * (R<1/6 * [1]^2)) + (((A<=0 * (A<=4 * R<1)) * (R<1/6 * [1]^2)) + (((A<=0 * (A<=2 * R<1)) * (R<1/6 * [1]^2)) + ((A<=0 * (A<=1 * R<1)) * (R<1/6 * [1]^2)))))))))))")
+  by (sos "((R<1 + ((R<1 * ((R<1 * [~1/6*x + ~1/6*y + ~1/6*z + 1]^2) + ((R<1/18 * [~1/2*x + ~1/2*y + z]^2) + (R<1/24 * [~1*x + y]^2)))) + (((A<0 * R<1) * (R<1/12 * [1]^2)) + (((A<=4 * (A<=5 * R<1)) * (R<1/6 * [1]^2)) + (((A<=2 * (A<=4 * R<1)) * (R<1/6 * [1]^2)) + (((A<=2 * (A<=3 * R<1)) * (R<1/6 * [1]^2)) + (((A<=0 * (A<=4 * R<1)) * (R<1/6 * [1]^2)) + (((A<=0 * (A<=2 * R<1)) * (R<1/6 * [1]^2)) + ((A<=0 * (A<=1 * R<1)) * (R<1/6 * [1]^2)))))))))))")
 
 
 text \<open>We can do 12. I think 12 is a sharp bound; see PP's certificate.\<close>
 
 lemma "2 <= (x::real) & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 --> 12 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)"
-  by (sos_cert "(((A<0 * R<1) + (((A<=4 * R<1) * (R<2/3 * [1]^2)) + (((A<=4 * (A<=5 * R<1)) * (R<1 * [1]^2)) + (((A<=3 * (A<=4 * R<1)) * (R<1/3 * [1]^2)) + (((A<=2 * R<1) * (R<2/3 * [1]^2)) + (((A<=2 * (A<=5 * R<1)) * (R<1/3 * [1]^2)) + (((A<=2 * (A<=4 * R<1)) * (R<8/3 * [1]^2)) + (((A<=2 * (A<=3 * R<1)) * (R<1 * [1]^2)) + (((A<=1 * (A<=4 * R<1)) * (R<1/3 * [1]^2)) + (((A<=1 * (A<=2 * R<1)) * (R<1/3 * [1]^2)) + (((A<=0 * R<1) * (R<2/3 * [1]^2)) + (((A<=0 * (A<=5 * R<1)) * (R<1/3 * [1]^2)) + (((A<=0 * (A<=4 * R<1)) * (R<8/3 * [1]^2)) + (((A<=0 * (A<=3 * R<1)) * (R<1/3 * [1]^2)) + (((A<=0 * (A<=2 * R<1)) * (R<8/3 * [1]^2)) + ((A<=0 * (A<=1 * R<1)) * (R<1 * [1]^2))))))))))))))))))")
+  by (sos "(((A<0 * R<1) + (((A<=4 * R<1) * (R<2/3 * [1]^2)) + (((A<=4 * (A<=5 * R<1)) * (R<1 * [1]^2)) + (((A<=3 * (A<=4 * R<1)) * (R<1/3 * [1]^2)) + (((A<=2 * R<1) * (R<2/3 * [1]^2)) + (((A<=2 * (A<=5 * R<1)) * (R<1/3 * [1]^2)) + (((A<=2 * (A<=4 * R<1)) * (R<8/3 * [1]^2)) + (((A<=2 * (A<=3 * R<1)) * (R<1 * [1]^2)) + (((A<=1 * (A<=4 * R<1)) * (R<1/3 * [1]^2)) + (((A<=1 * (A<=2 * R<1)) * (R<1/3 * [1]^2)) + (((A<=0 * R<1) * (R<2/3 * [1]^2)) + (((A<=0 * (A<=5 * R<1)) * (R<1/3 * [1]^2)) + (((A<=0 * (A<=4 * R<1)) * (R<8/3 * [1]^2)) + (((A<=0 * (A<=3 * R<1)) * (R<1/3 * [1]^2)) + (((A<=0 * (A<=2 * R<1)) * (R<8/3 * [1]^2)) + ((A<=0 * (A<=1 * R<1)) * (R<1 * [1]^2))))))))))))))))))")
 
 
 text \<open>Inequality from sci.math (see "Leon-Sotelo, por favor").\<close>
 
 lemma "0 <= (x::real) & 0 <= y & (x * y = 1) --> x + y <= x^2 + y^2"
-  by (sos_cert "(((A<0 * R<1) + (([1] * A=0) + (R<1 * ((R<1 * [~1/2*x + ~1/2*y + 1]^2) + (R<3/4 * [~1*x + y]^2))))))")
+  by (sos "(((A<0 * R<1) + (([1] * A=0) + (R<1 * ((R<1 * [~1/2*x + ~1/2*y + 1]^2) + (R<3/4 * [~1*x + y]^2))))))")
 
 lemma "0 <= (x::real) & 0 <= y & (x * y = 1) --> x * y * (x + y) <= x^2 + y^2"
-  by (sos_cert "(((A<0 * R<1) + (([~1*x + ~1*y + 1] * A=0) + (R<1 * ((R<1 * [~1/2*x + ~1/2*y + 1]^2) + (R<3/4 * [~1*x + y]^2))))))")
+  by (sos "(((A<0 * R<1) + (([~1*x + ~1*y + 1] * A=0) + (R<1 * ((R<1 * [~1/2*x + ~1/2*y + 1]^2) + (R<3/4 * [~1*x + y]^2))))))")
 
 lemma "0 <= (x::real) & 0 <= y --> x * y * (x + y)^2 <= (x^2 + y^2)^2"
-  by (sos_cert "(((A<0 * R<1) + (R<1 * ((R<1 * [~1/2*x^2 + y^2 + ~1/2*x*y]^2) + (R<3/4 * [~1*x^2 + x*y]^2)))))")
+  by (sos "(((A<0 * R<1) + (R<1 * ((R<1 * [~1/2*x^2 + y^2 + ~1/2*x*y]^2) + (R<3/4 * [~1*x^2 + x*y]^2)))))")
 
 lemma "(0::real) <= a & 0 <= b & 0 <= c & c * (2 * a + b)^3/ 27 <= x \<longrightarrow> c * a^2 * b <= x"
-  by (sos_cert "(((A<0 * R<1) + (((A<=3 * R<1) * (R<1 * [1]^2)) + (((A<=1 * (A<=2 * R<1)) * (R<1/27 * [~1*a + b]^2)) + ((A<=0 * (A<=2 * R<1)) * (R<8/27 * [~1*a + b]^2))))))")
+  by (sos "(((A<0 * R<1) + (((A<=3 * R<1) * (R<1 * [1]^2)) + (((A<=1 * (A<=2 * R<1)) * (R<1/27 * [~1*a + b]^2)) + ((A<=0 * (A<=2 * R<1)) * (R<8/27 * [~1*a + b]^2))))))")
 
 lemma "(0::real) < x --> 0 < 1 + x + x^2"
-  by (sos_cert "((R<1 + ((R<1 * (R<1 * [x]^2)) + (((A<0 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2))))))")
+  by (sos "((R<1 + ((R<1 * (R<1 * [x]^2)) + (((A<0 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2))))))")
 
 lemma "(0::real) <= x --> 0 < 1 + x + x^2"
-  by (sos_cert "((R<1 + ((R<1 * (R<1 * [x]^2)) + (((A<=1 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2))))))")
+  by (sos "((R<1 + ((R<1 * (R<1 * [x]^2)) + (((A<=1 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2))))))")
 
 lemma "(0::real) < 1 + x^2"
-  by (sos_cert "((R<1 + ((R<1 * (R<1 * [x]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2)))))")
+  by (sos "((R<1 + ((R<1 * (R<1 * [x]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2)))))")
 
 lemma "(0::real) <= 1 + 2 * x + x^2"
-  by (sos_cert "(((A<0 * R<1) + (R<1 * (R<1 * [x + 1]^2))))")
+  by (sos "(((A<0 * R<1) + (R<1 * (R<1 * [x + 1]^2))))")
 
 lemma "(0::real) < 1 + abs x"
-  by (sos_cert "((R<1 + (((A<=1 * R<1) * (R<1/2 * [1]^2)) + ((A<=0 * R<1) * (R<1/2 * [1]^2)))))")
+  by (sos "((R<1 + (((A<=1 * R<1) * (R<1/2 * [1]^2)) + ((A<=0 * R<1) * (R<1/2 * [1]^2)))))")
 
 lemma "(0::real) < 1 + (1 + x)^2 * (abs x)"
-  by (sos_cert "(((R<1 + (((A<=1 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [x + 1]^2))))) & ((R<1 + (((A<0 * R<1) * (R<1 * [x + 1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2))))))")
+  by (sos "(((R<1 + (((A<=1 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [x + 1]^2))))) & ((R<1 + (((A<0 * R<1) * (R<1 * [x + 1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2))))))")
 
 
 lemma "abs ((1::real) + x^2) = (1::real) + x^2"
-  by (sos_cert "(() & (((R<1 + ((R<1 * (R<1 * [x]^2)) + ((A<1 * R<1) * (R<1/2 * [1]^2))))) & ((R<1 + ((R<1 * (R<1 * [x]^2)) + ((A<0 * R<1) * (R<1 * [1]^2)))))))")
+  by (sos "(() & (((R<1 + ((R<1 * (R<1 * [x]^2)) + ((A<1 * R<1) * (R<1/2 * [1]^2))))) & ((R<1 + ((R<1 * (R<1 * [x]^2)) + ((A<0 * R<1) * (R<1 * [1]^2)))))))")
 lemma "(3::real) * x + 7 * a < 4 \<and> 3 < 2 * x \<longrightarrow> a < 0"
-  by (sos_cert "((R<1 + (((A<1 * R<1) * (R<2 * [1]^2)) + (((A<0 * R<1) * (R<3 * [1]^2)) + ((A<=0 * R<1) * (R<14 * [1]^2))))))")
+  by (sos "((R<1 + (((A<1 * R<1) * (R<2 * [1]^2)) + (((A<0 * R<1) * (R<3 * [1]^2)) + ((A<=0 * R<1) * (R<14 * [1]^2))))))")
 
 lemma "(0::real) < x --> 1 < y --> y * x <= z --> x < z"
-  by (sos_cert "((((A<0 * A<1) * R<1) + (((A<=1 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2)))))")
+  by (sos "((((A<0 * A<1) * R<1) + (((A<=1 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2)))))")
 lemma "(1::real) < x --> x^2 < y --> 1 < y"
-  by (sos_cert "((((A<0 * A<1) * R<1) + ((R<1 * ((R<1/10 * [~2*x + y + 1]^2) + (R<1/10 * [~1*x + y]^2))) + (((A<1 * R<1) * (R<1/2 * [1]^2)) + (((A<0 * R<1) * (R<1 * [x]^2)) + (((A<=0 * R<1) * ((R<1/10 * [x + 1]^2) + (R<1/10 * [x]^2))) + (((A<=0 * (A<1 * R<1)) * (R<1/5 * [1]^2)) + ((A<=0 * (A<0 * R<1)) * (R<1/5 * [1]^2)))))))))")
+  by (sos "((((A<0 * A<1) * R<1) + ((R<1 * ((R<1/10 * [~2*x + y + 1]^2) + (R<1/10 * [~1*x + y]^2))) + (((A<1 * R<1) * (R<1/2 * [1]^2)) + (((A<0 * R<1) * (R<1 * [x]^2)) + (((A<=0 * R<1) * ((R<1/10 * [x + 1]^2) + (R<1/10 * [x]^2))) + (((A<=0 * (A<1 * R<1)) * (R<1/5 * [1]^2)) + ((A<=0 * (A<0 * R<1)) * (R<1/5 * [1]^2)))))))))")
 lemma "(b::real)^2 < 4 * a * c --> ~(a * x^2 + b * x + c = 0)"
-  by (sos_cert "(((A<0 * R<1) + (R<1 * (R<1 * [2*a*x + b]^2))))")
+  by (sos "(((A<0 * R<1) + (R<1 * (R<1 * [2*a*x + b]^2))))")
 lemma "(b::real)^2 < 4 * a * c --> ~(a * x^2 + b * x + c = 0)"
-  by (sos_cert "(((A<0 * R<1) + (R<1 * (R<1 * [2*a*x + b]^2))))")
+  by (sos "(((A<0 * R<1) + (R<1 * (R<1 * [2*a*x + b]^2))))")
 lemma "((a::real) * x^2 + b * x + c = 0) --> b^2 >= 4 * a * c"
-  by (sos_cert "(((A<0 * R<1) + (R<1 * (R<1 * [2*a*x + b]^2))))")
+  by (sos "(((A<0 * R<1) + (R<1 * (R<1 * [2*a*x + b]^2))))")
 lemma "(0::real) <= b & 0 <= c & 0 <= x & 0 <= y & (x^2 = c) & (y^2 = a^2 * c + b) --> a * c <= y * x"
-  by (sos_cert "(((A<0 * (A<0 * R<1)) + (((A<=2 * (A<=3 * (A<0 * R<1))) * (R<2 * [1]^2)) + ((A<=0 * (A<=1 * R<1)) * (R<1 * [1]^2)))))")
+  by (sos "(((A<0 * (A<0 * R<1)) + (((A<=2 * (A<=3 * (A<0 * R<1))) * (R<2 * [1]^2)) + ((A<=0 * (A<=1 * R<1)) * (R<1 * [1]^2)))))")
 lemma "abs(x - z) <= e & abs(y - z) <= e & 0 <= u & 0 <= v & (u + v = 1) --> abs((u * x + v * y) - z) <= (e::real)"
-  by (sos_cert "((((A<0 * R<1) + (((A<=3 * (A<=6 * R<1)) * (R<1 * [1]^2)) + ((A<=1 * (A<=5 * R<1)) * (R<1 * [1]^2))))) & ((((A<0 * A<1) * R<1) + (((A<=3 * (A<=5 * (A<0 * R<1))) * (R<1 * [1]^2)) + ((A<=1 * (A<=4 * (A<0 * R<1))) * (R<1 * [1]^2))))))")
+  by (sos "((((A<0 * R<1) + (((A<=3 * (A<=6 * R<1)) * (R<1 * [1]^2)) + ((A<=1 * (A<=5 * R<1)) * (R<1 * [1]^2))))) & ((((A<0 * A<1) * R<1) + (((A<=3 * (A<=5 * (A<0 * R<1))) * (R<1 * [1]^2)) + ((A<=1 * (A<=4 * (A<0 * R<1))) * (R<1 * [1]^2))))))")
 
 
 (* lemma "((x::real) - y - 2 * x^4 = 0) & 0 <= x & x <= 2 & 0 <= y & y <= 3 --> y^2 - 7 * y - 12 * x + 17 >= 0" by sos *) (* Too hard?*)
 
 lemma "(0::real) <= x --> (1 + x + x^2)/(1 + x^2) <= 1 + x"
-  by (sos_cert "(((((A<0 * A<1) * R<1) + ((A<=0 * (A<0 * R<1)) * (R<1 * [x]^2)))) & ((R<1 + ((R<1 * (R<1 * [x]^2)) + ((A<0 * R<1) * (R<1 * [1]^2))))))")
+  by (sos "(((((A<0 * A<1) * R<1) + ((A<=0 * (A<0 * R<1)) * (R<1 * [x]^2)))) & ((R<1 + ((R<1 * (R<1 * [x]^2)) + ((A<0 * R<1) * (R<1 * [1]^2))))))")
 
 lemma "(0::real) <= x --> 1 - x <= 1 / (1 + x + x^2)"
-  by (sos_cert "(((R<1 + (([~4/3] * A=0) + ((R<1 * ((R<1/3 * [3/2*x + 1]^2) + (R<7/12 * [x]^2))) + ((A<=0 * R<1) * (R<1/3 * [1]^2)))))) & (((((A<0 * A<1) * R<1) + ((A<=0 * (A<0 * R<1)) * (R<1 * [x]^2)))) & ((R<1 + ((R<1 * (R<1 * [x]^2)) + (((A<0 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2))))))))")
+  by (sos "(((R<1 + (([~4/3] * A=0) + ((R<1 * ((R<1/3 * [3/2*x + 1]^2) + (R<7/12 * [x]^2))) + ((A<=0 * R<1) * (R<1/3 * [1]^2)))))) & (((((A<0 * A<1) * R<1) + ((A<=0 * (A<0 * R<1)) * (R<1 * [x]^2)))) & ((R<1 + ((R<1 * (R<1 * [x]^2)) + (((A<0 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2))))))))")
 
 lemma "(x::real) <= 1 / 2 --> - x - 2 * x^2 <= - x / (1 - x)"
-  by (sos_cert "((((A<0 * A<1) * R<1) + ((A<=0 * (A<0 * R<1)) * (R<1 * [x]^2))))")
+  by (sos "((((A<0 * A<1) * R<1) + ((A<=0 * (A<0 * R<1)) * (R<1 * [x]^2))))")
 
 lemma "4*r^2 = p^2 - 4*q & r >= (0::real) & x^2 + p*x + q = 0 --> 2*(x::real) = - p + 2*r | 2*x = -p - 2*r"
-  by (sos_cert "((((((A<0 * A<1) * R<1) + ([~4] * A=0))) & ((((A<0 * A<1) * R<1) + ([4] * A=0)))) & (((((A<0 * A<1) * R<1) + ([4] * A=0))) & ((((A<0 * A<1) * R<1) + ([~4] * A=0)))))")
+  by (sos "((((((A<0 * A<1) * R<1) + ([~4] * A=0))) & ((((A<0 * A<1) * R<1) + ([4] * A=0)))) & (((((A<0 * A<1) * R<1) + ([4] * A=0))) & ((((A<0 * A<1) * R<1) + ([~4] * A=0)))))")
 
 end
 
--- a/src/HOL/ex/SOS_Remote.thy	Wed Oct 08 10:15:04 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,37 +0,0 @@
-(*  Title:      HOL/ex/SOS_Remote.thy
-    Author:     Amine Chaieb, University of Cambridge
-    Author:     Philipp Meyer, TU Muenchen
-
-Examples for Sum_of_Squares: remote CSDP server.
-*)
-
-theory SOS_Remote
-imports "~~/src/HOL/Library/Sum_of_Squares"
-begin
-
-lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x \<Longrightarrow> a < 0"
-  by (sos remote_csdp)
-
-lemma "a1 >= 0 & a2 >= 0 \<and> (a1 * a1 + a2 * a2 = b1 * b1 + b2 * b2 + 2) \<and> (a1 * b1 + a2 * b2 = 0) --> a1 * a2 - b1 * b2 >= (0::real)"
-  by (sos remote_csdp)
-
-lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x --> a < 0"
-  by (sos remote_csdp)
-
-lemma "(0::real) <= x & x <= 1 & 0 <= y & y <= 1  --> x^2 + y^2 < 1 |(x - 1)^2 + y^2 < 1 | x^2 + (y - 1)^2 < 1 | (x - 1)^2 + (y - 1)^2 < 1"
-  by (sos remote_csdp)
-
-lemma "(0::real) <= x & 0 <= y & 0 <= z & x + y + z <= 3 --> x * y + x * z + y * z >= 3 * x * y * z"
-  by (sos remote_csdp)
-
-lemma "((x::real)^2 + y^2 + z^2 = 1) --> (x + y + z)^2 <= 3"
-  by (sos remote_csdp)
-
-lemma "(w^2 + x^2 + y^2 + z^2 = 1) --> (w + x + y + z)^2 <= (4::real)"
-  by (sos remote_csdp)
-
-lemma "(x::real) >= 1 & y >= 1 --> x * y >= x + y - 1"
-  by (sos remote_csdp)
-
-end
-