src/HOL/Library/Sum_of_Squares/positivstellensatz_tools.ML
changeset 55508 90c42b130652
parent 43946 ba88bb44c192
child 58629 a6a6cd499d4e
--- a/src/HOL/Library/Sum_of_Squares/positivstellensatz_tools.ML	Sat Feb 15 21:09:48 2014 +0100
+++ b/src/HOL/Library/Sum_of_Squares/positivstellensatz_tools.ML	Sat Feb 15 21:11:29 2014 +0100
@@ -8,11 +8,9 @@
 signature POSITIVSTELLENSATZ_TOOLS =
 sig
   val pss_tree_to_cert : RealArith.pss_tree -> string
-
   val cert_to_pss_tree : Proof.context -> string -> RealArith.pss_tree
 end
 
-
 structure PositivstellensatzTools : POSITIVSTELLENSATZ_TOOLS =
 struct
 
@@ -31,35 +29,34 @@
 fun string_of_varpow x k =
   let
     val term = term_of x
-    val name = case term of
-      Free (n, _) => n
-    | _ => error "Term in monomial not free variable"
+    val name =
+      (case term of
+        Free (n, _) => n
+      | _ => error "Term in monomial not free variable")
   in
-    if k = 1 then name else name ^ "^" ^ string_of_int k 
+    if k = 1 then name else name ^ "^" ^ string_of_int k
   end
 
-fun string_of_monomial m = 
- if FuncUtil.Ctermfunc.is_empty m then "1" 
- else 
-  let 
-   val m' = FuncUtil.dest_monomial m
-   val vps = fold_rev (fn (x,k) => cons (string_of_varpow x k)) m' [] 
-  in foldr1 (fn (s, t) => s ^ "*" ^ t) vps
-  end
+fun string_of_monomial m =
+  if FuncUtil.Ctermfunc.is_empty m then "1"
+  else
+    let
+      val m' = FuncUtil.dest_monomial m
+      val vps = fold_rev (fn (x,k) => cons (string_of_varpow x k)) m' []
+    in foldr1 (fn (s, t) => s ^ "*" ^ t) vps end
 
 fun string_of_cmonomial (m,c) =
   if FuncUtil.Ctermfunc.is_empty m then string_of_rat c
   else if c = Rat.one then string_of_monomial m
-  else (string_of_rat c) ^ "*" ^ (string_of_monomial m);
+  else string_of_rat c ^ "*" ^ string_of_monomial m;
 
-fun string_of_poly p = 
- if FuncUtil.Monomialfunc.is_empty p then "0" 
- else
-  let 
-   val cms = map string_of_cmonomial
-     (sort (prod_ord FuncUtil.monomial_order (K EQUAL)) (FuncUtil.Monomialfunc.dest p))
-  in foldr1 (fn (t1, t2) => t1 ^ " + " ^ t2) cms
-  end;
+fun string_of_poly p =
+  if FuncUtil.Monomialfunc.is_empty p then "0"
+  else
+    let
+      val cms = map string_of_cmonomial
+        (sort (prod_ord FuncUtil.monomial_order (K EQUAL)) (FuncUtil.Monomialfunc.dest p))
+    in foldr1 (fn (t1, t2) => t1 ^ " + " ^ t2) cms end;
 
 fun pss_to_cert (RealArith.Axiom_eq i) = "A=" ^ string_of_int i
   | pss_to_cert (RealArith.Axiom_le i) = "A<=" ^ string_of_int i
@@ -68,13 +65,18 @@
   | pss_to_cert (RealArith.Rational_le r) = "R<=" ^ string_of_rat r
   | pss_to_cert (RealArith.Rational_lt r) = "R<" ^ string_of_rat r
   | pss_to_cert (RealArith.Square p) = "[" ^ string_of_poly p ^ "]^2"
-  | pss_to_cert (RealArith.Eqmul (p, pss)) = "([" ^ string_of_poly p ^ "] * " ^ pss_to_cert pss ^ ")"
-  | pss_to_cert (RealArith.Sum (pss1, pss2)) = "(" ^ pss_to_cert pss1 ^ " + " ^ pss_to_cert pss2 ^ ")"
-  | pss_to_cert (RealArith.Product (pss1, pss2)) = "(" ^ pss_to_cert pss1 ^ " * " ^ pss_to_cert pss2 ^ ")"
+  | pss_to_cert (RealArith.Eqmul (p, pss)) =
+      "([" ^ string_of_poly p ^ "] * " ^ pss_to_cert pss ^ ")"
+  | pss_to_cert (RealArith.Sum (pss1, pss2)) =
+      "(" ^ pss_to_cert pss1 ^ " + " ^ pss_to_cert pss2 ^ ")"
+  | pss_to_cert (RealArith.Product (pss1, pss2)) =
+      "(" ^ pss_to_cert pss1 ^ " * " ^ pss_to_cert pss2 ^ ")"
 
 fun pss_tree_to_cert RealArith.Trivial = "()"
   | pss_tree_to_cert (RealArith.Cert pss) = "(" ^ pss_to_cert pss ^ ")"
