| author | wenzelm | 
| Thu, 25 Apr 2013 13:22:45 +0200 | |
| changeset 51776 | 8ea64fb16bae | 
| parent 49989 | 34d0ac1bdac6 | 
| child 55239 | 97921d23ebe3 | 
| 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: 
49985diff
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: 
49985diff
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: 
49985diff
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: 
49985diff
changeset | 15 | ML_file "Tools/sat_solver.ML" | 
| 48891 | 16 | ML_file "Tools/sat_funcs.ML" | 
| 17 | ||
| 32232 | 18 | ML {* structure sat = SATFunc(cnf) *}
 | 
| 17618 | 19 | |
| 32232 | 20 | method_setup sat = {* Scan.succeed (SIMPLE_METHOD' o sat.sat_tac) *}
 | 
| 17618 | 21 | "SAT solver" | 
| 22 | ||
| 32232 | 23 | method_setup satx = {* Scan.succeed (SIMPLE_METHOD' o sat.satx_tac) *}
 | 
| 17618 | 24 | "SAT solver (with definitional CNF)" | 
| 25 | ||
| 26 | end |