src/HOL/Tools/sat_solver.ML
author webertj
Wed, 21 Sep 2005 22:01:09 +0200
changeset 17577 e87bf1d8f17a
parent 17541 6a52083b71c3
child 17581 a50a53081808
permissions -rw-r--r--
zchaff_with_proofs does not delete zChaff\s resolve_trace file anymore
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
     1
(*  Title:      HOL/Tools/sat_solver.ML
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
     2
    ID:         $Id$
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
     3
    Author:     Tjark Weber
16911
20a139ca2f62 replaced calls to PropLogic.defcnf by PropLogic.auxcnf
webertj
parents: 16899
diff changeset
     4
    Copyright   2004-2005
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
     5
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
     6
Interface to external SAT solvers, and (simple) built-in SAT solvers.
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
     7
*)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
     8
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
     9
signature SAT_SOLVER =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    10
sig
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    11
	exception NOT_CONFIGURED
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    12
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    13
	type assignment = int -> bool option
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    14
	type proof      = int list Inttab.table * int
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    15
	datatype result = SATISFIABLE of assignment
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    16
	                | UNSATISFIABLE of proof option
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    17
	                | UNKNOWN
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    18
	type solver     = PropLogic.prop_formula -> result
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    19
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    20
	(* auxiliary functions to create external SAT solvers *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    21
	val write_dimacs_cnf_file : Path.T -> PropLogic.prop_formula -> unit
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    22
	val write_dimacs_sat_file : Path.T -> PropLogic.prop_formula -> unit
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    23
	val parse_std_result_file : Path.T -> string * string * string -> result
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    24
	val make_external_solver  : string -> (PropLogic.prop_formula -> unit) -> (unit -> result) -> solver
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    25
15040
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
    26
	val read_dimacs_cnf_file : Path.T -> PropLogic.prop_formula
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
    27
16899
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
    28
	(* generic solver interface *)
14487
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
    29
	val solvers       : (string * solver) list ref
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
    30
	val add_solver    : string * solver -> unit
15605
0c544d8b521f minor Library.option related modifications
webertj
parents: 15570
diff changeset
    31
	val invoke_solver : string -> solver  (* exception Option *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    32
end;
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    33
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    34
structure SatSolver : SAT_SOLVER =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    35
struct
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    36
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    37
	open PropLogic;
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    38
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    39
(* ------------------------------------------------------------------------- *)
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    40
(* should be raised by an external SAT solver to indicate that the solver is *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    41
(* not configured properly                                                   *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    42
(* ------------------------------------------------------------------------- *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    43
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    44
	exception NOT_CONFIGURED;
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    45
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    46
(* ------------------------------------------------------------------------- *)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
    47
(* type of partial (satisfying) assignments: 'a i = NONE' means that 'a' is  *)
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    48
(*      a satisfying assigment regardless of the value of variable 'i'       *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    49
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    50
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    51
	type assignment = int -> bool option;
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    52
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    53
(* ------------------------------------------------------------------------- *)
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    54
(* a proof of unsatisfiability, to be interpreted as follows: each integer   *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    55
(*      is a clause ID, each list 'xs' stored under the key 'x' in the table *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    56
(*      contains the IDs of clauses that must be resolved (in the given      *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    57
(*      order) to obtain the new clause 'x'.  Each list 'xs' must be         *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    58
(*      non-empty, and the literal to be resolved upon must always be unique *)
17494
e70600834f44 using curried Inttab.update_new function now
webertj
parents: 17493
diff changeset
    59
(*      (e.g. "A | ~B" must not be resolved with "~A | B").  Circular        *)
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    60
(*      dependencies of clauses are not allowed.  (At least) one of the      *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    61
(*      clauses in the table must be the empty clause (i.e. contain no       *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    62
(*      literals); its ID is given by the second component of the proof.     *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    63
(*      The clauses of the original problem passed to the SAT solver have    *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    64
(*      consecutive IDs starting with 0.  Clause IDs must be non-negative,   *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    65
(*      but do not need to be consecutive.                                   *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    66
(* ------------------------------------------------------------------------- *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    67
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    68
	type proof = int list Inttab.table * int;
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    69
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    70
(* ------------------------------------------------------------------------- *)
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    71
(* return type of SAT solvers: if the result is 'SATISFIABLE', a satisfying  *)
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    72
(*      assignment must be returned as well; if the result is                *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    73
(*      'UNSATISFIABLE', a proof of unsatisfiability may be returned         *)
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    74
(* ------------------------------------------------------------------------- *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    75
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    76
	datatype result = SATISFIABLE of assignment
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
    77
	                | UNSATISFIABLE of proof option
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    78
	                | UNKNOWN;
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    79
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    80
(* ------------------------------------------------------------------------- *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    81
(* type of SAT solvers: given a propositional formula, a satisfying          *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    82
(*      assignment may be returned                                           *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    83
(* ------------------------------------------------------------------------- *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    84
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    85
	type solver = prop_formula -> result;
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    86
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    87
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    88
(* write_dimacs_cnf_file: serializes a formula 'fm' of propositional logic   *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    89
(*      to a file in DIMACS CNF format (see "Satisfiability Suggested        *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    90
(*      Format", May 8 1993, Section 2.1)                                    *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    91
(* Note: 'fm' must not contain a variable index less than 1.                 *)
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
    92
(* Note: 'fm' must be given in CNF.                                          *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    93
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    94
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    95
	(* Path.T -> prop_formula -> unit *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    96
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    97
	fun write_dimacs_cnf_file path fm =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    98
	let
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
    99
		(* prop_formula -> prop_formula *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   100
		fun cnf_True_False_elim True =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   101
			Or (BoolVar 1, Not (BoolVar 1))
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   102
		  | cnf_True_False_elim False =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   103
			And (BoolVar 1, Not (BoolVar 1))
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   104
		  | cnf_True_False_elim fm =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   105
			fm  (* since 'fm' is in CNF, either 'fm'='True'/'False', or 'fm' does not contain 'True'/'False' at all *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   106
		(* prop_formula -> int *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   107
		fun cnf_number_of_clauses (And (fm1,fm2)) =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   108
			(cnf_number_of_clauses fm1) + (cnf_number_of_clauses fm2)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   109
		  | cnf_number_of_clauses _ =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   110
			1
16915
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   111
		(* prop_formula -> string list *)
14999
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   112
		fun cnf_string fm =
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   113
		let
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   114
			(* prop_formula -> string list -> string list *)
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   115
			fun cnf_string_acc True acc =
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   116
				error "formula is not in CNF"
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   117
			  | cnf_string_acc False acc =
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   118
				error "formula is not in CNF"
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   119
			  | cnf_string_acc (BoolVar i) acc =
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   120
				(assert (i>=1) "formula contains a variable index less than 1";
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   121
				string_of_int i :: acc)
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   122
			  | cnf_string_acc (Not (BoolVar i)) acc =
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   123
				"-" :: cnf_string_acc (BoolVar i) acc
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   124
			  | cnf_string_acc (Not _) acc =
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   125
				error "formula is not in CNF"
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   126
			  | cnf_string_acc (Or (fm1,fm2)) acc =
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   127
				cnf_string_acc fm1 (" " :: cnf_string_acc fm2 acc)
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   128
			  | cnf_string_acc (And (fm1,fm2)) acc =
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   129
				cnf_string_acc fm1 (" 0\n" :: cnf_string_acc fm2 acc)
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   130
		in
16915
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   131
			cnf_string_acc fm []
14999
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   132
		end
16915
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   133
		val fm'               = cnf_True_False_elim fm
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   134
		val number_of_vars    = maxidx fm'
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   135
		val number_of_clauses = cnf_number_of_clauses fm'
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   136
	in
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   137
		File.write path
16915
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   138
			("c This file was generated by SatSolver.write_dimacs_cnf_file\n" ^
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   139
			 "p cnf " ^ string_of_int number_of_vars ^ " " ^ string_of_int number_of_clauses ^ "\n");
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   140
		File.append_list path
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   141
			(cnf_string fm');
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   142
		File.append path
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   143
			" 0\n"
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   144
	end;
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   145
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   146
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   147
(* write_dimacs_sat_file: serializes a formula 'fm' of propositional logic   *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   148
(*      to a file in DIMACS SAT format (see "Satisfiability Suggested        *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   149
(*      Format", May 8 1993, Section 2.2)                                    *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   150
(* Note: 'fm' must not contain a variable index less than 1.                 *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   151
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   152
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   153
	(* Path.T -> prop_formula -> unit *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   154
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   155
	fun write_dimacs_sat_file path fm =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   156
	let
16915
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   157
		(* prop_formula -> string list *)
14999
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   158
		fun sat_string fm =
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   159
		let
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   160
			(* prop_formula -> string list -> string list *)
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   161
			fun sat_string_acc True acc =
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   162
				"*()" :: acc
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   163
			  | sat_string_acc False acc =
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   164
				"+()" :: acc
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   165
			  | sat_string_acc (BoolVar i) acc =
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   166
				(assert (i>=1) "formula contains a variable index less than 1";
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   167
				string_of_int i :: acc)
16899
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   168
			  | sat_string_acc (Not (BoolVar i)) acc =
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   169
				"-" :: sat_string_acc (BoolVar i) acc
14999
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   170
			  | sat_string_acc (Not fm) acc =
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   171
				"-(" :: sat_string_acc fm (")" :: acc)
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   172
			  | sat_string_acc (Or (fm1,fm2)) acc =
16899
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   173
				"+(" :: sat_string_acc_or fm1 (" " :: sat_string_acc_or fm2 (")" :: acc))
14999
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   174
			  | sat_string_acc (And (fm1,fm2)) acc =
16899
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   175
				"*(" :: sat_string_acc_and fm1 (" " :: sat_string_acc_and fm2 (")" :: acc))
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   176
			(* optimization to make use of n-ary disjunction/conjunction *)
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   177
			(* prop_formula -> string list -> string list *)
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   178
			and sat_string_acc_or (Or (fm1,fm2)) acc =
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   179
				sat_string_acc_or fm1 (" " :: sat_string_acc_or fm2 acc)
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   180
			  | sat_string_acc_or fm acc =
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   181
				sat_string_acc fm acc
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   182
			(* prop_formula -> string list -> string list *)
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   183
			and sat_string_acc_and (And (fm1,fm2)) acc =
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   184
				sat_string_acc_and fm1 (" " :: sat_string_acc_and fm2 acc)
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   185
			  | sat_string_acc_and fm acc =
ee4d620bcc1c write_dimacs_sat_file now generates slightly smaller files
webertj
parents: 16270
diff changeset
   186
				sat_string_acc fm acc
14999
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   187
		in
16915
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   188
			sat_string_acc fm []
14999
2c39efba8f67 faster conversion into DIMACS CNF and DIMACS SAT format
webertj
parents: 14965
diff changeset
   189
		end
16915
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   190
		val number_of_vars = Int.max (maxidx fm, 1)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   191
	in
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   192
		File.write path
16915
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   193
			("c This file was generated by SatSolver.write_dimacs_sat_file\n" ^
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   194
			 "p sat " ^ string_of_int number_of_vars ^ "\n" ^
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   195
			 "(");
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   196
		File.append_list path
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   197
			(sat_string fm);
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   198
		File.append path
bca4c3b1afca write_dimacs_sat_file and write_dimacs_cnf_file now write the file in chunks
webertj
parents: 16914
diff changeset
   199
			")\n"
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   200
	end;
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   201
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   202
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   203
(* parse_std_result_file: scans a SAT solver's output file for a satisfying  *)
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   204
(*      variable assignment.  Returns the assignment, or 'UNSATISFIABLE' if  *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   205
(*      the file contains 'unsatisfiable', or 'UNKNOWN' if the file contains *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   206
(*      neither 'satisfiable' nor 'unsatisfiable'.  Empty lines are ignored. *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   207
(*      The assignment must be given in one or more lines immediately after  *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   208
(*      the line that contains 'satisfiable'.  These lines must begin with   *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   209
(*      'assignment_prefix'.  Variables must be separated by " ".  Non-      *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   210
(*      integer strings are ignored.  If variable i is contained in the      *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   211
(*      assignment, then i is interpreted as 'true'.  If ~i is contained in  *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   212
(*      the assignment, then i is interpreted as 'false'.  Otherwise the     *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   213
(*      value of i is taken to be unspecified.                               *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   214
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   215
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   216
	(* Path.T -> string * string * string -> result *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   217
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   218
	fun parse_std_result_file path (satisfiable, assignment_prefix, unsatisfiable) =
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   219
	let
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   220
		(* string -> int list *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   221
		fun int_list_from_string s =
15605
0c544d8b521f minor Library.option related modifications
webertj
parents: 15570
diff changeset
   222
			List.mapPartial Int.fromString (space_explode " " s)
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   223
		(* int list -> assignment *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   224
		fun assignment_from_list [] i =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   225
			NONE  (* the SAT solver didn't provide a value for this variable *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   226
		  | assignment_from_list (x::xs) i =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   227
			if x=i then (SOME true)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   228
			else if x=(~i) then (SOME false)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   229
			else assignment_from_list xs i
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   230
		(* int list -> string list -> assignment *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   231
		fun parse_assignment xs [] =
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   232
			assignment_from_list xs
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   233
		  | parse_assignment xs (line::lines) =
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   234
			if String.isPrefix assignment_prefix line then
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   235
				parse_assignment (xs @ int_list_from_string line) lines
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   236
			else
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   237
				assignment_from_list xs
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   238
		(* string -> string -> bool *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   239
		fun is_substring needle haystack =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   240
		let
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   241
			val length1 = String.size needle
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   242
			val length2 = String.size haystack
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   243
		in
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   244
			if length2 < length1 then
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   245
				false
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   246
			else if needle = String.substring (haystack, 0, length1) then
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   247
				true
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   248
			else is_substring needle (String.substring (haystack, 1, length2-1))
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   249
		end
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   250
		(* string list -> result *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   251
		fun parse_lines [] =
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   252
			UNKNOWN
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   253
		  | parse_lines (line::lines) =
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   254
			if is_substring satisfiable line then
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   255
				SATISFIABLE (parse_assignment [] lines)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   256
			else if is_substring unsatisfiable line then
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   257
				UNSATISFIABLE NONE
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   258
			else
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   259
				parse_lines lines
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   260
	in
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   261
		(parse_lines o (List.filter (fn l => l <> "")) o split_lines o File.read) path
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   262
	end;
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   263
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   264
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   265
(* make_external_solver: call 'writefn', execute 'cmd', call 'readfn'        *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   266
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   267
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   268
	(* string -> (PropLogic.prop_formula -> unit) -> (unit -> result) -> solver *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   269
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   270
	fun make_external_solver cmd writefn readfn fm =
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   271
		(writefn fm; system cmd; readfn ());
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   272
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   273
(* ------------------------------------------------------------------------- *)
15040
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   274
(* read_dimacs_cnf_file: returns a propositional formula that corresponds to *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   275
(*      a SAT problem given in DIMACS CNF format                             *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   276
(* ------------------------------------------------------------------------- *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   277
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   278
	(* Path.T -> PropLogic.prop_formula *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   279
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   280
	fun read_dimacs_cnf_file path =
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   281
	let
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   282
		(* string list -> string list *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   283
		fun filter_preamble [] =
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   284
			error "problem line not found in DIMACS CNF file"
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   285
		  | filter_preamble (line::lines) =
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   286
			if String.isPrefix "c " line then
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   287
				(* ignore comments *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   288
				filter_preamble lines
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   289
			else if String.isPrefix "p " line then
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   290
				(* ignore the problem line (which must be the last line of the preamble) *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   291
				(* Ignoring the problem line implies that if the file contains more clauses *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   292
				(* or variables than specified in its preamble, we will accept it anyway.   *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   293
				lines
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   294
			else
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   295
				error "preamble in DIMACS CNF file contains a line that does not begin with \"c \" or \"p \""
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   296
		(* string -> int *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   297
		fun int_from_string s =
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   298
			case Int.fromString s of
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   299
			  SOME i => i
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   300
			| NONE   => error ("token " ^ quote s ^ "in DIMACS CNF file is not a number")
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   301
		(* int list -> int list list *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   302
		fun clauses xs =
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   303
			let
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   304
				val (xs1, xs2) = take_prefix (fn i => i <> 0) xs
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   305
			in
15128
da03f05815b0 bugfix in read_dimacs_cnf_file
webertj
parents: 15040
diff changeset
   306
				case xs2 of
da03f05815b0 bugfix in read_dimacs_cnf_file
webertj
parents: 15040
diff changeset
   307
				  []      => [xs1]
da03f05815b0 bugfix in read_dimacs_cnf_file
webertj
parents: 15040
diff changeset
   308
				| (0::[]) => [xs1]
da03f05815b0 bugfix in read_dimacs_cnf_file
webertj
parents: 15040
diff changeset
   309
				| (0::tl) => xs1 :: clauses tl
15329
bd94b0a71dd2 Removed a "Matches are not exhaustive" warning
webertj
parents: 15313
diff changeset
   310
				| _       => raise ERROR  (* this cannot happen *)
15040
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   311
			end
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   312
		(* int -> PropLogic.prop_formula *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   313
		fun literal_from_int i =
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   314
			(assert (i<>0) "variable index in DIMACS CNF file is 0";
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   315
			if i>0 then
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   316
				PropLogic.BoolVar i
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   317
			else
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   318
				PropLogic.Not (PropLogic.BoolVar (~i)))
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   319
		(* PropLogic.prop_formula list -> PropLogic.prop_formula *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   320
		fun disjunction [] =
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   321
			error "empty clause in DIMACS CNF file"
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   322
		  | disjunction (x::xs) =
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   323
			(case xs of
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   324
			  [] => x
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   325
			| _  => PropLogic.Or (x, disjunction xs))
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   326
		(* PropLogic.prop_formula list -> PropLogic.prop_formula *)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   327
		fun conjunction [] =
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   328
			error "no clause in DIMACS CNF file"
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   329
		  | conjunction (x::xs) =
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   330
			(case xs of
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   331
			  [] => x
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   332
			| _  => PropLogic.And (x, conjunction xs))
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   333
	in
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   334
		(conjunction
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   335
		o (map disjunction)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   336
		o (map (map literal_from_int))
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   337
		o clauses
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   338
		o (map int_from_string)
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   339
		o List.concat
15040
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   340
		o (map (String.fields (fn c => c mem [#" ", #"\t", #"\n"])))
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   341
		o filter_preamble
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   342
		o (List.filter (fn l => l <> ""))
15040
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   343
		o split_lines
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   344
		o File.read)
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   345
			path
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   346
	end;
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   347
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   348
(* ------------------------------------------------------------------------- *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   349
(* solvers: a (reference to a) table of all registered SAT solvers           *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   350
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   351
14487
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
   352
	val solvers = ref ([] : (string * solver) list);
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   353
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   354
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   355
(* add_solver: updates 'solvers' by adding a new solver                      *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   356
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   357
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   358
	(* string * solver -> unit *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   359
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   360
	fun add_solver (name, new_solver) =
17541
6a52083b71c3 introduced AList module
haftmann
parents: 17527
diff changeset
   361
		(solvers := update_warn (op =) ("SAT solver " ^ quote name ^ " was defined before") (name, new_solver) (!solvers));
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   362
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   363
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   364
(* invoke_solver: returns the solver associated with the given 'name'        *)
15605
0c544d8b521f minor Library.option related modifications
webertj
parents: 15570
diff changeset
   365
(* Note: If no solver is associated with 'name', exception 'Option' will be  *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   366
(*       raised.                                                             *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   367
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   368
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   369
	(* string -> solver *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   370
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   371
	fun invoke_solver name =
17541
6a52083b71c3 introduced AList module
haftmann
parents: 17527
diff changeset
   372
		(the o AList.lookup (op =) (!solvers)) name;
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   373
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   374
end;  (* SatSolver *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   375
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   376
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   377
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   378
(* Predefined SAT solvers                                                    *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   379
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   380
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   381
(* ------------------------------------------------------------------------- *)
14753
f40b45db8cf0 Comments fixed
webertj
parents: 14703
diff changeset
   382
(* Internal SAT solver, available as 'SatSolver.invoke_solver "enumerate"'   *)
f40b45db8cf0 Comments fixed
webertj
parents: 14703
diff changeset
   383
(* -- simply enumerates assignments until a satisfying assignment is found   *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   384
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   385
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   386
let
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   387
	fun enum_solver fm =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   388
	let
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   389
		(* int list *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   390
		val indices = PropLogic.indices fm
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   391
		(* int list -> int list -> int list option *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   392
		(* binary increment: list 'xs' of current bits, list 'ys' of all bits (lower bits first) *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   393
		fun next_list _ ([]:int list) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   394
			NONE
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   395
		  | next_list [] (y::ys) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   396
			SOME [y]
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   397
		  | next_list (x::xs) (y::ys) =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   398
			if x=y then
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   399
				(* reset the bit, continue *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   400
				next_list xs ys
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   401
			else
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   402
				(* set the lowest bit that wasn't set before, keep the higher bits *)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   403
				SOME (y::x::xs)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   404
		(* int list -> int -> bool *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   405
		fun assignment_from_list xs i =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   406
			i mem xs
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   407
		(* int list -> SatSolver.result *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   408
		fun solver_loop xs =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   409
			if PropLogic.eval (assignment_from_list xs) fm then
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   410
				SatSolver.SATISFIABLE (SOME o (assignment_from_list xs))
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   411
			else
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   412
				(case next_list xs indices of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   413
				  SOME xs' => solver_loop xs'
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   414
				| NONE     => SatSolver.UNSATISFIABLE NONE)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   415
	in
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   416
		(* start with the "first" assignment (all variables are interpreted as 'false') *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   417
		solver_loop []
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   418
	end
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   419
in
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   420
	SatSolver.add_solver ("enumerate", enum_solver)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   421
end;
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   422
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   423
(* ------------------------------------------------------------------------- *)
14753
f40b45db8cf0 Comments fixed
webertj
parents: 14703
diff changeset
   424
(* Internal SAT solver, available as 'SatSolver.invoke_solver "dpll"' -- a   *)
f40b45db8cf0 Comments fixed
webertj
parents: 14703
diff changeset
   425
(* simple implementation of the DPLL algorithm (cf. L. Zhang, S. Malik: "The *)
f40b45db8cf0 Comments fixed
webertj
parents: 14703
diff changeset
   426
(* Quest for Efficient Boolean Satisfiability Solvers", July 2002, Fig. 1).  *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   427
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   428
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   429
let
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   430
	local
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   431
		open PropLogic
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   432
	in
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   433
		fun dpll_solver fm =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   434
		let
16914
e68528b4fc0b replaced calls to PropLogic.auxcnf by PropLogic.defcnf again
webertj
parents: 16912
diff changeset
   435
			(* We could use 'PropLogic.defcnf fm' instead of 'PropLogic.nnf fm' *)
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   436
			(* but that sometimes leads to worse performance due to the         *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   437
			(* introduction of additional variables.                            *)
14514
15abb7d42e2e fixed dpll solver (now uses NNF)
webertj
parents: 14489
diff changeset
   438
			val fm' = PropLogic.nnf fm
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   439
			(* int list *)
14514
15abb7d42e2e fixed dpll solver (now uses NNF)
webertj
parents: 14489
diff changeset
   440
			val indices = PropLogic.indices fm'
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   441
			(* int list -> int -> prop_formula *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   442
			fun partial_var_eval []      i = BoolVar i
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   443
			  | partial_var_eval (x::xs) i = if x=i then True else if x=(~i) then False else partial_var_eval xs i
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   444
			(* int list -> prop_formula -> prop_formula *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   445
			fun partial_eval xs True             = True
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   446
			  | partial_eval xs False            = False
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   447
			  | partial_eval xs (BoolVar i)      = partial_var_eval xs i
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   448
			  | partial_eval xs (Not fm)         = SNot (partial_eval xs fm)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   449
			  | partial_eval xs (Or (fm1, fm2))  = SOr (partial_eval xs fm1, partial_eval xs fm2)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   450
			  | partial_eval xs (And (fm1, fm2)) = SAnd (partial_eval xs fm1, partial_eval xs fm2)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   451
			(* prop_formula -> int list *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   452
			fun forced_vars True              = []
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   453
			  | forced_vars False             = []
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   454
			  | forced_vars (BoolVar i)       = [i]
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   455
			  | forced_vars (Not (BoolVar i)) = [~i]
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   456
			  | forced_vars (Or (fm1, fm2))   = (forced_vars fm1) inter_int (forced_vars fm2)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   457
			  | forced_vars (And (fm1, fm2))  = (forced_vars fm1) union_int (forced_vars fm2)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   458
			  (* Above, i *and* ~i may be forced.  In this case the first occurrence takes   *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   459
			  (* precedence, and the next partial evaluation of the formula returns 'False'. *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   460
			  | forced_vars _                 = raise ERROR  (* formula is not in negation normal form *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   461
			(* int list -> prop_formula -> (int list * prop_formula) *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   462
			fun eval_and_force xs fm =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   463
			let
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   464
				val fm' = partial_eval xs fm
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   465
				val xs' = forced_vars fm'
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   466
			in
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   467
				if null xs' then
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   468
					(xs, fm')
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   469
				else
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   470
					eval_and_force (xs@xs') fm'  (* xs and xs' should be distinct, so '@' here should have *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   471
					                             (* the same effect as 'union_int'                         *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   472
			end
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   473
			(* int list -> int option *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   474
			fun fresh_var xs =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   475
				Library.find_first (fn i => not (i mem_int xs) andalso not ((~i) mem_int xs)) indices
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   476
			(* int list -> prop_formula -> int list option *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   477
			(* partial assignment 'xs' *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   478
			fun dpll xs fm =
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   479
			let
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   480
				val (xs', fm') = eval_and_force xs fm
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   481
			in
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   482
				case fm' of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   483
				  True  => SOME xs'
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   484
				| False => NONE
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   485
				| _     =>
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   486
					let
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   487
						val x = valOf (fresh_var xs')  (* a fresh variable must exist since 'fm' did not evaluate to 'True'/'False' *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   488
					in
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   489
						case dpll (x::xs') fm' of  (* passing fm' rather than fm should save some simplification work *)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   490
						  SOME xs'' => SOME xs''
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   491
						| NONE      => dpll ((~x)::xs') fm'  (* now try interpreting 'x' as 'False' *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   492
					end
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   493
			end
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   494
			(* int list -> assignment *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   495
			fun assignment_from_list [] i =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   496
				NONE  (* the DPLL procedure didn't provide a value for this variable *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   497
			  | assignment_from_list (x::xs) i =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   498
				if x=i then (SOME true)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   499
				else if x=(~i) then (SOME false)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   500
				else assignment_from_list xs i
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   501
		in
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   502
			(* initially, no variable is interpreted yet *)
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   503
			case dpll [] fm' of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15332
diff changeset
   504
			  SOME assignment => SatSolver.SATISFIABLE (assignment_from_list assignment)
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   505
			| NONE            => SatSolver.UNSATISFIABLE NONE
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   506
		end
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   507
	end  (* local *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   508
in
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   509
	SatSolver.add_solver ("dpll", dpll_solver)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   510
end;
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   511
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   512
(* ------------------------------------------------------------------------- *)
14753
f40b45db8cf0 Comments fixed
webertj
parents: 14703
diff changeset
   513
(* Internal SAT solver, available as 'SatSolver.invoke_solver "auto"': uses  *)
17577
e87bf1d8f17a zchaff_with_proofs does not delete zChaff\s resolve_trace file anymore
webertj
parents: 17541
diff changeset
   514
(* the last installed solver (other than "auto" itself) that does not raise  *)
15299
576fd0b65ed8 solver auto now returns the result of the first solver that does not raise NOT_CONFIGURED (which may be UNKNOWN)
webertj
parents: 15128
diff changeset
   515
(* 'NOT_CONFIGURED'.  (However, the solver may return 'UNKNOWN'.)            *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   516
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   517
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   518
let
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   519
	fun auto_solver fm =
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   520
	let
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   521
		fun loop [] =
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   522
			SatSolver.UNKNOWN
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   523
		  | loop ((name, solver)::solvers) =
14487
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
   524
			if name="auto" then
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   525
				(* do not call solver "auto" from within "auto" *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   526
				loop solvers
15299
576fd0b65ed8 solver auto now returns the result of the first solver that does not raise NOT_CONFIGURED (which may be UNKNOWN)
webertj
parents: 15128
diff changeset
   527
			else (
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   528
				(if name="dpll" orelse name="enumerate" then
14805
eff7b9df27e9 solver "auto" now issues a warning when it uses solver "enumerate"
webertj
parents: 14753
diff changeset
   529
					warning ("Using SAT solver " ^ quote name ^ "; for better performance, consider using an external solver.")
14487
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
   530
				else
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
   531
					());
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   532
				(* apply 'solver' to 'fm' *)
15299
576fd0b65ed8 solver auto now returns the result of the first solver that does not raise NOT_CONFIGURED (which may be UNKNOWN)
webertj
parents: 15128
diff changeset
   533
				solver fm
576fd0b65ed8 solver auto now returns the result of the first solver that does not raise NOT_CONFIGURED (which may be UNKNOWN)
webertj
parents: 15128
diff changeset
   534
					handle SatSolver.NOT_CONFIGURED => loop solvers
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   535
			)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   536
	in
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   537
		loop (rev (!SatSolver.solvers))
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   538
	end
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   539
in
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   540
	SatSolver.add_solver ("auto", auto_solver)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   541
end;
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   542
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   543
(* ------------------------------------------------------------------------- *)
15332
0dc05858a862 added ZCHAFF_VERSION
webertj
parents: 15329
diff changeset
   544
(* zChaff (http://www.princeton.edu/~chaff/zchaff.html)                      *)
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   545
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   546
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   547
(* ------------------------------------------------------------------------- *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   548
(* 'zchaff_with_proofs' applies the "zchaff" prover to a formula, and if     *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   549
(* zChaff finds that the formula is unsatisfiable, a proof of this is read   *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   550
(* from a file "resolve_trace" that was generated by zChaff.  See the code   *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   551
(* below for the expected format of the "resolve_trace" file.  Aside from    *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   552
(* some basic syntactic checks, no verification of the proof is performed.   *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   553
(* ------------------------------------------------------------------------- *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   554
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   555
(* add "zchaff_with_proofs" _before_ "zchaff" to the available solvers, so   *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   556
(* that the latter is preferred by the "auto" solver                         *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   557
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   558
let
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   559
	exception INVALID_PROOF of string
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   560
	fun zchaff_with_proofs fm =
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   561
	case SatSolver.invoke_solver "zchaff" fm of
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   562
	  SatSolver.UNSATISFIABLE NONE =>
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   563
		(let
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   564
			(* string list *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   565
			val proof_lines = ((split_lines o File.read) (Path.unpack "resolve_trace"))
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   566
				handle _ => raise INVALID_PROOF "Could not read file \"resolve_trace\""
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   567
			(* PropLogic.prop_formula -> int *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   568
			fun cnf_number_of_clauses (PropLogic.And (fm1, fm2)) = cnf_number_of_clauses fm1 + cnf_number_of_clauses fm2
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   569
			  | cnf_number_of_clauses _                          = 1
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   570
			(* string -> int *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   571
			fun int_from_string s = (
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   572
				case Int.fromString s of
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   573
				  SOME i => i
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   574
				| NONE   => raise INVALID_PROOF ("File format error: number expected (" ^ quote s ^ " encountered).")
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   575
			)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   576
			(* parse the "resolve_trace" file *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   577
			(* int ref *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   578
			val clause_offset = ref ~1
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   579
			(* string list -> proof -> proof *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   580
			fun process_tokens [] proof =
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   581
				proof
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   582
			  | process_tokens (tok::toks) (clause_table, conflict_id) =
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   583
				if tok="CL:" then (
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   584
					case toks of
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   585
					  id::sep::ids =>
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   586
						let
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   587
							val _   = if !clause_offset = ~1 then () else raise INVALID_PROOF ("File format error: \"CL:\" disallowed after \"VAR:\".")
17527
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   588
							val _   = if conflict_id = ~1 then () else raise INVALID_PROOF ("File format error: \"CL:\" disallowed after \"CONF:\".")
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   589
							val cid = int_from_string id
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   590
							val _   = if sep = "<=" then () else raise INVALID_PROOF ("File format error: \"<=\" expected (" ^ quote sep ^ " encountered).")
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   591
							val rs  = map int_from_string ids
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   592
						in
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   593
							(* update clause table *)
17494
e70600834f44 using curried Inttab.update_new function now
webertj
parents: 17493
diff changeset
   594
							(Inttab.update_new (cid, rs) clause_table, conflict_id)
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   595
								handle Inttab.DUP _ => raise INVALID_PROOF ("File format error: clause " ^ id ^ " defined more than once.")
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   596
						end
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   597
					| _ =>
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   598
						raise INVALID_PROOF "File format error: \"CL:\" followed by an insufficient number of tokens."
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   599
				) else if tok="VAR:" then (
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   600
					case toks of
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   601
					  id::levsep::levid::valsep::valid::antesep::anteid::litsep::lits =>
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   602
						let
17527
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   603
							val _   = if conflict_id = ~1 then () else raise INVALID_PROOF ("File format error: \"VAR:\" disallowed after \"CONF:\".")
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   604
							(* set 'clause_offset' to the largest used clause ID *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   605
							val _   = if !clause_offset = ~1 then clause_offset :=
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   606
								(case Inttab.max_key clause_table of
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   607
								  SOME id => id
17527
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   608
								| NONE    => cnf_number_of_clauses (PropLogic.defcnf fm) - 1  (* the first clause ID is 0, not 1 *))
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   609
								else
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   610
									()
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   611
							val vid = int_from_string id
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   612
							val _   = if levsep = "L:" then () else raise INVALID_PROOF ("File format error: \"L:\" expected (" ^ quote levsep ^ " encountered).")
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   613
							val _   = int_from_string levid
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   614
							val _   = if valsep = "V:" then () else raise INVALID_PROOF ("File format error: \"V:\" expected (" ^ quote valsep ^ " encountered).")
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   615
							val _   = int_from_string valid
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   616
							val _   = if antesep = "A:" then () else raise INVALID_PROOF ("File format error: \"A:\" expected (" ^ quote antesep ^ " encountered).")
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   617
							val aid = int_from_string anteid
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   618
							val _   = if litsep = "Lits:" then () else raise INVALID_PROOF ("File format error: \"Lits:\" expected (" ^ quote litsep ^ " encountered).")
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   619
							val ls  = map int_from_string lits
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   620
							(* convert the data provided by zChaff to our resolution-style proof format *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   621
							(* each "VAR:" line defines a unit clause, the resolvents are implicitly    *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   622
							(* given by the literals in the antecedent clause                           *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   623
							(* we use the sum of '!clause_offset' and the variable ID as clause ID for the unit clause *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   624
							val cid = !clause_offset + vid
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   625
							(* the low bit of each literal gives its sign (positive/negative), therefore  *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   626
							(* we have to divide each literal by 2 to obtain the proper variable ID; then *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   627
							(* we add '!clause_offset' to obtain the ID of the corresponding unit clause  *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   628
							val vids = filter (not_equal vid) (map (fn l => l div 2) ls)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   629
							val rs   = aid :: map (fn v => !clause_offset + v) vids
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   630
						in
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   631
							(* update clause table *)
17494
e70600834f44 using curried Inttab.update_new function now
webertj
parents: 17493
diff changeset
   632
							(Inttab.update_new (cid, rs) clause_table, conflict_id)
17527
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   633
								handle Inttab.DUP _ => raise INVALID_PROOF ("File format error: clause " ^ string_of_int cid ^ " (derived from antecedent for variable " ^ id ^ ") already defined.")
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   634
						end
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   635
					| _ =>
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   636
						raise INVALID_PROOF "File format error: \"VAR:\" followed by an insufficient number of tokens."
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   637
				) else if tok="CONF:" then (
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   638
					case toks of
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   639
					  id::sep::ids =>
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   640
						let
17527
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   641
							val _   = if conflict_id = ~1 then () else raise INVALID_PROOF "File format error: more than one conflicting clause specified."
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   642
							val cid = int_from_string id
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   643
							val _   = if sep = "==" then () else raise INVALID_PROOF ("File format error: \"==\" expected (" ^ quote sep ^ " encountered).")
17527
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   644
							val ls  = map int_from_string ids
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   645
							(* the conflict clause must be resolved with the unit clauses *)
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   646
							(* for its literals to obtain the empty clause                *)
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   647
							val vids     = map (fn l => l div 2) ls
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   648
							val rs       = cid :: map (fn v => !clause_offset + v) vids
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   649
							val empty_id = getOpt (Inttab.max_key clause_table, ~1) + 1
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   650
						in
17527
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   651
							(* update clause table and conflict id *)
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   652
							(Inttab.update_new (empty_id, rs) clause_table, empty_id)
5c25f27da4ca bugfix in "zchaff_with_proofs"
webertj
parents: 17494
diff changeset
   653
								handle Inttab.DUP _ => raise INVALID_PROOF ("File format error: clause " ^ string_of_int empty_id ^ " (empty clause derived from clause " ^ id ^ ") already defined.")
17493
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   654
						end
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   655
					| _ =>
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   656
						raise INVALID_PROOF "File format error: \"CONF:\" followed by an insufficient number of tokens."
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   657
				) else
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   658
					raise INVALID_PROOF ("File format error: unknown token " ^ quote tok ^ " encountered.")
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   659
			(* string list -> proof -> proof *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   660
			fun process_lines [] proof =
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   661
				proof
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   662
			  | process_lines (l::ls) proof =
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   663
				process_lines ls (process_tokens (String.tokens (fn c => c = #" " orelse c = #"\t") l) proof)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   664
			(* proof *)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   665
			val (clause_table, conflict_id) = process_lines proof_lines (Inttab.empty, ~1)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   666
			val _                           = if conflict_id <> ~1 then () else raise INVALID_PROOF "File format error: no conflicting clause specified."
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   667
		in
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   668
			SatSolver.UNSATISFIABLE (SOME (clause_table, conflict_id))
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   669
		end handle INVALID_PROOF reason => (warning reason; SatSolver.UNSATISFIABLE NONE))
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   670
	| result =>
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   671
		result
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   672
in
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   673
	SatSolver.add_solver ("zchaff_with_proofs", zchaff_with_proofs)
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   674
end;
cf8713d880b1 SAT solver interface modified to support proofs of unsatisfiability
webertj
parents: 16915
diff changeset
   675
14487
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
   676
let
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
   677
	fun zchaff fm =
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   678
	let
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   679
		val _          = if (getenv "ZCHAFF_HOME") = "" then raise SatSolver.NOT_CONFIGURED else ()
15332
0dc05858a862 added ZCHAFF_VERSION
webertj
parents: 15329
diff changeset
   680
		val _          = if (getenv "ZCHAFF_VERSION") <> "2004.5.13" andalso
0dc05858a862 added ZCHAFF_VERSION
webertj
parents: 15329
diff changeset
   681
		                    (getenv "ZCHAFF_VERSION") <> "2004.11.15" then raise SatSolver.NOT_CONFIGURED else ()
0dc05858a862 added ZCHAFF_VERSION
webertj
parents: 15329
diff changeset
   682
			(* both versions of zChaff appear to have the same interface, so we do *)
0dc05858a862 added ZCHAFF_VERSION
webertj
parents: 15329
diff changeset
   683
			(* not actually need to distinguish between them in the following code *)
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   684
		val inpath     = File.tmp_path (Path.unpack "isabelle.cnf")
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   685
		val outpath    = File.tmp_path (Path.unpack "result")
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   686
		val cmd        = (getenv "ZCHAFF_HOME") ^ "/zchaff " ^ (Path.pack inpath) ^ " > " ^ (Path.pack outpath)
16914
e68528b4fc0b replaced calls to PropLogic.auxcnf by PropLogic.defcnf again
webertj
parents: 16912
diff changeset
   687
		fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
15040
ed574b4f7e70 read_dimacs_cnf_file added
webertj
parents: 14999
diff changeset
   688
		fun readfn ()  = SatSolver.parse_std_result_file outpath ("Instance Satisfiable", "", "Instance Unsatisfiable")
16270
4280d6bbc1bb replaced File.sysify_path by Path.pack;
wenzelm
parents: 15605
diff changeset
   689
		val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.pack inpath)) else ()
4280d6bbc1bb replaced File.sysify_path by Path.pack;
wenzelm
parents: 15605
diff changeset
   690
		val _          = if File.exists outpath then warning ("overwriting existing file " ^ quote (Path.pack outpath)) else ()
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   691
		val result     = SatSolver.make_external_solver cmd writefn readfn fm
16270
4280d6bbc1bb replaced File.sysify_path by Path.pack;
wenzelm
parents: 15605
diff changeset
   692
		val _          = try File.rm inpath
4280d6bbc1bb replaced File.sysify_path by Path.pack;
wenzelm
parents: 15605
diff changeset
   693
		val _          = try File.rm outpath
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   694
	in
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   695
		result
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   696
	end
14487
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
   697
in
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   698
	SatSolver.add_solver ("zchaff", zchaff)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   699
end;
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   700
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   701
(* ------------------------------------------------------------------------- *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   702
(* BerkMin 561 (http://eigold.tripod.com/BerkMin.html)                       *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   703
(* ------------------------------------------------------------------------- *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   704
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   705
let
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   706
	fun berkmin fm =
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   707
	let
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   708
		val _          = if (getenv "BERKMIN_HOME") = "" orelse (getenv "BERKMIN_EXE") = "" then raise SatSolver.NOT_CONFIGURED else ()
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   709
		val inpath     = File.tmp_path (Path.unpack "isabelle.cnf")
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   710
		val outpath    = File.tmp_path (Path.unpack "result")
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   711
		val cmd        = (getenv "BERKMIN_HOME") ^ "/" ^ (getenv "BERKMIN_EXE") ^ " " ^ (Path.pack inpath) ^ " > " ^ (Path.pack outpath)
16914
e68528b4fc0b replaced calls to PropLogic.auxcnf by PropLogic.defcnf again
webertj
parents: 16912
diff changeset
   712
		fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   713
		fun readfn ()  = SatSolver.parse_std_result_file outpath ("Satisfiable          !!", "solution =", "UNSATISFIABLE          !!")
16270
4280d6bbc1bb replaced File.sysify_path by Path.pack;
wenzelm
parents: 15605
diff changeset
   714
		val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.pack inpath)) else ()
4280d6bbc1bb replaced File.sysify_path by Path.pack;
wenzelm
parents: 15605
diff changeset
   715
		val _          = if File.exists outpath then warning ("overwriting existing file " ^ quote (Path.pack outpath)) else ()
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   716
		val result     = SatSolver.make_external_solver cmd writefn readfn fm
16270
4280d6bbc1bb replaced File.sysify_path by Path.pack;
wenzelm
parents: 15605
diff changeset
   717
		val _          = try File.rm inpath
4280d6bbc1bb replaced File.sysify_path by Path.pack;
wenzelm
parents: 15605
diff changeset
   718
		val _          = try File.rm outpath
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   719
	in
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   720
		result
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   721
	end
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   722
in
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   723
	SatSolver.add_solver ("berkmin", berkmin)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   724
end;
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   725
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   726
(* ------------------------------------------------------------------------- *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   727
(* Jerusat 1.3 (http://www.cs.tau.ac.il/~ale1/)                              *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   728
(* ------------------------------------------------------------------------- *)
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   729
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   730
let
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   731
	fun jerusat fm =
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   732
	let
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   733
		val _          = if (getenv "JERUSAT_HOME") = "" then raise SatSolver.NOT_CONFIGURED else ()
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   734
		val inpath     = File.tmp_path (Path.unpack "isabelle.cnf")
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   735
		val outpath    = File.tmp_path (Path.unpack "result")
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   736
		val cmd        = (getenv "JERUSAT_HOME") ^ "/Jerusat1.3 " ^ (Path.pack inpath) ^ " > " ^ (Path.pack outpath)
16914
e68528b4fc0b replaced calls to PropLogic.auxcnf by PropLogic.defcnf again
webertj
parents: 16912
diff changeset
   737
		fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   738
		fun readfn ()  = SatSolver.parse_std_result_file outpath ("s SATISFIABLE", "v ", "s UNSATISFIABLE")
16270
4280d6bbc1bb replaced File.sysify_path by Path.pack;
wenzelm
parents: 15605
diff changeset
   739
		val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.pack inpath)) else ()
4280d6bbc1bb replaced File.sysify_path by Path.pack;
wenzelm
parents: 15605
diff changeset
   740
		val _          = if File.exists outpath then warning ("overwriting existing file " ^ quote (Path.pack outpath)) else ()
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   741
		val result     = SatSolver.make_external_solver cmd writefn readfn fm
16270
4280d6bbc1bb replaced File.sysify_path by Path.pack;
wenzelm
parents: 15605
diff changeset
   742
		val _          = try File.rm inpath
4280d6bbc1bb replaced File.sysify_path by Path.pack;
wenzelm
parents: 15605
diff changeset
   743
		val _          = try File.rm outpath
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   744
	in
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   745
		result
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   746
	end
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   747
in
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   748
	SatSolver.add_solver ("jerusat", jerusat)
14487
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
   749
end;
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   750
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   751
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   752
(* Add code for other SAT solvers below.                                     *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   753
(* ------------------------------------------------------------------------- *)
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   754
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   755
(*
14487
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
   756
let
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
   757
	fun mysolver fm = ...
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
   758
in
14965
7155b319eafa new SAT solver interface
webertj
parents: 14805
diff changeset
   759
	SatSolver.add_solver ("myname", (fn fm => if mysolver_is_configured then mysolver fm else raise SatSolver.NOT_CONFIGURED));  -- register the solver
14487
157d0ea7b2da Installed solvers now determined at call time (as opposed to compile time)
webertj
parents: 14460
diff changeset
   760
end;
14453
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   761
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   762
-- the solver is now available as SatSolver.invoke_solver "myname"
3397a69dfa4e Internal and external SAT solvers
webertj
parents:
diff changeset
   763
*)