refactor: get SAT properly again on some problems

This commit is contained in:
Simon Cruanes 2018-05-20 14:30:13 -05:00
parent 4a39192846
commit fade033458
21 changed files with 114 additions and 139 deletions

View file

@ -147,41 +147,4 @@ module Make(S : Sidekick_sat.S)(A : Arg with type atom := S.atom
end end
module Simple(S : Sidekick_sat.S) module Simple(S : Sidekick_sat.S) = Make(S)(Default(S))
(A : Arg with type atom := S.atom
and type hyp = S.atom list
and type lemma := S.lemma
and type assumption = S.atom) =
Make(S)(struct
module Atom = S.Atom
module Clause = S.Clause
module P = S.Proof
(* Some helpers *)
let get_form = Atom.get_formula
let get_assumption c =
match Clause.atoms_l c with
| [ x ] -> x
| _ -> assert false
let get_lemma c =
match P.expand (P.prove c) with
| {P.step=P.Lemma p;_} -> p
| _ -> assert false
(* Actual functions *)
let print_atom fmt a =
A.print_atom fmt a
let hyp_info c =
A.hyp_info (Clause.atoms_l c)
let lemma_info c =
A.lemma_info (get_lemma c)
let assumption_info c =
A.assumption_info (get_assumption c)
end)

View file

@ -60,10 +60,7 @@ module Make(S : Sidekick_sat.S)(A : Arg with type atom := S.atom
and type assumption := S.clause) : S with type t := S.proof and type assumption := S.clause) : S with type t := S.proof
(** Functor for making a module to export proofs to the DOT format. *) (** Functor for making a module to export proofs to the DOT format. *)
module Simple(S : Sidekick_sat.S)(A : Arg with type atom := S.atom module Simple(S : Sidekick_sat.S) : S with type t := S.proof
and type hyp = S.atom list
and type lemma := S.lemma
and type assumption = S.atom) : S with type t := S.proof
(** Functor for making a module to export proofs to the DOT format. (** Functor for making a module to export proofs to the DOT format.
The substitution of the hyp type is non-destructive due to a restriction The substitution of the hyp type is non-destructive due to a restriction
of destructive substitutions on earlier versions of ocaml. *) of destructive substitutions on earlier versions of ocaml. *)

View file

@ -15,7 +15,7 @@ rule token = parse
| '\n' { Lexing.new_line lexbuf; token lexbuf } | '\n' { Lexing.new_line lexbuf; token lexbuf }
| '0' { ZERO } | '0' { ZERO }
| '-'? number { LIT (int_of_string (Lexing.lexeme lexbuf)) } | '-'? number { LIT (int_of_string (Lexing.lexeme lexbuf)) }
| _ { Util.errorf "dimacs.lexer: unexpected char `%s`" (Lexing.lexeme lexbuf) } | _ { Error.errorf "dimacs.lexer: unexpected char `%s`" (Lexing.lexeme lexbuf) }
and comment = parse and comment = parse
| '\n' { Lexing.new_line lexbuf; token lexbuf } | '\n' { Lexing.new_line lexbuf; token lexbuf }

View file

@ -25,14 +25,14 @@ prelude:
| P CNF LIT LIT { () } | P CNF LIT LIT { () }
| error | error
{ {
Util.errorf "expected prelude %a" pp_pos ($startpos,$endpos) Error.errorf "expected prelude %a" pp_pos ($startpos,$endpos)
} }
clauses: clauses:
| l=clause* { l } | l=clause* { l }
| error | error
{ {
Util.errorf "expected list of clauses %a" Error.errorf "expected list of clauses %a"
pp_pos ($startpos,$endpos) pp_pos ($startpos,$endpos)
} }

View file

