mirror of
https://github.com/c-cube/sidekick.git
synced 2025-12-06 11:15:43 -05:00
fix: first version that seems to work on QF_UF
This commit is contained in:
parent
a47641ecea
commit
e3e964a4c6
4 changed files with 34 additions and 19 deletions
|
|
@ -604,7 +604,7 @@ module Make(CC_A: ARG) = struct
|
|||
let lits = explain_unfold cc e_ab in
|
||||
let lits = explain_eq_n ~init:lits cc a ra in
|
||||
let lits = explain_eq_n ~init:lits cc b rb in
|
||||
raise_conflict cc acts lits
|
||||
raise_conflict cc acts (List.rev_map Lit.neg lits)
|
||||
);
|
||||
(* We will merge [r_from] into [r_into].
|
||||
we try to ensure that [size ra <= size rb] in general, but always
|
||||
|
|
@ -801,8 +801,8 @@ module Make(CC_A: ARG) = struct
|
|||
(fun k->k "(@[cc.theory.raise-conflict@ :expl %a@])" Expl.pp expl);
|
||||
ps_clear cc;
|
||||
decompose_explain cc expl;
|
||||
let lits = cc.ps_lits in
|
||||
CC_A.Actions.raise_conflict acts lits A.Proof.default
|
||||
let lits = List.rev_map Lit.neg cc.ps_lits in
|
||||
raise_conflict cc acts lits
|
||||
|
||||
let merge cc n1 n2 expl =
|
||||
Log.debugf 5
|
||||
|
|
|
|||
|
|
@ -26,7 +26,7 @@ module CC_view = struct
|
|||
| Not of 't
|
||||
| Opaque of 't (* do not enter *)
|
||||
|
||||
let[@inline] map_view ~f_f ~f_t ~f_ts (v:_ t) : _ t =
|
||||
let map_view ~f_f ~f_t ~f_ts (v:_ t) : _ t =
|
||||
match v with
|
||||
| Bool b -> Bool b
|
||||
| App_fun (f, args) -> App_fun (f_f f, f_ts args)
|
||||
|
|
@ -431,7 +431,7 @@ module type SOLVER_INTERNAL = sig
|
|||
literals suitable for reasoning.
|
||||
Typically some clauses are also added to the solver. *)
|
||||
|
||||
type preprocess_hook = t -> actions -> term -> term option
|
||||
type preprocess_hook = t -> add_clause:(lit list -> unit) -> term -> term option
|
||||
(** Given a term, try to preprocess it. Return [None] if it didn't change.
|
||||
Can also add clauses to define new terms. *)
|
||||
|
||||
|
|
|
|||
|
|
@ -118,7 +118,7 @@ module Make(A : ARG)
|
|||
mutable on_partial_check: (t -> actions -> lit Iter.t -> unit) list;
|
||||
mutable on_final_check: (t -> actions -> lit Iter.t -> unit) list;
|
||||
}
|
||||
and preprocess_hook = t -> actions -> term -> term option
|
||||
and preprocess_hook = t -> add_clause:(lit list -> unit) -> term -> term option
|
||||
|
||||
type solver = t
|
||||
|
||||
|
|
@ -158,7 +158,7 @@ module Make(A : ARG)
|
|||
Stat.incr self.count_axiom;
|
||||
acts.Msat.acts_add_clause ~keep lits A.Proof.default
|
||||
|
||||
let preprocess_lit_ (self:t) (acts:actions) (lit:lit) : lit =
|
||||
let preprocess_lit_ (self:t) ~add_clause (lit:lit) : lit =
|
||||
(* compute and cache normal form of [t] *)
|
||||
let rec aux t =
|
||||
match T.Tbl.find self.preprocess_cache t with
|
||||
|
|
@ -174,15 +174,22 @@ module Make(A : ARG)
|
|||
and aux_rec t hooks = match hooks with
|
||||
| [] -> t
|
||||
| h :: hooks_tl ->
|
||||
match h self acts t with
|
||||
match h self ~add_clause t with
|
||||
| None -> aux_rec t hooks_tl
|
||||
| Some u -> aux u
|
||||
| Some u ->
|
||||
Log.debugf 30
|
||||
(fun k->k "(@[msat-solver.preprocess.step@ :from %a@ :to %a@])"
|
||||
T.pp t T.pp u);
|
||||
aux u
|
||||
in
|
||||
let t = aux (Lit.term lit) in
|
||||
Log.debugf 10
|
||||
(fun k->k "(@[msat-solver.preprocess@ :lit %a@ :into %a@])" Lit.pp lit T.pp t);
|
||||
Lit.atom self.tst ~sign:(Lit.sign lit) t
|
||||
|
||||
let[@inline] mk_lit self acts ?sign t =
|
||||
preprocess_lit_ self acts @@ Lit.atom self.tst ?sign t
|
||||
let add_clause lits = add_sat_clause_ self acts ~keep:true lits in
|
||||
preprocess_lit_ self ~add_clause @@ Lit.atom self.tst ?sign t
|
||||
|
||||
let[@inline] add_clause_temp self acts lits : unit =
|
||||
add_sat_clause_ self acts ~keep:false lits
|
||||
|
|
@ -228,7 +235,7 @@ module Make(A : ARG)
|
|||
(* handle a literal assumed by the SAT solver *)
|
||||
let assert_lits_ ~final (self:t) (acts:actions) (lits:Lit.t Iter.t) : unit =
|
||||
Msat.Log.debugf 2
|
||||
(fun k->k "(@[<hv1>@{<green>th_combine.assume_lits@}%s@ %a@])"
|
||||
(fun k->k "(@[<hv1>@{<green>msat-solver.assume_lits@}%s@ %a@])"
|
||||
(if final then "[final]" else "") (Util.pp_seq ~sep:"; " Lit.pp) lits);
|
||||
(* transmit to CC *)
|
||||
let cc = cc self in
|
||||
|
|
@ -255,7 +262,7 @@ module Make(A : ARG)
|
|||
let check_ ~final (self:t) (acts: msat_acts) =
|
||||
let iter = iter_atoms_ acts in
|
||||
(* TODO if Config.progress then print_progress(); *)
|
||||
Msat.Log.debugf 5 (fun k->k "(th_combine.assume :len %d)" (Iter.length iter));
|
||||
Msat.Log.debugf 5 (fun k->k "(msat-solver.assume :len %d)" (Iter.length iter));
|
||||
assert_lits_ ~final self acts iter
|
||||
|
||||
(* propagation from the bool solver *)
|
||||
|
|
@ -340,7 +347,7 @@ module Make(A : ARG)
|
|||
let add_theory (self:t) (th:theory) : unit =
|
||||
let (module Th) = th in
|
||||
Log.debugf 2
|
||||
(fun k-> k "(@[th_combine.add_th@ :name %S@])" Th.name);
|
||||
(fun k-> k "(@[msat-solver.add-theory@ :name %S@])" Th.name);
|
||||
let st = Th.create_and_setup self.si in
|
||||
(* add push/pop to the internal solver *)
|
||||
begin
|
||||
|
|
@ -358,7 +365,7 @@ module Make(A : ARG)
|
|||
|
||||
(* create a new solver *)
|
||||
let create ?(stat=Stat.global) ?size ?store_proof ~theories tst () : t =
|
||||
Log.debug 5 "th_combine.create";
|
||||
Log.debug 5 "msat-solver.create";
|
||||
let si = Solver_internal.create ~stat tst () in
|
||||
let self = {
|
||||
si;
|
||||
|
|
@ -405,10 +412,17 @@ module Make(A : ARG)
|
|||
| _ -> true)
|
||||
|> Iter.iter
|
||||
(fun sub ->
|
||||
Log.debugf 5 (fun k->k "(@[solver.map-to-lit@ :subterm %a@])" T.pp sub);
|
||||
Log.debugf 5 (fun k->k "(@[solver.map-bool-subterm-to-lit@ :subterm %a@])" T.pp sub);
|
||||
ignore (mk_atom_t_ self sub : Sat_solver.atom))
|
||||
|
||||
let mk_atom_lit self lit : Atom.t =
|
||||
let rec mk_atom_lit self lit : Atom.t =
|
||||
let lit =
|
||||
Solver_internal.preprocess_lit_
|
||||
~add_clause:(fun lits ->
|
||||
(* recursively add these sub-literals, so they're also properly processed *)
|
||||
let atoms = List.map (mk_atom_lit self) lits in
|
||||
Sat_solver.add_clause self.solver atoms A.Proof.default)
|
||||
self.si lit in
|
||||
add_bool_subterms_ self (Lit.term lit);
|
||||
Sat_solver.make_atom self.solver lit
|
||||
|
||||
|
|
@ -467,7 +481,7 @@ module Make(A : ARG)
|
|||
|
||||
(* print all terms reachable from watched literals *)
|
||||
let pp_term_graph _out (_:t) =
|
||||
()
|
||||
() (* TODO *)
|
||||
|
||||
let pp_stats out (self:t) : unit =
|
||||
Stat.pp_all out (Stat.all @@ stats self)
|
||||
|
|
|
|||
|
|
@ -127,8 +127,7 @@ module Make(A : ARG) : S with module A = A = struct
|
|||
Lit.atom self.tst t
|
||||
|
||||
(* TODO: polarity? *)
|
||||
let cnf (self:state) (solver:SI.t) (acts:SI.actions) (t:T.t) : T.t option =
|
||||
let add_clause lits = SI.add_clause_permanent solver acts lits in
|
||||
let cnf (self:state) (_solver:SI.t) ~add_clause (t:T.t) : T.t option =
|
||||
let rec get_lit (t:T.t) : Lit.t =
|
||||
match A.view_as_bool t with
|
||||
| B_bool b -> Lit.atom self.tst ~sign:b (T.bool self.tst true)
|
||||
|
|
@ -172,6 +171,8 @@ module Make(A : ARG) : S with module A = A = struct
|
|||
in
|
||||
let lit = get_lit t in
|
||||
let u = Lit.term lit in
|
||||
(* put sign back as a "not" *)
|
||||
let u = if Lit.sign lit then u else A.mk_bool self.tst (B_not u) in
|
||||
if T.equal u t then None else Some u
|
||||
|
||||
let create_and_setup si =
|
||||
|
|
|
|||
Loading…
Add table
Reference in a new issue