3
0
Fork 0
mirror of https://github.com/Z3Prover/z3 synced 2025-10-09 09:21:56 +00:00

Merge branch 'master' into polysat

This commit is contained in:
Jakob Rath 2022-07-01 16:11:17 +02:00
commit e5e79c1d4b
398 changed files with 24548 additions and 4983 deletions

View file

@ -112,6 +112,29 @@ static void read_clause(Buffer & in, std::ostream& err, sat::literal_vector & li
}
}
template<typename Buffer>
static void read_pragma(Buffer & in, std::ostream& err, std::string& p, sat::proof_hint& h) {
skip_whitespace(in);
if (*in != 'p')
return;
++in;
while (*in == ' ')
++in;
while (true) {
if (*in == EOF)
break;
if (*in == '\n') {
++in;
break;
}
p.push_back(*in);
++in;
}
if (!p.empty())
h.from_string(p);
}
template<typename Buffer>
static bool parse_dimacs_core(Buffer & in, std::ostream& err, sat::solver & solver) {
sat::literal_vector lits;
@ -156,7 +179,9 @@ namespace dimacs {
sat::status_pp pp(r.m_status, p.th);
switch (r.m_tag) {
case drat_record::tag_t::is_clause:
return out << pp << " " << r.m_lits << " 0\n";
if (!r.m_pragma.empty())
return out << pp << " " << r.m_lits << " 0 p " << r.m_pragma << "\n";
return out << pp << " " << r.m_lits << " 0\n";
case drat_record::tag_t::is_node:
return out << "e " << r.m_node_id << " " << r.m_name << " " << r.m_args << "0\n";
case drat_record::tag_t::is_sort:
@ -280,6 +305,8 @@ namespace dimacs {
try {
loop:
skip_whitespace(in);
m_record.m_pragma.clear();
m_record.m_hint.reset();
switch (*in) {
case EOF:
return false;
@ -304,6 +331,7 @@ namespace dimacs {
theory_id = read_theory_id();
skip_whitespace(in);
read_clause(in, err, m_record.m_lits);
read_pragma(in, err, m_record.m_pragma, m_record.m_hint);
m_record.m_tag = drat_record::tag_t::is_clause;
m_record.m_status = sat::status::th(false, theory_id);
break;

View file

@ -59,10 +59,12 @@ namespace dimacs {
// a node populates m_node_id, m_name, m_args
// a bool def populates m_node_id and one element in m_args
sat::literal_vector m_lits;
sat::status m_status{ sat::status::redundant() };
unsigned m_node_id{ 0 };
sat::status m_status = sat::status::redundant();
unsigned m_node_id = 0;
std::string m_name;
unsigned_vector m_args;
std::string m_pragma;
sat::proof_hint m_hint;
};
struct drat_pp {

View file

@ -818,11 +818,12 @@ namespace sat {
lbool r = s.check();
IF_VERBOSE(10, verbose_stream() << "check: " << r << "\n");
if (r == l_true) {
std::sort(vars.begin(), vars.end());
s.display(std::cout);
for (auto v : vars) std::cout << v << " := " << s.get_model()[v] << "\n";
std::string line;
std::getline(std::cin, line);
IF_VERBOSE(0,
std::sort(vars.begin(), vars.end());
s.display(verbose_stream());
for (auto v : vars) verbose_stream() << v << " := " << s.get_model()[v] << "\n";
);
UNREACHABLE();
}
}
};

View file

@ -197,6 +197,7 @@ namespace sat {
m_drat = (m_drat_check_unsat || m_drat_file.is_non_empty_string() || m_drat_check_sat) && p.threads() == 1;
m_drat_binary = p.drat_binary();
m_drat_activity = p.drat_activity();
m_drup_trim = p.drup_trim();
m_dyn_sub_res = p.dyn_sub_res();
// Parameters used in Liang, Ganesh, Poupart, Czarnecki AAAI 2016.

View file

@ -179,6 +179,7 @@ namespace sat {
symbol m_drat_file;
bool m_drat_check_unsat;
bool m_drat_check_sat;
bool m_drup_trim;
bool m_drat_activity;
bool m_card_solver;

View file

@ -75,8 +75,7 @@ namespace sat {
IF_VERBOSE(0,
verbose_stream() << "not validated: " << clause << "\n";
s.display(verbose_stream()););
std::string line;
std::getline(std::cin, line);
UNREACHABLE();
}
}
};

File diff suppressed because it is too large Load diff

View file

@ -62,10 +62,10 @@ namespace sat {
class drat {
struct stats {
unsigned m_num_drup { 0 };
unsigned m_num_drat { 0 };
unsigned m_num_add { 0 };
unsigned m_num_del { 0 };
unsigned m_num_drup = 0;
unsigned m_num_drat = 0;
unsigned m_num_add = 0;
unsigned m_num_del = 0;
};
struct watched_clause {
clause* m_clause;
@ -77,16 +77,19 @@ namespace sat {
typedef svector<unsigned> watch;
solver& s;
clause_allocator m_alloc;
std::ostream* m_out;
std::ostream* m_bout;
ptr_vector<clause> m_proof;
svector<status> m_status;
literal_vector m_units;
std::ostream* m_out = nullptr;
std::ostream* m_bout = nullptr;
svector<std::pair<clause&, status>> m_proof;
svector<std::pair<literal, clause*>> m_units;
vector<watch> m_watches;
svector<lbool> m_assignment;
vector<std::string> m_theory;
bool m_inconsistent;
bool m_check_unsat, m_check_sat, m_check, m_activity;
bool m_inconsistent = false;
bool m_check_unsat = false;
bool m_check_sat = false;
bool m_check = false;
bool m_activity = false;
bool m_trim = false;
stats m_stats;
void dump_activity();
@ -102,9 +105,9 @@ namespace sat {
status get_status(bool learned) const;
void declare(literal l);
void assign(literal l);
void assign(literal l, clause* c);
void propagate(literal l);
void assign_propagate(literal l);
void assign_propagate(literal l, clause* c);
void del_watch(clause& c, literal l);
bool is_drup(unsigned n, literal const* c);
bool is_drat(unsigned n, literal const* c);
@ -115,6 +118,10 @@ namespace sat {
void validate_propagation() const;
bool match(unsigned n, literal const* lits, clause const& c) const;
clause& mk_clause(clause& c);
clause& mk_clause(unsigned n, literal const* lits, bool is_learned);
public:
drat(solver& s);
@ -165,9 +172,12 @@ namespace sat {
void collect_statistics(statistics& st) const;
bool inconsistent() const { return m_inconsistent; }
literal_vector const& units() { return m_units; }
svector<std::pair<literal, clause*>> const& units() { return m_units; }
bool is_drup(unsigned n, literal const* c, literal_vector& units);
solver& get_solver() { return s; }
svector<std::pair<clause&, status>> trim();
};
}

View file

@ -91,8 +91,10 @@ namespace sat {
virtual double get_reward(literal l, ext_constraint_idx idx, literal_occs_fun& occs) const { return 0; }
virtual void get_antecedents(literal l, ext_justification_idx idx, literal_vector & r, bool probing) = 0;
virtual bool is_extended_binary(ext_justification_idx idx, literal_vector & r) { return false; }
virtual void asserted(literal l) {};
virtual void set_eliminated(bool_var v) {};
virtual bool decide(bool_var& var, lbool& phase) { return false; }
virtual bool get_case_split(bool_var& var, lbool& phase) { return false; }
virtual void asserted(literal l) {}
virtual void set_eliminated(bool_var v) {}
virtual check_result check() = 0;
virtual lbool resolve_conflict() { return l_undef; } // stores result in sat::solver::m_lemma
virtual void push() = 0;

View file

@ -398,7 +398,6 @@ namespace sat {
else
m_learned[j++] = &c;
}
std::cout << "gc: " << to_gc.size() << " " << m_learned.size() << " -> " << j << "\n";
SASSERT(m_learned.size() - j == to_gc.size());
m_learned.shrink(j);
}

View file

@ -50,6 +50,7 @@ def_module_params('sat',
('drat.binary', BOOL, False, 'use Binary DRAT output format'),
('drat.check_unsat', BOOL, False, 'build up internal proof and check'),
('drat.check_sat', BOOL, False, 'build up internal trace, check satisfying model'),
('drup.trim', BOOL, False, 'build and trim drup proof'),
('drat.activity', BOOL, False, 'dump variable activities'),
('cardinality.solver', BOOL, True, 'use cardinality solver'),
('pb.solver', SYMBOL, 'solver', 'method for handling Pseudo-Boolean constraints: circuit (arithmetical circuit), sorting (sorting circuit), totalizer (use totalizer encoding), binary_merge, segmented, solver (use native solver)'),

View file

@ -1831,8 +1831,8 @@ namespace sat {
if (not_l == l2)
continue;
if ((~l2).index() >= m_visited.size()) {
s.display(std::cout << l2 << " " << s.num_vars() << " " << m_visited.size() << "\n");
exit(0);
//s.display(std::cout << l2 << " " << s.num_vars() << " " << m_visited.size() << "\n");
UNREACHABLE();
}
if (m_visited[(~l2).index()]) {
res = false;

View file

@ -601,9 +601,8 @@ namespace sat {
else {
m_clauses.push_back(r);
}
if (m_config.m_drat) {
if (m_config.m_drat)
m_drat.add(*r, st);
}
for (literal l : *r) {
m_touched[l.var()] = m_touch_index;
}
@ -947,8 +946,8 @@ namespace sat {
if (j.level() == 0) {
if (m_config.m_drat)
drat_log_unit(l, j);
j = justification(0); // erase justification for level 0
if (!m_config.m_drup_trim)
j = justification(0); // erase justification for level 0
}
else {
VERIFY(!at_base_lvl());
@ -1663,49 +1662,66 @@ namespace sat {
return null_bool_var;
}
bool solver::decide() {
bool_var next = next_var();
if (next == null_bool_var)
return false;
push();
m_stats.m_decision++;
bool solver::guess(bool_var next) {
lbool lphase = m_ext ? m_ext->get_phase(next) : l_undef;
bool phase = lphase == l_true;
if (lphase == l_undef) {
switch (m_config.m_phase) {
if (lphase != l_undef)
return lphase == l_true;
switch (m_config.m_phase) {
case PS_ALWAYS_TRUE:
phase = true;
break;
return true;
case PS_ALWAYS_FALSE:
phase = false;
break;
return false;
case PS_BASIC_CACHING:
phase = m_phase[next];
break;
return m_phase[next];
case PS_FROZEN:
phase = m_best_phase[next];
break;
return m_best_phase[next];
case PS_SAT_CACHING:
if (m_search_state == s_unsat) {
phase = m_phase[next];
}
else {
phase = m_best_phase[next];
}
break;
if (m_search_state == s_unsat)
return m_phase[next];
return m_best_phase[next];
case PS_RANDOM:
phase = (m_rand() % 2) == 0;
break;
return (m_rand() % 2) == 0;
default:
UNREACHABLE();
phase = false;
break;
}
return false;
}
}
literal next_lit(next, !phase);
bool solver::decide() {
bool_var next;
lbool phase = l_undef;
bool is_pos;
bool used_queue = false;
if (!m_ext || !m_ext->get_case_split(next, phase)) {
used_queue = true;
next = next_var();
if (next == null_bool_var)
return false;
}
push();
m_stats.m_decision++;
if (phase == l_undef)
phase = guess(next) ? l_true: l_false;
literal next_lit(next, false);
if (m_ext && m_ext->decide(next, phase)) {
if (used_queue)
m_case_split_queue.unassign_var_eh(next);
next_lit = literal(next, false);
}
if (phase == l_undef)
is_pos = guess(next);
else
is_pos = phase == l_true;
if (!is_pos)
next_lit.neg();
TRACE("sat_decide", tout << scope_lvl() << ": next-case-split: " << next_lit << "\n";);
assign_scoped(next_lit);
return true;
@ -2014,17 +2030,18 @@ namespace sat {
sort_watch_lits();
CASSERT("sat_simplify_bug", check_invariant());
m_probing();
CASSERT("sat_missed_prop", check_missed_propagation());
CASSERT("sat_simplify_bug", check_invariant());
m_asymm_branch(false);
CASSERT("sat_missed_prop", check_missed_propagation());
CASSERT("sat_simplify_bug", check_invariant());
if (m_ext) {
m_ext->clauses_modifed();
m_ext->simplify();
}
m_probing();
CASSERT("sat_missed_prop", check_missed_propagation());
CASSERT("sat_simplify_bug", check_invariant());
m_asymm_branch(false);
if (m_config.m_lookahead_simplify && !m_ext) {
lookahead lh(*this);
lh.simplify(true);

View file

@ -288,7 +288,7 @@ namespace sat {
inline clause_allocator& cls_allocator() { return m_cls_allocator[m_cls_allocator_idx]; }
inline clause_allocator const& cls_allocator() const { return m_cls_allocator[m_cls_allocator_idx]; }
inline clause * alloc_clause(unsigned num_lits, literal const * lits, bool learned) { return cls_allocator().mk_clause(num_lits, lits, learned); }
inline void dealloc_clause(clause* c) { cls_allocator().del_clause(c); }
inline void dealloc_clause(clause* c) { cls_allocator().del_clause(c); }
struct cmp_activity;
void defrag_clauses();
bool should_defrag();
@ -541,6 +541,7 @@ namespace sat {
unsigned m_next_simplify { 0 };
bool m_simplify_enabled { true };
bool m_restart_enabled { true };
bool guess(bool_var next);
bool decide();
bool_var next_var();
lbool bounded_search();

View file

@ -698,6 +698,33 @@ public:
private:
lbool check_uninterpreted() {
func_decl_ref_vector funs(m);
m_goal2sat.get_interpreted_funs(funs);
if (!funs.empty()) {
m_has_uninterpreted = true;
std::stringstream strm;
strm << "(sat.giveup interpreted functions sent to SAT solver " << funs <<")";
TRACE("sat", tout << strm.str() << "\n";);
IF_VERBOSE(1, verbose_stream() << strm.str() << "\n";);
set_reason_unknown(strm.str());
return l_undef;
}
return l_true;
}
lbool internalize_goal(unsigned sz, expr* const* fmls) {
m_solver.pop_to_base_level();
if (m_solver.inconsistent())
return l_false;
m_pc.reset();
m_goal2sat(m, sz, fmls, m_params, m_solver, m_map, m_dep2asm, is_incremental());
if (!m_sat_mc) m_sat_mc = alloc(sat2goal::mc, m);
m_sat_mc->flush_smc(m_solver, m_map);
return check_uninterpreted();
}
lbool internalize_goal(goal_ref& g) {
m_solver.pop_to_base_level();
if (m_solver.inconsistent())
@ -712,12 +739,14 @@ private:
}
SASSERT(!g->proofs_enabled());
TRACE("sat", m_solver.display(tout); g->display(tout););
try {
if (m_is_cnf) {
m_subgoals.push_back(g.get());
}
else {
(*m_preprocess)(g, m_subgoals);
m_is_cnf = true;
}
}
catch (tactic_exception & ex) {
@ -737,8 +766,8 @@ private:
IF_VERBOSE(0, verbose_stream() << "size of subgoals is not 1, it is: " << m_subgoals.size() << "\n");
return l_undef;
}
g = m_subgoals[0];
func_decl_ref_vector funs(m);
m_pc = g->pc();
m_mcs.set(m_mcs.size()-1, concat(m_mcs.back(), g->mc()));
TRACE("sat", g->display_with_dependencies(tout););
@ -746,19 +775,10 @@ private:
// ensure that if goal is already internalized, then import mc from m_solver.
m_goal2sat(*g, m_params, m_solver, m_map, m_dep2asm, is_incremental());
m_goal2sat.get_interpreted_funs(funs);
if (!m_sat_mc) m_sat_mc = alloc(sat2goal::mc, m);
m_sat_mc->flush_smc(m_solver, m_map);
if (!funs.empty()) {
m_has_uninterpreted = true;
std::stringstream strm;
strm << "(sat.giveup interpreted functions sent to SAT solver " << funs <<")";
TRACE("sat", tout << strm.str() << "\n";);
IF_VERBOSE(1, verbose_stream() << strm.str() << "\n";);
set_reason_unknown(strm.str());
return l_undef;
}
return l_true;
return check_uninterpreted();
}
lbool internalize_assumptions(unsigned sz, expr* const* asms) {
@ -766,17 +786,29 @@ private:
m_asms.shrink(0);
return l_true;
}
for (unsigned i = 0; i < sz; ++i)
m_is_cnf &= is_literal(asms[i]);
for (unsigned i = 0; i < get_num_assumptions(); ++i)
m_is_cnf &= is_literal(get_assumption(i));
if (m_is_cnf) {
expr_ref_vector fmls(m);
fmls.append(sz, asms);
for (unsigned i = 0; i < get_num_assumptions(); ++i)
fmls.push_back(get_assumption(i));
m_goal2sat.assumptions(m, fmls.size(), fmls.data(), m_params, m_solver, m_map, m_dep2asm, is_incremental());
extract_assumptions(fmls.size(), fmls.data());
return l_true;
}
goal_ref g = alloc(goal, m, true, true); // models and cores are enabled.
for (unsigned i = 0; i < sz; ++i) {
for (unsigned i = 0; i < sz; ++i)
g->assert_expr(asms[i], m.mk_leaf(asms[i]));
}
for (unsigned i = 0; i < get_num_assumptions(); ++i) {
for (unsigned i = 0; i < get_num_assumptions(); ++i)
g->assert_expr(get_assumption(i), m.mk_leaf(get_assumption(i)));
}
lbool res = internalize_goal(g);
if (res == l_true) {
if (res == l_true)
extract_assumptions(sz, asms);
}
return res;
}
@ -890,33 +922,40 @@ private:
}
bool is_clause(expr* fml) {
if (is_literal(fml)) {
if (get_depth(fml) > 4)
return false;
if (is_literal(fml))
return true;
if (m.is_or(fml) || m.is_and(fml) || m.is_implies(fml) || m.is_not(fml) || m.is_iff(fml)) {
for (expr* n : *to_app(fml))
if (!is_clause(n))
return false;
return true;
}
if (!m.is_or(fml)) {
return false;
}
for (expr* n : *to_app(fml)) {
if (!is_literal(n)) {
return false;
}
}
return true;
return false;
}
lbool internalize_formulas() {
if (m_fmls_head == m_fmls.size()) {
if (m_fmls_head == m_fmls.size())
return l_true;
lbool res;
if (m_is_cnf) {
res = internalize_goal(m_fmls.size() - m_fmls_head, m_fmls.data() + m_fmls_head);
}
goal_ref g = alloc(goal, m, true, false); // models, maybe cores are enabled
for (unsigned i = m_fmls_head ; i < m_fmls.size(); ++i) {
expr* fml = m_fmls.get(i);
g->assert_expr(fml);
else {
goal_ref g = alloc(goal, m, true, false); // models, maybe cores are enabled
for (unsigned i = m_fmls_head ; i < m_fmls.size(); ++i) {
expr* fml = m_fmls.get(i);
g->assert_expr(fml);
}
res = internalize_goal(g);
}
lbool res = internalize_goal(g);
if (res != l_undef) {
if (res != l_undef)
m_fmls_head = m_fmls.size();
}
m_internalized_converted = false;
return res;
}

View file

@ -28,6 +28,7 @@ Revision History:
#include "util/stopwatch.h"
#include "util/symbol.h"
#include "util/sat_literal.h"
#include "util/rational.h"
class params_ref;
class reslimit;
@ -93,27 +94,47 @@ namespace sat {
};
enum class hint_type {
null_h,
farkas_h,
bound_h,
implied_eq_h,
};
struct proof_hint {
hint_type m_ty = hint_type::null_h;
vector<std::pair<rational, literal>> m_literals;
vector<std::pair<unsigned, unsigned>> m_eqs;
vector<std::pair<unsigned, unsigned>> m_diseqs;
void reset() { m_ty = hint_type::null_h; m_literals.reset(); m_eqs.reset(); m_diseqs.reset(); }
std::string to_string() const;
void from_string(char const* s);
void from_string(std::string const& s) { from_string(s.c_str()); }
};
class status {
public:
enum class st { input, asserted, redundant, deleted };
st m_st;
int m_orig;
const proof_hint* m_hint;
public:
status(st s, int o) : m_st(s), m_orig(o) {};
status(status const& s) : m_st(s.m_st), m_orig(s.m_orig) {}
status(status&& s) noexcept { m_st = st::asserted; m_orig = -1; std::swap(m_st, s.m_st); std::swap(m_orig, s.m_orig); }
status(st s, int o, proof_hint const* ps = nullptr) : m_st(s), m_orig(o), m_hint(ps) {};
status(status const& s) : m_st(s.m_st), m_orig(s.m_orig), m_hint(s.m_hint) {}
status(status&& s) noexcept { m_st = st::asserted; m_orig = -1; std::swap(m_st, s.m_st); std::swap(m_orig, s.m_orig); std::swap(m_hint, s.m_hint); }
status& operator=(status const& other) { m_st = other.m_st; m_orig = other.m_orig; return *this; }
static status redundant() { return status(status::st::redundant, -1); }
static status asserted() { return status(status::st::asserted, -1); }
static status deleted() { return status(status::st::deleted, -1); }
static status input() { return status(status::st::input, -1); }
static status th(bool redundant, int id) { return status(redundant ? st::redundant : st::asserted, id); }
static status th(bool redundant, int id, proof_hint const* ps = nullptr) { return status(redundant ? st::redundant : st::asserted, id, ps); }
bool is_input() const { return st::input == m_st; }
bool is_redundant() const { return st::redundant == m_st; }
bool is_asserted() const { return st::asserted == m_st; }
bool is_deleted() const { return st::deleted == m_st; }
proof_hint const* get_hint() const { return m_hint; }
bool is_sat() const { return -1 == m_orig; }
int get_th() const { return m_orig; }

View file

@ -128,6 +128,22 @@ namespace arith {
}
else {
expr_ref mone(a.mk_int(-1), m);
expr_ref abs_q(m.mk_ite(a.mk_ge(q, zero), q, a.mk_uminus(q)), m);
literal eqz = mk_literal(m.mk_eq(q, zero));
literal mod_ge_0 = mk_literal(a.mk_ge(mod, zero));
literal mod_lt_q = mk_literal(a.mk_le(a.mk_sub(mod, abs_q), mone));
// q = 0 or p = (p mod q) + q * (p div q)
// q = 0 or (p mod q) >= 0
// q = 0 or (p mod q) < abs(q)
add_clause(eqz, eq);
add_clause(eqz, mod_ge_0);
add_clause(eqz, mod_lt_q);
#if 0
/*literal div_ge_0 = */ mk_literal(a.mk_ge(div, zero));
/*literal div_le_0 = */ mk_literal(a.mk_le(div, zero));
/*literal p_ge_0 = */ mk_literal(a.mk_ge(p, zero));
@ -139,6 +155,8 @@ namespace arith {
// q <= 0 or (p mod q) >= 0
// q <= 0 or (p mod q) < q
// q >= 0 or (p mod q) < -q
literal q_ge_0 = mk_literal(a.mk_ge(q, zero));
literal q_le_0 = mk_literal(a.mk_le(q, zero));
@ -148,6 +166,17 @@ namespace arith {
add_clause(q_le_0, mod_ge_0);
add_clause(q_le_0, ~mk_literal(a.mk_ge(a.mk_sub(mod, q), zero)));
add_clause(q_ge_0, ~mk_literal(a.mk_ge(a.mk_add(mod, q), zero)));
#endif
if (a.is_zero(p)) {
add_clause(eqz, mk_literal(m.mk_eq(mod, zero)));
add_clause(eqz, mk_literal(m.mk_eq(div, zero)));
}
else if (!a.is_numeral(q)) {
// (or (= y 0) (<= (* y (div x y)) x))
add_clause(eqz, mk_literal(a.mk_le(a.mk_mul(q, div), p)));
}
}
if (get_config().m_arith_enum_const_mod && k.is_pos() && k < rational(8)) {
@ -234,44 +263,54 @@ namespace arith {
if (k1 == k2 && kind1 == kind2) return;
SASSERT(k1 != k2 || kind1 != kind2);
auto bin_clause = [&](sat::literal l1, sat::literal l2) {
sat::proof_hint* bound_params = nullptr;
if (ctx.use_drat()) {
bound_params = &m_farkas2;
m_farkas2.m_literals[0] = std::make_pair(rational(1), ~l1);
m_farkas2.m_literals[1] = std::make_pair(rational(1), ~l2);
}
add_clause(l1, l2, bound_params);
};
if (kind1 == lp_api::lower_t) {
if (kind2 == lp_api::lower_t) {
if (k2 <= k1)
add_clause(~l1, l2);
bin_clause(~l1, l2);
else
add_clause(l1, ~l2);
bin_clause(l1, ~l2);
}
else if (k1 <= k2)
// k1 <= k2, k1 <= x or x <= k2
add_clause(l1, l2);
bin_clause(l1, l2);
else {
// k1 > hi_inf, k1 <= x => ~(x <= hi_inf)
add_clause(~l1, ~l2);
bin_clause(~l1, ~l2);
if (v_is_int && k1 == k2 + rational(1))
// k1 <= x or x <= k1-1
add_clause(l1, l2);
bin_clause(l1, l2);
}
}
else if (kind2 == lp_api::lower_t) {
if (k1 >= k2)
// k1 >= lo_inf, k1 >= x or lo_inf <= x
add_clause(l1, l2);
bin_clause(l1, l2);
else {
// k1 < k2, k2 <= x => ~(x <= k1)
add_clause(~l1, ~l2);
bin_clause(~l1, ~l2);
if (v_is_int && k1 == k2 - rational(1))
// x <= k1 or k1+l <= x
add_clause(l1, l2);
bin_clause(l1, l2);
}
}
else {
// kind1 == A_UPPER, kind2 == A_UPPER
if (k1 >= k2)
// k1 >= k2, x <= k2 => x <= k1
add_clause(l1, ~l2);
bin_clause(l1, ~l2);
else
// k1 <= hi_sup , x <= k1 => x <= hi_sup
add_clause(~l1, l2);
bin_clause(~l1, l2);
}
}
@ -498,8 +537,8 @@ namespace arith {
if (x->get_root() == y->get_root())
return;
reset_evidence();
set_evidence(ci1, m_core, m_eqs);
set_evidence(ci2, m_core, m_eqs);
set_evidence(ci1);
set_evidence(ci2);
++m_stats.m_fixed_eqs;
auto* jst = euf::th_explain::propagate(*this, m_core, m_eqs, x, y);
ctx.propagate(x, y, jst->to_index());

View file

@ -79,4 +79,58 @@ namespace arith {
lp().settings().stats().collect_statistics(st);
if (m_nla) m_nla->collect_statistics(st);
}
void solver::explain_assumptions() {
m_arith_hint.reset();
unsigned i = 0;
for (auto const & ev : m_explanation) {
++i;
auto idx = ev.ci();
if (UINT_MAX == idx)
continue;
switch (m_constraint_sources[idx]) {
case inequality_source: {
literal lit = m_inequalities[idx];
m_arith_hint.m_literals.push_back({ev.coeff(), lit});
break;
}
case equality_source: {
auto [u, v] = m_equalities[idx];
ctx.drat_log_expr(u->get_expr());
ctx.drat_log_expr(v->get_expr());
m_arith_hint.m_eqs.push_back({u->get_expr_id(), v->get_expr_id()});
break;
}
default:
break;
}
}
}
/**
* It may be necessary to use the following assumption when checking Farkas claims
* generated from bounds propagation:
* A bound literal ax <= b is explained by a set of weighted literals
* r1*(a1*x <= b1) + .... + r_k*(a_k*x <= b_k), where r_i > 0
* such that there is a r >= 1
* (r1*a1+..+r_k*a_k) = r*a, (r1*b1+..+r_k*b_k) <= r*b
*/
sat::proof_hint const* solver::explain(sat::hint_type ty, sat::literal lit) {
if (!ctx.use_drat())
return nullptr;
m_arith_hint.m_ty = ty;
explain_assumptions();
if (lit != sat::null_literal)
m_arith_hint.m_literals.push_back({rational(1), ~lit});
return &m_arith_hint;
}
sat::proof_hint const* solver::explain_implied_eq(euf::enode* a, euf::enode* b) {
if (!ctx.use_drat())
return nullptr;
m_arith_hint.m_ty = sat::hint_type::implied_eq_h;
explain_assumptions();
m_arith_hint.m_diseqs.push_back({a->get_expr_id(), b->get_expr_id()});
return &m_arith_hint;
}
}

View file

@ -0,0 +1,356 @@
/*++
Copyright (c) 2020 Microsoft Corporation
Module Name:
arith_proof_checker.h
Abstract:
Plugin for checking arithmetic lemmas
Author:
Nikolaj Bjorner (nbjorner) 2020-09-08
--*/
#pragma once
#include "util/obj_pair_set.h"
#include "ast/ast_trail.h"
#include "ast/arith_decl_plugin.h"
namespace arith {
class proof_checker {
struct row {
obj_map<expr, rational> m_coeffs;
rational m_coeff;
void reset() {
m_coeffs.reset();
m_coeff = 0;
}
};
ast_manager& m;
arith_util a;
vector<std::pair<rational, expr*>> m_todo;
bool m_strict = false;
row m_ineq;
row m_conseq;
vector<row> m_eqs;
vector<row> m_ineqs;
vector<row> m_diseqs;
void add(row& r, expr* v, rational const& coeff) {
rational coeff1;
if (coeff.is_zero())
return;
if (r.m_coeffs.find(v, coeff1)) {
coeff1 += coeff;
if (coeff1.is_zero())
r.m_coeffs.erase(v);
else
r.m_coeffs[v] = coeff1;
}
else
r.m_coeffs.insert(v, coeff);
}
void mul(row& r, rational const& coeff) {
if (coeff == 1)
return;
for (auto & [v, c] : r.m_coeffs)
c *= coeff;
r.m_coeff *= coeff;
}
// dst <- dst + mul*src
void add(row& dst, row const& src, rational const& mul) {
for (auto const& [v, c] : src.m_coeffs)
add(dst, v, c*mul);
dst.m_coeff += mul*src.m_coeff;
}
// dst <- X*dst + Y*src
// where
// X = lcm(a,b)/b, Y = -lcm(a,b)/a if v is integer
// X = 1/b, Y = -1/a if v is real
//
void resolve(expr* v, row& dst, rational const& A, row const& src) {
rational B, x, y;
if (!dst.m_coeffs.find(v, B))
return;
if (a.is_int(v)) {
rational lc = lcm(abs(A), abs(B));
x = lc / abs(B);
y = lc / abs(A);
}
else {
x = rational(1) / abs(B);
y = rational(1) / abs(A);
}
if (A < 0 && B < 0)
y.neg();
if (A > 0 && B > 0)
y.neg();
mul(dst, x);
add(dst, src, y);
}
void cut(row& r) {
if (r.m_coeffs.empty())
return;
auto const& [v, coeff] = *r.m_coeffs.begin();
if (!a.is_int(v))
return;
rational lc = denominator(r.m_coeff);
for (auto const& [v, coeff] : r.m_coeffs)
lc = lcm(lc, denominator(coeff));
if (lc != 1) {
r.m_coeff *= lc;
for (auto & [v, coeff] : r.m_coeffs)
coeff *= lc;
}
rational g(0);
for (auto const& [v, coeff] : r.m_coeffs)
g = gcd(coeff, g);
if (g == 1)
return;
rational m = mod(r.m_coeff, g);
if (m == 0)
return;
r.m_coeff += g - m;
}
/**
* \brief populate m_coeffs, m_coeff based on mul*e
*/
void linearize(row& r, rational const& mul, expr* e) {
SASSERT(m_todo.empty());
m_todo.push_back({ mul, e });
rational coeff1;
expr* e1, *e2;
for (unsigned i = 0; i < m_todo.size(); ++i) {
auto [coeff, e] = m_todo[i];
if (a.is_mul(e, e1, e2) && a.is_numeral(e1, coeff1))
m_todo.push_back({coeff*coeff1, e2});
else if (a.is_mul(e, e1, e2) && a.is_numeral(e2, coeff1))
m_todo.push_back({coeff*coeff1, e1});
else if (a.is_add(e))
for (expr* arg : *to_app(e))
m_todo.push_back({coeff, arg});
else if (a.is_uminus(e, e1))
m_todo.push_back({-coeff, e1});
else if (a.is_sub(e, e1, e2)) {
m_todo.push_back({coeff, e1});
m_todo.push_back({-coeff, e2});
}
else if (a.is_numeral(e, coeff1))
r.m_coeff += coeff*coeff1;
else
add(r, e, coeff);
}
m_todo.reset();
}
bool check_ineq(row& r) {
if (r.m_coeffs.empty() && r.m_coeff > 0)
return true;
if (r.m_coeffs.empty() && m_strict && r.m_coeff == 0)
return true;
return false;
}
// triangulate equalities, substitute results into m_ineq, m_conseq.
void reduce_eq() {
for (unsigned i = 0; i < m_eqs.size(); ++i) {
auto& r = m_eqs[i];
if (r.m_coeffs.empty())
continue;
auto [v, coeff] = *r.m_coeffs.begin();
for (unsigned j = i + 1; j < m_eqs.size(); ++j)
resolve(v, m_eqs[j], coeff, r);
resolve(v, m_ineq, coeff, r);
resolve(v, m_conseq, coeff, r);
}
}
bool add_literal(row& r, rational const& coeff, expr* e, bool sign) {
expr* e1, *e2 = nullptr;
if ((a.is_le(e, e1, e2) || a.is_ge(e, e2, e1)) && !sign) {
linearize(r, coeff, e1);
linearize(r, -coeff, e2);
}
else if ((a.is_lt(e, e1, e2) || a.is_gt(e, e2, e1)) && sign) {
linearize(r, coeff, e2);
linearize(r, -coeff, e1);
}
else if ((a.is_le(e, e1, e2) || a.is_ge(e, e2, e1)) && sign) {
linearize(r, coeff, e2);
linearize(r, -coeff, e1);
if (a.is_int(e1))
r.m_coeff += coeff;
else
m_strict = true;
}
else if ((a.is_lt(e, e1, e2) || a.is_gt(e, e2, e1)) && !sign) {
linearize(r, coeff, e1);
linearize(r, -coeff, e2);
if (a.is_int(e1))
r.m_coeff += coeff;
else
m_strict = true;
}
else
return false;
// display_row(std::cout << coeff << " * " << (sign?"~":"") << mk_pp(e, m) << "\n", r) << "\n";
return true;
}
bool check_farkas() {
if (check_ineq(m_ineq))
return true;
reduce_eq();
if (check_ineq(m_ineq))
return true;
// convert to expression, maybe follows from a cut.
return false;
}
//
// farkas coefficient is computed for m_conseq
// after all inequalities in ineq have been added up
//
bool check_bound() {
reduce_eq();
if (check_ineq(m_conseq))
return true;
if (m_ineq.m_coeffs.empty() ||
m_conseq.m_coeffs.empty())
return false;
cut(m_ineq);
auto const& [v, coeff1] = *m_ineq.m_coeffs.begin();
rational coeff2;
if (!m_conseq.m_coeffs.find(v, coeff2))
return false;
add(m_conseq, m_ineq, abs(coeff2/coeff1));
if (check_ineq(m_conseq))
return true;
return false;
}
//
// checking disequalities is TBD.
// it has to select only a subset of bounds to justify each inequality.
// example
// c <= x <= c, c <= y <= c => x = y
// for the proof of x <= y use the inequalities x <= c <= y
// for the proof of y <= x use the inequalities y <= c <= x
// example
// x <= y, y <= z, z <= u, u <= x => x = z
// for the proof of x <= z use the inequalities x <= y, y <= z
// for the proof of z <= x use the inequalities z <= u, u <= x
//
// so when m_diseqs is non-empty we can't just add inequalities with Farkas coefficients
// into m_ineq, since coefficients of the usable subset vanish.
//
bool check_diseq() {
return false;
}
std::ostream& display_row(std::ostream& out, row const& r) {
bool first = true;
for (auto const& [v, coeff] : r.m_coeffs) {
if (!first)
out << " + ";
if (coeff != 1)
out << coeff << " * ";
out << mk_pp(v, m);
first = false;
}
if (r.m_coeff != 0)
out << " + " << r.m_coeff;
return out;
}
void display_eq(std::ostream& out, row const& r) {
display_row(out, r);
out << " = 0\n";
}
void display_ineq(std::ostream& out, row const& r) {
display_row(out, r);
if (m_strict)
out << " < 0\n";
else
out << " <= 0\n";
}
row& fresh(vector<row>& rows) {
rows.push_back(row());
return rows.back();
}
public:
proof_checker(ast_manager& m): m(m), a(m) {}
void reset() {
m_ineq.reset();
m_conseq.reset();
m_eqs.reset();
m_ineqs.reset();
m_diseqs.reset();
m_strict = false;
}
bool add_ineq(rational const& coeff, expr* e, bool sign) {
if (!m_diseqs.empty())
return add_literal(fresh(m_ineqs), abs(coeff), e, sign);
return add_literal(m_ineq, abs(coeff), e, sign);
}
bool add_conseq(rational const& coeff, expr* e, bool sign) {
return add_literal(m_conseq, abs(coeff), e, sign);
}
void add_eq(expr* a, expr* b) {
row& r = fresh(m_eqs);
linearize(r, rational(1), a);
linearize(r, rational(-1), b);
}
void add_diseq(expr* a, expr* b) {
row& r = fresh(m_diseqs);
linearize(r, rational(1), a);
linearize(r, rational(-1), b);
}
bool check() {
if (!m_diseqs.empty())
return check_diseq();
else if (!m_conseq.m_coeffs.empty())
return check_bound();
else
return check_farkas();
}
std::ostream& display(std::ostream& out) {
for (auto & r : m_eqs)
display_eq(out, r);
display_ineq(out, m_ineq);
if (!m_conseq.m_coeffs.empty())
display_ineq(out, m_conseq);
return out;
}
};
}

View file

@ -39,6 +39,8 @@ namespace arith {
lp().settings().set_random_seed(get_config().m_random_seed);
m_lia = alloc(lp::int_solver, *m_solver.get());
m_farkas2.m_ty = sat::hint_type::farkas_h;
m_farkas2.m_literals.resize(2);
}
solver::~solver() {
@ -194,11 +196,14 @@ namespace arith {
++m_stats.m_bound_propagations2;
reset_evidence();
m_core.push_back(lit1);
m_params.push_back(parameter(m_farkas));
m_params.push_back(parameter(rational(1)));
m_params.push_back(parameter(rational(1)));
TRACE("arith", tout << lit2 << " <- " << m_core << "\n";);
assign(lit2, m_core, m_eqs, m_params);
sat::proof_hint* ph = nullptr;
if (ctx.use_drat()) {
ph = &m_farkas2;
m_farkas2.m_literals[0] = std::make_pair(rational(1), lit1);
m_farkas2.m_literals[1] = std::make_pair(rational(1), ~lit2);
}
assign(lit2, m_core, m_eqs, ph);
++m_stats.m_bounds_propagations;
}
@ -227,26 +232,23 @@ namespace arith {
if (v == euf::null_theory_var)
return;
TRACE("arith", tout << "v" << v << " " << be.kind() << " " << be.m_bound << "\n";);
reserve_bounds(v);
if (m_unassigned_bounds[v] == 0 && !should_refine_bounds()) {
TRACE("arith", tout << "return\n";);
if (m_unassigned_bounds[v] == 0 && !should_refine_bounds())
return;
}
TRACE("arith", tout << "lp bound v" << v << " " << be.kind() << " " << be.m_bound << "\n";);
lp_bounds const& bounds = m_bounds[v];
bool first = true;
for (unsigned i = 0; i < bounds.size(); ++i) {
api_bound* b = bounds[i];
if (s().value(b->get_lit()) != l_undef) {
if (s().value(b->get_lit()) != l_undef)
continue;
}
literal lit = is_bound_implied(be.kind(), be.m_bound, *b);
if (lit == sat::null_literal) {
if (lit == sat::null_literal)
continue;
}
TRACE("arith", tout << lit << " bound: " << *b << " first: " << first << "\n";);
TRACE("arith", tout << "lp bound " << lit << " bound: " << *b << " first: " << first << "\n";);
lp().settings().stats().m_num_of_implied_bounds++;
if (first) {
@ -260,7 +262,7 @@ namespace arith {
TRACE("arith", for (auto lit : m_core) tout << lit << ": " << s().value(lit) << "\n";);
DEBUG_CODE(for (auto lit : m_core) { VERIFY(s().value(lit) == l_true); });
++m_stats.m_bound_propagations1;
assign(lit, m_core, m_eqs, m_params);
assign(lit, m_core, m_eqs, explain(sat::hint_type::bound_h, lit));
}
if (should_refine_bounds() && first)
@ -297,7 +299,7 @@ namespace arith {
void solver::consume(rational const& v, lp::constraint_index j) {
set_evidence(j, m_core, m_eqs);
set_evidence(j);
m_explanation.add_pair(j, v);
}
@ -318,8 +320,9 @@ namespace arith {
return false;
reset_evidence();
for (auto ev : e)
set_evidence(ev.ci(), m_core, m_eqs);
auto* jst = euf::th_explain::propagate(*this, m_core, m_eqs, n1, n2);
set_evidence(ev.ci());
auto* ex = explain_implied_eq(n1, n2);
auto* jst = euf::th_explain::propagate(*this, m_core, m_eqs, n1, n2, ex);
ctx.propagate(n1, n2, jst->to_index());
return true;
}
@ -375,7 +378,7 @@ namespace arith {
reset_evidence();
m_explanation.clear();
lp().explain_implied_bound(be, m_bp);
assign(bound, m_core, m_eqs, m_params);
assign(bound, m_core, m_eqs, explain(sat::hint_type::farkas_h, bound));
}
@ -748,13 +751,14 @@ namespace arith {
++m_stats.m_fixed_eqs;
reset_evidence();
set_evidence(ci1, m_core, m_eqs);
set_evidence(ci2, m_core, m_eqs);
set_evidence(ci3, m_core, m_eqs);
set_evidence(ci4, m_core, m_eqs);
set_evidence(ci1);
set_evidence(ci2);
set_evidence(ci3);
set_evidence(ci4);
enode* x = var2enode(v1);
enode* y = var2enode(v2);
auto* jst = euf::th_explain::propagate(*this, m_core, m_eqs, x, y);
auto* ex = explain_implied_eq(x, y);
auto* jst = euf::th_explain::propagate(*this, m_core, m_eqs, x, y, ex);
ctx.propagate(x, y, jst->to_index());
}
@ -1019,6 +1023,9 @@ namespace arith {
return sat::check_result::CR_CONTINUE;
case l_undef:
TRACE("arith", tout << "check-lia giveup\n";);
if (ctx.get_config().m_arith_ignore_int)
return sat::check_result::CR_GIVEUP;
st = sat::check_result::CR_CONTINUE;
break;
}
@ -1132,7 +1139,11 @@ namespace arith {
if (!check_idiv_bounds())
return l_false;
switch (m_lia->check(&m_explanation)) {
auto cr = m_lia->check(&m_explanation);
if (cr != lp::lia_move::sat && ctx.get_config().m_arith_ignore_int)
return l_undef;
switch (cr) {
case lp::lia_move::sat:
lia_check = l_true;
break;
@ -1161,13 +1172,13 @@ namespace arith {
// m_explanation implies term <= k
reset_evidence();
for (auto ev : m_explanation)
set_evidence(ev.ci(), m_core, m_eqs);
set_evidence(ev.ci());
// The call mk_bound() can set the m_infeasible_column in lar_solver
// so the explanation is safer to take before this call.
app_ref b = mk_bound(m_lia->get_term(), m_lia->get_offset(), !m_lia->is_upper());
IF_VERBOSE(4, verbose_stream() << "cut " << b << "\n");
literal lit = expr2literal(b);
assign(lit, m_core, m_eqs, m_params);
assign(lit, m_core, m_eqs, explain(sat::hint_type::bound_h, lit));
lia_check = l_false;
break;
}
@ -1189,16 +1200,16 @@ namespace arith {
return lia_check;
}
void solver::assign(literal lit, literal_vector const& core, svector<enode_pair> const& eqs, vector<parameter> const& params) {
void solver::assign(literal lit, literal_vector const& core, svector<enode_pair> const& eqs, sat::proof_hint const* pma) {
if (core.size() < small_lemma_size() && eqs.empty()) {
m_core2.reset();
for (auto const& c : core)
m_core2.push_back(~c);
m_core2.push_back(lit);
add_clause(m_core2);
add_clause(m_core2, pma);
}
else {
auto* jst = euf::th_explain::propagate(*this, core, eqs, lit);
auto* jst = euf::th_explain::propagate(*this, core, eqs, lit, pma);
ctx.propagate(lit, jst->to_index());
}
}
@ -1221,7 +1232,7 @@ namespace arith {
++m_num_conflicts;
++m_stats.m_conflicts;
for (auto ev : m_explanation)
set_evidence(ev.ci(), m_core, m_eqs);
set_evidence(ev.ci());
TRACE("arith",
tout << "Lemma - " << (is_conflict ? "conflict" : "propagation") << "\n";
for (literal c : m_core) tout << literal2expr(c) << "\n";
@ -1236,14 +1247,14 @@ namespace arith {
for (literal& c : m_core)
c.neg();
add_clause(m_core);
add_clause(m_core, explain(sat::hint_type::farkas_h));
}
bool solver::is_infeasible() const {
return lp().get_status() == lp::lp_status::INFEASIBLE;
}
void solver::set_evidence(lp::constraint_index idx, literal_vector& core, svector<enode_pair>& eqs) {
void solver::set_evidence(lp::constraint_index idx) {
if (idx == UINT_MAX) {
return;
}
@ -1251,7 +1262,7 @@ namespace arith {
case inequality_source: {
literal lit = m_inequalities[idx];
SASSERT(lit != sat::null_literal);
core.push_back(lit);
m_core.push_back(lit);
break;
}
case equality_source:

View file

@ -413,12 +413,18 @@ namespace arith {
void get_infeasibility_explanation_and_set_conflict();
void set_conflict();
void set_conflict_or_lemma(literal_vector const& core, bool is_conflict);
void set_evidence(lp::constraint_index idx, literal_vector& core, svector<enode_pair>& eqs);
void assign(literal lit, literal_vector const& core, svector<enode_pair> const& eqs, vector<parameter> const& params);
void set_evidence(lp::constraint_index idx);
void assign(literal lit, literal_vector const& core, svector<enode_pair> const& eqs, sat::proof_hint const* pma);
void false_case_of_check_nla(const nla::lemma& l);
void dbg_finalize_model(model& mdl);
sat::proof_hint m_arith_hint;
sat::proof_hint m_farkas2;
sat::proof_hint const* explain(sat::hint_type ty, sat::literal lit = sat::null_literal);
sat::proof_hint const* explain_implied_eq(euf::enode* a, euf::enode* b);
void explain_assumptions();
public:
solver(euf::solver& ctx, theory_id id);

View file

@ -59,13 +59,17 @@ namespace array {
axiom_record& r = m_axiom_trail[idx];
switch (r.m_kind) {
case axiom_record::kind_t::is_store:
return assert_store_axiom(to_app(r.n->get_expr()));
return assert_store_axiom(r.n->get_app());
case axiom_record::kind_t::is_select:
return assert_select(idx, r);
case axiom_record::kind_t::is_default:
return assert_default(r);
case axiom_record::kind_t::is_extensionality:
return assert_extensionality(r.n->get_expr(), r.select->get_expr());
case axiom_record::kind_t::is_diff:
return assert_diff(r.n->get_app());
case axiom_record::kind_t::is_diffselect:
return assert_diff_select(r.n->get_app(), r.select->get_app());
case axiom_record::kind_t::is_congruence:
return assert_congruent_axiom(r.n->get_expr(), r.select->get_expr());
default:
@ -273,6 +277,54 @@ namespace array {
return add_clause(lit1, ~lit2);
}
/**
* a = b or default(a) != default(b) or a[md(a,b)] != b[md(a,b)]
*/
bool solver::assert_diff(expr* md) {
expr* x = nullptr, *y = nullptr;
VERIFY(a.is_maxdiff(md, x, y) || a.is_mindiff(md, x, y));
expr* args1[2] = { x, md };
expr* args2[2] = { y, md };
literal eq = eq_internalize(x, y);
literal eq_default = eq_internalize(a.mk_default(x), a.mk_default(y));
literal eq_md = eq_internalize(a.mk_select(2, args1), a.mk_select(2, args2));
return add_clause(eq, ~eq_default, ~eq_md);
}
/**
* a = b and a[i] != c[i] => i <= md(b, c) or default(b) != default(c)
* a = c and a[i] != b[i] => i <= md(b, c) or default(b) != default(c)
* where ai = a[i], md = md(b, c)
*/
bool solver::assert_diff_select(app* md, app* ai) {
SASSERT(a.is_select(ai));
SASSERT(ai->get_num_args() == 2);
expr* A = ai->get_arg(0);
expr* i = ai->get_arg(1);
expr* B = md->get_arg(0);
expr* C = md->get_arg(1);
literal eq_default = eq_internalize(a.mk_default(B), a.mk_default(C));
arith_util autil(m);
literal ineq = mk_literal(a.is_maxdiff(md) ? autil.mk_le(i, md) : autil.mk_le(md, i));
bool is_new = false;
if (ctx.get_enode(A)->get_root() == ctx.get_enode(B)->get_root()) {
literal eq_ab = eq_internalize(A, B);
expr* args[2] = { C, i };
literal eq_select = eq_internalize(ai, a.mk_select(2, args));
if (add_clause(~eq_ab, eq_select, ineq, ~eq_default))
is_new = true;
}
if (ctx.get_enode(A)->get_root() == ctx.get_enode(C)->get_root()) {
literal eq_ac = eq_internalize(A, C);
expr* args[2] = { B, i };
literal eq_select = eq_internalize(ai, a.mk_select(2, args));
if (add_clause(~eq_ac, eq_select, ineq, ~eq_default))
is_new = true;
}
return is_new;
}
bool solver::is_map_combinator(expr* map) const {
return a.is_map(map) || a.is_union(map) || a.is_intersect(map) || a.is_difference(map) || a.is_complement(map);
}
@ -629,12 +681,14 @@ namespace array {
euf::enode * n = var2enode(i);
if (!is_array(n))
continue;
CTRACE("array", !ctx.is_relevant(n), tout << "not relevant: " << ctx.bpp(n) << "\n");
if (!ctx.is_relevant(n))
continue;
euf::enode * r = n->get_root();
if (r->is_marked1())
continue;
// arrays used as indices in other arrays have to be treated as shared issue #3532, #3529
// arrays used as indices in other arrays have to be treated as shared issue #3532, #3529
CTRACE("array", !ctx.is_shared(r) && !is_shared_arg(r), tout << "not shared: " << ctx.bpp(r) << "\n");
if (ctx.is_shared(r) || is_shared_arg(r))
roots.push_back(r->get_th_var(get_id()));
r->mark1();
@ -655,10 +709,33 @@ namespace array {
return true;
if (a.is_const(e))
return true;
if (a.is_ext(e))
return true;
}
return false;
}
bool solver::add_diff_select_axioms() {
bool added = false;
auto add_diff_select = [&](euf::enode* md, euf::enode* a) {
var_data const& d = get_var_data(find(get_th_var(a)));
for (euf::enode* select : d.m_parent_selects) {
if (assert_diff_select(md->get_app(), select->get_app()))
added = true;
}
};
for (euf::enode* md : m_minmaxdiffs) {
euf::enode* a = md->get_arg(0);
euf::enode* b = md->get_arg(1);
add_diff_select(md, a);
add_diff_select(md, b);
}
return added;
}
}

View file

@ -93,8 +93,10 @@ namespace array {
validate_extensionality(n, k);
}
expr* x = nullptr, *y = nullptr;
#if 0
if (m.is_eq(n->get_expr(), x, y) && a.is_array(x))
std::cout << ctx.bpp(n) << " " << s().value(n->bool_var()) << "\n";
#endif
if (m.is_eq(n->get_expr(), x, y) && a.is_array(x) && s().value(n->bool_var()) == l_false)
validate_extensionality(expr2enode(x), expr2enode(y));
}

View file

@ -99,7 +99,8 @@ namespace array {
}
void solver::internalize_eh(euf::enode* n) {
switch (n->get_decl()->get_decl_kind()) {
auto k = n->get_decl()->get_decl_kind();
switch (k) {
case OP_STORE:
ctx.push_vec(get_var_data(find(n)).m_lambdas, n);
push_axiom(store_axiom(n));
@ -114,6 +115,12 @@ namespace array {
SASSERT(is_array(n->get_arg(0)));
push_axiom(extensionality_axiom(n->get_arg(0), n->get_arg(1)));
break;
case OP_ARRAY_MINDIFF:
case OP_ARRAY_MAXDIFF:
push_axiom(diff_axiom(n));
m_minmaxdiffs.push_back(n);
ctx.push(push_back_vector(m_minmaxdiffs));
break;
case OP_ARRAY_DEFAULT:
add_parent_default(find(n->get_arg(0)), n);
break;
@ -169,6 +176,10 @@ namespace array {
break;
case OP_ARRAY_EXT:
break;
case OP_ARRAY_MINDIFF:
case OP_ARRAY_MAXDIFF:
// todo
break;
case OP_ARRAY_DEFAULT:
set_prop_upward(find(n->get_arg(0)));
break;
@ -236,6 +247,14 @@ namespace array {
return false;
}
bool solver::is_beta_redex(euf::enode* p, euf::enode* n) const {
if (a.is_select(p->get_expr()))
return p->get_arg(0)->get_root() == n->get_root();
if (a.is_map(p->get_expr()))
return true;
return false;
}
func_decl_ref_vector const& solver::sort2diff(sort* s) {
func_decl_ref_vector* result = nullptr;
if (m_sort2diff.find(s, result))

View file

@ -54,7 +54,7 @@ namespace array {
euf::enode* d = get_default(v);
if (d)
dep.add(n, d);
if (!dep.deps().contains(n))
if (!dep.contains_dep(n))
dep.insert(n, nullptr);
return true;
}

View file

@ -101,9 +101,14 @@ namespace array {
else if (!turn[idx] && add_interface_equalities())
return sat::check_result::CR_CONTINUE;
}
if (m_delay_qhead < m_axiom_trail.size())
if (add_diff_select_axioms())
return sat::check_result::CR_CONTINUE;
if (m_delay_qhead < m_axiom_trail.size())
return sat::check_result::CR_CONTINUE;
// validate_check();
return sat::check_result::CR_DONE;
}

View file

@ -83,6 +83,8 @@ namespace array {
is_store,
is_select,
is_extensionality,
is_diff,
is_diffselect,
is_default,
is_congruence
};
@ -145,9 +147,9 @@ namespace array {
axiom_record::eq m_eq;
axiom_table_t m_axioms;
svector<axiom_record> m_axiom_trail;
unsigned m_qhead { 0 };
unsigned m_delay_qhead { 0 };
bool m_enable_delay { true };
unsigned m_qhead = 0;
unsigned m_delay_qhead = 0;
bool m_enable_delay = true;
struct reset_new;
void push_axiom(axiom_record const& r);
bool propagate_axiom(unsigned idx);
@ -163,6 +165,9 @@ namespace array {
axiom_record store_axiom(euf::enode* n) { return axiom_record(axiom_record::kind_t::is_store, n); }
axiom_record extensionality_axiom(euf::enode* x, euf::enode* y) { return axiom_record(axiom_record::kind_t::is_extensionality, x, y); }
axiom_record congruence_axiom(euf::enode* a, euf::enode* b) { return axiom_record(axiom_record::kind_t::is_congruence, a, b); }
axiom_record diff_axiom(euf::enode* md) { return axiom_record(axiom_record::kind_t::is_diff, md); }
euf::enode_vector m_minmaxdiffs;
axiom_record diff_select_axiom(euf::enode* md, euf::enode* ai) { return axiom_record(axiom_record::kind_t::is_diffselect, md, ai); }
scoped_ptr<sat::constraint_base> m_constraint;
@ -175,12 +180,15 @@ namespace array {
bool assert_select_map_axiom(app* select, app* map);
bool assert_select_lambda_axiom(app* select, expr* lambda);
bool assert_extensionality(expr* e1, expr* e2);
bool assert_diff(expr* md);
bool assert_diff_select(app* ai, app* md);
bool assert_default_map_axiom(app* map);
bool assert_default_const_axiom(app* cnst);
bool assert_default_store_axiom(app* store);
bool assert_congruent_axiom(expr* e1, expr* e2);
bool add_delayed_axioms();
bool add_as_array_eqs(euf::enode* n);
bool add_diff_select_axioms();
expr_ref apply_map(app* map, unsigned n, expr* const* args);
bool is_map_combinator(expr* e) const;
@ -291,6 +299,7 @@ namespace array {
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;
bool is_beta_redex(euf::enode* p, euf::enode* n) const override;
bool enable_self_propagate() const override { return true; }
void relevant_eh(euf::enode* n) override;
bool enable_ackerman_axioms(euf::enode* n) const override { return !a.is_array(n->get_sort()); }

View file

@ -93,6 +93,7 @@ namespace bv {
} while (curr != v);
zero_one_bits const& _bits = m_zero_one_bits[v];
#if 0
if (_bits.size() != num_bits) {
std::cout << "v" << v << " " << _bits.size() << " " << num_bits << "\n";
std::cout << "true: " << mk_true() << "\n";
@ -102,6 +103,7 @@ namespace bv {
}
while (curr != v);
}
#endif
SASSERT(_bits.size() == num_bits);
VERIFY(_bits.size() == num_bits);
bool_vector already_found;

View file

@ -29,6 +29,7 @@ namespace dt {
th_euf_solver(ctx, ctx.get_manager().get_family_name(id), id),
dt(m),
m_autil(m),
m_sutil(m),
m_find(*this),
m_args(m)
{}
@ -496,13 +497,41 @@ namespace dt {
}
ptr_vector<euf::enode> const& solver::get_array_args(enode* n) {
m_array_args.reset();
m_nodes.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);
m_nodes.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;
m_nodes.push_back(ctx.get_enode(def));
return m_nodes;
}
ptr_vector<euf::enode> const& solver::get_seq_args(enode* n) {
m_nodes.reset();
m_todo.reset();
auto add_todo = [&](enode* n) {
if (!n->is_marked1()) {
n->mark1();
m_todo.push_back(n);
}
};
for (enode* sib : euf::enode_class(n))
add_todo(sib);
for (unsigned i = 0; i < m_todo.size(); ++i) {
enode* n = m_todo[i];
expr* e = n->get_expr();
if (m_sutil.str.is_unit(e))
m_nodes.push_back(n->get_arg(0));
else if (m_sutil.str.is_concat(e))
for (expr* arg : *to_app(e))
add_todo(ctx.get_enode(arg));
}
for (enode* n : m_todo)
n->unmark1();
return m_nodes;
}
// Assuming `app` is equal to a constructor term, return the constructor enode
@ -536,6 +565,12 @@ namespace dt {
for (enode* aarg : get_array_args(arg))
add(aarg);
}
sort* se;
if (m_sutil.is_seq(child->get_sort(), se) && dt.is_datatype(se)) {
for (enode* aarg : get_seq_args(child))
add(aarg);
}
VERIFY(found);
}
@ -575,6 +610,21 @@ namespace dt {
return false;
enode* parent = d->m_constructor;
oc_mark_on_stack(parent);
auto process_arg = [&](enode* aarg) {
if (oc_cycle_free(aarg))
return false;
if (oc_on_stack(aarg)) {
occurs_check_explain(parent, aarg);
return true;
}
if (dt.is_datatype(aarg->get_sort())) {
m_parent.insert(aarg->get_root(), parent);
oc_push_stack(aarg);
}
return false;
};
for (enode* arg : euf::enode_args(parent)) {
if (oc_cycle_free(arg))
continue;
@ -585,24 +635,20 @@ namespace dt {
}
// explore `arg` (with parent)
expr* earg = arg->get_expr();
sort* s = earg->get_sort();
sort* s = earg->get_sort(), *se;
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);
else if (m_sutil.is_seq(s, se) && dt.is_datatype(se)) {
for (enode* sarg : get_seq_args(arg))
if (process_arg(sarg))
return true;
}
else if (m_autil.is_array(s) && dt.is_datatype(get_array_range(s))) {
for (enode* sarg : get_array_args(arg))
if (process_arg(sarg))
return true;
}
if (is_datatype(aarg)) {
m_parent.insert(aarg->get_root(), parent);
oc_push_stack(aarg);
}
}
}
}
return false;

View file

@ -19,6 +19,7 @@ Author:
#include "sat/smt/sat_th.h"
#include "ast/datatype_decl_plugin.h"
#include "ast/array_decl_plugin.h"
#include "ast/seq_decl_plugin.h"
namespace euf {
class solver;
@ -62,6 +63,7 @@ namespace dt {
mutable datatype_util dt;
array_util m_autil;
seq_util m_sutil;
stats m_stats;
ptr_vector<var_data> m_var_data;
dt_union_find m_find;
@ -108,8 +110,9 @@ namespace dt {
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> m_nodes, m_todo;
ptr_vector<enode> const& get_array_args(enode* n);
ptr_vector<enode> const& get_seq_args(enode* n);
void pop_core(unsigned n) override;

View file

@ -358,6 +358,7 @@ namespace euf {
for (auto const& p : euf::enode_th_vars(n)) {
family_id id = p.get_id();
if (m.get_basic_family_id() != id) {
if (th_id != m.get_basic_family_id())
return true;
th_id = id;
@ -369,6 +370,8 @@ namespace euf {
for (enode* parent : euf::enode_parents(n)) {
app* p = to_app(parent->get_expr());
family_id fid = p->get_family_id();
if (is_beta_redex(parent, n))
continue;
if (fid != th_id && fid != m.get_basic_family_id())
return true;
}
@ -407,6 +410,13 @@ namespace euf {
return false;
}
bool solver::is_beta_redex(enode* p, enode* n) const {
for (auto const& th : enode_th_vars(p))
if (fid2solver(th.get_id())->is_beta_redex(p, n))
return true;
return false;
}
expr_ref solver::mk_eq(expr* e1, expr* e2) {
expr_ref _e1(e1, m);
expr_ref _e2(e2, m);

View file

@ -128,12 +128,14 @@ namespace euf {
n->unmark1();
TRACE("model",
for (auto const& d : deps.deps())
if (d.m_value) {
tout << bpp(d.m_key) << ":\n";
for (auto* n : *d.m_value)
for (auto * t : deps.deps()) {
auto* v = deps.get_dep(t);
if (v) {
tout << bpp(t) << ":\n";
for (auto* n : *v)
tout << " " << bpp(n) << "\n";
}
}
);
}
@ -335,8 +337,8 @@ namespace euf {
continue;
if (!tt && !mdl.is_true(e))
continue;
IF_VERBOSE(0, display_validation_failure(verbose_stream(), mdl, n););
CTRACE("euf", first, display_validation_failure(tout, mdl, n););
IF_VERBOSE(0, display_validation_failure(verbose_stream(), mdl, n););
(void)first;
first = false;
exit(1);

View file

@ -104,7 +104,7 @@ namespace euf {
std::ostringstream strm;
smt2_pp_environment_dbg env(m);
ast_smt2_pp(strm, f, env);
get_drat().def_begin('f', f->get_decl_id(), strm.str());
get_drat().def_begin('f', f->get_small_id(), strm.str());
get_drat().def_end();
}
@ -166,7 +166,7 @@ namespace euf {
lits.push_back(jst.lit_consequent());
if (jst.eq_consequent().first != nullptr)
lits.push_back(add_lit(jst.eq_consequent()));
get_drat().add(lits, sat::status::th(m_is_redundant, jst.ext().get_id()));
get_drat().add(lits, sat::status::th(m_is_redundant, jst.ext().get_id(), jst.get_pragma()));
}
void solver::drat_eq_def(literal lit, expr* eq) {

View file

@ -614,7 +614,7 @@ namespace euf {
if (si.is_bool_op(e))
lit = literal(replay.m[e], false);
else
lit = si.internalize(e, true);
lit = si.internalize(e, false);
VERIFY(lit.var() == v);
if (!m_egraph.find(e) && (!m.is_iff(e) && !m.is_or(e) && !m.is_and(e) && !m.is_not(e))) {
ptr_buffer<euf::enode> args;

View file

@ -176,7 +176,6 @@ namespace euf {
void log_antecedents(literal l, literal_vector const& r);
void log_justification(literal l, th_explain const& jst);
void drat_log_decl(func_decl* f);
void drat_log_expr(expr* n);
void drat_log_expr1(expr* n);
ptr_vector<expr> m_drat_todo;
obj_hashtable<ast> m_drat_asts;
@ -345,6 +344,7 @@ namespace euf {
sat::drat& get_drat() { return s().get_drat(); }
void drat_bool_def(sat::bool_var v, expr* n);
void drat_eq_def(sat::literal lit, expr* eq);
void drat_log_expr(expr* n);
// decompile
bool extract_pb(std::function<void(unsigned sz, literal const* c, unsigned k)>& card,
@ -371,6 +371,7 @@ namespace euf {
th_rewriter& get_rewriter() { return m_rewriter; }
void rewrite(expr_ref& e) { m_rewriter(e); }
bool is_shared(euf::enode* n) const;
bool is_beta_redex(euf::enode* p, euf::enode* n) const;
bool enable_ackerman_axioms(expr* n) const;
bool is_fixed(euf::enode* n, expr_ref& val, sat::literal_vector& explain);

View file

@ -34,7 +34,7 @@ namespace pb {
void solver::set_conflict(constraint& c, literal lit) {
m_stats.m_num_conflicts++;
TRACE("ba", display(tout, c, true); );
TRACE("pb", display(tout, c, true); );
if (!validate_conflict(c)) {
IF_VERBOSE(0, display(verbose_stream(), c, true));
UNREACHABLE();
@ -56,7 +56,7 @@ namespace pb {
default:
m_stats.m_num_propagations++;
m_num_propagations_since_pop++;
//TRACE("ba", tout << "#prop: " << m_stats.m_num_propagations << " - " << c.lit() << " => " << lit << "\n";);
//TRACE("pb", tout << "#prop: " << m_stats.m_num_propagations << " - " << c.lit() << " => " << lit << "\n";);
SASSERT(validate_unit_propagation(c, lit));
assign(lit, sat::justification::mk_ext_justification(s().scope_lvl(), c.cindex()));
break;
@ -69,7 +69,7 @@ namespace pb {
void solver::simplify(constraint& p) {
SASSERT(s().at_base_lvl());
if (p.lit() != sat::null_literal && value(p.lit()) == l_false) {
TRACE("ba", tout << "pb: flip sign " << p << "\n";);
TRACE("pb", tout << "pb: flip sign " << p << "\n";);
IF_VERBOSE(2, verbose_stream() << "sign is flipped " << p << "\n";);
return;
}
@ -280,7 +280,7 @@ namespace pb {
*/
lbool solver::add_assign(pbc& p, literal alit) {
BADLOG(display(verbose_stream() << "assign: " << alit << " watch: " << p.num_watch() << " size: " << p.size(), p, true));
TRACE("ba", display(tout << "assign: " << alit << "\n", p, true););
TRACE("pb", display(tout << "assign: " << alit << "\n", p, true););
SASSERT(!inconsistent());
unsigned sz = p.size();
unsigned bound = p.k();
@ -290,6 +290,7 @@ namespace pb {
SASSERT(p.lit() == sat::null_literal || value(p.lit()) == l_true);
SASSERT(num_watch <= sz);
SASSERT(num_watch > 0);
SASSERT(validate_watch(p, sat::null_literal));
unsigned index = 0;
m_a_max = 0;
m_pb_undef.reset();
@ -311,8 +312,6 @@ namespace pb {
return l_undef;
}
SASSERT(validate_watch(p, sat::null_literal));
// SASSERT(validate_watch(p, sat::null_literal));
SASSERT(index < num_watch);
unsigned index1 = index + 1;
@ -349,7 +348,7 @@ namespace pb {
SASSERT(validate_watch(p, sat::null_literal));
BADLOG(display(verbose_stream() << "conflict: " << alit << " watch: " << p.num_watch() << " size: " << p.size(), p, true));
SASSERT(bound <= slack);
TRACE("ba", tout << "conflict " << alit << "\n";);
TRACE("pb", tout << "conflict " << alit << "\n";);
set_conflict(p, alit);
return l_false;
}
@ -366,13 +365,14 @@ namespace pb {
p.swap(num_watch, index);
//
// slack >= bound, but slack - w(l) < bound
// l must be true.
//
if (slack < bound + m_a_max) {
BADLOG(verbose_stream() << "slack " << slack << " " << bound << " " << m_a_max << "\n";);
TRACE("ba", tout << p << "\n"; for(auto j : m_pb_undef) tout << j << " "; tout << "\n";);
TRACE("pb", tout << p << "\n"; for(auto j : m_pb_undef) tout << j << " "; tout << "\n";);
for (unsigned index1 : m_pb_undef) {
if (index1 == num_watch) {
index1 = index;
@ -389,7 +389,7 @@ namespace pb {
SASSERT(validate_watch(p, alit)); // except that alit is still watched.
TRACE("ba", display(tout << "assign: " << alit << "\n", p, true););
TRACE("pb", display(tout << "assign: " << alit << "\n", p, true););
BADLOG(verbose_stream() << "unwatch " << alit << " watch: " << p.num_watch() << " size: " << p.size() << " slack: " << p.slack() << " " << inconsistent() << "\n");
@ -547,7 +547,7 @@ namespace pb {
literal l = literal(v, c1 < 0);
c1 = std::abs(c1);
unsigned c = static_cast<unsigned>(c1);
// TRACE("ba", tout << l << " " << c << "\n";);
// TRACE("pb", tout << l << " " << c << "\n";);
m_overflow |= c != c1;
return wliteral(c, l);
}
@ -596,12 +596,13 @@ namespace pb {
s().reset_mark(v);
--m_num_marks;
}
if (idx == 0 && !_debug_conflict) {
if (idx == 0 && !_debug_conflict && m_num_marks > 0) {
_debug_conflict = true;
_debug_var2position.reserve(s().num_vars());
for (unsigned i = 0; i < lits.size(); ++i) {
_debug_var2position[lits[i].var()] = i;
}
IF_VERBOSE(0, verbose_stream() << "num marks: " << m_num_marks << "\n");
IF_VERBOSE(0,
active2pb(m_A);
uint64_t c = 0;
@ -617,20 +618,19 @@ namespace pb {
}
}
m_num_marks = 0;
resolve_conflict();
resolve_conflict();
exit(0);
}
--idx;
}
}
lbool solver::resolve_conflict() {
if (0 == m_num_propagations_since_pop) {
if (0 == m_num_propagations_since_pop)
return l_undef;
}
if (s().m_config.m_pb_resolve == sat::PB_ROUNDING) {
if (s().m_config.m_pb_resolve == sat::PB_ROUNDING)
return resolve_conflict_rs();
}
m_overflow = false;
reset_coeffs();
@ -638,7 +638,7 @@ namespace pb {
m_bound = 0;
literal consequent = s().m_not_l;
sat::justification js = s().m_conflict;
TRACE("ba", tout << consequent << " " << js << "\n";);
TRACE("pb", tout << consequent << " " << js << "\n";);
bool unique_max;
m_conflict_lvl = s().get_max_lvl(consequent, js, unique_max);
if (m_conflict_lvl == 0) {
@ -667,7 +667,7 @@ namespace pb {
}
DEBUG_CODE(TRACE("sat_verbose", display(tout, m_A);););
TRACE("ba", tout << "process consequent: " << consequent << " : "; s().display_justification(tout, js) << "\n";);
TRACE("pb", tout << "process consequent: " << consequent << " : "; s().display_justification(tout, js) << "\n";);
SASSERT(offset > 0);
DEBUG_CODE(justification2pb(js, consequent, offset, m_B););
@ -741,7 +741,7 @@ namespace pb {
inc_bound(offset);
inc_coeff(consequent, offset);
get_antecedents(consequent, p, m_lemma);
TRACE("ba", display(tout, p, true); tout << m_lemma << "\n";);
TRACE("pb", display(tout, p, true); tout << m_lemma << "\n";);
if (_debug_conflict) {
verbose_stream() << consequent << " ";
verbose_stream() << "antecedents: " << m_lemma << "\n";
@ -766,7 +766,7 @@ namespace pb {
active2pb(m_C);
VERIFY(validate_resolvent());
m_A = m_C;
TRACE("ba", display(tout << "conflict: ", m_A);););
TRACE("pb", display(tout << "conflict: ", m_A);););
cut();
@ -887,7 +887,7 @@ namespace pb {
}
}
ineq.divide(c);
TRACE("ba", display(tout << "var: " << v << " " << c << ": ", ineq, true););
TRACE("pb", display(tout << "var: " << v << " " << c << ": ", ineq, true););
}
void solver::round_to_one(bool_var w) {
@ -905,7 +905,7 @@ namespace pb {
SASSERT(validate_lemma());
divide(c);
SASSERT(validate_lemma());
TRACE("ba", active2pb(m_B); display(tout, m_B, true););
TRACE("pb", active2pb(m_B); display(tout, m_B, true););
}
void solver::divide(unsigned c) {
@ -935,14 +935,14 @@ namespace pb {
}
void solver::resolve_with(ineq const& ineq) {
TRACE("ba", display(tout, ineq, true););
TRACE("pb", display(tout, ineq, true););
inc_bound(ineq.m_k);
TRACE("ba", tout << "bound: " << m_bound << "\n";);
TRACE("pb", tout << "bound: " << m_bound << "\n";);
for (unsigned i = ineq.size(); i-- > 0; ) {
literal l = ineq.lit(i);
inc_coeff(l, static_cast<unsigned>(ineq.coeff(i)));
TRACE("ba", tout << "bound: " << m_bound << " lit: " << l << " coeff: " << ineq.coeff(i) << "\n";);
TRACE("pb", tout << "bound: " << m_bound << " lit: " << l << " coeff: " << ineq.coeff(i) << "\n";);
}
}
@ -995,11 +995,11 @@ namespace pb {
consequent.neg();
process_antecedent(consequent, 1);
}
TRACE("ba", tout << consequent << " " << js << "\n";);
TRACE("pb", tout << consequent << " " << js << "\n";);
unsigned idx = s().m_trail.size() - 1;
do {
TRACE("ba", s().display_justification(tout << "process consequent: " << consequent << " : ", js) << "\n";
TRACE("pb", s().display_justification(tout << "process consequent: " << consequent << " : ", js) << "\n";
if (consequent != sat::null_literal) { active2pb(m_A); display(tout, m_A, true); }
);
@ -1069,7 +1069,7 @@ namespace pb {
}
else {
SASSERT(k > c);
TRACE("ba", tout << "visited: " << l << "\n";);
TRACE("pb", tout << "visited: " << l << "\n";);
k -= c;
}
}
@ -1118,7 +1118,7 @@ namespace pb {
}
}
if (idx == 0) {
TRACE("ba", tout << "there is no consequent\n";);
TRACE("pb", tout << "there is no consequent\n";);
goto bail_out;
}
--idx;
@ -1131,7 +1131,7 @@ namespace pb {
js = s().m_justification[v];
}
while (m_num_marks > 0 && !m_overflow);
TRACE("ba", active2pb(m_A); display(tout, m_A, true););
TRACE("pb", active2pb(m_A); display(tout, m_A, true););
// TBD: check if this is useful
if (!m_overflow && consequent != sat::null_literal) {
@ -1143,7 +1143,7 @@ namespace pb {
}
bail_out:
TRACE("ba", tout << "bail " << m_overflow << "\n";);
TRACE("pb", tout << "bail " << m_overflow << "\n";);
if (m_overflow) {
++m_stats.m_num_overflow;
m_overflow = false;
@ -1199,23 +1199,23 @@ namespace pb {
}
}
if (slack >= 0) {
TRACE("ba", tout << "slack is non-negative\n";);
TRACE("pb", tout << "slack is non-negative\n";);
IF_VERBOSE(20, verbose_stream() << "(sat.card slack: " << slack << " skipped: " << num_skipped << ")\n";);
return false;
}
if (m_overflow) {
TRACE("ba", tout << "overflow\n";);
TRACE("pb", tout << "overflow\n";);
return false;
}
if (m_lemma[0] == sat::null_literal) {
if (m_lemma.size() == 1) {
s().set_conflict(sat::justification(0));
}
TRACE("ba", tout << "no asserting literal\n";);
TRACE("pb", tout << "no asserting literal\n";);
return false;
}
TRACE("ba", tout << m_lemma << "\n";);
TRACE("pb", tout << m_lemma << "\n";);
if (get_config().m_drat && m_solver) {
s().m_drat.add(m_lemma, sat::status::th(true, get_id()));
@ -1224,7 +1224,7 @@ namespace pb {
s().m_lemma.reset();
s().m_lemma.append(m_lemma);
for (unsigned i = 1; i < m_lemma.size(); ++i) {
CTRACE("ba", s().is_marked(m_lemma[i].var()), tout << "marked: " << m_lemma[i] << "\n";);
CTRACE("pb", s().is_marked(m_lemma[i].var()), tout << "marked: " << m_lemma[i] << "\n";);
s().mark(m_lemma[i].var());
}
return true;
@ -1346,11 +1346,11 @@ namespace pb {
}
solver::solver(ast_manager& m, sat::sat_internalizer& si, euf::theory_id id)
: euf::th_solver(m, symbol("ba"), id),
: euf::th_solver(m, symbol("pb"), id),
si(si), m_pb(m),
m_lookahead(nullptr),
m_constraint_id(0), m_ba(*this), m_sort(m_ba) {
TRACE("ba", tout << this << "\n";);
TRACE("pb", tout << this << "\n";);
m_num_propagations_since_pop = 0;
}
@ -1418,6 +1418,8 @@ namespace pb {
}
else if (lit == sat::null_literal) {
init_watch(*c);
if (c->is_pb())
validate_watch(c->to_pb(), sat::null_literal);
}
else {
if (m_solver) m_solver->set_external(lit.var());
@ -1569,7 +1571,7 @@ namespace pb {
}
void solver::get_antecedents(literal l, pbc const& p, literal_vector& r) {
TRACE("ba", display(tout << l << " level: " << s().scope_lvl() << " ", p, true););
TRACE("pb", display(tout << l << " level: " << s().scope_lvl() << " ", p, true););
SASSERT(p.lit() == sat::null_literal || value(p.lit()) == l_true);
if (p.lit() != sat::null_literal) {
@ -1621,10 +1623,10 @@ namespace pb {
if (j < p.num_watch()) {
j = p.num_watch();
}
CTRACE("ba", coeff == 0, display(tout << l << " coeff: " << coeff << "\n", p, true););
CTRACE("pb", coeff == 0, display(tout << l << " coeff: " << coeff << "\n", p, true););
if (_debug_conflict) {
std::cout << "coeff " << coeff << "\n";
IF_VERBOSE(0, verbose_stream() << "coeff " << coeff << "\n";);
}
SASSERT(coeff > 0);
@ -1672,7 +1674,7 @@ namespace pb {
for (unsigned i = 0; !found && i < c.k(); ++i) {
found = c[i] == l;
}
CTRACE("ba",!found, s().display(tout << l << ":" << c << "\n"););
CTRACE("pb",!found, s().display(tout << l << ":" << c << "\n"););
SASSERT(found););
VERIFY(c.lit() == sat::null_literal || value(c.lit()) != l_false);
@ -1712,7 +1714,7 @@ namespace pb {
}
void solver::remove_constraint(constraint& c, char const* reason) {
TRACE("ba", display(tout << "remove ", c, true) << " " << reason << "\n";);
TRACE("pb", display(tout << "remove ", c, true) << " " << reason << "\n";);
IF_VERBOSE(21, display(verbose_stream() << "remove " << reason << " ", c, true););
c.nullify_tracking_literal(*this);
clear_watch(c);
@ -1886,7 +1888,7 @@ namespace pb {
}
void solver::gc_half(char const* st_name) {
TRACE("ba", tout << "gc\n";);
TRACE("pb", tout << "gc\n";);
unsigned sz = m_learned.size();
unsigned new_sz = sz/2;
unsigned removed = 0;
@ -1933,7 +1935,7 @@ namespace pb {
// literal is assigned to false.
unsigned sz = c.size();
unsigned bound = c.k();
TRACE("ba", tout << "assign: " << c.lit() << ": " << ~alit << "@" << lvl(~alit) << " " << c << "\n";);
TRACE("pb", tout << "assign: " << c.lit() << ": " << ~alit << "@" << lvl(~alit) << " " << c << "\n";);
SASSERT(0 < bound && bound <= sz);
if (bound == sz) {
@ -1971,7 +1973,7 @@ namespace pb {
// conflict
if (bound != index && value(c[bound]) == l_false) {
TRACE("ba", tout << "conflict " << c[bound] << " " << alit << "\n";);
TRACE("pb", tout << "conflict " << c[bound] << " " << alit << "\n";);
if (c.lit() != sat::null_literal && value(c.lit()) == l_undef) {
if (index + 1 < bound) c.swap(index, bound - 1);
assign(c, ~c.lit());
@ -1985,7 +1987,7 @@ namespace pb {
c.swap(index, bound);
}
// TRACE("ba", tout << "no swap " << index << " " << alit << "\n";);
// TRACE("pb", tout << "no swap " << index << " " << alit << "\n";);
// there are no literals to swap with,
// prepare for unit propagation by swapping the false literal into
// position bound. Then literals in positions 0..bound-1 have to be
@ -2256,7 +2258,7 @@ namespace pb {
SASSERT(c.lit() == sat::null_literal || c.is_watched(*this, c.lit()));
// pre-condition is that the literals, except c.lit(), in c are unwatched.
if (c.id() == _bad_id) std::cout << "recompile: " << c << "\n";
//if (c.id() == _bad_id) std::cout << "recompile: " << c << "\n";
m_weights.resize(2*s().num_vars(), 0);
for (literal l : c) {
++m_weights[l.index()];
@ -2351,7 +2353,7 @@ namespace pb {
}
if (!all_units) {
TRACE("ba", tout << "replacing by pb: " << c << "\n";);
TRACE("pb", tout << "replacing by pb: " << c << "\n";);
m_wlits.reset();
for (unsigned i = 0; i < sz; ++i) {
m_wlits.push_back(wliteral(coeffs[i], c[i]));
@ -2699,7 +2701,7 @@ namespace pb {
}
}
++m_stats.m_num_big_strengthenings;
constraint* c = add_pb_ge(sat::null_literal, wlits, b, p.learned());
add_pb_ge(sat::null_literal, wlits, b, p.learned());
p.set_removed();
return;
}
@ -2914,13 +2916,13 @@ namespace pb {
SASSERT(&c1 != &c2);
if (subsumes(c1, c2, slit)) {
if (slit.empty()) {
TRACE("ba", tout << "subsume cardinality\n" << c1 << "\n" << c2 << "\n";);
TRACE("pb", tout << "subsume cardinality\n" << c1 << "\n" << c2 << "\n";);
remove_constraint(c2, "subsumed");
++m_stats.m_num_pb_subsumes;
set_non_learned(c1);
}
else {
TRACE("ba", tout << "self subsume cardinality\n";);
TRACE("pb", tout << "self subsume cardinality\n";);
IF_VERBOSE(11,
verbose_stream() << "self-subsume cardinality\n";
verbose_stream() << c1 << "\n";
@ -2952,7 +2954,7 @@ namespace pb {
// self-subsumption is TBD
}
else {
TRACE("ba", tout << "remove\n" << c1 << "\n" << c2 << "\n";);
TRACE("pb", tout << "remove\n" << c1 << "\n" << c2 << "\n";);
removed_clauses.push_back(&c2);
++m_stats.m_num_clause_subsumes;
set_non_learned(c1);
@ -3284,7 +3286,7 @@ namespace pb {
val += wl.first;
}
}
CTRACE("ba", val >= 0, active2pb(m_A); display(tout, m_A, true););
CTRACE("pb", val >= 0, active2pb(m_A); display(tout, m_A, true););
return val < 0;
}
@ -3297,7 +3299,7 @@ namespace pb {
if (!is_false(wl.second))
k += wl.first;
}
CTRACE("ba", k > 0, display(tout, ineq, true););
CTRACE("pb", k > 0, display(tout, ineq, true););
return k <= 0;
}
@ -3356,7 +3358,7 @@ namespace pb {
return nullptr;
}
constraint* c = add_pb_ge(sat::null_literal, m_wlits, m_bound, true);
TRACE("ba", if (c) display(tout, *c, true););
TRACE("pb", if (c) display(tout, *c, true););
++m_stats.m_num_lemmas;
return c;
}
@ -3587,7 +3589,7 @@ namespace pb {
s0.assign_scoped(l2);
s0.assign_scoped(l3);
lbool is_sat = s0.check();
TRACE("ba", s0.display(tout << "trying sat encoding"););
TRACE("pb", s0.display(tout << "trying sat encoding"););
if (is_sat == l_false) return true;
IF_VERBOSE(0,
@ -3698,11 +3700,11 @@ namespace pb {
bool solver::validate_conflict(literal_vector const& lits, ineq& p) {
for (literal l : lits) {
if (value(l) != l_false) {
TRACE("ba", tout << "literal " << l << " is not false\n";);
TRACE("pb", tout << "literal " << l << " is not false\n";);
return false;
}
if (!p.contains(l)) {
TRACE("ba", tout << "lemma contains literal " << l << " not in inequality\n";);
TRACE("pb", tout << "lemma contains literal " << l << " not in inequality\n";);
return false;
}
}
@ -3713,7 +3715,7 @@ namespace pb {
value += coeff;
}
}
CTRACE("ba", value >= p.m_k, tout << "slack: " << value << " bound " << p.m_k << "\n";
CTRACE("pb", value >= p.m_k, tout << "slack: " << value << " bound " << p.m_k << "\n";
display(tout, p);
tout << lits << "\n";);
return value < p.m_k;

View file

@ -341,7 +341,6 @@ namespace q {
return false;
if (ctx.s().inconsistent())
return true;
TRACE("q", c.display(ctx, tout) << "\n";);
unsigned idx = UINT_MAX;
m_evidence.reset();
lbool ev = m_eval(binding, c, idx, m_evidence);
@ -611,39 +610,49 @@ namespace q {
return ctx.get_config().m_ematching && propagate(false);
}
void ematch::propagate(clause& c, bool flush, bool& propagated) {
ptr_buffer<binding> to_remove;
binding* b = c.m_bindings;
if (!b)
return;
do {
if (propagate(true, b->m_nodes, b->m_max_generation, c, propagated))
to_remove.push_back(b);
else if (flush) {
instantiate(*b);
to_remove.push_back(b);
propagated = true;
}
b = b->next();
}
while (b != c.m_bindings);
for (auto* b : to_remove) {
SASSERT(binding::contains(c.m_bindings, b));
binding::remove_from(c.m_bindings, b);
binding::detach(b);
ctx.push(insert_binding(ctx, c, b));
}
}
bool ematch::propagate(bool flush) {
m_mam->propagate();
bool propagated = flush_prop_queue();
if (m_qhead >= m_clause_queue.size())
return m_inst_queue.propagate() || propagated;
ctx.push(value_trail<unsigned>(m_qhead));
ptr_buffer<binding> to_remove;
for (; m_qhead < m_clause_queue.size() && m.inc(); ++m_qhead) {
unsigned idx = m_clause_queue[m_qhead];
clause& c = *m_clauses[idx];
binding* b = c.m_bindings;
if (!b)
continue;
do {
if (propagate(true, b->m_nodes, b->m_max_generation, c, propagated))
to_remove.push_back(b);
else if (flush) {
instantiate(*b);
to_remove.push_back(b);
propagated = true;
}
b = b->next();
}
while (b != c.m_bindings);
for (auto* b : to_remove) {
SASSERT(binding::contains(c.m_bindings, b));
binding::remove_from(c.m_bindings, b);
binding::detach(b);
ctx.push(insert_binding(ctx, c, b));
if (flush) {
for (auto* c : m_clauses)
propagate(*c, flush, propagated);
}
else {
if (m_qhead >= m_clause_queue.size())
return m_inst_queue.propagate() || propagated;
ctx.push(value_trail<unsigned>(m_qhead));
for (; m_qhead < m_clause_queue.size() && m.inc(); ++m_qhead) {
unsigned idx = m_clause_queue[m_qhead];
clause& c = *m_clauses[idx];
propagate(c, flush, propagated);
}
to_remove.reset();
}
m_clause_in_queue.reset();
m_node_in_queue.reset();
@ -662,15 +671,18 @@ namespace q {
if (propagate(false))
return true;
for (unsigned i = 0; i < m_clauses.size(); ++i)
if (m_clauses[i]->m_bindings)
if (m_clauses[i]->m_bindings)
insert_clause_in_queue(i);
if (propagate(true))
return true;
if (m_inst_queue.lazy_propagate())
return true;
for (unsigned i = 0; i < m_clauses.size(); ++i)
if (m_clauses[i]->m_bindings)
if (m_clauses[i]->m_bindings) {
IF_VERBOSE(0, verbose_stream() << "missed propagation " << i << "\n");
TRACE("q", display(tout << "missed propagation\n"));
break;
}
TRACE("q", tout << "no more propagation\n";);
return false;

View file

@ -121,6 +121,7 @@ namespace q {
void propagate(bool is_conflict, unsigned idx, sat::ext_justification_idx j_idx);
bool propagate(bool flush);
void propagate(clause& c, bool flush, bool& propagated);
expr_ref_vector m_new_defs;
proof_ref_vector m_new_proofs;

View file

@ -104,7 +104,7 @@ namespace q {
public:
unsigned char operator()(func_decl * lbl) {
unsigned lbl_id = lbl->get_decl_id();
unsigned lbl_id = lbl->get_small_id();
if (lbl_id >= m_lbl2hash.size())
m_lbl2hash.resize(lbl_id + 1, -1);
if (m_lbl2hash[lbl_id] == -1) {
@ -2868,7 +2868,7 @@ namespace q {
SASSERT(first_idx < mp->get_num_args());
app * p = to_app(mp->get_arg(first_idx));
func_decl * lbl = p->get_decl();
unsigned lbl_id = lbl->get_decl_id();
unsigned lbl_id = lbl->get_small_id();
m_trees.reserve(lbl_id+1, nullptr);
if (m_trees[lbl_id] == nullptr) {
m_trees[lbl_id] = m_compiler.mk_tree(qa, mp, first_idx, false);
@ -2898,7 +2898,7 @@ namespace q {
}
code_tree * get_code_tree_for(func_decl * lbl) const {
unsigned lbl_id = lbl->get_decl_id();
unsigned lbl_id = lbl->get_small_id();
if (lbl_id < m_trees.size())
return m_trees[lbl_id];
else
@ -3112,11 +3112,11 @@ namespace q {
}
bool is_plbl(func_decl * lbl) const {
unsigned lbl_id = lbl->get_decl_id();
unsigned lbl_id = lbl->get_small_id();
return lbl_id < m_is_plbl.size() && m_is_plbl[lbl_id];
}
bool is_clbl(func_decl * lbl) const {
unsigned lbl_id = lbl->get_decl_id();
unsigned lbl_id = lbl->get_small_id();
return lbl_id < m_is_clbl.size() && m_is_clbl[lbl_id];
}
@ -3129,7 +3129,7 @@ namespace q {
}
void update_clbls(func_decl * lbl) {
unsigned lbl_id = lbl->get_decl_id();
unsigned lbl_id = lbl->get_small_id();
m_is_clbl.reserve(lbl_id+1, false);
TRACE("trigger_bug", tout << "update_clbls: " << lbl->get_name() << " is already clbl: " << m_is_clbl[lbl_id] << "\n";);
TRACE("mam_bug", tout << "update_clbls: " << lbl->get_name() << " is already clbl: " << m_is_clbl[lbl_id] << "\n";);
@ -3169,7 +3169,7 @@ namespace q {
}
void update_plbls(func_decl * lbl) {
unsigned lbl_id = lbl->get_decl_id();
unsigned lbl_id = lbl->get_small_id();
m_is_plbl.reserve(lbl_id+1, false);
TRACE("trigger_bug", tout << "update_plbls: " << lbl->get_name() << " is already plbl: " << m_is_plbl[lbl_id] << ", lbl_id: " << lbl_id << "\n";
tout << "mam: " << this << "\n";);
@ -3698,7 +3698,7 @@ namespace q {
app * p = to_app(mp->get_arg(0));
func_decl * lbl = p->get_decl();
if (!m_egraph.enodes_of(lbl).empty()) {
unsigned lbl_id = lbl->get_decl_id();
unsigned lbl_id = lbl->get_small_id();
m_tmp_trees.reserve(lbl_id+1, 0);
if (m_tmp_trees[lbl_id] == 0) {
m_tmp_trees[lbl_id] = m_compiler.mk_tree(qa, mp, 0, false);
@ -3711,7 +3711,7 @@ namespace q {
}
for (func_decl * lbl : m_tmp_trees_to_delete) {
unsigned lbl_id = lbl->get_decl_id();
unsigned lbl_id = lbl->get_small_id();
code_tree * tmp_tree = m_tmp_trees[lbl_id];
SASSERT(tmp_tree != 0);
SASSERT(!m_egraph.enodes_of(lbl).empty());
@ -3843,7 +3843,7 @@ namespace q {
unsigned h = m_lbl_hasher(lbl);
TRACE("trigger_bug", tout << "lbl: " << lbl->get_name() << " is_clbl(lbl): " << is_clbl(lbl)
<< ", is_plbl(lbl): " << is_plbl(lbl) << ", h: " << h << "\n";
tout << "lbl_id: " << lbl->get_decl_id() << "\n";);
tout << "lbl_id: " << lbl->get_small_id() << "\n";);
if (is_clbl(lbl))
update_lbls(n, h);
if (is_plbl(lbl))

View file

@ -302,9 +302,10 @@ namespace q {
return md->v2t[md->values[j]];
};
#if 0
for (unsigned j = 0; j < sz; ++j)
std::cout << mk_pp(md->values[j], m) << "\n";
#endif
expr* arg = t->get_arg(i);

View file

@ -333,6 +333,11 @@ namespace recfun {
return found;
}
bool solver::is_beta_redex(euf::enode* p, euf::enode* n) const {
return is_defined(p) || is_case_pred(p);
}
bool solver::add_dep(euf::enode* n, top_sort<euf::enode>& dep) {
if (n->num_args() == 0)
dep.insert(n, nullptr);

View file

@ -108,6 +108,7 @@ namespace recfun {
bool is_shared(euf::theory_var v) const override { return true; }
void init_search() override {}
bool should_research(sat::literal_vector const& core) override;
bool is_beta_redex(euf::enode* p, euf::enode* n) const override;
void add_assumptions(sat::literal_set& assumptions) override;
bool tracking_assumptions() override { return true; }
};

View file

@ -125,8 +125,8 @@ namespace euf {
pop_core(n);
}
sat::status th_euf_solver::mk_status() {
return sat::status::th(m_is_redundant, get_id());
sat::status th_euf_solver::mk_status(sat::proof_hint const* ps) {
return sat::status::th(m_is_redundant, get_id(), ps);
}
bool th_euf_solver::add_unit(sat::literal lit) {
@ -149,6 +149,11 @@ namespace euf {
return add_clause(2, lits);
}
bool th_euf_solver::add_clause(sat::literal a, sat::literal b, sat::proof_hint const* ps) {
sat::literal lits[2] = { a, b };
return add_clause(2, lits, ps);
}
bool th_euf_solver::add_clause(sat::literal a, sat::literal b, sat::literal c) {
sat::literal lits[3] = { a, b, c };
return add_clause(3, lits);
@ -159,12 +164,12 @@ namespace euf {
return add_clause(4, lits);
}
bool th_euf_solver::add_clause(unsigned n, sat::literal* lits) {
bool th_euf_solver::add_clause(unsigned n, sat::literal* lits, sat::proof_hint const* ps) {
bool was_true = false;
for (unsigned i = 0; i < n; ++i)
was_true |= is_true(lits[i]);
ctx.add_root(n, lits);
s().add_clause(n, lits, mk_status());
s().add_clause(n, lits, mk_status(ps));
return !was_true;
}
@ -221,37 +226,48 @@ namespace euf {
return ctx.s().rand()();
}
size_t th_explain::get_obj_size(unsigned num_lits, unsigned num_eqs) {
return sat::constraint_base::obj_size(sizeof(th_explain) + sizeof(sat::literal) * num_lits + sizeof(enode_pair) * num_eqs);
size_t th_explain::get_obj_size(unsigned num_lits, unsigned num_eqs, sat::proof_hint const* pma) {
return sat::constraint_base::obj_size(sizeof(th_explain) + sizeof(sat::literal) * num_lits + sizeof(enode_pair) * num_eqs + (pma?pma->to_string().length()+1:1));
}
th_explain::th_explain(unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode_pair const& p) {
th_explain::th_explain(unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode_pair const& p, sat::proof_hint const* pma) {
m_consequent = c;
m_eq = p;
m_num_literals = n_lits;
m_num_eqs = n_eqs;
m_literals = reinterpret_cast<literal*>(reinterpret_cast<char*>(this) + sizeof(th_explain));
for (unsigned i = 0; i < n_lits; ++i)
char * base_ptr = reinterpret_cast<char*>(this) + sizeof(th_explain);
m_literals = reinterpret_cast<literal*>(base_ptr);
unsigned i;
for (i = 0; i < n_lits; ++i)
m_literals[i] = lits[i];
m_eqs = reinterpret_cast<enode_pair*>(reinterpret_cast<char*>(this) + sizeof(th_explain) + sizeof(literal) * n_lits);
for (unsigned i = 0; i < n_eqs; ++i)
m_eqs[i] = eqs[i];
base_ptr += sizeof(literal) * n_lits;
m_eqs = reinterpret_cast<enode_pair*>(base_ptr);
for (i = 0; i < n_eqs; ++i)
m_eqs[i] = eqs[i];
base_ptr += sizeof(enode_pair) * n_eqs;
m_pragma = reinterpret_cast<char*>(base_ptr);
i = 0;
if (pma) {
std::string s = pma->to_string();
for (i = 0; s[i]; ++i)
m_pragma[i] = s[i];
}
m_pragma[i] = 0;
}
th_explain* th_explain::mk(th_euf_solver& th, unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode* x, enode* y) {
th_explain* th_explain::mk(th_euf_solver& th, unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode* x, enode* y, sat::proof_hint const* pma) {
region& r = th.ctx.get_region();
void* mem = r.allocate(get_obj_size(n_lits, n_eqs));
void* mem = r.allocate(get_obj_size(n_lits, n_eqs, pma));
sat::constraint_base::initialize(mem, &th);
return new (sat::constraint_base::ptr2mem(mem)) th_explain(n_lits, lits, n_eqs, eqs, c, enode_pair(x, y));
}
th_explain* th_explain::propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, sat::literal consequent) {
return mk(th, lits.size(), lits.data(), eqs.size(), eqs.data(), consequent, nullptr, nullptr);
th_explain* th_explain::propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, sat::literal consequent, sat::proof_hint const* pma) {
return mk(th, lits.size(), lits.data(), eqs.size(), eqs.data(), consequent, nullptr, nullptr, pma);
}
th_explain* th_explain::propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, euf::enode* x, euf::enode* y) {
return mk(th, lits.size(), lits.data(), eqs.size(), eqs.data(), sat::null_literal, x, y);
th_explain* th_explain::propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, euf::enode* x, euf::enode* y, sat::proof_hint const* pma) {
return mk(th, lits.size(), lits.data(), eqs.size(), eqs.data(), sat::null_literal, x, y, pma);
}
th_explain* th_explain::propagate(th_euf_solver& th, sat::literal lit, euf::enode* x, euf::enode* y) {
@ -293,6 +309,8 @@ namespace euf {
out << "--> " << m_consequent;
if (m_eq.first != nullptr)
out << "--> " << m_eq.first->get_expr_id() << " == " << m_eq.second->get_expr_id();
if (m_pragma != nullptr)
out << " p " << m_pragma;
return out;
}

View file

@ -127,6 +127,13 @@ namespace euf {
*/
virtual bool is_shared(theory_var v) const { return false; }
/**
\brief Determine if argument n of parent p is a beta redex position
*/
virtual bool is_beta_redex(euf::enode* p, euf::enode* n) const { return false; }
sat::status status() const { return sat::status::th(m_is_redundant, get_id()); }
};
@ -143,15 +150,16 @@ namespace euf {
region& get_region();
sat::status mk_status();
sat::status mk_status(sat::proof_hint const* ps = nullptr);
bool add_unit(sat::literal lit);
bool add_units(sat::literal_vector const& lits);
bool add_clause(sat::literal lit) { return add_unit(lit); }
bool add_clause(sat::literal a, sat::literal b);
bool add_clause(sat::literal a, sat::literal b, sat::proof_hint const* ps);
bool add_clause(sat::literal a, sat::literal b, sat::literal c);
bool add_clause(sat::literal a, sat::literal b, sat::literal c, sat::literal d);
bool add_clause(sat::literal_vector const& lits) { return add_clause(lits.size(), lits.data()); }
bool add_clause(unsigned n, sat::literal* lits);
bool add_clause(sat::literal_vector const& lits, sat::proof_hint const* ps = nullptr) { return add_clause(lits.size(), lits.data(), ps); }
bool add_clause(unsigned n, sat::literal* lits, sat::proof_hint const* ps = nullptr);
void add_equiv(sat::literal a, sat::literal b);
void add_equiv_and(sat::literal a, sat::literal_vector const& bs);
@ -213,15 +221,16 @@ namespace euf {
* that retrieve literals on demand.
*/
class th_explain {
sat::literal m_consequent { sat::null_literal }; // literal consequent for propagations
enode_pair m_eq { enode_pair() }; // equality consequent for propagations
sat::literal m_consequent = sat::null_literal; // literal consequent for propagations
enode_pair m_eq = enode_pair(); // equality consequent for propagations
unsigned m_num_literals;
unsigned m_num_eqs;
sat::literal* m_literals;
enode_pair* m_eqs;
static size_t get_obj_size(unsigned num_lits, unsigned num_eqs);
th_explain(unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode_pair const& eq);
static th_explain* mk(th_euf_solver& th, unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode* x, enode* y);
char* m_pragma = nullptr;
static size_t get_obj_size(unsigned num_lits, unsigned num_eqs, sat::proof_hint const* pma);
th_explain(unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode_pair const& eq, sat::proof_hint const* pma = nullptr);
static th_explain* mk(th_euf_solver& th, unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode* x, enode* y, sat::proof_hint const* pma = nullptr);
public:
static th_explain* conflict(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs);
@ -232,8 +241,8 @@ namespace euf {
static th_explain* conflict(th_euf_solver& th, sat::literal lit, euf::enode* x, euf::enode* y);
static th_explain* conflict(th_euf_solver& th, euf::enode* x, euf::enode* y);
static th_explain* propagate(th_euf_solver& th, sat::literal lit, euf::enode* x, euf::enode* y);
static th_explain* propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, sat::literal consequent);
static th_explain* propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, euf::enode* x, euf::enode* y);
static th_explain* propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, sat::literal consequent, sat::proof_hint const* pma = nullptr);
static th_explain* propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, euf::enode* x, euf::enode* y, sat::proof_hint const* pma = nullptr);
sat::ext_constraint_idx to_index() const {
return sat::constraint_base::mem2base(this);
@ -268,6 +277,8 @@ namespace euf {
enode_pair eq_consequent() const { return m_eq; }
sat::proof_hint const* get_pragma() const { return nullptr; } //*m_pragma ? m_pragma : nullptr; }
};

View file

@ -56,6 +56,18 @@ namespace user_solver {
void solver::register_cb(expr* e) {
add_expr(e);
}
void solver::next_split_cb(expr* e, unsigned idx, lbool phase) {
if (e == nullptr) {
m_next_split_expr = nullptr;
return;
}
force_push();
ctx.internalize(e, false);
m_next_split_expr = e;
m_next_split_idx = idx;
m_next_split_phase = phase;
}
sat::check_result solver::check() {
if (!(bool)m_final_eh)
@ -72,6 +84,41 @@ namespace user_solver {
m_id2justification.setx(v, sat::literal_vector(num_lits, jlits), sat::literal_vector());
m_fixed_eh(m_user_context, this, var2expr(v), value);
}
bool solver::decide(sat::bool_var& var, lbool& phase) {
if (!m_decide_eh)
return false;
euf::enode* original_enode = bool_var2enode(var);
if (!is_attached_to_var(original_enode))
return false;
unsigned new_bit = 0; // ignored; currently no bv-support
expr* e = bool_var2expr(var);
m_decide_eh(m_user_context, this, &e, &new_bit, &phase);
euf::enode* new_enode = ctx.get_enode(e);
if (original_enode == new_enode)
return false;
var = new_enode->bool_var();
return true;
}
bool solver::get_case_split(sat::bool_var& var, lbool &phase){
if (!m_next_split_expr)
return false;
euf::enode* n = ctx.get_enode(m_next_split_expr);
var = n->bool_var();
phase = m_next_split_phase;
m_next_split_expr = nullptr;
return true;
}
void solver::asserted(sat::literal lit) {
if (!m_fixed_eh)

View file

@ -56,24 +56,28 @@ namespace user_solver {
void reset() { memset(this, 0, sizeof(*this)); }
};
void* m_user_context;
user_propagator::push_eh_t m_push_eh;
user_propagator::pop_eh_t m_pop_eh;
user_propagator::fresh_eh_t m_fresh_eh;
user_propagator::final_eh_t m_final_eh;
user_propagator::fixed_eh_t m_fixed_eh;
user_propagator::eq_eh_t m_eq_eh;
user_propagator::eq_eh_t m_diseq_eh;
user_propagator::created_eh_t m_created_eh;
void* m_user_context;
user_propagator::push_eh_t m_push_eh = nullptr;
user_propagator::pop_eh_t m_pop_eh = nullptr;
user_propagator::fresh_eh_t m_fresh_eh = nullptr;
user_propagator::final_eh_t m_final_eh = nullptr;
user_propagator::fixed_eh_t m_fixed_eh = nullptr;
user_propagator::eq_eh_t m_eq_eh = nullptr;
user_propagator::eq_eh_t m_diseq_eh = nullptr;
user_propagator::created_eh_t m_created_eh = nullptr;
user_propagator::decide_eh_t m_decide_eh = nullptr;
user_propagator::context_obj* m_api_context = nullptr;
unsigned m_qhead = 0;
vector<prop_info> m_prop;
unsigned_vector m_prop_lim;
vector<sat::literal_vector> m_id2justification;
sat::literal_vector m_lits;
euf::enode_pair_vector m_eqs;
unsigned_vector m_fixed_ids;
stats m_stats;
unsigned m_qhead = 0;
vector<prop_info> m_prop;
unsigned_vector m_prop_lim;
vector<sat::literal_vector> m_id2justification;
sat::literal_vector m_lits;
euf::enode_pair_vector m_eqs;
unsigned_vector m_fixed_ids;
stats m_stats;
expr* m_next_split_expr = nullptr;
unsigned m_next_split_idx;
lbool m_next_split_phase;
struct justification {
unsigned m_propagation_index { 0 };
@ -94,7 +98,7 @@ namespace user_solver {
void propagate_consequence(prop_info const& prop);
void propagate_new_fixed(prop_info const& prop);
void validate_propagation();
void validate_propagation();
bool visit(expr* e) override;
bool visited(expr* e) override;
@ -126,14 +130,19 @@ namespace user_solver {
void register_eq(user_propagator::eq_eh_t& eq_eh) { m_eq_eh = eq_eh; }
void register_diseq(user_propagator::eq_eh_t& diseq_eh) { m_diseq_eh = diseq_eh; }
void register_created(user_propagator::created_eh_t& created_eh) { m_created_eh = created_eh; }
void register_decide(user_propagator::decide_eh_t& decide_eh) { m_decide_eh = decide_eh; }
bool has_fixed() const { return (bool)m_fixed_eh; }
void propagate_cb(unsigned num_fixed, expr* const* fixed_ids, unsigned num_eqs, expr* const* lhs, expr* const* rhs, expr* conseq) override;
void register_cb(expr* e) override;
void next_split_cb(expr* e, unsigned idx, lbool phase) override;
void new_fixed_eh(euf::theory_var v, expr* value, unsigned num_lits, sat::literal const* jlits);
bool decide(sat::bool_var& var, lbool& phase) override;
bool get_case_split(sat::bool_var& var, lbool &phase) override;
void asserted(sat::literal lit) override;
sat::check_result check() override;
void push_core() override;

View file

@ -894,16 +894,36 @@ struct goal2sat::imp : public sat::sat_internalizer {
m_result_stack.pop_back();
}
struct scoped_reset {
imp& i;
scoped_reset(imp& i) :i(i) {}
~scoped_reset() {
i.m_interface_vars.reset();
i.m_app2lit.reset();
i.m_lit2app.reset();
}
};
void operator()(unsigned n, expr* const* fmls) {
scoped_reset _reset(*this);
// collect_boolean_interface(g, m_interface_vars);
for (unsigned i = 0; i < n; ++i)
process(fmls[i]);
}
void assumptions(unsigned n, expr* const* fmls) {
scoped_reset _reset(*this);
// collect_boolean_interface(g, m_interface_vars);
for (unsigned i = 0; i < n; ++i) {
expr* f = fmls[i];
expr* f1 = f;
bool sign = m.is_not(f, f1);
insert_dep(f, f1, sign);
}
}
void operator()(goal const & g) {
struct scoped_reset {
imp& i;
scoped_reset(imp& i) :i(i) {}
~scoped_reset() {
i.m_interface_vars.reset();
i.m_app2lit.reset();
i.m_lit2app.reset();
}
};
scoped_reset _reset(*this);
collect_boolean_interface(g, m_interface_vars);
unsigned size = g.size();
@ -1002,16 +1022,30 @@ void goal2sat::collect_param_descrs(param_descrs & r) {
r.insert("ite_extra", CPK_BOOL, "(default: true) add redundant clauses (that improve unit propagation) when encoding if-then-else formulas");
}
void goal2sat::operator()(goal const & g, params_ref const & p, sat::solver_core & t, atom2bool_var & m, dep2asm_map& dep2asm, bool default_external) {
void goal2sat::init(ast_manager& m, params_ref const & p, sat::solver_core & t, atom2bool_var & map, dep2asm_map& dep2asm, bool default_external) {
if (!m_imp) {
m_imp = alloc(imp, g.m(), p, t, m, dep2asm, default_external);
m_imp = alloc(imp, m, p, t, map, dep2asm, default_external);
for (unsigned i = 0; i < m_scopes; ++i)
m_imp->user_push();
}
}
void goal2sat::operator()(goal const & g, params_ref const & p, sat::solver_core & t, atom2bool_var & m, dep2asm_map& dep2asm, bool default_external) {
init(g.m(), p, t, m, dep2asm, default_external);
(*m_imp)(g);
}
void goal2sat::operator()(ast_manager& m, unsigned n, expr* const* fmls, params_ref const & p, sat::solver_core & t, atom2bool_var & map, dep2asm_map& dep2asm, bool default_external) {
init(m, p, t, map, dep2asm, default_external);
(*m_imp)(n, fmls);
}
void goal2sat::assumptions(ast_manager& m, unsigned n, expr* const* fmls, params_ref const & p, sat::solver_core & t, atom2bool_var & map, dep2asm_map& dep2asm, bool default_external) {
init(m, p, t, map, dep2asm, default_external);
m_imp->assumptions(n, fmls);
}
void goal2sat::get_interpreted_funs(func_decl_ref_vector& funs) {
if (m_imp)
funs.append(m_imp->interpreted_funs());

View file

@ -34,15 +34,19 @@ Notes:
#include "sat/smt/sat_internalizer.h"
class goal2sat {
public:
typedef obj_map<expr, sat::literal> dep2asm_map;
private:
struct imp;
imp * m_imp;
unsigned m_scopes { 0 };
unsigned m_scopes = 0;
void init(ast_manager& m, params_ref const & p, sat::solver_core & t, atom2bool_var & map, dep2asm_map& dep2asm, bool default_external);
public:
goal2sat();
~goal2sat();
typedef obj_map<expr, sat::literal> dep2asm_map;
static void collect_param_descrs(param_descrs & r);
@ -60,6 +64,10 @@ public:
*/
void operator()(goal const & g, params_ref const & p, sat::solver_core & t, atom2bool_var & m, dep2asm_map& dep2asm, bool default_external = false);
void operator()(ast_manager& m, unsigned n, expr* const* fmls, params_ref const & p, sat::solver_core & t, atom2bool_var & map, dep2asm_map& dep2asm, bool default_external = false);
void assumptions(ast_manager& m, unsigned n, expr* const* fmls, params_ref const & p, sat::solver_core & t, atom2bool_var & map, dep2asm_map& dep2asm, bool default_external = false);
void get_interpreted_funs(func_decl_ref_vector& funs);
bool has_interpreted_funs() const;