src/HOL/Tools/sat_solver.ML
author blanchet
Thu, 05 Mar 2009 10:19:51 +0100
changeset 30275 381ce8d88cb8
parent 30240 5b25fee0362c
child 31219 034f23104635
permissions -rw-r--r--
Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
     1 (*  Title:      HOL/Tools/sat_solver.ML
     2     ID:         $Id$
     3     Author:     Tjark Weber
     4     Copyright   2004-2006
     5 
     6 Interface to external SAT solvers, and (simple) built-in SAT solvers.
     7 *)
     8 
     9 signature SAT_SOLVER =
    10 sig
    11   exception NOT_CONFIGURED
    12 
    13   type assignment = int -> bool option
    14   type proof      = int list Inttab.table * int
    15   datatype result = SATISFIABLE of assignment
    16                   | UNSATISFIABLE of proof option
    17                   | UNKNOWN
    18   type solver     = PropLogic.prop_formula -> result
    19 
    20   (* auxiliary functions to create external SAT solvers *)
    21   val write_dimacs_cnf_file : Path.T -> PropLogic.prop_formula -> unit
    22   val write_dimacs_sat_file : Path.T -> PropLogic.prop_formula -> unit
    23   val read_std_result_file  : Path.T -> string * string * string -> result
    24   val make_external_solver  : string -> (PropLogic.prop_formula -> unit) -> (unit -> result) -> solver
    25 
    26   val read_dimacs_cnf_file : Path.T -> PropLogic.prop_formula
    27 
    28   (* generic solver interface *)
    29   val solvers       : (string * solver) list ref
    30   val add_solver    : string * solver -> unit
    31   val invoke_solver : string -> solver  (* exception Option *)
    32 end;
    33 
    34 structure SatSolver : SAT_SOLVER =
    35 struct
    36 
    37   open PropLogic;
    38 
    39 (* ------------------------------------------------------------------------- *)
    40 (* should be raised by an external SAT solver to indicate that the solver is *)
    41 (* not configured properly                                                   *)
    42 (* ------------------------------------------------------------------------- *)
    43 
    44   exception NOT_CONFIGURED;
    45 
    46 (* ------------------------------------------------------------------------- *)
    47 (* type of partial (satisfying) assignments: 'a i = NONE' means that 'a' is  *)
    48 (*      a satisfying assignment regardless of the value of variable 'i'      *)
    49 (* ------------------------------------------------------------------------- *)
    50 
    51   type assignment = int -> bool option;
    52 
    53 (* ------------------------------------------------------------------------- *)
    54 (* a proof of unsatisfiability, to be interpreted as follows: each integer   *)
    55 (*      is a clause ID, each list 'xs' stored under the key 'x' in the table *)
    56 (*      contains the IDs of clauses that must be resolved (in the given      *)
    57 (*      order) to obtain the new clause 'x'.  Each list 'xs' must be         *)
    58 (*      non-empty, and the literal to be resolved upon must always be unique *)
    59 (*      (e.g. "A | ~B" must not be resolved with "~A | B").  Circular        *)
    60 (*      dependencies of clauses are not allowed.  (At least) one of the      *)
    61 (*      clauses in the table must be the empty clause (i.e. contain no       *)
    62 (*      literals); its ID is given by the second component of the proof.     *)
    63 (*      The clauses of the original problem passed to the SAT solver have    *)
    64 (*      consecutive IDs starting with 0.  Clause IDs must be non-negative,   *)
    65 (*      but do not need to be consecutive.                                   *)
    66 (* ------------------------------------------------------------------------- *)
    67 
    68   type proof = int list Inttab.table * int;
    69 
    70 (* ------------------------------------------------------------------------- *)
    71 (* return type of SAT solvers: if the result is 'SATISFIABLE', a satisfying  *)
    72 (*      assignment must be returned as well; if the result is                *)
    73 (*      'UNSATISFIABLE', a proof of unsatisfiability may be returned         *)
    74 (* ------------------------------------------------------------------------- *)
    75 
    76   datatype result = SATISFIABLE of assignment
    77                   | UNSATISFIABLE of proof option
    78                   | UNKNOWN;
    79 
    80 (* ------------------------------------------------------------------------- *)
    81 (* type of SAT solvers: given a propositional formula, a satisfying          *)
    82 (*      assignment may be returned                                           *)
    83 (* ------------------------------------------------------------------------- *)
    84 
    85   type solver = prop_formula -> result;
    86 
    87 (* ------------------------------------------------------------------------- *)
    88 (* write_dimacs_cnf_file: serializes a formula 'fm' of propositional logic   *)
    89 (*      to a file in DIMACS CNF format (see "Satisfiability Suggested        *)
    90 (*      Format", May 8 1993, Section 2.1)                                    *)
    91 (* Note: 'fm' must not contain a variable index less than 1.                 *)
    92 (* Note: 'fm' must be given in CNF.                                          *)
    93 (* ------------------------------------------------------------------------- *)
    94 
    95   (* Path.T -> prop_formula -> unit *)
    96 
    97   fun write_dimacs_cnf_file path fm =
    98   let
    99     (* prop_formula -> prop_formula *)
   100     fun cnf_True_False_elim True =
   101       Or (BoolVar 1, Not (BoolVar 1))
   102       | cnf_True_False_elim False =
   103       And (BoolVar 1, Not (BoolVar 1))
   104       | cnf_True_False_elim fm =
   105       fm  (* since 'fm' is in CNF, either 'fm'='True'/'False', or 'fm' does not contain 'True'/'False' at all *)
   106     (* prop_formula -> int *)
   107     fun cnf_number_of_clauses (And (fm1,fm2)) =
   108       (cnf_number_of_clauses fm1) + (cnf_number_of_clauses fm2)
   109       | cnf_number_of_clauses _ =
   110       1
   111     (* prop_formula -> string list *)
   112     fun cnf_string fm =
   113     let
   114       (* prop_formula -> string list -> string list *)
   115       fun cnf_string_acc True acc =
   116         error "formula is not in CNF"
   117         | cnf_string_acc False acc =
   118         error "formula is not in CNF"
   119         | cnf_string_acc (BoolVar i) acc =
   120         (i>=1 orelse error "formula contains a variable index less than 1";
   121         string_of_int i :: acc)
   122         | cnf_string_acc (Not (BoolVar i)) acc =
   123         "-" :: cnf_string_acc (BoolVar i) acc
   124         | cnf_string_acc (Not _) acc =
   125         error "formula is not in CNF"
   126         | cnf_string_acc (Or (fm1,fm2)) acc =
   127         cnf_string_acc fm1 (" " :: cnf_string_acc fm2 acc)
   128         | cnf_string_acc (And (fm1,fm2)) acc =
   129         cnf_string_acc fm1 (" 0\n" :: cnf_string_acc fm2 acc)
   130     in
   131       cnf_string_acc fm []
   132     end
   133     val fm'               = cnf_True_False_elim fm
   134     val number_of_vars    = maxidx fm'
   135     val number_of_clauses = cnf_number_of_clauses fm'
   136   in
   137     File.write path
   138       ("c This file was generated by SatSolver.write_dimacs_cnf_file\n" ^
   139        "p cnf " ^ string_of_int number_of_vars ^ " " ^ string_of_int number_of_clauses ^ "\n");
   140     File.append_list path
   141       (cnf_string fm');
   142     File.append path
   143       " 0\n"
   144   end;
   145 
   146 (* ------------------------------------------------------------------------- *)
   147 (* write_dimacs_sat_file: serializes a formula 'fm' of propositional logic   *)
   148 (*      to a file in DIMACS SAT format (see "Satisfiability Suggested        *)
   149 (*      Format", May 8 1993, Section 2.2)                                    *)
   150 (* Note: 'fm' must not contain a variable index less than 1.                 *)
   151 (* ------------------------------------------------------------------------- *)
   152 
   153   (* Path.T -> prop_formula -> unit *)
   154 
   155   fun write_dimacs_sat_file path fm =
   156   let
   157     (* prop_formula -> string list *)
   158     fun sat_string fm =
   159     let
   160       (* prop_formula -> string list -> string list *)
   161       fun sat_string_acc True acc =
   162         "*()" :: acc
   163         | sat_string_acc False acc =
   164         "+()" :: acc
   165         | sat_string_acc (BoolVar i) acc =
   166         (i>=1 orelse error "formula contains a variable index less than 1";
   167         string_of_int i :: acc)
   168         | sat_string_acc (Not (BoolVar i)) acc =
   169         "-" :: sat_string_acc (BoolVar i) acc
   170         | sat_string_acc (Not fm) acc =
   171         "-(" :: sat_string_acc fm (")" :: acc)
   172         | sat_string_acc (Or (fm1,fm2)) acc =
   173         "+(" :: sat_string_acc_or fm1 (" " :: sat_string_acc_or fm2 (")" :: acc))
   174         | sat_string_acc (And (fm1,fm2)) acc =
   175         "*(" :: sat_string_acc_and fm1 (" " :: sat_string_acc_and fm2 (")" :: acc))
   176       (* optimization to make use of n-ary disjunction/conjunction *)
   177       (* prop_formula -> string list -> string list *)
   178       and sat_string_acc_or (Or (fm1,fm2)) acc =
   179         sat_string_acc_or fm1 (" " :: sat_string_acc_or fm2 acc)
   180         | sat_string_acc_or fm acc =
   181         sat_string_acc fm acc
   182       (* prop_formula -> string list -> string list *)
   183       and sat_string_acc_and (And (fm1,fm2)) acc =
   184         sat_string_acc_and fm1 (" " :: sat_string_acc_and fm2 acc)
   185         | sat_string_acc_and fm acc =
   186         sat_string_acc fm acc
   187     in
   188       sat_string_acc fm []
   189     end
   190     val number_of_vars = Int.max (maxidx fm, 1)
   191   in
   192     File.write path
   193       ("c This file was generated by SatSolver.write_dimacs_sat_file\n" ^
   194        "p sat " ^ string_of_int number_of_vars ^ "\n" ^
   195        "(");
   196     File.append_list path
   197       (sat_string fm);
   198     File.append path
   199       ")\n"
   200   end;
   201 
   202 (* ------------------------------------------------------------------------- *)
   203 (* read_std_result_file: scans a SAT solver's output file for a satisfying   *)
   204 (*      variable assignment.  Returns the assignment, or 'UNSATISFIABLE' if  *)
   205 (*      the file contains 'unsatisfiable', or 'UNKNOWN' if the file contains *)
   206 (*      neither 'satisfiable' nor 'unsatisfiable'.  Empty lines are ignored. *)
   207 (*      The assignment must be given in one or more lines immediately after  *)
   208 (*      the line that contains 'satisfiable'.  These lines must begin with   *)
   209 (*      'assignment_prefix'.  Variables must be separated by " ".  Non-      *)
   210 (*      integer strings are ignored.  If variable i is contained in the      *)
   211 (*      assignment, then i is interpreted as 'true'.  If ~i is contained in  *)
   212 (*      the assignment, then i is interpreted as 'false'.  Otherwise the     *)
   213 (*      value of i is taken to be unspecified.                               *)
   214 (* ------------------------------------------------------------------------- *)
   215 
   216   (* Path.T -> string * string * string -> result *)
   217 
   218   fun read_std_result_file path (satisfiable, assignment_prefix, unsatisfiable) =
   219   let
   220     (* string -> int list *)
   221     fun int_list_from_string s =
   222       List.mapPartial Int.fromString (space_explode " " s)
   223     (* int list -> assignment *)
   224     fun assignment_from_list [] i =
   225       NONE  (* the SAT solver didn't provide a value for this variable *)
   226       | assignment_from_list (x::xs) i =
   227       if x=i then (SOME true)
   228       else if x=(~i) then (SOME false)
   229       else assignment_from_list xs i
   230     (* int list -> string list -> assignment *)
   231     fun parse_assignment xs [] =
   232       assignment_from_list xs
   233       | parse_assignment xs (line::lines) =
   234       if String.isPrefix assignment_prefix line then
   235         parse_assignment (xs @ int_list_from_string line) lines
   236       else
   237         assignment_from_list xs
   238     (* string -> string -> bool *)
   239     fun is_substring needle haystack =
   240     let
   241       val length1 = String.size needle
   242       val length2 = String.size haystack
   243     in
   244       if length2 < length1 then
   245         false
   246       else if needle = String.substring (haystack, 0, length1) then
   247         true
   248       else is_substring needle (String.substring (haystack, 1, length2-1))
   249     end
   250     (* string list -> result *)
   251     fun parse_lines [] =
   252       UNKNOWN
   253       | parse_lines (line::lines) =
   254       if is_substring unsatisfiable line then
   255         UNSATISFIABLE NONE
   256       else if is_substring satisfiable line then
   257         SATISFIABLE (parse_assignment [] lines)
   258       else
   259         parse_lines lines
   260   in
   261     (parse_lines o (List.filter (fn l => l <> "")) o split_lines o File.read) path
   262   end;
   263 
   264 (* ------------------------------------------------------------------------- *)
   265 (* make_external_solver: call 'writefn', execute 'cmd', call 'readfn'        *)
   266 (* ------------------------------------------------------------------------- *)
   267 
   268   (* string -> (PropLogic.prop_formula -> unit) -> (unit -> result) -> solver *)
   269 
   270   fun make_external_solver cmd writefn readfn fm =
   271     (writefn fm; system cmd; readfn ());
   272 
   273 (* ------------------------------------------------------------------------- *)
   274 (* read_dimacs_cnf_file: returns a propositional formula that corresponds to *)
   275 (*      a SAT problem given in DIMACS CNF format                             *)
   276 (* ------------------------------------------------------------------------- *)
   277 
   278   (* Path.T -> PropLogic.prop_formula *)
   279 
   280   fun read_dimacs_cnf_file path =
   281   let
   282     (* string list -> string list *)
   283     fun filter_preamble [] =
   284       error "problem line not found in DIMACS CNF file"
   285       | filter_preamble (line::lines) =
   286       if String.isPrefix "c " line orelse line = "c" then
   287         (* ignore comments *)
   288         filter_preamble lines
   289       else if String.isPrefix "p " line then
   290         (* ignore the problem line (which must be the last line of the preamble) *)
   291         (* Ignoring the problem line implies that if the file contains more clauses *)
   292         (* or variables than specified in its preamble, we will accept it anyway.   *)
   293         lines
   294       else
   295         error "preamble in DIMACS CNF file contains a line that does not begin with \"c \" or \"p \""
   296     (* string -> int *)
   297     fun int_from_string s =
   298       case Int.fromString s of
   299         SOME i => i
   300       | NONE   => error ("token " ^ quote s ^ "in DIMACS CNF file is not a number")
   301     (* int list -> int list list *)
   302     fun clauses xs =
   303       let
   304         val (xs1, xs2) = take_prefix (fn i => i <> 0) xs
   305       in
   306         case xs2 of
   307           []      => [xs1]
   308         | (0::[]) => [xs1]
   309         | (0::tl) => xs1 :: clauses tl
   310         | _       => sys_error "this cannot happen"
   311       end
   312     (* int -> PropLogic.prop_formula *)
   313     fun literal_from_int i =
   314       (i<>0 orelse error "variable index in DIMACS CNF file is 0";
   315       if i>0 then
   316         PropLogic.BoolVar i
   317       else
   318         PropLogic.Not (PropLogic.BoolVar (~i)))
   319     (* PropLogic.prop_formula list -> PropLogic.prop_formula *)
   320     fun disjunction [] =
   321       error "empty clause in DIMACS CNF file"
   322       | disjunction (x::xs) =
   323       (case xs of
   324         [] => x
   325       | _  => PropLogic.Or (x, disjunction xs))
   326     (* PropLogic.prop_formula list -> PropLogic.prop_formula *)
   327     fun conjunction [] =
   328       error "no clause in DIMACS CNF file"
   329       | conjunction (x::xs) =
   330       (case xs of
   331         [] => x
   332       | _  => PropLogic.And (x, conjunction xs))
   333   in
   334     (conjunction
   335     o (map disjunction)
   336     o (map (map literal_from_int))
   337     o clauses
   338     o (map int_from_string)
   339     o List.concat
   340     o (map (String.fields (fn c => c mem [#" ", #"\t", #"\n"])))
   341     o filter_preamble
   342     o (List.filter (fn l => l <> ""))
   343     o split_lines
   344     o File.read)
   345       path
   346   end;
   347 
   348 (* ------------------------------------------------------------------------- *)
   349 (* solvers: a (reference to a) table of all registered SAT solvers           *)
   350 (* ------------------------------------------------------------------------- *)
   351 
   352   val solvers = ref ([] : (string * solver) list);
   353 
   354 (* ------------------------------------------------------------------------- *)
   355 (* add_solver: updates 'solvers' by adding a new solver                      *)
   356 (* ------------------------------------------------------------------------- *)
   357 
   358   (* string * solver -> unit *)
   359 
   360     fun add_solver (name, new_solver) =
   361       let
   362         val the_solvers = !solvers;
   363         val _ = if AList.defined (op =) the_solvers name
   364           then warning ("SAT solver " ^ quote name ^ " was defined before")
   365           else ();
   366       in solvers := AList.update (op =) (name, new_solver) the_solvers end;
   367 
   368 (* ------------------------------------------------------------------------- *)
   369 (* invoke_solver: returns the solver associated with the given 'name'        *)
   370 (* Note: If no solver is associated with 'name', exception 'Option' will be  *)
   371 (*       raised.                                                             *)
   372 (* ------------------------------------------------------------------------- *)
   373 
   374   (* string -> solver *)
   375 
   376   fun invoke_solver name =
   377     (the o AList.lookup (op =) (!solvers)) name;
   378 
   379 end;  (* SatSolver *)
   380 
   381 
   382 (* ------------------------------------------------------------------------- *)
   383 (* Predefined SAT solvers                                                    *)
   384 (* ------------------------------------------------------------------------- *)
   385 
   386 (* ------------------------------------------------------------------------- *)
   387 (* Internal SAT solver, available as 'SatSolver.invoke_solver "enumerate"'   *)
   388 (* -- simply enumerates assignments until a satisfying assignment is found   *)
   389 (* ------------------------------------------------------------------------- *)
   390 
   391 let
   392   fun enum_solver fm =
   393   let
   394     (* int list *)
   395     val indices = PropLogic.indices fm
   396     (* int list -> int list -> int list option *)
   397     (* binary increment: list 'xs' of current bits, list 'ys' of all bits (lower bits first) *)
   398     fun next_list _ ([]:int list) =
   399       NONE
   400       | next_list [] (y::ys) =
   401       SOME [y]
   402       | next_list (x::xs) (y::ys) =
   403       if x=y then
   404         (* reset the bit, continue *)
   405         next_list xs ys
   406       else
   407         (* set the lowest bit that wasn't set before, keep the higher bits *)
   408         SOME (y::x::xs)
   409     (* int list -> int -> bool *)
   410     fun assignment_from_list xs i =
   411       i mem xs
   412     (* int list -> SatSolver.result *)
   413     fun solver_loop xs =
   414       if PropLogic.eval (assignment_from_list xs) fm then
   415         SatSolver.SATISFIABLE (SOME o (assignment_from_list xs))
   416       else
   417         (case next_list xs indices of
   418           SOME xs' => solver_loop xs'
   419         | NONE     => SatSolver.UNSATISFIABLE NONE)
   420   in
   421     (* start with the "first" assignment (all variables are interpreted as 'false') *)
   422     solver_loop []
   423   end
   424 in
   425   SatSolver.add_solver ("enumerate", enum_solver)
   426 end;
   427 
   428 (* ------------------------------------------------------------------------- *)
   429 (* Internal SAT solver, available as 'SatSolver.invoke_solver "dpll"' -- a   *)
   430 (* simple implementation of the DPLL algorithm (cf. L. Zhang, S. Malik: "The *)
   431 (* Quest for Efficient Boolean Satisfiability Solvers", July 2002, Fig. 1).  *)
   432 (* ------------------------------------------------------------------------- *)
   433 
   434 let
   435   local
   436     open PropLogic
   437   in
   438     fun dpll_solver fm =
   439     let
   440       (* We could use 'PropLogic.defcnf fm' instead of 'PropLogic.nnf fm' *)
   441       (* but that sometimes leads to worse performance due to the         *)
   442       (* introduction of additional variables.                            *)
   443       val fm' = PropLogic.nnf fm
   444       (* int list *)
   445       val indices = PropLogic.indices fm'
   446       (* int list -> int -> prop_formula *)
   447       fun partial_var_eval []      i = BoolVar i
   448         | partial_var_eval (x::xs) i = if x=i then True else if x=(~i) then False else partial_var_eval xs i
   449       (* int list -> prop_formula -> prop_formula *)
   450       fun partial_eval xs True             = True
   451         | partial_eval xs False            = False
   452         | partial_eval xs (BoolVar i)      = partial_var_eval xs i
   453         | partial_eval xs (Not fm)         = SNot (partial_eval xs fm)
   454         | partial_eval xs (Or (fm1, fm2))  = SOr (partial_eval xs fm1, partial_eval xs fm2)
   455         | partial_eval xs (And (fm1, fm2)) = SAnd (partial_eval xs fm1, partial_eval xs fm2)
   456       (* prop_formula -> int list *)
   457       fun forced_vars True              = []
   458         | forced_vars False             = []
   459         | forced_vars (BoolVar i)       = [i]
   460         | forced_vars (Not (BoolVar i)) = [~i]
   461         | forced_vars (Or (fm1, fm2))   = (forced_vars fm1) inter_int (forced_vars fm2)
   462         | forced_vars (And (fm1, fm2))  = (forced_vars fm1) union_int (forced_vars fm2)
   463         (* Above, i *and* ~i may be forced.  In this case the first occurrence takes   *)
   464         (* precedence, and the next partial evaluation of the formula returns 'False'. *)
   465         | forced_vars _                 = error "formula is not in negation normal form"
   466       (* int list -> prop_formula -> (int list * prop_formula) *)
   467       fun eval_and_force xs fm =
   468       let
   469         val fm' = partial_eval xs fm
   470         val xs' = forced_vars fm'
   471       in
   472         if null xs' then
   473           (xs, fm')
   474         else
   475           eval_and_force (xs@xs') fm'  (* xs and xs' should be distinct, so '@' here should have *)
   476                                        (* the same effect as 'union_int'                         *)
   477       end
   478       (* int list -> int option *)
   479       fun fresh_var xs =
   480         Library.find_first (fn i => not (i mem_int xs) andalso not ((~i) mem_int xs)) indices
   481       (* int list -> prop_formula -> int list option *)
   482       (* partial assignment 'xs' *)
   483       fun dpll xs fm =
   484       let
   485         val (xs', fm') = eval_and_force xs fm
   486       in
   487         case fm' of
   488           True  => SOME xs'
   489         | False => NONE
   490         | _     =>
   491           let
   492             val x = valOf (fresh_var xs')  (* a fresh variable must exist since 'fm' did not evaluate to 'True'/'False' *)
   493           in
   494             case dpll (x::xs') fm' of  (* passing fm' rather than fm should save some simplification work *)
   495               SOME xs'' => SOME xs''
   496             | NONE      => dpll ((~x)::xs') fm'  (* now try interpreting 'x' as 'False' *)
   497           end
   498       end
   499       (* int list -> assignment *)
   500       fun assignment_from_list [] i =
   501         NONE  (* the DPLL procedure didn't provide a value for this variable *)
   502         | assignment_from_list (x::xs) i =
   503         if x=i then (SOME true)
   504         else if x=(~i) then (SOME false)
   505         else assignment_from_list xs i
   506     in
   507       (* initially, no variable is interpreted yet *)
   508       case dpll [] fm' of
   509         SOME assignment => SatSolver.SATISFIABLE (assignment_from_list assignment)
   510       | NONE            => SatSolver.UNSATISFIABLE NONE
   511     end
   512   end  (* local *)
   513 in
   514   SatSolver.add_solver ("dpll", dpll_solver)
   515 end;
   516 
   517 (* ------------------------------------------------------------------------- *)
   518 (* Internal SAT solver, available as 'SatSolver.invoke_solver "auto"': uses  *)
   519 (* the last installed solver (other than "auto" itself) that does not raise  *)
   520 (* 'NOT_CONFIGURED'.  (However, the solver may return 'UNKNOWN'.)            *)
   521 (* ------------------------------------------------------------------------- *)
   522 
   523 let
   524   fun auto_solver fm =
   525   let
   526     fun loop [] =
   527       SatSolver.UNKNOWN
   528       | loop ((name, solver)::solvers) =
   529       if name="auto" then
   530         (* do not call solver "auto" from within "auto" *)
   531         loop solvers
   532       else (
   533         (if name="dpll" orelse name="enumerate" then
   534           warning ("Using SAT solver " ^ quote name ^ "; for better performance, consider using an external solver.")
   535         else
   536           tracing ("Using SAT solver " ^ quote name ^ "."));
   537         (* apply 'solver' to 'fm' *)
   538         solver fm
   539           handle SatSolver.NOT_CONFIGURED => loop solvers
   540       )
   541   in
   542     loop (!SatSolver.solvers)
   543   end
   544 in
   545   SatSolver.add_solver ("auto", auto_solver)
   546 end;
   547 
   548 (* ------------------------------------------------------------------------- *)
   549 (* MiniSat 1.14                                                              *)
   550 (* (http://www.cs.chalmers.se/Cs/Research/FormalMethods/MiniSat/)            *)
   551 (* ------------------------------------------------------------------------- *)
   552 
   553 (* ------------------------------------------------------------------------- *)
   554 (* "minisat_with_proofs" requires a modified version of MiniSat 1.14 by John *)
   555 (* Matthews, which can output ASCII proof traces.  Replaying binary proof    *)
   556 (* traces generated by MiniSat-p_v1.14 has _not_ been implemented.           *)
   557 (* ------------------------------------------------------------------------- *)
   558 
   559 (* add "minisat_with_proofs" _before_ "minisat" to the available solvers, so *)
   560 (* that the latter is preferred by the "auto" solver                         *)
   561 
   562 (* There is a complication that is dealt with in the code below: MiniSat     *)
   563 (* introduces IDs for original clauses in the proof trace.  It does not (in  *)
   564 (* general) follow the convention that the original clauses are numbered     *)
   565 (* from 0 to n-1 (where n is the number of clauses in the formula).          *)
   566 
   567 let
   568   exception INVALID_PROOF of string
   569   fun minisat_with_proofs fm =
   570   let
   571     val _          = if (getenv "MINISAT_HOME") = "" then raise SatSolver.NOT_CONFIGURED else ()
   572     val serial_str = serial_string ()
   573     val inpath     = File.tmp_path (Path.explode ("isabelle" ^ serial_str ^ ".cnf"))
   574     val outpath    = File.tmp_path (Path.explode ("result" ^ serial_str))
   575     val proofpath  = File.tmp_path (Path.explode ("result" ^ serial_str ^ ".prf"))
   576     val cmd        = (getenv "MINISAT_HOME") ^ "/minisat " ^ (Path.implode inpath) ^ " -r " ^ (Path.implode outpath) ^ " -t " ^ (Path.implode proofpath) ^ "> /dev/null"
   577     fun writefn fm = SatSolver.write_dimacs_cnf_file inpath fm
   578     fun readfn ()  = SatSolver.read_std_result_file outpath ("SAT", "", "UNSAT")
   579     val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
   580     val _          = if File.exists outpath then warning ("overwriting existing file " ^ quote (Path.implode outpath)) else ()
   581     val cnf        = PropLogic.defcnf fm
   582     val result     = SatSolver.make_external_solver cmd writefn readfn cnf
   583     val _          = try File.rm inpath
   584     val _          = try File.rm outpath
   585   in  case result of
   586     SatSolver.UNSATISFIABLE NONE =>
   587     (let
   588       (* string list *)
   589       val proof_lines = (split_lines o File.read) proofpath
   590         handle IO.Io _ => raise INVALID_PROOF "Could not read file \"result.prf\""
   591       (* representation of clauses as ordered lists of literals (with duplicates removed) *)
   592       (* prop_formula -> int list *)
   593       fun clause_to_lit_list (PropLogic.Or (fm1, fm2)) =
   594         OrdList.union int_ord (clause_to_lit_list fm1) (clause_to_lit_list fm2)
   595         | clause_to_lit_list (PropLogic.BoolVar i) =
   596         [i]
   597         | clause_to_lit_list (PropLogic.Not (PropLogic.BoolVar i)) =
   598         [~i]
   599         | clause_to_lit_list _ =
   600         raise INVALID_PROOF "Error: invalid clause in CNF formula."
   601       (* prop_formula -> int *)
   602       fun cnf_number_of_clauses (PropLogic.And (fm1, fm2)) =
   603         cnf_number_of_clauses fm1 + cnf_number_of_clauses fm2
   604         | cnf_number_of_clauses _ =
   605         1
   606       val number_of_clauses = cnf_number_of_clauses cnf
   607       (* int list array *)
   608       val clauses = Array.array (number_of_clauses, [])
   609       (* initialize the 'clauses' array *)
   610       (* prop_formula * int -> int *)
   611       fun init_array (PropLogic.And (fm1, fm2), n) =
   612         init_array (fm2, init_array (fm1, n))
   613         | init_array (fm, n) =
   614         (Array.update (clauses, n, clause_to_lit_list fm); n+1)
   615       val _ = init_array (cnf, 0)
   616       (* optimization for the common case where MiniSat "R"s clauses in their *)
   617       (* original order:                                                      *)
   618       val last_ref_clause = ref (number_of_clauses - 1)
   619       (* search the 'clauses' array for the given list of literals 'lits', *)
   620       (* starting at index '!last_ref_clause + 1'                          *)
   621       (* int list -> int option *)
   622       fun original_clause_id lits =
   623       let
   624         fun original_clause_id_from index =
   625           if index = number_of_clauses then
   626             (* search from beginning again *)
   627             original_clause_id_from 0
   628           (* both 'lits' and the list of literals used in 'clauses' are sorted, so *)
   629           (* testing for equality should suffice -- barring duplicate literals     *)
   630           else if Array.sub (clauses, index) = lits then (
   631             (* success *)
   632             last_ref_clause := index;
   633             SOME index
   634           ) else if index = !last_ref_clause then
   635             (* failure *)
   636             NONE
   637           else
   638             (* continue search *)
   639             original_clause_id_from (index + 1)
   640       in
   641         original_clause_id_from (!last_ref_clause + 1)
   642       end
   643       (* string -> int *)
   644       fun int_from_string s = (
   645         case Int.fromString s of
   646           SOME i => i
   647         | NONE   => raise INVALID_PROOF ("File format error: number expected (" ^ quote s ^ " encountered).")
   648       )
   649       (* parse the proof file *)
   650       val clause_table  = ref (Inttab.empty : int list Inttab.table)
   651       val empty_id      = ref ~1
   652       (* contains a mapping from clause IDs as used by MiniSat to clause IDs in *)
   653       (* our proof format, where original clauses are numbered starting from 0  *)
   654       val clause_id_map = ref (Inttab.empty : int Inttab.table)
   655       fun sat_to_proof id = (
   656         case Inttab.lookup (!clause_id_map) id of
   657           SOME id' => id'
   658         | NONE     => raise INVALID_PROOF ("Clause ID " ^ Int.toString id ^ " used, but not defined.")
   659       )
   660       val next_id = ref (number_of_clauses - 1)
   661       (* string list -> unit *)
   662       fun process_tokens [] =
   663         ()
   664         | process_tokens (tok::toks) =
   665         if tok="R" then (
   666           case toks of
   667             id::sep::lits =>
   668             let
   669               val _        = if !empty_id = ~1 then () else raise INVALID_PROOF "File format error: \"R\" disallowed after \"X\"."
   670               val cid      = int_from_string id
   671               val _        = if sep = "<=" then () else raise INVALID_PROOF ("File format error: \"<=\" expected (" ^ quote sep ^ " encountered).")
   672               val ls       = sort int_ord (map int_from_string lits)
   673               val proof_id = case original_clause_id ls of
   674                                SOME orig_id => orig_id
   675                              | NONE         => raise INVALID_PROOF ("Original clause (new ID is " ^ id ^ ") not found.")
   676             in
   677               (* extend the mapping of clause IDs with this newly defined ID *)
   678               clause_id_map := Inttab.update_new (cid, proof_id) (!clause_id_map)
   679                 handle Inttab.DUP _ => raise INVALID_PROOF ("File format error: clause " ^ id ^ " defined more than once (in \"R\").")
   680               (* the proof itself doesn't change *)
   681             end
   682           | _ =>
   683             raise INVALID_PROOF "File format error: \"R\" followed by an insufficient number of tokens."
   684         ) else if tok="C" then (
   685           case toks of
   686             id::sep::ids =>
   687             let
   688               val _        = if !empty_id = ~1 then () else raise INVALID_PROOF "File format error: \"C\" disallowed after \"X\"."
   689               val cid      = int_from_string id
   690               val _        = if sep = "<=" then () else raise INVALID_PROOF ("File format error: \"<=\" expected (" ^ quote sep ^ " encountered).")
   691               (* ignore the pivot literals in MiniSat's trace *)
   692               fun unevens []             = raise INVALID_PROOF "File format error: \"C\" followed by an even number of IDs."
   693                 | unevens (x :: [])      = x :: []
   694                 | unevens (x :: _ :: xs) = x :: unevens xs
   695               val rs       = (map sat_to_proof o unevens o map int_from_string) ids
   696               (* extend the mapping of clause IDs with this newly defined ID *)
   697               val proof_id = inc next_id
   698               val _        = clause_id_map := Inttab.update_new (cid, proof_id) (!clause_id_map)
   699                                handle Inttab.DUP _ => raise INVALID_PROOF ("File format error: clause " ^ id ^ " defined more than once (in \"C\").")
   700             in
   701               (* update clause table *)
   702               clause_table := Inttab.update_new (proof_id, rs) (!clause_table)
   703                 handle Inttab.DUP _ => raise INVALID_PROOF ("Error: internal ID for clause " ^ id ^ " already used.")
   704             end
   705           | _ =>
   706             raise INVALID_PROOF "File format error: \"C\" followed by an insufficient number of tokens."
   707         ) else if tok="D" then (
   708           case toks of
   709             [id] =>
   710             let
   711               val _ = if !empty_id = ~1 then () else raise INVALID_PROOF "File format error: \"D\" disallowed after \"X\"."
   712               val _ = sat_to_proof (int_from_string id)
   713             in
   714               (* simply ignore "D" *)
   715               ()
   716             end
   717           | _ =>
   718             raise INVALID_PROOF "File format error: \"D\" followed by an illegal number of tokens."
   719         ) else if tok="X" then (
   720           case toks of
   721             [id1, id2] =>
   722             let
   723               val _            = if !empty_id = ~1 then () else raise INVALID_PROOF "File format error: more than one end-of-proof statement."
   724               val _            = sat_to_proof (int_from_string id1)
   725               val new_empty_id = sat_to_proof (int_from_string id2)
   726             in
   727               (* update conflict id *)
   728               empty_id := new_empty_id
   729             end
   730           | _ =>
   731             raise INVALID_PROOF "File format error: \"X\" followed by an illegal number of tokens."
   732         ) else
   733           raise INVALID_PROOF ("File format error: unknown token " ^ quote tok ^ " encountered.")
   734       (* string list -> unit *)
   735       fun process_lines [] =
   736         ()
   737         | process_lines (l::ls) = (
   738           process_tokens (String.tokens (fn c => c = #" " orelse c = #"\t") l);
   739           process_lines ls
   740         )
   741       (* proof *)
   742       val _ = process_lines proof_lines
   743       val _ = if !empty_id <> ~1 then () else raise INVALID_PROOF "File format error: no conflicting clause specified."
   744     in
   745       SatSolver.UNSATISFIABLE (SOME (!clause_table, !empty_id))
   746     end handle INVALID_PROOF reason => (warning reason; SatSolver.UNSATISFIABLE NONE))
   747   | result =>
   748     result
   749   end
   750 in
   751   SatSolver.add_solver ("minisat_with_proofs", minisat_with_proofs)
   752 end;
   753 
   754 let
   755   fun minisat fm =
   756   let
   757     val _          = if (getenv "MINISAT_HOME") = "" then raise SatSolver.NOT_CONFIGURED else ()
   758     val serial_str = serial_string ()
   759     val inpath     = File.tmp_path (Path.explode ("isabelle" ^ serial_str ^ ".cnf"))
   760     val outpath    = File.tmp_path (Path.explode ("result" ^ serial_str))
   761     val cmd        = (getenv "MINISAT_HOME") ^ "/minisat " ^ (Path.implode inpath) ^ " -r " ^ (Path.implode outpath) ^ " > /dev/null"
   762     fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
   763     fun readfn ()  = SatSolver.read_std_result_file outpath ("SAT", "", "UNSAT")
   764     val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
   765     val _          = if File.exists outpath then warning ("overwriting existing file " ^ quote (Path.implode outpath)) else ()
   766     val result     = SatSolver.make_external_solver cmd writefn readfn fm
   767     val _          = try File.rm inpath
   768     val _          = try File.rm outpath
   769   in
   770     result
   771   end
   772 in
   773   SatSolver.add_solver ("minisat", minisat)
   774 end;
   775 
   776 (* ------------------------------------------------------------------------- *)
   777 (* zChaff (http://www.princeton.edu/~chaff/zchaff.html)                      *)
   778 (* ------------------------------------------------------------------------- *)
   779 
   780 (* ------------------------------------------------------------------------- *)
   781 (* 'zchaff_with_proofs' applies the "zchaff" prover to a formula, and if     *)
   782 (* zChaff finds that the formula is unsatisfiable, a proof of this is read   *)
   783 (* from a file "resolve_trace" that was generated by zChaff.  See the code   *)
   784 (* below for the expected format of the "resolve_trace" file.  Aside from    *)
   785 (* some basic syntactic checks, no verification of the proof is performed.   *)
   786 (* ------------------------------------------------------------------------- *)
   787 
   788 (* add "zchaff_with_proofs" _before_ "zchaff" to the available solvers, so   *)
   789 (* that the latter is preferred by the "auto" solver                         *)
   790 
   791 let
   792   exception INVALID_PROOF of string
   793   fun zchaff_with_proofs fm =
   794   case SatSolver.invoke_solver "zchaff" fm of
   795     SatSolver.UNSATISFIABLE NONE =>
   796     (let
   797       (* string list *)
   798       val proof_lines = ((split_lines o File.read) (Path.explode "resolve_trace"))
   799         handle IO.Io _ => raise INVALID_PROOF "Could not read file \"resolve_trace\""
   800       (* PropLogic.prop_formula -> int *)
   801       fun cnf_number_of_clauses (PropLogic.And (fm1, fm2)) = cnf_number_of_clauses fm1 + cnf_number_of_clauses fm2
   802         | cnf_number_of_clauses _                          = 1
   803       (* string -> int *)
   804       fun int_from_string s = (
   805         case Int.fromString s of
   806           SOME i => i
   807         | NONE   => raise INVALID_PROOF ("File format error: number expected (" ^ quote s ^ " encountered).")
   808       )
   809       (* parse the "resolve_trace" file *)
   810       val clause_offset = ref ~1
   811       val clause_table  = ref (Inttab.empty : int list Inttab.table)
   812       val empty_id      = ref ~1
   813       (* string list -> unit *)
   814       fun process_tokens [] =
   815         ()
   816         | process_tokens (tok::toks) =
   817         if tok="CL:" then (
   818           case toks of
   819             id::sep::ids =>
   820             let
   821               val _   = if !clause_offset = ~1 then () else raise INVALID_PROOF ("File format error: \"CL:\" disallowed after \"VAR:\".")
   822               val _   = if !empty_id = ~1 then () else raise INVALID_PROOF ("File format error: \"CL:\" disallowed after \"CONF:\".")
   823               val cid = int_from_string id
   824               val _   = if sep = "<=" then () else raise INVALID_PROOF ("File format error: \"<=\" expected (" ^ quote sep ^ " encountered).")
   825               val rs  = map int_from_string ids
   826             in
   827               (* update clause table *)
   828               clause_table := Inttab.update_new (cid, rs) (!clause_table)
   829                 handle Inttab.DUP _ => raise INVALID_PROOF ("File format error: clause " ^ id ^ " defined more than once.")
   830             end
   831           | _ =>
   832             raise INVALID_PROOF "File format error: \"CL:\" followed by an insufficient number of tokens."
   833         ) else if tok="VAR:" then (
   834           case toks of
   835             id::levsep::levid::valsep::valid::antesep::anteid::litsep::lits =>
   836             let
   837               val _   = if !empty_id = ~1 then () else raise INVALID_PROOF ("File format error: \"VAR:\" disallowed after \"CONF:\".")
   838               (* set 'clause_offset' to the largest used clause ID *)
   839               val _   = if !clause_offset = ~1 then clause_offset :=
   840                 (case Inttab.max_key (!clause_table) of
   841                   SOME id => id
   842                 | NONE    => cnf_number_of_clauses (PropLogic.defcnf fm) - 1  (* the first clause ID is 0, not 1 *))
   843                 else
   844                   ()
   845               val vid = int_from_string id
   846               val _   = if levsep = "L:" then () else raise INVALID_PROOF ("File format error: \"L:\" expected (" ^ quote levsep ^ " encountered).")
   847               val _   = int_from_string levid
   848               val _   = if valsep = "V:" then () else raise INVALID_PROOF ("File format error: \"V:\" expected (" ^ quote valsep ^ " encountered).")
   849               val _   = int_from_string valid
   850               val _   = if antesep = "A:" then () else raise INVALID_PROOF ("File format error: \"A:\" expected (" ^ quote antesep ^ " encountered).")
   851               val aid = int_from_string anteid
   852               val _   = if litsep = "Lits:" then () else raise INVALID_PROOF ("File format error: \"Lits:\" expected (" ^ quote litsep ^ " encountered).")
   853               val ls  = map int_from_string lits
   854               (* convert the data provided by zChaff to our resolution-style proof format *)
   855               (* each "VAR:" line defines a unit clause, the resolvents are implicitly    *)
   856               (* given by the literals in the antecedent clause                           *)
   857               (* we use the sum of '!clause_offset' and the variable ID as clause ID for the unit clause *)
   858               val cid = !clause_offset + vid
   859               (* the low bit of each literal gives its sign (positive/negative), therefore  *)
   860               (* we have to divide each literal by 2 to obtain the proper variable ID; then *)
   861               (* we add '!clause_offset' to obtain the ID of the corresponding unit clause  *)
   862               val vids = filter (not_equal vid) (map (fn l => l div 2) ls)
   863               val rs   = aid :: map (fn v => !clause_offset + v) vids
   864             in
   865               (* update clause table *)
   866               clause_table := Inttab.update_new (cid, rs) (!clause_table)
   867                 handle Inttab.DUP _ => raise INVALID_PROOF ("File format error: clause " ^ string_of_int cid ^ " (derived from antecedent for variable " ^ id ^ ") already defined.")
   868             end
   869           | _ =>
   870             raise INVALID_PROOF "File format error: \"VAR:\" followed by an insufficient number of tokens."
   871         ) else if tok="CONF:" then (
   872           case toks of
   873             id::sep::ids =>
   874             let
   875               val _   = if !empty_id = ~1 then () else raise INVALID_PROOF "File format error: more than one conflicting clause specified."
   876               val cid = int_from_string id
   877               val _   = if sep = "==" then () else raise INVALID_PROOF ("File format error: \"==\" expected (" ^ quote sep ^ " encountered).")
   878               val ls  = map int_from_string ids
   879               (* the conflict clause must be resolved with the unit clauses *)
   880               (* for its literals to obtain the empty clause                *)
   881               val vids         = map (fn l => l div 2) ls
   882               val rs           = cid :: map (fn v => !clause_offset + v) vids
   883               val new_empty_id = getOpt (Inttab.max_key (!clause_table), !clause_offset) + 1
   884             in
   885               (* update clause table and conflict id *)
   886               clause_table := Inttab.update_new (new_empty_id, rs) (!clause_table)
   887                 handle Inttab.DUP _ => raise INVALID_PROOF ("File format error: clause " ^ string_of_int new_empty_id ^ " (empty clause derived from clause " ^ id ^ ") already defined.");
   888               empty_id     := new_empty_id
   889             end
   890           | _ =>
   891             raise INVALID_PROOF "File format error: \"CONF:\" followed by an insufficient number of tokens."
   892         ) else
   893           raise INVALID_PROOF ("File format error: unknown token " ^ quote tok ^ " encountered.")
   894       (* string list -> unit *)
   895       fun process_lines [] =
   896         ()
   897         | process_lines (l::ls) = (
   898           process_tokens (String.tokens (fn c => c = #" " orelse c = #"\t") l);
   899           process_lines ls
   900         )
   901       (* proof *)
   902       val _ = process_lines proof_lines
   903       val _ = if !empty_id <> ~1 then () else raise INVALID_PROOF "File format error: no conflicting clause specified."
   904     in
   905       SatSolver.UNSATISFIABLE (SOME (!clause_table, !empty_id))
   906     end handle INVALID_PROOF reason => (warning reason; SatSolver.UNSATISFIABLE NONE))
   907   | result =>
   908     result
   909 in
   910   SatSolver.add_solver ("zchaff_with_proofs", zchaff_with_proofs)
   911 end;
   912 
   913 let
   914   fun zchaff fm =
   915   let
   916     val _          = if (getenv "ZCHAFF_HOME") = "" then raise SatSolver.NOT_CONFIGURED else ()
   917     val serial_str = serial_string ()
   918     val inpath     = File.tmp_path (Path.explode ("isabelle" ^ serial_str ^ ".cnf"))
   919     val outpath    = File.tmp_path (Path.explode ("result" ^ serial_str))
   920     val cmd        = (getenv "ZCHAFF_HOME") ^ "/zchaff " ^ (Path.implode inpath) ^ " > " ^ (Path.implode outpath)
   921     fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
   922     fun readfn ()  = SatSolver.read_std_result_file outpath ("Instance Satisfiable", "", "Instance Unsatisfiable")
   923     val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
   924     val _          = if File.exists outpath then warning ("overwriting existing file " ^ quote (Path.implode outpath)) else ()
   925     val result     = SatSolver.make_external_solver cmd writefn readfn fm
   926     val _          = try File.rm inpath
   927     val _          = try File.rm outpath
   928   in
   929     result
   930   end
   931 in
   932   SatSolver.add_solver ("zchaff", zchaff)
   933 end;
   934 
   935 (* ------------------------------------------------------------------------- *)
   936 (* BerkMin 561 (http://eigold.tripod.com/BerkMin.html)                       *)
   937 (* ------------------------------------------------------------------------- *)
   938 
   939 let
   940   fun berkmin fm =
   941   let
   942     val _          = if (getenv "BERKMIN_HOME") = "" then raise SatSolver.NOT_CONFIGURED else ()
   943     val serial_str = serial_string ()
   944     val inpath     = File.tmp_path (Path.explode ("isabelle" ^ serial_str ^ ".cnf"))
   945     val outpath    = File.tmp_path (Path.explode ("result" ^ serial_str))
   946     val exec       = getenv "BERKMIN_EXE"
   947     val cmd        = (getenv "BERKMIN_HOME") ^ "/" ^ (if exec = "" then "BerkMin561" else exec) ^ " " ^ (Path.implode inpath) ^ " > " ^ (Path.implode outpath)
   948     fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
   949     fun readfn ()  = SatSolver.read_std_result_file outpath ("Satisfiable          !!", "solution =", "UNSATISFIABLE          !!")
   950     val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
   951     val _          = if File.exists outpath then warning ("overwriting existing file " ^ quote (Path.implode outpath)) else ()
   952     val result     = SatSolver.make_external_solver cmd writefn readfn fm
   953     val _          = try File.rm inpath
   954     val _          = try File.rm outpath
   955   in
   956     result
   957   end
   958 in
   959   SatSolver.add_solver ("berkmin", berkmin)
   960 end;
   961 
   962 (* ------------------------------------------------------------------------- *)
   963 (* Jerusat 1.3 (http://www.cs.tau.ac.il/~ale1/)                              *)
   964 (* ------------------------------------------------------------------------- *)
   965 
   966 let
   967   fun jerusat fm =
   968   let
   969     val _          = if (getenv "JERUSAT_HOME") = "" then raise SatSolver.NOT_CONFIGURED else ()
   970     val serial_str = serial_string ()
   971     val inpath     = File.tmp_path (Path.explode ("isabelle" ^ serial_str ^ ".cnf"))
   972     val outpath    = File.tmp_path (Path.explode ("result" ^ serial_str))
   973     val cmd        = (getenv "JERUSAT_HOME") ^ "/Jerusat1.3 " ^ (Path.implode inpath) ^ " > " ^ (Path.implode outpath)
   974     fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
   975     fun readfn ()  = SatSolver.read_std_result_file outpath ("s SATISFIABLE", "v ", "s UNSATISFIABLE")
   976     val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
   977     val _          = if File.exists outpath then warning ("overwriting existing file " ^ quote (Path.implode outpath)) else ()
   978     val result     = SatSolver.make_external_solver cmd writefn readfn fm
   979     val _          = try File.rm inpath
   980     val _          = try File.rm outpath
   981   in
   982     result
   983   end
   984 in
   985   SatSolver.add_solver ("jerusat", jerusat)
   986 end;
   987 
   988 (* ------------------------------------------------------------------------- *)
   989 (* Add code for other SAT solvers below.                                     *)
   990 (* ------------------------------------------------------------------------- *)
   991 
   992 (*
   993 let
   994   fun mysolver fm = ...
   995 in
   996   SatSolver.add_solver ("myname", (fn fm => if mysolver_is_configured then mysolver fm else raise SatSolver.NOT_CONFIGURED));  -- register the solver
   997 end;
   998 
   999 -- the solver is now available as SatSolver.invoke_solver "myname"
  1000 *)