mirror of
https://github.com/Z3Prover/z3
synced 2025-04-24 09:35:32 +00:00
merging master to unit_prop_on_monomials
This commit is contained in:
parent
a297a2b25c
commit
7de06c4350
19 changed files with 333 additions and 375 deletions
|
@ -71,8 +71,6 @@ def_module_params(module_name='smt',
|
|||
('arith.nl.grobner_row_length_limit', UINT, 10, 'row is disregarded by the heuristic if its length is longer than the value'),
|
||||
('arith.nl.grobner_frequency', UINT, 4, 'grobner\'s call frequency'),
|
||||
('arith.nl.grobner', BOOL, True, 'run grobner\'s basis heuristic'),
|
||||
('arith.nl.use_lemmas_in_unit_prop', BOOL, False, 'use lemmas in monomial unit propagation'),
|
||||
('arith.nl.throttle_unit_prop', BOOL, True, 'unit propogate a monomial only once per scope'),
|
||||
('arith.nl.grobner_eqs_growth', UINT, 10, 'grobner\'s number of equalities growth '),
|
||||
('arith.nl.grobner_expr_size_growth', UINT, 2, 'grobner\'s maximum expr size growth'),
|
||||
('arith.nl.grobner_expr_degree_growth', UINT, 2, 'grobner\'s maximum expr degree growth'),
|
||||
|
|
|
@ -90,14 +90,14 @@ namespace smt {
|
|||
return proof_ref(m);
|
||||
}
|
||||
|
||||
void clause_proof::add(clause& c) {
|
||||
void clause_proof::add(clause& c, literal_buffer const* simp_lits) {
|
||||
if (!is_enabled())
|
||||
return;
|
||||
justification* j = c.get_justification();
|
||||
auto st = kind2st(c.get_kind());
|
||||
auto pr = justification2proof(st, j);
|
||||
CTRACE("mk_clause", pr.get(), tout << mk_bounded_pp(pr, m, 4) << "\n";);
|
||||
update(c, st, pr);
|
||||
update(c, st, pr, simp_lits);
|
||||
}
|
||||
|
||||
void clause_proof::add(unsigned n, literal const* lits, clause_kind k, justification* j) {
|
||||
|
@ -137,12 +137,15 @@ namespace smt {
|
|||
update(st, m_lits, pr);
|
||||
}
|
||||
|
||||
void clause_proof::add(literal lit1, literal lit2, clause_kind k, justification* j) {
|
||||
void clause_proof::add(literal lit1, literal lit2, clause_kind k, justification* j, literal_buffer const* simp_lits) {
|
||||
if (!is_enabled())
|
||||
return;
|
||||
m_lits.reset();
|
||||
m_lits.push_back(ctx.literal2expr(lit1));
|
||||
m_lits.push_back(ctx.literal2expr(lit2));
|
||||
if (simp_lits)
|
||||
for (auto lit : *simp_lits)
|
||||
m_lits.push_back(ctx.literal2expr(~lit));
|
||||
auto st = kind2st(k);
|
||||
auto pr = justification2proof(st, j);
|
||||
update(st, m_lits, pr);
|
||||
|
@ -160,7 +163,7 @@ namespace smt {
|
|||
}
|
||||
|
||||
void clause_proof::del(clause& c) {
|
||||
update(c, status::deleted, justification2proof(status::deleted, nullptr));
|
||||
update(c, status::deleted, justification2proof(status::deleted, nullptr), nullptr);
|
||||
}
|
||||
|
||||
std::ostream& clause_proof::display_literals(std::ostream& out, expr_ref_vector const& v) {
|
||||
|
@ -190,7 +193,9 @@ namespace smt {
|
|||
if (ctx.get_fparams().m_clause_proof)
|
||||
m_trail.push_back(info(st, v, p));
|
||||
if (m_on_clause_eh)
|
||||
m_on_clause_eh(m_on_clause_ctx, p, 0, nullptr, v.size(), v.data());
|
||||
m_on_clause_eh(m_on_clause_ctx, p, 0, nullptr, v.size(), v.data());
|
||||
static unsigned s_count = 0;
|
||||
|
||||
if (m_has_log) {
|
||||
init_pp_out();
|
||||
auto& out = *m_pp_out;
|
||||
|
@ -220,12 +225,15 @@ namespace smt {
|
|||
}
|
||||
}
|
||||
|
||||
void clause_proof::update(clause& c, status st, proof* p) {
|
||||
void clause_proof::update(clause& c, status st, proof* p, literal_buffer const* simp_lits) {
|
||||
if (!is_enabled())
|
||||
return;
|
||||
m_lits.reset();
|
||||
for (literal lit : c)
|
||||
m_lits.push_back(ctx.literal2expr(lit));
|
||||
m_lits.push_back(ctx.literal2expr(lit));
|
||||
if (simp_lits)
|
||||
for (auto lit : *simp_lits)
|
||||
m_lits.push_back(ctx.literal2expr(~lit));
|
||||
update(st, m_lits, p);
|
||||
}
|
||||
|
||||
|
|
|
@ -68,7 +68,7 @@ namespace smt {
|
|||
void init_pp_out();
|
||||
|
||||
void update(status st, expr_ref_vector& v, proof* p);
|
||||
void update(clause& c, status st, proof* p);
|
||||
void update(clause& c, status st, proof* p, literal_buffer const* simp_lits);
|
||||
status kind2st(clause_kind k);
|
||||
proof_ref justification2proof(status st, justification* j);
|
||||
void log(status st, proof* p);
|
||||
|
@ -79,8 +79,8 @@ namespace smt {
|
|||
clause_proof(context& ctx);
|
||||
void shrink(clause& c, unsigned new_size);
|
||||
void add(literal lit, clause_kind k, justification* j);
|
||||
void add(literal lit1, literal lit2, clause_kind k, justification* j);
|
||||
void add(clause& c);
|
||||
void add(literal lit1, literal lit2, clause_kind k, justification* j, literal_buffer const* simp_lits = nullptr);
|
||||
void add(clause& c, literal_buffer const* simp_lits = nullptr);
|
||||
void add(unsigned n, literal const* lits, clause_kind k, justification* j);
|
||||
void propagate(literal lit, justification const& j, literal_vector const& ante);
|
||||
void del(clause& c);
|
||||
|
|
|
@ -601,6 +601,7 @@ namespace smt {
|
|||
|
||||
finalize_resolve(conflict, not_l);
|
||||
|
||||
|
||||
return true;
|
||||
}
|
||||
|
||||
|
|
|
@ -1378,12 +1378,12 @@ namespace smt {
|
|||
clause * context::mk_clause(unsigned num_lits, literal * lits, justification * j, clause_kind k, clause_del_eh * del_eh) {
|
||||
TRACE("mk_clause", display_literals_verbose(tout << "creating clause: " << literal_vector(num_lits, lits) << "\n", num_lits, lits) << "\n";);
|
||||
m_clause_proof.add(num_lits, lits, k, j);
|
||||
literal_buffer simp_lits;
|
||||
switch (k) {
|
||||
case CLS_TH_AXIOM:
|
||||
dump_axiom(num_lits, lits);
|
||||
Z3_fallthrough;
|
||||
case CLS_AUX: {
|
||||
literal_buffer simp_lits;
|
||||
if (m_searching)
|
||||
dump_lemma(num_lits, lits);
|
||||
if (!simplify_aux_clause_literals(num_lits, lits, simp_lits)) {
|
||||
|
@ -1451,7 +1451,7 @@ namespace smt {
|
|||
else if (get_assignment(l2) == l_false) {
|
||||
assign(l1, b_justification(~l2));
|
||||
}
|
||||
m_clause_proof.add(l1, l2, k, j);
|
||||
m_clause_proof.add(l1, l2, k, j, &simp_lits);
|
||||
m_stats.m_num_mk_bin_clause++;
|
||||
return nullptr;
|
||||
}
|
||||
|
@ -1464,7 +1464,7 @@ namespace smt {
|
|||
bool reinit = save_atoms;
|
||||
SASSERT(!lemma || j == 0 || !j->in_region());
|
||||
clause * cls = clause::mk(m, num_lits, lits, k, j, del_eh, save_atoms, m_bool_var2expr.data());
|
||||
m_clause_proof.add(*cls);
|
||||
m_clause_proof.add(*cls, &simp_lits);
|
||||
if (lemma) {
|
||||
cls->set_activity(activity);
|
||||
if (k == CLS_LEARNED) {
|
||||
|
|
|
@ -1535,7 +1535,8 @@ namespace smt {
|
|||
m_stats.m_max_min++;
|
||||
unsigned best_efforts = 0;
|
||||
bool inc = false;
|
||||
|
||||
|
||||
|
||||
SASSERT(!maintain_integrality || valid_assignment());
|
||||
SASSERT(satisfy_bounds());
|
||||
|
||||
|
|
|
@ -765,10 +765,8 @@ typename theory_arith<Ext>::numeral theory_arith<Ext>::get_monomial_fixed_var_pr
|
|||
template<typename Ext>
|
||||
expr * theory_arith<Ext>::get_monomial_non_fixed_var(expr * m) const {
|
||||
SASSERT(is_pure_monomial(m));
|
||||
for (unsigned i = 0; i < to_app(m)->get_num_args(); i++) {
|
||||
expr * arg = to_app(m)->get_arg(i);
|
||||
theory_var _var = expr2var(arg);
|
||||
if (!is_fixed(_var))
|
||||
for (expr* arg : *to_app(m)) {
|
||||
if (!is_fixed(expr2var(arg)))
|
||||
return arg;
|
||||
}
|
||||
return nullptr;
|
||||
|
@ -780,7 +778,7 @@ expr * theory_arith<Ext>::get_monomial_non_fixed_var(expr * m) const {
|
|||
*/
|
||||
template<typename Ext>
|
||||
bool theory_arith<Ext>::propagate_linear_monomial(theory_var v) {
|
||||
TRACE("non_linear", tout << "checking whether v" << v << " became linear...\n";);
|
||||
TRACE("non_linear_verbose", tout << "checking whether v" << v << " became linear...\n";);
|
||||
if (m_data[v].m_nl_propagated)
|
||||
return false; // already propagated this monomial.
|
||||
expr * m = var2expr(v);
|
||||
|
@ -819,6 +817,11 @@ bool theory_arith<Ext>::propagate_linear_monomial(theory_var v) {
|
|||
ctx.mark_as_relevant(rhs);
|
||||
}
|
||||
TRACE("non_linear_bug", tout << "enode: " << ctx.get_enode(rhs) << " enode_id: " << ctx.get_enode(rhs)->get_owner_id() << "\n";);
|
||||
IF_VERBOSE(3,
|
||||
for (auto* arg : *to_app(m))
|
||||
if (is_fixed(expr2var(arg)))
|
||||
verbose_stream() << mk_pp(arg, get_manager()) << " = " << -k << "\n");
|
||||
|
||||
theory_var new_v = expr2var(rhs);
|
||||
SASSERT(new_v != null_theory_var);
|
||||
new_lower = alloc(derived_bound, new_v, inf_numeral(0), B_LOWER);
|
||||
|
@ -906,7 +909,7 @@ bool theory_arith<Ext>::propagate_linear_monomials() {
|
|||
return false;
|
||||
if (!reflection_enabled())
|
||||
return false;
|
||||
TRACE("non_linear", tout << "propagating linear monomials...\n";);
|
||||
TRACE("non_linear_verbose", tout << "propagating linear monomials...\n";);
|
||||
bool p = false;
|
||||
// CMW: m_nl_monomials can grow during this loop, so
|
||||
// don't use iterators.
|
||||
|
|
|
@ -264,7 +264,7 @@ class theory_lra::imp {
|
|||
|
||||
void ensure_nla() {
|
||||
if (!m_nla) {
|
||||
m_nla = alloc(nla::solver, *m_solver.get(), ctx().get_params(), m.limit(), m_implied_bounds);
|
||||
m_nla = alloc(nla::solver, *m_solver.get(), ctx().get_params(), m.limit());
|
||||
for (auto const& _s : m_scopes) {
|
||||
(void)_s;
|
||||
m_nla->push();
|
||||
|
@ -1528,14 +1528,12 @@ public:
|
|||
unsigned old_sz = m_assume_eq_candidates.size();
|
||||
unsigned num_candidates = 0;
|
||||
int start = ctx().get_random_value();
|
||||
unsigned num_relevant = 0;
|
||||
for (theory_var i = 0; i < sz; ++i) {
|
||||
theory_var v = (i + start) % sz;
|
||||
enode* n1 = get_enode(v);
|
||||
if (!th.is_relevant_and_shared(n1)) {
|
||||
continue;
|
||||
}
|
||||
++num_relevant;
|
||||
ensure_column(v);
|
||||
if (!is_registered_var(v))
|
||||
continue;
|
||||
|
@ -1553,7 +1551,7 @@ public:
|
|||
num_candidates++;
|
||||
}
|
||||
}
|
||||
|
||||
|
||||
if (num_candidates > 0) {
|
||||
ctx().push_trail(restore_vector(m_assume_eq_candidates, old_sz));
|
||||
}
|
||||
|
@ -1605,8 +1603,7 @@ public:
|
|||
case l_true:
|
||||
return FC_DONE;
|
||||
case l_false:
|
||||
for (const nla::lemma & l : m_nla->lemmas())
|
||||
false_case_of_check_nla(l);
|
||||
add_lemmas();
|
||||
return FC_CONTINUE;
|
||||
case l_undef:
|
||||
return FC_GIVEUP;
|
||||
|
@ -1803,8 +1800,7 @@ public:
|
|||
if (!m_nla)
|
||||
return true;
|
||||
m_nla->check_bounded_divisions();
|
||||
for (auto & lemma : m_nla->lemmas())
|
||||
false_case_of_check_nla(lemma);
|
||||
add_lemmas();
|
||||
return m_nla->lemmas().empty();
|
||||
}
|
||||
|
||||
|
@ -2003,7 +1999,7 @@ public:
|
|||
// create term >= 0 (or term <= 0)
|
||||
atom = mk_bound(ineq.term(), ineq.rs(), is_lower);
|
||||
return literal(ctx().get_bool_var(atom), pos);
|
||||
}
|
||||
}
|
||||
|
||||
void false_case_of_check_nla(const nla::lemma & l) {
|
||||
m_lemma = l; //todo avoid the copy
|
||||
|
@ -2024,14 +2020,11 @@ public:
|
|||
|
||||
final_check_status check_nla_continue() {
|
||||
m_a1 = nullptr; m_a2 = nullptr;
|
||||
lbool r = m_nla->check(m_nla_literals);
|
||||
lbool r = m_nla->check();
|
||||
|
||||
switch (r) {
|
||||
case l_false:
|
||||
for (const nla::ineq& i : m_nla_literals)
|
||||
assume_literal(i);
|
||||
for (const nla::lemma & l : m_nla->lemmas())
|
||||
false_case_of_check_nla(l);
|
||||
add_lemmas();
|
||||
return FC_CONTINUE;
|
||||
case l_true:
|
||||
return assume_eqs()? FC_CONTINUE: FC_DONE;
|
||||
|
@ -2120,6 +2113,8 @@ public:
|
|||
bool propagate_core() {
|
||||
m_model_is_initialized = false;
|
||||
flush_bound_axioms();
|
||||
// disabled in master:
|
||||
propagate_nla();
|
||||
if (!can_propagate_core())
|
||||
return false;
|
||||
m_new_def = false;
|
||||
|
@ -2151,7 +2146,6 @@ public:
|
|||
break;
|
||||
case l_true:
|
||||
propagate_basic_bounds();
|
||||
propagate_bounds_with_nlp();
|
||||
propagate_bounds_with_lp_solver();
|
||||
break;
|
||||
case l_undef:
|
||||
|
@ -2161,6 +2155,47 @@ public:
|
|||
return true;
|
||||
}
|
||||
|
||||
void propagate_nla() {
|
||||
if (m_nla) {
|
||||
m_nla->propagate();
|
||||
add_lemmas();
|
||||
add_equalities();
|
||||
}
|
||||
}
|
||||
|
||||
void add_equalities() {
|
||||
if (!propagate_eqs())
|
||||
return;
|
||||
for (auto const& [v,k,e] : m_nla->fixed_equalities())
|
||||
add_equality(v, k, e);
|
||||
for (auto const& [i,j,e] : m_nla->equalities())
|
||||
add_eq(i,j,e,false);
|
||||
}
|
||||
|
||||
void add_equality(lpvar j, rational const& k, lp::explanation const& exp) {
|
||||
//verbose_stream() << "equality " << j << " " << k << "\n";
|
||||
TRACE("arith", tout << "equality " << j << " " << k << "\n");
|
||||
theory_var v;
|
||||
if (k == 1)
|
||||
v = m_one_var;
|
||||
else if (k == 0)
|
||||
v = m_zero_var;
|
||||
else if (!m_value2var.find(k, v))
|
||||
return;
|
||||
theory_var w = lp().local_to_external(j);
|
||||
if (w < 0)
|
||||
return;
|
||||
lpvar i = register_theory_var_in_lar_solver(v);
|
||||
add_eq(i, j, exp, true);
|
||||
}
|
||||
|
||||
void add_lemmas() {
|
||||
for (const nla::ineq& i : m_nla->literals())
|
||||
assume_literal(i);
|
||||
for (const nla::lemma & l : m_nla->lemmas())
|
||||
false_case_of_check_nla(l);
|
||||
}
|
||||
|
||||
bool should_propagate() const {
|
||||
return bound_prop_mode::BP_NONE != propagation_mode();
|
||||
}
|
||||
|
@ -2173,50 +2208,33 @@ public:
|
|||
set_evidence(j, m_core, m_eqs);
|
||||
m_explanation.add_pair(j, v);
|
||||
}
|
||||
|
||||
void propagate_bounds_with_lp_solver() {
|
||||
if (!should_propagate())
|
||||
return;
|
||||
|
||||
m_bp.init();
|
||||
lp().propagate_bounds_for_touched_rows(m_bp);
|
||||
|
||||
if (!m.inc())
|
||||
return;
|
||||
|
||||
void finish_bound_propagation() {
|
||||
if (is_infeasible()) {
|
||||
get_infeasibility_explanation_and_set_conflict();
|
||||
// verbose_stream() << "unsat\n";
|
||||
}
|
||||
else {
|
||||
for (auto &ib : m_bp.ibounds()) {
|
||||
unsigned count = 0, prop = 0;
|
||||
for (auto& ib : m_bp.ibounds()) {
|
||||
m.inc();
|
||||
if (ctx().inconsistent())
|
||||
break;
|
||||
propagate_lp_solver_bound(ib);
|
||||
++prop;
|
||||
count += propagate_lp_solver_bound(ib);
|
||||
}
|
||||
}
|
||||
}
|
||||
|
||||
void propagate_bounds_with_lp_solver() {
|
||||
if (!should_propagate())
|
||||
return;
|
||||
m_bp.init();
|
||||
lp().propagate_bounds_for_touched_rows(m_bp);
|
||||
|
||||
if (m.inc())
|
||||
finish_bound_propagation();
|
||||
}
|
||||
|
||||
void propagate_bounds_for_monomials() {
|
||||
m_nla->propagate_bounds_for_touched_monomials();
|
||||
for (const auto & l : m_nla->lemmas())
|
||||
false_case_of_check_nla(l);
|
||||
}
|
||||
|
||||
void propagate_bounds_with_nlp() {
|
||||
if (!m_nla)
|
||||
return;
|
||||
if (is_infeasible() || !should_propagate())
|
||||
return;
|
||||
|
||||
propagate_bounds_for_monomials();
|
||||
|
||||
if (m.inc())
|
||||
finish_bound_propagation();
|
||||
}
|
||||
|
||||
bool bound_is_interesting(unsigned vi, lp::lconstraint_kind kind, const rational & bval) const {
|
||||
theory_var v = lp().local_to_external(vi);
|
||||
if (v == null_theory_var)
|
||||
|
@ -3161,8 +3179,7 @@ public:
|
|||
std::function<expr*(void)> fn = [&]() { return m.mk_eq(x->get_expr(), y->get_expr()); };
|
||||
scoped_trace_stream _sts(th, fn);
|
||||
|
||||
|
||||
// SASSERT(validate_eq(x, y));
|
||||
//VERIFY(validate_eq(x, y));
|
||||
ctx().assign_eq(x, y, eq_justification(js));
|
||||
}
|
||||
|
||||
|
@ -3206,12 +3223,11 @@ public:
|
|||
}
|
||||
|
||||
lp::explanation m_explanation;
|
||||
vector<nla::ineq> m_nla_literals;
|
||||
literal_vector m_core;
|
||||
svector<enode_pair> m_eqs;
|
||||
vector<parameter> m_params;
|
||||
|
||||
void reset_evidence() {
|
||||
void reset_evidence() {
|
||||
m_core.reset();
|
||||
m_eqs.reset();
|
||||
m_params.reset();
|
||||
|
@ -3278,6 +3294,7 @@ public:
|
|||
display(tout << "is-conflict: " << is_conflict << "\n"););
|
||||
for (auto ev : m_explanation)
|
||||
set_evidence(ev.ci(), m_core, m_eqs);
|
||||
|
||||
|
||||
// SASSERT(validate_conflict(m_core, m_eqs));
|
||||
if (is_conflict) {
|
||||
|
@ -3533,6 +3550,8 @@ public:
|
|||
lbool r = nctx.check();
|
||||
if (r == l_true) {
|
||||
nctx.display_asserted_formulas(std::cout);
|
||||
std::cout.flush();
|
||||
std::cout.flush();
|
||||
}
|
||||
return l_true != r;
|
||||
}
|
||||
|
@ -3882,6 +3901,7 @@ public:
|
|||
IF_VERBOSE(1, verbose_stream() << enode_pp(n, ctx()) << " evaluates to " << r2 << " but arith solver has " << r1 << "\n");
|
||||
}
|
||||
}
|
||||
|
||||
};
|
||||
|
||||
theory_lra::theory_lra(context& ctx):
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue