| author | wenzelm | 
| Fri, 25 Apr 2014 21:31:39 +0200 | |
| changeset 56730 | e723f041b6d0 | 
| parent 55239 | 97921d23ebe3 | 
| child 58889 | 5b7a9633cfa8 | 
| permissions | -rw-r--r-- | 
| 17618 | 1  | 
(* Title: HOL/SAT.thy  | 
2  | 
Author: Alwen Tiu, Tjark Weber  | 
|
3  | 
Copyright 2005  | 
|
4  | 
||
| 
49989
 
34d0ac1bdac6
moved "SAT" before "FunDef" and moved back all SAT-related ML files to where they belong
 
blanchet 
parents: 
49985 
diff
changeset
 | 
5  | 
Basic setup for the 'sat' and 'satx' tactics.  | 
| 17618 | 6  | 
*)  | 
7  | 
||
8  | 
header {* Reconstructing external resolution proofs for propositional logic *}
 | 
|
9  | 
||
| 26521 | 10  | 
theory SAT  | 
| 
49989
 
34d0ac1bdac6
moved "SAT" before "FunDef" and moved back all SAT-related ML files to where they belong
 
blanchet 
parents: 
49985 
diff
changeset
 | 
11  | 
imports HOL  | 
| 17618 | 12  | 
begin  | 
13  | 
||
| 
49989
 
34d0ac1bdac6
moved "SAT" before "FunDef" and moved back all SAT-related ML files to where they belong
 
blanchet 
parents: 
49985 
diff
changeset
 | 
14  | 
ML_file "Tools/prop_logic.ML"  | 
| 
 
34d0ac1bdac6
moved "SAT" before "FunDef" and moved back all SAT-related ML files to where they belong
 
blanchet 
parents: 
49985 
diff
changeset
 | 
15  | 
ML_file "Tools/sat_solver.ML"  | 
| 55239 | 16  | 
ML_file "Tools/sat.ML"  | 
| 48891 | 17  | 
|
| 55239 | 18  | 
method_setup sat = {* Scan.succeed (SIMPLE_METHOD' o SAT.sat_tac) *}
 | 
| 17618 | 19  | 
"SAT solver"  | 
20  | 
||
| 55239 | 21  | 
method_setup satx = {* Scan.succeed (SIMPLE_METHOD' o SAT.satx_tac) *}
 | 
| 17618 | 22  | 
"SAT solver (with definitional CNF)"  | 
23  | 
||
24  | 
end  |