author | wenzelm |
Sat, 03 Jan 2015 21:24:16 +0100 | |
changeset 59253 | 9448f4fc95e0 |
parent 56853 | a265e41cc33b |
child 60982 | 67e389f67073 |
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. |
|
51940 | 6 |
|
7 |
Relevant Isabelle environment settings: |
|
8 |
||
9 |
# MiniSat 1.14 |
|
10 |
#MINISAT_HOME=/usr/local/bin |
|
11 |
||
12 |
# zChaff |
|
13 |
#ZCHAFF_HOME=/usr/local/bin |
|
14 |
||
15 |
# BerkMin561 |
|
16 |
#BERKMIN_HOME=/usr/local/bin |
|
17 |
#BERKMIN_EXE=BerkMin561-linux |
|
18 |
#BERKMIN_EXE=BerkMin561-solaris |
|
19 |
||
20 |
# Jerusat 1.3 |
|
21 |
#JERUSAT_HOME=/usr/local/bin |
|
14453 | 22 |
*) |
23 |
||
24 |
signature SAT_SOLVER = |
|
25 |
sig |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
26 |
exception NOT_CONFIGURED |
14453 | 27 |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
28 |
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
|
29 |
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
|
30 |
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
|
31 |
| UNSATISFIABLE of proof option |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
32 |
| UNKNOWN |
41471 | 33 |
type solver = Prop_Logic.prop_formula -> result |
14965 | 34 |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
35 |
(* auxiliary functions to create external SAT solvers *) |
41471 | 36 |
val write_dimacs_cnf_file : Path.T -> Prop_Logic.prop_formula -> unit |
37 |
val write_dimacs_sat_file : Path.T -> Prop_Logic.prop_formula -> unit |
|
38 |
val read_std_result_file : Path.T -> string * string * string -> result |
|
39 |
val make_external_solver : string -> (Prop_Logic.prop_formula -> unit) -> |
|
40 |
(unit -> result) -> solver |
|
14453 | 41 |
|
41471 | 42 |
val read_dimacs_cnf_file : Path.T -> Prop_Logic.prop_formula |
15040 | 43 |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
44 |
(* generic solver interface *) |
56147 | 45 |
val get_solvers : unit -> (string * solver) list |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
46 |
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
|
47 |
val invoke_solver : string -> solver (* exception Option *) |
14453 | 48 |
end; |
49 |
||
56853 | 50 |
structure SAT_Solver : SAT_SOLVER = |
14453 | 51 |
struct |
52 |
||
41471 | 53 |
open Prop_Logic; |
14453 | 54 |
|
55 |
(* ------------------------------------------------------------------------- *) |
|
14965 | 56 |
(* should be raised by an external SAT solver to indicate that the solver is *) |
57 |
(* not configured properly *) |
|
58 |
(* ------------------------------------------------------------------------- *) |
|
59 |
||
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
60 |
exception NOT_CONFIGURED; |
14965 | 61 |
|
62 |
(* ------------------------------------------------------------------------- *) |
|
15531 | 63 |
(* type of partial (satisfying) assignments: 'a i = NONE' means that 'a' is *) |
19190 | 64 |
(* a satisfying assignment regardless of the value of variable 'i' *) |
14453 | 65 |
(* ------------------------------------------------------------------------- *) |
66 |
||
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
67 |
type assignment = int -> bool option; |
14965 | 68 |
|
69 |
(* ------------------------------------------------------------------------- *) |
|
17493
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
70 |
(* 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
|
71 |
(* 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
|
72 |
(* 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
|
73 |
(* 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
|
74 |
(* non-empty, and the literal to be resolved upon must always be unique *) |
17494 | 75 |
(* (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
|
76 |
(* 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
|
77 |
(* 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
|
78 |
(* 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
|
79 |
(* 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
|
80 |
(* 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
|
81 |
(* but do not need to be consecutive. *) |
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
82 |
(* ------------------------------------------------------------------------- *) |
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
83 |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
84 |
type proof = int list Inttab.table * int; |
17493
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
85 |
|
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
86 |
(* ------------------------------------------------------------------------- *) |
14965 | 87 |
(* 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
|
88 |
(* 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
|
89 |
(* 'UNSATISFIABLE', a proof of unsatisfiability may be returned *) |
14965 | 90 |
(* ------------------------------------------------------------------------- *) |
91 |
||
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
92 |
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
|
93 |
| UNSATISFIABLE of proof option |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
94 |
| UNKNOWN; |
14965 | 95 |
|
96 |
(* ------------------------------------------------------------------------- *) |
|
97 |
(* type of SAT solvers: given a propositional formula, a satisfying *) |
|
98 |
(* assignment may be returned *) |
|
99 |
(* ------------------------------------------------------------------------- *) |
|
100 |
||
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
101 |
type solver = prop_formula -> result; |
14453 | 102 |
|
103 |
(* ------------------------------------------------------------------------- *) |
|
104 |
(* write_dimacs_cnf_file: serializes a formula 'fm' of propositional logic *) |
|
105 |
(* to a file in DIMACS CNF format (see "Satisfiability Suggested *) |
|
106 |
(* Format", May 8 1993, Section 2.1) *) |
|
107 |
(* Note: 'fm' must not contain a variable index less than 1. *) |
|
14965 | 108 |
(* Note: 'fm' must be given in CNF. *) |
14453 | 109 |
(* ------------------------------------------------------------------------- *) |
110 |
||
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
111 |
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
|
112 |
let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
113 |
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
|
114 |
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
|
115 |
| 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
|
116 |
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
|
117 |
| 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
|
118 |
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
|
119 |
or 'fm' does not contain 'True'/'False' at all *) |
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 |
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
|
121 |
(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
|
122 |
| cnf_number_of_clauses _ = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
123 |
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
|
124 |
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
|
125 |
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
|
126 |
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
|
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 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
|
129 |
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
|
130 |
| 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
|
131 |
(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
|
132 |
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
|
133 |
| 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
|
134 |
(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
|
135 |
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
|
136 |
| 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
|
137 |
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
|
138 |
| 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
|
139 |
(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
|
140 |
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
|
141 |
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
|
142 |
| 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
|
143 |
(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
|
144 |
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
|
145 |
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
|
146 |
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
|
147 |
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
|
148 |
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
|
149 |
in |
56853 | 150 |
TextIO.output (out, "c This file was generated by SAT_Solver.write_dimacs_cnf_file\n"); |
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
|
151 |
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
|
152 |
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
|
153 |
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
|
154 |
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
|
155 |
end |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
156 |
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
|
157 |
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
|
158 |
end; |
14453 | 159 |
|
160 |
(* ------------------------------------------------------------------------- *) |
|
161 |
(* write_dimacs_sat_file: serializes a formula 'fm' of propositional logic *) |
|
162 |
(* to a file in DIMACS SAT format (see "Satisfiability Suggested *) |
|
163 |
(* Format", May 8 1993, Section 2.2) *) |
|
164 |
(* Note: 'fm' must not contain a variable index less than 1. *) |
|
165 |
(* ------------------------------------------------------------------------- *) |
|
166 |
||
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
167 |
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
|
168 |
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
|
169 |
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
|
170 |
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
|
171 |
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
|
172 |
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
|
173 |
| 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
|
174 |
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
|
175 |
| 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
|
176 |
(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
|
177 |
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
|
178 |
| 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
|
179 |
(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
|
180 |
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
|
181 |
| 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
|
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 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
|
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 (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
|
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_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
|
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_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
|
190 |
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
|
191 |
| 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
|
192 |
(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
|
193 |
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
|
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_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
|
196 |
TextIO.output (out, ")")) |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
197 |
(* 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
|
198 |
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
|
199 |
(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
|
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_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
|
202 |
| 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
|
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 |
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
|
205 |
(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
|
206 |
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
|
207 |
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
|
208 |
| 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
|
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 |
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
|
211 |
in |
56853 | 212 |
TextIO.output (out, "c This file was generated by SAT_Solver.write_dimacs_sat_file\n"); |
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 |
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
|
214 |
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
|
215 |
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
|
216 |
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
|
217 |
end |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
218 |
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
|
219 |
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
|
220 |
end; |
14453 | 221 |
|
222 |
(* ------------------------------------------------------------------------- *) |
|
17620
49568e5e0450
parse_std_result_file renamed to read_std_result_file
webertj
parents:
17581
diff
changeset
|
223 |
(* read_std_result_file: scans a SAT solver's output file for a satisfying *) |
14965 | 224 |
(* variable assignment. Returns the assignment, or 'UNSATISFIABLE' if *) |
225 |
(* the file contains 'unsatisfiable', or 'UNKNOWN' if the file contains *) |
|
226 |
(* neither 'satisfiable' nor 'unsatisfiable'. Empty lines are ignored. *) |
|
227 |
(* The assignment must be given in one or more lines immediately after *) |
|
228 |
(* the line that contains 'satisfiable'. These lines must begin with *) |
|
229 |
(* 'assignment_prefix'. Variables must be separated by " ". Non- *) |
|
230 |
(* integer strings are ignored. If variable i is contained in the *) |
|
231 |
(* assignment, then i is interpreted as 'true'. If ~i is contained in *) |
|
232 |
(* the assignment, then i is interpreted as 'false'. Otherwise the *) |
|
233 |
(* value of i is taken to be unspecified. *) |
|
14453 | 234 |
(* ------------------------------------------------------------------------- *) |
235 |
||
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
236 |
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
|
237 |
let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
238 |
fun int_list_from_string s = |
32952 | 239 |
map_filter Int.fromString (space_explode " " s) |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
240 |
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
|
241 |
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
|
242 |
| 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
|
243 |
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
|
244 |
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
|
245 |
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
|
246 |
fun parse_assignment xs [] = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
247 |
assignment_from_list xs |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
248 |
| 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
|
249 |
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
|
250 |
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
|
251 |
else |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
252 |
assignment_from_list xs |
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 |
fun parse_lines [] = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
265 |
UNKNOWN |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
266 |
| parse_lines (line::lines) = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
267 |
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
|
268 |
UNSATISFIABLE NONE |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
269 |
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
|
270 |
SATISFIABLE (parse_assignment [] lines) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
271 |
else |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
272 |
parse_lines lines |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
273 |
in |
33317 | 274 |
(parse_lines o filter (fn l => l <> "") o split_lines o File.read) path |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
275 |
end; |
14453 | 276 |
|
277 |
(* ------------------------------------------------------------------------- *) |
|
278 |
(* make_external_solver: call 'writefn', execute 'cmd', call 'readfn' *) |
|
279 |
(* ------------------------------------------------------------------------- *) |
|
280 |
||
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
281 |
fun make_external_solver cmd writefn readfn fm = |
43850
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
wenzelm
parents:
43701
diff
changeset
|
282 |
(writefn fm; Isabelle_System.bash cmd; readfn ()); |
14453 | 283 |
|
284 |
(* ------------------------------------------------------------------------- *) |
|
15040 | 285 |
(* read_dimacs_cnf_file: returns a propositional formula that corresponds to *) |
286 |
(* a SAT problem given in DIMACS CNF format *) |
|
287 |
(* ------------------------------------------------------------------------- *) |
|
288 |
||
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
289 |
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
|
290 |
let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
291 |
fun filter_preamble [] = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
292 |
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
|
293 |
| filter_preamble (line::lines) = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
294 |
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
|
295 |
(* ignore comments *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
296 |
filter_preamble lines |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
297 |
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
|
298 |
(* 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
|
299 |
(* 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
|
300 |
(* 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
|
301 |
lines |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
302 |
else |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
303 |
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
|
304 |
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
|
305 |
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
|
306 |
SOME i => i |
33937
b5ca587d0885
read_dimacs_cnf_file can now read DIMACS files that contain successive
webertj
parents:
33576
diff
changeset
|
307 |
| NONE => error ("token " ^ quote s ^ " in DIMACS CNF file is not a number") |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
308 |
fun clauses xs = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
309 |
let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
310 |
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
|
311 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
312 |
case xs2 of |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
313 |
[] => [xs1] |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
314 |
| (0::[]) => [xs1] |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
315 |
| (0::tl) => xs1 :: clauses tl |
56853 | 316 |
| _ => raise Fail "SAT_Solver.clauses" |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
317 |
end |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
318 |
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
|
319 |
(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
|
320 |
if i>0 then |
41471 | 321 |
Prop_Logic.BoolVar i |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
322 |
else |
41471 | 323 |
Prop_Logic.Not (Prop_Logic.BoolVar (~i))) |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
324 |
fun disjunction [] = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
325 |
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
|
326 |
| disjunction (x::xs) = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
327 |
(case xs of |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
328 |
[] => x |
41471 | 329 |
| _ => Prop_Logic.Or (x, disjunction xs)) |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
330 |
fun conjunction [] = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
331 |
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
|
332 |
| conjunction (x::xs) = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
333 |
(case xs of |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
334 |
[] => x |
41471 | 335 |
| _ => Prop_Logic.And (x, conjunction xs)) |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
336 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
337 |
(conjunction |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
338 |
o (map disjunction) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
339 |
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
|
340 |
o clauses |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
341 |
o (map int_from_string) |
36692
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents:
35011
diff
changeset
|
342 |
o (maps (String.tokens (member (op =) [#" ", #"\t", #"\n"]))) |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
343 |
o filter_preamble |
33317 | 344 |
o filter (fn l => l <> "") |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
345 |
o split_lines |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
346 |
o File.read) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
347 |
path |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
348 |
end; |
15040 | 349 |
|
350 |
(* ------------------------------------------------------------------------- *) |
|
56147 | 351 |
(* solvers: a table of all registered SAT solvers *) |
14453 | 352 |
(* ------------------------------------------------------------------------- *) |
353 |
||
56147 | 354 |
val solvers = Synchronized.var "solvers" ([] : (string * solver) list); |
355 |
||
356 |
fun get_solvers () = Synchronized.value solvers; |
|
14453 | 357 |
|
358 |
(* ------------------------------------------------------------------------- *) |
|
359 |
(* add_solver: updates 'solvers' by adding a new solver *) |
|
360 |
(* ------------------------------------------------------------------------- *) |
|
361 |
||
56147 | 362 |
fun add_solver (name, new_solver) = |
363 |
Synchronized.change solvers (fn the_solvers => |
|
22220 | 364 |
let |
365 |
val _ = if AList.defined (op =) the_solvers name |
|
366 |
then warning ("SAT solver " ^ quote name ^ " was defined before") |
|
367 |
else (); |
|
56147 | 368 |
in AList.update (op =) (name, new_solver) the_solvers end); |
14453 | 369 |
|
370 |
(* ------------------------------------------------------------------------- *) |
|
371 |
(* invoke_solver: returns the solver associated with the given 'name' *) |
|
15605 | 372 |
(* Note: If no solver is associated with 'name', exception 'Option' will be *) |
14453 | 373 |
(* raised. *) |
374 |
(* ------------------------------------------------------------------------- *) |
|
375 |
||
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
376 |
fun invoke_solver name = |
56147 | 377 |
the (AList.lookup (op =) (get_solvers ()) name); |
14453 | 378 |
|
56853 | 379 |
end; (* SAT_Solver *) |
14453 | 380 |
|
381 |
||
382 |
(* ------------------------------------------------------------------------- *) |
|
383 |
(* Predefined SAT solvers *) |
|
384 |
(* ------------------------------------------------------------------------- *) |
|
385 |
||
386 |
(* ------------------------------------------------------------------------- *) |
|
56853 | 387 |
(* Internal SAT solver, available as 'SAT_Solver.invoke_solver "cdclite"' -- *) |
56845 | 388 |
(* a simplified implementation of the conflict-driven clause-learning *) |
389 |
(* algorithm (cf. L. Zhang, S. Malik: "The Quest for Efficient Boolean *) |
|
390 |
(* Satisfiability Solvers", July 2002, Fig. 2). This solver produces models *) |
|
391 |
(* and proof traces. *) |
|
56815 | 392 |
(* ------------------------------------------------------------------------- *) |
393 |
||
394 |
let |
|
395 |
type clause = int list * int |
|
396 |
type value = bool option |
|
397 |
datatype reason = Decided | Implied of clause | Level0 of int |
|
398 |
type variable = bool option * reason * int * int |
|
399 |
type proofs = int * int list Inttab.table |
|
400 |
type state = |
|
401 |
int * int list * variable Inttab.table * clause list Inttab.table * proofs |
|
402 |
exception CONFLICT of clause * state |
|
403 |
exception UNSAT of clause * state |
|
404 |
||
405 |
fun neg i = ~i |
|
406 |
||
407 |
fun lit_value lit value = if lit > 0 then value else Option.map not value |
|
408 |
||
409 |
fun var_of vars lit: variable = the (Inttab.lookup vars (abs lit)) |
|
410 |
fun value_of vars lit = lit_value lit (#1 (var_of vars lit)) |
|
411 |
fun reason_of vars lit = #2 (var_of vars lit) |
|
412 |
fun level_of vars lit = #3 (var_of vars lit) |
|
413 |
||
414 |
fun is_true vars lit = (value_of vars lit = SOME true) |
|
415 |
fun is_false vars lit = (value_of vars lit = SOME false) |
|
416 |
fun is_unassigned vars lit = (value_of vars lit = NONE) |
|
417 |
fun assignment_of vars lit = the_default NONE (try (value_of vars) lit) |
|
418 |
||
419 |
fun put_var value reason level (_, _, _, rank) = (value, reason, level, rank) |
|
420 |
fun incr_rank (value, reason, level, rank) = (value, reason, level, rank + 1) |
|
421 |
fun update_var lit f = Inttab.map_entry (abs lit) f |
|
422 |
fun add_var lit = Inttab.update (abs lit, (NONE, Decided, ~1, 0)) |
|
423 |
||
424 |
fun assign lit r l = update_var lit (put_var (SOME (lit > 0)) r l) |
|
425 |
fun unassign lit = update_var lit (put_var NONE Decided ~1) |
|
426 |
||
427 |
fun add_proof [] (idx, ptab) = (idx, (idx + 1, ptab)) |
|
428 |
| add_proof ps (idx, ptab) = (idx, (idx + 1, Inttab.update (idx, ps) ptab)) |
|
429 |
||
430 |
fun level0_proof_of (Level0 idx) = SOME idx |
|
431 |
| level0_proof_of _ = NONE |
|
432 |
||
433 |
fun level0_proofs_of vars = map_filter (level0_proof_of o reason_of vars) |
|
434 |
fun prems_of vars (lits, p) = p :: level0_proofs_of vars lits |
|
435 |
fun mk_proof vars cls proofs = add_proof (prems_of vars cls) proofs |
|
436 |
||
437 |
fun push lit cls (level, trail, vars, clss, proofs) = |
|
438 |
let |
|
439 |
val (reason, proofs) = |
|
440 |
if level = 0 then apfst Level0 (mk_proof vars cls proofs) |
|
441 |
else (Implied cls, proofs) |
|
442 |
in (level, lit :: trail, assign lit reason level vars, clss, proofs) end |
|
443 |
||
444 |
fun push_decided lit (level, trail, vars, clss, proofs) = |
|
445 |
let val vars' = assign lit Decided (level + 1) vars |
|
446 |
in (level + 1, lit :: 0 :: trail, vars', clss, proofs) end |
|
447 |
||
448 |
fun prop (cls as (lits, _)) (cx as (units, state as (level, _, vars, _, _))) = |
|
449 |
if exists (is_true vars) lits then cx |
|
450 |
else if forall (is_false vars) lits then |
|
451 |
if level = 0 then raise UNSAT (cls, state) |
|
452 |
else raise CONFLICT (cls, state) |
|
453 |
else |
|
454 |
(case filter (is_unassigned vars) lits of |
|
455 |
[lit] => (lit :: units, push lit cls state) |
|
456 |
| _ => cx) |
|
457 |
||
458 |
fun propagate units (state as (_, _, _, clss, _)) = |
|
459 |
(case fold (fold prop o Inttab.lookup_list clss) units ([], state) of |
|
460 |
([], state') => (NONE, state') |
|
461 |
| (units', state') => propagate units' state') |
|
462 |
handle CONFLICT (cls, state') => (SOME cls, state') |
|
463 |
||
464 |
fun max_unassigned (v, (NONE, _, _, rank)) (x as (_, r)) = |
|
465 |
if rank > r then (SOME v, rank) else x |
|
466 |
| max_unassigned _ x = x |
|
467 |
||
468 |
fun decide (state as (_, _, vars, _, _)) = |
|
469 |
(case Inttab.fold max_unassigned vars (NONE, 0) of |
|
470 |
(SOME lit, _) => SOME (lit, push_decided lit state) |
|
471 |
| (NONE, _) => NONE) |
|
472 |
||
473 |
fun mark lit = Inttab.update (abs lit, true) |
|
474 |
fun marked ms lit = the_default false (Inttab.lookup ms (abs lit)) |
|
475 |
fun ignore l ms lit = ((lit = l) orelse marked ms lit) |
|
476 |
||
477 |
fun first_lit _ [] = raise Empty |
|
478 |
| first_lit _ (0 :: _) = raise Empty |
|
479 |
| first_lit pred (lit :: lits) = |
|
480 |
if pred lit then (lit, lits) else first_lit pred lits |
|
481 |
||
482 |
fun reason_cls_of vars lit = |
|
483 |
(case reason_of vars lit of |
|
484 |
Implied cls => cls |
|
485 |
| _ => raise Option) |
|
486 |
||
487 |
fun analyze conflicting_cls (level, trail, vars, _, _) = |
|
488 |
let |
|
489 |
fun back i lit (lits, p) trail ms ls ps = |
|
490 |
let |
|
491 |
val (lits0, lits') = List.partition (equal 0 o level_of vars) lits |
|
492 |
val lits1 = filter_out (ignore lit ms) lits' |
|
493 |
val lits2 = filter_out (equal level o level_of vars) lits1 |
|
494 |
val i' = length lits1 - length lits2 + i |
|
495 |
val ms' = fold mark lits1 ms |
|
496 |
val ls' = lits2 @ ls |
|
497 |
val ps' = level0_proofs_of vars lits0 @ (p :: ps) |
|
498 |
val (lit', trail') = first_lit (marked ms') trail |
|
499 |
in |
|
500 |
if i' = 1 then (neg lit', ls', rev ps') |
|
501 |
else back (i' - 1) lit' (reason_cls_of vars lit') trail' ms' ls' ps' |
|
502 |
end |
|
503 |
in back 0 0 conflicting_cls trail Inttab.empty [] [] end |
|
504 |
||
505 |
fun keep_clause (cls as (lits, _)) (level, trail, vars, clss, proofs) = |
|
506 |
let |
|
507 |
val vars' = fold (fn lit => update_var lit incr_rank) lits vars |
|
508 |
val clss' = fold (fn lit => Inttab.cons_list (neg lit, cls)) lits clss |
|
509 |
in (level, trail, vars', clss', proofs) end |
|
510 |
||
511 |
fun learn (cls as (lits, _)) = (length lits <= 2) ? keep_clause cls |
|
512 |
||
513 |
fun backjump _ (state as (_, [], _, _, _)) = state |
|
514 |
| backjump i (level, 0 :: trail, vars, clss, proofs) = |
|
515 |
(level - 1, trail, vars, clss, proofs) |> (i > 1) ? backjump (i - 1) |
|
516 |
| backjump i (level, lit :: trail, vars, clss, proofs) = |
|
517 |
backjump i (level, trail, unassign lit vars, clss, proofs) |
|
518 |
||
519 |
fun search units state = |
|
520 |
(case propagate units state of |
|
521 |
(NONE, state' as (_, _, vars, _, _)) => |
|
522 |
(case decide state' of |
|
56853 | 523 |
NONE => SAT_Solver.SATISFIABLE (assignment_of vars) |
56815 | 524 |
| SOME (lit, state'') => search [lit] state'') |
525 |
| (SOME conflicting_cls, state' as (level, trail, vars, clss, proofs)) => |
|
526 |
let |
|
527 |
val (lit, lits, ps) = analyze conflicting_cls state' |
|
528 |
val (idx, proofs') = add_proof ps proofs |
|
529 |
val cls = (lit :: lits, idx) |
|
530 |
in |
|
531 |
(level, trail, vars, clss, proofs') |
|
532 |
|> backjump (level - fold (Integer.max o level_of vars) lits 0) |
|
533 |
|> learn cls |
|
534 |
|> push lit cls |
|
535 |
|> search [lit] |
|
536 |
end) |
|
537 |
||
538 |
fun has_opposing_lits [] = false |
|
539 |
| has_opposing_lits (lit :: lits) = |
|
540 |
member (op =) lits (neg lit) orelse has_opposing_lits lits |
|
541 |
||
542 |
fun add_clause (cls as ([_], _)) (units, state) = |
|
543 |
let val (units', state') = prop cls (units, state) |
|
544 |
in (units', state') end |
|
545 |
| add_clause (cls as (lits, _)) (cx as (units, state)) = |
|
546 |
if has_opposing_lits lits then cx |
|
547 |
else (units, keep_clause cls state) |
|
548 |
||
549 |
fun mk_clause lits proofs = |
|
550 |
apfst (pair (distinct (op =) lits)) (add_proof [] proofs) |
|
551 |
||
552 |
fun solve litss = |
|
553 |
let |
|
554 |
val (clss, proofs) = fold_map mk_clause litss (0, Inttab.empty) |
|
555 |
val vars = fold (fold add_var) litss Inttab.empty |
|
556 |
val state = (0, [], vars, Inttab.empty, proofs) |
|
557 |
in uncurry search (fold add_clause clss ([], state)) end |
|
558 |
handle UNSAT (conflicting_cls, (_, _, vars, _, proofs)) => |
|
559 |
let val (idx, (_, ptab)) = mk_proof vars conflicting_cls proofs |
|
56853 | 560 |
in SAT_Solver.UNSATISFIABLE (SOME (ptab, idx)) end |
56815 | 561 |
|
562 |
fun variable_of (Prop_Logic.BoolVar 0) = error "bad propositional variable" |
|
563 |
| variable_of (Prop_Logic.BoolVar i) = i |
|
564 |
| variable_of _ = error "expected formula in CNF" |
|
565 |
fun literal_of (Prop_Logic.Not fm) = neg (variable_of fm) |
|
566 |
| literal_of fm = variable_of fm |
|
567 |
fun clause_of (Prop_Logic.Or (fm1, fm2)) = clause_of fm1 @ clause_of fm2 |
|
568 |
| clause_of fm = [literal_of fm] |
|
569 |
fun clauses_of (Prop_Logic.And (fm1, fm2)) = clauses_of fm1 @ clauses_of fm2 |
|
570 |
| clauses_of Prop_Logic.True = [[1, ~1]] |
|
571 |
| clauses_of Prop_Logic.False = [[1], [~1]] |
|
572 |
| clauses_of fm = [clause_of fm] |
|
573 |
||
574 |
fun dpll_solver fm = |
|
575 |
let val fm' = if Prop_Logic.is_cnf fm then fm else Prop_Logic.defcnf fm |
|
576 |
in solve (clauses_of fm') end |
|
577 |
in |
|
56853 | 578 |
SAT_Solver.add_solver ("cdclite", dpll_solver) |
56815 | 579 |
end; |
580 |
||
581 |
(* ------------------------------------------------------------------------- *) |
|
56853 | 582 |
(* Internal SAT solver, available as 'SAT_Solver.invoke_solver "auto"': uses *) |
17577
e87bf1d8f17a
zchaff_with_proofs does not delete zChaff\s resolve_trace file anymore
webertj
parents:
17541
diff
changeset
|
583 |
(* 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
|
584 |
(* 'NOT_CONFIGURED'. (However, the solver may return 'UNKNOWN'.) *) |
14453 | 585 |
(* ------------------------------------------------------------------------- *) |
586 |
||
587 |
let |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
588 |
fun auto_solver fm = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
589 |
let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
590 |
fun loop [] = |
56853 | 591 |
SAT_Solver.UNKNOWN |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
592 |
| loop ((name, solver)::solvers) = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
593 |
if name="auto" then |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
594 |
(* 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
|
595 |
loop solvers |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
596 |
else ( |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
597 |
(* apply 'solver' to 'fm' *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
598 |
solver fm |
56853 | 599 |
handle SAT_Solver.NOT_CONFIGURED => loop solvers |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
600 |
) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
601 |
in |
56853 | 602 |
loop (SAT_Solver.get_solvers ()) |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
603 |
end |
14453 | 604 |
in |
56853 | 605 |
SAT_Solver.add_solver ("auto", auto_solver) |
14453 | 606 |
end; |
607 |
||
608 |
(* ------------------------------------------------------------------------- *) |
|
20033 | 609 |
(* MiniSat 1.14 *) |
610 |
(* (http://www.cs.chalmers.se/Cs/Research/FormalMethods/MiniSat/) *) |
|
611 |
(* ------------------------------------------------------------------------- *) |
|
612 |
||
20135 | 613 |
(* ------------------------------------------------------------------------- *) |
614 |
(* "minisat_with_proofs" requires a modified version of MiniSat 1.14 by John *) |
|
615 |
(* Matthews, which can output ASCII proof traces. Replaying binary proof *) |
|
616 |
(* traces generated by MiniSat-p_v1.14 has _not_ been implemented. *) |
|
617 |
(* ------------------------------------------------------------------------- *) |
|
618 |
||
619 |
(* add "minisat_with_proofs" _before_ "minisat" to the available solvers, so *) |
|
620 |
(* that the latter is preferred by the "auto" solver *) |
|
621 |
||
20152
b6373fe199e1
MiniSat proof trace format changed; MiniSat is now expected to produce a proof also for "trivial" problems
webertj
parents:
20137
diff
changeset
|
622 |
(* 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
|
623 |
(* 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
|
624 |
(* 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
|
625 |
(* from 0 to n-1 (where n is the number of clauses in the formula). *) |
20135 | 626 |
|
627 |
let |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
628 |
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
|
629 |
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
|
630 |
let |
56853 | 631 |
val _ = if (getenv "MINISAT_HOME") = "" then raise SAT_Solver.NOT_CONFIGURED else () |
29872
14e208d607af
Added serial_string to SAT solver input and output files, to prevent multithreading chaos.
blanchet
parents:
22567
diff
changeset
|
632 |
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
|
633 |
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
|
634 |
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
|
635 |
val proofpath = File.tmp_path (Path.explode ("result" ^ serial_str ^ ".prf")) |
35011
9e55e87434ff
proper treatment of paths passed to the shell -- to allow spaces in file names as usual;
wenzelm
parents:
35010
diff
changeset
|
636 |
val cmd = getenv "MINISAT_HOME" ^ "/minisat " ^ File.shell_path inpath ^ " -r " ^ File.shell_path outpath ^ " -t " ^ File.shell_path proofpath ^ "> /dev/null" |
56853 | 637 |
fun writefn fm = SAT_Solver.write_dimacs_cnf_file inpath fm |
638 |
fun readfn () = SAT_Solver.read_std_result_file outpath ("SAT", "", "UNSAT") |
|
41944
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
wenzelm
parents:
41491
diff
changeset
|
639 |
val _ = if File.exists inpath then warning ("overwriting existing file " ^ Path.print inpath) else () |
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
wenzelm
parents:
41491
diff
changeset
|
640 |
val _ = if File.exists outpath then warning ("overwriting existing file " ^ Path.print outpath) else () |
41471 | 641 |
val cnf = Prop_Logic.defcnf fm |
56853 | 642 |
val result = SAT_Solver.make_external_solver cmd writefn readfn cnf |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
643 |
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
|
644 |
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
|
645 |
in case result of |
56853 | 646 |
SAT_Solver.UNSATISFIABLE NONE => |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
647 |
(let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
648 |
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
|
649 |
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
|
650 |
(* representation of clauses as ordered lists of literals (with duplicates removed) *) |
41471 | 651 |
fun clause_to_lit_list (Prop_Logic.Or (fm1, fm2)) = |
39687 | 652 |
Ord_List.union int_ord (clause_to_lit_list fm1) (clause_to_lit_list fm2) |
41471 | 653 |
| clause_to_lit_list (Prop_Logic.BoolVar i) = |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
654 |
[i] |
41471 | 655 |
| clause_to_lit_list (Prop_Logic.Not (Prop_Logic.BoolVar i)) = |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
656 |
[~i] |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
657 |
| clause_to_lit_list _ = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
658 |
raise INVALID_PROOF "Error: invalid clause in CNF formula." |
41471 | 659 |
fun cnf_number_of_clauses (Prop_Logic.And (fm1, fm2)) = |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
660 |
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
|
661 |
| cnf_number_of_clauses _ = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
662 |
1 |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
663 |
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
|
664 |
(* int list array *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
665 |
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
|
666 |
(* initialize the 'clauses' array *) |
41471 | 667 |
fun init_array (Prop_Logic.And (fm1, fm2), n) = |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
668 |
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
|
669 |
| init_array (fm, n) = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
670 |
(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
|
671 |
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
|
672 |
(* 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
|
673 |
(* original order: *) |
32740 | 674 |
val last_ref_clause = Unsynchronized.ref (number_of_clauses - 1) |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
675 |
(* 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
|
676 |
(* 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
|
677 |
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
|
678 |
let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
679 |
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
|
680 |
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
|
681 |
(* search from beginning again *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
682 |
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
|
683 |
(* 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
|
684 |
(* 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
|
685 |
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
|
686 |
(* success *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
687 |
last_ref_clause := index; |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
688 |
SOME index |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
689 |
) 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
|
690 |
(* failure *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
691 |
NONE |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
692 |
else |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
693 |
(* continue search *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
694 |
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
|
695 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
696 |
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
|
697 |
end |
55436
9781e17dcc23
removed odd comments -- inferred types are shown by Prover IDE;
wenzelm
parents:
52049
diff
changeset
|
698 |
fun int_from_string s = |
9781e17dcc23
removed odd comments -- inferred types are shown by Prover IDE;
wenzelm
parents:
52049
diff
changeset
|
699 |
(case Int.fromString s of |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
700 |
SOME i => i |
55436
9781e17dcc23
removed odd comments -- inferred types are shown by Prover IDE;
wenzelm
parents:
52049
diff
changeset
|
701 |
| NONE => raise INVALID_PROOF ("File format error: number expected (" ^ quote s ^ " encountered).")) |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
702 |
(* parse the proof file *) |
32740 | 703 |
val clause_table = Unsynchronized.ref (Inttab.empty : int list Inttab.table) |
704 |
val empty_id = Unsynchronized.ref ~1 |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
705 |
(* 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
|
706 |
(* our proof format, where original clauses are numbered starting from 0 *) |
32740 | 707 |
val clause_id_map = Unsynchronized.ref (Inttab.empty : int Inttab.table) |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
708 |
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
|
709 |
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
|
710 |
SOME id' => id' |
41491 | 711 |
| NONE => raise INVALID_PROOF ("Clause ID " ^ string_of_int id ^ " used, but not defined.") |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
712 |
) |
32740 | 713 |
val next_id = Unsynchronized.ref (number_of_clauses - 1) |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
714 |
fun process_tokens [] = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
715 |
() |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
716 |
| process_tokens (tok::toks) = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
717 |
if tok="R" then ( |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
718 |
case toks of |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
719 |
id::sep::lits => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
720 |
let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
721 |
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
|
722 |
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
|
723 |
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
|
724 |
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
|
725 |
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
|
726 |
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
|
727 |
| 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
|
728 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
729 |
(* 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
|
730 |
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
|
731 |
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
|
732 |
(* 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
|
733 |
end |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
734 |
| _ => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
735 |
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
|
736 |
) 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
|
737 |
case toks of |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
738 |
id::sep::ids => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
739 |
let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
740 |
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
|
741 |
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
|
742 |
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
|
743 |
(* 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
|
744 |
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
|
745 |
| unevens (x :: []) = x :: [] |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
746 |
| 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
|
747 |
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
|
748 |
(* extend the mapping of clause IDs with this newly defined ID *) |
32740 | 749 |
val proof_id = Unsynchronized.inc next_id |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
750 |
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
|
751 |
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
|
752 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
753 |
(* update clause table *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
754 |
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
|
755 |
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
|
756 |
end |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
757 |
| _ => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
758 |
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
|
759 |
) 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
|
760 |
case toks of |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
761 |
[id] => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
762 |
let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
763 |
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
|
764 |
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
|
765 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
766 |
(* simply ignore "D" *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
767 |
() |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
768 |
end |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
769 |
| _ => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
770 |
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
|
771 |
) 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
|
772 |
case toks of |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
773 |
[id1, id2] => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
774 |
let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
775 |
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
|
776 |
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
|
777 |
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
|
778 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
779 |
(* update conflict id *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
780 |
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
|
781 |
end |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
782 |
| _ => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
783 |
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
|
784 |
) else |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
785 |
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
|
786 |
fun process_lines [] = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
787 |
() |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
788 |
| process_lines (l::ls) = ( |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
789 |
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
|
790 |
process_lines ls |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
791 |
) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
792 |
(* proof *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
793 |
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
|
794 |
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
|
795 |
in |
56853 | 796 |
SAT_Solver.UNSATISFIABLE (SOME (!clause_table, !empty_id)) |
797 |
end handle INVALID_PROOF reason => (warning reason; SAT_Solver.UNSATISFIABLE NONE)) |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
798 |
| result => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
799 |
result |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
800 |
end |
20135 | 801 |
in |
56853 | 802 |
SAT_Solver.add_solver ("minisat_with_proofs", minisat_with_proofs) |
20135 | 803 |
end; |
804 |
||
20033 | 805 |
let |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
806 |
fun minisat fm = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
807 |
let |
56853 | 808 |
val _ = if getenv "MINISAT_HOME" = "" then raise SAT_Solver.NOT_CONFIGURED else () |
29872
14e208d607af
Added serial_string to SAT solver input and output files, to prevent multithreading chaos.
blanchet
parents:
22567
diff
changeset
|
809 |
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
|
810 |
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
|
811 |
val outpath = File.tmp_path (Path.explode ("result" ^ serial_str)) |
35011
9e55e87434ff
proper treatment of paths passed to the shell -- to allow spaces in file names as usual;
wenzelm
parents:
35010
diff
changeset
|
812 |
val cmd = getenv "MINISAT_HOME" ^ "/minisat " ^ File.shell_path inpath ^ " -r " ^ File.shell_path outpath ^ " > /dev/null" |
56853 | 813 |
fun writefn fm = SAT_Solver.write_dimacs_cnf_file inpath (Prop_Logic.defcnf fm) |
814 |
fun readfn () = SAT_Solver.read_std_result_file outpath ("SAT", "", "UNSAT") |
|
41944
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
wenzelm
parents:
41491
diff
changeset
|
815 |
val _ = if File.exists inpath then warning ("overwriting existing file " ^ Path.print inpath) else () |
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
wenzelm
parents:
41491
diff
changeset
|
816 |
val _ = if File.exists outpath then warning ("overwriting existing file " ^ Path.print outpath) else () |
56853 | 817 |
val result = SAT_Solver.make_external_solver cmd writefn readfn fm |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
818 |
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
|
819 |
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
|
820 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
821 |
result |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
822 |
end |
20033 | 823 |
in |
56853 | 824 |
SAT_Solver.add_solver ("minisat", minisat) |
20033 | 825 |
end; |
826 |
||
827 |
(* ------------------------------------------------------------------------- *) |
|
15332 | 828 |
(* zChaff (http://www.princeton.edu/~chaff/zchaff.html) *) |
14453 | 829 |
(* ------------------------------------------------------------------------- *) |
830 |
||
17493
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
831 |
(* ------------------------------------------------------------------------- *) |
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
832 |
(* '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
|
833 |
(* 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
|
834 |
(* 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
|
835 |
(* 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
|
836 |
(* 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
|
837 |
(* ------------------------------------------------------------------------- *) |
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
838 |
|
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
839 |
(* 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
|
840 |
(* that the latter is preferred by the "auto" solver *) |
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
841 |
|
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
842 |
let |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
843 |
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
|
844 |
fun zchaff_with_proofs fm = |
56853 | 845 |
case SAT_Solver.invoke_solver "zchaff" fm of |
846 |
SAT_Solver.UNSATISFIABLE NONE => |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
847 |
(let |
43701 | 848 |
(* FIXME File.tmp_path (!?) *) |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
849 |
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
|
850 |
handle IO.Io _ => raise INVALID_PROOF "Could not read file \"resolve_trace\"" |
41471 | 851 |
fun cnf_number_of_clauses (Prop_Logic.And (fm1, fm2)) = |
852 |
cnf_number_of_clauses fm1 + cnf_number_of_clauses fm2 |
|
853 |
| cnf_number_of_clauses _ = 1 |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
854 |
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
|
855 |
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
|
856 |
SOME i => i |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
857 |
| 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
|
858 |
) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
859 |
(* parse the "resolve_trace" file *) |
32740 | 860 |
val clause_offset = Unsynchronized.ref ~1 |
861 |
val clause_table = Unsynchronized.ref (Inttab.empty : int list Inttab.table) |
|
862 |
val empty_id = Unsynchronized.ref ~1 |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
863 |
fun process_tokens [] = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
864 |
() |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
865 |
| process_tokens (tok::toks) = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
866 |
if tok="CL:" then ( |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
867 |
case toks of |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
868 |
id::sep::ids => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
869 |
let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
870 |
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
|
871 |
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
|
872 |
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
|
873 |
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
|
874 |
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
|
875 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
876 |
(* update clause table *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
877 |
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
|
878 |
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
|
879 |
end |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
880 |
| _ => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
881 |
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
|
882 |
) 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
|
883 |
case toks of |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
884 |
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
|
885 |
let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
886 |
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
|
887 |
(* 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
|
888 |
val _ = if !clause_offset = ~1 then clause_offset := |
52049 | 889 |
(case Inttab.max (!clause_table) of |
890 |
SOME (id, _) => id |
|
41471 | 891 |
| NONE => cnf_number_of_clauses (Prop_Logic.defcnf fm) - 1 (* the first clause ID is 0, not 1 *)) |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
892 |
else |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
893 |
() |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
894 |
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
|
895 |
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
|
896 |
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
|
897 |
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
|
898 |
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
|
899 |
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
|
900 |
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
|
901 |
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
|
902 |
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
|
903 |
(* 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
|
904 |
(* 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
|
905 |
(* 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
|
906 |
(* 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
|
907 |
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
|
908 |
(* 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
|
909 |
(* 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
|
910 |
(* 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
|
911 |
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
|
912 |
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
|
913 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
914 |
(* update clause table *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
915 |
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
|
916 |
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
|
917 |
end |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
918 |
| _ => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
919 |
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
|
920 |
) 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
|
921 |
case toks of |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
922 |
id::sep::ids => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
923 |
let |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
924 |
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
|
925 |
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
|
926 |
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
|
927 |
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
|
928 |
(* 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
|
929 |
(* 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
|
930 |
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
|
931 |
val rs = cid :: map (fn v => !clause_offset + v) vids |
52049 | 932 |
val new_empty_id = the_default (!clause_offset) (Option.map fst (Inttab.max (!clause_table))) + 1 |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
933 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
934 |
(* 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
|
935 |
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
|
936 |
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
|
937 |
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
|
938 |
end |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
939 |
| _ => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
940 |
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
|
941 |
) else |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
942 |
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
|
943 |
fun process_lines [] = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
944 |
() |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
945 |
| process_lines (l::ls) = ( |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
946 |
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
|
947 |
process_lines ls |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
948 |
) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
949 |
(* proof *) |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
950 |
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
|
951 |
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
|
952 |
in |
56853 | 953 |
SAT_Solver.UNSATISFIABLE (SOME (!clause_table, !empty_id)) |
954 |
end handle INVALID_PROOF reason => (warning reason; SAT_Solver.UNSATISFIABLE NONE)) |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
955 |
| result => |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
956 |
result |
17493
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
957 |
in |
56853 | 958 |
SAT_Solver.add_solver ("zchaff_with_proofs", zchaff_with_proofs) |
17493
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
959 |
end; |
cf8713d880b1
SAT solver interface modified to support proofs of unsatisfiability
webertj
parents:
16915
diff
changeset
|
960 |
|
14487
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
webertj
parents:
14460
diff
changeset
|
961 |
let |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
962 |
fun zchaff fm = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
963 |
let |
56853 | 964 |
val _ = if getenv "ZCHAFF_HOME" = "" then raise SAT_Solver.NOT_CONFIGURED else () |
29872
14e208d607af
Added serial_string to SAT solver input and output files, to prevent multithreading chaos.
blanchet
parents:
22567
diff
changeset
|
965 |
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
|
966 |
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
|
967 |
val outpath = File.tmp_path (Path.explode ("result" ^ serial_str)) |
35011
9e55e87434ff
proper treatment of paths passed to the shell -- to allow spaces in file names as usual;
wenzelm
parents:
35010
diff
changeset
|
968 |
val cmd = getenv "ZCHAFF_HOME" ^ "/zchaff " ^ File.shell_path inpath ^ " > " ^ File.shell_path outpath |
56853 | 969 |
fun writefn fm = SAT_Solver.write_dimacs_cnf_file inpath (Prop_Logic.defcnf fm) |
970 |
fun readfn () = SAT_Solver.read_std_result_file outpath ("Instance Satisfiable", "", "Instance Unsatisfiable") |
|
41944
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
wenzelm
parents:
41491
diff
changeset
|
971 |
val _ = if File.exists inpath then warning ("overwriting existing file " ^ Path.print inpath) else () |
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
wenzelm
parents:
41491
diff
changeset
|
972 |
val _ = if File.exists outpath then warning ("overwriting existing file " ^ Path.print outpath) else () |
56853 | 973 |
val result = SAT_Solver.make_external_solver cmd writefn readfn fm |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
974 |
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
|
975 |
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
|
976 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
977 |
result |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
978 |
end |
14487
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
webertj
parents:
14460
diff
changeset
|
979 |
in |
56853 | 980 |
SAT_Solver.add_solver ("zchaff", zchaff) |
14965 | 981 |
end; |
982 |
||
983 |
(* ------------------------------------------------------------------------- *) |
|
984 |
(* BerkMin 561 (http://eigold.tripod.com/BerkMin.html) *) |
|
985 |
(* ------------------------------------------------------------------------- *) |
|
986 |
||
987 |
let |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
988 |
fun berkmin fm = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
989 |
let |
56853 | 990 |
val _ = if (getenv "BERKMIN_HOME") = "" then raise SAT_Solver.NOT_CONFIGURED else () |
29872
14e208d607af
Added serial_string to SAT solver input and output files, to prevent multithreading chaos.
blanchet
parents:
22567
diff
changeset
|
991 |
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
|
992 |
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
|
993 |
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
|
994 |
val exec = getenv "BERKMIN_EXE" |
35011
9e55e87434ff
proper treatment of paths passed to the shell -- to allow spaces in file names as usual;
wenzelm
parents:
35010
diff
changeset
|
995 |
val cmd = getenv "BERKMIN_HOME" ^ "/" ^ (if exec = "" then "BerkMin561" else exec) ^ " " ^ File.shell_path inpath ^ " > " ^ File.shell_path outpath |
56853 | 996 |
fun writefn fm = SAT_Solver.write_dimacs_cnf_file inpath (Prop_Logic.defcnf fm) |
997 |
fun readfn () = SAT_Solver.read_std_result_file outpath ("Satisfiable !!", "solution =", "UNSATISFIABLE !!") |
|
41944
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
wenzelm
parents:
41491
diff
changeset
|
998 |
val _ = if File.exists inpath then warning ("overwriting existing file " ^ Path.print inpath) else () |
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
wenzelm
parents:
41491
diff
changeset
|
999 |
val _ = if File.exists outpath then warning ("overwriting existing file " ^ Path.print outpath) else () |
56853 | 1000 |
val result = SAT_Solver.make_external_solver cmd writefn readfn fm |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
1001 |
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
|
1002 |
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
|
1003 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
1004 |
result |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
1005 |
end |
14965 | 1006 |
in |
56853 | 1007 |
SAT_Solver.add_solver ("berkmin", berkmin) |
14965 | 1008 |
end; |
1009 |
||
1010 |
(* ------------------------------------------------------------------------- *) |
|
1011 |
(* Jerusat 1.3 (http://www.cs.tau.ac.il/~ale1/) *) |
|
1012 |
(* ------------------------------------------------------------------------- *) |
|
1013 |
||
1014 |
let |
|
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
1015 |
fun jerusat fm = |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
1016 |
let |
56853 | 1017 |
val _ = if (getenv "JERUSAT_HOME") = "" then raise SAT_Solver.NOT_CONFIGURED else () |
29872
14e208d607af
Added serial_string to SAT solver input and output files, to prevent multithreading chaos.
blanchet
parents:
22567
diff
changeset
|
1018 |
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
|
1019 |
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
|
1020 |
val outpath = File.tmp_path (Path.explode ("result" ^ serial_str)) |
35011
9e55e87434ff
proper treatment of paths passed to the shell -- to allow spaces in file names as usual;
wenzelm
parents:
35010
diff
changeset
|
1021 |
val cmd = getenv "JERUSAT_HOME" ^ "/Jerusat1.3 " ^ File.shell_path inpath ^ " > " ^ File.shell_path outpath |
56853 | 1022 |
fun writefn fm = SAT_Solver.write_dimacs_cnf_file inpath (Prop_Logic.defcnf fm) |
1023 |
fun readfn () = SAT_Solver.read_std_result_file outpath ("s SATISFIABLE", "v ", "s UNSATISFIABLE") |
|
41944
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
wenzelm
parents:
41491
diff
changeset
|
1024 |
val _ = if File.exists inpath then warning ("overwriting existing file " ^ Path.print inpath) else () |
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
wenzelm
parents:
41491
diff
changeset
|
1025 |
val _ = if File.exists outpath then warning ("overwriting existing file " ^ Path.print outpath) else () |
56853 | 1026 |
val result = SAT_Solver.make_external_solver cmd writefn readfn fm |
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
1027 |
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
|
1028 |
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
|
1029 |
in |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
1030 |
result |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22220
diff
changeset
|
1031 |
end |
14965 | 1032 |
in |
56853 | 1033 |
SAT_Solver.add_solver ("jerusat", jerusat) |
14487
157d0ea7b2da
Installed solvers now determined at call time (as opposed to compile time)
webertj
parents:
14460
diff
changeset
|
1034 |
end; |