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/cnf_funcs.ML"
|
|
14 |
"Tools/sat_funcs.ML"
|
17618
|
15 |
begin
|
|
16 |
|
17722
|
17 |
text {* \medskip Late package setup: default values for refute, see
|
21588
|
18 |
also theory @{theory Refute}. *}
|
17722
|
19 |
|
|
20 |
refute_params
|
|
21 |
["itself"=1,
|
|
22 |
minsize=1,
|
|
23 |
maxsize=8,
|
|
24 |
maxvars=10000,
|
|
25 |
maxtime=60,
|
|
26 |
satsolver="auto"]
|
|
27 |
|
17627
|
28 |
ML {* structure sat = SATFunc(structure cnf = cnf); *}
|
17618
|
29 |
|
30549
|
30 |
method_setup sat = {* Scan.succeed (K (SIMPLE_METHOD' sat.sat_tac)) *}
|
17618
|
31 |
"SAT solver"
|
|
32 |
|
30549
|
33 |
method_setup satx = {* Scan.succeed (K (SIMPLE_METHOD' sat.satx_tac)) *}
|
17618
|
34 |
"SAT solver (with definitional CNF)"
|
|
35 |
|
|
36 |
end
|