refactor: model building in smtlib, for smtlib

- sidekick.model removed, now just smtlib.Model (specific to it)
- use function entries for models, not just term->term
- re-building models in smtlib driver
- asolver.solve, in Check_res.t, does not return a concrete model, but a
  bundle of functions to query the solver
- store constants in smtlib typechecker AST (so we can directly map them
  to values in model construction)
This commit is contained in:
Simon Cruanes 2022-10-15 22:42:10 -04:00
parent 5ca966a827
commit 08541613af
No known key found for this signature in database
GPG key ID: EBFFF6F283F3A2B4
30 changed files with 350 additions and 206 deletions

View file

@ -1,16 +1,29 @@
(** Result of solving for the current set of clauses *) (** Result of solving for the current set of clauses *)
module Model = Sidekick_model type value = Term.t
type sat_result = {
get_value: Term.t -> value option; (** Value for this term *)
iter_classes: (Term.t Iter.t * value) Iter.t;
(** All equivalence classes in the congruence closure *)
eval_lit: Lit.t -> bool option; (** Evaluate literal *)
iter_true_lits: Lit.t Iter.t;
(** Iterate on literals that are true in the trail *)
}
(** Satisfiable *)
type unsat_result = {
unsat_core: unit -> Lit.t Iter.t;
(** Unsat core (subset of assumptions), or empty *)
unsat_proof: unit -> Sidekick_proof.Step.id option;
(** Proof step for the empty clause *)
}
(** Unsatisfiable *)
(** Result of calling "check" *) (** Result of calling "check" *)
type t = type t =
| Sat of Model.t (** Satisfiable *) | Sat of sat_result
| Unsat of { | Unsat of unsat_result
unsat_core: unit -> Lit.t Iter.t;
(** Unsat core (subset of assumptions), or empty *)
unsat_proof: unit -> Sidekick_proof.Step.id option;
(** Proof step for the empty clause *)
} (** Unsatisfiable *)
| Unknown of Unknown.t | Unknown of Unknown.t
(** Unknown, obtained after a timeout, memory limit, etc. *) (** Unknown, obtained after a timeout, memory limit, etc. *)

View file

@ -2,4 +2,4 @@
(name sidekick_abstract_solver) (name sidekick_abstract_solver)
(public_name sidekick.abstract-solver) (public_name sidekick.abstract-solver)
(flags :standard -open Sidekick_util -open Sidekick_core) (flags :standard -open Sidekick_util -open Sidekick_core)
(libraries containers iter sidekick.core sidekick.model sidekick.proof)) (libraries containers iter sidekick.core sidekick.proof))

View file

@ -5,8 +5,9 @@ type t = statement =
| Stmt_set_option of string list | Stmt_set_option of string list
| Stmt_set_info of string * string | Stmt_set_info of string * string
| Stmt_data of data list | Stmt_data of data list
| Stmt_ty_decl of ID.t * int (* new atomic cstor *) | Stmt_ty_decl of { name: ID.t; arity: int; ty_const: ty }
| Stmt_decl of ID.t * ty list * ty (** new atomic cstor *)
| Stmt_decl of { name: ID.t; ty_args: ty list; ty_ret: ty; const: term }
| Stmt_define of definition list | Stmt_define of definition list
| Stmt_assert of term | Stmt_assert of term
| Stmt_assert_clause of term list | Stmt_assert_clause of term list
@ -30,10 +31,11 @@ let pp out = function
Fmt.fprintf out "(@[not %a@])" Term.pp_debug t Fmt.fprintf out "(@[not %a@])" Term.pp_debug t
in in
Fmt.fprintf out "(@[check-sat-assuming@ (@[%a@])@])" (Fmt.list pp_pair) l Fmt.fprintf out "(@[check-sat-assuming@ (@[%a@])@])" (Fmt.list pp_pair) l
| Stmt_ty_decl (s, n) -> Fmt.fprintf out "(@[declare-sort@ %a %d@])" ID.pp s n | Stmt_ty_decl { name; arity; ty_const = _ } ->
| Stmt_decl (id, args, ret) -> Fmt.fprintf out "(@[declare-sort@ %a %d@])" ID.pp name arity
Fmt.fprintf out "(@[<1>declare-fun@ %a (@[%a@])@ %a@])" ID.pp id | Stmt_decl { name; ty_args; ty_ret; const = _ } ->
(Util.pp_list Ty.pp) args Ty.pp ret Fmt.fprintf out "(@[<1>declare-fun@ %a (@[%a@])@ %a@])" ID.pp name
(Util.pp_list Ty.pp) ty_args Ty.pp ty_ret
| Stmt_assert t -> Fmt.fprintf out "(@[assert@ %a@])" Term.pp_debug t | Stmt_assert t -> Fmt.fprintf out "(@[assert@ %a@])" Term.pp_debug t
| Stmt_assert_clause c -> | Stmt_assert_clause c ->
Fmt.fprintf out "(@[assert-clause@ %a@])" (Util.pp_list Term.pp_debug) c Fmt.fprintf out "(@[assert-clause@ %a@])" (Util.pp_list Term.pp_debug) c

View file

@ -11,8 +11,9 @@ type t = statement =
| Stmt_set_option of string list | Stmt_set_option of string list
| Stmt_set_info of string * string | Stmt_set_info of string * string
| Stmt_data of data list | Stmt_data of data list
| Stmt_ty_decl of ID.t * int (* new atomic cstor *) | Stmt_ty_decl of { name: ID.t; arity: int; ty_const: ty }
| Stmt_decl of ID.t * ty list * ty (** new atomic cstor *)
| Stmt_decl of { name: ID.t; ty_args: ty list; ty_ret: ty; const: term }
| Stmt_define of definition list | Stmt_define of definition list
| Stmt_assert of term | Stmt_assert of term
| Stmt_assert_clause of term list | Stmt_assert_clause of term list