@ -34,26 +34,6 @@ let p_progress = ref false
let hyps : Ast.term list ref = ref [] let hyps : Ast.term list ref = ref []
module Dot = Sidekick_backend.Dot.Make(Solver.Sat_solver)
let check_model _state =
let check_clause _c =
true
(* FIXME
let l =
List.map
(fun a ->
Sidekick_sat.Log.debugf 99
(fun k -> k "Checking value of %a" Term.pp (Sat.Atom.term a));
Sat_solver.Sat_state.eval state a)
c
in
List.exists (fun x -> x) l
*)
in
let l = List.map check_clause !hyps in
List.for_all (fun x -> x) l
(* Arguments parsing *) (* Arguments parsing *)
let int_arg r arg = let int_arg r arg =
let l = String.length arg in let l = String.length arg in
@ -170,7 +150,7 @@ let () = match main() with
| exception e -> | exception e ->
let b = Printexc.get_backtrace () in let b = Printexc.get_backtrace () in
begin match e with begin match e with
| Util.Error msg -> | Error.Error msg ->
Format.printf "@{<Red>Error@}: %s@." msg; Format.printf "@{<Red>Error@}: %s@." msg;
ignore @@ exit 1 ignore @@ exit 1
| Out_of_time -> | Out_of_time ->

View file

