| author | aspinall | 
| Fri, 01 Oct 2004 11:51:55 +0200 | |
| changeset 15220 | cc88c8ee4d2f | 
| parent 15128 | da03f05815b0 | 
| child 15299 | 576fd0b65ed8 | 
| permissions | -rw-r--r-- | 
| 14453 | 1 | (* Title: HOL/Tools/sat_solver.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Tjark Weber | |
| 4 | Copyright 2004 | |
| 5 | ||
| 6 | Interface to external SAT solvers, and (simple) built-in SAT solvers. | |
| 7 | *) | |
| 8 | ||
| 9 | signature SAT_SOLVER = | |
| 10 | sig | |
| 14965 | 11 | exception NOT_CONFIGURED | 
| 14453 | 12 | |
| 14965 | 13 | type assignment = int -> bool option | 
| 14 | datatype result = SATISFIABLE of assignment | |
| 15 | | UNSATISFIABLE | |
| 16 | | UNKNOWN | |
| 17 | type solver = PropLogic.prop_formula -> result | |
| 18 | ||
| 19 | (* auxiliary functions to create external SAT solvers *) | |
| 14453 | 20 | val write_dimacs_cnf_file : Path.T -> PropLogic.prop_formula -> unit | 
| 21 | val write_dimacs_sat_file : Path.T -> PropLogic.prop_formula -> unit | |
| 14965 | 22 | val parse_std_result_file : Path.T -> string * string * string -> result | 
| 23 | val make_external_solver : string -> (PropLogic.prop_formula -> unit) -> (unit -> result) -> solver | |
| 14453 | 24 | |
| 15040 | 25 | val read_dimacs_cnf_file : Path.T -> PropLogic.prop_formula | 
| 26 | ||
| 14453 | 27 | (* generic interface *) | 
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 28 | val solvers : (string * solver) list ref | 
| 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 29 | val add_solver : string * solver -> unit | 
| 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 30 | val invoke_solver : string -> solver (* exception OPTION *) | 
| 14453 | 31 | end; | 
| 32 | ||
| 33 | structure SatSolver : SAT_SOLVER = | |
| 34 | struct | |
| 35 | ||
| 36 | open PropLogic; | |
| 37 | ||
| 38 | (* ------------------------------------------------------------------------- *) | |
| 14965 | 39 | (* should be raised by an external SAT solver to indicate that the solver is *) | 
| 40 | (* not configured properly *) | |
| 41 | (* ------------------------------------------------------------------------- *) | |
| 42 | ||
| 43 | exception NOT_CONFIGURED; | |
| 44 | ||
| 45 | (* ------------------------------------------------------------------------- *) | |
| 46 | (* type of partial (satisfying) assignments: 'a i = None' means that 'a' is *) | |
| 47 | (* a satisfying assigment regardless of the value of variable 'i' *) | |
| 14453 | 48 | (* ------------------------------------------------------------------------- *) | 
| 49 | ||
| 14965 | 50 | type assignment = int -> bool option; | 
| 51 | ||
| 52 | (* ------------------------------------------------------------------------- *) | |
| 53 | (* return type of SAT solvers: if the result is 'SATISFIABLE', a satisfying *) | |
| 54 | (* assignment must be returned as well *) | |
| 55 | (* ------------------------------------------------------------------------- *) | |
| 56 | ||
| 57 | datatype result = SATISFIABLE of assignment | |
| 58 | | UNSATISFIABLE | |
| 59 | | UNKNOWN; | |
| 60 | ||
| 61 | (* ------------------------------------------------------------------------- *) | |
| 62 | (* type of SAT solvers: given a propositional formula, a satisfying *) | |
| 63 | (* assignment may be returned *) | |
| 64 | (* ------------------------------------------------------------------------- *) | |
| 65 | ||
| 66 | type solver = prop_formula -> result; | |
| 14453 | 67 | |
| 68 | (* ------------------------------------------------------------------------- *) | |
| 69 | (* write_dimacs_cnf_file: serializes a formula 'fm' of propositional logic *) | |
| 70 | (* to a file in DIMACS CNF format (see "Satisfiability Suggested *) | |
| 71 | (* Format", May 8 1993, Section 2.1) *) | |
| 72 | (* Note: 'fm' must not contain a variable index less than 1. *) | |
| 14965 | 73 | (* Note: 'fm' must be given in CNF. *) | 
| 14453 | 74 | (* ------------------------------------------------------------------------- *) | 
| 75 | ||
| 76 | (* Path.T -> prop_formula -> unit *) | |
| 77 | ||
| 78 | fun write_dimacs_cnf_file path fm = | |
| 79 | let | |
| 80 | (* prop_formula -> prop_formula *) | |
| 81 | fun cnf_True_False_elim True = | |
| 82 | Or (BoolVar 1, Not (BoolVar 1)) | |
| 83 | | cnf_True_False_elim False = | |
| 84 | And (BoolVar 1, Not (BoolVar 1)) | |
| 85 | | cnf_True_False_elim fm = | |
| 86 | fm (* since 'fm' is in CNF, either 'fm'='True'/'False', or 'fm' does not contain 'True'/'False' at all *) | |
| 87 | (* prop_formula -> int *) | |
| 88 | fun cnf_number_of_clauses (And (fm1,fm2)) = | |
| 89 | (cnf_number_of_clauses fm1) + (cnf_number_of_clauses fm2) | |
| 90 | | cnf_number_of_clauses _ = | |
| 91 | 1 | |
| 92 | (* prop_formula -> string *) | |
| 14999 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 93 | fun cnf_string fm = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 94 | let | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 95 | (* prop_formula -> string list -> string list *) | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 96 | fun cnf_string_acc True acc = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 97 | error "formula is not in CNF" | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 98 | | cnf_string_acc False acc = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 99 | error "formula is not in CNF" | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 100 | | cnf_string_acc (BoolVar i) acc = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 101 | (assert (i>=1) "formula contains a variable index less than 1"; | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 102 | string_of_int i :: acc) | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 103 | | cnf_string_acc (Not (BoolVar i)) acc = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 104 | "-" :: cnf_string_acc (BoolVar i) acc | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 105 | | cnf_string_acc (Not _) acc = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 106 | error "formula is not in CNF" | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 107 | | cnf_string_acc (Or (fm1,fm2)) acc = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 108 | 				cnf_string_acc fm1 (" " :: cnf_string_acc fm2 acc)
 | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 109 | | cnf_string_acc (And (fm1,fm2)) acc = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 110 | 				cnf_string_acc fm1 (" 0\n" :: cnf_string_acc fm2 acc)
 | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 111 | in | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 112 | concat (cnf_string_acc fm []) | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 113 | end | 
| 14453 | 114 | in | 
| 115 | File.write path | |
| 116 | (let | |
| 14965 | 117 | val fm' = cnf_True_False_elim fm | 
| 118 | val number_of_vars = maxidx fm' | |
| 119 | val number_of_clauses = cnf_number_of_clauses fm' | |
| 14453 | 120 | in | 
| 121 | "c This file was generated by SatSolver.write_dimacs_cnf_file\n" ^ | |
| 122 | "c (c) Tjark Weber\n" ^ | |
| 14965 | 123 | "p cnf " ^ string_of_int number_of_vars ^ " " ^ string_of_int number_of_clauses ^ "\n" ^ | 
| 14999 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 124 | cnf_string fm' ^ " 0\n" | 
| 14453 | 125 | end) | 
| 126 | end; | |
| 127 | ||
| 128 | (* ------------------------------------------------------------------------- *) | |
| 129 | (* write_dimacs_sat_file: serializes a formula 'fm' of propositional logic *) | |
| 130 | (* to a file in DIMACS SAT format (see "Satisfiability Suggested *) | |
| 131 | (* Format", May 8 1993, Section 2.2) *) | |
| 132 | (* Note: 'fm' must not contain a variable index less than 1. *) | |
| 133 | (* ------------------------------------------------------------------------- *) | |
| 134 | ||
| 135 | (* Path.T -> prop_formula -> unit *) | |
| 136 | ||
| 137 | fun write_dimacs_sat_file path fm = | |
| 138 | let | |
| 139 | (* prop_formula -> string *) | |
| 14999 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 140 | fun sat_string fm = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 141 | let | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 142 | (* prop_formula -> string list -> string list *) | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 143 | fun sat_string_acc True acc = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 144 | "*()" :: acc | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 145 | | sat_string_acc False acc = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 146 | "+()" :: acc | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 147 | | sat_string_acc (BoolVar i) acc = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 148 | (assert (i>=1) "formula contains a variable index less than 1"; | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 149 | string_of_int i :: acc) | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 150 | | sat_string_acc (Not fm) acc = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 151 | 				"-(" :: sat_string_acc fm (")" :: acc)
 | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 152 | | sat_string_acc (Or (fm1,fm2)) acc = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 153 | 				"+(" :: sat_string_acc fm1 (" " :: sat_string_acc fm2 (")" :: acc))
 | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 154 | | sat_string_acc (And (fm1,fm2)) acc = | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 155 | 				"*(" :: sat_string_acc fm1 (" " :: sat_string_acc fm2 (")" :: acc))
 | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 156 | in | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 157 | concat (sat_string_acc fm []) | 
| 
2c39efba8f67
faster conversion into DIMACS CNF and DIMACS SAT format
 webertj parents: 
14965diff
changeset | 158 | end | 
| 14453 | 159 | in | 
| 160 | File.write path | |
| 161 | (let | |
| 162 | val number_of_vars = Int.max (maxidx fm, 1) | |
| 163 | in | |
| 164 | "c This file was generated by SatSolver.write_dimacs_sat_file\n" ^ | |
| 165 | "c (c) Tjark Weber\n" ^ | |
| 14965 | 166 | "p sat " ^ string_of_int number_of_vars ^ "\n" ^ | 
| 167 | 				"(" ^ sat_string fm ^ ")\n"
 | |
| 14453 | 168 | end) | 
| 169 | end; | |
| 170 | ||
| 171 | (* ------------------------------------------------------------------------- *) | |
| 172 | (* parse_std_result_file: scans a SAT solver's output file for a satisfying *) | |
| 14965 | 173 | (* variable assignment. Returns the assignment, or 'UNSATISFIABLE' if *) | 
| 174 | (* the file contains 'unsatisfiable', or 'UNKNOWN' if the file contains *) | |
| 175 | (* neither 'satisfiable' nor 'unsatisfiable'. Empty lines are ignored. *) | |
| 176 | (* The assignment must be given in one or more lines immediately after *) | |
| 177 | (* the line that contains 'satisfiable'. These lines must begin with *) | |
| 178 | (* 'assignment_prefix'. Variables must be separated by " ". Non- *) | |
| 179 | (* integer strings are ignored. If variable i is contained in the *) | |
| 180 | (* assignment, then i is interpreted as 'true'. If ~i is contained in *) | |
| 181 | (* the assignment, then i is interpreted as 'false'. Otherwise the *) | |
| 182 | (* value of i is taken to be unspecified. *) | |
| 14453 | 183 | (* ------------------------------------------------------------------------- *) | 
| 184 | ||
| 14965 | 185 | (* Path.T -> string * string * string -> result *) | 
| 14453 | 186 | |
| 14965 | 187 | fun parse_std_result_file path (satisfiable, assignment_prefix, unsatisfiable) = | 
| 14453 | 188 | let | 
| 189 | (* 'a option -> 'a Library.option *) | |
| 190 | fun option (SOME a) = | |
| 191 | Some a | |
| 192 | | option NONE = | |
| 193 | None | |
| 194 | (* string -> int list *) | |
| 195 | fun int_list_from_string s = | |
| 196 | mapfilter (option o Int.fromString) (space_explode " " s) | |
| 14965 | 197 | (* int list -> assignment *) | 
| 14453 | 198 | fun assignment_from_list [] i = | 
| 14965 | 199 | None (* the SAT solver didn't provide a value for this variable *) | 
| 14453 | 200 | | assignment_from_list (x::xs) i = | 
| 14965 | 201 | if x=i then (Some true) | 
| 202 | else if x=(~i) then (Some false) | |
| 14453 | 203 | else assignment_from_list xs i | 
| 14965 | 204 | (* int list -> string list -> assignment *) | 
| 205 | fun parse_assignment xs [] = | |
| 206 | assignment_from_list xs | |
| 207 | | parse_assignment xs (line::lines) = | |
| 208 | if String.isPrefix assignment_prefix line then | |
| 209 | parse_assignment (xs @ int_list_from_string line) lines | |
| 210 | else | |
| 211 | assignment_from_list xs | |
| 14453 | 212 | (* string -> string -> bool *) | 
| 213 | fun is_substring needle haystack = | |
| 214 | let | |
| 215 | val length1 = String.size needle | |
| 216 | val length2 = String.size haystack | |
| 217 | in | |
| 218 | if length2 < length1 then | |
| 219 | false | |
| 220 | else if needle = String.substring (haystack, 0, length1) then | |
| 221 | true | |
| 222 | else is_substring needle (String.substring (haystack, 1, length2-1)) | |
| 223 | end | |
| 14965 | 224 | (* string list -> result *) | 
| 14453 | 225 | fun parse_lines [] = | 
| 14965 | 226 | UNKNOWN | 
| 14453 | 227 | | parse_lines (line::lines) = | 
| 14965 | 228 | if is_substring satisfiable line then | 
| 229 | SATISFIABLE (parse_assignment [] lines) | |
| 230 | else if is_substring unsatisfiable line then | |
| 231 | UNSATISFIABLE | |
| 14453 | 232 | else | 
| 233 | parse_lines lines | |
| 234 | in | |
| 14965 | 235 | (parse_lines o (filter (fn l => l <> "")) o split_lines o File.read) path | 
| 14453 | 236 | end; | 
| 237 | ||
| 238 | (* ------------------------------------------------------------------------- *) | |
| 239 | (* make_external_solver: call 'writefn', execute 'cmd', call 'readfn' *) | |
| 240 | (* ------------------------------------------------------------------------- *) | |
| 241 | ||
| 14965 | 242 | (* string -> (PropLogic.prop_formula -> unit) -> (unit -> result) -> solver *) | 
| 14453 | 243 | |
| 244 | fun make_external_solver cmd writefn readfn fm = | |
| 14965 | 245 | (writefn fm; system cmd; readfn ()); | 
| 14453 | 246 | |
| 247 | (* ------------------------------------------------------------------------- *) | |
| 15040 | 248 | (* read_dimacs_cnf_file: returns a propositional formula that corresponds to *) | 
| 249 | (* a SAT problem given in DIMACS CNF format *) | |
| 250 | (* ------------------------------------------------------------------------- *) | |
| 251 | ||
| 252 | (* Path.T -> PropLogic.prop_formula *) | |
| 253 | ||
| 254 | fun read_dimacs_cnf_file path = | |
| 255 | let | |
| 256 | (* string list -> string list *) | |
| 257 | fun filter_preamble [] = | |
| 258 | error "problem line not found in DIMACS CNF file" | |
| 259 | | filter_preamble (line::lines) = | |
| 260 | if String.isPrefix "c " line then | |
| 261 | (* ignore comments *) | |
| 262 | filter_preamble lines | |
| 263 | else if String.isPrefix "p " line then | |
| 264 | (* ignore the problem line (which must be the last line of the preamble) *) | |
| 265 | (* Ignoring the problem line implies that if the file contains more clauses *) | |
| 266 | (* or variables than specified in its preamble, we will accept it anyway. *) | |
| 267 | lines | |
| 268 | else | |
| 269 | error "preamble in DIMACS CNF file contains a line that does not begin with \"c \" or \"p \"" | |
| 270 | (* string -> int *) | |
| 271 | fun int_from_string s = | |
| 272 | case Int.fromString s of | |
| 273 | SOME i => i | |
| 274 | 			| NONE   => error ("token " ^ quote s ^ "in DIMACS CNF file is not a number")
 | |
| 275 | (* int list -> int list list *) | |
| 276 | fun clauses xs = | |
| 277 | let | |
| 278 | val (xs1, xs2) = take_prefix (fn i => i <> 0) xs | |
| 279 | in | |
| 15128 | 280 | case xs2 of | 
| 281 | [] => [xs1] | |
| 282 | | (0::[]) => [xs1] | |
| 283 | | (0::tl) => xs1 :: clauses tl | |
| 15040 | 284 | end | 
| 285 | (* int -> PropLogic.prop_formula *) | |
| 286 | fun literal_from_int i = | |
| 287 | (assert (i<>0) "variable index in DIMACS CNF file is 0"; | |
| 288 | if i>0 then | |
| 289 | PropLogic.BoolVar i | |
| 290 | else | |
| 291 | PropLogic.Not (PropLogic.BoolVar (~i))) | |
| 292 | (* PropLogic.prop_formula list -> PropLogic.prop_formula *) | |
| 293 | fun disjunction [] = | |
| 294 | error "empty clause in DIMACS CNF file" | |
| 295 | | disjunction (x::xs) = | |
| 296 | (case xs of | |
| 297 | [] => x | |
| 298 | | _ => PropLogic.Or (x, disjunction xs)) | |
| 299 | (* PropLogic.prop_formula list -> PropLogic.prop_formula *) | |
| 300 | fun conjunction [] = | |
| 301 | error "no clause in DIMACS CNF file" | |
| 302 | | conjunction (x::xs) = | |
| 303 | (case xs of | |
| 304 | [] => x | |
| 305 | | _ => PropLogic.And (x, conjunction xs)) | |
| 306 | in | |
| 307 | (conjunction | |
| 308 | o (map disjunction) | |
| 309 | o (map (map literal_from_int)) | |
| 310 | o clauses | |
| 311 | o (map int_from_string) | |
| 312 | o flat | |
| 313 | o (map (String.fields (fn c => c mem [#" ", #"\t", #"\n"]))) | |
| 314 | o filter_preamble | |
| 315 | o (filter (fn l => l <> "")) | |
| 316 | o split_lines | |
| 317 | o File.read) | |
| 318 | path | |
| 319 | end; | |
| 320 | ||
| 321 | (* ------------------------------------------------------------------------- *) | |
| 14453 | 322 | (* solvers: a (reference to a) table of all registered SAT solvers *) | 
| 323 | (* ------------------------------------------------------------------------- *) | |
| 324 | ||
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 325 | val solvers = ref ([] : (string * solver) list); | 
| 14453 | 326 | |
| 327 | (* ------------------------------------------------------------------------- *) | |
| 328 | (* add_solver: updates 'solvers' by adding a new solver *) | |
| 329 | (* ------------------------------------------------------------------------- *) | |
| 330 | ||
| 331 | (* string * solver -> unit *) | |
| 332 | ||
| 333 | fun add_solver (name, new_solver) = | |
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 334 | 		(solvers := overwrite_warn (!solvers, (name, new_solver)) ("SAT solver " ^ quote name ^ " was defined before"));
 | 
| 14453 | 335 | |
| 336 | (* ------------------------------------------------------------------------- *) | |
| 337 | (* invoke_solver: returns the solver associated with the given 'name' *) | |
| 338 | (* Note: If no solver is associated with 'name', exception 'OPTION' will be *) | |
| 339 | (* raised. *) | |
| 340 | (* ------------------------------------------------------------------------- *) | |
| 341 | ||
| 342 | (* string -> solver *) | |
| 343 | ||
| 344 | fun invoke_solver name = | |
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 345 | (the o assoc) (!solvers, name); | 
| 14453 | 346 | |
| 347 | end; (* SatSolver *) | |
| 348 | ||
| 349 | ||
| 350 | (* ------------------------------------------------------------------------- *) | |
| 351 | (* Predefined SAT solvers *) | |
| 352 | (* ------------------------------------------------------------------------- *) | |
| 353 | ||
| 354 | (* ------------------------------------------------------------------------- *) | |
| 14753 | 355 | (* Internal SAT solver, available as 'SatSolver.invoke_solver "enumerate"' *) | 
| 356 | (* -- simply enumerates assignments until a satisfying assignment is found *) | |
| 14453 | 357 | (* ------------------------------------------------------------------------- *) | 
| 358 | ||
| 359 | let | |
| 360 | fun enum_solver fm = | |
| 361 | let | |
| 362 | (* int list *) | |
| 363 | val indices = PropLogic.indices fm | |
| 364 | (* int list -> int list -> int list option *) | |
| 365 | (* binary increment: list 'xs' of current bits, list 'ys' of all bits (lower bits first) *) | |
| 366 | fun next_list _ ([]:int list) = | |
| 367 | None | |
| 368 | | next_list [] (y::ys) = | |
| 369 | Some [y] | |
| 370 | | next_list (x::xs) (y::ys) = | |
| 371 | if x=y then | |
| 372 | (* reset the bit, continue *) | |
| 373 | next_list xs ys | |
| 374 | else | |
| 375 | (* set the lowest bit that wasn't set before, keep the higher bits *) | |
| 376 | Some (y::x::xs) | |
| 377 | (* int list -> int -> bool *) | |
| 378 | fun assignment_from_list xs i = | |
| 379 | i mem xs | |
| 14965 | 380 | (* int list -> SatSolver.result *) | 
| 14453 | 381 | fun solver_loop xs = | 
| 382 | if PropLogic.eval (assignment_from_list xs) fm then | |
| 14965 | 383 | SatSolver.SATISFIABLE (Some o (assignment_from_list xs)) | 
| 14453 | 384 | else | 
| 385 | (case next_list xs indices of | |
| 386 | Some xs' => solver_loop xs' | |
| 14965 | 387 | | None => SatSolver.UNSATISFIABLE) | 
| 14453 | 388 | in | 
| 14965 | 389 | (* start with the "first" assignment (all variables are interpreted as 'false') *) | 
| 14453 | 390 | solver_loop [] | 
| 391 | end | |
| 392 | in | |
| 14965 | 393 | 	SatSolver.add_solver ("enumerate", enum_solver)
 | 
| 14453 | 394 | end; | 
| 395 | ||
| 396 | (* ------------------------------------------------------------------------- *) | |
| 14753 | 397 | (* Internal SAT solver, available as 'SatSolver.invoke_solver "dpll"' -- a *) | 
| 398 | (* simple implementation of the DPLL algorithm (cf. L. Zhang, S. Malik: "The *) | |
| 399 | (* Quest for Efficient Boolean Satisfiability Solvers", July 2002, Fig. 1). *) | |
| 14453 | 400 | (* ------------------------------------------------------------------------- *) | 
| 401 | ||
| 402 | let | |
| 403 | local | |
| 404 | open PropLogic | |
| 405 | in | |
| 406 | fun dpll_solver fm = | |
| 407 | let | |
| 14965 | 408 | (* We could use 'PropLogic.defcnf fm' instead of 'PropLogic.nnf fm' *) | 
| 409 | (* but that sometimes introduces more boolean variables than we can *) | |
| 410 | (* handle efficiently. *) | |
| 14514 | 411 | val fm' = PropLogic.nnf fm | 
| 14453 | 412 | (* int list *) | 
| 14514 | 413 | val indices = PropLogic.indices fm' | 
| 14453 | 414 | (* int list -> int -> prop_formula *) | 
| 415 | fun partial_var_eval [] i = BoolVar i | |
| 416 | | partial_var_eval (x::xs) i = if x=i then True else if x=(~i) then False else partial_var_eval xs i | |
| 417 | (* int list -> prop_formula -> prop_formula *) | |
| 418 | fun partial_eval xs True = True | |
| 419 | | partial_eval xs False = False | |
| 420 | | partial_eval xs (BoolVar i) = partial_var_eval xs i | |
| 421 | | partial_eval xs (Not fm) = SNot (partial_eval xs fm) | |
| 422 | | partial_eval xs (Or (fm1, fm2)) = SOr (partial_eval xs fm1, partial_eval xs fm2) | |
| 423 | | partial_eval xs (And (fm1, fm2)) = SAnd (partial_eval xs fm1, partial_eval xs fm2) | |
| 424 | (* prop_formula -> int list *) | |
| 425 | fun forced_vars True = [] | |
| 426 | | forced_vars False = [] | |
| 427 | | forced_vars (BoolVar i) = [i] | |
| 428 | | forced_vars (Not (BoolVar i)) = [~i] | |
| 429 | | forced_vars (Or (fm1, fm2)) = (forced_vars fm1) inter_int (forced_vars fm2) | |
| 430 | | forced_vars (And (fm1, fm2)) = (forced_vars fm1) union_int (forced_vars fm2) | |
| 431 | (* Above, i *and* ~i may be forced. In this case the first occurrence takes *) | |
| 432 | (* precedence, and the next partial evaluation of the formula returns 'False'. *) | |
| 433 | | forced_vars _ = raise ERROR (* formula is not in negation normal form *) | |
| 434 | (* int list -> prop_formula -> (int list * prop_formula) *) | |
| 435 | fun eval_and_force xs fm = | |
| 436 | let | |
| 437 | val fm' = partial_eval xs fm | |
| 438 | val xs' = forced_vars fm' | |
| 439 | in | |
| 440 | if null xs' then | |
| 441 | (xs, fm') | |
| 442 | else | |
| 443 | eval_and_force (xs@xs') fm' (* xs and xs' should be distinct, so '@' here should have *) | |
| 444 | (* the same effect as 'union_int' *) | |
| 445 | end | |
| 446 | (* int list -> int option *) | |
| 447 | fun fresh_var xs = | |
| 448 | Library.find_first (fn i => not (i mem_int xs) andalso not ((~i) mem_int xs)) indices | |
| 449 | (* int list -> prop_formula -> int list option *) | |
| 450 | (* partial assignment 'xs' *) | |
| 451 | fun dpll xs fm = | |
| 452 | let | |
| 453 | val (xs', fm') = eval_and_force xs fm | |
| 454 | in | |
| 455 | case fm' of | |
| 456 | True => Some xs' | |
| 457 | | False => None | |
| 458 | | _ => | |
| 459 | let | |
| 460 | val x = the (fresh_var xs') (* a fresh variable must exist since 'fm' did not evaluate to 'True'/'False' *) | |
| 461 | in | |
| 462 | case dpll (x::xs') fm' of (* passing fm' rather than fm should save some simplification work *) | |
| 463 | Some xs'' => Some xs'' | |
| 464 | | None => dpll ((~x)::xs') fm' (* now try interpreting 'x' as 'False' *) | |
| 465 | end | |
| 466 | end | |
| 14965 | 467 | (* int list -> assignment *) | 
| 14453 | 468 | fun assignment_from_list [] i = | 
| 14965 | 469 | None (* the DPLL procedure didn't provide a value for this variable *) | 
| 14453 | 470 | | assignment_from_list (x::xs) i = | 
| 14965 | 471 | if x=i then (Some true) | 
| 472 | else if x=(~i) then (Some false) | |
| 14453 | 473 | else assignment_from_list xs i | 
| 474 | in | |
| 475 | (* initially, no variable is interpreted yet *) | |
| 14965 | 476 | case dpll [] fm' of | 
| 477 | Some assignment => SatSolver.SATISFIABLE (assignment_from_list assignment) | |
| 478 | | None => SatSolver.UNSATISFIABLE | |
| 14453 | 479 | end | 
| 480 | end (* local *) | |
| 481 | in | |
| 14965 | 482 | 	SatSolver.add_solver ("dpll", dpll_solver)
 | 
| 14453 | 483 | end; | 
| 484 | ||
| 485 | (* ------------------------------------------------------------------------- *) | |
| 14753 | 486 | (* Internal SAT solver, available as 'SatSolver.invoke_solver "auto"': uses *) | 
| 14965 | 487 | (* all installed solvers (other than "auto" itself) until one returns either *) | 
| 488 | (* SATISFIABLE or UNSATISFIABLE *) | |
| 14453 | 489 | (* ------------------------------------------------------------------------- *) | 
| 490 | ||
| 491 | let | |
| 492 | fun auto_solver fm = | |
| 14965 | 493 | let | 
| 494 | fun loop [] = | |
| 495 | SatSolver.UNKNOWN | |
| 496 | | loop ((name, solver)::solvers) = | |
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 497 | if name="auto" then | 
| 14965 | 498 | (* do not call solver "auto" from within "auto" *) | 
| 499 | loop solvers | |
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 500 | else | 
| 14965 | 501 | ( | 
| 502 | (if name="dpll" orelse name="enumerate" then | |
| 14805 
eff7b9df27e9
solver "auto" now issues a warning when it uses solver "enumerate"
 webertj parents: 
14753diff
changeset | 503 | 					warning ("Using SAT solver " ^ quote name ^ "; for better performance, consider using an external solver.")
 | 
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 504 | else | 
| 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 505 | ()); | 
| 14965 | 506 | (* apply 'solver' to 'fm' *) | 
| 507 | (case solver fm of | |
| 508 | SatSolver.SATISFIABLE a => SatSolver.SATISFIABLE a | |
| 509 | | SatSolver.UNSATISFIABLE => SatSolver.UNSATISFIABLE | |
| 510 | | SatSolver.UNKNOWN => loop solvers) | |
| 511 | handle SatSolver.NOT_CONFIGURED => loop solvers | |
| 512 | ) | |
| 513 | in | |
| 514 | loop (rev (!SatSolver.solvers)) | |
| 515 | end | |
| 14453 | 516 | in | 
| 517 | 	SatSolver.add_solver ("auto", auto_solver)
 | |
| 518 | end; | |
| 519 | ||
| 520 | (* ------------------------------------------------------------------------- *) | |
| 15040 | 521 | (* ZChaff, Version 2004.05.13 (http://www.princeton.edu/~chaff/zchaff.html) *) | 
| 14453 | 522 | (* ------------------------------------------------------------------------- *) | 
| 523 | ||
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 524 | let | 
| 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 525 | fun zchaff fm = | 
| 14453 | 526 | let | 
| 14965 | 527 | val _ = if (getenv "ZCHAFF_HOME") = "" then raise SatSolver.NOT_CONFIGURED else () | 
| 528 | val inpath = File.tmp_path (Path.unpack "isabelle.cnf") | |
| 529 | val outpath = File.tmp_path (Path.unpack "result") | |
| 14453 | 530 | val cmd = (getenv "ZCHAFF_HOME") ^ "/zchaff " ^ (Path.pack inpath) ^ " > " ^ (Path.pack outpath) | 
| 14965 | 531 | fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm) | 
| 15040 | 532 | 		fun readfn ()  = SatSolver.parse_std_result_file outpath ("Instance Satisfiable", "", "Instance Unsatisfiable")
 | 
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 533 | 		val _          = assert (not (File.exists inpath)) ("file " ^ quote (Path.pack inpath) ^ " exists, please delete (will not overwrite)")
 | 
| 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 534 | 		val _          = assert (not (File.exists outpath)) ("file " ^ quote (Path.pack outpath) ^ " exists, please delete (will not overwrite)")
 | 
| 14965 | 535 | val result = SatSolver.make_external_solver cmd writefn readfn fm | 
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 536 | val _ = (File.rm inpath handle _ => ()) | 
| 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 537 | val _ = (File.rm outpath handle _ => ()) | 
| 14453 | 538 | in | 
| 14965 | 539 | result | 
| 14453 | 540 | end | 
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 541 | in | 
| 14965 | 542 | 	SatSolver.add_solver ("zchaff", zchaff)
 | 
| 543 | end; | |
| 544 | ||
| 545 | (* ------------------------------------------------------------------------- *) | |
| 546 | (* BerkMin 561 (http://eigold.tripod.com/BerkMin.html) *) | |
| 547 | (* ------------------------------------------------------------------------- *) | |
| 548 | ||
| 549 | let | |
| 550 | fun berkmin fm = | |
| 551 | let | |
| 552 | val _ = if (getenv "BERKMIN_HOME") = "" orelse (getenv "BERKMIN_EXE") = "" then raise SatSolver.NOT_CONFIGURED else () | |
| 553 | val inpath = File.tmp_path (Path.unpack "isabelle.cnf") | |
| 554 | val outpath = File.tmp_path (Path.unpack "result") | |
| 555 | val cmd = (getenv "BERKMIN_HOME") ^ "/" ^ (getenv "BERKMIN_EXE") ^ " " ^ (Path.pack inpath) ^ " > " ^ (Path.pack outpath) | |
| 556 | fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm) | |
| 557 | 		fun readfn ()  = SatSolver.parse_std_result_file outpath ("Satisfiable          !!", "solution =", "UNSATISFIABLE          !!")
 | |
| 558 | 		val _          = assert (not (File.exists inpath)) ("file " ^ quote (Path.pack inpath) ^ " exists, please delete (will not overwrite)")
 | |
| 559 | 		val _          = assert (not (File.exists outpath)) ("file " ^ quote (Path.pack outpath) ^ " exists, please delete (will not overwrite)")
 | |
| 560 | val result = SatSolver.make_external_solver cmd writefn readfn fm | |
| 561 | val _ = (File.rm inpath handle _ => ()) | |
| 562 | val _ = (File.rm outpath handle _ => ()) | |
| 563 | in | |
| 564 | result | |
| 565 | end | |
| 566 | in | |
| 567 | 	SatSolver.add_solver ("berkmin", berkmin)
 | |
| 568 | end; | |
| 569 | ||
| 570 | (* ------------------------------------------------------------------------- *) | |
| 571 | (* Jerusat 1.3 (http://www.cs.tau.ac.il/~ale1/) *) | |
| 572 | (* ------------------------------------------------------------------------- *) | |
| 573 | ||
| 574 | let | |
| 575 | fun jerusat fm = | |
| 576 | let | |
| 577 | val _ = if (getenv "JERUSAT_HOME") = "" then raise SatSolver.NOT_CONFIGURED else () | |
| 578 | val inpath = File.tmp_path (Path.unpack "isabelle.cnf") | |
| 579 | val outpath = File.tmp_path (Path.unpack "result") | |
| 580 | val cmd = (getenv "JERUSAT_HOME") ^ "/Jerusat1.3 " ^ (Path.pack inpath) ^ " > " ^ (Path.pack outpath) | |
| 581 | fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm) | |
| 582 | 		fun readfn ()  = SatSolver.parse_std_result_file outpath ("s SATISFIABLE", "v ", "s UNSATISFIABLE")
 | |
| 583 | 		val _          = assert (not (File.exists inpath)) ("file " ^ quote (Path.pack inpath) ^ " exists, please delete (will not overwrite)")
 | |
| 584 | 		val _          = assert (not (File.exists outpath)) ("file " ^ quote (Path.pack outpath) ^ " exists, please delete (will not overwrite)")
 | |
| 585 | val result = SatSolver.make_external_solver cmd writefn readfn fm | |
| 586 | val _ = (File.rm inpath handle _ => ()) | |
| 587 | val _ = (File.rm outpath handle _ => ()) | |
| 588 | in | |
| 589 | result | |
| 590 | end | |
| 591 | in | |
| 592 | 	SatSolver.add_solver ("jerusat", jerusat)
 | |
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 593 | end; | 
| 14453 | 594 | |
| 595 | (* ------------------------------------------------------------------------- *) | |
| 596 | (* Add code for other SAT solvers below. *) | |
| 597 | (* ------------------------------------------------------------------------- *) | |
| 598 | ||
| 599 | (* | |
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 600 | let | 
| 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 601 | fun mysolver fm = ... | 
| 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 602 | in | 
| 14965 | 603 | 	SatSolver.add_solver ("myname", (fn fm => if mysolver_is_configured then mysolver fm else raise SatSolver.NOT_CONFIGURED));  -- register the solver
 | 
| 14487 
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
 webertj parents: 
14460diff
changeset | 604 | end; | 
| 14453 | 605 | |
| 606 | -- the solver is now available as SatSolver.invoke_solver "myname" | |
| 607 | *) |