src/HOL/Tools/sat_solver.ML
author haftmann
Wed, 31 Jan 2007 16:05:13 +0100
changeset 22220 6dc8d0dca678
parent 21858 05f57309170c
child 22567 1565d476a9e2
permissions -rw-r--r--
dropped Output.update_warn
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
webertj@14965
    11
	exception NOT_CONFIGURED
webertj@14453
    12
webertj@14965
    13
	type assignment = int -> bool option
webertj@17493
    14
	type proof      = int list Inttab.table * int
webertj@14965
    15
	datatype result = SATISFIABLE of assignment
webertj@17493
    16
	                | UNSATISFIABLE of proof option
webertj@14965
    17
	                | UNKNOWN
webertj@14965
    18
	type solver     = PropLogic.prop_formula -> result
webertj@14965
    19
webertj@14965
    20
	(* auxiliary functions to create external SAT solvers *)
webertj@14453
    21
	val write_dimacs_cnf_file : Path.T -> PropLogic.prop_formula -> unit
webertj@14453
    22
	val write_dimacs_sat_file : Path.T -> PropLogic.prop_formula -> unit
webertj@17620
    23
	val read_std_result_file  : Path.T -> string * string * string -> result
webertj@14965
    24
	val make_external_solver  : string -> (PropLogic.prop_formula -> unit) -> (unit -> result) -> solver
webertj@14453
    25
webertj@15040
    26
	val read_dimacs_cnf_file : Path.T -> PropLogic.prop_formula
webertj@15040
    27
webertj@16899
    28
	(* generic solver interface *)
webertj@14487
    29
	val solvers       : (string * solver) list ref
webertj@14487
    30
	val add_solver    : string * solver -> unit
webertj@15605
    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
webertj@14453
    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
webertj@14965
    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
webertj@14965
    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
webertj@17493
    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
webertj@14965
    76
	datatype result = SATISFIABLE of assignment
webertj@17493
    77
	                | UNSATISFIABLE of proof option
webertj@14965
    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
webertj@14965
    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
webertj@14453
    95
	(* Path.T -> prop_formula -> unit *)
webertj@14453
    96
webertj@14453
    97
	fun write_dimacs_cnf_file path fm =
webertj@14453
    98
	let
webertj@14453
    99
		(* prop_formula -> prop_formula *)
webertj@14453
   100
		fun cnf_True_False_elim True =
webertj@14453
   101
			Or (BoolVar 1, Not (BoolVar 1))
webertj@14453
   102
		  | cnf_True_False_elim False =
webertj@14453
   103
			And (BoolVar 1, Not (BoolVar 1))
webertj@14453
   104
		  | cnf_True_False_elim fm =
webertj@14453
   105
			fm  (* since 'fm' is in CNF, either 'fm'='True'/'False', or 'fm' does not contain 'True'/'False' at all *)
webertj@14453
   106
		(* prop_formula -> int *)
webertj@14453
   107
		fun cnf_number_of_clauses (And (fm1,fm2)) =
webertj@14453
   108
			(cnf_number_of_clauses fm1) + (cnf_number_of_clauses fm2)
webertj@14453
   109
		  | cnf_number_of_clauses _ =
webertj@14453
   110
			1
webertj@16915
   111
		(* prop_formula -> string list *)
webertj@14999
   112
		fun cnf_string fm =
webertj@14999
   113
		let
webertj@14999
   114
			(* prop_formula -> string list -> string list *)
webertj@14999
   115
			fun cnf_string_acc True acc =
webertj@14999
   116
				error "formula is not in CNF"
webertj@14999
   117
			  | cnf_string_acc False acc =
webertj@14999
   118
				error "formula is not in CNF"
webertj@14999
   119
			  | cnf_string_acc (BoolVar i) acc =
webertj@14999
   120
				(assert (i>=1) "formula contains a variable index less than 1";
webertj@14999
   121
				string_of_int i :: acc)
webertj@14999
   122
			  | cnf_string_acc (Not (BoolVar i)) acc =
webertj@14999
   123
				"-" :: cnf_string_acc (BoolVar i) acc
webertj@14999
   124
			  | cnf_string_acc (Not _) acc =
webertj@14999
   125
				error "formula is not in CNF"
webertj@14999
   126
			  | cnf_string_acc (Or (fm1,fm2)) acc =
webertj@14999
   127
				cnf_string_acc fm1 (" " :: cnf_string_acc fm2 acc)
webertj@14999
   128
			  | cnf_string_acc (And (fm1,fm2)) acc =
webertj@14999
   129
				cnf_string_acc fm1 (" 0\n" :: cnf_string_acc fm2 acc)
webertj@14999
   130
		in
webertj@16915
   131
			cnf_string_acc fm []
webertj@14999
   132
		end
webertj@16915
   133
		val fm'               = cnf_True_False_elim fm
webertj@16915
   134
		val number_of_vars    = maxidx fm'
webertj@16915
   135
		val number_of_clauses = cnf_number_of_clauses fm'
webertj@14453
   136
	in
webertj@14453
   137
		File.write path
webertj@16915
   138
			("c This file was generated by SatSolver.write_dimacs_cnf_file\n" ^
webertj@16915
   139
			 "p cnf " ^ string_of_int number_of_vars ^ " " ^ string_of_int number_of_clauses ^ "\n");
webertj@16915
   140
		File.append_list path
