| 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 | 
 | 
| 32232 |     16 | ML {* structure sat = SATFunc(cnf) *}
 | 
| 17618 |     17 | 
 | 
| 32232 |     18 | method_setup sat = {* Scan.succeed (SIMPLE_METHOD' o sat.sat_tac) *}
 | 
| 17618 |     19 |   "SAT solver"
 | 
|  |     20 | 
 | 
| 32232 |     21 | method_setup satx = {* Scan.succeed (SIMPLE_METHOD' o sat.satx_tac) *}
 | 
| 17618 |     22 |   "SAT solver (with definitional CNF)"
 | 
|  |     23 | 
 | 
|  |     24 | end
 |