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