View file

@ -66,8 +66,9 @@ type statement =
| Stmt_set_option of string list | Stmt_set_option of string list
| Stmt_set_info of string * string | Stmt_set_info of string * string
| Stmt_data of data list | Stmt_data of data list
| Stmt_ty_decl of ID.t * int (* new atomic cstor *) | Stmt_ty_decl of { name: ID.t; arity: int; ty_const: ty }
| Stmt_decl of ID.t * ty list * ty (** new atomic cstor *)
| Stmt_decl of { name: ID.t; ty_args: ty list; ty_ret: ty; const: term }
| Stmt_define of definition list | Stmt_define of definition list
| Stmt_assert of term | Stmt_assert of term
| Stmt_assert_clause of term list | Stmt_assert_clause of term list

View file

@ -25,7 +25,6 @@ let p_model = ref false
let check = ref false let check = ref false
let time_limit = ref 300. let time_limit = ref 300.
let mem_limit = ref 1_000_000_000. let mem_limit = ref 1_000_000_000.
let restarts = ref true
let p_stat = ref false let p_stat = ref false
let p_gc_stat = ref false let p_gc_stat = ref false
let p_progress = ref false let p_progress = ref false
@ -74,8 +73,6 @@ let argspec =
Arg.Set check, Arg.Set check,
" build, check and print the proof (if output is set), if unsat" ); " build, check and print the proof (if output is set), if unsat" );
"--no-check", Arg.Clear check, " inverse of -check"; "--no-check", Arg.Clear check, " inverse of -check";
"--restarts", Arg.Set restarts, " enable restarts";
"--no-restarts", Arg.Clear restarts, " disable restarts";
"--stat", Arg.Set p_stat, " print statistics"; "--stat", Arg.Set p_stat, " print statistics";
"--proof", Arg.Set p_proof, " print proof"; "--proof", Arg.Set p_proof, " print proof";
"--no-proof", Arg.Clear p_proof, " do not print proof"; "--no-proof", Arg.Clear p_proof, " do not print proof";
@ -220,9 +217,8 @@ let main_smt ~config () : _ result =
parse_res >>= fun input -> parse_res >>= fun input ->
let driver = let driver =
let asolver = Solver.Smt_solver.Solver.as_asolver solver in let asolver = Solver.Smt_solver.Solver.as_asolver solver in
Driver.create ~restarts:!restarts ~pp_cnf:!p_cnf ~time:!time_limit Driver.create ~pp_cnf:!p_cnf ~time:!time_limit ~memory:!mem_limit
~memory:!mem_limit ~pp_model:!p_model ?proof_file ~check:!check ~pp_model:!p_model ?proof_file ~check:!check ~progress:!p_progress asolver
~progress:!p_progress asolver
in in
(* process statements *) (* process statements *)

View file

@ -1,6 +0,0 @@
(library
(name sidekick_model)
(public_name sidekick.model)
(synopsis "Finite models for Sidekick")
(libraries sidekick.sigs sidekick.util sidekick.core)
(flags :standard -w +32 -open Sidekick_util))

View file

@ -1,24 +0,0 @@
open Sidekick_core
module T = Term
type t = { m: Term.t T.Map.t } [@@unboxed]
let empty : t = { m = T.Map.empty }
let is_empty self = T.Map.is_empty self.m
let mem self t = T.Map.mem t self.m
let find self t = T.Map.find_opt t self.m
let eval = find
let add t v self : t = { m = T.Map.add t v self.m }
let keys self = T.Map.keys self.m
let to_iter self = T.Map.to_iter self.m
let pp out (self : t) =
if is_empty self then
Fmt.string out "(model)"
else (
let pp_pair out (t, v) =
Fmt.fprintf out "(@[<1>%a@ := %a@])" Term.pp t Term.pp v
in
Fmt.fprintf out "(@[<hv>model@ %a@])" (Util.pp_iter pp_pair)
(T.Map.to_iter self.m)
)

View file

@ -1,19 +0,0 @@
(** Models
A model can be produced when the solver is found to be in a
satisfiable state after a call to {!solve}. *)
open Sidekick_core
type t
val empty : t
val is_empty : t -> bool
val mem : t -> Term.t -> bool
val find : t -> Term.t -> Term.t option
val eval : t -> Term.t -> Term.t option
val add : Term.t -> Term.t -> t -> t
val keys : t -> Term.t Iter.t
val to_iter : t -> (Term.t * Term.t) Iter.t
include Sidekick_sigs.PRINT with type t := t

View file

@ -9,6 +9,8 @@ Copyright 2016 Simon Cruanes
open Sidekick_core open Sidekick_core
module Proof = Sidekick_proof module Proof = Sidekick_proof
exception UndecidedLit
(** Solver in a "SATISFIABLE" state *) (** Solver in a "SATISFIABLE" state *)
module type SAT_STATE = sig module type SAT_STATE = sig
val eval : Lit.t -> bool val eval : Lit.t -> bool

View file

@ -36,7 +36,7 @@ type unsat_cause =
exception E_sat exception E_sat
exception E_unsat of unsat_cause exception E_unsat of unsat_cause
exception UndecidedLit exception UndecidedLit = Sigs.UndecidedLit
exception Restart exception Restart
exception Conflict of clause exception Conflict of clause

View file

@ -7,7 +7,6 @@
*) *)
module Sigs = Sigs module Sigs = Sigs
module Model = Sidekick_model
module Model_builder = Model_builder module Model_builder = Model_builder
module Registry = Registry module Registry = Registry
module Solver_internal = Solver_internal module Solver_internal = Solver_internal