@ -53,7 +53,7 @@ module Make (Th : Theory_intf.S) = struct
and reason = and reason =
| Decision | Decision
| Bcp of clause | Bcp of clause
| Local | Local of clause
and premise = and premise =
| Hyp | Hyp
@ -352,7 +352,7 @@ module Make (Th : Theory_intf.S) = struct
Format.fprintf fmt "%%" Format.fprintf fmt "%%"
| n, None -> | n, None ->
Format.fprintf fmt "%d" n Format.fprintf fmt "%d" n
| n, Some Local -> | n, Some Local _ ->
Format.fprintf fmt "L%d" n Format.fprintf fmt "L%d" n
| n, Some Decision -> | n, Some Decision ->
Format.fprintf fmt "@@%d" n Format.fprintf fmt "@@%d" n
@ -362,17 +362,17 @@ module Make (Th : Theory_intf.S) = struct
let pp_level fmt a = let pp_level fmt a =
debug_reason fmt (a.var.v_level, a.var.reason) debug_reason fmt (a.var.v_level, a.var.reason)
let debug_value fmt a = let debug_value out a =
if a.is_true then if a.is_true then (
Format.fprintf fmt "T%a" pp_level a Format.fprintf out "[T%a]" pp_level a
else if a.neg.is_true then ) else if a.neg.is_true then (
Format.fprintf fmt "F%a" pp_level a Format.fprintf out "[F%a]" pp_level a
else ) else (
Format.fprintf fmt "" CCFormat.string out "[?]"
)
let debug out a = let debug out a =
Format.fprintf out "@[%s%d[%a][@[%a@]]@]" Format.fprintf out "@[@[%a@]%a@]" Th.Form.print a.lit debug_value a
(sign a) (a.var.vid+1) debug_value a Th.Form.print a.lit
let debug_a out vec = let debug_a out vec =
Array.iteri (fun i a -> if i>0 then Format.fprintf out "@ "; debug out a) vec Array.iteri (fun i a -> if i>0 then Format.fprintf out "@ "; debug out a) vec
@ -615,14 +615,14 @@ module Make (Th : Theory_intf.S) = struct
let[@inline] conclusion (p:proof) : clause = p let[@inline] conclusion (p:proof) : clause = p
let expand conclusion = let expand conclusion =
Log.debugf 5 (fun k -> k "Expanding : @[%a@]" Clause.debug conclusion); Log.debugf 5 (fun k -> k "(@[proof.expand@ @[%a@]@])" Clause.debug conclusion);
match conclusion.c_premise with match conclusion.c_premise with
| Lemma l -> | Lemma l ->
{conclusion; step = Lemma l; } {conclusion; step = Lemma l; }
| Hyp -> | Hyp ->
{ conclusion; step = Hypothesis; } { conclusion; step = Hypothesis; }
| History [] -> | History [] ->
Log.debugf 1 (fun k -> k "Empty history for clause: %a" Clause.debug conclusion); Log.debugf 1 (fun k -> k "(@[proof.empty_history@ %a@]" Clause.debug conclusion);
raise (Resolution_error "Empty history") raise (Resolution_error "Empty history")
| Simplified c -> | Simplified c ->
let duplicates, res = analyze (list c) in let duplicates, res = analyze (list c) in
@ -974,7 +974,7 @@ module Make (Th : Theory_intf.S) = struct
Wrapper function for adding a new propagated formula. *) Wrapper function for adding a new propagated formula. *)
let enqueue_bool st a reason : unit = let enqueue_bool st a reason : unit =
if a.neg.is_true then ( if a.neg.is_true then (
Util.errorf "(@[sat.enqueue_bool.1@ :false-literal %a@])" Atom.debug a Error.errorf "(@[sat.enqueue_bool.1@ :false-literal %a@])" Atom.debug a
); );
let level = decision_level st in let level = decision_level st in
Log.debugf 5 Log.debugf 5
@ -1000,11 +1000,17 @@ module Make (Th : Theory_intf.S) = struct
a.(j) <- tmp; a.(j) <- tmp;
) )
(* move atoms assigned at high levels first *) (* move atoms that are not assigned first,
else put atoms assigned at high levels first *)
let put_high_level_atoms_first (arr:atom array) : unit = let put_high_level_atoms_first (arr:atom array) : unit =
(* move [a] in front of [b]? *)
let[@inline] put_before a b =
if Atom.level a < 0 then Atom.level b >= 0
else (Atom.level b >= 0 && Atom.level a > Atom.level b)
in
Array.iteri Array.iteri
(fun i a -> (fun i a ->
if i>0 && Atom.level a > Atom.level arr.(0) then ( if i>0 && put_before a arr.(0) then (
(* move first to second, [i]-th to first, second to [i] *) (* move first to second, [i]-th to first, second to [i] *)
if i=1 then ( if i=1 then (
swap_arr arr 0 1; swap_arr arr 0 1;
@ -1014,7 +1020,7 @@ module Make (Th : Theory_intf.S) = struct
arr.(0) <- arr.(i); arr.(0) <- arr.(i);
arr.(i) <- tmp; arr.(i) <- tmp;
); );
) else if i>1 && Atom.level a > Atom.level arr.(1) then ( ) else if i>1 && put_before a arr.(1) then (
swap_arr arr 1 i; swap_arr arr 1 i;
)) ))
arr arr
@ -1214,7 +1220,9 @@ module Make (Th : Theory_intf.S) = struct
- report unsat if conflict at level 0 - report unsat if conflict at level 0
*) *)
and add_boolean_conflict st (confl:clause): unit = and add_boolean_conflict st (confl:clause): unit =
Log.debugf 3 (fun k -> k "(@[@{<Yellow>sat.boolean-conflict@}@ %a@])" Clause.debug confl); Log.debugf 3
(fun k -> k "(@[@{<Yellow>sat.boolean-conflict@}@ %a@ :decision-lvl %d@])"
Clause.debug confl (decision_level st));
st.next_decision <- None; st.next_decision <- None;
assert (decision_level st >= base_level st); assert (decision_level st >= base_level st);
if decision_level st = base_level st || if decision_level st = base_level st ||
@ -1336,7 +1344,7 @@ module Make (Th : Theory_intf.S) = struct
enqueue_bool st p (Bcp c) enqueue_bool st p (Bcp c)
) )
) else ( ) else (
Util.errorf "(@[sat.act_propagate.invalid_guard@ :guard %a@ \ Error.errorf "(@[sat.act_propagate.invalid_guard@ :guard %a@ \
:1 all lits are not true@])" :1 all lits are not true@])"
(Util.pp_list Atom.debug) l (Util.pp_list Atom.debug) l
) )
@ -1502,21 +1510,20 @@ module Make (Th : Theory_intf.S) = struct
let a = Atom.make st lit in let a = Atom.make st lit in
Log.debugf 3 (fun k-> k "(@[sat.local_assumption@ %a@])" Atom.debug a); Log.debugf 3 (fun k-> k "(@[sat.local_assumption@ %a@])" Atom.debug a);
assert (decision_level st = base_level st); assert (decision_level st = base_level st);
let c = Clause.make [|a|] Hyp in
if not a.is_true then ( if not a.is_true then (
if a.neg.is_true then ( if a.neg.is_true then (
(* conflict between assumptions: UNSAT *) (* conflict between assumptions: UNSAT *)
let c = Clause.make [|a|] Hyp in
report_unsat st c; report_unsat st c;
) else ( ) else (
(* make a decision, propagate *) (* make a decision, propagate *)
enqueue_bool st a Local; enqueue_bool st a (Local c);
) )
) )
in in
match st.unsat_conflict with match st.unsat_conflict with
| None -> | None ->
Log.debug 3 "(sat.adding_local_assumptions)"; Log.debug 3 "(sat.adding_local_assumptions)";
cancel_until st (base_level st);
List.iter add_lit assumptions List.iter add_lit assumptions
| Some _ -> | Some _ ->
Log.debug 2 "(sat.local_assumptions.1: already unsat)" Log.debug 2 "(sat.local_assumptions.1: already unsat)"
@ -1568,7 +1575,9 @@ module Make (Th : Theory_intf.S) = struct
end end
in in
try try
cancel_until st 0;
local st assumptions; local st assumptions;
new_decision_level st;
add_clauses() add_clauses()
with Sat -> () with Sat -> ()
@ -1609,6 +1618,15 @@ module Make (Th : Theory_intf.S) = struct
(* Unsafe access to internal data *) (* Unsafe access to internal data *)
let trail st = st.trail let trail st = st.trail
let check_model st : unit =
Log.debug 5 "(sat.check-model)";
Vec.iter
(fun c ->
if not (Array.exists Atom.is_true c.atoms) then (
Error.errorf "(@[sat.check-model.invalid-model@ :clause %a@])" Clause.debug c
))
st.clauses
end end
[@@inline] [@@inline]

