| author | desharna | 
| Mon, 10 Jun 2024 13:44:46 +0200 | |
| changeset 80321 | 31b9dfbe534c | 
| parent 69605 | a96320074298 | 
| child 83153 | d5715946bc3f | 
| 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 | ||
| 60758 | 8 | section \<open>Reconstructing external resolution proofs for propositional logic\<close> | 
| 17618 | 9 | |
| 26521 | 10 | theory SAT | 
| 63962 
83a625d06e91
use argo as additional SAT solver with models but no proofs, since the proof trace formats are not easily translatable
 boehmes parents: 
60758diff
changeset | 11 | imports Argo | 
| 17618 | 12 | begin | 
| 13 | ||
| 69605 | 14 | ML_file \<open>Tools/prop_logic.ML\<close> | 
| 15 | ML_file \<open>Tools/sat_solver.ML\<close> | |
| 16 | ML_file \<open>Tools/sat.ML\<close> | |
| 17 | ML_file \<open>Tools/Argo/argo_sat_solver.ML\<close> | |
| 48891 | 18 | |
| 60758 | 19 | method_setup sat = \<open>Scan.succeed (SIMPLE_METHOD' o SAT.sat_tac)\<close> | 
| 17618 | 20 | "SAT solver" | 
| 21 | ||
| 60758 | 22 | method_setup satx = \<open>Scan.succeed (SIMPLE_METHOD' o SAT.satx_tac)\<close> | 
| 17618 | 23 | "SAT solver (with definitional CNF)" | 
| 24 | ||
| 25 | end |