View file

@ -3,6 +3,6 @@
(public_name sidekick.smt-solver) (public_name sidekick.smt-solver)
(synopsis "main SMT solver") (synopsis "main SMT solver")
(libraries containers iter sidekick.core sidekick.util sidekick.cc (libraries containers iter sidekick.core sidekick.util sidekick.cc
sidekick.sat sidekick.abstract-solver sidekick.simplify sidekick.model sidekick.sat sidekick.abstract-solver sidekick.simplify sidekick.proof
sidekick.proof sidekick.trace) sidekick.trace)
(flags :standard -w +32 -open Sidekick_util)) (flags :standard -w +32 -open Sidekick_util))

View file

@ -1,24 +1,23 @@
open Sigs open Sigs
module T = Term module T = Term
module TM = Term.Map
type t = { type t = {
tst: Term.store; tst: Term.store;
mutable m: Model.t; mutable m: Term.t TM.t;
required: Term.t Queue.t; required: Term.t Queue.t;
gensym: Gensym.t; gensym: Gensym.t;
} }
let create tst : t = let create tst : t =
{ { tst; m = TM.empty; required = Queue.create (); gensym = Gensym.create tst }
tst;
m = Model.empty;
required = Queue.create ();
gensym = Gensym.create tst;
}
let pp out (self : t) : unit = let pp out (self : t) : unit =
Fmt.fprintf out "(@[model-builder@ :model %a@ :q (@[%a@])@])" Model.pp self.m let pp_kv out (k, v) =
(Util.pp_iter T.pp) Fmt.fprintf out "(@[%a@ := %a@])" Term.pp k Term.pp v
in
Fmt.fprintf out "(@[model-builder@ :model %a@ :q (@[%a@])@])"
(Util.pp_iter pp_kv) (TM.to_iter self.m) (Util.pp_iter T.pp)
(Iter.of_queue self.required) (Iter.of_queue self.required)
let gensym self ~pre ~ty : Term.t = Gensym.fresh_term self.gensym ~pre ty let gensym self ~pre ~ty : Term.t = Gensym.fresh_term self.gensym ~pre ty
@ -26,33 +25,33 @@ let gensym self ~pre ~ty : Term.t = Gensym.fresh_term self.gensym ~pre ty
let rec pop_required (self : t) : _ option = let rec pop_required (self : t) : _ option =
match Queue.take_opt self.required with match Queue.take_opt self.required with
| None -> None | None -> None
| Some t when Model.mem self.m t -> pop_required self | Some t when TM.mem t self.m -> pop_required self
| Some t -> Some t | Some t -> Some t
let require_eval (self : t) t : unit = let require_eval (self : t) t : unit =
if not @@ Model.mem self.m t then Queue.push t self.required if not @@ TM.mem t self.m then Queue.push t self.required
let[@inline] mem self t : bool = Model.mem self.m t let[@inline] mem self t : bool = TM.mem t self.m
let add (self : t) ?(subs = []) t v : unit = let add (self : t) ?(subs = []) t v : unit =
if not @@ mem self t then ( if not @@ mem self t then (
self.m <- Model.add t v self.m; self.m <- TM.add t v self.m;
List.iter (fun u -> require_eval self u) subs List.iter (fun u -> require_eval self u) subs
) )
type eval_cache = Term.Internal_.cache type eval_cache = Term.Internal_.cache
let eval ?(cache = Term.Internal_.create_cache 8) (self : t) (t : Term.t) = let eval ?(cache = Term.Internal_.create_cache 8) (self : t) (t : Term.t) =
let t = Model.find self.m t |> Option.value ~default:t in let t = TM.get t self.m |> Option.value ~default:t in
T.Internal_.replace_ ~cache self.tst ~recursive:true t ~f:(fun ~recurse:_ u -> T.Internal_.replace_ ~cache self.tst ~recursive:true t ~f:(fun ~recurse:_ u ->
Model.find self.m u) TM.get u self.m)
let to_model (self : t) : Model.t = let to_map (self : t) : _ TM.t =
(* ensure we evaluate each term only once *) (* ensure we evaluate each term only once *)
let cache = T.Internal_.create_cache 8 in let cache = T.Internal_.create_cache 8 in
let m = let m =
Model.keys self.m TM.keys self.m
|> Iter.map (fun t -> t, eval ~cache self t) |> Iter.map (fun t -> t, eval ~cache self t)
|> Iter.fold (fun m (t, v) -> Model.add t v m) Model.empty |> Iter.fold (fun m (t, v) -> TM.add t v m) TM.empty
in in
m m

View file

@ -33,4 +33,4 @@ val eval : ?cache:eval_cache -> t -> Term.t -> value
val pop_required : t -> Term.t option val pop_required : t -> Term.t option
(** gives the next subterm that is required but has no value yet *) (** gives the next subterm that is required but has no value yet *)
val to_model : t -> Model.t val to_map : t -> Term.t Term.Map.t

View file

@ -13,7 +13,6 @@
*) *)
include Sidekick_core include Sidekick_core
module Model = Sidekick_model
module Simplify = Sidekick_simplify module Simplify = Sidekick_simplify
module CC = Sidekick_cc.CC module CC = Sidekick_cc.CC
module E_node = Sidekick_cc.E_node module E_node = Sidekick_cc.E_node

View file