View file

@ -99,6 +99,8 @@ module Make (Th : Theory_intf.S) = struct
let clauses = st.S.clauses in let clauses = st.S.clauses in
{clauses} {clauses}
let check_model = S.check_model
module Atom = S.Atom module Atom = S.Atom
module Clause = struct module Clause = struct

View file

@ -138,6 +138,8 @@ module type S = sig
val export : t -> clause export val export : t -> clause export
val check_model : t -> unit
(** {2 Re-export some functions} *) (** {2 Re-export some functions} *)
type solver = t type solver = t
@ -235,7 +237,7 @@ module type S = sig
(** Check the contents of a proof. Mainly for internal use *) (** Check the contents of a proof. Mainly for internal use *)
module Tbl : Hashtbl.S with type key = t module Tbl : Hashtbl.S with type key = t
end end
module Clause : sig module Clause : sig
type t = clause type t = clause

View file

@ -93,7 +93,7 @@ module Ty = struct
end) end)
let ill_typed fmt = let ill_typed fmt =
Util.errorf ("ill-typed: " ^^ fmt) Error.errorf ("ill-typed: " ^^ fmt)
end end
type var = Ty.t Var.t type var = Ty.t Var.t
@ -393,7 +393,7 @@ let num_z ty z = mk_ (Num_z z) ty
let parse_num ~where (s:string) : [`Q of Q.t | `Z of Z.t] = let parse_num ~where (s:string) : [`Q of Q.t | `Z of Z.t] =
let fail() = let fail() =
Util.errorf "%sexpected number, got `%s`" (Lazy.force where) s Error.errorf "%sexpected number, got `%s`" (Lazy.force where) s
in in
begin match Z.of_string s with begin match Z.of_string s with
| n -> `Z n | n -> `Z n

View file

@ -18,6 +18,8 @@ type t = {
(* constant -> its value *) (* constant -> its value *)
} }
let empty : t = { env=A.env_empty; domains=A.Ty.Map.empty; consts=ID.Map.empty}
let make ~env ~consts ~domains = let make ~env ~consts ~domains =
(* also add domains to [env] *) (* also add domains to [env] *)
let env = let env =
@ -169,7 +171,7 @@ let rec eval_whnf (m:t) (st:term list) (subst:subst) (t:term): term =
let st = t :: st in let st = t :: st in
try try
eval_whnf_rec m st subst t eval_whnf_rec m st subst t
with Util.Error msg -> with Error.Error msg ->
errorf "@[<2>Model:@ internal type error `%s`@ in %a@]" msg pp_stack st errorf "@[<2>Model:@ internal type error `%s`@ in %a@]" msg pp_stack st
and eval_whnf_rec m st subst t = match A.term_view t with and eval_whnf_rec m st subst t = match A.term_view t with
| A.Num_q _ | A.Num_q _

View file

@ -16,6 +16,10 @@ type t = private {
(* constant -> its value *) (* constant -> its value *)
} }
(* TODO: remove *)
(** Trivial model *)
val empty : t
val make : val make :
env:Ast.env -> env:Ast.env ->
consts:term ID.Map.t -> consts:term ID.Map.t ->

View file

