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