webertj@16915
   141
			(cnf_string fm');
webertj@16915
   142
		File.append path
webertj@16915
   143
			" 0\n"
webertj@14453
   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
webertj@14453
   153
	(* Path.T -> prop_formula -> unit *)
webertj@14453
   154
webertj@14453
   155
	fun write_dimacs_sat_file path fm =
webertj@14453
   156
	let
webertj@16915
   157
		(* prop_formula -> string list *)
webertj@14999
   158
		fun sat_string fm =
webertj@14999
   159
		let
webertj@14999
   160
			(* prop_formula -> string list -> string list *)
webertj@14999
   161
			fun sat_string_acc True acc =
webertj@14999
   162
				"*()" :: acc
webertj@14999
   163
			  | sat_string_acc False acc =
webertj@14999
   164
				"+()" :: acc
webertj@14999
   165
			  | sat_string_acc (BoolVar i) acc =
webertj@14999
   166
				(assert (i>=1) "formula contains a variable index less than 1";
webertj@14999
   167
				string_of_int i :: acc)
webertj@16899
   168
			  | sat_string_acc (Not (BoolVar i)) acc =
webertj@16899
   169
				"-" :: sat_string_acc (BoolVar i) acc
webertj@14999
   170
			  | sat_string_acc (Not fm) acc =
webertj@14999
   171
				"-(" :: sat_string_acc fm (")" :: acc)
webertj@14999
   172
			  | sat_string_acc (Or (fm1,fm2)) acc =
webertj@16899
   173
				"+(" :: sat_string_acc_or fm1 (" " :: sat_string_acc_or fm2 (")" :: acc))
webertj@14999
   174
			  | sat_string_acc (And (fm1,fm2)) acc =
webertj@16899
   175
				"*(" :: sat_string_acc_and fm1 (" " :: sat_string_acc_and fm2 (")" :: acc))
webertj@16899
   176
			(* optimization to make use of n-ary disjunction/conjunction *)
webertj@16899
   177
			(* prop_formula -> string list -> string list *)
webertj@16899
   178
			and sat_string_acc_or (Or (fm1,fm2)) acc =
webertj@16899
   179
				sat_string_acc_or fm1 (" " :: sat_string_acc_or fm2 acc)
webertj@16899
   180
			  | sat_string_acc_or fm acc =
webertj@16899
   181
				sat_string_acc fm acc
webertj@16899
   182
			(* prop_formula -> string list -> string list *)
webertj@16899
   183
			and sat_string_acc_and (And (fm1,fm2)) acc =
webertj@16899
   184
				sat_string_acc_and fm1 (" " :: sat_string_acc_and fm2 acc)
webertj@16899
   185
			  | sat_string_acc_and fm acc =
webertj@16899
   186
				sat_string_acc fm acc
webertj@14999
   187
		in
webertj@16915
   188
			sat_string_acc fm []
webertj@14999
   189
		end
webertj@16915
   190
		val number_of_vars = Int.max (maxidx fm, 1)
webertj@14453
   191
	in
webertj@14453
   192
		File.write path
webertj@16915
   193
			("c This file was generated by SatSolver.write_dimacs_sat_file\n" ^
webertj@16915
   194
			 "p sat " ^ string_of_int number_of_vars ^ "\n" ^
webertj@16915
   195
			 "(");
webertj@16915
   196
		File.append_list path
webertj@16915
   197
			(sat_string fm);
webertj@16915
   198
		File.append path
webertj@16915
   199
			")\n"
webertj@14453
   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
webertj@14965
   216
	(* Path.T -> string * string * string -> result *)
webertj@14453
   217
webertj@17620
   218
	fun read_std_result_file path (satisfiable, assignment_prefix, unsatisfiable) =
webertj@14453
   219
	let
webertj@14453
   220
		(* string -> int list *)
webertj@14453
   221
		fun int_list_from_string s =
webertj@15605
   222
			List.mapPartial Int.fromString (space_explode " " s)
webertj@14965
   223
		(* int list -> assignment *)
webertj@14453
   224
		fun assignment_from_list [] i =
skalberg@15531
   225
			NONE  (* the SAT solver didn't provide a value for this variable *)
webertj@14453
   226
		  | assignment_from_list (x::xs) i =
skalberg@15531
   227
			if x=i then (SOME true)
skalberg@15531
   228
			else if x=(~i) then (SOME false)
webertj@14453
   229
			else assignment_from_list xs i
webertj@14965
   230
		(* int list -> string list -> assignment *)
webertj@14965
   231
		fun parse_assignment xs [] =
webertj@14965
   232
			assignment_from_list xs
webertj@14965
   233
		  | parse_assignment xs (line::lines) =
webertj@14965
   234
			if String.isPrefix assignment_prefix line then
webertj@14965
   235
				parse_assignment (xs @ int_list_from_string line) lines
webertj@14965
   236
			else
webertj@14965
   237
				assignment_from_list xs
webertj@14453
   238
		(* string -> string -> bool *)
webertj@14453
   239
		fun is_substring needle haystack =
webertj@14453
   240
		let
webertj@14453
   241
			val length1 = String.size needle
webertj@14453
   242
			val length2 = String.size haystack
webertj@14453
   243
		in
webertj@14453
   244
			if length2 < length1 then
webertj@14453
   245
				false
webertj@14453
   246
			else if needle = String.substring (haystack, 0, length1) then
webertj@14453
   247
				true
webertj@14453
   248
			else is_substring needle (String.substring (haystack, 1, length2-1))
webertj@14453
   249
		end
webertj@14965
   250
		(* string list -> result *)
webertj@14453
   251
		fun parse_lines [] =
webertj@14965
   252
			UNKNOWN
webertj@14453
   253
		  | parse_lines (line::lines) =
webertj@20033
   254
			if is_substring unsatisfiable line then
webertj@20033
   255
				UNSATISFIABLE NONE
webertj@20033
   256
			else if is_substring satisfiable line then
webertj@14965
   257
				SATISFIABLE (parse_assignment [] lines)
webertj@14453
   258
			else
webertj@14453
   259
				parse_lines lines
webertj@14453
   260
	in
skalberg@15570
   261
		(parse_lines o (List.filter (fn l => l <> "")) o split_lines o File.read) path
webertj@14453
   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
webertj@14965
   268
	(* string -> (PropLogic.prop_formula -> unit) -> (unit -> result) -> solver *)
webertj@14453
   269
webertj@14453
   270
	fun make_external_solver cmd writefn readfn fm =
webertj@14965
   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
webertj@15040
   278
	(* Path.T -> PropLogic.prop_formula *)
webertj@15040
   279
webertj@15040
   280
	fun read_dimacs_cnf_file path =
webertj@15040
   281
	let
webertj@15040
   282
		(* string list -> string list *)
webertj@15040
   283
		fun filter_preamble [] =
webertj@15040
   284
			error "problem line not found in DIMACS CNF file"
webertj@15040
   285
		  | filter_preamble (line::lines) =
webertj@20441
   286
			if String.isPrefix "c " line orelse line = "c" then
webertj@15040
   287
				(* ignore comments *)
webertj@15040
   288
				filter_preamble lines
webertj@15040
   289
			else if String.isPrefix "p " line then
webertj@15040
   290
				(* ignore the problem line (which must be the last line of the preamble) *)
webertj@15040
   291
				(* Ignoring the problem line implies that if the file contains more clauses *)
webertj@15040
   292
				(* or variables than specified in its preamble, we will accept it anyway.   *)
webertj@15040
   293
				lines
webertj@15040
   294
			else
webertj@15040
   295
				error "preamble in DIMACS CNF file contains a line that does not begin with \"c \" or \"p \""
webertj@15040
   296
		(* string -> int *)
webertj@15040
   297
		fun int_from_string s =
webertj@15040
   298
			case Int.fromString s of
webertj@15040
   299
			  SOME i => i
webertj@15040
   300
			| NONE   => error ("token " ^ quote s ^ "in DIMACS CNF file is not a number")
webertj@15040
   301
		(* int list -> int list list *)
webertj@15040
   302
		fun clauses xs =
webertj@15040
   303
			let
webertj@15040
   304
				val (xs1, xs2) = take_prefix (fn i => i <> 0) xs
webertj@15040
   305
			in
webertj@15128
   306
				case xs2 of
webertj@15128
   307
				  []      => [xs1]
webertj@15128
   308
				| (0::[]) => [xs1]
webertj@15128
   309
				| (0::tl) => xs1 :: clauses tl
wenzelm@18678
   310
				| _       => sys_error "this cannot happen"
webertj@15040
   311
			end
webertj@15040
   312
		(* int -> PropLogic.prop_formula *)
webertj@15040
   313
		fun literal_from_int i =
webertj@15040
   314
			(assert (i<>0) "variable index in DIMACS CNF file is 0";
webertj@15040
   315
			if i>0 then
webertj@15040
   316
				PropLogic.BoolVar i
webertj@15040
   317
			else
webertj@15040
   318
				PropLogic.Not (PropLogic.BoolVar (~i)))
webertj@15040
   319
		(* PropLogic.prop_formula list -> PropLogic.prop_formula *)
webertj@15040
   320
		fun disjunction [] =
webertj@15040
   321
			error "empty clause in DIMACS CNF file"
webertj@15040
   322
		  | disjunction (x::xs) =
webertj@15040
   323
			(case xs of
webertj@15040
   324
			  [] => x
webertj@15040
   325
			| _  => PropLogic.Or (x, disjunction xs))
webertj@15040
   326
		(* PropLogic.prop_formula list -> PropLogic.prop_formula *)
webertj@15040
   327
		fun conjunction [] =
webertj@15040
   328
			error "no clause in DIMACS CNF file"
webertj@15040
   329
		  | conjunction (x::xs) =
webertj@15040
   330
			(case xs of
webertj@15040
   331
			  [] => x
webertj@15040
   332
			| _  => PropLogic.And (x, conjunction xs))
webertj@15040
   333
	in
webertj@15040
   334
		(conjunction
webertj@15040
   335
		o (map disjunction)
webertj@15040
   336
		o (map (map literal_from_int))
webertj@15040
   337
		o clauses
webertj@15040
   338
		o (map int_from_string)
skalberg@15570
   339
		o List.concat
webertj@15040
   340
		o (map (String.fields (fn c => c mem [#" ", #"\t", #"\n"])))
webertj@15040
   341
		o filter_preamble
skalberg@15570
   342
		o (List.filter (fn l => l <> ""))
webertj@15040
   343
		o split_lines
webertj@15040
   344
		o File.read)
webertj@15040
   345
			path
webertj@15040
   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
webertj@14487
   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
webertj@14453
   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
webertj@14453
   374
	(* string -> solver *)
webertj@14453
   375
webertj@14453
   376
	fun invoke_solver name =
haftmann@17541
   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
webertj@14453
   392
	fun enum_solver fm =
webertj@14453
   393
	let
webertj@14453
   394
		(* int list *)
webertj@14453
   395
		val indices = PropLogic.indices fm
webertj@14453
   396
		(* int list -> int list -> int list option *)
webertj@14453
   397
		(* binary increment: list 'xs' of current bits, list 'ys' of all bits (lower bits first) *)
webertj@14453
   398
		fun next_list _ ([]:int list) =
skalberg@15531
   399
			NONE
webertj@14453
   400
		  | next_list [] (y::ys) =
skalberg@15531
   401
			SOME [y]
webertj@14453
   402
		  | next_list (x::xs) (y::ys) =
webertj@14453
   403
			if x=y then
webertj@14453
   404
				(* reset the bit, continue *)
webertj@14453
   405
				next_list xs ys
webertj@14453
   406
			else
webertj@14453
   407
				(* set the lowest bit that wasn't set before, keep the higher bits *)
skalberg@15531
   408
				SOME (y::x::xs)
webertj@14453
   409
		(* int list -> int -> bool *)
webertj@14453
   410
		fun assignment_from_list xs i =
webertj@14453
   411
			i mem xs
webertj@14965
   412
		(* int list -> SatSolver.result *)
webertj@14453
   413
		fun solver_loop xs =
webertj@14453
   414
			if PropLogic.eval (assignment_from_list xs) fm then
skalberg@15531
   415
				SatSolver.SATISFIABLE (SOME o (assignment_from_list xs))
webertj@14453
   416
			else
webertj@14453
   417
				(case next_list xs indices of
skalberg@15531
   418
				  SOME xs' => solver_loop xs'
webertj@17493
   419
				| NONE     => SatSolver.UNSATISFIABLE NONE)
webertj@14453
   420
	in
webertj@14965
   421
		(* start with the "first" assignment (all variables are interpreted as 'false') *)
webertj@14453
   422
		solver_loop []
webertj@14453
   423
	end
webertj@14453
   424
in
webertj@14965
   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
webertj@14453
   435
	local
webertj@14453
   436
		open PropLogic
webertj@14453
   437
	in
webertj@14453
   438
		fun dpll_solver fm =
webertj@14453
   439
		let
webertj@16914
   440
			(* We could use 'PropLogic.defcnf fm' instead of 'PropLogic.nnf fm' *)
webertj@17493
   441
			(* but that sometimes leads to worse performance due to the         *)
webertj@17493
   442
			(* introduction of additional variables.                            *)
webertj@14514
   443
			val fm' = PropLogic.nnf fm
webertj@14453
   444
			(* int list *)
webertj@14514
   445
			val indices = PropLogic.indices fm'
webertj@14453
   446
			(* int list -> int -> prop_formula *)
webertj@14453
   447
			fun partial_var_eval []      i = BoolVar i
webertj@14453
   448
			  | partial_var_eval (x::xs) i = if x=i then True else if x=(~i) then False else partial_var_eval xs i
webertj@14453
   449
			(* int list -> prop_formula -> prop_formula *)
webertj@14453
   450
			fun partial_eval xs True             = True
webertj@14453
   451
			  | partial_eval xs False            = False
webertj@14453
   452
			  | partial_eval xs (BoolVar i)      = partial_var_eval xs i
webertj@14453
   453
			  | partial_eval xs (Not fm)         = SNot (partial_eval xs fm)
webertj@14453
   454
			  | partial_eval xs (Or (fm1, fm2))  = SOr (partial_eval xs fm1, partial_eval xs fm2)
webertj@14453
   455
			  | partial_eval xs (And (fm1, fm2)) = SAnd (partial_eval xs fm1, partial_eval xs fm2)
webertj@14453
   456
			(* prop_formula -> int list *)
webertj@14453
   457
			fun forced_vars True              = []
webertj@14453
   458
			  | forced_vars False             = []
webertj@14453
   459
			  | forced_vars (BoolVar i)       = [i]
webertj@14453
   460
			  | forced_vars (Not (BoolVar i)) = [~i]
webertj@14453
   461
			  | forced_vars (Or (fm1, fm2))   = (forced_vars fm1) inter_int (forced_vars fm2)
webertj@14453
   462
			  | forced_vars (And (fm1, fm2))  = (forced_vars fm1) union_int (forced_vars fm2)
webertj@14453
   463
			  (* Above, i *and* ~i may be forced.  In this case the first occurrence takes   *)
webertj@14453
   464
			  (* precedence, and the next partial evaluation of the formula returns 'False'. *)
wenzelm@18678
   465
			  | forced_vars _                 = error "formula is not in negation normal form"
webertj@14453
   466
			(* int list -> prop_formula -> (int list * prop_formula) *)
webertj@14453
   467
			fun eval_and_force xs fm =
webertj@14453
   468
			let
webertj@14453
   469
				val fm' = partial_eval xs fm
webertj@14453
   470
				val xs' = forced_vars fm'
webertj@14453
   471
			in
webertj@14453
   472
				if null xs' then
webertj@14453
   473
					(xs, fm')
webertj@14453
   474
				else
webertj@14453
   475
					eval_and_force (xs@xs') fm'  (* xs and xs' should be distinct, so '@' here should have *)
webertj@14453
   476
					                             (* the same effect as 'union_int'                         *)
webertj@14453
   477
			end
webertj@14453
   478
			(* int list -> int option *)
webertj@14453
   479
			fun fresh_var xs =
webertj@14453
   480
				Library.find_first (fn i => not (i mem_int xs) andalso not ((~i) mem_int xs)) indices
webertj@14453
   481
			(* int list -> prop_formula -> int list option *)
webertj@14453
   482
			(* partial assignment 'xs' *)
webertj@14453
   483
			fun dpll xs fm =
webertj@14453
   484
			let
webertj@14453
   485
				val (xs', fm') = eval_and_force xs fm
webertj@14453
   486
			in
webertj@14453
   487
				case fm' of
skalberg@15531
   488
				  True  => SOME xs'
skalberg@15531
   489
				| False => NONE
webertj@14453
   490
				| _     =>
webertj@14453
   491
					let
skalberg@15570
   492
						val x = valOf (fresh_var xs')  (* a fresh variable must exist since 'fm' did not evaluate to 'True'/'False' *)
webertj@14453
   493
					in
webertj@14453
   494
						case dpll (x::xs') fm' of  (* passing fm' rather than fm should save some simplification work *)
skalberg@15531
   495
						  SOME xs'' => SOME xs''
skalberg@15531
   496
						| NONE      => dpll ((~x)::xs') fm'  (* now try interpreting 'x' as 'False' *)
webertj@14453
   497
					end
webertj@14453
   498
			end
webertj@14965
   499
			(* int list -> assignment *)
webertj@14453
   500
			fun assignment_from_list [] i =
skalberg@15531
   501
				NONE  (* the DPLL procedure didn't provide a value for this variable *)
webertj@14453
   502
			  | assignment_from_list (x::xs) i =
skalberg@15531
   503
				if x=i then (SOME true)
skalberg@15531
   504
				else if x=(~i) then (SOME false)
webertj@14453
   505
				else assignment_from_list xs i
webertj@14453
   506
		in
webertj@14453
   507
			(* initially, no variable is interpreted yet *)
webertj@14965
   508
			case dpll [] fm' of
skalberg@15531
   509
			  SOME assignment => SatSolver.SATISFIABLE (assignment_from_list assignment)
webertj@17493
   510
			| NONE            => SatSolver.UNSATISFIABLE NONE
webertj@14453
   511
		end
webertj@14453
   512
	end  (* local *)
webertj@14453
   513
in
webertj@14965
   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
webertj@14453
   524
	fun auto_solver fm =
webertj@14965
   525
	let
webertj@14965
   526
		fun loop [] =
webertj@14965
   527
			SatSolver.UNKNOWN
webertj@14965
   528
		  | loop ((name, solver)::solvers) =
webertj@14487
   529
			if name="auto" then
webertj@14965
   530
				(* do not call solver "auto" from within "auto" *)
webertj@14965
   531
				loop solvers
webertj@15299
   532
			else (
webertj@14965
   533
				(if name="dpll" orelse name="enumerate" then
webertj@14805
   534
					warning ("Using SAT solver " ^ quote name ^ "; for better performance, consider using an external solver.")
webertj@14487
   535
				else
webertj@20135
   536
					tracing ("Using SAT solver " ^ quote name ^ "."));
webertj@14965
   537
				(* apply 'solver' to 'fm' *)
webertj@15299
   538
				solver fm
webertj@15299
   539
					handle SatSolver.NOT_CONFIGURED => loop solvers
webertj@14965
   540
			)
webertj@14965
   541
	in
webertj@17581
   542
		loop (!SatSolver.solvers)
webertj@14965
   543
	end
webertj@14453
   544
in
webertj@14453
   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
webertj@20135
   568
	exception INVALID_PROOF of string
webertj@20135
   569
	fun minisat_with_proofs fm =
webertj@20135
   570
	let
webertj@20135
   571
		val _          = if (getenv "MINISAT_HOME") = "" then raise SatSolver.NOT_CONFIGURED else ()
wenzelm@21858
   572
		val inpath     = File.tmp_path (Path.explode "isabelle.cnf")
wenzelm@21858
   573
		val outpath    = File.tmp_path (Path.explode "result")
wenzelm@21858
   574
		val proofpath  = File.tmp_path (Path.explode "result.prf")
wenzelm@21858
   575
		val cmd        = (getenv "MINISAT_HOME") ^ "/minisat " ^ (Path.implode inpath) ^ " -r " ^ (Path.implode outpath) ^ " -t " ^ (Path.implode proofpath) ^ "> /dev/null"
webertj@20135
   576
		fun writefn fm = SatSolver.write_dimacs_cnf_file inpath fm
webertj@20135
   577
		fun readfn ()  = SatSolver.read_std_result_file outpath ("SAT", "", "UNSAT")
wenzelm@21858
   578
		val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
wenzelm@21858
   579
		val _          = if File.exists outpath then warning ("overwriting existing file " ^ quote (Path.implode outpath)) else ()
webertj@20135
   580
		val cnf        = PropLogic.defcnf fm
webertj@20135
   581
		val result     = SatSolver.make_external_solver cmd writefn readfn cnf
webertj@20135
   582
		val _          = try File.rm inpath
webertj@20135
   583
		val _          = try File.rm outpath
webertj@20135
   584
	in  case result of
webertj@20135
   585
	  SatSolver.UNSATISFIABLE NONE =>
webertj@20135
   586
		(let
webertj@20152
   587
			(* string list *)
webertj@21267
   588
			val proof_lines = (split_lines o File.read) proofpath
webertj@20152
   589
				handle IO.Io _ => raise INVALID_PROOF "Could not read file \"result.prf\""
webertj@21267
   590
			(* representation of clauses as ordered lists of literals (with duplicates removed) *)
webertj@20135
   591
			(* prop_formula -> int list *)
webertj@21267
   592
			fun clause_to_lit_list (PropLogic.Or (fm1, fm2)) =
webertj@21267
   593
				OrdList.union int_ord (clause_to_lit_list fm1) (clause_to_lit_list fm2)
webertj@21267
   594
			  | clause_to_lit_list (PropLogic.BoolVar i) =
webertj@21267
   595
				[i]
webertj@21267
   596
			  | clause_to_lit_list (PropLogic.Not (PropLogic.BoolVar i)) =
webertj@21267
   597
				[~i]
webertj@21267
   598
			  | clause_to_lit_list _ =
webertj@21267
   599
				raise INVALID_PROOF "Error: invalid clause in CNF formula."
webertj@21267
   600
			(* prop_formula -> int *)
webertj@21267
   601
			fun cnf_number_of_clauses (PropLogic.And (fm1, fm2)) =
webertj@21267
   602
				cnf_number_of_clauses fm1 + cnf_number_of_clauses fm2
webertj@21267
   603
			  | cnf_number_of_clauses _ =
webertj@21267
   604
				1
webertj@21267
   605
			val number_of_clauses = cnf_number_of_clauses cnf
webertj@21267
   606
			(* int list array *)
webertj@21267
   607
			val clauses = Array.array (number_of_clauses, [])
webertj@21267
   608
			(* initialize the 'clauses' array *)
webertj@21267
   609
			(* prop_formula * int -> int *)
webertj@21267
   610
			fun init_array (PropLogic.And (fm1, fm2), n) =
webertj@21267
   611
				init_array (fm2, init_array (fm1, n))
webertj@21267
   612
			  | init_array (fm, n) =
webertj@21267
   613
				(Array.update (clauses, n, clause_to_lit_list fm); n+1)
webertj@21267
   614
			val _ = init_array (cnf, 0)
webertj@21267
   615
			(* optimization for the common case where MiniSat "R"s clauses in their *)
webertj@21267
   616
			(* original order:                                                      *)
webertj@21267
   617
			val last_ref_clause = ref (number_of_clauses - 1)
webertj@21267
   618
			(* search the 'clauses' array for the given list of literals 'lits', *)
webertj@21267
   619
			(* starting at index '!last_ref_clause + 1'                          *)
webertj@21267
   620
			(* int list -> int option *)
webertj@21267
   621
			fun original_clause_id lits =
webertj@21267
   622
			let
webertj@21267
   623
				fun original_clause_id_from index =
webertj@21267
   624
					if index = number_of_clauses then
webertj@21267
   625
						(* search from beginning again *)
webertj@21267
   626
						original_clause_id_from 0
webertj@21267
   627
					(* both 'lits' and the list of literals used in 'clauses' are sorted, so *)
webertj@21267
   628
					(* testing for equality should suffice -- barring duplicate literals     *)
webertj@21267
   629
					else if Array.sub (clauses, index) = lits then (
webertj@21267
   630
						(* success *)
webertj@21267
   631
						last_ref_clause := index;
webertj@21267
   632
						SOME index
webertj@21267
   633
					) else if index = !last_ref_clause then
webertj@21267
   634
						(* failure *)
webertj@21267
   635
						NONE
webertj@21267
   636
					else
webertj@21267
   637
						(* continue search *)
webertj@21267
   638
						original_clause_id_from (index + 1)
webertj@21267
   639
			in
webertj@21267
   640
				original_clause_id_from (!last_ref_clause + 1)
webertj@21267
   641
			end
webertj@20135
   642
			(* string -> int *)
webertj@20135
   643
			fun int_from_string s = (
webertj@20135
   644
				case Int.fromString s of
webertj@20135
   645
				  SOME i => i
webertj@20135
   646
				| NONE   => raise INVALID_PROOF ("File format error: number expected (" ^ quote s ^ " encountered).")
webertj@20135
   647
			)
webertj@20135
   648
			(* parse the proof file *)
webertj@21267
   649
			val clause_table  = ref (Inttab.empty : int list Inttab.table)
webertj@21267
   650
			val empty_id      = ref ~1
webertj@20135
   651
			(* contains a mapping from clause IDs as used by MiniSat to clause IDs in *)
webertj@20135
   652
			(* our proof format, where original clauses are numbered starting from 0  *)
webertj@20135
   653
			val clause_id_map = ref (Inttab.empty : int Inttab.table)
webertj@20135
   654
			fun sat_to_proof id = (
webertj@20135
   655
				case Inttab.lookup (!clause_id_map) id of
webertj@20135
   656
				  SOME id' => id'
webertj@20135
   657
				| NONE     => raise INVALID_PROOF ("Clause ID " ^ Int.toString id ^ " used, but not defined.")
webertj@20135
   658
			)
webertj@21267
   659
			val next_id = ref (number_of_clauses - 1)
webertj@21267
   660
			(* string list -> unit *)
webertj@21267
   661
			fun process_tokens [] =
webertj@21267
   662
				()
webertj@21267
   663
			  | process_tokens (tok::toks) =
webertj@20135
   664
				if tok="R" then (
webertj@20135
   665
					case toks of
webertj@20135
   666
					  id::sep::lits =>
webertj@20135
   667
						let
webertj@21267
   668
							val _        = if !empty_id = ~1 then () else raise INVALID_PROOF "File format error: \"R\" disallowed after \"X\"."
webertj@20135
   669
							val cid      = int_from_string id
webertj@20135
   670
							val _        = if sep = "<=" then () else raise INVALID_PROOF ("File format error: \"<=\" expected (" ^ quote sep ^ " encountered).")
webertj@20152
   671
							val ls       = sort int_ord (map int_from_string lits)
webertj@21267
   672
							val proof_id = case original_clause_id ls of
webertj@20135
   673
							                 SOME orig_id => orig_id
webertj@21267
   674
							               | NONE         => raise INVALID_PROOF ("Original clause (new ID is " ^ id ^ ") not found.")
webertj@21267
   675
						in
webertj@20135
   676
							(* extend the mapping of clause IDs with this newly defined ID *)
webertj@21267
   677
							clause_id_map := Inttab.update_new (cid, proof_id) (!clause_id_map)
webertj@21267
   678
								handle Inttab.DUP _ => raise INVALID_PROOF ("File format error: clause " ^ id ^ " defined more than once (in \"R\").")
webertj@20135
   679
							(* the proof itself doesn't change *)
webertj@20135
   680
						end
webertj@20135
   681
					| _ =>
webertj@20135
   682
						raise INVALID_PROOF "File format error: \"R\" followed by an insufficient number of tokens."
webertj@20135
   683
				) else if tok="C" then (
webertj@20135
   684
					case toks of
webertj@20135
   685
					  id::sep::ids =>
webertj@20135
   686
						let
webertj@21267
   687
							val _        = if !empty_id = ~1 then () else raise INVALID_PROOF "File format error: \"C\" disallowed after \"X\"."
webertj@20135
   688
							val cid      = int_from_string id
webertj@20135
   689
							val _        = if sep = "<=" then () else raise INVALID_PROOF ("File format error: \"<=\" expected (" ^ quote sep ^ " encountered).")
webertj@20135
   690
							(* ignore the pivot literals in MiniSat's trace *)
webertj@20135
   691
							fun unevens []             = raise INVALID_PROOF "File format error: \"C\" followed by an even number of IDs."
webertj@20135
   692
							  | unevens (x :: [])      = x :: []
webertj@20135
   693
							  | unevens (x :: _ :: xs) = x :: unevens xs
webertj@20152
   694
							val rs       = (map sat_to_proof o unevens o map int_from_string) ids
webertj@20135
   695
							(* extend the mapping of clause IDs with this newly defined ID *)
webertj@20135
   696
							val proof_id = inc next_id
webertj@20135
   697
							val _        = clause_id_map := Inttab.update_new (cid, proof_id) (!clause_id_map)
webertj@20135
   698
							                 handle Inttab.DUP _ => raise INVALID_PROOF ("File format error: clause " ^ id ^ " defined more than once (in \"C\").")
webertj@20135
   699
						in
webertj@20135
   700
							(* update clause table *)
webertj@21267
   701
							clause_table := Inttab.update_new (proof_id, rs) (!clause_table)
webertj@20135
   702
								handle Inttab.DUP _ => raise INVALID_PROOF ("Error: internal ID for clause " ^ id ^ " already used.")
webertj@20135
   703
						end
webertj@20135
   704
					| _ =>
webertj@20135
   705
						raise INVALID_PROOF "File format error: \"C\" followed by an insufficient number of tokens."
webertj@20135
   706
				) else if tok="D" then (
webertj@20135
   707
					case toks of
webertj@20135
   708
					  [id] =>
webertj@20135
   709
						let
webertj@21267
   710
							val _ = if !empty_id = ~1 then () else raise INVALID_PROOF "File format error: \"D\" disallowed after \"X\"."
webertj@20135
   711
							val _ = sat_to_proof (int_from_string id)
webertj@20135
   712
						in
webertj@20135
   713
							(* simply ignore "D" *)
webertj@21267
   714
							()
webertj@20135
   715
						end
webertj@20135
   716
					| _ =>
webertj@20135
   717
						raise INVALID_PROOF "File format error: \"D\" followed by an illegal number of tokens."
webertj@20135
   718
				) else if tok="X" then (
webertj@20135
   719
					case toks of
webertj@20135
   720
					  [id1, id2] =>
webertj@20135
   721
						let
webertj@21267
   722
							val _            = if !empty_id = ~1 then () else raise INVALID_PROOF "File format error: more than one end-of-proof statement."
webertj@20135
   723
							val _            = sat_to_proof (int_from_string id1)
webertj@20135
   724
							val new_empty_id = sat_to_proof (int_from_string id2)
webertj@20135
   725
						in
webertj@20135
   726
							(* update conflict id *)
webertj@21267
   727
							empty_id := new_empty_id
webertj@20135
   728
						end
webertj@20135
   729
					| _ =>
webertj@20135
   730
						raise INVALID_PROOF "File format error: \"X\" followed by an illegal number of tokens."
webertj@20135
   731
				) else
webertj@20135
   732
					raise INVALID_PROOF ("File format error: unknown token " ^ quote tok ^ " encountered.")
webertj@21267
   733
			(* string list -> unit *)
webertj@21267
   734
			fun process_lines [] =
webertj@21267
   735
				()
webertj@21267
   736
			  | process_lines (l::ls) = (
webertj@21267
   737
					process_tokens (String.tokens (fn c => c = #" " orelse c = #"\t") l);
webertj@21267
   738
					process_lines ls
webertj@21267
   739
				)
webertj@20135
   740
			(* proof *)
webertj@21267
   741
			val _ = process_lines proof_lines
webertj@21267
   742
			val _ = if !empty_id <> ~1 then () else raise INVALID_PROOF "File format error: no conflicting clause specified."
webertj@20135
   743
		in
webertj@21267
   744
			SatSolver.UNSATISFIABLE (SOME (!clause_table, !empty_id))
webertj@20152
   745
		end handle INVALID_PROOF reason => (warning reason; SatSolver.UNSATISFIABLE NONE))
webertj@20135
   746
	| result =>
webertj@20135
   747
		result
webertj@20135
   748
	end
webertj@20135
   749
in
webertj@20135
   750
	SatSolver.add_solver ("minisat_with_proofs", minisat_with_proofs)
webertj@20135
   751
end;
webertj@20135
   752
webertj@20033
   753
let
webertj@20033
   754
	fun minisat fm =
webertj@20033
   755
	let
webertj@20033
   756
		val _          = if (getenv "MINISAT_HOME") = "" then raise SatSolver.NOT_CONFIGURED else ()
wenzelm@21858
   757
		val inpath     = File.tmp_path (Path.explode "isabelle.cnf")
wenzelm@21858
   758
		val outpath    = File.tmp_path (Path.explode "result")
wenzelm@21858
   759
		val cmd        = (getenv "MINISAT_HOME") ^ "/minisat " ^ (Path.implode inpath) ^ " -r " ^ (Path.implode outpath) ^ " > /dev/null"
webertj@20033
   760
		fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
webertj@20033
   761
		fun readfn ()  = SatSolver.read_std_result_file outpath ("SAT", "", "UNSAT")
wenzelm@21858
   762
		val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
wenzelm@21858
   763
		val _          = if File.exists outpath then warning ("overwriting existing file " ^ quote (Path.implode outpath)) else ()
webertj@20033
   764
		val result     = SatSolver.make_external_solver cmd writefn readfn fm
webertj@20033
   765
		val _          = try File.rm inpath
webertj@20033
   766
		val _          = try File.rm outpath
webertj@20033
   767
	in
webertj@20033
   768
		result
webertj@20033
   769
	end
webertj@20033
   770
in
webertj@20033
   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
webertj@17493
   790
	exception INVALID_PROOF of string
webertj@17493
   791
	fun zchaff_with_proofs fm =
webertj@17493
   792
	case SatSolver.invoke_solver "zchaff" fm of
webertj@17493
   793
	  SatSolver.UNSATISFIABLE NONE =>
webertj@17493
   794
		(let
webertj@17493
   795
			(* string list *)
wenzelm@21858
   796
			val proof_lines = ((split_lines o File.read) (Path.explode "resolve_trace"))
webertj@20135
   797
				handle IO.Io _ => raise INVALID_PROOF "Could not read file \"resolve_trace\""
webertj@17493
   798
			(* PropLogic.prop_formula -> int *)
webertj@17493
   799
			fun cnf_number_of_clauses (PropLogic.And (fm1, fm2)) = cnf_number_of_clauses fm1 + cnf_number_of_clauses fm2
webertj@17493
   800
			  | cnf_number_of_clauses _                          = 1
webertj@17493
   801
			(* string -> int *)
webertj@17493
   802
			fun int_from_string s = (
webertj@17493
   803
				case Int.fromString s of
webertj@17493
   804
				  SOME i => i
webertj@17493
   805
				| NONE   => raise INVALID_PROOF ("File format error: number expected (" ^ quote s ^ " encountered).")
webertj@17493
   806
			)
webertj@17493
   807
			(* parse the "resolve_trace" file *)
webertj@17493
   808
			val clause_offset = ref ~1
webertj@20463
   809
			val clause_table  = ref (Inttab.empty : int list Inttab.table)
webertj@20463
   810
			val empty_id      = ref ~1
webertj@20463
   811
			(* string list -> unit *)
webertj@20463
   812
			fun process_tokens [] =
webertj@20463
   813
				()
webertj@20463
   814
			  | process_tokens (tok::toks) =
webertj@17493
   815
				if tok="CL:" then (
webertj@17493
   816
					case toks of
webertj@17493
   817
					  id::sep::ids =>
webertj@17493
   818
						let
webertj@17493
   819
							val _   = if !clause_offset = ~1 then () else raise INVALID_PROOF ("File format error: \"CL:\" disallowed after \"VAR:\".")
webertj@20463
   820
							val _   = if !empty_id = ~1 then () else raise INVALID_PROOF ("File format error: \"CL:\" disallowed after \"CONF:\".")
webertj@17493
   821
							val cid = int_from_string id
webertj@17493
   822
							val _   = if sep = "<=" then () else raise INVALID_PROOF ("File format error: \"<=\" expected (" ^ quote sep ^ " encountered).")
webertj@17493
   823
							val rs  = map int_from_string ids
webertj@17493
   824
						in
webertj@17493
   825
							(* update clause table *)
webertj@20463
   826
							clause_table := Inttab.update_new (cid, rs) (!clause_table)
webertj@17493
   827
								handle Inttab.DUP _ => raise INVALID_PROOF ("File format error: clause " ^ id ^ " defined more than once.")
webertj@17493
   828
						end
webertj@17493
   829
					| _ =>
webertj@17493
   830
						raise INVALID_PROOF "File format error: \"CL:\" followed by an insufficient number of tokens."
webertj@17493
   831
				) else if tok="VAR:" then (
webertj@17493
   832
					case toks of
webertj@17493
   833
					  id::levsep::levid::valsep::valid::antesep::anteid::litsep::lits =>
webertj@17493
   834
						let
webertj@20463
   835
							val _   = if !empty_id = ~1 then () else raise INVALID_PROOF ("File format error: \"VAR:\" disallowed after \"CONF:\".")
webertj@17493
   836
							(* set 'clause_offset' to the largest used clause ID *)
webertj@17493
   837
							val _   = if !clause_offset = ~1 then clause_offset :=
webertj@20463
   838
								(case Inttab.max_key (!clause_table) of
webertj@17493
   839
								  SOME id => id
webertj@17527
   840
								| NONE    => cnf_number_of_clauses (PropLogic.defcnf fm) - 1  (* the first clause ID is 0, not 1 *))
webertj@17493
   841
								else
webertj@17493
   842
									()
webertj@17493
   843
							val vid = int_from_string id
webertj@17493
   844
							val _   = if levsep = "L:" then () else raise INVALID_PROOF ("File format error: \"L:\" expected (" ^ quote levsep ^ " encountered).")
webertj@17493
   845
							val _   = int_from_string levid
webertj@17493
   846
							val _   = if valsep = "V:" then () else raise INVALID_PROOF ("File format error: \"V:\" expected (" ^ quote valsep ^ " encountered).")
webertj@17493
   847
							val _   = int_from_string valid
webertj@17493
   848
							val _   = if antesep = "A:" then () else raise INVALID_PROOF ("File format error: \"A:\" expected (" ^ quote antesep ^ " encountered).")
webertj@17493
   849
							val aid = int_from_string anteid
webertj@17493
   850
							val _   = if litsep = "Lits:" then () else raise INVALID_PROOF ("File format error: \"Lits:\" expected (" ^ quote litsep ^ " encountered).")
webertj@17493
   851
							val ls  = map int_from_string lits
webertj@17493
   852
							(* convert the data provided by zChaff to our resolution-style proof format *)
webertj@17493
   853
							(* each "VAR:" line defines a unit clause, the resolvents are implicitly    *)
webertj@17493
   854
							(* given by the literals in the antecedent clause                           *)
webertj@17493
   855
							(* we use the sum of '!clause_offset' and the variable ID as clause ID for the unit clause *)
webertj@17493
   856
							val cid = !clause_offset + vid
webertj@17493
   857
							(* the low bit of each literal gives its sign (positive/negative), therefore  *)
webertj@17493
   858
							(* we have to divide each literal by 2 to obtain the proper variable ID; then *)
webertj@17493
   859
							(* we add '!clause_offset' to obtain the ID of the corresponding unit clause  *)
webertj@17493
   860
							val vids = filter (not_equal vid) (map (fn l => l div 2) ls)
webertj@17493
   861
							val rs   = aid :: map (fn v => !clause_offset + v) vids
webertj@17493
   862
						in
webertj@17493
   863
							(* update clause table *)
webertj@20463
   864
							clause_table := Inttab.update_new (cid, rs) (!clause_table)
webertj@17527
   865
								handle Inttab.DUP _ => raise INVALID_PROOF ("File format error: clause " ^ string_of_int cid ^ " (derived from antecedent for variable " ^ id ^ ") already defined.")
webertj@17493
   866
						end
webertj@17493
   867
					| _ =>
webertj@17493
   868
						raise INVALID_PROOF "File format error: \"VAR:\" followed by an insufficient number of tokens."
webertj@17493
   869
				) else if tok="CONF:" then (
webertj@17493
   870
					case toks of
webertj@17493
   871
					  id::sep::ids =>
webertj@17493
   872
						let
webertj@20463
   873
							val _   = if !empty_id = ~1 then () else raise INVALID_PROOF "File format error: more than one conflicting clause specified."
webertj@17493
   874
							val cid = int_from_string id
webertj@17493
   875
							val _   = if sep = "==" then () else raise INVALID_PROOF ("File format error: \"==\" expected (" ^ quote sep ^ " encountered).")
webertj@17527
   876
							val ls  = map int_from_string ids
webertj@17527
   877
							(* the conflict clause must be resolved with the unit clauses *)
webertj@17527
   878
							(* for its literals to obtain the empty clause                *)
webertj@17621
   879
							val vids         = map (fn l => l div 2) ls
webertj@17621
   880
							val rs           = cid :: map (fn v => !clause_offset + v) vids
webertj@20463
   881
							val new_empty_id = getOpt (Inttab.max_key (!clause_table), !clause_offset) + 1
webertj@17493
   882
						in
webertj@17527
   883
							(* update clause table and conflict id *)
webertj@20463
   884
							clause_table := Inttab.update_new (new_empty_id, rs) (!clause_table)
webertj@20463
   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.");
webertj@20463
   886
							empty_id     := new_empty_id
webertj@17493
   887
						end
webertj@17493
   888
					| _ =>
webertj@17493
   889
						raise INVALID_PROOF "File format error: \"CONF:\" followed by an insufficient number of tokens."
webertj@17493
   890
				) else
webertj@17493
   891
					raise INVALID_PROOF ("File format error: unknown token " ^ quote tok ^ " encountered.")
webertj@20463
   892
			(* string list -> unit *)
webertj@20463
   893
			fun process_lines [] =
webertj@20463
   894
				()
webertj@20463
   895
			  | process_lines (l::ls) = (
webertj@20463
   896
					process_tokens (String.tokens (fn c => c = #" " orelse c = #"\t") l);
webertj@20463
   897
					process_lines ls
webertj@20463
   898
				)
webertj@17493
   899
			(* proof *)
webertj@21268
   900
			val _ = process_lines proof_lines
webertj@20463
   901
			val _ = if !empty_id <> ~1 then () else raise INVALID_PROOF "File format error: no conflicting clause specified."
webertj@17493
   902
		in
webertj@20463
   903
			SatSolver.UNSATISFIABLE (SOME (!clause_table, !empty_id))
webertj@17493
   904
		end handle INVALID_PROOF reason => (warning reason; SatSolver.UNSATISFIABLE NONE))
webertj@17493
   905
	| result =>
webertj@17493
   906
		result
webertj@17493
   907
in
webertj@17493
   908
	SatSolver.add_solver ("zchaff_with_proofs", zchaff_with_proofs)
webertj@17493
   909
end;
webertj@17493
   910
webertj@14487
   911
let
webertj@14487
   912
	fun zchaff fm =
webertj@14453
   913
	let
webertj@14965
   914
		val _          = if (getenv "ZCHAFF_HOME") = "" then raise SatSolver.NOT_CONFIGURED else ()
webertj@15332
   915
		val _          = if (getenv "ZCHAFF_VERSION") <> "2004.5.13" andalso
webertj@15332
   916
		                    (getenv "ZCHAFF_VERSION") <> "2004.11.15" then raise SatSolver.NOT_CONFIGURED else ()
webertj@15332
   917
			(* both versions of zChaff appear to have the same interface, so we do *)
webertj@15332
   918
			(* not actually need to distinguish between them in the following code *)
wenzelm@21858
   919
		val inpath     = File.tmp_path (Path.explode "isabelle.cnf")
wenzelm@21858
   920
		val outpath    = File.tmp_path (Path.explode "result")
wenzelm@21858
   921
		val cmd        = (getenv "ZCHAFF_HOME") ^ "/zchaff " ^ (Path.implode inpath) ^ " > " ^ (Path.implode outpath)
webertj@16914
   922
		fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
webertj@17620
   923
		fun readfn ()  = SatSolver.read_std_result_file outpath ("Instance Satisfiable", "", "Instance Unsatisfiable")
wenzelm@21858
   924
		val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
wenzelm@21858
   925
		val _          = if File.exists outpath then warning ("overwriting existing file " ^ quote (Path.implode outpath)) else ()
webertj@14965
   926
		val result     = SatSolver.make_external_solver cmd writefn readfn fm
wenzelm@16270
   927
		val _          = try File.rm inpath
wenzelm@16270
   928
		val _          = try File.rm outpath
webertj@14453
   929
	in
webertj@14965
   930
		result
webertj@14453
   931
	end
webertj@14487
   932
in
webertj@14965
   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
webertj@14965
   941
	fun berkmin fm =
webertj@14965
   942
	let
webertj@14965
   943
		val _          = if (getenv "BERKMIN_HOME") = "" orelse (getenv "BERKMIN_EXE") = "" then raise SatSolver.NOT_CONFIGURED else ()
wenzelm@21858
   944
		val inpath     = File.tmp_path (Path.explode "isabelle.cnf")
wenzelm@21858
   945
		val outpath    = File.tmp_path (Path.explode "result")
wenzelm@21858
   946
		val cmd        = (getenv "BERKMIN_HOME") ^ "/" ^ (getenv "BERKMIN_EXE") ^ " " ^ (Path.implode inpath) ^ " > " ^ (Path.implode outpath)
webertj@16914
   947
		fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
webertj@17620
   948
		fun readfn ()  = SatSolver.read_std_result_file outpath ("Satisfiable          !!", "solution =", "UNSATISFIABLE          !!")
wenzelm@21858
   949
		val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
wenzelm@21858
   950
		val _          = if File.exists outpath then warning ("overwriting existing file " ^ quote (Path.implode outpath)) else ()
webertj@14965
   951
		val result     = SatSolver.make_external_solver cmd writefn readfn fm
wenzelm@16270
   952
		val _          = try File.rm inpath
wenzelm@16270
   953
		val _          = try File.rm outpath
webertj@14965
   954
	in
webertj@14965
   955
		result
webertj@14965
   956
	end
webertj@14965
   957
in
webertj@14965
   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
webertj@14965
   966
	fun jerusat fm =
webertj@14965
   967
	let
webertj@14965
   968
		val _          = if (getenv "JERUSAT_HOME") = "" then raise SatSolver.NOT_CONFIGURED else ()
wenzelm@21858
   969
		val inpath     = File.tmp_path (Path.explode "isabelle.cnf")
wenzelm@21858
   970
		val outpath    = File.tmp_path (Path.explode "result")
wenzelm@21858
   971
		val cmd        = (getenv "JERUSAT_HOME") ^ "/Jerusat1.3 " ^ (Path.implode inpath) ^ " > " ^ (Path.implode outpath)
webertj@16914
   972
		fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
webertj@17620
   973
		fun readfn ()  = SatSolver.read_std_result_file outpath ("s SATISFIABLE", "v ", "s UNSATISFIABLE")
wenzelm@21858
   974
		val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
wenzelm@21858
   975
		val _          = if File.exists outpath then warning ("overwriting existing file " ^ quote (Path.implode outpath)) else ()
webertj@14965
   976
		val result     = SatSolver.make_external_solver cmd writefn readfn fm
wenzelm@16270
   977
		val _          = try File.rm inpath
wenzelm@16270
   978
		val _          = try File.rm outpath
webertj@14965
   979
	in
webertj@14965
   980
		result
webertj@14965
   981
	end
webertj@14965
   982
in
webertj@14965
   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
webertj@14487
   992
	fun mysolver fm = ...
webertj@14487
   993
in
webertj@14965
   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
*)