@ -19,6 +19,8 @@ let[@inline] clause_of_mclause (c:Sat_solver.clause): Lit.t IArray.t =
module Proof = struct module Proof = struct
type t = Sat_solver.Proof.t type t = Sat_solver.Proof.t
let check = Sat_solver.Proof.check
let pp out (p:t) : unit = let pp out (p:t) : unit =
let pp_step_res out p = let pp_step_res out p =
let {Sat_solver.Proof.conclusion; _ } = Sat_solver.Proof.expand p in let {Sat_solver.Proof.conclusion; _ } = Sat_solver.Proof.expand p in
@ -398,6 +400,8 @@ let[@inline] assume_eq self t u expl : unit =
let[@inline] assume_distinct self l ~neq expl : unit = let[@inline] assume_distinct self l ~neq expl : unit =
Congruence_closure.assert_distinct (cc self) l (E_lit expl) ~neq Congruence_closure.assert_distinct (cc self) l (E_lit expl) ~neq
let check_model (s:t) = Sat_solver.check_model s.solver
(* (*
type unsat_core = Sat.clause list type unsat_core = Sat.clause list
*) *)
@ -409,11 +413,13 @@ let solve ?on_exit:(_=[]) ?check:(_=true) ~assumptions (self:t) : res =
match r with match r with
| Sat_solver.Sat (Sidekick_sat.Sat_state _st) -> | Sat_solver.Sat (Sidekick_sat.Sat_state _st) ->
Log.debugf 0 (fun k->k "SAT"); Log.debugf 0 (fun k->k "SAT");
Unknown U_incomplete (* TODO *) Sat Model.empty
(* (*
let env = Ast.env_empty in let env = Ast.env_empty in
let m = Model.make ~env let m = Model.make ~env in
Sat m *)
Unknown U_incomplete (* TODO *)
*)
| Sat_solver.Unsat (Sidekick_sat.Unsat_state us) -> | Sat_solver.Unsat (Sidekick_sat.Unsat_state us) ->
let pr = us.get_proof () in let pr = us.get_proof () in
Unsat pr Unsat pr

View file

@ -17,6 +17,7 @@ type model = Model.t
module Proof : sig module Proof : sig
type t = Sat_solver.Proof.t type t = Sat_solver.Proof.t
val check : t -> unit
val pp : t CCFormat.printer val pp : t CCFormat.printer
end end
@ -63,6 +64,8 @@ val solve :
@param check if true, the model is checked before returning @param check if true, the model is checked before returning
@param on_exit functions to be run before this returns *) @param on_exit functions to be run before this returns *)
val check_model : t -> unit
val pp_term_graph: t CCFormat.printer val pp_term_graph: t CCFormat.printer
val pp_stats : t CCFormat.printer val pp_stats : t CCFormat.printer
val pp_unknown : unknown CCFormat.printer val pp_unknown : unknown CCFormat.printer

View file