-  | pss_tree_to_cert (RealArith.Branch (t1, t2)) = "(" ^ pss_tree_to_cert t1 ^ " & " ^ pss_tree_to_cert t2 ^ ")"
+  | pss_tree_to_cert (RealArith.Branch (t1, t2)) =
+      "(" ^ pss_tree_to_cert t1 ^ " & " ^ pss_tree_to_cert t2 ^ ")"
+
 
 (*** certificate parsing ***)
 
@@ -82,15 +84,16 @@
 
 val str = Scan.this_string
 
-val number = Scan.repeat1 (Scan.one Symbol.is_ascii_digit >>
-  (fn s => ord s - ord "0")) >>
-  foldl1 (fn (n, d) => n * 10 + d)
+val number =
+  Scan.repeat1 (Scan.one Symbol.is_ascii_digit >> (fn s => ord s - ord "0"))
+    >> foldl1 (fn (n, d) => n * 10 + d)
 
 val nat = number
 val int = Scan.optional (str "~" >> K ~1) 1 -- nat >> op *;
 val rat = int --| str "/" -- int >> Rat.rat_of_quotient
 val rat_int = rat || int >> Rat.rat_of_int
 
+
 (* polynomial parser *)
 
 fun repeat_sep s f = f ::: Scan.repeat (str s |-- f)
@@ -98,7 +101,7 @@
 val parse_id = Scan.one Symbol.is_letter ::: Scan.many Symbol.is_letdig >> implode
 
 fun parse_varpow ctxt = parse_id -- Scan.optional (str "^" |-- nat) 1 >>
-  (fn (x, k) => (cterm_of (Proof_Context.theory_of ctxt) (Free (x, @{typ real})), k)) 
+  (fn (x, k) => (cterm_of (Proof_Context.theory_of ctxt) (Free (x, @{typ real})), k))
 
 fun parse_monomial ctxt = repeat_sep "*" (parse_varpow ctxt) >>
   (fn xs => fold FuncUtil.Ctermfunc.update xs FuncUtil.Ctermfunc.empty)
@@ -111,6 +114,7 @@
 fun parse_poly ctxt = repeat_sep "+" (parse_cmonomial ctxt) >>
   (fn xs => fold FuncUtil.Monomialfunc.update xs FuncUtil.Monomialfunc.empty)
 
+
 (* positivstellensatz parser *)
 
 val parse_axiom =
@@ -128,12 +132,12 @@
     val pc = parse_cert ctxt
     val pp = parse_poly ctxt
   in
-  (parse_axiom ||
-   parse_rational ||
-   str "[" |-- pp --| str "]^2" >> RealArith.Square ||
-   str "([" |-- pp --| str "]*" -- pc --| str ")" >> RealArith.Eqmul ||
-   str "(" |-- pc --| str "*" -- pc --| str ")" >> RealArith.Product ||
-   str "(" |-- pc --| str "+" -- pc --| str ")" >> RealArith.Sum) input
+    (parse_axiom ||
+     parse_rational ||
+     str "[" |-- pp --| str "]^2" >> RealArith.Square ||
+     str "([" |-- pp --| str "]*" -- pc --| str ")" >> RealArith.Eqmul ||
+     str "(" |-- pc --| str "*" -- pc --| str ")" >> RealArith.Product ||
+     str "(" |-- pc --| str "+" -- pc --| str ")" >> RealArith.Sum) input
   end
 
 fun parse_cert_tree ctxt input =
@@ -141,11 +145,12 @@
     val pc = parse_cert ctxt
     val pt = parse_cert_tree ctxt
   in
-  (str "()" >> K RealArith.Trivial ||
-   str "(" |-- pc --| str ")" >> RealArith.Cert ||
-   str "(" |-- pt --| str "&" -- pt --| str ")" >> RealArith.Branch) input
+    (str "()" >> K RealArith.Trivial ||
+     str "(" |-- pc --| str ")" >> RealArith.Cert ||
+     str "(" |-- pt --| str "&" -- pt --| str ")" >> RealArith.Branch) input
   end
 
+
 (* scanner *)
 
 fun cert_to_pss_tree ctxt input_str =