wip: use new sigs

This commit is contained in:
Simon Cruanes 2022-07-15 23:51:53 -04:00
parent ab6bcf8cbe
commit 633a658e0c
No known key found for this signature in database
GPG key ID: EBFFF6F283F3A2B4
8 changed files with 326 additions and 1149 deletions

View file

@ -1,41 +1,37 @@
open Sidekick_core include Sidekick_sigs_cc
module View = Sidekick_core.CC_view open View
type ('f, 't, 'ts) view = ('f, 't, 'ts) View.t = module Make (A : ARG) :
| Bool of bool
| App_fun of 'f * 'ts
| App_ho of 't * 't
| If of 't * 't * 't
| Eq of 't * 't
| Not of 't
| Opaque of 't
(* do not enter *)
module type S = Sidekick_core.CC_S
module Make (A : CC_ARG) :
S S
with module T = A.T with module T = A.T
and module Lit = A.Lit and module Lit = A.Lit
and type proof = A.proof and module Proof_trace = A.Proof_trace = struct
and type proof_step = A.proof_step
and module Actions = A.Actions = struct
module T = A.T module T = A.T
module Lit = A.Lit module Lit = A.Lit
module Actions = A.Actions module Proof_trace = A.Proof_trace
module P = Actions.P module Term = T.Term
module Fun = T.Fun
open struct
(* proof rules *)
module Rules_ = A.CC.Proof_rules
module P = Sidekick_sigs_proof_trace.Utils_ (Proof_trace)
end
type term = T.Term.t type term = T.Term.t
type value = term type value = term
type term_store = T.Term.store type term_store = T.Term.store
type lit = Lit.t type lit = Lit.t
type fun_ = T.Fun.t type fun_ = T.Fun.t
type proof = A.proof type proof = A.Proof_trace.t
type proof_step = A.proof_step type proof_step = A.Proof_trace.step_id
type actions = Actions.t
module Term = T.Term type actions =
module Fun = T.Fun (module ACTIONS
with type term = T.Term.t
and type lit = Lit.t
and type proof = proof
and type proof_step = proof_step)
module Bits : sig module Bits : sig
type t = private int type t = private int
@ -97,7 +93,7 @@ module Make (A : CC_ARG) :
An equivalence class is represented by its "root" element, An equivalence class is represented by its "root" element,
the representative. *) the representative. *)
and signature = (fun_, node, node list) view and signature = (fun_, node, node list) View.t
and explanation_forest_link = and explanation_forest_link =
| FL_none | FL_none
@ -117,7 +113,7 @@ module Make (A : CC_ARG) :
type repr = node type repr = node
module N = struct module Class = struct
type t = node type t = node
let[@inline] equal (n1 : t) n2 = n1 == n2 let[@inline] equal (n1 : t) n2 = n1 == n2
@ -171,11 +167,11 @@ module Make (A : CC_ARG) :
(* non-recursive, inlinable function for [find] *) (* non-recursive, inlinable function for [find] *)
let[@inline] find_ (n : node) : repr = let[@inline] find_ (n : node) : repr =
let n2 = n.n_root in let n2 = n.n_root in
assert (N.is_root n2); assert (Class.is_root n2);
n2 n2
let[@inline] same_class (n1 : node) (n2 : node) : bool = let[@inline] same_class (n1 : node) (n2 : node) : bool =
N.equal (find_ n1) (find_ n2) Class.equal (find_ n1) (find_ n2)
let[@inline] find _ n = find_ n let[@inline] find _ n = find_ n
@ -187,8 +183,9 @@ module Make (A : CC_ARG) :
| E_trivial -> Fmt.string out "reduction" | E_trivial -> Fmt.string out "reduction"
| E_lit lit -> Lit.pp out lit | E_lit lit -> Lit.pp out lit
| E_congruence (n1, n2) -> | E_congruence (n1, n2) ->
Fmt.fprintf out "(@[congruence@ %a@ %a@])" N.pp n1 N.pp n2 Fmt.fprintf out "(@[congruence@ %a@ %a@])" Class.pp n1 Class.pp n2
| E_merge (a, b) -> Fmt.fprintf out "(@[merge@ %a@ %a@])" N.pp a N.pp b | E_merge (a, b) ->
Fmt.fprintf out "(@[merge@ %a@ %a@])" Class.pp a Class.pp b
| E_merge_t (a, b) -> | E_merge_t (a, b) ->
Fmt.fprintf out "(@[<hv>merge@ @[:n1 %a@]@ @[:n2 %a@]@])" Term.pp a Fmt.fprintf out "(@[<hv>merge@ @[:n1 %a@]@ @[:n2 %a@]@])" Term.pp a
Term.pp b Term.pp b
@ -199,13 +196,13 @@ module Make (A : CC_ARG) :
es es
| E_and (a, b) -> Format.fprintf out "(@[<hv1>and@ %a@ %a@])" pp a pp b | E_and (a, b) -> Format.fprintf out "(@[<hv1>and@ %a@ %a@])" pp a pp b
| E_same_val (n1, n2) -> | E_same_val (n1, n2) ->
Fmt.fprintf out "(@[same-value@ %a@ %a@])" N.pp n1 N.pp n2 Fmt.fprintf out "(@[same-value@ %a@ %a@])" Class.pp n1 Class.pp n2
let mk_trivial : t = E_trivial let mk_trivial : t = E_trivial
let[@inline] mk_congruence n1 n2 : t = E_congruence (n1, n2) let[@inline] mk_congruence n1 n2 : t = E_congruence (n1, n2)
let[@inline] mk_merge a b : t = let[@inline] mk_merge a b : t =
if N.equal a b then if Class.equal a b then
mk_trivial mk_trivial
else else
E_merge (a, b) E_merge (a, b)
@ -220,7 +217,7 @@ module Make (A : CC_ARG) :
let[@inline] mk_theory t u es pr = E_theory (t, u, es, pr) let[@inline] mk_theory t u es pr = E_theory (t, u, es, pr)
let[@inline] mk_same_value t u = let[@inline] mk_same_value t u =
if N.equal t u then if Class.equal t u then
mk_trivial mk_trivial
else else
E_same_val (t, u) E_same_val (t, u)
@ -239,7 +236,7 @@ module Make (A : CC_ARG) :
module Resolved_expl = struct module Resolved_expl = struct
type t = { type t = {
lits: lit list; lits: lit list;
same_value: (N.t * N.t) list; same_value: (Class.t * Class.t) list;
pr: proof -> proof_step; pr: proof -> proof_step;
} }
@ -256,7 +253,7 @@ module Make (A : CC_ARG) :
let { lits; same_value; pr = _ } = self in let { lits; same_value; pr = _ } = self in
Fmt.fprintf out "(@[resolved-expl@ (@[%a@])@ :same-val (@[%a@])@])" Fmt.fprintf out "(@[resolved-expl@ (@[%a@])@ :same-val (@[%a@])@])"
(Util.pp_list Lit.pp) lits (Util.pp_list Lit.pp) lits
(Util.pp_list @@ Fmt.Dump.pair N.pp N.pp) (Util.pp_list @@ Fmt.Dump.pair Class.pp Class.pp)
same_value same_value
) )
end end
@ -271,13 +268,14 @@ module Make (A : CC_ARG) :
| Bool b1, Bool b2 -> b1 = b2 | Bool b1, Bool b2 -> b1 = b2
| App_fun (f1, []), App_fun (f2, []) -> Fun.equal f1 f2 | App_fun (f1, []), App_fun (f2, []) -> Fun.equal f1 f2
| App_fun (f1, l1), App_fun (f2, l2) -> | App_fun (f1, l1), App_fun (f2, l2) ->
Fun.equal f1 f2 && CCList.equal N.equal l1 l2 Fun.equal f1 f2 && CCList.equal Class.equal l1 l2
| App_ho (f1, a1), App_ho (f2, a2) -> N.equal f1 f2 && N.equal a1 a2 | App_ho (f1, a1), App_ho (f2, a2) ->
| Not a, Not b -> N.equal a b Class.equal f1 f2 && Class.equal a1 a2
| Not a, Not b -> Class.equal a b
| If (a1, b1, c1), If (a2, b2, c2) -> | If (a1, b1, c1), If (a2, b2, c2) ->
N.equal a1 a2 && N.equal b1 b2 && N.equal c1 c2 Class.equal a1 a2 && Class.equal b1 b2 && Class.equal c1 c2
| Eq (a1, b1), Eq (a2, b2) -> N.equal a1 a2 && N.equal b1 b2 | Eq (a1, b1), Eq (a2, b2) -> Class.equal a1 a2 && Class.equal b1 b2
| Opaque u1, Opaque u2 -> N.equal u1 u2 | Opaque u1, Opaque u2 -> Class.equal u1 u2
| Bool _, _ | Bool _, _
| App_fun _, _ | App_fun _, _
| App_ho _, _ | App_ho _, _
@ -291,24 +289,25 @@ module Make (A : CC_ARG) :
let module H = CCHash in let module H = CCHash in
match s with match s with
| Bool b -> H.combine2 10 (H.bool b) | Bool b -> H.combine2 10 (H.bool b)
| App_fun (f, l) -> H.combine3 20 (Fun.hash f) (H.list N.hash l) | App_fun (f, l) -> H.combine3 20 (Fun.hash f) (H.list Class.hash l)
| App_ho (f, a) -> H.combine3 30 (N.hash f) (N.hash a) | App_ho (f, a) -> H.combine3 30 (Class.hash f) (Class.hash a)
| Eq (a, b) -> H.combine3 40 (N.hash a) (N.hash b) | Eq (a, b) -> H.combine3 40 (Class.hash a) (Class.hash b)
| Opaque u -> H.combine2 50 (N.hash u) | Opaque u -> H.combine2 50 (Class.hash u)
| If (a, b, c) -> H.combine4 60 (N.hash a) (N.hash b) (N.hash c) | If (a, b, c) ->
| Not u -> H.combine2 70 (N.hash u) H.combine4 60 (Class.hash a) (Class.hash b) (Class.hash c)
| Not u -> H.combine2 70 (Class.hash u)
let pp out = function let pp out = function
| Bool b -> Fmt.bool out b | Bool b -> Fmt.bool out b
| App_fun (f, []) -> Fun.pp out f | App_fun (f, []) -> Fun.pp out f
| App_fun (f, l) -> | App_fun (f, l) ->
Fmt.fprintf out "(@[%a@ %a@])" Fun.pp f (Util.pp_list N.pp) l Fmt.fprintf out "(@[%a@ %a@])" Fun.pp f (Util.pp_list Class.pp) l
| App_ho (f, a) -> Fmt.fprintf out "(@[%a@ %a@])" N.pp f N.pp a | App_ho (f, a) -> Fmt.fprintf out "(@[%a@ %a@])" Class.pp f Class.pp a
| Opaque t -> N.pp out t | Opaque t -> Class.pp out t
| Not u -> Fmt.fprintf out "(@[not@ %a@])" N.pp u | Not u -> Fmt.fprintf out "(@[not@ %a@])" Class.pp u
| Eq (a, b) -> Fmt.fprintf out "(@[=@ %a@ %a@])" N.pp a N.pp b | Eq (a, b) -> Fmt.fprintf out "(@[=@ %a@ %a@])" Class.pp a Class.pp b
| If (a, b, c) -> | If (a, b, c) ->
Fmt.fprintf out "(@[ite@ %a@ %a@ %a@])" N.pp a N.pp b N.pp c Fmt.fprintf out "(@[ite@ %a@ %a@ %a@])" Class.pp a Class.pp b Class.pp c
end end
module Sig_tbl = CCHashtbl.Make (Signature) module Sig_tbl = CCHashtbl.Make (Signature)
@ -360,12 +359,12 @@ module Make (A : CC_ARG) :
several times. several times.
See "fast congruence closure and extensions", Nieuwenhuis&al, page 14 *) See "fast congruence closure and extensions", Nieuwenhuis&al, page 14 *)
and ev_on_pre_merge = t -> actions -> N.t -> N.t -> Expl.t -> unit and ev_on_pre_merge = t -> actions -> Class.t -> Class.t -> Expl.t -> unit
and ev_on_post_merge = t -> actions -> N.t -> N.t -> unit and ev_on_post_merge = t -> actions -> Class.t -> Class.t -> unit
and ev_on_new_term = t -> N.t -> term -> unit and ev_on_new_term = t -> Class.t -> term -> unit
and ev_on_conflict = t -> th:bool -> lit list -> unit and ev_on_conflict = t -> th:bool -> lit list -> unit
and ev_on_propagate = t -> lit -> (unit -> lit list * proof_step) -> unit and ev_on_propagate = t -> lit -> (unit -> lit list * proof_step) -> unit
and ev_on_is_subterm = N.t -> term -> unit and ev_on_is_subterm = Class.t -> term -> unit
let[@inline] size_ (r : repr) = r.n_size let[@inline] size_ (r : repr) = r.n_size
let[@inline] n_true cc = Lazy.force cc.true_ let[@inline] n_true cc = Lazy.force cc.true_
@ -387,13 +386,13 @@ module Make (A : CC_ARG) :
let[@inline] on_backtrack cc f : unit = let[@inline] on_backtrack cc f : unit =
Backtrack_stack.push_if_nonzero_level cc.undo f Backtrack_stack.push_if_nonzero_level cc.undo f
let[@inline] get_bitfield _cc field n = N.get_field field n let[@inline] get_bitfield _cc field n = Class.get_field field n
let set_bitfield cc field b n = let set_bitfield cc field b n =
let old = N.get_field field n in let old = Class.get_field field n in
if old <> b then ( if old <> b then (
on_backtrack cc (fun () -> N.set_field field old n); on_backtrack cc (fun () -> Class.set_field field old n);
N.set_field field b n Class.set_field field b n
) )
(* check if [t] is in the congruence closure. (* check if [t] is in the congruence closure.
@ -402,24 +401,26 @@ module Make (A : CC_ARG) :
(* print full state *) (* print full state *)
let pp_full out (cc : t) : unit = let pp_full out (cc : t) : unit =
let pp_next out n = Fmt.fprintf out "@ :next %a" N.pp n.n_next in let pp_next out n = Fmt.fprintf out "@ :next %a" Class.pp n.n_next in
let pp_root out n = let pp_root out n =
if N.is_root n then if Class.is_root n then
Fmt.string out " :is-root" Fmt.string out " :is-root"
else else
Fmt.fprintf out "@ :root %a" N.pp n.n_root Fmt.fprintf out "@ :root %a" Class.pp n.n_root
in in
let pp_expl out n = let pp_expl out n =
match n.n_expl with match n.n_expl with
| FL_none -> () | FL_none -> ()
| FL_some e -> | FL_some e ->
Fmt.fprintf out " (@[:forest %a :expl %a@])" N.pp e.next Expl.pp e.expl Fmt.fprintf out " (@[:forest %a :expl %a@])" Class.pp e.next Expl.pp
e.expl
in in
let pp_n out n = let pp_n out n =
Fmt.fprintf out "(@[%a%a%a%a@])" Term.pp n.n_term pp_root n pp_next n Fmt.fprintf out "(@[%a%a%a%a@])" Term.pp n.n_term pp_root n pp_next n
pp_expl n pp_expl n
and pp_sig_e out (s, n) = and pp_sig_e out (s, n) =
Fmt.fprintf out "(@[<1>%a@ ~~> %a%a@])" Signature.pp s N.pp n pp_root n Fmt.fprintf out "(@[<1>%a@ ~~> %a%a@])" Signature.pp s Class.pp n pp_root
n
in in
Fmt.fprintf out Fmt.fprintf out
"(@[@{<yellow>cc.state@}@ (@[<hv>:nodes@ %a@])@ (@[<hv>:sig-tbl@ %a@])@])" "(@[@{<yellow>cc.state@}@ (@[<hv>:nodes@ %a@])@ (@[<hv>:sig-tbl@ %a@])@])"
@ -441,19 +442,19 @@ module Make (A : CC_ARG) :
let add_signature cc (s : signature) (n : node) : unit = let add_signature cc (s : signature) (n : node) : unit =
assert (not @@ Sig_tbl.mem cc.signatures_tbl s); assert (not @@ Sig_tbl.mem cc.signatures_tbl s);
Log.debugf 50 (fun k -> Log.debugf 50 (fun k ->
k "(@[cc.add-sig@ %a@ ~~> %a@])" Signature.pp s N.pp n); k "(@[cc.add-sig@ %a@ ~~> %a@])" Signature.pp s Class.pp n);
on_backtrack cc (fun () -> Sig_tbl.remove cc.signatures_tbl s); on_backtrack cc (fun () -> Sig_tbl.remove cc.signatures_tbl s);
Sig_tbl.add cc.signatures_tbl s n Sig_tbl.add cc.signatures_tbl s n
let push_pending cc t : unit = let push_pending cc t : unit =
Log.debugf 50 (fun k -> k "(@[<hv1>cc.push-pending@ %a@])" N.pp t); Log.debugf 50 (fun k -> k "(@[<hv1>cc.push-pending@ %a@])" Class.pp t);
Vec.push cc.pending t Vec.push cc.pending t
let merge_classes cc t u e : unit = let merge_classes cc t u e : unit =
if t != u && not (same_class t u) then ( if t != u && not (same_class t u) then (
Log.debugf 50 (fun k -> Log.debugf 50 (fun k ->
k "(@[<hv1>cc.push-combine@ %a ~@ %a@ :expl %a@])" N.pp t N.pp u k "(@[<hv1>cc.push-combine@ %a ~@ %a@ :expl %a@])" Class.pp t Class.pp
Expl.pp e); u Expl.pp e);
Vec.push cc.combine @@ CT_merge (t, u, e) Vec.push cc.combine @@ CT_merge (t, u, e)
) )
@ -477,10 +478,11 @@ module Make (A : CC_ARG) :
Vec.clear cc.combine; Vec.clear cc.combine;
List.iter (fun f -> f cc ~th e) cc.on_conflict; List.iter (fun f -> f cc ~th e) cc.on_conflict;
Stat.incr cc.count_conflict; Stat.incr cc.count_conflict;
Actions.raise_conflict acts e p let (module A) = acts in
A.raise_conflict e p
let[@inline] all_classes cc : repr Iter.t = let[@inline] all_classes cc : repr Iter.t =
T_tbl.values cc.tbl |> Iter.filter N.is_root T_tbl.values cc.tbl |> Iter.filter Class.is_root
(* find the closest common ancestor of [a] and [b] in the proof forest. (* find the closest common ancestor of [a] and [b] in the proof forest.
@ -494,7 +496,7 @@ module Make (A : CC_ARG) :
let find_common_ancestor cc (a : node) (b : node) : node = let find_common_ancestor cc (a : node) (b : node) : node =
(* catch up to the other node *) (* catch up to the other node *)
let rec find1 a = let rec find1 a =
if N.get_field cc.field_marked_explain a then if Class.get_field cc.field_marked_explain a then
a a
else ( else (
match a.n_expl with match a.n_expl with
@ -503,15 +505,15 @@ module Make (A : CC_ARG) :
) )
in in
let rec find2 a b = let rec find2 a b =
if N.equal a b then if Class.equal a b then
a a
else if N.get_field cc.field_marked_explain a then else if Class.get_field cc.field_marked_explain a then
a a
else if N.get_field cc.field_marked_explain b then else if Class.get_field cc.field_marked_explain b then
b b
else ( else (
N.set_field cc.field_marked_explain true a; Class.set_field cc.field_marked_explain true a;
N.set_field cc.field_marked_explain true b; Class.set_field cc.field_marked_explain true b;
match a.n_expl, b.n_expl with match a.n_expl, b.n_expl with
| FL_some r1, FL_some r2 -> find2 r1.next r2.next | FL_some r1, FL_some r2 -> find2 r1.next r2.next
| FL_some r, FL_none -> find1 r.next | FL_some r, FL_none -> find1 r.next
@ -523,8 +525,8 @@ module Make (A : CC_ARG) :
(* cleanup tags on nodes traversed in [find2] *) (* cleanup tags on nodes traversed in [find2] *)
let rec cleanup_ n = let rec cleanup_ n =
if N.get_field cc.field_marked_explain n then ( if Class.get_field cc.field_marked_explain n then (
N.set_field cc.field_marked_explain false n; Class.set_field cc.field_marked_explain false n;
match n.n_expl with match n.n_expl with
| FL_none -> () | FL_none -> ()
| FL_some { next; _ } -> cleanup_ next | FL_some { next; _ } -> cleanup_ next
@ -538,7 +540,7 @@ module Make (A : CC_ARG) :
module Expl_state = struct module Expl_state = struct
type t = { type t = {
mutable lits: Lit.t list; mutable lits: Lit.t list;
mutable same_val: (N.t * N.t) list; mutable same_val: (Class.t * Class.t) list;
mutable th_lemmas: (Lit.t * (Lit.t * Lit.t list) list * proof_step) list; mutable th_lemmas: (Lit.t * (Lit.t * Lit.t list) list * proof_step) list;
} }
@ -572,7 +574,9 @@ module Make (A : CC_ARG) :
Iter.of_list self.th_lemmas Iter.of_list self.th_lemmas
|> Iter.map (fun (lit_t_u, _, _) -> Lit.neg lit_t_u) |> Iter.map (fun (lit_t_u, _, _) -> Lit.neg lit_t_u)
in in
let p_cc = P.lemma_cc (Iter.append p_lits1 p_lits2) proof in let p_cc =
P.add_step proof @@ Rules_.lemma_cc (Iter.append p_lits1 p_lits2)
in
let resolve_with_th_proof pr (lit_t_u, sub_proofs, pr_th) = let resolve_with_th_proof pr (lit_t_u, sub_proofs, pr_th) =
(* pr_th: [sub_proofs |- t=u]. (* pr_th: [sub_proofs |- t=u].
now resolve away [sub_proofs] to get literals that were now resolve away [sub_proofs] to get literals that were
@ -582,15 +586,16 @@ module Make (A : CC_ARG) :
(fun pr_th (lit_i, hyps_i) -> (fun pr_th (lit_i, hyps_i) ->
(* [hyps_i |- lit_i] *) (* [hyps_i |- lit_i] *)
let lemma_i = let lemma_i =
P.lemma_cc P.add_step proof
Iter.(cons lit_i (of_list hyps_i |> map Lit.neg)) @@ Rules_.lemma_cc
proof Iter.(cons lit_i (of_list hyps_i |> map Lit.neg))
in in
(* resolve [lit_i] away. *) (* resolve [lit_i] away. *)
P.proof_res ~pivot:(Lit.term lit_i) lemma_i pr_th proof) P.add_step proof
@@ Rules_.proof_res ~pivot:(Lit.term lit_i) lemma_i pr_th)
pr_th sub_proofs pr_th sub_proofs
in in
P.proof_res ~pivot:(Lit.term lit_t_u) pr_th pr proof P.add_step proof @@ Rules_.proof_res ~pivot:(Lit.term lit_t_u) pr_th pr
in in
(* resolve with theory proofs responsible for some merges, if any. *) (* resolve with theory proofs responsible for some merges, if any. *)
List.fold_left resolve_with_th_proof p_cc self.th_lemmas List.fold_left resolve_with_th_proof p_cc self.th_lemmas
@ -629,14 +634,14 @@ module Make (A : CC_ARG) :
let sub_proofs = let sub_proofs =
List.map List.map
(fun (t_i, u_i, expls_i) -> (fun (t_i, u_i, expls_i) ->
let lit_i = A.mk_lit_eq cc.tst t_i u_i in let lit_i = A.CC.mk_lit_eq cc.tst t_i u_i in
(* use a separate call to [explain_expls] for each set *) (* use a separate call to [explain_expls] for each set *)
let sub = explain_expls cc expls_i in let sub = explain_expls cc expls_i in
Expl_state.merge st sub; Expl_state.merge st sub;
lit_i, sub.lits) lit_i, sub.lits)
expl_sets expl_sets
in in
let lit_t_u = A.mk_lit_eq cc.tst t u in let lit_t_u = A.CC.mk_lit_eq cc.tst t u in
Expl_state.add_th st lit_t_u sub_proofs pr Expl_state.add_th st lit_t_u sub_proofs pr
| E_merge (a, b) -> explain_equal_rec_ cc st a b | E_merge (a, b) -> explain_equal_rec_ cc st a b
| E_merge_t (a, b) -> | E_merge_t (a, b) ->
@ -657,8 +662,8 @@ module Make (A : CC_ARG) :
and explain_equal_rec_ (cc : t) (st : Expl_state.t) (a : node) (b : node) : and explain_equal_rec_ (cc : t) (st : Expl_state.t) (a : node) (b : node) :
unit = unit =
Log.debugf 5 (fun k -> Log.debugf 5 (fun k ->
k "(@[cc.explain_loop.at@ %a@ =?= %a@])" N.pp a N.pp b); k "(@[cc.explain_loop.at@ %a@ =?= %a@])" Class.pp a Class.pp b);
assert (N.equal (find_ a) (find_ b)); assert (Class.equal (find_ a) (find_ b));
let ancestor = find_common_ancestor cc a b in let ancestor = find_common_ancestor cc a b in
explain_along_path cc st a ancestor; explain_along_path cc st a ancestor;
explain_along_path cc st b ancestor explain_along_path cc st b ancestor
@ -689,7 +694,7 @@ module Make (A : CC_ARG) :
and add_new_term_ cc (t : term) : node = and add_new_term_ cc (t : term) : node =
assert (not @@ mem cc t); assert (not @@ mem cc t);
Log.debugf 15 (fun k -> k "(@[cc.add-term@ %a@])" Term.pp t); Log.debugf 15 (fun k -> k "(@[cc.add-term@ %a@])" Term.pp t);
let n = N.make t in let n = Class.make t in
(* register sub-terms, add [t] to their parent list, and return the (* register sub-terms, add [t] to their parent list, and return the
corresponding initial signature *) corresponding initial signature *)
let sig0 = compute_sig0 cc n in let sig0 = compute_sig0 cc n in
@ -724,7 +729,7 @@ module Make (A : CC_ARG) :
sub sub
in in
let[@inline] return x = Some x in let[@inline] return x = Some x in
match A.cc_view n.n_term with match A.CC.view n.n_term with
| Bool _ | Opaque _ -> None | Bool _ | Opaque _ -> None
| Eq (a, b) -> | Eq (a, b) ->
let a = deref_sub a in let a = deref_sub a in
@ -750,13 +755,14 @@ module Make (A : CC_ARG) :
match n.n_as_lit with match n.n_as_lit with
| Some _ -> () | Some _ -> ()
| None -> | None ->
Log.debugf 15 (fun k -> k "(@[cc.set-as-lit@ %a@ %a@])" N.pp n Lit.pp lit); Log.debugf 15 (fun k ->
k "(@[cc.set-as-lit@ %a@ %a@])" Class.pp n Lit.pp lit);
on_backtrack cc (fun () -> n.n_as_lit <- None); on_backtrack cc (fun () -> n.n_as_lit <- None);
n.n_as_lit <- Some lit n.n_as_lit <- Some lit
(* is [n] true or false? *) (* is [n] true or false? *)
let n_is_bool_value (self : t) n : bool = let n_is_bool_value (self : t) n : bool =
N.equal n (n_true self) || N.equal n (n_false self) Class.equal n (n_true self) || Class.equal n (n_false self)
(* gather a pair [lits, pr], where [lits] is the set of (* gather a pair [lits, pr], where [lits] is the set of
asserted literals needed in the explanation (which is useful for asserted literals needed in the explanation (which is useful for
@ -790,16 +796,17 @@ module Make (A : CC_ARG) :
if same_class a b then ( if same_class a b then (
let expl = Expl.mk_merge a b in let expl = Expl.mk_merge a b in
Log.debugf 5 (fun k -> Log.debugf 5 (fun k ->
k "(@[cc.pending.eq@ %a@ :r1 %a@ :r2 %a@])" N.pp n N.pp a N.pp b); k "(@[cc.pending.eq@ %a@ :r1 %a@ :r2 %a@])" Class.pp n Class.pp a
Class.pp b);
merge_classes cc n (n_true cc) expl merge_classes cc n (n_true cc) expl
) )
| Some (Not u) -> | Some (Not u) ->
(* [u = bool ==> not u = not bool] *) (* [u = bool ==> not u = not bool] *)
let r_u = find_ u in let r_u = find_ u in
if N.equal r_u (n_true cc) then ( if Class.equal r_u (n_true cc) then (
let expl = Expl.mk_merge u (n_true cc) in let expl = Expl.mk_merge u (n_true cc) in
merge_classes cc n (n_false cc) expl merge_classes cc n (n_false cc) expl
) else if N.equal r_u (n_false cc) then ( ) else if Class.equal r_u (n_false cc) then (
let expl = Expl.mk_merge u (n_false cc) in let expl = Expl.mk_merge u (n_false cc) in
merge_classes cc n (n_true cc) expl merge_classes cc n (n_true cc) expl
) )
@ -810,7 +817,7 @@ module Make (A : CC_ARG) :
| None -> | None ->
(* add to the signature table [sig(n) --> n] *) (* add to the signature table [sig(n) --> n] *)
add_signature cc s n add_signature cc s n
| Some u when N.equal n u -> () | Some u when Class.equal n u -> ()
| Some u -> | Some u ->
(* [t1] and [t2] must be applications of the same symbol to (* [t1] and [t2] must be applications of the same symbol to
arguments that are pairwise equal *) arguments that are pairwise equal *)
@ -841,10 +848,11 @@ module Make (A : CC_ARG) :
k k
"(@[cc.semantic-conflict.set-val@ (@[set-val %a@ := %a@])@ \ "(@[cc.semantic-conflict.set-val@ (@[set-val %a@ := %a@])@ \
(@[existing-val %a@ := %a@])@])" (@[existing-val %a@ := %a@])@])"
N.pp n Term.pp v N.pp n' Term.pp v'); Class.pp n Term.pp v Class.pp n' Term.pp v');
Stat.incr cc.count_semantic_conflict; Stat.incr cc.count_semantic_conflict;
Actions.raise_semantic_conflict acts lits tuples let (module A) = acts in
A.raise_semantic_conflict lits tuples
| Some _ -> () | Some _ -> ()
| None -> T_b_tbl.add cc.t_to_val repr_n.n_term (n, v)); | None -> T_b_tbl.add cc.t_to_val repr_n.n_term (n, v));
(* now for the reverse map, look in self.val_to_t for [v]. (* now for the reverse map, look in self.val_to_t for [v].
@ -861,20 +869,20 @@ module Make (A : CC_ARG) :
and task_merge_ cc acts a b e_ab : unit = and task_merge_ cc acts a b e_ab : unit =
let ra = find_ a in let ra = find_ a in
let rb = find_ b in let rb = find_ b in
if not @@ N.equal ra rb then ( if not @@ Class.equal ra rb then (
assert (N.is_root ra); assert (Class.is_root ra);
assert (N.is_root rb); assert (Class.is_root rb);
Stat.incr cc.count_merge; Stat.incr cc.count_merge;
(* check we're not merging [true] and [false] *) (* check we're not merging [true] and [false] *)
if if
(N.equal ra (n_true cc) && N.equal rb (n_false cc)) (Class.equal ra (n_true cc) && Class.equal rb (n_false cc))
|| (N.equal rb (n_true cc) && N.equal ra (n_false cc)) || (Class.equal rb (n_true cc) && Class.equal ra (n_false cc))
then ( then (
Log.debugf 5 (fun k -> Log.debugf 5 (fun k ->
k k
"(@[<hv>cc.merge.true_false_conflict@ @[:r1 %a@ :t1 %a@]@ @[:r2 \ "(@[<hv>cc.merge.true_false_conflict@ @[:r1 %a@ :t1 %a@]@ @[:r2 \
%a@ :t2 %a@]@ :e_ab %a@])" %a@ :t2 %a@]@ :e_ab %a@])"
N.pp ra N.pp a N.pp rb N.pp b Expl.pp e_ab); Class.pp ra Class.pp a Class.pp rb Class.pp b Expl.pp e_ab);
let th = ref false in let th = ref false in
(* TODO: (* TODO:
C1: P.true_neq_false C1: P.true_neq_false
@ -891,11 +899,12 @@ module Make (A : CC_ARG) :
let lits = expl_st.lits in let lits = expl_st.lits in
let same_val = let same_val =
expl_st.same_val expl_st.same_val
|> List.rev_map (fun (t, u) -> true, N.term t, N.term u) |> List.rev_map (fun (t, u) -> true, Class.term t, Class.term u)
in in
assert (same_val <> []); assert (same_val <> []);
Stat.incr cc.count_semantic_conflict; Stat.incr cc.count_semantic_conflict;
Actions.raise_semantic_conflict acts lits same_val let (module A) = acts in
A.raise_semantic_conflict lits same_val
) else ( ) else (
(* regular conflict *) (* regular conflict *)
let lits, pr = lits_and_proof_of_expl cc expl_st in let lits, pr = lits_and_proof_of_expl cc expl_st in
@ -917,9 +926,9 @@ module Make (A : CC_ARG) :
in in
(* when merging terms with [true] or [false], possibly propagate them to SAT *) (* when merging terms with [true] or [false], possibly propagate them to SAT *)
let merge_bool r1 t1 r2 t2 = let merge_bool r1 t1 r2 t2 =
if N.equal r1 (n_true cc) then if Class.equal r1 (n_true cc) then
propagate_bools cc acts r2 t2 r1 t1 e_ab true propagate_bools cc acts r2 t2 r1 t1 e_ab true
else if N.equal r1 (n_false cc) then else if Class.equal r1 (n_false cc) then
propagate_bools cc acts r2 t2 r1 t1 e_ab false propagate_bools cc acts r2 t2 r1 t1 e_ab false
in in
@ -930,7 +939,7 @@ module Make (A : CC_ARG) :
(* perform [union r_from r_into] *) (* perform [union r_from r_into] *)
Log.debugf 15 (fun k -> Log.debugf 15 (fun k ->
k "(@[cc.merge@ :from %a@ :into %a@])" N.pp r_from N.pp r_into); k "(@[cc.merge@ :from %a@ :into %a@])" Class.pp r_from Class.pp r_into);
(* call [on_pre_merge] functions, and merge theory data items *) (* call [on_pre_merge] functions, and merge theory data items *)
if not cc.model_mode then ( if not cc.model_mode then (
@ -942,9 +951,9 @@ module Make (A : CC_ARG) :
); );
((* parents might have a different signature, check for collisions *) ((* parents might have a different signature, check for collisions *)
N.iter_parents r_from (fun parent -> push_pending cc parent); Class.iter_parents r_from (fun parent -> push_pending cc parent);
(* for each node in [r_from]'s class, make it point to [r_into] *) (* for each node in [r_from]'s class, make it point to [r_into] *)
N.iter_class r_from (fun u -> Class.iter_class r_from (fun u ->
assert (u.n_root == r_from); assert (u.n_root == r_from);
u.n_root <- r_into); u.n_root <- r_into);
(* capture current state *) (* capture current state *)
@ -961,15 +970,15 @@ module Make (A : CC_ARG) :
(* on backtrack, unmerge classes and restore the pointers to [r_from] *) (* on backtrack, unmerge classes and restore the pointers to [r_from] *)
on_backtrack cc (fun () -> on_backtrack cc (fun () ->
Log.debugf 30 (fun k -> Log.debugf 30 (fun k ->
k "(@[cc.undo_merge@ :from %a@ :into %a@])" N.pp r_from N.pp k "(@[cc.undo_merge@ :from %a@ :into %a@])" Class.pp r_from
r_into); Class.pp r_into);
r_into.n_bits <- r_into_old_bits; r_into.n_bits <- r_into_old_bits;
r_into.n_next <- r_into_old_next; r_into.n_next <- r_into_old_next;
r_from.n_next <- r_from_old_next; r_from.n_next <- r_from_old_next;
r_into.n_parents <- r_into_old_parents; r_into.n_parents <- r_into_old_parents;
(* NOTE: this must come after the restoration of [next] pointers, (* NOTE: this must come after the restoration of [next] pointers,
otherwise we'd iterate on too big a class *) otherwise we'd iterate on too big a class *)
N.iter_class_ r_from (fun u -> u.n_root <- r_from); Class.iter_class_ r_from (fun u -> u.n_root <- r_from);
r_into.n_size <- r_into.n_size - r_from.n_size)); r_into.n_size <- r_into.n_size - r_from.n_size));
(* check for semantic values, update the one of [r_into] (* check for semantic values, update the one of [r_into]
@ -997,10 +1006,11 @@ module Make (A : CC_ARG) :
k k
"(@[cc.semantic-conflict.post-merge@ (@[n-from %a@ := %a@])@ \ "(@[cc.semantic-conflict.post-merge@ (@[n-from %a@ := %a@])@ \
(@[n-into %a@ := %a@])@])" (@[n-into %a@ := %a@])@])"
N.pp n_from Term.pp v_from N.pp n_into Term.pp v_into); Class.pp n_from Term.pp v_from Class.pp n_into Term.pp v_into);
Stat.incr cc.count_semantic_conflict; Stat.incr cc.count_semantic_conflict;
Actions.raise_semantic_conflict acts lits tuples let (module A) = acts in
A.raise_semantic_conflict lits tuples
| Some _ -> ())); | Some _ -> ()));
(* update explanations (a -> b), arbitrarily. (* update explanations (a -> b), arbitrarily.
@ -1012,8 +1022,8 @@ module Make (A : CC_ARG) :
that bridges between [a] and [b] *) that bridges between [a] and [b] *)
on_backtrack cc (fun () -> on_backtrack cc (fun () ->
match a.n_expl, b.n_expl with match a.n_expl, b.n_expl with
| FL_some e, _ when N.equal e.next b -> a.n_expl <- FL_none | FL_some e, _ when Class.equal e.next b -> a.n_expl <- FL_none
| _, FL_some e when N.equal e.next a -> b.n_expl <- FL_none | _, FL_some e when Class.equal e.next a -> b.n_expl <- FL_none
| _ -> assert false); | _ -> assert false);
a.n_expl <- FL_some { next = b; expl = e_ab }; a.n_expl <- FL_some { next = b; expl = e_ab };
(* call [on_post_merge] *) (* call [on_post_merge] *)
@ -1036,14 +1046,14 @@ module Make (A : CC_ARG) :
in in
(* TODO: flag per class, `or`-ed on merge, to indicate if the class (* TODO: flag per class, `or`-ed on merge, to indicate if the class
contains at least one lit *) contains at least one lit *)
N.iter_class r1 (fun u1 -> Class.iter_class r1 (fun u1 ->
(* propagate if: (* propagate if:
- [u1] is a proper literal - [u1] is a proper literal
- [t2 != r2], because that can only happen - [t2 != r2], because that can only happen
after an explicit merge (no way to obtain that by propagation) after an explicit merge (no way to obtain that by propagation)
*) *)
match N.as_lit u1 with match Class.as_lit u1 with
| Some lit when not (N.equal r2 t2) -> | Some lit when not (Class.equal r2 t2) ->
let lit = let lit =
if sign then if sign then
lit lit
@ -1070,7 +1080,8 @@ module Make (A : CC_ARG) :
in in
List.iter (fun f -> f cc lit reason) cc.on_propagate; List.iter (fun f -> f cc lit reason) cc.on_propagate;
Stat.incr cc.count_props; Stat.incr cc.count_props;
Actions.propagate acts lit ~reason let (module A) = acts in
A.propagate lit ~reason
) )
| _ -> ()) | _ -> ())
@ -1078,9 +1089,7 @@ module Make (A : CC_ARG) :
let pp out _ = Fmt.string out "cc" let pp out _ = Fmt.string out "cc"
end end
let add_seq cc seq = let add_iter cc it : unit = it (fun t -> ignore @@ add_term_rec_ cc t)
seq (fun t -> ignore @@ add_term_rec_ cc t);
()
let[@inline] push_level (self : t) : unit = let[@inline] push_level (self : t) : unit =
Backtrack_stack.push_level self.undo; Backtrack_stack.push_level self.undo;
@ -1112,7 +1121,7 @@ module Make (A : CC_ARG) :
all_classes self all_classes self
|> Iter.filter_map (fun repr -> |> Iter.filter_map (fun repr ->
match T_b_tbl.get self.t_to_val repr.n_term with match T_b_tbl.get self.t_to_val repr.n_term with
| Some (_, v) -> Some (repr, N.iter_class repr, v) | Some (_, v) -> Some (repr, Class.iter_class repr, v)
| None -> None) | None -> None)
(* assert that this boolean literal holds. (* assert that this boolean literal holds.
@ -1122,7 +1131,7 @@ module Make (A : CC_ARG) :
let t = Lit.term lit in let t = Lit.term lit in
Log.debugf 15 (fun k -> k "(@[cc.assert-lit@ %a@])" Lit.pp lit); Log.debugf 15 (fun k -> k "(@[cc.assert-lit@ %a@])" Lit.pp lit);
let sign = Lit.sign lit in let sign = Lit.sign lit in
match A.cc_view t with match A.CC.view t with
| Eq (a, b) when sign -> | Eq (a, b) when sign ->
let a = add_term cc a in let a = add_term cc a in
let b = add_term cc b in let b = add_term cc b in
@ -1159,8 +1168,8 @@ module Make (A : CC_ARG) :
let merge cc n1 n2 expl = let merge cc n1 n2 expl =
Log.debugf 5 (fun k -> Log.debugf 5 (fun k ->
k "(@[cc.theory.merge@ :n1 %a@ :n2 %a@ :expl %a@])" N.pp n1 N.pp n2 k "(@[cc.theory.merge@ :n1 %a@ :n2 %a@ :expl %a@])" Class.pp n1 Class.pp
Expl.pp expl); n2 Expl.pp expl);
assert (T.Ty.equal (T.Term.ty n1.n_term) (T.Term.ty n2.n_term)); assert (T.Ty.equal (T.Term.ty n1.n_term) (T.Term.ty n2.n_term));
merge_classes cc n1 n2 expl merge_classes cc n1 n2 expl
@ -1245,7 +1254,8 @@ module Make (A : CC_ARG) :
let check_inv_ (self : t) : unit = let check_inv_ (self : t) : unit =
if check_inv_enabled_ then ( if check_inv_enabled_ then (
Log.debug 2 "(cc.check-invariants)"; Log.debug 2 "(cc.check-invariants)";
all_classes self |> Iter.flat_map N.iter_class all_classes self
|> Iter.flat_map Class.iter_class
|> Iter.iter (fun n -> |> Iter.iter (fun n ->
match n.n_sig0 with match n.n_sig0 with
| None -> () | None -> ()
@ -1254,16 +1264,143 @@ module Make (A : CC_ARG) :
let ok = let ok =
match find_signature self s' with match find_signature self s' with
| None -> false | None -> false
| Some r -> N.equal r n.n_root | Some r -> Class.equal r n.n_root
in in
if not ok then if not ok then
Log.debugf 0 (fun k -> Log.debugf 0 (fun k ->
k "(@[cc.check.fail@ :n %a@ :sig %a@ :actual-sig %a@])" k "(@[cc.check.fail@ :n %a@ :sig %a@ :actual-sig %a@])"
N.pp n Signature.pp s Signature.pp s')) Class.pp n Signature.pp s Signature.pp s'))
) )
(* model: return all the classes *) (* model: return all the classes *)
let get_model (cc : t) : repr Iter.t Iter.t = let get_model (cc : t) : repr Iter.t Iter.t =
check_inv_ cc; check_inv_ cc;
all_classes cc |> Iter.map N.iter_class all_classes cc |> Iter.map Class.iter_class
end
module Make_plugin (M : MONOID_ARG) : PLUGIN_BUILDER with module M = M = struct
module M = M
module CC = M.CC
module Class = CC.Class
module N_tbl = Backtrackable_tbl.Make (Class)
module Expl = CC.Expl
type term = CC.term
module type PL = PLUGIN with module CC = M.CC and module M = M
type plugin = (module PL)
module Make (A : sig
val size : int option
val cc : CC.t
end) : PL = struct
module M = M
module CC = CC
open A
(* repr -> value for the class *)
let values : M.t N_tbl.t = N_tbl.create ?size ()
(* bit in CC to filter out quickly classes without value *)
let field_has_value : Class.bitfield =
CC.allocate_bitfield ~descr:("monoid." ^ M.name ^ ".has-value") cc
let push_level () = N_tbl.push_level values
let pop_levels n = N_tbl.pop_levels values n
let n_levels () = N_tbl.n_levels values
let mem n =
let res = CC.get_bitfield cc field_has_value n in
assert (
if res then
N_tbl.mem values n
else
true);
res
let get n =
if CC.get_bitfield cc field_has_value n then
N_tbl.get values n
else
None
let on_new_term cc n (t : term) : unit =
(*Log.debugf 50 (fun k->k "(@[monoid[%s].on-new-term.try@ %a@])" M.name N.pp n);*)
let maybe_m, l = M.of_term cc n t in
(match maybe_m with
| Some v ->
Log.debugf 20 (fun k ->
k "(@[monoid[%s].on-new-term@ :n %a@ :value %a@])" M.name Class.pp n
M.pp v);
CC.set_bitfield cc field_has_value true n;
N_tbl.add values n v
| None -> ());
List.iter
(fun (n_u, m_u) ->
Log.debugf 20 (fun k ->
k "(@[monoid[%s].on-new-term.sub@ :n %a@ :sub-t %a@ :value %a@])"
M.name Class.pp n Class.pp n_u M.pp m_u);
let n_u = CC.find cc n_u in
if CC.get_bitfield cc field_has_value n_u then (
let m_u' =
try N_tbl.find values n_u
with Not_found ->
Error.errorf "node %a has bitfield but no value" Class.pp n_u
in
match M.merge cc n_u m_u n_u m_u' (Expl.mk_list []) with
| Error expl ->
Error.errorf
"when merging@ @[for node %a@],@ values %a and %a:@ conflict %a"
Class.pp n_u M.pp m_u M.pp m_u' CC.Expl.pp expl
| Ok m_u_merged ->
Log.debugf 20 (fun k ->
k
"(@[monoid[%s].on-new-term.sub.merged@ :n %a@ :sub-t %a@ \
:value %a@])"
M.name Class.pp n Class.pp n_u M.pp m_u_merged);
N_tbl.add values n_u m_u_merged
) else (
(* just add to [n_u] *)
CC.set_bitfield cc field_has_value true n_u;
N_tbl.add values n_u m_u
))
l;
()
let iter_all : _ Iter.t = N_tbl.to_iter values
let on_pre_merge cc acts n1 n2 e_n1_n2 : unit =
match get n1, get n2 with
| Some v1, Some v2 ->
Log.debugf 5 (fun k ->
k
"(@[monoid[%s].on_pre_merge@ (@[:n1 %a@ :val1 %a@])@ (@[:n2 %a@ \
:val2 %a@])@])"
M.name Class.pp n1 M.pp v1 Class.pp n2 M.pp v2);
(match M.merge cc n1 v1 n2 v2 e_n1_n2 with
| Ok v' ->
N_tbl.remove values n2;
(* only keep repr *)
N_tbl.add values n1 v'
| Error expl -> CC.raise_conflict_from_expl cc acts expl)
| None, Some cr ->
CC.set_bitfield cc field_has_value true n1;
N_tbl.add values n1 cr;
N_tbl.remove values n2 (* only keep reprs *)
| Some _, None -> () (* already there on the left *)
| None, None -> ()
(* setup *)
let () =
CC.on_new_term cc on_new_term;
CC.on_pre_merge cc on_pre_merge;
()
end
let create_and_setup ?size (cc : CC.t) : plugin =
(module Make (struct
let size = size
let cc = cc
end))
end end

View file

@ -1,13 +1,20 @@
(** {2 Congruence Closure} *) (** Congruence Closure Implementation *)
open Sidekick_core module View = Sidekick_sigs_cc.View
module type S = Sidekick_core.CC_S module type TERM = Sidekick_sigs_cc.TERM
module type LIT = Sidekick_sigs_cc.LIT
module type ARG = Sidekick_sigs_cc.ARG
module type S = Sidekick_sigs_cc.S
module type MONOID_ARG = Sidekick_sigs_cc.MONOID_ARG
module type PLUGIN = Sidekick_sigs_cc.PLUGIN
module type PLUGIN_BUILDER = Sidekick_sigs_cc.PLUGIN_BUILDER
module Make (A : CC_ARG) : module Make (A : ARG) :
S S
with module T = A.T with module T = A.T
and module Lit = A.Lit and module Lit = A.Lit
and type proof = A.proof and module Proof_trace = A.Proof_trace
and type proof_step = A.proof_step
and module Actions = A.Actions (** Create a plugin builder from the given per-class monoid *)
module Make_plugin (M : MONOID_ARG) : PLUGIN_BUILDER with module M = M

View file

@ -1,5 +1,5 @@
(library (library
(name Sidekick_cc) (name Sidekick_cc)
(public_name sidekick.cc) (public_name sidekick.cc)
(libraries containers iter sidekick.core sidekick.util) (libraries containers iter sidekick.sigs sidekick.sigs.cc sidekick.util)
(flags :standard -warn-error -a+8 -w -32 -open Sidekick_util)) (flags :standard -warn-error -a+8 -w -32 -open Sidekick_util))

File diff suppressed because it is too large Load diff

View file

@ -2,4 +2,6 @@
(name Sidekick_core) (name Sidekick_core)
(public_name sidekick.core) (public_name sidekick.core)
(flags :standard -open Sidekick_util) (flags :standard -open Sidekick_util)
(libraries containers iter sidekick.util)) (libraries containers iter sidekick.util sidekick.sigs.proof-trace
sidekick.sigs.term sidekick.sigs.lit sidekick.sigs.proof.sat
sidekick.sigs.proof.core sidekick.sigs.cc))

View file

@ -1,7 +1,9 @@
module CC_view = Sidekick_core.CC_view module CC_view = Sidekick_sigs_cc.View
module type TERM = Sidekick_sigs_term.S
module type ARG = sig module type ARG = sig
module T : Sidekick_core.TERM module T : TERM
val cc_view : T.Term.t -> (T.Fun.t, T.Term.t, T.Term.t Iter.t) CC_view.t val cc_view : T.Term.t -> (T.Fun.t, T.Term.t, T.Term.t Iter.t) CC_view.t
end end

View file

@ -1,17 +1,19 @@
(** {1 Mini congruence closure} (** Mini congruence closure
This implementation is as simple as possible, and doesn't provide This implementation is as simple as possible, and doesn't provide
backtracking, theories, or explanations. backtracking, theories, or explanations.
It just decides the satisfiability of a set of (dis)equations. It just decides the satisfiability of a set of (dis)equations.
*) *)
module CC_view = Sidekick_core.CC_view module CC_view = Sidekick_sigs_cc.View
module type TERM = Sidekick_sigs_term.S
(** Argument for the functor {!Make} (** Argument for the functor {!Make}
It only requires a term structure, and a congruence-oriented view. *) It only requires a term structure, and a congruence-oriented view. *)
module type ARG = sig module type ARG = sig
module T : Sidekick_core.TERM module T : TERM
val cc_view : T.Term.t -> (T.Fun.t, T.Term.t, T.Term.t Iter.t) CC_view.t val cc_view : T.Term.t -> (T.Fun.t, T.Term.t, T.Term.t Iter.t) CC_view.t
end end

View file

@ -1,5 +1,5 @@
(library (library
(name Sidekick_mini_cc) (name Sidekick_mini_cc)
(public_name sidekick.mini-cc) (public_name sidekick.mini-cc)
(libraries containers iter sidekick.core sidekick.util) (libraries containers iter sidekick.sigs.cc sidekick.sigs.term sidekick.util)
(flags :standard -warn-error -a+8 -w -32 -open Sidekick_util)) (flags :standard -warn-error -a+8 -w -32 -open Sidekick_util))