@ -264,7 +264,7 @@ let tseitin (self:t) (lit:Lit.t) (lit_t:term) (b:term builtin) : unit =
self.acts.Theory.propagate_distinct l ~neq:lit_t (Explanation.lit lit) self.acts.Theory.propagate_distinct l ~neq:lit_t (Explanation.lit lit)
) else ( ) else (
(* TODO: propagate pairwise equalities? *) (* TODO: propagate pairwise equalities? *)
Util.errorf "cannot process negative distinct lit %a" Lit.pp lit; Error.errorf "cannot process negative distinct lit %a" Lit.pp lit;
) )
| B_and subs -> | B_and subs ->
if Lit.sign lit then ( if Lit.sign lit then (

View file

@ -9,6 +9,7 @@ module E = CCResult
module A = Ast module A = Ast
module Form = Sidekick_th_bool module Form = Sidekick_th_bool
module Fmt = CCFormat module Fmt = CCFormat
module Dot = Sidekick_backend.Dot.Simple(Solver.Sat_solver)
module Subst = struct module Subst = struct
type 'a t = 'a ID.Map.t type 'a t = 'a ID.Map.t
@ -17,7 +18,7 @@ module Subst = struct
let pp pp_x out = ID.Map.pp ~arrow:"" ID.pp pp_x out let pp pp_x out = ID.Map.pp ~arrow:"" ID.pp pp_x out
let add subst v t = let add subst v t =
if mem subst v then ( if mem subst v then (
Util.errorf "%a already bound" A.Var.pp v; Error.errorf "%a already bound" A.Var.pp v;
); );
ID.Map.add (A.Var.id v) t subst ID.Map.add (A.Var.id v) t subst
let find subst v = ID.Map.get (A.Var.id v) subst let find subst v = ID.Map.get (A.Var.id v) subst
@ -33,9 +34,9 @@ module Conv = struct
(* | A.Ty.Rat -> Reg.find_exn reg Mc2_lra.k_rat *) (* | A.Ty.Rat -> Reg.find_exn reg Mc2_lra.k_rat *)
| A.Ty.App (id, []) -> mk_ty id | A.Ty.App (id, []) -> mk_ty id
| A.Ty.App (_, _) -> | A.Ty.App (_, _) ->
Util.errorf "cannot convert parametrized type %a" A.Ty.pp ty Error.errorf "cannot convert parametrized type %a" A.Ty.pp ty
| A.Ty.Arrow _ -> | A.Ty.Arrow _ ->
Util.errorf "cannot convert arrow type `%a`" A.Ty.pp ty Error.errorf "cannot convert arrow type `%a`" A.Ty.pp ty
in in
aux_ty ty aux_ty ty
@ -83,7 +84,7 @@ module Conv = struct
begin match A.term_view t with begin match A.term_view t with
| A.Var v -> | A.Var v ->
begin match Subst.find subst v with begin match Subst.find subst v with
| None -> Util.errorf "variable %a not bound" A.Var.pp v | None -> Error.errorf "variable %a not bound" A.Var.pp v
| Some t -> t | Some t -> t
end end
| A.Const id -> | A.Const id ->
@ -94,7 +95,7 @@ module Conv = struct
| A.Const id -> | A.Const id ->
(* TODO: lookup definition of [f] *) (* TODO: lookup definition of [f] *)
mk_app (Cst.make_undef id (A.ty f |> conv_ty)) l mk_app (Cst.make_undef id (A.ty f |> conv_ty)) l
| _ -> Util.errorf "cannot process HO application %a" A.pp_term t | _ -> Error.errorf "cannot process HO application %a" A.pp_term t
end end
| A.If (a,b,c) -> | A.If (a,b,c) ->
let a = aux subst a in let a = aux subst a in
@ -140,7 +141,7 @@ module Conv = struct
begin match op, l with begin match op, l with
| A.Minus, [a] -> RLE.neg a |> ret_rat | A.Minus, [a] -> RLE.neg a |> ret_rat
| _, [] | _, [_] -> | _, [] | _, [_] ->
Util.errorf "ill-formed arith expr:@ %a@ (need ≥ 2 args)" A.pp_term t Error.errorf "ill-formed arith expr:@ %a@ (need ≥ 2 args)" A.pp_term t
| A.Leq, [a;b] -> | A.Leq, [a;b] ->
let e = RLE.diff a b in let e = RLE.diff a b in
mk_lra_pred Mc2_lra.Leq0 e |> ret_any mk_lra_pred Mc2_lra.Leq0 e |> ret_any
@ -154,7 +155,7 @@ module Conv = struct
let e = RLE.diff b a in let e = RLE.diff b a in
mk_lra_pred Mc2_lra.Lt0 e |> ret_any mk_lra_pred Mc2_lra.Lt0 e |> ret_any
| (A.Leq | A.Lt | A.Geq | A.Gt), _ -> | (A.Leq | A.Lt | A.Geq | A.Gt), _ ->
Util.errorf "ill-formed arith expr:@ %a@ (binary operator)" A.pp_term t Error.errorf "ill-formed arith expr:@ %a@ (binary operator)" A.pp_term t
| A.Add, _ -> | A.Add, _ ->
let e = List.fold_left (fun n t -> RLE.add t n) RLE.empty l in let e = List.fold_left (fun n t -> RLE.add t n) RLE.empty l in
mk_lra_expr e |> ret_t mk_lra_expr e |> ret_t
@ -179,7 +180,7 @@ module Conv = struct
| _, [t] -> | _, [t] ->
List.fold_right RLE.mult coeffs t |> ret_rat List.fold_right RLE.mult coeffs t |> ret_rat
| _ -> | _ ->
Util.errorf "non-linear expr:@ `%a`" A.pp_term t Error.errorf "non-linear expr:@ `%a`" A.pp_term t
end end
| A.Div, (first::l) -> | A.Div, (first::l) ->
(* support t/a/b/c where only [t] is a rational *) (* support t/a/b/c where only [t] is a rational *)
@ -187,7 +188,7 @@ module Conv = struct
List.map List.map
(fun c -> match RLE.as_const c with (fun c -> match RLE.as_const c with
| None -> | None ->
Util.errorf "non-linear expr:@ `%a`" A.pp_term t Error.errorf "non-linear expr:@ `%a`" A.pp_term t
| Some c -> Q.inv c) | Some c -> Q.inv c)
l l
in in
@ -259,7 +260,7 @@ end
let solve let solve
?gc:_ ?gc:_
?restarts:_ ?restarts:_
?dot_proof:_ ?dot_proof
?(pp_model=false) ?(pp_model=false)
?(check=false) ?(check=false)
?time:_ ?memory:_ ?progress:_ ?time:_ ?memory:_ ?progress:_
@ -275,17 +276,10 @@ let solve
if pp_model then ( if pp_model then (
Format.printf "(@[<hv1>model@ %a@])@." Model.pp m Format.printf "(@[<hv1>model@ %a@])@." Model.pp m
); );
if check then ( if check then Solver.check_model s;
(* TODO
if not (check_model m) then (
Util.error "invalid model"
)
*)
);
let t3 = Sys.time () -. t2 in let t3 = Sys.time () -. t2 in
Format.printf "Sat (%.3f/%.3f/%.3f)@." t1 (t2-.t1) t3; Format.printf "Sat (%.3f/%.3f/%.3f)@." t1 (t2-.t1) t3;
| Solver.Unsat _pr -> | Solver.Unsat p ->
(*
if check then ( if check then (
Solver.Proof.check p; Solver.Proof.check p;
begin match dot_proof with begin match dot_proof with
@ -299,7 +293,6 @@ let solve
Format.pp_print_flush fmt (); flush oc) Format.pp_print_flush fmt (); flush oc)
end end
); );
*)
let t3 = Sys.time () -. t2 in let t3 = Sys.time () -. t2 in
Format.printf "Unsat (%.3f/%.3f/%.3f)@." t1 (t2-.t1) t3; Format.printf "Unsat (%.3f/%.3f/%.3f)@." t1 (t2-.t1) t3;
| Solver.Unknown reas -> | Solver.Unknown reas ->
@ -365,11 +358,11 @@ let process_stmt
Solver.assume solver (IArray.of_list c); Solver.assume solver (IArray.of_list c);
E.return () E.return ()
| A.Goal (_, _) -> | A.Goal (_, _) ->
Util.errorf "cannot deal with goals yet" Error.errorf "cannot deal with goals yet"
| A.Data _ -> | A.Data _ ->
Util.errorf "cannot deal with datatypes yet" Error.errorf "cannot deal with datatypes yet"
| A.Define _ -> | A.Define _ ->
Util.errorf "cannot deal with definitions yet" Error.errorf "cannot deal with definitions yet"
end end
@ -403,7 +396,7 @@ end = struct
| Ast.Ty.Prop -> Ty.prop | Ast.Ty.Prop -> Ty.prop
| Ast.Ty.Const id -> | Ast.Ty.Const id ->
begin try ID.Tbl.find ty_tbl_ id |> Lazy.force begin try ID.Tbl.find ty_tbl_ id |> Lazy.force
with Not_found -> Util.errorf "type %a not in ty_tbl" ID.pp id with Not_found -> Error.errorf "type %a not in ty_tbl" ID.pp id
end end
| Ast.Ty.Arrow (a,b) -> Ty.arrow (conv_ty a) (conv_ty b) | Ast.Ty.Arrow (a,b) -> Ty.arrow (conv_ty a) (conv_ty b)

