author | huffman |
Fri, 16 Mar 2012 15:51:53 +0100 | |
changeset 46962 | 5bdcdb28be83 |
parent 46096 | a00685a18e55 |
child 48891 | c0eafbd55de3 |
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 uses "Tools/sat_funcs.ML" begin 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