@ -11,14 +11,30 @@ module Sat_solver = Sidekick_sat
module Unknown = Sidekick_abstract_solver.Unknown [@@ocaml.warning "-37"] module Unknown = Sidekick_abstract_solver.Unknown [@@ocaml.warning "-37"]
type value = Term.t
type sat_result = Check_res.sat_result = {
get_value: Term.t -> value option; (** Value for this term *)
iter_classes: (Term.t Iter.t * value) Iter.t;
(** All equivalence classes in the congruence closure *)
eval_lit: Lit.t -> bool option; (** Evaluate literal *)
iter_true_lits: Lit.t Iter.t;
(** Iterate on literals that are true in the trail *)
}
(** Satisfiable *)
type unsat_result = Check_res.unsat_result = {
unsat_core: unit -> Lit.t Iter.t;
(** Unsat core (subset of assumptions), or empty *)
unsat_proof: unit -> Sidekick_proof.Step.id option;
(** Proof step for the empty clause *)
}
(** Unsatisfiable *)
(** Result of solving for the current set of clauses *)
type res = Check_res.t = type res = Check_res.t =
| Sat of Model.t (** Satisfiable *) | Sat of sat_result
| Unsat of { | Unsat of unsat_result
unsat_core: unit -> lit Iter.t;
(** Unsat core (subset of assumptions), or empty *)
unsat_proof: unit -> step_id option;
(** Proof step for the empty clause *)
} (** Unsatisfiable *)
| Unknown of Unknown.t | Unknown of Unknown.t
(** Unknown, obtained after a timeout, memory limit, etc. *) (** Unknown, obtained after a timeout, memory limit, etc. *)
@ -188,26 +204,57 @@ let solve ?(on_exit = []) ?(on_progress = fun _ -> ())
"(@[sidekick.smt-solver: SAT@ actual: UNKNOWN@ :reason \ "(@[sidekick.smt-solver: SAT@ actual: UNKNOWN@ :reason \
incomplete-fragment@])"); incomplete-fragment@])");
Unknown Unknown.U_incomplete Unknown Unknown.U_incomplete
| Sat_solver.Sat _ -> | Sat_solver.Sat (module SAT) ->
Log.debug 1 "(sidekick.smt-solver: SAT)"; Log.debug 1 "(sidekick.smt-solver: SAT)";
Log.debugf 5 (fun k ->
let ppc out n =
Fmt.fprintf out "{@[<hv1>class@ %a@]}"
(Util.pp_iter ~sep:";" E_node.pp)
(E_node.iter_class n)
in
k "(@[sidekick.smt-solver.classes@ (@[%a@])@])" (Util.pp_iter ppc)
(CC.all_classes @@ Solver_internal.cc self.si));
let m = let m =
match SI.last_model self.si with match SI.last_model self.si with
| Some m -> m | Some m -> m
| None -> assert false | None -> assert false
in in
let iter_classes =
CC.all_classes (Solver_internal.cc self.si)
|> Iter.filter (fun repr ->
not @@ Term.is_pi (Term.ty @@ E_node.term repr))
|> Iter.map (fun repr ->
let v =
match
(* find value for this class *)
Iter.find_map
(fun en -> Term.Map.get (E_node.term en) m)
(E_node.iter_class repr)
with
| None ->
Error.errorf
"(@[solver.mk-model.no-value-for-repr@ %a@ :ty %a@])"
E_node.pp repr Term.pp
(Term.ty @@ E_node.term repr)
| Some v -> v
in
let terms = E_node.iter_class repr |> Iter.map E_node.term in
terms, v)
in
Log.debugf 5 (fun k ->
let ppc out (cls, v) =
Fmt.fprintf out "{@[<hv1>class@ :value %a@ %a@]}" Term.pp v
(Util.pp_iter ~sep:";" Term.pp)
cls
in
k "(@[sidekick.smt-solver.classes@ (@[%a@])@])" (Util.pp_iter ppc)
iter_classes);
do_on_exit (); do_on_exit ();
Sat m Sat
{
get_value = (fun t -> Term.Map.get t m);
iter_classes;
eval_lit =
(fun l ->
try Some (SAT.eval l) with Sat_solver.UndecidedLit -> None);
iter_true_lits = SAT.iter_trail;
}
| Sat_solver.Unsat (module UNSAT) -> | Sat_solver.Unsat (module UNSAT) ->
let unsat_core () = UNSAT.unsat_assumptions () in let unsat_core () = UNSAT.unsat_assumptions () in
let unsat_proof () = Some (UNSAT.unsat_proof ()) in let unsat_proof () = Some (UNSAT.unsat_proof ()) in

View file

@ -102,15 +102,30 @@ val assert_term : t -> term -> unit
val add_ty : t -> ty -> unit val add_ty : t -> ty -> unit
type value = Term.t
type sat_result = Check_res.sat_result = {
get_value: Term.t -> value option; (** Value for this term *)
iter_classes: (Term.t Iter.t * value) Iter.t;
(** All equivalence classes in the congruence closure *)
eval_lit: Lit.t -> bool option; (** Evaluate literal *)
iter_true_lits: Lit.t Iter.t;
(** Iterate on literals that are true in the trail *)
}
(** Satisfiable *)
type unsat_result = Check_res.unsat_result = {
unsat_core: unit -> Lit.t Iter.t;
(** Unsat core (subset of assumptions), or empty *)
unsat_proof: unit -> Sidekick_proof.Step.id option;
(** Proof step for the empty clause *)
}
(** Unsatisfiable *)
(** Result of solving for the current set of clauses *) (** Result of solving for the current set of clauses *)
type res = Check_res.t = type res = Check_res.t =
| Sat of Model.t (** Satisfiable *) | Sat of sat_result
| Unsat of { | Unsat of unsat_result
unsat_core: unit -> lit Iter.t;
(** Unsat core (subset of assumptions), or empty *)
unsat_proof: unit -> step_id option;
(** Proof step for the empty clause *)
} (** Unsatisfiable *)
| Unknown of Unknown.t | Unknown of Unknown.t
(** Unknown, obtained after a timeout, memory limit, etc. *) (** Unknown, obtained after a timeout, memory limit, etc. *)

