author | bulwahn |
Fri, 18 Mar 2011 18:19:42 +0100 | |
changeset 42031 | 2de57cda5b24 |
parent 39036 | dff91b90d74c |
child 46096 | a00685a18e55 |
permissions | -rw-r--r-- |
17618 | 1 |
(* Title: HOL/SAT.thy |
2 |
Author: Alwen Tiu, Tjark Weber |
|
3 |
Copyright 2005 |
|
4 |
||
17627 | 5 |
Basic setup for the 'sat' and 'satx' tactic. |
17618 | 6 |
*) |
7 |
||
8 |
header {* Reconstructing external resolution proofs for propositional logic *} |
|
9 |
||
26521 | 10 |
theory SAT |
11 |
imports Refute |
|
17722 | 12 |
uses |
26521 | 13 |
"Tools/sat_funcs.ML" |
17618 | 14 |
begin |
15 |
||
17722 | 16 |
text {* \medskip Late package setup: default values for refute, see |
21588 | 17 |
also theory @{theory Refute}. *} |
17722 | 18 |
|
19 |
refute_params |
|
20 |
["itself"=1, |
|
21 |
minsize=1, |
|
22 |
maxsize=8, |
|
23 |
maxvars=10000, |
|
24 |
maxtime=60, |
|
34120
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents:
32232
diff
changeset
|
25 |
satsolver="auto", |
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents:
32232
diff
changeset
|
26 |
no_assms="false"] |
17722 | 27 |
|
32232 | 28 |
ML {* structure sat = SATFunc(cnf) *} |
17618 | 29 |
|
32232 | 30 |
method_setup sat = {* Scan.succeed (SIMPLE_METHOD' o sat.sat_tac) *} |
17618 | 31 |
"SAT solver" |
32 |
||
32232 | 33 |
method_setup satx = {* Scan.succeed (SIMPLE_METHOD' o sat.satx_tac) *} |
17618 | 34 |
"SAT solver (with definitional CNF)" |
35 |
||
36 |
end |