View file

@ -78,15 +78,15 @@ module Ctx = struct
let find_kind t (id:ID.t) : kind = let find_kind t (id:ID.t) : kind =
try ID.Tbl.find t.kinds id try ID.Tbl.find t.kinds id
with Not_found -> Util.errorf "did not find kind of ID `%a`" ID.pp id with Not_found -> Error.errorf "did not find kind of ID `%a`" ID.pp id
let as_data t (ty:A.Ty.t) : (ID.t * A.Ty.t) list = match ty with let as_data t (ty:A.Ty.t) : (ID.t * A.Ty.t) list = match ty with
| A.Ty.App (id,_) -> | A.Ty.App (id,_) ->
begin match ID.Tbl.get t.data id with begin match ID.Tbl.get t.data id with
| Some l -> l | Some l -> l
| None -> Util.errorf "expected %a to be a datatype" A.Ty.pp ty | None -> Error.errorf "expected %a to be a datatype" A.Ty.pp ty
end end
| _ -> Util.errorf "expected %a to be a constant type" A.Ty.pp ty | _ -> Error.errorf "expected %a to be a constant type" A.Ty.pp ty
let pp_kind out = function let pp_kind out = function
| K_ty _ -> Format.fprintf out "type" | K_ty _ -> Format.fprintf out "type"
@ -107,7 +107,7 @@ end
let error_loc ctx : string = Fmt.sprintf "at %a: " pp_loc_opt (Ctx.loc ctx) let error_loc ctx : string = Fmt.sprintf "at %a: " pp_loc_opt (Ctx.loc ctx)
let errorf_ctx ctx msg = let errorf_ctx ctx msg =
Util.errorf ("at %a:@ " ^^ msg) pp_loc_opt (Ctx.loc ctx) Error.errorf ("at %a:@ " ^^ msg) pp_loc_opt (Ctx.loc ctx)
let check_bool_ t = let check_bool_ t =
if not (A.Ty.equal (A.ty t) A.Ty.prop) then ( if not (A.Ty.equal (A.ty t) A.Ty.prop) then (

View file

@ -7,7 +7,8 @@
((name sidekick_smtlib) ((name sidekick_smtlib)
(public_name sidekick.smtlib) (public_name sidekick.smtlib)
(optional) ; only if deps present (optional) ; only if deps present
(libraries (containers sidekick.smt sidekick.util sidekick.smt.th_bool zarith)) (libraries (containers zarith sidekick.smt sidekick.util
sidekick.smt.th_bool sidekick.backend))
(flags (:standard -w +a-4-42-44-48-50-58-32-60@8 (flags (:standard -w +a-4-42-44-48-50-58-32-60@8
-safe-string -color always -open Sidekick_util)) -safe-string -color always -open Sidekick_util))
(ocamlopt_flags (:standard -O3 -color always -bin-annot (ocamlopt_flags (:standard -O3 -color always -bin-annot

12
src/util/Error.ml Normal file
View file

@ -0,0 +1,12 @@
module Fmt = CCFormat
exception Error of string
let () = Printexc.register_printer
(function
| Error msg -> Some ("internal error: " ^ msg)
| _ -> None)
let errorf msg = Fmt.ksprintf msg ~f:(fun s -> raise (Error s))

6
src/util/Error.mli Normal file
View file

@ -0,0 +1,6 @@
exception Error of string
val errorf : ('a, Format.formatter, unit, 'b) format4 -> 'a
(** @raise Error when called *)

View file

@ -24,15 +24,6 @@ let pp_array ?(sep=" ") pp out l =
let pp_iarray ?(sep=" ") pp out a = let pp_iarray ?(sep=" ") pp out a =
Fmt.seq ~sep:(pp_sep sep) pp out (IArray.to_seq a) Fmt.seq ~sep:(pp_sep sep) pp out (IArray.to_seq a)
exception Error of string
let () = Printexc.register_printer
(function
| Error msg -> Some ("internal error: " ^ msg)
| _ -> None)
let errorf msg = Fmt.ksprintf msg ~f:(fun s -> raise (Error s))
let setup_gc () = let setup_gc () =
let g = Gc.get () in let g = Gc.get () in
g.Gc.space_overhead <- 3_000; (* major gc *) g.Gc.space_overhead <- 3_000; (* major gc *)

View file

@ -15,11 +15,6 @@ val pp_pair : ?sep:string -> 'a printer -> 'b printer -> ('a * 'b) printer
val pp_iarray : ?sep:string -> 'a CCFormat.printer -> 'a IArray.t CCFormat.printer val pp_iarray : ?sep:string -> 'a CCFormat.printer -> 'a IArray.t CCFormat.printer
exception Error of string
val errorf : ('a, Format.formatter, unit, 'b) format4 -> 'a
(** @raise Error when called *)
val setup_gc : unit -> unit val setup_gc : unit -> unit
(** Change parameters of the GC *) (** Change parameters of the GC *)