use argo as additional SAT solver with models but no proofs, since the proof trace formats are not easily translatable
authorboehmes
Thu, 29 Sep 2016 20:54:45 +0200
changeset 63962 83a625d06e91
parent 63961 2fd9656c4c82
child 63963 ed98a055b9a5
use argo as additional SAT solver with models but no proofs, since the proof trace formats are not easily translatable
src/HOL/Presburger.thy
src/HOL/SAT.thy
src/HOL/Tools/Argo/argo_sat_solver.ML
--- a/src/HOL/Presburger.thy	Thu Sep 29 20:54:45 2016 +0200
+++ b/src/HOL/Presburger.thy	Thu Sep 29 20:54:45 2016 +0200
@@ -5,7 +5,7 @@
 section \<open>Decision Procedure for Presburger Arithmetic\<close>
 
 theory Presburger
-imports Groebner_Basis Set_Interval Argo
+imports Groebner_Basis Set_Interval
 keywords "try0" :: diag
 begin
 
--- a/src/HOL/SAT.thy	Thu Sep 29 20:54:45 2016 +0200
+++ b/src/HOL/SAT.thy	Thu Sep 29 20:54:45 2016 +0200
@@ -8,12 +8,13 @@
 section \<open>Reconstructing external resolution proofs for propositional logic\<close>
 
 theory SAT
-imports HOL
+imports Argo
 begin
 
 ML_file "Tools/prop_logic.ML"
 ML_file "Tools/sat_solver.ML"
 ML_file "Tools/sat.ML"
+ML_file "Tools/Argo/argo_sat_solver.ML"
 
 method_setup sat = \<open>Scan.succeed (SIMPLE_METHOD' o SAT.sat_tac)\<close>
   "SAT solver"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Argo/argo_sat_solver.ML	Thu Sep 29 20:54:45 2016 +0200
@@ -0,0 +1,27 @@
+(*  Title:      HOL/Tools/argo_sat_solver.ML
+    Author:     Sascha Boehme
+
+A SAT solver based on the Argo solver.
+This SAT solver produces models. For proofs use Argo_Tactic.prove instead.
+*)
+
+structure Argo_SAT_Solver: sig end =
+struct
+
+fun con_of i = (string_of_int i, Argo_Expr.Bool)
+
+fun expr_of Prop_Logic.True = Argo_Expr.true_expr
+  | expr_of Prop_Logic.False = Argo_Expr.false_expr
+  | expr_of (Prop_Logic.Not p) = Argo_Expr.mk_not (expr_of p)
+  | expr_of (Prop_Logic.Or (p1, p2)) = Argo_Expr.mk_or2 (expr_of p1) (expr_of p2)
+  | expr_of (Prop_Logic.And (p1, p2)) = Argo_Expr.mk_and2 (expr_of p1) (expr_of p2)
+  | expr_of (Prop_Logic.BoolVar i) = Argo_Expr.mk_con (con_of i)
+
+fun argo_solver p =
+  let val argo = Argo_Solver.assert [expr_of p] Argo_Solver.context
+  in SAT_Solver.SATISFIABLE (Argo_Solver.model_of argo o con_of) end
+  handle Argo_Proof.UNSAT _ => SAT_Solver.UNSATISFIABLE NONE
+
+val () = SAT_Solver.add_solver ("argo", argo_solver)
+
+end