| author | Andreas Lochbihler | 
| Wed, 10 Oct 2012 16:41:19 +0200 | |
| changeset 49811 | 3fc6b3289c31 | 
| parent 48891 | c0eafbd55de3 | 
| child 49985 | 5b4b0e4e5205 | 
| permissions | -rw-r--r-- | 
(* Title: HOL/SAT.thy Author: Alwen Tiu, Tjark Weber Copyright 2005 Basic setup for the 'sat' and 'satx' tactic. *) header {* Reconstructing external resolution proofs for propositional logic *} theory SAT imports Refute begin ML_file "Tools/sat_funcs.ML" ML {* structure sat = SATFunc(cnf) *} method_setup sat = {* Scan.succeed (SIMPLE_METHOD' o sat.sat_tac) *} "SAT solver" method_setup satx = {* Scan.succeed (SIMPLE_METHOD' o sat.satx_tac) *} "SAT solver (with definitional CNF)" end