| author | haftmann | 
| Wed, 17 Feb 2016 21:51:56 +0100 | |
| changeset 62345 | e66d7841d5a2 | 
| parent 60758 | d8d85a8172b5 | 
| child 63962 | 83a625d06e91 | 
| 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  | 
||
| 60758 | 8  | 
section \<open>Reconstructing external resolution proofs for propositional logic\<close>  | 
| 17618 | 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  | 
|
| 60758 | 18  | 
method_setup sat = \<open>Scan.succeed (SIMPLE_METHOD' o SAT.sat_tac)\<close>  | 
| 17618 | 19  | 
"SAT solver"  | 
20  | 
||
| 60758 | 21  | 
method_setup satx = \<open>Scan.succeed (SIMPLE_METHOD' o SAT.satx_tac)\<close>  | 
| 17618 | 22  | 
"SAT solver (with definitional CNF)"  | 
23  | 
||
24  | 
end  |