mirror of
https://github.com/Z3Prover/z3
synced 2025-08-23 11:37:54 +00:00
adding dt-solver (#4739)
* adding dt-solver Signed-off-by: Nikolaj Bjorner <nbjorner@microsoft.com> * dt Signed-off-by: Nikolaj Bjorner <nbjorner@microsoft.com> * move mbp to self-contained module Signed-off-by: Nikolaj Bjorner <nbjorner@microsoft.com> * files Signed-off-by: Nikolaj Bjorner <nbjorner@microsoft.com> * Create CMakeLists.txt * dt Signed-off-by: Nikolaj Bjorner <nbjorner@microsoft.com> * rename to bool_var2expr to indicate type class * mbp * na
This commit is contained in:
parent
b77c57451f
commit
2f756da294
62 changed files with 2309 additions and 1257 deletions
|
@ -76,7 +76,7 @@ namespace sat {
|
|||
~scoped_drating() { ext.m_drating = false; }
|
||||
};
|
||||
virtual void init_search() {}
|
||||
virtual bool propagate(sat::literal l, sat::ext_constraint_idx idx) { UNREACHABLE(); return false; }
|
||||
virtual bool propagated(sat::literal l, sat::ext_constraint_idx idx) { UNREACHABLE(); return false; }
|
||||
virtual bool unit_propagate() = 0;
|
||||
virtual bool is_external(bool_var v) { return false; }
|
||||
virtual double get_reward(literal l, ext_constraint_idx idx, literal_occs_fun& occs) const { return 0; }
|
||||
|
|
|
@ -1354,7 +1354,7 @@ namespace sat {
|
|||
watch_list::iterator it = wlist.begin(), it2 = it, end = wlist.end();
|
||||
for (; it != end && !inconsistent(); ++it) {
|
||||
SASSERT(it->get_kind() == watched::EXT_CONSTRAINT);
|
||||
bool keep = m_s.m_ext->propagate(l, it->get_ext_constraint_idx());
|
||||
bool keep = m_s.m_ext->propagated(l, it->get_ext_constraint_idx());
|
||||
if (m_search_mode == lookahead_mode::lookahead1 && !m_inconsistent) {
|
||||
lookahead_literal_occs_fun literal_occs_fn(*this);
|
||||
m_lookahead_reward += m_s.m_ext->get_reward(l, it->get_ext_constraint_idx(), literal_occs_fn);
|
||||
|
|
|
@ -1180,7 +1180,7 @@ namespace sat {
|
|||
}
|
||||
case watched::EXT_CONSTRAINT:
|
||||
SASSERT(m_ext);
|
||||
keep = m_ext->propagate(l, it->get_ext_constraint_idx());
|
||||
keep = m_ext->propagated(l, it->get_ext_constraint_idx());
|
||||
if (m_inconsistent) {
|
||||
if (!keep) {
|
||||
++it;
|
||||
|
|
|
@ -20,6 +20,7 @@ z3_add_component(sat_smt
|
|||
bv_internalize.cpp
|
||||
bv_invariant.cpp
|
||||
bv_solver.cpp
|
||||
dt_solver.cpp
|
||||
euf_ackerman.cpp
|
||||
euf_internalize.cpp
|
||||
euf_invariant.cpp
|
||||
|
|
|
@ -53,12 +53,7 @@ namespace arith {
|
|||
}
|
||||
|
||||
std::ostream& solver::display_justification(std::ostream& out, sat::ext_justification_idx idx) const {
|
||||
auto& jst = euf::th_propagation::from_index(idx);
|
||||
for (auto lit : euf::th_propagation::lits(jst))
|
||||
out << lit << " ";
|
||||
for (auto eq : euf::th_propagation::eqs(jst))
|
||||
out << eq.first->get_expr_id() << " == " << eq.second->get_expr_id() << " ";
|
||||
return out;
|
||||
return euf::th_propagation::from_index(idx).display(out);
|
||||
}
|
||||
|
||||
std::ostream& solver::display_constraint(std::ostream& out, sat::ext_constraint_idx idx) const {
|
||||
|
|
|
@ -61,14 +61,10 @@ namespace arith {
|
|||
m_asserted.push_back(l);
|
||||
}
|
||||
|
||||
euf::th_solver* solver::clone(sat::solver* s, euf::solver& ctx) {
|
||||
arith::solver* result = alloc(arith::solver, ctx, get_id());
|
||||
ast_translation tr(m, ctx.get_manager());
|
||||
for (unsigned i = result->get_num_vars(); i < get_num_vars(); ++i) {
|
||||
expr* e1 = var2expr(i);
|
||||
expr* e2 = tr(e1);
|
||||
result->mk_evar(e2);
|
||||
}
|
||||
euf::th_solver* solver::clone(euf::solver& dst_ctx) {
|
||||
arith::solver* result = alloc(arith::solver, dst_ctx, get_id());
|
||||
for (unsigned i = result->get_num_vars(); i < get_num_vars(); ++i)
|
||||
result->mk_evar(ctx.copy(dst_ctx, var2enode(i))->get_expr());
|
||||
|
||||
unsigned v = 0;
|
||||
result->m_bounds.resize(m_bounds.size());
|
||||
|
@ -1425,19 +1421,7 @@ namespace arith {
|
|||
|
||||
void solver::get_antecedents(literal l, sat::ext_justification_idx idx, literal_vector& r, bool probing) {
|
||||
auto& jst = euf::th_propagation::from_index(idx);
|
||||
for (auto lit : euf::th_propagation::lits(jst))
|
||||
r.push_back(lit);
|
||||
for (auto eq : euf::th_propagation::eqs(jst))
|
||||
ctx.add_antecedent(eq.first, eq.second);
|
||||
|
||||
if (!probing && ctx.use_drat()) {
|
||||
literal_vector lits;
|
||||
for (auto lit : euf::th_propagation::lits(jst))
|
||||
lits.push_back(~lit);
|
||||
lits.push_back(l);
|
||||
ctx.get_drat().add(lits, status());
|
||||
for (auto eq : euf::th_propagation::eqs(jst))
|
||||
IF_VERBOSE(0, verbose_stream() << "drat-log with equalities is TBD " << eq.first->get_expr_id() << "\n");
|
||||
}
|
||||
ctx.get_antecedents(l, jst, r, probing);
|
||||
}
|
||||
|
||||
}
|
||||
|
|
|
@ -410,7 +410,6 @@ namespace arith {
|
|||
solver(euf::solver& ctx, theory_id id);
|
||||
~solver() override;
|
||||
bool is_external(bool_var v) override { return false; }
|
||||
bool propagate(literal l, sat::ext_constraint_idx idx) override { UNREACHABLE(); return false; }
|
||||
void get_antecedents(literal l, sat::ext_justification_idx idx, literal_vector& r, bool probing) override;
|
||||
void asserted(literal l) override;
|
||||
sat::check_result check() override;
|
||||
|
@ -419,7 +418,7 @@ namespace arith {
|
|||
std::ostream& display_justification(std::ostream& out, sat::ext_justification_idx idx) const override;
|
||||
std::ostream& display_constraint(std::ostream& out, sat::ext_constraint_idx idx) const override;
|
||||
void collect_statistics(statistics& st) const override;
|
||||
euf::th_solver* clone(sat::solver* s, euf::solver& ctx) override;
|
||||
euf::th_solver* clone(euf::solver& ctx) override;
|
||||
bool use_diseqs() const override { return true; }
|
||||
void new_eq_eh(euf::th_eq const& eq) override { mk_eq_axiom(eq.v1(), eq.v2()); }
|
||||
void new_diseq_eh(euf::th_eq const& de) override { mk_eq_axiom(de.v1(), de.v2()); }
|
||||
|
|
|
@ -26,7 +26,10 @@ namespace array {
|
|||
TRACE("array", tout << mk_pp(e, m) << "\n";);
|
||||
return sat::null_literal;
|
||||
}
|
||||
return expr2literal(e);
|
||||
auto lit = expr2literal(e);
|
||||
if (sign)
|
||||
lit.neg();
|
||||
return lit;
|
||||
}
|
||||
|
||||
void solver::internalize(expr* e, bool redundant) {
|
||||
|
|
|
@ -151,15 +151,10 @@ namespace array {
|
|||
st.update("array splits", m_stats.m_num_eq_splits);
|
||||
}
|
||||
|
||||
euf::th_solver* solver::clone(sat::solver* s, euf::solver& ctx) {
|
||||
auto* result = alloc(solver, ctx, get_id());
|
||||
ast_translation tr(m, ctx.get_manager());
|
||||
for (unsigned i = 0; i < get_num_vars(); ++i) {
|
||||
expr* e1 = var2expr(i);
|
||||
expr* e2 = tr(e1);
|
||||
euf::enode* n = ctx.get_enode(e2);
|
||||
result->mk_var(n);
|
||||
}
|
||||
euf::th_solver* solver::clone(euf::solver& dst_ctx) {
|
||||
auto* result = alloc(solver, dst_ctx, get_id());
|
||||
for (unsigned i = 0; i < get_num_vars(); ++i)
|
||||
result->mk_var(ctx.copy(dst_ctx, var2enode(i)));
|
||||
return result;
|
||||
}
|
||||
|
||||
|
|
|
@ -192,7 +192,6 @@ namespace array {
|
|||
solver(euf::solver& ctx, theory_id id);
|
||||
~solver() override;
|
||||
bool is_external(bool_var v) override { return false; }
|
||||
bool propagate(literal l, sat::ext_constraint_idx idx) override { UNREACHABLE(); return false; }
|
||||
void get_antecedents(literal l, sat::ext_justification_idx idx, literal_vector& r, bool probing) override {}
|
||||
void asserted(literal l) override {}
|
||||
sat::check_result check() override;
|
||||
|
@ -201,7 +200,7 @@ namespace array {
|
|||
std::ostream& display_justification(std::ostream& out, sat::ext_justification_idx idx) const override;
|
||||
std::ostream& display_constraint(std::ostream& out, sat::ext_constraint_idx idx) const override;
|
||||
void collect_statistics(statistics& st) const override;
|
||||
euf::th_solver* clone(sat::solver* s, euf::solver& ctx) override;
|
||||
euf::th_solver* clone(euf::solver& ctx) override;
|
||||
void new_eq_eh(euf::th_eq const& eq) override;
|
||||
bool use_diseqs() const override { return true; }
|
||||
void new_diseq_eh(euf::th_eq const& eq) override;
|
||||
|
@ -217,5 +216,7 @@ namespace array {
|
|||
void merge_eh(theory_var, theory_var, theory_var v1, theory_var v2);
|
||||
void after_merge_eh(theory_var r1, theory_var r2, theory_var v1, theory_var v2) {}
|
||||
void unmerge_eh(theory_var v1, theory_var v2) {}
|
||||
|
||||
euf::enode_vector const& parent_selects(euf::enode* n) const { return m_var_data[n->get_th_var(get_id())]->m_parent_selects; }
|
||||
};
|
||||
}
|
||||
|
|
|
@ -1484,7 +1484,7 @@ namespace sat {
|
|||
/*
|
||||
\brief return true to keep watching literal.
|
||||
*/
|
||||
bool ba_solver::propagate(literal l, ext_constraint_idx idx) {
|
||||
bool ba_solver::propagated(literal l, ext_constraint_idx idx) {
|
||||
SASSERT(value(l) == l_true);
|
||||
constraint& c = index2constraint(idx);
|
||||
if (c.lit() != null_literal && l.var() == c.lit().var()) {
|
||||
|
@ -3137,16 +3137,16 @@ namespace sat {
|
|||
}
|
||||
|
||||
extension* ba_solver::copy(solver* s) {
|
||||
return clone_aux(s, m, si, m_id);
|
||||
return clone_aux(m, *s, si, m_id);
|
||||
}
|
||||
|
||||
euf::th_solver* ba_solver::clone(solver* new_s, euf::solver& new_ctx) {
|
||||
return clone_aux(new_s, new_ctx.get_manager(), new_ctx.get_si(), get_id());
|
||||
euf::th_solver* ba_solver::clone(euf::solver& new_ctx) {
|
||||
return clone_aux(new_ctx.get_manager(), new_ctx.s(), new_ctx.get_si(), get_id());
|
||||
}
|
||||
|
||||
euf::th_solver* ba_solver::clone_aux(solver* new_s, ast_manager& m, sat::sat_internalizer& si, euf::theory_id id) {
|
||||
euf::th_solver* ba_solver::clone_aux(ast_manager& m, sat::solver& s, sat::sat_internalizer& si, euf::theory_id id) {
|
||||
ba_solver* result = alloc(ba_solver, m, si, id);
|
||||
result->set_solver(new_s);
|
||||
result->set_solver(&s);
|
||||
copy_constraints(result, m_constraints);
|
||||
return result;
|
||||
}
|
||||
|
|
|
@ -151,7 +151,7 @@ namespace sat {
|
|||
unsigned_vector m_weights;
|
||||
svector<wliteral> m_wlits;
|
||||
|
||||
euf::th_solver* clone_aux(sat::solver* new_s, ast_manager& m, sat::sat_internalizer& si, euf::theory_id id);
|
||||
euf::th_solver* clone_aux(ast_manager& m, sat::solver& s, sat::sat_internalizer& si, euf::theory_id id);
|
||||
|
||||
bool subsumes(card& c1, card& c2, literal_vector& comp);
|
||||
bool subsumes(card& c1, clause& c2, bool& self);
|
||||
|
@ -402,7 +402,7 @@ namespace sat {
|
|||
void add_xr(literal_vector const& lits);
|
||||
|
||||
bool is_external(bool_var v) override;
|
||||
bool propagate(literal l, ext_constraint_idx idx) override;
|
||||
bool propagated(literal l, ext_constraint_idx idx) override;
|
||||
bool unit_propagate() override { return false; }
|
||||
lbool resolve_conflict() override;
|
||||
void get_antecedents(literal l, ext_justification_idx idx, literal_vector & r, bool probing) override;
|
||||
|
@ -433,7 +433,7 @@ namespace sat {
|
|||
literal internalize(expr* e, bool sign, bool root, bool redundant) override;
|
||||
void internalize(expr* e, bool redundant) override;
|
||||
bool to_formulas(std::function<expr_ref(sat::literal)>& l2e, expr_ref_vector& fmls) override;
|
||||
euf::th_solver* clone(solver* s, euf::solver& ctx) override;
|
||||
euf::th_solver* clone(euf::solver& ctx) override;
|
||||
|
||||
ptr_vector<constraint> const & constraints() const { return m_constraints; }
|
||||
std::ostream& display(std::ostream& out, constraint const& c, bool values) const;
|
||||
|
|
|
@ -260,7 +260,6 @@ namespace bv {
|
|||
double solver::get_reward(literal l, sat::ext_constraint_idx idx, sat::literal_occs_fun& occs) const { return 0; }
|
||||
bool solver::is_extended_binary(sat::ext_justification_idx idx, literal_vector& r) { return false; }
|
||||
bool solver::is_external(bool_var v) { return true; }
|
||||
bool solver::propagate(literal l, sat::ext_constraint_idx idx) { return false; }
|
||||
|
||||
void solver::get_antecedents(literal l, sat::ext_justification_idx idx, literal_vector& r, bool probing) {
|
||||
auto& c = bv_justification::from_index(idx);
|
||||
|
@ -649,7 +648,7 @@ namespace bv {
|
|||
|
||||
sat::extension* solver::copy(sat::solver* s) { UNREACHABLE(); return nullptr; }
|
||||
|
||||
euf::th_solver* solver::clone(sat::solver* s, euf::solver& ctx) {
|
||||
euf::th_solver* solver::clone(euf::solver& ctx) {
|
||||
bv::solver* result = alloc(bv::solver, ctx, get_id());
|
||||
ast_translation tr(m, ctx.get_manager());
|
||||
for (unsigned i = 0; i < get_num_vars(); ++i) {
|
||||
|
|
|
@ -309,7 +309,6 @@ namespace bv {
|
|||
double get_reward(literal l, sat::ext_constraint_idx idx, sat::literal_occs_fun& occs) const override;
|
||||
bool is_extended_binary(sat::ext_justification_idx idx, literal_vector& r) override;
|
||||
bool is_external(bool_var v) override;
|
||||
bool propagate(literal l, sat::ext_constraint_idx idx) override;
|
||||
void get_antecedents(literal l, sat::ext_justification_idx idx, literal_vector & r, bool probing) override;
|
||||
void asserted(literal l) override;
|
||||
sat::check_result check() override;
|
||||
|
@ -324,7 +323,7 @@ namespace bv {
|
|||
std::ostream& display_justification(std::ostream& out, sat::ext_justification_idx idx) const override;
|
||||
std::ostream& display_constraint(std::ostream& out, sat::ext_constraint_idx idx) const override;
|
||||
void collect_statistics(statistics& st) const override;
|
||||
euf::th_solver* clone(sat::solver* s, euf::solver& ctx) override;
|
||||
euf::th_solver* clone(euf::solver& ctx) override;
|
||||
extension* copy(sat::solver* s) override;
|
||||
void find_mutexes(literal_vector& lits, vector<literal_vector> & mutexes) override {}
|
||||
void gc() override {}
|
||||
|
|
785
src/sat/smt/dt_solver.cpp
Normal file
785
src/sat/smt/dt_solver.cpp
Normal file
|
@ -0,0 +1,785 @@
|
|||
/*++
|
||||
Copyright (c) 2020 Microsoft Corporation
|
||||
|
||||
Module Name:
|
||||
|
||||
dt_solver.h
|
||||
|
||||
Abstract:
|
||||
|
||||
Theory plugin for altegraic datatypes
|
||||
|
||||
Author:
|
||||
|
||||
Nikolaj Bjorner (nbjorner) 2020-09-08
|
||||
|
||||
--*/
|
||||
|
||||
#include "sat/smt/dt_solver.h"
|
||||
#include "sat/smt/euf_solver.h"
|
||||
#include "sat/smt/array_solver.h"
|
||||
|
||||
namespace euf {
|
||||
class solver;
|
||||
}
|
||||
|
||||
namespace dt {
|
||||
|
||||
solver::solver(euf::solver& ctx, theory_id id) :
|
||||
th_euf_solver(ctx, m.get_family_name(id), id),
|
||||
dt(m),
|
||||
m_autil(m),
|
||||
m_find(*this),
|
||||
m_args(m)
|
||||
{}
|
||||
|
||||
solver::~solver() {
|
||||
std::for_each(m_var_data.begin(), m_var_data.end(), delete_proc<var_data>());
|
||||
m_var_data.reset();
|
||||
}
|
||||
|
||||
void solver::clone_var(solver& src, theory_var v) {
|
||||
enode* n = src.ctx.copy(ctx, src.var2enode(v));
|
||||
VERIFY(v == th_euf_solver::mk_var(n));
|
||||
m_var_data.push_back(alloc(var_data));
|
||||
var_data* d_dst = m_var_data[v];
|
||||
var_data* d_src = src.m_var_data[v];
|
||||
ctx.attach_th_var(n, this, v);
|
||||
if (d_src->m_constructor && !d_dst->m_constructor)
|
||||
d_dst->m_constructor = src.ctx.copy(ctx, d_src->m_constructor);
|
||||
for (auto* r : d_src->m_recognizers)
|
||||
d_dst->m_recognizers.push_back(src.ctx.copy(ctx, r));
|
||||
}
|
||||
|
||||
euf::th_solver* solver::clone(euf::solver& dst_ctx) {
|
||||
auto* result = alloc(solver, dst_ctx, get_id());
|
||||
for (unsigned v = 0; v < get_num_vars(); ++v)
|
||||
result->clone_var(*this, v);
|
||||
return result;
|
||||
}
|
||||
|
||||
solver::final_check_st::final_check_st(solver& s) : s(s) {
|
||||
SASSERT(s.m_to_unmark1.empty());
|
||||
SASSERT(s.m_to_unmark2.empty());
|
||||
s.m_used_eqs.reset();
|
||||
s.m_dfs.reset();
|
||||
}
|
||||
|
||||
solver::final_check_st::~final_check_st() {
|
||||
s.clear_mark();
|
||||
}
|
||||
|
||||
void solver::clear_mark() {
|
||||
for (enode* n : m_to_unmark1)
|
||||
n->unmark1();
|
||||
for (enode* n : m_to_unmark2)
|
||||
n->unmark2();
|
||||
m_to_unmark1.reset();
|
||||
m_to_unmark2.reset();
|
||||
}
|
||||
|
||||
void solver::oc_mark_on_stack(enode* n) {
|
||||
n = n->get_root();
|
||||
n->mark1();
|
||||
m_to_unmark1.push_back(n);
|
||||
}
|
||||
|
||||
void solver::oc_mark_cycle_free(enode* n) {
|
||||
n = n->get_root();
|
||||
n->mark2();
|
||||
m_to_unmark2.push_back(n);
|
||||
}
|
||||
|
||||
void solver::oc_push_stack(enode* n) {
|
||||
m_dfs.push_back(std::make_pair(EXIT, n));
|
||||
m_dfs.push_back(std::make_pair(ENTER, n));
|
||||
}
|
||||
|
||||
/**
|
||||
\brief Assert the axiom (antecedent => lhs = rhs)
|
||||
antecedent may be null_literal
|
||||
*/
|
||||
void solver::assert_eq_axiom(enode* lhs, expr* rhs, literal antecedent) {
|
||||
if (antecedent == sat::null_literal)
|
||||
add_unit(eq_internalize(lhs->get_expr(), rhs));
|
||||
else if (s().value(antecedent) == l_true) {
|
||||
euf::th_propagation* jst = euf::th_propagation::mk(*this, antecedent);
|
||||
ctx.propagate(lhs, e_internalize(rhs), jst);
|
||||
}
|
||||
else
|
||||
add_clause(~antecedent, eq_internalize(lhs->get_expr(), rhs));
|
||||
}
|
||||
|
||||
/**
|
||||
\brief Assert the equality (= n (c (acc_1 n) ... (acc_m n))) where
|
||||
where acc_i are the accessors of constructor c.
|
||||
*/
|
||||
void solver::assert_is_constructor_axiom(enode* n, func_decl* c, literal antecedent) {
|
||||
expr* e = n->get_expr();
|
||||
TRACE("dt", tout << "creating axiom (= n (c (acc_1 n) ... (acc_m n))) for\n"
|
||||
<< mk_pp(c, m) << " " << mk_pp(e, m) << "\n";);
|
||||
m_stats.m_assert_cnstr++;
|
||||
SASSERT(dt.is_constructor(c));
|
||||
SASSERT(is_datatype(e));
|
||||
SASSERT(c->get_range() == m.get_sort(e));
|
||||
m_args.reset();
|
||||
ptr_vector<func_decl> const& accessors = *dt.get_constructor_accessors(c);
|
||||
SASSERT(c->get_arity() == accessors.size());
|
||||
for (func_decl* d : accessors)
|
||||
m_args.push_back(m.mk_app(d, e));
|
||||
expr_ref con(m.mk_app(c, m_args), m);
|
||||
assert_eq_axiom(n, con, antecedent);
|
||||
}
|
||||
|
||||
/**
|
||||
\brief Given a constructor n := (c a_1 ... a_m) assert the axioms
|
||||
(= (acc_1 n) a_1)
|
||||
...
|
||||
(= (acc_m n) a_m)
|
||||
*/
|
||||
void solver::assert_accessor_axioms(enode* n) {
|
||||
m_stats.m_assert_accessor++;
|
||||
expr* e = n->get_expr();
|
||||
SASSERT(is_constructor(n));
|
||||
func_decl* d = n->get_decl();
|
||||
ptr_vector<func_decl> const& accessors = *dt.get_constructor_accessors(d);
|
||||
SASSERT(n->num_args() == accessors.size());
|
||||
unsigned i = 0;
|
||||
for (func_decl* acc : accessors) {
|
||||
app_ref acc_app(m.mk_app(acc, e), m);
|
||||
assert_eq_axiom(n->get_arg(i), acc_app);
|
||||
++i;
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
\brief Sign a conflict for r := is_mk(a), c := mk(...), not(r), and c == a.
|
||||
*/
|
||||
void solver::sign_recognizer_conflict(enode* c, enode* r) {
|
||||
SASSERT(is_constructor(c));
|
||||
SASSERT(is_recognizer(r));
|
||||
SASSERT(dt.get_recognizer_constructor(r->get_decl()) == c->get_decl());
|
||||
SASSERT(c->get_root() == r->get_arg(0)->get_root());
|
||||
TRACE("dt", tout << ctx.bpp(c) << "\n" << ctx.bpp(r) << "\n";);
|
||||
literal l = ctx.enode2literal(r);
|
||||
SASSERT(s().value(l) == l_false);
|
||||
clear_mark();
|
||||
auto* jst = euf::th_propagation::mk(*this, ~l, c, r->get_arg(0));
|
||||
ctx.set_conflict(jst);
|
||||
}
|
||||
|
||||
/**
|
||||
\brief Given a field update n := { r with field := v } for constructor C, assert the axioms:
|
||||
(=> (is-C r) (= (acc_j n) (acc_j r))) for acc_j != field
|
||||
(=> (is-C r) (= (field n) v)) for acc_j != field
|
||||
(=> (not (is-C r)) (= n r))
|
||||
(=> (is-C r) (is-C n))
|
||||
*/
|
||||
void solver::assert_update_field_axioms(enode* n) {
|
||||
m_stats.m_assert_update_field++;
|
||||
SASSERT(is_update_field(n));
|
||||
expr* own = n->get_expr();
|
||||
expr* arg1 = n->get_arg(0)->get_expr();
|
||||
func_decl* upd = n->get_decl();
|
||||
func_decl* acc = to_func_decl(upd->get_parameter(0).get_ast());
|
||||
func_decl* con = dt.get_accessor_constructor(acc);
|
||||
func_decl* rec = dt.get_constructor_is(con);
|
||||
ptr_vector<func_decl> const& accessors = *dt.get_constructor_accessors(con);
|
||||
app_ref rec_app(m.mk_app(rec, arg1), m);
|
||||
app_ref acc_app(m);
|
||||
literal is_con = b_internalize(rec_app);
|
||||
for (func_decl* acc1 : accessors) {
|
||||
enode* arg;
|
||||
if (acc1 == acc) {
|
||||
arg = n->get_arg(1);
|
||||
}
|
||||
else {
|
||||
acc_app = m.mk_app(acc1, arg1);
|
||||
arg = e_internalize(acc_app);
|
||||
}
|
||||
app_ref acc_own(m.mk_app(acc1, own), m);
|
||||
assert_eq_axiom(arg, acc_own, is_con);
|
||||
}
|
||||
// update_field is identity if 'n' is not created by a matching constructor.
|
||||
assert_eq_axiom(n, arg1, ~is_con);
|
||||
app_ref n_is_con(m.mk_app(rec, own), m);
|
||||
add_clause(~is_con, mk_literal(n_is_con));
|
||||
}
|
||||
|
||||
euf::theory_var solver::mk_var(enode* n) {
|
||||
if (is_attached_to_var(n))
|
||||
return n->get_th_var(get_id());
|
||||
euf::theory_var r = th_euf_solver::mk_var(n);
|
||||
VERIFY(r == static_cast<theory_var>(m_find.mk_var()));
|
||||
SASSERT(r == static_cast<int>(m_var_data.size()));
|
||||
m_var_data.push_back(alloc(var_data));
|
||||
var_data* d = m_var_data[r];
|
||||
ctx.attach_th_var(n, this, r);
|
||||
if (is_constructor(n)) {
|
||||
d->m_constructor = n;
|
||||
assert_accessor_axioms(n);
|
||||
}
|
||||
else if (is_update_field(n)) {
|
||||
assert_update_field_axioms(n);
|
||||
}
|
||||
else {
|
||||
sort* s = m.get_sort(n->get_expr());
|
||||
if (dt.get_datatype_num_constructors(s) == 1)
|
||||
assert_is_constructor_axiom(n, dt.get_datatype_constructors(s)->get(0));
|
||||
else if (get_config().m_dt_lazy_splits == 0 || (get_config().m_dt_lazy_splits == 1 && !s->is_infinite()))
|
||||
mk_split(r);
|
||||
}
|
||||
return r;
|
||||
}
|
||||
|
||||
|
||||
/**
|
||||
\brief Create a new case split for v. That is, create the atom (is_mk v) and mark it as relevant.
|
||||
If first is true, it means that v does not have recognizer yet.
|
||||
*/
|
||||
void solver::mk_split(theory_var v) {
|
||||
m_stats.m_splits++;
|
||||
|
||||
v = m_find.find(v);
|
||||
enode* n = var2enode(v);
|
||||
sort* srt = m.get_sort(n->get_expr());
|
||||
func_decl* non_rec_c = dt.get_non_rec_constructor(srt);
|
||||
unsigned non_rec_idx = dt.get_constructor_idx(non_rec_c);
|
||||
var_data* d = m_var_data[v];
|
||||
SASSERT(d->m_constructor == nullptr);
|
||||
func_decl* r = nullptr;
|
||||
|
||||
TRACE("dt", tout << "non_rec_c: " << non_rec_c->get_name() << " #rec: " << d->m_recognizers.size() << "\n";);
|
||||
|
||||
enode* recognizer = d->m_recognizers.get(non_rec_idx, nullptr);
|
||||
if (recognizer == nullptr)
|
||||
r = dt.get_constructor_is(non_rec_c);
|
||||
else if (ctx.value(recognizer) != l_false)
|
||||
// if is l_true, then we are done
|
||||
// otherwise wait for recognizer to be assigned.
|
||||
return;
|
||||
else {
|
||||
// look for a slot of d->m_recognizers that is 0, or it is not marked as relevant and is unassigned.
|
||||
unsigned idx = 0;
|
||||
ptr_vector<func_decl> const& constructors = *dt.get_datatype_constructors(srt);
|
||||
for (enode* curr : d->m_recognizers) {
|
||||
if (curr == nullptr) {
|
||||
// found empty slot...
|
||||
r = dt.get_constructor_is(constructors[idx]);
|
||||
break;
|
||||
}
|
||||
else if (ctx.value(curr) != l_false)
|
||||
return;
|
||||
++idx;
|
||||
}
|
||||
if (r == nullptr)
|
||||
return; // all recognizers are asserted to false... conflict will be detected...
|
||||
}
|
||||
SASSERT(r != nullptr);
|
||||
app_ref r_app(m.mk_app(r, n->get_expr()), m);
|
||||
TRACE("dt", tout << "creating split: " << mk_pp(r_app, m) << "\n";);
|
||||
b_internalize(r_app);
|
||||
}
|
||||
|
||||
void solver::apply_sort_cnstr(enode* n, sort* s) {
|
||||
force_push();
|
||||
// Remark: If s is an infinite sort, then it is not necessary to create
|
||||
// a theory variable.
|
||||
//
|
||||
// Actually, when the logical context has quantifiers, it is better to
|
||||
// disable this optimization.
|
||||
// Example:
|
||||
//
|
||||
// (forall (l list) (a Int) (= (len (cons a l)) (+ (len l) 1)))
|
||||
// (assert (> (len a) 1)
|
||||
//
|
||||
// If the theory variable is not created for 'a', then a wrong model will be generated.
|
||||
TRACE("dt", tout << "apply_sort_cnstr: #" << n->get_expr_id() << " " << mk_pp(n->get_expr(), m) << "\n";);
|
||||
TRACE("dt_bug",
|
||||
tout << "apply_sort_cnstr:\n" << mk_pp(n->get_expr(), m) << " ";
|
||||
tout << dt.is_datatype(s) << " ";
|
||||
if (dt.is_datatype(s)) tout << "is-infinite: " << s->is_infinite() << " ";
|
||||
if (dt.is_datatype(s)) tout << "attached: " << is_attached_to_var(n) << " ";
|
||||
tout << "\n";);
|
||||
|
||||
if (!is_attached_to_var(n) &&
|
||||
(/*ctx.has_quantifiers()*/ true ||
|
||||
(dt.is_datatype(s) && dt.has_nested_arrays()) ||
|
||||
(dt.is_datatype(s) && !s->is_infinite()))) {
|
||||
mk_var(n);
|
||||
}
|
||||
}
|
||||
|
||||
|
||||
void solver::new_eq_eh(euf::th_eq const& eq) {
|
||||
force_push();
|
||||
m_find.merge(eq.v1(), eq.v2());
|
||||
}
|
||||
|
||||
void solver::asserted(literal lit) {
|
||||
force_push();
|
||||
enode* n = bool_var2enode(lit.var());
|
||||
if (!is_recognizer(n))
|
||||
return;
|
||||
TRACE("dt", tout << "assigning recognizer: #" << n->get_expr_id() << " " << ctx.bpp(n) << "\n";);
|
||||
SASSERT(n->num_args() == 1);
|
||||
enode* arg = n->get_arg(0);
|
||||
theory_var tv = arg->get_th_var(get_id());
|
||||
tv = m_find.find(tv);
|
||||
var_data* d = m_var_data[tv];
|
||||
func_decl* r = n->get_decl();
|
||||
func_decl* c = dt.get_recognizer_constructor(r);
|
||||
if (!lit.sign()) {
|
||||
SASSERT(tv != euf::null_theory_var);
|
||||
if (d->m_constructor != nullptr && d->m_constructor->get_decl() == c)
|
||||
return; // do nothing
|
||||
assert_is_constructor_axiom(arg, c, lit);
|
||||
}
|
||||
else if (d->m_constructor == nullptr) // make sure a constructor is attached
|
||||
propagate_recognizer(tv, n);
|
||||
else if (d->m_constructor->get_decl() == c) // conflict
|
||||
sign_recognizer_conflict(d->m_constructor, n);
|
||||
}
|
||||
|
||||
void solver::add_recognizer(theory_var v, enode* recognizer) {
|
||||
SASSERT(is_recognizer(recognizer));
|
||||
v = m_find.find(v);
|
||||
var_data* d = m_var_data[v];
|
||||
sort* s = recognizer->get_decl()->get_domain(0);
|
||||
if (d->m_recognizers.empty()) {
|
||||
SASSERT(dt.is_datatype(s));
|
||||
d->m_recognizers.resize(dt.get_datatype_num_constructors(s), nullptr);
|
||||
}
|
||||
SASSERT(d->m_recognizers.size() == dt.get_datatype_num_constructors(s));
|
||||
unsigned c_idx = dt.get_recognizer_constructor_idx(recognizer->get_decl());
|
||||
if (d->m_recognizers[c_idx] == nullptr) {
|
||||
lbool val = ctx.value(recognizer);
|
||||
TRACE("dt", tout << "adding recognizer to v" << v << " rec: #" << recognizer->get_expr_id() << " val: " << val << "\n";);
|
||||
if (val == l_true) {
|
||||
// do nothing...
|
||||
// If recognizer assignment was already processed, then
|
||||
// d->m_constructor is already set.
|
||||
// Otherwise, it will be set when asserted is invoked.
|
||||
return;
|
||||
}
|
||||
if (val == l_false && d->m_constructor != nullptr) {
|
||||
func_decl* c_decl = dt.get_recognizer_constructor(recognizer->get_decl());
|
||||
if (d->m_constructor->get_decl() == c_decl) {
|
||||
// conflict
|
||||
sign_recognizer_conflict(d->m_constructor, recognizer);
|
||||
}
|
||||
return;
|
||||
}
|
||||
SASSERT(val == l_undef || (val == l_false && d->m_constructor == nullptr));
|
||||
d->m_recognizers[c_idx] = recognizer;
|
||||
ctx.push(set_vector_idx_trail<euf::solver, enode>(d->m_recognizers, c_idx));
|
||||
if (val == l_false)
|
||||
propagate_recognizer(v, recognizer);
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
\brief Propagate a recognizer assigned to false.
|
||||
*/
|
||||
void solver::propagate_recognizer(theory_var v, enode* recognizer) {
|
||||
SASSERT(is_recognizer(recognizer));
|
||||
SASSERT(static_cast<int>(m_find.find(v)) == v);
|
||||
SASSERT(ctx.value(recognizer) == l_false);
|
||||
unsigned num_unassigned = 0;
|
||||
unsigned unassigned_idx = UINT_MAX;
|
||||
enode* n = var2enode(v);
|
||||
sort* srt = m.get_sort(n->get_expr());
|
||||
var_data* d = m_var_data[v];
|
||||
if (d->m_recognizers.empty()) {
|
||||
theory_var w = recognizer->get_arg(0)->get_th_var(get_id());
|
||||
SASSERT(w != euf::null_theory_var);
|
||||
add_recognizer(w, recognizer);
|
||||
}
|
||||
CTRACE("dt", d->m_recognizers.empty(), ctx.display(tout););
|
||||
SASSERT(!d->m_recognizers.empty());
|
||||
literal_vector lits;
|
||||
enode_pair_vector eqs;
|
||||
unsigned idx = 0;
|
||||
for (enode* r : d->m_recognizers) {
|
||||
if (!r) {
|
||||
if (num_unassigned == 0)
|
||||
unassigned_idx = idx;
|
||||
num_unassigned++;
|
||||
}
|
||||
else if (ctx.value(r) == l_true)
|
||||
return; // nothing to be propagated
|
||||
else if (ctx.value(r) == l_false) {
|
||||
SASSERT(r->num_args() == 1);
|
||||
lits.push_back(~ctx.enode2literal(r));
|
||||
if (n != r->get_arg(0)) {
|
||||
// Argument of the current recognizer is not necessarily equal to n.
|
||||
// This can happen when n and r->get_arg(0) are in the same equivalence class.
|
||||
// We must add equality as an assumption to the conflict or propagation
|
||||
SASSERT(n->get_root() == r->get_arg(0)->get_root());
|
||||
eqs.push_back(euf::enode_pair(n, r->get_arg(0)));
|
||||
}
|
||||
}
|
||||
++idx;
|
||||
}
|
||||
TRACE("dt", tout << "propagate " << num_unassigned << " eqs: " << eqs.size() << "\n";);
|
||||
if (num_unassigned == 0)
|
||||
ctx.set_conflict(euf::th_propagation::mk(*this, lits, eqs));
|
||||
else if (num_unassigned == 1) {
|
||||
// propagate remaining recognizer
|
||||
SASSERT(!lits.empty());
|
||||
enode* r = d->m_recognizers[unassigned_idx];
|
||||
literal consequent;
|
||||
if (!r) {
|
||||
ptr_vector<func_decl> const& constructors = *dt.get_datatype_constructors(srt);
|
||||
func_decl* rec = dt.get_constructor_is(constructors[unassigned_idx]);
|
||||
app_ref rec_app(m.mk_app(rec, n->get_expr()), m);
|
||||
consequent = b_internalize(rec_app);
|
||||
}
|
||||
else
|
||||
consequent = ctx.enode2literal(r);
|
||||
ctx.propagate(consequent, euf::th_propagation::mk(*this, lits, eqs));
|
||||
}
|
||||
else if (get_config().m_dt_lazy_splits == 0 || (!srt->is_infinite() && get_config().m_dt_lazy_splits == 1))
|
||||
// there are more than 2 unassigned recognizers...
|
||||
// if eager splits are enabled... create new case split
|
||||
mk_split(v);
|
||||
}
|
||||
|
||||
void solver::merge_eh(theory_var v1, theory_var v2, theory_var, theory_var) {
|
||||
// v1 is the new root
|
||||
TRACE("dt", tout << "merging v" << v1 << " v" << v2 << "\n";);
|
||||
SASSERT(v1 == static_cast<int>(m_find.find(v1)));
|
||||
var_data* d1 = m_var_data[v1];
|
||||
var_data* d2 = m_var_data[v2];
|
||||
auto* con1 = d1->m_constructor;
|
||||
auto* con2 = d2->m_constructor;
|
||||
if (con2 != nullptr) {
|
||||
if (con1 == nullptr) {
|
||||
ctx.push(set_ptr_trail<euf::solver, enode>(con1));
|
||||
// check whether there is a recognizer in d1 that conflicts with con2;
|
||||
if (!d1->m_recognizers.empty()) {
|
||||
unsigned c_idx = dt.get_constructor_idx(con2->get_decl());
|
||||
enode* recognizer = d1->m_recognizers[c_idx];
|
||||
if (recognizer != nullptr && ctx.value(recognizer) == l_false) {
|
||||
sign_recognizer_conflict(con2, recognizer);
|
||||
return;
|
||||
}
|
||||
}
|
||||
d1->m_constructor = con2;
|
||||
}
|
||||
else if (con1->get_decl() != con2->get_decl())
|
||||
add_unit(~eq_internalize(con1->get_expr(), con2->get_expr()));
|
||||
}
|
||||
for (enode* e : d2->m_recognizers)
|
||||
if (e)
|
||||
add_recognizer(v1, e);
|
||||
}
|
||||
|
||||
ptr_vector<euf::enode> const& solver::get_array_args(enode* n) {
|
||||
m_array_args.reset();
|
||||
array::solver* th = dynamic_cast<array::solver*>(ctx.fid2solver(m_autil.get_family_id()));
|
||||
for (enode* p : th->parent_selects(n))
|
||||
m_array_args.push_back(p);
|
||||
app_ref def(m_autil.mk_default(n->get_expr()), m);
|
||||
m_array_args.push_back(ctx.get_enode(def));
|
||||
return m_array_args;
|
||||
}
|
||||
|
||||
// Assuming `app` is equal to a constructor term, return the constructor enode
|
||||
inline euf::enode* solver::oc_get_cstor(enode* app) {
|
||||
theory_var v = app->get_root()->get_th_var(get_id());
|
||||
SASSERT(v != euf::null_theory_var);
|
||||
v = m_find.find(v);
|
||||
var_data* d = m_var_data[v];
|
||||
SASSERT(d->m_constructor);
|
||||
return d->m_constructor;
|
||||
}
|
||||
|
||||
void solver::explain_is_child(enode* parent, enode* child) {
|
||||
enode* parentc = oc_get_cstor(parent);
|
||||
if (parent != parentc)
|
||||
m_used_eqs.push_back(enode_pair(parent, parentc));
|
||||
|
||||
// collect equalities on all children that may have been used.
|
||||
bool found = false;
|
||||
auto add = [&](enode* arg) {
|
||||
if (arg->get_root() == child->get_root()) {
|
||||
if (arg != child)
|
||||
m_used_eqs.push_back(enode_pair(arg, child));
|
||||
found = true;
|
||||
}
|
||||
};
|
||||
for (enode* arg : euf::enode_args(parentc)) {
|
||||
add(arg);
|
||||
sort* s = m.get_sort(arg->get_expr());
|
||||
if (m_autil.is_array(s) && dt.is_datatype(get_array_range(s)))
|
||||
for (enode* aarg : get_array_args(arg))
|
||||
add(aarg);
|
||||
}
|
||||
VERIFY(found);
|
||||
}
|
||||
|
||||
// explain the cycle root -> ... -> app -> root
|
||||
void solver::occurs_check_explain(enode* app, enode* root) {
|
||||
TRACE("dt", tout << "occurs_check_explain " << ctx.bpp(app) << " <-> " << ctx.bpp(root) << "\n";);
|
||||
|
||||
// first: explain that root=v, given that app=cstor(...,v,...)
|
||||
|
||||
explain_is_child(app, root);
|
||||
|
||||
// now explain app=cstor(..,v,..) where v=root, and recurse with parent of app
|
||||
while (app->get_root() != root->get_root()) {
|
||||
enode* parent_app = m_parent[app->get_root()];
|
||||
explain_is_child(parent_app, app);
|
||||
SASSERT(is_constructor(parent_app));
|
||||
app = parent_app;
|
||||
}
|
||||
|
||||
SASSERT(app->get_root() == root->get_root());
|
||||
if (app != root)
|
||||
m_used_eqs.push_back(enode_pair(app, root));
|
||||
|
||||
TRACE("dt",
|
||||
tout << "occurs_check\n"; for (enode_pair const& p : m_used_eqs) tout << ctx.bpp(p.first) << " - " << ctx.bpp(p.second) << " ";);
|
||||
}
|
||||
|
||||
// start exploring subgraph below `app`
|
||||
bool solver::occurs_check_enter(enode* app) {
|
||||
app = app->get_root();
|
||||
theory_var v = app->get_th_var(get_id());
|
||||
if (v == euf::null_theory_var)
|
||||
return false;
|
||||
v = m_find.find(v);
|
||||
var_data* d = m_var_data[v];
|
||||
if (!d->m_constructor)
|
||||
return false;
|
||||
enode* parent = d->m_constructor;
|
||||
oc_mark_on_stack(parent);
|
||||
for (enode* arg : euf::enode_args(parent)) {
|
||||
if (oc_cycle_free(arg))
|
||||
continue;
|
||||
if (oc_on_stack(arg)) {
|
||||
// arg was explored before app, and is still on the stack: cycle
|
||||
occurs_check_explain(parent, arg);
|
||||
return true;
|
||||
}
|
||||
// explore `arg` (with parent)
|
||||
expr* earg = arg->get_expr();
|
||||
sort* s = m.get_sort(earg);
|
||||
if (dt.is_datatype(s)) {
|
||||
m_parent.insert(arg->get_root(), parent);
|
||||
oc_push_stack(arg);
|
||||
}
|
||||
else if (m_autil.is_array(s) && dt.is_datatype(get_array_range(s))) {
|
||||
for (enode* aarg : get_array_args(arg)) {
|
||||
if (oc_cycle_free(aarg))
|
||||
continue;
|
||||
if (oc_on_stack(aarg)) {
|
||||
occurs_check_explain(parent, aarg);
|
||||
return true;
|
||||
}
|
||||
if (is_datatype(aarg)) {
|
||||
m_parent.insert(aarg->get_root(), parent);
|
||||
oc_push_stack(aarg);
|
||||
}
|
||||
}
|
||||
}
|
||||
}
|
||||
return false;
|
||||
}
|
||||
|
||||
/**
|
||||
\brief Check if n can be reached starting from n and following equalities and constructors.
|
||||
For example, occur_check(a1) returns true in the following set of equalities:
|
||||
a1 = cons(v1, a2)
|
||||
a2 = cons(v2, a3)
|
||||
a3 = cons(v3, a1)
|
||||
*/
|
||||
bool solver::occurs_check(enode* n) {
|
||||
TRACE("dt", tout << "occurs check: " << ctx.bpp(n) << "\n";);
|
||||
m_stats.m_occurs_check++;
|
||||
|
||||
bool res = false;
|
||||
oc_push_stack(n);
|
||||
|
||||
// DFS traversal from `n`. Look at top element and explore it.
|
||||
while (!res && !m_dfs.empty()) {
|
||||
stack_op op = m_dfs.back().first;
|
||||
enode* app = m_dfs.back().second;
|
||||
m_dfs.pop_back();
|
||||
|
||||
if (oc_cycle_free(app))
|
||||
continue;
|
||||
|
||||
TRACE("dt", tout << "occurs check loop: " << ctx.bpp(app) << (op == ENTER ? " enter" : " exit") << "\n";);
|
||||
|
||||
switch (op) {
|
||||
case ENTER:
|
||||
res = occurs_check_enter(app);
|
||||
break;
|
||||
|
||||
case EXIT:
|
||||
oc_mark_cycle_free(app);
|
||||
break;
|
||||
}
|
||||
}
|
||||
|
||||
if (res) {
|
||||
clear_mark();
|
||||
ctx.set_conflict(euf::th_propagation::mk(*this, m_used_eqs));
|
||||
}
|
||||
return res;
|
||||
}
|
||||
|
||||
sat::check_result solver::check() {
|
||||
force_push();
|
||||
int num_vars = get_num_vars();
|
||||
sat::check_result r = sat::check_result::CR_DONE;
|
||||
final_check_st _guard(*this);
|
||||
int start = s().rand()();
|
||||
for (int i = 0; i < num_vars; i++) {
|
||||
theory_var v = (i + start) % num_vars;
|
||||
if (v == static_cast<int>(m_find.find(v))) {
|
||||
enode* node = var2enode(v);
|
||||
if (!is_datatype(node))
|
||||
continue;
|
||||
if (!oc_cycle_free(node) && occurs_check(node))
|
||||
// conflict was detected...
|
||||
return sat::check_result::CR_CONTINUE;
|
||||
if (get_config().m_dt_lazy_splits > 0) {
|
||||
// using lazy case splits...
|
||||
var_data* d = m_var_data[v];
|
||||
if (d->m_constructor == nullptr) {
|
||||
clear_mark();
|
||||
mk_split(v);
|
||||
r = sat::check_result::CR_CONTINUE;
|
||||
}
|
||||
}
|
||||
}
|
||||
}
|
||||
return r;
|
||||
}
|
||||
|
||||
void solver::pop_core(unsigned num_scopes) {
|
||||
th_euf_solver::pop_core(num_scopes);
|
||||
std::for_each(m_var_data.begin() + get_num_vars(), m_var_data.end(), delete_proc<var_data>());
|
||||
m_var_data.shrink(get_num_vars());
|
||||
SASSERT(m_find.get_num_vars() == m_var_data.size());
|
||||
SASSERT(m_find.get_num_vars() == get_num_vars());
|
||||
}
|
||||
|
||||
void solver::get_antecedents(literal l, sat::ext_justification_idx idx, literal_vector& r, bool probing) {
|
||||
auto& jst = euf::th_propagation::from_index(idx);
|
||||
ctx.get_antecedents(l, jst, r, probing);
|
||||
}
|
||||
|
||||
void solver::add_value(euf::enode* n, model& mdl, expr_ref_vector& values) {
|
||||
theory_var v = n->get_th_var(get_id());
|
||||
v = m_find.find(v);
|
||||
SASSERT(v != euf::null_theory_var);
|
||||
enode* con = m_var_data[v]->m_constructor;
|
||||
func_decl* c_decl = con->get_decl();
|
||||
m_args.reset();
|
||||
for (enode* arg : euf::enode_args(m_var_data[v]->m_constructor))
|
||||
m_args.push_back(values.get(arg->get_root_id()));
|
||||
values.set(n->get_root_id(), m.mk_app(c_decl, m_args));
|
||||
}
|
||||
|
||||
void solver::add_dep(euf::enode* n, top_sort<euf::enode>& dep) {
|
||||
theory_var v = n->get_th_var(get_id());
|
||||
for (enode* arg : euf::enode_args(m_var_data[m_find.find(v)]->m_constructor))
|
||||
dep.add(n, arg);
|
||||
}
|
||||
|
||||
sat::literal solver::internalize(expr* e, bool sign, bool root, bool redundant) {
|
||||
if (!visit_rec(m, e, sign, root, redundant)) {
|
||||
TRACE("dt", tout << mk_pp(e, m) << "\n";);
|
||||
return sat::null_literal;
|
||||
}
|
||||
auto lit = ctx.expr2literal(e);
|
||||
if (sign)
|
||||
lit.neg();
|
||||
return lit;
|
||||
}
|
||||
|
||||
void solver::internalize(expr* e, bool redundant) {
|
||||
visit_rec(m, e, false, false, redundant);
|
||||
}
|
||||
|
||||
bool solver::visit(expr* e) {
|
||||
if (visited(e))
|
||||
return true;
|
||||
if (!is_app(e) || to_app(e)->get_family_id() != get_id()) {
|
||||
ctx.internalize(e, m_is_redundant);
|
||||
if (is_datatype(e))
|
||||
mk_var(expr2enode(e));
|
||||
return true;
|
||||
}
|
||||
m_stack.push_back(sat::eframe(e));
|
||||
return false;
|
||||
}
|
||||
|
||||
bool solver::visited(expr* e) {
|
||||
euf::enode* n = expr2enode(e);
|
||||
return n && n->is_attached_to(get_id());
|
||||
}
|
||||
|
||||
bool solver::post_visit(expr* term, bool sign, bool root) {
|
||||
euf::enode* n = expr2enode(term);
|
||||
SASSERT(!n || !n->is_attached_to(get_id()));
|
||||
if (!n)
|
||||
n = mk_enode(term);
|
||||
SASSERT(!n->is_attached_to(get_id()));
|
||||
if (is_constructor(term) || is_update_field(term)) {
|
||||
for (enode* arg : euf::enode_args(n)) {
|
||||
sort* s = m.get_sort(arg->get_expr());
|
||||
if (dt.is_datatype(s))
|
||||
mk_var(arg);
|
||||
else if (m_autil.is_array(s) && dt.is_datatype(get_array_range(s))) {
|
||||
app_ref def(m_autil.mk_default(arg->get_expr()), m);
|
||||
mk_var(e_internalize(def));
|
||||
}
|
||||
}
|
||||
mk_var(n);
|
||||
}
|
||||
else if (is_recognizer(term)) {
|
||||
enode* arg = n->get_arg(0);
|
||||
theory_var v = mk_var(arg);
|
||||
add_recognizer(v, n);
|
||||
mk_var(n);
|
||||
}
|
||||
else {
|
||||
SASSERT(is_accessor(term));
|
||||
SASSERT(n->num_args() == 1);
|
||||
mk_var(n->get_arg(0));
|
||||
}
|
||||
return true;
|
||||
}
|
||||
|
||||
void solver::collect_statistics(::statistics& st) const {
|
||||
st.update("datatype occurs check", m_stats.m_occurs_check);
|
||||
st.update("datatype splits", m_stats.m_splits);
|
||||
st.update("datatype constructor ax", m_stats.m_assert_cnstr);
|
||||
st.update("datatype accessor ax", m_stats.m_assert_accessor);
|
||||
st.update("datatype update ax", m_stats.m_assert_update_field);
|
||||
}
|
||||
|
||||
std::ostream& solver::display(std::ostream& out) const {
|
||||
unsigned num_vars = get_num_vars();
|
||||
if (num_vars > 0)
|
||||
out << "Theory datatype:\n";
|
||||
for (unsigned v = 0; v < num_vars; v++)
|
||||
display_var(out, v);
|
||||
return out;
|
||||
}
|
||||
|
||||
void solver::display_var(std::ostream& out, theory_var v) const {
|
||||
var_data* d = m_var_data[v];
|
||||
out << "v" << v << " #" << var2expr(v)->get_id() << " -> v" << m_find.find(v) << " ";
|
||||
if (d->m_constructor)
|
||||
out << ctx.bpp(d->m_constructor);
|
||||
else
|
||||
out << "(null)";
|
||||
out << "\n";
|
||||
}
|
||||
}
|
159
src/sat/smt/dt_solver.h
Normal file
159
src/sat/smt/dt_solver.h
Normal file
|
@ -0,0 +1,159 @@
|
|||
/*++
|
||||
Copyright (c) 2020 Microsoft Corporation
|
||||
|
||||
Module Name:
|
||||
|
||||
dt_solver.h
|
||||
|
||||
Abstract:
|
||||
|
||||
Theory plugin for altegraic datatypes
|
||||
|
||||
Author:
|
||||
|
||||
Nikolaj Bjorner (nbjorner) 2020-09-08
|
||||
|
||||
--*/
|
||||
#pragma once
|
||||
|
||||
#include "sat/smt/sat_th.h"
|
||||
#include "ast/datatype_decl_plugin.h"
|
||||
#include "ast/array_decl_plugin.h"
|
||||
|
||||
namespace euf {
|
||||
class solver;
|
||||
}
|
||||
|
||||
namespace dt {
|
||||
|
||||
class solver : public euf::th_euf_solver {
|
||||
typedef euf::theory_var theory_var;
|
||||
typedef euf::theory_id theory_id;
|
||||
typedef euf::enode enode;
|
||||
typedef euf::enode_pair enode_pair;
|
||||
typedef euf::enode_pair_vector enode_pair_vector;
|
||||
typedef sat::bool_var bool_var;
|
||||
typedef sat::literal literal;
|
||||
typedef sat::literal_vector literal_vector;
|
||||
typedef union_find<solver, euf::solver> dt_union_find;
|
||||
|
||||
struct var_data {
|
||||
ptr_vector<enode> m_recognizers; //!< recognizers of this equivalence class that are being watched.
|
||||
enode * m_constructor; //!< constructor of this equivalence class, 0 if there is no constructor in the eqc.
|
||||
var_data():
|
||||
m_constructor(nullptr) {
|
||||
}
|
||||
};
|
||||
|
||||
// class for managing state of final_check
|
||||
class final_check_st {
|
||||
solver& s;
|
||||
public:
|
||||
final_check_st(solver& s);
|
||||
~final_check_st();
|
||||
};
|
||||
|
||||
struct stats {
|
||||
unsigned m_occurs_check, m_splits;
|
||||
unsigned m_assert_cnstr, m_assert_accessor, m_assert_update_field;
|
||||
void reset() { memset(this, 0, sizeof(*this)); }
|
||||
stats() { reset(); }
|
||||
};
|
||||
|
||||
datatype_util dt;
|
||||
array_util m_autil;
|
||||
stats m_stats;
|
||||
ptr_vector<var_data> m_var_data;
|
||||
dt_union_find m_find;
|
||||
expr_ref_vector m_args;
|
||||
|
||||
bool is_constructor(expr * f) const { return dt.is_constructor(f); }
|
||||
bool is_recognizer(expr * f) const { return dt.is_recognizer(f); }
|
||||
bool is_accessor(expr * f) const { return dt.is_accessor(f); }
|
||||
bool is_update_field(expr * f) const { return dt.is_update_field(f); }
|
||||
|
||||
bool is_constructor(enode * n) const { return is_constructor(n->get_expr()); }
|
||||
bool is_recognizer(enode * n) const { return is_recognizer(n->get_expr()); }
|
||||
bool is_accessor(enode * n) const { return is_accessor(n->get_expr()); }
|
||||
bool is_update_field(enode * n) const { return dt.is_update_field(n->get_expr()); }
|
||||
|
||||
bool is_datatype(expr* e) const { return dt.is_datatype(m.get_sort(e)); }
|
||||
bool is_datatype(enode* n) const { return is_datatype(n->get_expr()); }
|
||||
|
||||
void assert_eq_axiom(enode * lhs, expr * rhs, literal antecedent = sat::null_literal);
|
||||
void assert_is_constructor_axiom(enode * n, func_decl * c, literal antecedent = sat::null_literal);
|
||||
void assert_accessor_axioms(enode * n);
|
||||
void assert_update_field_axioms(enode * n);
|
||||
void add_recognizer(theory_var v, enode * recognizer);
|
||||
void propagate_recognizer(theory_var v, enode * r);
|
||||
void sign_recognizer_conflict(enode * c, enode * r);
|
||||
|
||||
typedef enum { ENTER, EXIT } stack_op;
|
||||
typedef obj_map<enode, enode*> parent_tbl;
|
||||
typedef std::pair<stack_op, enode*> stack_entry;
|
||||
|
||||
ptr_vector<enode> m_to_unmark1;
|
||||
ptr_vector<enode> m_to_unmark2;
|
||||
enode_pair_vector m_used_eqs; // conflict, if any
|
||||
parent_tbl m_parent; // parent explanation for occurs_check
|
||||
svector<stack_entry> m_dfs; // stack for DFS for occurs_check
|
||||
|
||||
void clear_mark();
|
||||
|
||||
void oc_mark_on_stack(enode * n);
|
||||
bool oc_on_stack(enode * n) const { return n->get_root()->is_marked1(); }
|
||||
|
||||
void oc_mark_cycle_free(enode * n);
|
||||
bool oc_cycle_free(enode * n) const { return n->get_root()->is_marked2(); }
|
||||
|
||||
void oc_push_stack(enode * n);
|
||||
ptr_vector<enode> m_array_args;
|
||||
ptr_vector<enode> const& get_array_args(enode* n);
|
||||
|
||||
void pop_core(unsigned n) override;
|
||||
|
||||
enode * oc_get_cstor(enode * n);
|
||||
bool occurs_check(enode * n);
|
||||
bool occurs_check_enter(enode * n);
|
||||
void occurs_check_explain(enode * top, enode * root);
|
||||
void explain_is_child(enode* parent, enode* child);
|
||||
|
||||
void mk_split(theory_var v);
|
||||
|
||||
void display_var(std::ostream & out, theory_var v) const;
|
||||
|
||||
// internalize
|
||||
bool visit(expr* e) override;
|
||||
bool visited(expr* e) override;
|
||||
bool post_visit(expr* e, bool sign, bool root) override;
|
||||
void clone_var(solver& src, theory_var v);
|
||||
|
||||
public:
|
||||
solver(euf::solver& ctx, theory_id id);
|
||||
~solver() override;
|
||||
|
||||
bool is_external(bool_var v) override { return false; }
|
||||
void get_antecedents(literal l, sat::ext_justification_idx idx, literal_vector& r, bool probing) override;
|
||||
void asserted(literal l) override;
|
||||
sat::check_result check() override;
|
||||
|
||||
std::ostream& display(std::ostream& out) const override;
|
||||
std::ostream& display_justification(std::ostream& out, sat::ext_justification_idx idx) const override { return euf::th_propagation::from_index(idx).display(out); }
|
||||
std::ostream& display_constraint(std::ostream& out, sat::ext_constraint_idx idx) const override { return display_justification(out, idx); }
|
||||
void collect_statistics(statistics& st) const override;
|
||||
euf::th_solver* clone(euf::solver& ctx) override;
|
||||
void new_eq_eh(euf::th_eq const& eq) override;
|
||||
bool unit_propagate() override { return false; }
|
||||
void add_value(euf::enode* n, model& mdl, expr_ref_vector& values) override;
|
||||
void add_dep(euf::enode* n, top_sort<euf::enode>& dep) override;
|
||||
sat::literal internalize(expr* e, bool sign, bool root, bool redundant) override;
|
||||
void internalize(expr* e, bool redundant) override;
|
||||
euf::theory_var mk_var(euf::enode* n) override;
|
||||
void apply_sort_cnstr(euf::enode* n, sort* s) override;
|
||||
bool is_shared(theory_var v) const override { return false; }
|
||||
|
||||
void merge_eh(theory_var, theory_var, theory_var v1, theory_var v2);
|
||||
void after_merge_eh(theory_var r1, theory_var r2, theory_var v1, theory_var v2) {}
|
||||
void unmerge_eh(theory_var v1, theory_var v2) {}
|
||||
};
|
||||
}
|
|
@ -134,14 +134,14 @@ namespace euf {
|
|||
lit = lit2;
|
||||
}
|
||||
|
||||
m_var2expr.reserve(v + 1, nullptr);
|
||||
if (m_var2expr[v]) {
|
||||
m_bool_var2expr.reserve(v + 1, nullptr);
|
||||
if (m_bool_var2expr[v]) {
|
||||
SASSERT(m_egraph.find(e));
|
||||
SASSERT(m_egraph.find(e)->bool_var() == v);
|
||||
return lit;
|
||||
}
|
||||
TRACE("euf", tout << "attach " << v << " " << mk_bounded_pp(e, m) << "\n";);
|
||||
m_var2expr[v] = e;
|
||||
m_bool_var2expr[v] = e;
|
||||
m_var_trail.push_back(v);
|
||||
enode* n = m_egraph.find(e);
|
||||
if (!n)
|
||||
|
@ -367,12 +367,15 @@ namespace euf {
|
|||
return true;
|
||||
|
||||
return false;
|
||||
|
||||
}
|
||||
|
||||
expr_ref solver::mk_eq(expr* e1, expr* e2) {
|
||||
if (e1 == e2)
|
||||
expr_ref _e1(e1, m);
|
||||
expr_ref _e2(e2, m);
|
||||
if (m.are_equal(e1, e2))
|
||||
return expr_ref(m.mk_true(), m);
|
||||
if (m.are_distinct(e1, e2))
|
||||
return expr_ref(m.mk_false(), m);
|
||||
expr_ref r(m.mk_eq(e2, e1), m);
|
||||
if (!m_egraph.find(r))
|
||||
r = m.mk_eq(e1, e2);
|
||||
|
|
|
@ -84,7 +84,7 @@ namespace euf {
|
|||
|
||||
void solver::log_antecedents(std::ostream& out, literal l, literal_vector const& r) {
|
||||
for (sat::literal l : r) {
|
||||
expr* n = m_var2expr[l.var()];
|
||||
expr* n = m_bool_var2expr[l.var()];
|
||||
out << ~l << ": ";
|
||||
if (!l.sign()) out << "! ";
|
||||
out << mk_bounded_pp(n, m) << "\n";
|
||||
|
@ -93,7 +93,7 @@ namespace euf {
|
|||
if (l != sat::null_literal) {
|
||||
out << l << ": ";
|
||||
if (l.sign()) out << "! ";
|
||||
expr* n = m_var2expr[l.var()];
|
||||
expr* n = m_bool_var2expr[l.var()];
|
||||
out << mk_bounded_pp(n, m) << "\n";
|
||||
}
|
||||
}
|
||||
|
|
|
@ -64,7 +64,7 @@ namespace euf {
|
|||
m_relevant_expr_ids.resize(max_id + 1, false);
|
||||
auto const& core = m_dual_solver->core();
|
||||
for (auto lit : core) {
|
||||
expr* e = m_var2expr.get(lit.var(), nullptr);
|
||||
expr* e = m_bool_var2expr.get(lit.var(), nullptr);
|
||||
if (e)
|
||||
todo.push_back(e);
|
||||
}
|
||||
|
|
|
@ -26,6 +26,7 @@ Author:
|
|||
#include "sat/smt/arith_solver.h"
|
||||
#include "sat/smt/q_solver.h"
|
||||
#include "sat/smt/fpa_solver.h"
|
||||
#include "sat/smt/dt_solver.h"
|
||||
|
||||
namespace euf {
|
||||
|
||||
|
@ -60,9 +61,9 @@ namespace euf {
|
|||
* retrieve extension that is associated with Boolean variable.
|
||||
*/
|
||||
th_solver* solver::bool_var2solver(sat::bool_var v) {
|
||||
if (v >= m_var2expr.size())
|
||||
if (v >= m_bool_var2expr.size())
|
||||
return nullptr;
|
||||
expr* e = m_var2expr[v];
|
||||
expr* e = m_bool_var2expr[v];
|
||||
if (!e)
|
||||
return nullptr;
|
||||
return expr2solver(e);
|
||||
|
@ -81,10 +82,8 @@ namespace euf {
|
|||
auto* ext = m_id2solver.get(fid, nullptr);
|
||||
if (ext)
|
||||
return ext;
|
||||
ext = alloc(q::solver, *this);
|
||||
ext->set_solver(m_solver);
|
||||
ext->push_scopes(s().num_scopes());
|
||||
add_solver(fid, ext);
|
||||
ext = alloc(q::solver, *this, fid);
|
||||
add_solver(ext);
|
||||
return ext;
|
||||
}
|
||||
|
||||
|
@ -101,34 +100,37 @@ namespace euf {
|
|||
array_util au(m);
|
||||
fpa_util fpa(m);
|
||||
arith_util arith(m);
|
||||
if (pb.get_family_id() == fid)
|
||||
datatype_util dt(m);
|
||||
if (pb.get_family_id() == fid)
|
||||
ext = alloc(sat::ba_solver, *this, fid);
|
||||
else if (bvu.get_family_id() == fid)
|
||||
else if (bvu.get_family_id() == fid)
|
||||
ext = alloc(bv::solver, *this, fid);
|
||||
else if (au.get_family_id() == fid)
|
||||
else if (au.get_family_id() == fid)
|
||||
ext = alloc(array::solver, *this, fid);
|
||||
else if (fpa.get_family_id() == fid)
|
||||
else if (fpa.get_family_id() == fid)
|
||||
ext = alloc(fpa::solver, *this);
|
||||
else if (arith.get_family_id() == fid)
|
||||
ext = alloc(arith::solver, *this, fid);
|
||||
else if (dt.get_family_id() == fid)
|
||||
ext = alloc(dt::solver, *this, fid);
|
||||
|
||||
if (ext) {
|
||||
if (use_drat())
|
||||
s().get_drat().add_theory(fid, ext->name());
|
||||
ext->set_solver(m_solver);
|
||||
ext->push_scopes(s().num_scopes());
|
||||
add_solver(fid, ext);
|
||||
if (ext->use_diseqs())
|
||||
m_egraph.set_th_propagates_diseqs(fid);
|
||||
}
|
||||
if (ext)
|
||||
add_solver(ext);
|
||||
else if (f)
|
||||
unhandled_function(f);
|
||||
return ext;
|
||||
}
|
||||
|
||||
void solver::add_solver(family_id fid, th_solver* th) {
|
||||
void solver::add_solver(th_solver* th) {
|
||||
family_id fid = th->get_id();
|
||||
if (use_drat())
|
||||
s().get_drat().add_theory(fid, th->name());
|
||||
th->set_solver(m_solver);
|
||||
th->push_scopes(s().num_scopes());
|
||||
m_solvers.push_back(th);
|
||||
m_id2solver.setx(fid, th, nullptr);
|
||||
if (th->use_diseqs())
|
||||
m_egraph.set_th_propagates_diseqs(fid);
|
||||
}
|
||||
|
||||
void solver::unhandled_function(func_decl* f) {
|
||||
|
@ -146,7 +148,7 @@ namespace euf {
|
|||
}
|
||||
|
||||
bool solver::is_external(bool_var v) {
|
||||
if (nullptr != m_var2expr.get(v, nullptr))
|
||||
if (nullptr != m_bool_var2expr.get(v, nullptr))
|
||||
return true;
|
||||
for (auto* s : m_solvers)
|
||||
if (s->is_external(v))
|
||||
|
@ -154,10 +156,18 @@ namespace euf {
|
|||
return false;
|
||||
}
|
||||
|
||||
bool solver::propagate(literal l, ext_constraint_idx idx) {
|
||||
bool solver::propagated(literal l, ext_constraint_idx idx) {
|
||||
auto* ext = sat::constraint_base::to_extension(idx);
|
||||
SASSERT(ext != this);
|
||||
return ext->propagate(l, idx);
|
||||
return ext->propagated(l, idx);
|
||||
}
|
||||
|
||||
void solver::set_conflict(ext_constraint_idx idx) {
|
||||
s().set_conflict(sat::justification::mk_ext_justification(s().scope_lvl(), idx));
|
||||
}
|
||||
|
||||
void solver::propagate(literal lit, ext_justification_idx idx) {
|
||||
s().assign(lit, sat::justification::mk_ext_justification(s().scope_lvl(), idx));
|
||||
}
|
||||
|
||||
void solver::get_antecedents(literal l, ext_justification_idx idx, literal_vector& r, bool probing) {
|
||||
|
@ -192,6 +202,23 @@ namespace euf {
|
|||
log_antecedents(l, r);
|
||||
}
|
||||
|
||||
void solver::get_antecedents(literal l, th_propagation& jst, literal_vector& r, bool probing) {
|
||||
for (auto lit : euf::th_propagation::lits(jst))
|
||||
r.push_back(lit);
|
||||
for (auto eq : euf::th_propagation::eqs(jst))
|
||||
add_antecedent(eq.first, eq.second);
|
||||
|
||||
if (!probing && use_drat()) {
|
||||
literal_vector lits;
|
||||
for (auto lit : euf::th_propagation::lits(jst))
|
||||
lits.push_back(~lit);
|
||||
lits.push_back(l);
|
||||
get_drat().add(lits, sat::status::th(m_is_redundant, jst.ext().get_id()));
|
||||
for (auto eq : euf::th_propagation::eqs(jst))
|
||||
IF_VERBOSE(0, verbose_stream() << "drat-log with equalities is TBD " << eq.first->get_expr_id() << "\n");
|
||||
}
|
||||
}
|
||||
|
||||
void solver::add_antecedent(enode* a, enode* b) {
|
||||
m_egraph.explain_eq<size_t>(m_explain, a, b);
|
||||
}
|
||||
|
@ -217,7 +244,7 @@ namespace euf {
|
|||
m_egraph.explain<size_t>(m_explain);
|
||||
break;
|
||||
case constraint::kind_t::eq:
|
||||
e = m_var2expr[l.var()];
|
||||
e = m_bool_var2expr[l.var()];
|
||||
n = m_egraph.find(e);
|
||||
SASSERT(n);
|
||||
SASSERT(n->is_equality());
|
||||
|
@ -225,7 +252,7 @@ namespace euf {
|
|||
m_egraph.explain_eq<size_t>(m_explain, n->get_arg(0), n->get_arg(1));
|
||||
break;
|
||||
case constraint::kind_t::lit:
|
||||
e = m_var2expr[l.var()];
|
||||
e = m_bool_var2expr[l.var()];
|
||||
n = m_egraph.find(e);
|
||||
SASSERT(n);
|
||||
SASSERT(m.is_bool(n->get_expr()));
|
||||
|
@ -238,7 +265,7 @@ namespace euf {
|
|||
}
|
||||
|
||||
void solver::asserted(literal l) {
|
||||
expr* e = m_var2expr.get(l.var(), nullptr);
|
||||
expr* e = m_bool_var2expr.get(l.var(), nullptr);
|
||||
if (!e) {
|
||||
TRACE("euf", tout << "asserted: " << l << "@" << s().scope_lvl() << "\n";);
|
||||
return;
|
||||
|
@ -421,7 +448,7 @@ namespace euf {
|
|||
m_egraph.pop(n);
|
||||
scope const & s = m_scopes[m_scopes.size() - n];
|
||||
for (unsigned i = m_var_trail.size(); i-- > s.m_var_lim; )
|
||||
m_var2expr[m_var_trail[i]] = nullptr;
|
||||
m_bool_var2expr[m_var_trail[i]] = nullptr;
|
||||
m_var_trail.shrink(s.m_var_lim);
|
||||
m_scopes.shrink(m_scopes.size() - n);
|
||||
SASSERT(m_egraph.num_scopes() == m_scopes.size());
|
||||
|
@ -534,7 +561,7 @@ namespace euf {
|
|||
m_egraph.display(out);
|
||||
out << "bool-vars\n";
|
||||
for (unsigned v : m_var_trail) {
|
||||
expr* e = m_var2expr[v];
|
||||
expr* e = m_bool_var2expr[v];
|
||||
out << v << ": " << e->get_id() << " " << m_solver->value(v) << " " << mk_bounded_pp(e, m, 1) << "\n";
|
||||
}
|
||||
for (auto* e : m_solvers)
|
||||
|
@ -584,6 +611,17 @@ namespace euf {
|
|||
st.update("euf ackerman", m_stats.m_ackerman);
|
||||
}
|
||||
|
||||
enode* solver::copy(solver& dst_ctx, enode* src_n) {
|
||||
if (!src_n)
|
||||
return nullptr;
|
||||
ast_translation tr(m, dst_ctx.get_manager(), false);
|
||||
expr* e1 = src_n->get_expr();
|
||||
expr* e2 = tr(e1);
|
||||
euf::enode* n2 = dst_ctx.get_enode(e2);
|
||||
SASSERT(n2);
|
||||
return n2;
|
||||
}
|
||||
|
||||
sat::extension* solver::copy(sat::solver* s) {
|
||||
auto* r = alloc(solver, *m_to_m, *m_to_si);
|
||||
r->m_config = m_config;
|
||||
|
@ -596,10 +634,12 @@ namespace euf {
|
|||
};
|
||||
r->m_egraph.copy_from(m_egraph, copy_justification);
|
||||
r->set_solver(s);
|
||||
for (unsigned i = 0; i < m_id2solver.size(); ++i) {
|
||||
auto* e = m_id2solver[i];
|
||||
if (e)
|
||||
r->add_solver(i, e->clone(s, *r));
|
||||
for (auto* s_orig : m_id2solver) {
|
||||
if (s_orig) {
|
||||
auto* s_clone = s_orig->clone(*r);
|
||||
r->add_solver(s_clone);
|
||||
s_clone->set_solver(s);
|
||||
}
|
||||
}
|
||||
return r;
|
||||
}
|
||||
|
@ -653,8 +693,8 @@ namespace euf {
|
|||
unsigned solver::max_var(unsigned w) const {
|
||||
for (auto* e : m_solvers)
|
||||
w = e->max_var(w);
|
||||
for (unsigned sz = m_var2expr.size(); sz-- > 0; ) {
|
||||
expr* n = m_var2expr[sz];
|
||||
for (unsigned sz = m_bool_var2expr.size(); sz-- > 0; ) {
|
||||
expr* n = m_bool_var2expr[sz];
|
||||
if (n && m.is_bool(n)) {
|
||||
w = std::max(w, sz);
|
||||
break;
|
||||
|
|
|
@ -94,7 +94,7 @@ namespace euf {
|
|||
scoped_ptr<sat::dual_solver> m_dual_solver;
|
||||
user::solver* m_user_propagator{ nullptr };
|
||||
|
||||
ptr_vector<expr> m_var2expr;
|
||||
ptr_vector<expr> m_bool_var2expr;
|
||||
ptr_vector<size_t> m_explain;
|
||||
unsigned m_num_scopes{ 0 };
|
||||
unsigned_vector m_var_trail;
|
||||
|
@ -132,8 +132,7 @@ namespace euf {
|
|||
th_solver* quantifier2solver();
|
||||
th_solver* expr2solver(expr* e);
|
||||
th_solver* bool_var2solver(sat::bool_var v);
|
||||
th_solver* fid2solver(family_id fid) const { return m_id2solver.get(fid, nullptr); }
|
||||
void add_solver(family_id fid, th_solver* th);
|
||||
void add_solver(th_solver* th);
|
||||
void init_ackerman();
|
||||
|
||||
// model building
|
||||
|
@ -214,10 +213,13 @@ namespace euf {
|
|||
ast_manager& get_manager() { return m; }
|
||||
enode* get_enode(expr* e) const { return m_egraph.find(e); }
|
||||
sat::literal expr2literal(expr* e) const { return enode2literal(get_enode(e)); }
|
||||
sat::literal enode2literal(enode* e) const { return sat::literal(e->bool_var(), false); }
|
||||
sat::literal enode2literal(enode* n) const { return sat::literal(n->bool_var(), false); }
|
||||
lbool value(enode* n) const { return s().value(enode2literal(n)); }
|
||||
smt_params const& get_config() const { return m_config; }
|
||||
region& get_region() { return m_trail.get_region(); }
|
||||
egraph& get_egraph() { return m_egraph; }
|
||||
th_solver* fid2solver(family_id fid) const { return m_id2solver.get(fid, nullptr); }
|
||||
|
||||
template <typename C>
|
||||
void push(C const& c) { m_trail.push(c); }
|
||||
template <typename V>
|
||||
|
@ -238,13 +240,22 @@ namespace euf {
|
|||
double get_reward(literal l, ext_constraint_idx idx, sat::literal_occs_fun& occs) const override;
|
||||
bool is_extended_binary(ext_justification_idx idx, literal_vector& r) override;
|
||||
bool is_external(bool_var v) override;
|
||||
bool propagate(literal l, ext_constraint_idx idx) override;
|
||||
bool propagated(literal l, ext_constraint_idx idx) override;
|
||||
bool unit_propagate() override;
|
||||
bool propagate(enode* a, enode* b, ext_justification_idx);
|
||||
|
||||
void propagate(literal lit, ext_justification_idx idx);
|
||||
bool propagate(enode* a, enode* b, ext_justification_idx idx);
|
||||
void set_conflict(ext_justification_idx idx);
|
||||
|
||||
void propagate(literal lit, th_propagation* p) { propagate(lit, p->to_index()); }
|
||||
bool propagate(enode* a, enode* b, th_propagation* p) { return propagate(a, b, p->to_index()); }
|
||||
void set_conflict(th_propagation* p) { set_conflict(p->to_index()); }
|
||||
|
||||
bool set_root(literal l, literal r) override;
|
||||
void flush_roots() override;
|
||||
|
||||
void get_antecedents(literal l, ext_justification_idx idx, literal_vector& r, bool probing) override;
|
||||
void get_antecedents(literal l, th_propagation& jst, literal_vector& r, bool probing);
|
||||
void add_antecedent(enode* a, enode* b);
|
||||
void asserted(literal l) override;
|
||||
sat::check_result check() override;
|
||||
|
@ -260,8 +271,10 @@ namespace euf {
|
|||
std::ostream& display(std::ostream& out) const override;
|
||||
std::ostream& display_justification(std::ostream& out, ext_justification_idx idx) const override;
|
||||
std::ostream& display_constraint(std::ostream& out, ext_constraint_idx idx) const override;
|
||||
euf::egraph::b_pp bpp(enode* n) { return m_egraph.bpp(n); }
|
||||
void collect_statistics(statistics& st) const override;
|
||||
extension* copy(sat::solver* s) override;
|
||||
enode* copy(solver& dst_ctx, enode* src_n);
|
||||
void find_mutexes(literal_vector& lits, vector<literal_vector>& mutexes) override;
|
||||
void gc() override;
|
||||
void pop_reinit() override;
|
||||
|
@ -288,7 +301,7 @@ namespace euf {
|
|||
expr_ref mk_eq(expr* e1, expr* e2);
|
||||
expr_ref mk_eq(euf::enode* n1, euf::enode* n2) { return mk_eq(n1->get_expr(), n2->get_expr()); }
|
||||
euf::enode* mk_enode(expr* e, unsigned n, enode* const* args) { return m_egraph.mk(e, n, args); }
|
||||
expr* bool_var2expr(sat::bool_var v) { return m_var2expr.get(v, nullptr); }
|
||||
expr* bool_var2expr(sat::bool_var v) { return m_bool_var2expr.get(v, nullptr); }
|
||||
sat::literal attach_lit(sat::literal lit, expr* e);
|
||||
void unhandled_function(func_decl* f);
|
||||
th_rewriter& get_rewriter() { return m_rewriter; }
|
||||
|
|
|
@ -42,8 +42,7 @@ namespace fpa {
|
|||
}
|
||||
|
||||
|
||||
expr_ref solver::convert(expr* e)
|
||||
{
|
||||
expr_ref solver::convert(expr* e) {
|
||||
expr_ref res(m);
|
||||
expr* ccnv;
|
||||
TRACE("t_fpa", tout << "converting " << mk_ismt2_pp(e, m) << std::endl;);
|
||||
|
@ -105,6 +104,8 @@ namespace fpa {
|
|||
}
|
||||
|
||||
bool solver::visit(expr* e) {
|
||||
if (visited(e))
|
||||
return true;
|
||||
if (!is_app(e) || to_app(e)->get_family_id() != get_id()) {
|
||||
ctx.internalize(e, m_is_redundant);
|
||||
return true;
|
||||
|
@ -161,12 +162,12 @@ namespace fpa {
|
|||
SASSERT(m_fpa_util.is_float(n->get_expr()) || m_fpa_util.is_rm(n->get_expr()));
|
||||
SASSERT(n->get_decl()->get_range() == s);
|
||||
|
||||
expr* owner = n->get_expr();
|
||||
|
||||
if (is_attached_to_var(n))
|
||||
return;
|
||||
attach_new_th_var(n);
|
||||
|
||||
expr* owner = n->get_expr();
|
||||
|
||||
if (m_fpa_util.is_rm(s) && !m_fpa_util.is_bv2rm(owner)) {
|
||||
// For every RM term, we need to make sure that it's
|
||||
// associated bit-vector is within the valid range.
|
||||
|
|
|
@ -71,7 +71,7 @@ namespace fpa {
|
|||
void get_antecedents(sat::literal l, sat::ext_justification_idx idx, sat::literal_vector& r, bool probing) override { UNREACHABLE(); }
|
||||
sat::check_result check() override { return sat::check_result::CR_DONE; }
|
||||
|
||||
euf::th_solver* clone(sat::solver*, euf::solver& ctx) override { return alloc(solver, ctx); }
|
||||
euf::th_solver* clone(euf::solver& ctx) override { return alloc(solver, ctx); }
|
||||
|
||||
};
|
||||
|
||||
|
|
|
@ -94,7 +94,8 @@ namespace q {
|
|||
init_solver();
|
||||
::solver::scoped_push _sp(*m_solver);
|
||||
expr_ref_vector vars(m);
|
||||
expr_ref body = specialize(q, vars);
|
||||
quantifier* q_flat = qs.flatten(q);
|
||||
expr_ref body = specialize(q_flat, vars);
|
||||
m_solver->assert_expr(body);
|
||||
lbool r = m_solver->check_sat(0, nullptr);
|
||||
if (r == l_undef)
|
||||
|
@ -105,17 +106,17 @@ namespace q {
|
|||
m_solver->get_model(mdl0);
|
||||
expr_ref proj(m);
|
||||
auto add_projection = [&](model& mdl, bool inv) {
|
||||
proj = project(mdl, q, vars, inv);
|
||||
proj = project(mdl, q_flat, vars, inv);
|
||||
if (!proj)
|
||||
return;
|
||||
if (is_forall(q))
|
||||
qs.add_clause(~ctx.expr2literal(q), ctx.b_internalize(proj));
|
||||
else
|
||||
qs.add_clause(ctx.expr2literal(q), ~ctx.b_internalize(proj));
|
||||
qs.add_clause(ctx.expr2literal(q), ~ctx.b_internalize(proj));
|
||||
};
|
||||
bool added = false;
|
||||
#if 0
|
||||
m_model_finder.restrict_instantiations(*m_solver, *mdl0, q, vars);
|
||||
m_model_finder.restrict_instantiations(*m_solver, *mdl0, q_flat, vars);
|
||||
for (unsigned i = 0; i < m_max_cex && l_true == m_solver->check_sat(0, nullptr); ++i) {
|
||||
m_solver->get_model(mdl1);
|
||||
add_projection(*mdl1, true);
|
||||
|
|
|
@ -24,8 +24,8 @@ Author:
|
|||
|
||||
namespace q {
|
||||
|
||||
solver::solver(euf::solver& ctx):
|
||||
th_euf_solver(ctx, symbol("quant"), ctx.get_manager().get_family_id(symbol("quant"))),
|
||||
solver::solver(euf::solver& ctx, family_id fid) :
|
||||
th_euf_solver(ctx, ctx.get_manager().get_family_name(fid), fid),
|
||||
m_mbqi(ctx, *this)
|
||||
{}
|
||||
|
||||
|
@ -61,8 +61,9 @@ namespace q {
|
|||
st.update("quantifier asserts", m_stats.m_num_quantifier_asserts);
|
||||
}
|
||||
|
||||
euf::th_solver* solver::clone(sat::solver* s, euf::solver& ctx) {
|
||||
return alloc(solver, ctx);
|
||||
euf::th_solver* solver::clone(euf::solver& ctx) {
|
||||
family_id fid = ctx.get_manager().mk_family_id(symbol("quant"));
|
||||
return alloc(solver, ctx, fid);
|
||||
}
|
||||
|
||||
bool solver::unit_propagate() {
|
||||
|
|
|
@ -53,7 +53,7 @@ namespace q {
|
|||
|
||||
public:
|
||||
|
||||
solver(euf::solver& ctx);
|
||||
solver(euf::solver& ctx, family_id fid);
|
||||
~solver() override {}
|
||||
bool is_external(sat::bool_var v) override { return false; }
|
||||
void get_antecedents(sat::literal l, sat::ext_justification_idx idx, sat::literal_vector& r, bool probing) override {}
|
||||
|
@ -64,7 +64,7 @@ namespace q {
|
|||
std::ostream& display_justification(std::ostream& out, sat::ext_justification_idx idx) const override { UNREACHABLE(); return out; }
|
||||
std::ostream& display_constraint(std::ostream& out, sat::ext_constraint_idx idx) const override { UNREACHABLE(); return out; }
|
||||
void collect_statistics(statistics& st) const override;
|
||||
euf::th_solver* clone(sat::solver* s, euf::solver& ctx) override;
|
||||
euf::th_solver* clone(euf::solver& ctx) override;
|
||||
bool unit_propagate() override;
|
||||
sat::literal internalize(expr* e, bool sign, bool root, bool learned) override;
|
||||
void internalize(expr* e, bool redundant) override { UNREACHABLE(); }
|
||||
|
|
|
@ -201,30 +201,61 @@ namespace euf {
|
|||
}
|
||||
|
||||
sat::literal th_euf_solver::eq_internalize(expr* a, expr* b) {
|
||||
return b_internalize(ctx.mk_eq(a, b));
|
||||
expr_ref eq(ctx.mk_eq(a, b), m);
|
||||
return b_internalize(eq);
|
||||
}
|
||||
|
||||
unsigned th_propagation::get_obj_size(unsigned num_lits, unsigned num_eqs) {
|
||||
return sizeof(th_propagation) + sizeof(sat::literal) * num_lits + sizeof(enode_pair) * num_eqs;
|
||||
}
|
||||
|
||||
th_propagation::th_propagation(sat::literal_vector const& lits, enode_pair_vector const& eqs) {
|
||||
m_num_literals = lits.size();
|
||||
m_num_eqs = eqs.size();
|
||||
th_propagation::th_propagation(unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs) {
|
||||
m_num_literals = n_lits;
|
||||
m_num_eqs = n_eqs;
|
||||
m_literals = reinterpret_cast<literal*>(reinterpret_cast<char*>(this) + sizeof(th_propagation));
|
||||
unsigned i = 0;
|
||||
for (sat::literal lit : lits)
|
||||
m_literals[i++] = lit;
|
||||
m_eqs = reinterpret_cast<enode_pair*>(reinterpret_cast<char*>(this) + sizeof(th_propagation) + sizeof(literal) * m_num_literals);
|
||||
i = 0;
|
||||
for (auto eq : eqs)
|
||||
m_eqs[i++] = eq;
|
||||
for (unsigned i = 0; i < n_lits; ++i)
|
||||
m_literals[i] = lits[i];
|
||||
m_eqs = reinterpret_cast<enode_pair*>(reinterpret_cast<char*>(this) + sizeof(th_propagation) + sizeof(literal) * n_lits);
|
||||
for (unsigned i = 0; i < n_eqs; ++i)
|
||||
m_eqs[i] = eqs[i];
|
||||
}
|
||||
|
||||
th_propagation* th_propagation::mk(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs) {
|
||||
region& r = th.ctx.get_region();
|
||||
void* mem = r.allocate(get_obj_size(lits.size(), eqs.size()));
|
||||
sat::constraint_base::initialize(mem, &th);
|
||||
return new (sat::constraint_base::ptr2mem(mem)) th_propagation(lits, eqs);
|
||||
return mk(th, lits.size(), lits.c_ptr(), eqs.size(), eqs.c_ptr());
|
||||
}
|
||||
|
||||
th_propagation* th_propagation::mk(th_euf_solver& th, unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs) {
|
||||
region& r = th.ctx.get_region();
|
||||
void* mem = r.allocate(get_obj_size(n_lits, n_eqs));
|
||||
sat::constraint_base::initialize(mem, &th);
|
||||
return new (sat::constraint_base::ptr2mem(mem)) th_propagation(n_lits, lits, n_eqs, eqs);
|
||||
}
|
||||
|
||||
th_propagation* th_propagation::mk(th_euf_solver& th, enode_pair_vector const& eqs) {
|
||||
return mk(th, 0, nullptr, eqs.size(), eqs.c_ptr());
|
||||
}
|
||||
|
||||
th_propagation* th_propagation::mk(th_euf_solver& th, sat::literal lit) {
|
||||
return mk(th, 1, &lit, 0, nullptr);
|
||||
}
|
||||
|
||||
th_propagation* th_propagation::mk(th_euf_solver& th, sat::literal lit, euf::enode* x, euf::enode* y) {
|
||||
enode_pair eq(x, y);
|
||||
return mk(th, 1, &lit, 1, &eq);
|
||||
}
|
||||
|
||||
th_propagation* th_propagation::mk(th_euf_solver& th, euf::enode* x, euf::enode* y) {
|
||||
enode_pair eq(x, y);
|
||||
return mk(th, 0, nullptr, 1, &eq);
|
||||
}
|
||||
|
||||
std::ostream& th_propagation::display(std::ostream& out) const {
|
||||
for (auto lit : euf::th_propagation::lits(*this))
|
||||
out << lit << " ";
|
||||
for (auto eq : euf::th_propagation::eqs(*this))
|
||||
out << eq.first->get_expr_id() << " == " << eq.second->get_expr_id() << " ";
|
||||
return out;
|
||||
}
|
||||
|
||||
|
||||
}
|
||||
|
|
|
@ -101,7 +101,7 @@ namespace euf {
|
|||
public:
|
||||
th_solver(ast_manager& m, symbol const& name, euf::theory_id id) : extension(name, id), m(m) {}
|
||||
|
||||
virtual th_solver* clone(sat::solver* s, euf::solver& ctx) = 0;
|
||||
virtual th_solver* clone(euf::solver& ctx) = 0;
|
||||
|
||||
virtual void new_eq_eh(euf::th_eq const& eq) {}
|
||||
|
||||
|
@ -150,7 +150,7 @@ namespace euf {
|
|||
sat::literal eq_internalize(expr* a, expr* b);
|
||||
|
||||
euf::enode* e_internalize(expr* e) { internalize(e, m_is_redundant); return expr2enode(e); }
|
||||
euf::enode* mk_enode(expr* e, bool suppress_args);
|
||||
euf::enode* mk_enode(expr* e, bool suppress_args = false);
|
||||
expr_ref mk_eq(expr* e1, expr* e2);
|
||||
expr_ref mk_var_eq(theory_var v1, theory_var v2) { return mk_eq(var2expr(v1), var2expr(v2)); }
|
||||
|
||||
|
@ -187,14 +187,19 @@ namespace euf {
|
|||
|
||||
|
||||
class th_propagation {
|
||||
unsigned m_num_literals;
|
||||
unsigned m_num_eqs;
|
||||
sat::literal* m_literals;
|
||||
enode_pair* m_eqs;
|
||||
unsigned m_num_literals;
|
||||
unsigned m_num_eqs;
|
||||
sat::literal* m_literals;
|
||||
enode_pair* m_eqs;
|
||||
static unsigned get_obj_size(unsigned num_lits, unsigned num_eqs);
|
||||
th_propagation(sat::literal_vector const& lits, enode_pair_vector const& eqs);
|
||||
th_propagation(unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs);
|
||||
public:
|
||||
static th_propagation* mk(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs);
|
||||
static th_propagation* mk(th_euf_solver& th, unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs);
|
||||
static th_propagation* mk(th_euf_solver& th, enode_pair_vector const& eqs);
|
||||
static th_propagation* mk(th_euf_solver& th, sat::literal lit);
|
||||
static th_propagation* mk(th_euf_solver& th, sat::literal lit, euf::enode* x, euf::enode* y);
|
||||
static th_propagation* mk(th_euf_solver& th, euf::enode* x, euf::enode* y);
|
||||
|
||||
sat::ext_constraint_idx to_index() const {
|
||||
return sat::constraint_base::mem2base(this);
|
||||
|
@ -203,18 +208,24 @@ namespace euf {
|
|||
return *reinterpret_cast<th_propagation*>(sat::constraint_base::from_index(idx)->mem());
|
||||
}
|
||||
|
||||
sat::extension& ext() const {
|
||||
return *sat::constraint_base::to_extension(to_index());
|
||||
}
|
||||
|
||||
std::ostream& display(std::ostream& out) const;
|
||||
|
||||
class lits {
|
||||
th_propagation& th;
|
||||
th_propagation const& th;
|
||||
public:
|
||||
lits(th_propagation& th) : th(th) {}
|
||||
lits(th_propagation const& th) : th(th) {}
|
||||
sat::literal const* begin() const { return th.m_literals; }
|
||||
sat::literal const* end() const { return th.m_literals + th.m_num_literals; }
|
||||
};
|
||||
|
||||
class eqs {
|
||||
th_propagation& th;
|
||||
th_propagation const& th;
|
||||
public:
|
||||
eqs(th_propagation& th) : th(th) {}
|
||||
eqs(th_propagation const& th) : th(th) {}
|
||||
enode_pair const* begin() const { return th.m_eqs; }
|
||||
enode_pair const* end() const { return th.m_eqs + th.m_num_eqs; }
|
||||
};
|
||||
|
|
|
@ -146,12 +146,10 @@ namespace user {
|
|||
return display_justification(out, idx);
|
||||
}
|
||||
|
||||
euf::th_solver* solver::clone(sat::solver* dst_s, euf::solver& dst_ctx) {
|
||||
euf::th_solver* solver::clone(euf::solver& dst_ctx) {
|
||||
auto* result = alloc(solver, dst_ctx);
|
||||
result->set_solver(dst_s);
|
||||
ast_translation tr(m, dst_ctx.get_manager(), false);
|
||||
for (unsigned i = 0; i < get_num_vars(); ++i)
|
||||
result->add_expr(tr(var2expr(i)));
|
||||
for (unsigned i = 0; i < get_num_vars(); ++i)
|
||||
result->add_expr(ctx.copy(dst_ctx, var2enode(i))->get_expr());
|
||||
return result;
|
||||
}
|
||||
|
||||
|
|
|
@ -124,7 +124,7 @@ namespace user {
|
|||
std::ostream& display(std::ostream& out) const override;
|
||||
std::ostream& display_justification(std::ostream& out, sat::ext_justification_idx idx) const override;
|
||||
std::ostream& display_constraint(std::ostream& out, sat::ext_constraint_idx idx) const override;
|
||||
euf::th_solver* clone(sat::solver* s, euf::solver& ctx) override;
|
||||
euf::th_solver* clone(euf::solver& ctx) override;
|
||||
|
||||
};
|
||||
};
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue