3
0
Fork 0
mirror of https://github.com/Z3Prover/z3 synced 2025-04-24 09:35:32 +00:00

merge with master

Signed-off-by: Nikolaj Bjorner <nbjorner@microsoft.com>
This commit is contained in:
Nikolaj Bjorner 2022-04-16 18:30:03 +02:00
commit 3533bf486f
223 changed files with 7175 additions and 2167 deletions

View file

@ -17,6 +17,7 @@ Notes:
--*/
#include "tactic/arith/bv2real_rewriter.h"
#include "tactic/tactic_exception.h"
#include "ast/rewriter/rewriter_def.h"
#include "ast/ast_pp.h"
#include "ast/for_each_expr.h"
@ -40,6 +41,7 @@ bv2real_util::bv2real_util(ast_manager& m, rational const& default_root, rationa
m_pos_le = m.mk_fresh_func_decl("<=","",2,domain,m.mk_bool_sort());
m_decls.push_back(m_pos_lt);
m_decls.push_back(m_pos_le);
m_max_memory = std::max((1ull << 31ull), 3*memory::get_allocation_size());
}
bool bv2real_util::is_bv2real(func_decl* f) const {
@ -178,12 +180,10 @@ void bv2real_util::align_divisors(expr_ref& s1, expr_ref& s2, expr_ref& t1, expr
expr* bv2real_util::mk_bv_mul(expr* s, expr* t) {
SASSERT(m_bv.is_bv(s));
SASSERT(m_bv.is_bv(t));
if (is_zero(s)) {
if (is_zero(s))
return s;
}
if (is_zero(t)) {
if (is_zero(t))
return t;
}
expr_ref s1(s, m()), t1(t, m());
align_sizes(s1, t1);
unsigned n = m_bv.get_bv_size(t1);
@ -343,6 +343,10 @@ bool bv2real_util::mk_is_divisible_by(expr_ref& s, rational const& _overflow) {
}
bool bv2real_util::memory_exceeded() const {
return m_max_memory <= memory::get_allocation_size();
}
// ---------------------------------------------------------------------
// bv2real_rewriter
@ -362,6 +366,11 @@ br_status bv2real_rewriter::mk_app_core(func_decl * f, unsigned num_args, expr *
tout << mk_pp(args[i], m()) << " ";
}
tout << "\n";);
if (u().memory_exceeded()) {
std::cout << "tactic exception\n";
throw tactic_exception("bv2real-memory exceeded");
}
if(f->get_family_id() == m_arith.get_family_id()) {
switch (f->get_decl_kind()) {
case OP_NUM: return BR_FAILED;

View file

@ -65,6 +65,7 @@ class bv2real_util {
rational m_default_divisor;
rational m_max_divisor;
unsigned m_max_num_bits;
uint64_t m_max_memory;
class contains_bv2real_proc;
@ -81,6 +82,8 @@ public:
bool contains_bv2real(expr* e) const;
bool memory_exceeded() const;
bool mk_bv2real(expr* s, expr* t, rational& d, rational& r, expr_ref& result);
expr* mk_bv2real_c(expr* s, expr* t, rational const& d, rational const& r);
expr* mk_bv2real(expr* n, expr* m) { return mk_bv2real_c(n, m, default_divisor(), default_root()); }

View file

@ -216,11 +216,10 @@ public:
}
// IF_VERBOSE(0, verbose_stream() << mk_pp(g->form(i), m) << "\n--->\n" << new_curr << "\n";);
g->update(i, new_curr, new_pr, g->dep(i));
}
for (expr* a : axioms) {
for (expr* a : axioms)
g->assert_expr(a);
}
if (m_mc) g->add(m_mc.get());
g->inc_depth();
result.push_back(g.get());

View file

@ -100,20 +100,19 @@ class nla2bv_tactic : public tactic {
return;
}
substitute_vars(g);
TRACE("nla2bv", g.display(tout << "substitute vars\n"););
TRACE("nla2bv", g.display(tout << "substitute vars\n"));
reduce_bv2int(g);
reduce_bv2real(g);
TRACE("nla2bv", g.display(tout << "after reduce\n"););
TRACE("nla2bv", g.display(tout << "after reduce\n"));
mc = m_fmc.get();
for (unsigned i = 0; i < m_vars.size(); ++i) {
m_fmc->add(m_vars[i].get(), m_defs[i].get());
}
for (unsigned i = 0; i < m_vars.size(); ++i)
m_fmc->add(m_vars.get(i), m_defs.get(i));
for (unsigned i = 0; i < m_bv2real.num_aux_decls(); ++i) {
m_fmc->hide(m_bv2real.get_aux_decl(i));
}
IF_VERBOSE(TACTIC_VERBOSITY_LVL, verbose_stream() << "(nla->bv :sat-preserving " << m_is_sat_preserving << ")\n";);
TRACE("nla2bv_verbose", g.display(tout););
TRACE("nla2bv", tout << "Muls: " << count_mul(g) << "\n";);
TRACE("nla2bv_verbose", g.display(tout));
TRACE("nla2bv", tout << "Muls: " << count_mul(g) << "\n");
g.inc_depth();
if (!is_sat_preserving())
g.updt_prec(goal::UNDER);

View file

@ -21,6 +21,7 @@ Notes:
#include "tactic/model_converter.h"
#include "ast/bv_decl_plugin.h"
#include "ast/ast_smt2_pp.h"
#include "ast/ast_pp.h"
#include "ast/ast_util.h"
/**
@ -40,9 +41,7 @@ struct bit_blaster_model_converter : public model_converter {
obj_map<func_decl, expr*> const & const2bits,
ptr_vector<func_decl> const& newbits):
m_vars(m), m_bits(m), m_newbits(m) {
for (auto const& kv : const2bits) {
func_decl * v = kv.m_key;
expr * bits = kv.m_value;
for (auto const& [v, bits] : const2bits) {
SASSERT(!TO_BOOL || is_app_of(bits, m.get_family_id("bv"), OP_MKBV));
SASSERT(TO_BOOL || is_app_of(bits, m.get_family_id("bv"), OP_CONCAT));
m_vars.push_back(v);
@ -70,7 +69,7 @@ struct bit_blaster_model_converter : public model_converter {
bits.insert(to_app(bit)->get_decl());
}
}
TRACE("blaster_mc",
TRACE("model_converter",
tout << "bits that should not be included in the model:\n";
for (func_decl* f : bits) {
tout << f->get_name() << " ";
@ -85,14 +84,14 @@ struct bit_blaster_model_converter : public model_converter {
func_decl * f = old_model->get_constant(i);
if (bits.contains(f))
continue;
TRACE("blaster_mc", tout << "non-bit: " << f->get_name() << "\n";);
TRACE("model_converter", tout << "non-bit: " << f->get_name() << "\n";);
expr * fi = old_model->get_const_interp(f);
new_model->register_decl(f, fi);
}
TRACE("blaster_mc", tout << "after copy non bits:\n"; model_pp(tout, *new_model););
TRACE("model_converter", tout << "after copy non bits:\n"; model_pp(tout, *new_model););
new_model->copy_func_interps(*old_model);
new_model->copy_usort_interps(*old_model);
TRACE("blaster_mc", tout << "after copying functions and sorts:\n"; model_pp(tout, *new_model););
TRACE("model_converter", tout << "after copying functions and sorts:\n"; model_pp(tout, *new_model););
}
void mk_bvs(model * old_model, model * new_model) {
@ -121,7 +120,9 @@ struct bit_blaster_model_converter : public model_converter {
SASSERT(is_uninterp_const(bit));
func_decl * bit_decl = to_app(bit)->get_decl();
expr * bit_val = old_model->get_const_interp(bit_decl);
if (bit_val != nullptr && m().is_true(bit_val))
if (bit_val && !m().is_true(bit_val) && !m().is_false(bit_val))
goto bail;
if (bit_val && m().is_true(bit_val))
val++;
}
}
@ -136,11 +137,27 @@ struct bit_blaster_model_converter : public model_converter {
func_decl * bit_decl = to_app(bit)->get_decl();
expr * bit_val = old_model->get_const_interp(bit_decl);
// remark: if old_model does not assign bit_val, then assume it is false.
if (bit_val != nullptr && !util.is_zero(bit_val))
if (bit_val && !util.is_one(bit_val) && !util.is_zero(bit_val))
goto bail;
if (bit_val && util.is_one(bit_val))
val++;
}
}
new_val = util.mk_numeral(val, bv_sz);
new_val = util.mk_numeral(val, bv_sz);
new_model->register_decl(m_vars.get(i), new_val);
continue;
bail:
expr_ref_vector vals(m());
for (expr* bit : *to_app(bs)) {
func_decl * bit_decl = to_app(bit)->get_decl();
expr * bit_val = old_model->get_const_interp(bit_decl);
SASSERT(bit_val);
vals.push_back(bit_val);
}
if (TO_BOOL)
new_val = util.mk_bv(vals.size(), vals.data());
else
new_val = util.mk_concat(vals);
new_model->register_decl(m_vars.get(i), new_val);
}
}

View file

@ -37,7 +37,7 @@ class bv1_blaster_tactic : public tactic {
bv_util m_util;
obj_map<func_decl, expr*> m_const2bits;
ptr_vector<func_decl> m_newbits;
expr_ref_vector m_saved;
ast_ref_vector m_saved;
expr_ref m_bit1;
expr_ref m_bit0;
@ -108,9 +108,11 @@ class bv1_blaster_tactic : public tactic {
for (unsigned i = 0; i < bv_size; i++) {
bits.push_back(m().mk_fresh_const(nullptr, b));
m_newbits.push_back(to_app(bits.back())->get_decl());
m_saved.push_back(m_newbits.back());
}
r = butil().mk_concat(bits.size(), bits.data());
m_saved.push_back(r);
m_saved.push_back(f);
m_const2bits.insert(f, r);
result = r;
}

View file

@ -110,7 +110,6 @@ protected:
void operator()(quantifier * q) {
m_stats["quantifiers"]++;
SASSERT(is_app(q->get_expr()));
app * body = to_app(q->get_expr());
switch (q->get_kind()) {
case forall_k:
m_stats["forall-variables"] += q->get_num_decls();

View file

@ -892,9 +892,8 @@ public:
m_num_elim_apps = 0;
}
unsigned user_propagate_register_expr(expr* e) override {
void user_propagate_register_expr(expr* e) override {
m_nonvars.insert(e);
return 0;
}
void user_propagate_clear() override {

View file

@ -119,42 +119,38 @@ public:
}
generic_model_converter* pp = alloc(generic_model_converter, m, "pb-preprocess");
g->add(pp);
g->inc_depth();
result.push_back(g.get());
while (simplify(g, *pp));
g->add(pp);
// decompose(g);
}
bool simplify(goal_ref const& g, generic_model_converter& mc) {
reset();
normalize(g);
if (g->inconsistent()) {
if (g->inconsistent())
return false;
}
for (unsigned i = 0; i < g->size(); ++i) {
process_vars(i, g);
}
if (m_ge.empty()) {
for (unsigned i = 0; i < g->size(); ++i)
process_vars(i, g);
if (m_ge.empty())
return false;
}
for (unsigned i = 0; i < m_ge.size(); ++i) {
for (unsigned i = 0; i < m_ge.size(); ++i)
if (!classify_vars(i, to_app(g->form(m_ge[i]))))
return false;
}
declassifier dcl(m_vars);
expr_mark visited;
for (unsigned i = 0; !m_vars.empty() && i < m_other.size(); ++i) {
for (unsigned i = 0; !m_vars.empty() && i < m_other.size(); ++i)
for_each_expr(dcl, visited, g->form(m_other[i]));
}
if (m_vars.empty()) {
if (m_vars.empty())
return false;
}
// display_annotation(tout, g);
m_progress = false;
@ -172,24 +168,23 @@ public:
replace(r.pos, e, m.mk_true(), g);
set_value(mc, e, true);
}
if (g->inconsistent()) return false;
if (g->inconsistent())
return false;
++it;
it = next_resolvent(it);
}
// now resolve clauses.
it = next_resolvent(m_vars.begin());
while (it != m_vars.end()) {
while (it != m_vars.end()) {
app * e = it->m_key;
SASSERT(is_uninterp_const(e));
rec const& r = it->m_value;
if (r.pos.size() == 1 && !r.neg.empty()) {
if (r.pos.size() == 1 && !r.neg.empty())
resolve(mc, r.pos[0], r.neg, e, true, g);
}
else if (r.neg.size() == 1 && !r.pos.empty()) {
else if (r.neg.size() == 1 && !r.pos.empty())
resolve(mc, r.neg[0], r.pos, e, false, g);
}
if (g->inconsistent()) return false;
if (g->inconsistent())
return false;
++it;
it = next_resolvent(it);
}
@ -201,20 +196,29 @@ public:
vector<rational> coeffs1, coeffs2;
rational k1, k2;
expr* fml = g->form(m_ge[i]);
if (!to_ge(fml, args1, coeffs1, k1)) continue;
if (args1.empty()) continue;
if (!to_ge(fml, args1, coeffs1, k1))
continue;
if (args1.empty())
continue;
expr* arg = args1.get(0);
bool neg = m.is_not(arg, arg);
if (!is_uninterp_const(arg)) continue;
if (!m_vars.contains(to_app(arg))) continue;
if (!is_uninterp_const(arg))
continue;
if (!m_vars.contains(to_app(arg)))
continue;
rec const& r = m_vars.find(to_app(arg));
unsigned_vector const& pos = neg?r.neg:r.pos;
for (unsigned j = 0; j < pos.size(); ++j) {
unsigned k = pos[j];
if (k == m_ge[i]) continue;
if (!to_ge(g->form(k), args2, coeffs2, k2)) continue;
if (k == m_ge[i])
continue;
coeffs2.reset();
args2.reset();
if (!to_ge(g->form(k), args2, coeffs2, k2))
continue;
if (subsumes(args1, coeffs1, k1, args2, coeffs2, k2)) {
IF_VERBOSE(3, verbose_stream() << "replace " << mk_pp(g->form(k), m) << "\n";);
IF_VERBOSE(3, verbose_stream() << "subsumes " << mk_pp(fml, m) << "\n");
IF_VERBOSE(3, verbose_stream() << "replace " << mk_pp(g->form(k), m) << "\n");
g->update(k, m.mk_true(), nullptr, m.mk_join(g->dep(m_ge[i]), g->dep(k)));
m_progress = true;
}
@ -417,12 +421,10 @@ private:
m_trail.push_back(e);
m_vars.insert(e, rec());
}
if (pos) {
if (pos)
m_vars.find(e).pos.push_back(i);
}
else {
else
m_vars.find(e).neg.push_back(i);
}
}
bool pure_args(app* a) const {
@ -631,16 +633,19 @@ private:
vector<rational> const& coeffs1, rational const& k1,
expr_ref_vector const& args2,
vector<rational> const& coeffs2, rational const& k2) {
if (k2 > k1) return false;
if (k2 > k1)
return false;
for (unsigned i = 0; i < args1.size(); ++i) {
bool found = false;
for (unsigned j = 0; !found && j < args2.size(); ++j) {
if (args1[i] == args2[j]) {
if (coeffs1[i] > coeffs2[j]) return false;
if (coeffs1[i] > coeffs2[j])
return false;
found = true;
}
}
if (!found) return false;
if (!found)
return false;
}
return true;
}

View file

@ -78,7 +78,7 @@ public:
void operator()(goal_ref const & g, goal_ref_buffer & result) override;
void cleanup() override;
unsigned user_propagate_register_expr(expr* e) override;
void user_propagate_register_expr(expr* e) override;
void user_propagate_clear() override;
};
@ -502,9 +502,8 @@ void reduce_args_tactic::cleanup() {
m_imp->m_vars.append(vars);
}
unsigned reduce_args_tactic::user_propagate_register_expr(expr* e) {
void reduce_args_tactic::user_propagate_register_expr(expr* e) {
m_imp->m_vars.push_back(e);
return 0;
}
void reduce_args_tactic::user_propagate_clear() {

View file

@ -65,7 +65,8 @@ class solve_eqs_tactic : public tactic {
m_a_util(m),
m_num_steps(0),
m_num_eliminated_vars(0),
m_marked_candidates(m) {
m_marked_candidates(m),
m_var_trail(m) {
updt_params(p);
if (m_r == nullptr)
m_r = mk_default_expr_replacer(m, true);
@ -524,7 +525,14 @@ class solve_eqs_tactic : public tactic {
}
}
expr_mark m_compatible_tried;
expr_ref_vector m_var_trail;
bool is_compatible(goal const& g, unsigned idx, vector<nnf_context> const & path, expr* v, expr* eq) {
if (m_compatible_tried.is_marked(v))
return false;
m_compatible_tried.mark(v);
m_var_trail.push_back(v);
expr_mark occ;
svector<lbool> cache;
mark_occurs(occ, g, v);
@ -649,7 +657,7 @@ class solve_eqs_tactic : public tactic {
for (unsigned i = 0; i < args.size(); ++i) {
pr = nullptr;
expr* arg = args.get(i), *lhs = nullptr, *rhs = nullptr;
if (m().is_eq(arg, lhs, rhs)) {
if (m().is_eq(arg, lhs, rhs) && !m().is_bool(lhs)) {
if (trivial_solve1(lhs, rhs, var, def, pr) && is_compatible(g, idx, path, var, arg)) {
insert_solution(g, idx, arg, var, def, pr);
}

View file

@ -165,8 +165,8 @@ public:
void get_levels(ptr_vector<expr> const& vars, unsigned_vector& depth) override {
m_solver->get_levels(vars, depth);
}
expr_ref_vector get_trail() override {
return m_solver->get_trail();
expr_ref_vector get_trail(unsigned max_level) override {
return m_solver->get_trail(max_level);
}
model_converter* external_model_converter() const {

View file

@ -189,8 +189,8 @@ public:
m_solver->get_levels(vars, depth);
}
expr_ref_vector get_trail() override {
return m_solver->get_trail();
expr_ref_vector get_trail(unsigned max_level) override {
return m_solver->get_trail(max_level);
}
unsigned get_num_assertions() const override {

View file

@ -105,8 +105,8 @@ public:
m_solver->get_levels(vars, depth);
}
expr_ref_vector get_trail() override {
return m_solver->get_trail();
expr_ref_vector get_trail(unsigned max_level) override {
return m_solver->get_trail(max_level);
}
model_converter* external_model_converter() const{

View file

@ -2098,9 +2098,9 @@ namespace smtfd {
m_fd_sat_solver->get_levels(vars, depth);
}
expr_ref_vector get_trail() override {
expr_ref_vector get_trail(unsigned max_level) override {
init();
return m_fd_sat_solver->get_trail();
return m_fd_sat_solver->get_trail(max_level);
}
unsigned get_num_assertions() const override {

View file

@ -80,8 +80,8 @@ bool horn_subsume_model_converter::mk_horn(
if (w >= subst.size()) {
subst.resize(w+1);
}
if (subst[w].get()) {
conjs.push_back(m.mk_eq(v, subst[w].get()));
if (subst.get(w)) {
conjs.push_back(m.mk_eq(v, subst.get(w)));
}
else {
subst[w] = v;
@ -92,11 +92,11 @@ bool horn_subsume_model_converter::mk_horn(
}
}
expr_ref body_expr(m);
body_expr = m.mk_and(conjs.size(), conjs.data());
body_expr = m.mk_and(conjs);
// substitute variables directly.
if (!subst.empty()) {
body_expr = vs(body_expr, subst.size(), subst.data());
body_expr = vs(body_expr, subst);
}
if (fv.empty()) {
@ -174,17 +174,16 @@ void horn_subsume_model_converter::operator()(model_ref& mr) {
func_decl_ref pred(m);
expr_ref body_res(m);
for (unsigned i = 0; i < m_delay_head.size(); ++i) {
VERIFY(mk_horn(m_delay_head[i].get(), m_delay_body[i].get(), pred, body_res));
VERIFY(mk_horn(m_delay_head.get(i), m_delay_body.get(i), pred, body_res));
insert(pred.get(), body_res.get());
}
m_delay_head.reset();
m_delay_body.reset();
TRACE("mc", tout << m_funcs.size() << "\n"; model_smt2_pp(tout, m, *mr, 0););
for (unsigned i = m_funcs.size(); i > 0; ) {
--i;
func_decl* h = m_funcs[i].get();
expr_ref body(m_bodies[i].get(), m);
for (unsigned i = m_funcs.size(); i-- > 0; ) {
func_decl* h = m_funcs.get(i);
expr_ref body(m_bodies.get(i), m);
unsigned arity = h->get_arity();
add_default_false_interpretation(body, mr);
SASSERT(m.is_bool(body));

View file

@ -143,10 +143,9 @@ public:
tactic_ref t;
if (tp.default_tactic() != symbol::null &&
!tp.default_tactic().is_numerical() &&
tp.default_tactic().bare_str() &&
tp.default_tactic().bare_str()[0]) {
tp.default_tactic().str()[0]) {
cmd_context ctx(false, &m, l);
std::istringstream is(tp.default_tactic().bare_str());
std::istringstream is(tp.default_tactic().str());
char const* file_name = "";
sexpr_ref se = parse_sexpr(ctx, is, p, file_name);
if (se) {

View file

@ -52,6 +52,7 @@ struct tactic_report::imp {
<< " :before-memory " << std::fixed << std::setprecision(2) << m_start_memory
<< " :after-memory " << std::fixed << std::setprecision(2) << end_memory
<< ")" << std::endl);
IF_VERBOSE(20, m_goal.display(verbose_stream() << m_id << "\n"));
SASSERT(m_goal.is_well_formed());
}
};

View file

@ -85,7 +85,7 @@ public:
throw default_exception("tactic does not support user propagation");
}
unsigned user_propagate_register_expr(expr* e) override { return 0; }
void user_propagate_register_expr(expr* e) override { }
virtual char const* name() const = 0;
protected:

View file

@ -190,9 +190,9 @@ public:
m_t2->user_propagate_register_diseq(diseq_eh);
}
unsigned user_propagate_register_expr(expr* e) override {
void user_propagate_register_expr(expr* e) override {
m_t1->user_propagate_register_expr(e);
return m_t2->user_propagate_register_expr(e);
m_t2->user_propagate_register_expr(e);
}
void user_propagate_clear() override {
@ -204,6 +204,10 @@ public:
m_t2->user_propagate_register_created(created_eh);
}
void user_propagate_register_decide(user_propagator::decide_eh_t& decide_eh) override {
m_t2->user_propagate_register_decide(decide_eh);
}
};
tactic * and_then(tactic * t1, tactic * t2) {
@ -336,6 +340,9 @@ public:
catch (tactic_exception &) {
result.reset();
}
catch (rewriter_exception&) {
result.reset();
}
catch (z3_error & ex) {
IF_VERBOSE(10, verbose_stream() << "z3 error: " << ex.error_code() << " in or-else\n");
throw;
@ -848,7 +855,7 @@ public:
void reset() override { m_t->reset(); }
void set_logic(symbol const& l) override { m_t->set_logic(l); }
void set_progress_callback(progress_callback * callback) override { m_t->set_progress_callback(callback); }
unsigned user_propagate_register_expr(expr* e) override { return m_t->user_propagate_register_expr(e); }
void user_propagate_register_expr(expr* e) override { m_t->user_propagate_register_expr(e); }
void user_propagate_clear() override { m_t->user_propagate_clear(); }
protected:
@ -1019,7 +1026,6 @@ public:
void operator()(goal_ref const & in, goal_ref_buffer& result) override {
cancel_eh<reslimit> eh(in->m().limit());
{
// Warning: scoped_timer is not thread safe in Linux.
scoped_timer timer(m_timeout, &eh);
m_t->operator()(in, result);
}

View file

@ -2,14 +2,15 @@
#pragma once
#include "ast/ast.h"
#include "util/lbool.h"
namespace user_propagator {
class callback {
public:
virtual ~callback() = default;
virtual void propagate_cb(unsigned num_fixed, unsigned const* fixed_ids, unsigned num_eqs, unsigned const* eq_lhs, unsigned const* eq_rhs, expr* conseq) = 0;
virtual unsigned register_cb(expr* e) = 0;
virtual void propagate_cb(unsigned num_fixed, expr* const* fixed_ids, unsigned num_eqs, expr* const* eq_lhs, expr* const* eq_rhs, expr* conseq) = 0;
virtual void register_cb(expr* e) = 0;
};
class context_obj {
@ -17,14 +18,14 @@ namespace user_propagator {
virtual ~context_obj() = default;
};
typedef std::function<void(void*, callback*)> final_eh_t;
typedef std::function<void(void*, callback*, unsigned, expr*)> fixed_eh_t;
typedef std::function<void(void*, callback*, unsigned, unsigned)> eq_eh_t;
typedef std::function<void*(void*, ast_manager&, context_obj*&)> fresh_eh_t;
typedef std::function<void(void*)> push_eh_t;
typedef std::function<void(void*,unsigned)> pop_eh_t;
typedef std::function<void(void*, callback*, expr*, unsigned)> created_eh_t;
typedef std::function<void(void*, callback*)> final_eh_t;
typedef std::function<void(void*, callback*, expr*, expr*)> fixed_eh_t;
typedef std::function<void(void*, callback*, expr*, expr*)> eq_eh_t;
typedef std::function<void*(void*, ast_manager&, context_obj*&)> fresh_eh_t;
typedef std::function<void(void*, callback*)> push_eh_t;
typedef std::function<void(void*, callback*, unsigned)> pop_eh_t;
typedef std::function<void(void*, callback*, expr*)> created_eh_t;
typedef std::function<void(void*, callback*, expr**, unsigned*, lbool*)> decide_eh_t;
class plugin : public decl_plugin {
public:
@ -77,7 +78,7 @@ namespace user_propagator {
throw default_exception("user-propagators are only supported on the SMT solver");
}
virtual unsigned user_propagate_register_expr(expr* e) {
virtual void user_propagate_register_expr(expr* e) {
throw default_exception("user-propagators are only supported on the SMT solver");
}
@ -85,6 +86,10 @@ namespace user_propagator {
throw default_exception("user-propagators are only supported on the SMT solver");
}
virtual void user_propagate_register_decide(decide_eh_t& r) {
throw default_exception("user-propagators are only supported on the SMT solver");
}
virtual void user_propagate_clear() {
}