View file

@ -53,7 +53,7 @@ type t = {
mutable model_complete: model_completion_hook list; mutable model_complete: model_completion_hook list;
simp: Simplify.t; simp: Simplify.t;
delayed_actions: delayed_action Queue.t; delayed_actions: delayed_action Queue.t;
mutable last_model: Model.t option; mutable last_model: Term.t Term.Map.t option;
mutable th_states: th_states; (** Set of theories *) mutable th_states: th_states; (** Set of theories *)
mutable level: int; mutable level: int;
mutable complete: bool; mutable complete: bool;
@ -327,7 +327,7 @@ let rec pop_lvls_theories_ n = function
(** {2 Model construction and theory combination} *) (** {2 Model construction and theory combination} *)
(* make model from the congruence closure *) (* make model from the congruence closure *)
let mk_model_ (self : t) (lits : lit Iter.t) : Model.t = let mk_model_ (self : t) (lits : lit Iter.t) : Term.t Term.Map.t =
let@ () = Profile.with_ "smt-solver.mk-model" in let@ () = Profile.with_ "smt-solver.mk-model" in
Log.debug 1 "(smt.solver.mk-model)"; Log.debug 1 "(smt.solver.mk-model)";
let module MB = Model_builder in let module MB = Model_builder in
@ -395,7 +395,7 @@ let mk_model_ (self : t) (lits : lit Iter.t) : Model.t =
in in
compute_fixpoint (); compute_fixpoint ();
MB.to_model model MB.to_map model
(* call congruence closure, perform the actions it scheduled *) (* call congruence closure, perform the actions it scheduled *)
let check_cc_with_acts_ (self : t) (acts : theory_actions) = let check_cc_with_acts_ (self : t) (acts : theory_actions) =

View file

@ -270,7 +270,7 @@ val on_progress : t -> (unit, unit) Event.t
val is_complete : t -> bool val is_complete : t -> bool
(** Are we still in a complete logic fragment? *) (** Are we still in a complete logic fragment? *)
val last_model : t -> Model.t option val last_model : t -> Term.t Term.Map.t option
(** {2 Delayed actions} *) (** {2 Delayed actions} *)

View file

@ -1,4 +1,3 @@
open Sidekick_core
module Profile = Sidekick_util.Profile module Profile = Sidekick_util.Profile
module Proof = Sidekick_proof module Proof = Sidekick_proof
module Asolver = Solver.Asolver module Asolver = Solver.Asolver
@ -12,52 +11,10 @@ type 'a or_error = ('a, string) CCResult.t
module E = CCResult module E = CCResult
module Fmt = CCFormat module Fmt = CCFormat
(* TODO: use external proof checker instead: check-sat(φ + model)
(* check SMT model *)
let check_smt_model (solver:Solver.Sat_solver.t) (hyps:_ Vec.t) (m:Model.t) : unit =
Log.debug 1 "(smt.check-smt-model)";
let module S = Solver.Sat_solver in
let check_atom (lit:Lit.t) : Msat.lbool =
Log.debugf 5 (fun k->k "(@[smt.check-smt-model.atom@ %a@])" Lit.pp lit);
let a = S.make_atom solver lit in
let sat_value = S.eval_atom solver a in
let t, sign = Lit.as_atom lit in
begin match Model.eval m t with
| Some (V_bool b) ->
let b = if sign then b else not b in
if (sat_value <> Msat.L_undefined) &&
((b && sat_value=Msat.L_false) || (not b && sat_value=Msat.L_true)) then (
Error.errorf "(@[check-model.error@ :atom %a@ :model-val %B@ :sat-val %a@])"
S.Atom.pp a b Msat.pp_lbool sat_value
) else (
Log.debugf 5
(fun k->k "(@[check-model@ :atom %a@ :model-val %B@ :sat-val %a@])"
S.Atom.pp a b Msat.pp_lbool sat_value);
)
| Some v ->
Error.errorf "(@[check-model.error@ :atom %a@ :non-bool-value %a@])"
S.Atom.pp a Value.pp v
| None ->
if sat_value <> Msat.L_undefined then (
Error.errorf "(@[check-model.error@ :atom %a@ :no-smt-value@ :sat-val %a@])"
S.Atom.pp a Msat.pp_lbool sat_value
);
end;
sat_value
in
let check_c c =
let bs = List.map check_atom c in
if List.for_all (function Msat.L_true -> false | _ -> true) bs then (
Error.errorf "(@[check-model.error.none-true@ :clause %a@ :vals %a@])"
(Fmt.Dump.list Lit.pp) c Fmt.(Dump.list @@ Msat.pp_lbool) bs
);
in
Vec.iter check_c hyps
*)
type t = { type t = {
progress: Progress_bar.t option; progress: Progress_bar.t option;
solver: Asolver.t; solver: Asolver.t;
build_model: Build_model.t;
time_start: float; time_start: float;
time_limit: float; time_limit: float;
memory_limit: float; memory_limit: float;
@ -65,13 +22,11 @@ type t = {
pp_model: bool; pp_model: bool;
pp_cnf: bool; pp_cnf: bool;
check: bool; check: bool;
restarts: bool;
} }
(* call the solver to check-sat *) (* call the solver to check-sat *)
let create ?(restarts = true) ?(pp_cnf = false) ?proof_file ?(pp_model = false) let create ?(pp_cnf = false) ?proof_file ?(pp_model = false) ?(check = false)
?(check = false) ?time ?memory ?(progress = false) (solver : Asolver.t) : t ?time ?memory ?(progress = false) (solver : Asolver.t) : t =
=
let time_start = now () in let time_start = now () in
let progress = let progress =
if progress then if progress then
@ -86,7 +41,7 @@ let create ?(restarts = true) ?(pp_cnf = false) ?proof_file ?(pp_model = false)
{ {
time_start; time_start;
restarts; build_model = Build_model.create ();
progress; progress;
solver; solver;
proof_file; proof_file;
@ -97,15 +52,20 @@ let create ?(restarts = true) ?(pp_cnf = false) ?proof_file ?(pp_model = false)
memory_limit; memory_limit;
} }
let decl_sort (_self : t) c n : unit = let decl_sort (self : t) c n ty_const : unit =
(* TODO: more? pass to abstract solver? *)
Log.debugf 1 (fun k -> k "(@[declare-sort %a@ :arity %d@])" ID.pp c n)
let decl_fun (_self : t) id args ret : unit =
(* TODO: more? record for model building *)
Log.debugf 1 (fun k -> Log.debugf 1 (fun k ->
k "(@[declare-fun %a@ :args (@[%a@])@ :ret %a@])" ID.pp id k "(@[declare-sort %a@ :arity %d@ :ty-const %a@])" ID.pp c n Term.pp
(Util.pp_list Ty.pp) args Ty.pp ret) ty_const);
Build_model.add_ty self.build_model ty_const
let decl_fun (self : t) f args ret const : unit =
Log.debugf 1 (fun k ->
k "(@[declare-fun %a@ :args (@[%a@])@ :ret %a@ :const %a@])" ID.pp f
(Util.pp_list Ty.pp) args Ty.pp ret Term.pp const);
Build_model.add_fun self.build_model const
let build_model (self : t) (sat : Solver.sat_result) : Model.t =
Build_model.build self.build_model sat
(* call the solver to check satisfiability *) (* call the solver to check satisfiability *)
let solve (self : t) ~assumptions () : Solver.res = let solve (self : t) ~assumptions () : Solver.res =
@ -133,10 +93,12 @@ let solve (self : t) ~assumptions () : Solver.res =
let t2 = now () in let t2 = now () in
flush stdout; flush stdout;
(match res with (match res with
| Asolver.Check_res.Sat m -> | Asolver.Check_res.Sat sat ->
if self.pp_model then if self.pp_model then (
let m = build_model self sat in
(* TODO: use actual {!Model} in the solver? or build it afterwards *) (* TODO: use actual {!Model} in the solver? or build it afterwards *)
Format.printf "(@[<hv1>model@ %a@])@." Sidekick_smt_solver.Model.pp m; Format.printf "(@[<hv1>model@ %a@])@." Model.pp m
);
(* TODO (* TODO
if check then ( if check then (
Solver.check_model s; Solver.check_model s;
@ -194,8 +156,6 @@ let process_stmt (self : t) (stmt : Statement.t) : unit or_error =
Log.debugf 5 (fun k -> Log.debugf 5 (fun k ->
k "(@[smtlib.process-statement@ %a@])" Statement.pp stmt); k "(@[smtlib.process-statement@ %a@])" Statement.pp stmt);
let add_step r = Proof.Tracer.add_step (Asolver.proof self.solver) r in
match stmt with match stmt with
| Statement.Stmt_set_logic logic -> | Statement.Stmt_set_logic logic ->
if not @@ List.mem logic known_logics then if not @@ List.mem logic known_logics then
@ -216,17 +176,17 @@ let process_stmt (self : t) (stmt : Statement.t) : unit or_error =
in in
ignore (solve self ~assumptions () : Solver.res); ignore (solve self ~assumptions () : Solver.res);
E.return () E.return ()
| Statement.Stmt_ty_decl (id, n) -> | Statement.Stmt_ty_decl { name; arity; ty_const } ->
decl_sort self id n; decl_sort self name arity ty_const;
E.return () E.return ()
| Statement.Stmt_decl (f, ty_args, ty_ret) -> | Statement.Stmt_decl { name; ty_args; ty_ret; const } ->
decl_fun self f ty_args ty_ret; decl_fun self name ty_args ty_ret const;
E.return () E.return ()
| Statement.Stmt_assert t -> | Statement.Stmt_assert t ->
if self.pp_cnf then Format.printf "(@[<hv1>assert@ %a@])@." Term.pp t; if self.pp_cnf then Format.printf "(@[<hv1>assert@ %a@])@." Term.pp t;
let lit = Asolver.lit_of_term self.solver t in let lit = Asolver.lit_of_term self.solver t in
Asolver.assert_clause self.solver [| lit |] Asolver.assert_clause self.solver [| lit |] (fun () ->
(fun () -> Proof.Sat_rules.sat_input_clause [ lit ]); Proof.Sat_rules.sat_input_clause [ lit ]);
E.return () E.return ()
| Statement.Stmt_assert_clause c_ts -> | Statement.Stmt_assert_clause c_ts ->
if self.pp_cnf then if self.pp_cnf then
@ -235,8 +195,7 @@ let process_stmt (self : t) (stmt : Statement.t) : unit or_error =
let c = CCList.map (fun t -> Asolver.lit_of_term self.solver t) c_ts in let c = CCList.map (fun t -> Asolver.lit_of_term self.solver t) c_ts in
(* proof of assert-input + preprocessing *) (* proof of assert-input + preprocessing *)
let pr = let pr () =
fun () ->
let lits = List.map (Asolver.lit_of_term self.solver) c_ts in let lits = List.map (Asolver.lit_of_term self.solver) c_ts in
Proof.Sat_rules.sat_input_clause lits Proof.Sat_rules.sat_input_clause lits
in in
@ -245,16 +204,19 @@ let process_stmt (self : t) (stmt : Statement.t) : unit or_error =
E.return () E.return ()
| Statement.Stmt_get_model -> | Statement.Stmt_get_model ->
(match Asolver.last_res self.solver with (match Asolver.last_res self.solver with
| Some (Solver.Sat m) -> Fmt.printf "%a@." Sidekick_smt_solver.Model.pp m | Some (Solver.Sat sat) ->
let m = build_model self sat in
Fmt.printf "%a@." Model.pp m
| _ -> Error.errorf "cannot access model"); | _ -> Error.errorf "cannot access model");
E.return () E.return ()
| Statement.Stmt_get_value l -> | Statement.Stmt_get_value l ->
(match Asolver.last_res self.solver with (match Asolver.last_res self.solver with
| Some (Solver.Sat m) -> | Some (Solver.Sat sat) ->
let m = build_model self sat in
let l = let l =
List.map List.map
(fun t -> (fun t ->
match Sidekick_smt_solver.Model.eval m t with match Model.eval t m with
| None -> Error.errorf "cannot evaluate %a" Term.pp t | None -> Error.errorf "cannot evaluate %a" Term.pp t
| Some u -> t, u) | Some u -> t, u)
l l

View file

@ -21,7 +21,6 @@ type t
(** The SMTLIB driver *) (** The SMTLIB driver *)
val create : val create :
?restarts:bool ->
?pp_cnf:bool -> ?pp_cnf:bool ->
?proof_file:string -> ?proof_file:string ->
?pp_model:bool -> ?pp_model:bool ->

View file

@ -5,6 +5,7 @@ module Solver = Solver
module Term = Sidekick_base.Term module Term = Sidekick_base.Term
module Stmt = Sidekick_base.Statement module Stmt = Sidekick_base.Statement
module Check_cc = Check_cc module Check_cc = Check_cc
module Model = Model
type 'a or_error = ('a, string) CCResult.t type 'a or_error = ('a, string) CCResult.t

View file

@ -11,6 +11,7 @@ module Stmt = Sidekick_base.Statement
module Driver = Driver module Driver = Driver
module Solver = Solver module Solver = Solver
module Check_cc = Check_cc module Check_cc = Check_cc
module Model = Model
val parse : Term.store -> string -> Stmt.t list or_error val parse : Term.store -> string -> Stmt.t list or_error
val parse_stdin : Term.store -> Stmt.t list or_error val parse_stdin : Term.store -> Stmt.t list or_error

View file

@ -457,13 +457,13 @@ and conv_statement_aux (ctx : Ctx.t) (stmt : PA.statement) : Stmt.t list =
let id = ID.make s in let id = ID.make s in
let ty = Ty.uninterpreted tst id in let ty = Ty.uninterpreted tst id in
Ctx.add_id_ ctx s id (Ctx.K_ty (Ctx.K_atomic ty)); Ctx.add_id_ ctx s id (Ctx.K_ty (Ctx.K_atomic ty));
[ Stmt.Stmt_ty_decl (id, n) ] [ Stmt.Stmt_ty_decl { name = id; arity = n; ty_const = ty } ]
| PA.Stmt_decl fr -> | PA.Stmt_decl fr ->
let f, args, ret = conv_fun_decl ctx fr in let f, args, ret = conv_fun_decl ctx fr in
let id = ID.make f in let id = ID.make f in
let c_f = Uconst.uconst_of_id' tst id args ret in let c_f = Uconst.uconst_of_id' tst id args ret in
Ctx.add_id_ ctx f id (Ctx.K_fun c_f); Ctx.add_id_ ctx f id (Ctx.K_fun c_f);
[ Stmt.Stmt_decl (id, args, ret) ] [ Stmt.Stmt_decl { name = id; ty_args = args; ty_ret = ret; const = c_f } ]
| PA.Stmt_data l -> | PA.Stmt_data l ->
(* first, read and declare each datatype (it can occur in the other (* first, read and declare each datatype (it can occur in the other
datatypes' constructors) *) datatypes' constructors) *)
@ -572,7 +572,10 @@ and conv_statement_aux (ctx : Ctx.t) (stmt : PA.statement) : Stmt.t list =
let f = Uconst.uconst_of_id tst id ret in let f = Uconst.uconst_of_id tst id ret in
Ctx.add_id_ ctx fun_name id (Ctx.K_fun f); Ctx.add_id_ ctx fun_name id (Ctx.K_fun f);
let rhs = conv_term ctx fr_body in let rhs = conv_term ctx fr_body in
[ Stmt.Stmt_decl (id, [], ret); Stmt.Stmt_assert (Form.eq tst f rhs) ] [
Stmt.Stmt_decl { name = id; ty_args = []; ty_ret = ret; const = f };
Stmt.Stmt_assert (Form.eq tst f rhs);
]
| PA.Stmt_fun_rec _ | PA.Stmt_fun_def _ -> | PA.Stmt_fun_rec _ | PA.Stmt_fun_def _ ->
errorf_ctx ctx "unsupported definition: %a" PA.pp_stmt stmt errorf_ctx ctx "unsupported definition: %a" PA.pp_stmt stmt
| PA.Stmt_assert t -> | PA.Stmt_assert t ->

46
src/smtlib/build_model.ml Normal file
View file

@ -0,0 +1,46 @@
open Common_
open! Sidekick_base
type t = { fun_symbols: unit Term.Tbl.t; ty_symbols: unit Term.Tbl.t }
let create () : t =
{ fun_symbols = Term.Tbl.create 32; ty_symbols = Term.Tbl.create 32 }
let add_ty (self : t) ty_const : unit = Term.Tbl.add self.ty_symbols ty_const ()
let add_fun (self : t) const : unit = Term.Tbl.add self.fun_symbols const ()
let build (self : t) (sat : Solver.sat_result) : Model.t =
let m = ref Model.empty in
(* process [t], whose value is [value] in the model *)
let examine_term t ~value : unit =
let f, args = Term.unfold_app t in
(* TODO: types *)
if Term.Tbl.mem self.fun_symbols f then (
(* add entry for [f], to build a if-then-else tree *)
match List.map (fun t -> sat.get_value t |> Option.get) args with
| exception _ ->
Log.debugf 1 (fun k ->
k "(@[build-model.warn@ :no-entry-for %a@])" Term.pp t);
() (* TODO: warning? *)
| v_args ->
(* see if [v_args] already maps to a value *)
let other_v = Model.get_fun_entry f v_args !m in
Option.iter
(fun v' ->
if not (Term.equal value v') then
Error.errorf
"Inconsistent model@ for fun `%a`,@ values %a@ map to `%a` and \
`%a`"
Term.pp f (Fmt.Dump.list Term.pp) v_args Term.pp value Term.pp
v')
other_v;
(* save mapping *)
m := Model.add_fun_entry f v_args value !m
)
in
sat.iter_classes (fun (cls, v) -> cls (fun t -> examine_term t ~value:v));
!m

View file

@ -0,0 +1,9 @@
open Common_
open! Sidekick_base
type t
val create : unit -> t
val add_ty : t -> Term.t -> unit
val add_fun : t -> Term.t -> unit
val build : t -> Solver.sat_result -> Model.t

View file

@ -4,6 +4,5 @@
(synopsis "SMTLIB 2.6 driver for Sidekick") (synopsis "SMTLIB 2.6 driver for Sidekick")
(private_modules Common_) (private_modules Common_)
(libraries containers zarith sidekick.core sidekick.util sidekick-base (libraries containers zarith sidekick.core sidekick.util sidekick-base
sidekick.abstract-solver sidekick.abstract-solver sidekick.mini-cc smtlib-utils sidekick.tef)
sidekick.mini-cc smtlib-utils sidekick.tef)
(flags :standard -warn-error -a+8 -open Sidekick_util)) (flags :standard -warn-error -a+8 -open Sidekick_util))

77
src/smtlib/model.ml Normal file
View file

@ -0,0 +1,77 @@
open Sidekick_core
module T = Term
module TM = Term.Map
type value = Term.t
type fun_ = Term.t
module TL_map = CCMap.Make (struct
type t = value list
let compare = CCList.compare Term.compare
end)
type t = { m: value TL_map.t TM.t } [@@unboxed]
let empty : t = { m = T.Map.empty }
let is_empty self = T.Map.is_empty self.m
let iter_fun_entries (self : t) = TM.to_iter self.m
let get_fun_entries f self = TM.get f self.m
let get_fun_entry f vs self =
match get_fun_entries f self with
| None -> None
| Some tm -> TL_map.get vs tm
let add_fun_entry f vs v self =
let m = TM.get_or ~default:TL_map.empty f self.m in
{ m = TM.add f (TL_map.add vs v m) self.m }
let rec eval t (self : t) : value option =
let eval_exn t =
match eval t self with
| Some v -> v
| None -> raise Not_found
in
let f, args = Term.unfold_app t in
match List.map eval_exn args with
| exception Not_found -> None
| v_args -> get_fun_entry f v_args self
let pp out (self : t) =
if is_empty self then
Fmt.string out "(model)"
else (
let rec pp_entries out = function
| [] -> ()
| ([], v) :: _ | [ (_, v) ] -> Term.pp out v
| ((_ :: _ as vs), v) :: tl ->
let pp_guard out () =
match vs with
| [] -> ()
| [ t ] -> Fmt.fprintf out "(@[= x0 %a@])" Term.pp t
| _ ->
Fmt.fprintf out "(@[and@ ";
List.iteri
(fun i t -> Fmt.fprintf out "(@[= x%d %a@])" i Term.pp t)
vs;
Fmt.fprintf out "@])"
in
Fmt.fprintf out "(@[ite %a@ %a@ %a@])" pp_guard () Term.pp v pp_entries
tl
in
let pp_fun out (f, entries) =
match TL_map.choose_opt entries with
| None -> ()
| Some (args, v) ->
let pp_arg out (i, ty) = Fmt.fprintf out "(@[x%d %a@])" i Term.pp ty in
Fmt.fprintf out "(@[<1>define-fun %a (@[%a@])@ %a@ %a@])" Term.pp f
(Util.pp_list ~sep:"" pp_arg)
(List.mapi (fun i v -> i, Term.ty v) args)
Term.pp (Term.ty v) pp_entries (TL_map.to_list entries)
in
Fmt.fprintf out "(@[<hv>model@ %a@])" (Util.pp_iter pp_fun)
(TM.to_iter self.m)
)

22
src/smtlib/model.mli Normal file
View file

@ -0,0 +1,22 @@
(** Models
A model can be produced when the solver is found to be in a
satisfiable state after a call to {!solve}. *)
open Sidekick_core
type t
type value = Term.t
type fun_ = Term.t
module TL_map : CCMap.S with type key = value list
val empty : t
val is_empty : t -> bool
val add_fun_entry : fun_ -> value list -> value -> t -> t
val get_fun_entry : fun_ -> value list -> t -> value option
val get_fun_entries : fun_ -> t -> value TL_map.t option
val iter_fun_entries : t -> (fun_ * value TL_map.t) Iter.t
val eval : Term.t -> t -> value option
include Sidekick_sigs.PRINT with type t := t