3
0
Fork 0
mirror of https://github.com/Z3Prover/z3 synced 2025-04-23 09:05:31 +00:00

Use override rather than virtual.

This commit is contained in:
Bruce Mitchener 2018-02-10 09:15:12 +07:00
parent ce123d9dbc
commit 7167fda1dc
220 changed files with 2546 additions and 2548 deletions

View file

@ -42,7 +42,7 @@ namespace smt {
m_n2(n2) {
}
virtual void undo(context & ctx) {
void undo(context & ctx) override {
m_already_processed.erase(m_n1, m_n2);
TRACE("arith_eq_adapter_profile", tout << "del #" << m_n1->get_owner_id() << " #" << m_n2->get_owner_id() << "\n";);
}
@ -67,9 +67,9 @@ namespace smt {
m_ge(ge) {
}
virtual ~arith_eq_relevancy_eh() {}
~arith_eq_relevancy_eh() override {}
virtual void operator()(relevancy_propagator & rp) {
void operator()(relevancy_propagator & rp) override {
if (!rp.is_relevant(m_n1))
return;
if (!rp.is_relevant(m_n2))

View file

@ -82,80 +82,80 @@ class asserted_formulas {
class reduce_asserted_formulas_fn : public simplify_fmls {
public:
reduce_asserted_formulas_fn(asserted_formulas& af): simplify_fmls(af, "reduce-asserted") {}
virtual void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) { af.m_rewriter(j.get_fml(), n, p); }
void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) override { af.m_rewriter(j.get_fml(), n, p); }
};
class find_macros_fn : public simplify_fmls {
public:
find_macros_fn(asserted_formulas& af): simplify_fmls(af, "find-macros") {}
virtual void operator()() { af.find_macros_core(); }
virtual bool should_apply() const { return af.m_params.m_macro_finder && af.has_quantifiers(); }
virtual void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) { UNREACHABLE(); }
void operator()() override { af.find_macros_core(); }
bool should_apply() const override { return af.m_params.m_macro_finder && af.has_quantifiers(); }
void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) override { UNREACHABLE(); }
};
class apply_quasi_macros_fn : public simplify_fmls {
public:
apply_quasi_macros_fn(asserted_formulas& af): simplify_fmls(af, "find-quasi-macros") {}
virtual void operator()() { af.apply_quasi_macros(); }
virtual bool should_apply() const { return af.m_params.m_quasi_macros && af.has_quantifiers(); }
virtual void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) { UNREACHABLE(); }
void operator()() override { af.apply_quasi_macros(); }
bool should_apply() const override { return af.m_params.m_quasi_macros && af.has_quantifiers(); }
void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) override { UNREACHABLE(); }
};
class nnf_cnf_fn : public simplify_fmls {
public:
nnf_cnf_fn(asserted_formulas& af): simplify_fmls(af, "nnf-cnf") {}
virtual void operator()() { af.nnf_cnf(); }
virtual bool should_apply() const { return af.m_params.m_nnf_cnf || (af.m_params.m_mbqi && af.has_quantifiers()); }
virtual void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) { UNREACHABLE(); }
void operator()() override { af.nnf_cnf(); }
bool should_apply() const override { return af.m_params.m_nnf_cnf || (af.m_params.m_mbqi && af.has_quantifiers()); }
void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) override { UNREACHABLE(); }
};
class propagate_values_fn : public simplify_fmls {
public:
propagate_values_fn(asserted_formulas& af): simplify_fmls(af, "propagate-values") {}
virtual void operator()() { af.propagate_values(); }
virtual bool should_apply() const { return af.m_params.m_propagate_values; }
virtual void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) { UNREACHABLE(); }
void operator()() override { af.propagate_values(); }
bool should_apply() const override { return af.m_params.m_propagate_values; }
void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) override { UNREACHABLE(); }
};
class distribute_forall_fn : public simplify_fmls {
distribute_forall m_functor;
public:
distribute_forall_fn(asserted_formulas& af): simplify_fmls(af, "distribute-forall"), m_functor(af.m) {}
virtual void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) { m_functor(j.get_fml(), n); }
virtual bool should_apply() const { return af.m_params.m_distribute_forall && af.has_quantifiers(); }
virtual void post_op() { af.reduce_and_solve(); TRACE("asserted_formulas", af.display(tout);); }
void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) override { m_functor(j.get_fml(), n); }
bool should_apply() const override { return af.m_params.m_distribute_forall && af.has_quantifiers(); }
void post_op() override { af.reduce_and_solve(); TRACE("asserted_formulas", af.display(tout);); }
};
class pattern_inference_fn : public simplify_fmls {
pattern_inference_rw m_infer;
public:
pattern_inference_fn(asserted_formulas& af): simplify_fmls(af, "pattern-inference"), m_infer(af.m, af.m_params) {}
virtual void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) { m_infer(j.get_fml(), n, p); }
virtual bool should_apply() const { return af.m_params.m_ematching && af.has_quantifiers(); }
void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) override { m_infer(j.get_fml(), n, p); }
bool should_apply() const override { return af.m_params.m_ematching && af.has_quantifiers(); }
};
class refine_inj_axiom_fn : public simplify_fmls {
public:
refine_inj_axiom_fn(asserted_formulas& af): simplify_fmls(af, "refine-injectivity") {}
virtual void simplify(justified_expr const& j, expr_ref& n, proof_ref& p);
virtual bool should_apply() const { return af.m_params.m_refine_inj_axiom && af.has_quantifiers(); }
void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) override;
bool should_apply() const override { return af.m_params.m_refine_inj_axiom && af.has_quantifiers(); }
};
class max_bv_sharing_fn : public simplify_fmls {
public:
max_bv_sharing_fn(asserted_formulas& af): simplify_fmls(af, "maximizing-bv-sharing") {}
virtual void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) { af.m_bv_sharing(j.get_fml(), n, p); }
virtual bool should_apply() const { return af.m_params.m_max_bv_sharing; }
virtual void post_op() { af.m_reduce_asserted_formulas(); }
void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) override { af.m_bv_sharing(j.get_fml(), n, p); }
bool should_apply() const override { return af.m_params.m_max_bv_sharing; }
void post_op() override { af.m_reduce_asserted_formulas(); }
};
class elim_term_ite_fn : public simplify_fmls {
elim_term_ite_rw m_elim;
public:
elim_term_ite_fn(asserted_formulas& af): simplify_fmls(af, "elim-term-ite"), m_elim(af.m, af.m_defined_names) {}
virtual void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) { m_elim(j.get_fml(), n, p); }
virtual bool should_apply() const { return af.m_params.m_eliminate_term_ite && af.m_params.m_lift_ite != LI_FULL; }
virtual void post_op() { af.m_formulas.append(m_elim.new_defs()); af.reduce_and_solve(); m_elim.reset(); }
void simplify(justified_expr const& j, expr_ref& n, proof_ref& p) override { m_elim(j.get_fml(), n, p); }
bool should_apply() const override { return af.m_params.m_eliminate_term_ite && af.m_params.m_lift_ite != LI_FULL; }
void post_op() override { af.m_formulas.append(m_elim.new_defs()); af.reduce_and_solve(); m_elim.reset(); }
};
#define MK_SIMPLIFIERA(NAME, FUNCTOR, MSG, APP, ARG, REDUCE) \

View file

@ -39,12 +39,12 @@ namespace smt {
SASSERT(m_app1->get_id() < m_app2->get_id());
}
virtual char const * get_name() const { return "dyn-ack"; }
char const * get_name() const override { return "dyn-ack"; }
virtual void get_antecedents(conflict_resolution & cr) {
void get_antecedents(conflict_resolution & cr) override {
}
virtual void display_debug_info(conflict_resolution & cr, std::ostream & out) {
void display_debug_info(conflict_resolution & cr, std::ostream & out) override {
ast_manager & m = cr.get_manager();
out << "m_app1:\n" << mk_pp(m_app1, m) << "\n";
out << "m_app2:\n" << mk_pp(m_app2, m) << "\n";
@ -70,7 +70,7 @@ namespace smt {
}
}
virtual proof * mk_proof(conflict_resolution & cr) {
proof * mk_proof(conflict_resolution & cr) override {
ast_manager & m = cr.get_manager();
context & ctx = cr.get_context();
unsigned num_args = m_app1->get_num_args();
@ -288,8 +288,8 @@ namespace smt {
dyn_ack_clause_del_eh(dyn_ack_manager & m):
m_manager(m) {
}
virtual ~dyn_ack_clause_del_eh() {}
virtual void operator()(ast_manager & m, clause * cls) {
~dyn_ack_clause_del_eh() override {}
void operator()(ast_manager & m, clause * cls) override {
m_manager.del_clause_eh(cls);
dealloc(this);
}

View file

@ -2844,7 +2844,7 @@ namespace smt {
unsigned m_lbl_id;
public:
mk_tree_trail(ptr_vector<code_tree> & t, unsigned id):m_trees(t), m_lbl_id(id) {}
virtual void undo(mam_impl & m) {
void undo(mam_impl & m) override {
dealloc(m_trees[m_lbl_id]);
m_trees[m_lbl_id] = 0;
}
@ -3104,7 +3104,7 @@ namespace smt {
enode * m_enode;
public:
add_shared_enode_trail(enode * n):m_enode(n) {}
virtual void undo(mam_impl & m) { m.m_shared_enodes.erase(m_enode); }
void undo(mam_impl & m) override { m.m_shared_enodes.erase(m_enode); }
};
#ifdef Z3DEBUG
@ -3819,11 +3819,11 @@ namespace smt {
reset_pp_pc();
}
virtual ~mam_impl() {
~mam_impl() override {
m_trail_stack.reset();
}
virtual void add_pattern(quantifier * qa, app * mp) {
void add_pattern(quantifier * qa, app * mp) override {
SASSERT(m_ast_manager.is_pattern(mp));
TRACE("trigger_bug", tout << "adding pattern\n" << mk_ismt2_pp(qa, m_ast_manager) << "\n" << mk_ismt2_pp(mp, m_ast_manager) << "\n";);
TRACE("mam_bug", tout << "adding pattern\n" << mk_pp(qa, m_ast_manager) << "\n" << mk_pp(mp, m_ast_manager) << "\n";);
@ -3846,11 +3846,11 @@ namespace smt {
m_trees.add_pattern(qa, mp, i);
}
virtual void push_scope() {
void push_scope() override {
m_trail_stack.push_scope();
}
virtual void pop_scope(unsigned num_scopes) {
void pop_scope(unsigned num_scopes) override {
if (!m_to_match.empty()) {
ptr_vector<code_tree>::iterator it = m_to_match.begin();
ptr_vector<code_tree>::iterator end = m_to_match.end();
@ -3864,7 +3864,7 @@ namespace smt {
m_trail_stack.pop_scope(num_scopes);
}
virtual void reset() {
void reset() override {
m_trail_stack.reset();
m_trees.reset();
m_to_match.reset();
@ -3875,7 +3875,7 @@ namespace smt {
m_tmp_region.reset();
}
virtual void display(std::ostream& out) {
void display(std::ostream& out) override {
out << "mam:\n";
m_lbl_hasher.display(out);
ptr_vector<code_tree>::iterator it = m_trees.begin_code_trees();
@ -3886,7 +3886,7 @@ namespace smt {
}
}
virtual void match() {
void match() override {
TRACE("trigger_bug", tout << "match\n"; display(tout););
ptr_vector<code_tree>::iterator it = m_to_match.begin();
ptr_vector<code_tree>::iterator end = m_to_match.end();
@ -3903,7 +3903,7 @@ namespace smt {
}
}
virtual void rematch(bool use_irrelevant) {
void rematch(bool use_irrelevant) override {
ptr_vector<code_tree>::iterator it = m_trees.begin_code_trees();
ptr_vector<code_tree>::iterator end = m_trees.end_code_trees();
unsigned lbl = 0;
@ -3932,7 +3932,7 @@ namespace smt {
}
#endif
virtual void on_match(quantifier * qa, app * pat, unsigned num_bindings, enode * const * bindings, unsigned max_generation, ptr_vector<enode> & used_enodes) {
void on_match(quantifier * qa, app * pat, unsigned num_bindings, enode * const * bindings, unsigned max_generation, ptr_vector<enode> & used_enodes) override {
TRACE("trigger_bug", tout << "found match " << mk_pp(qa, m_ast_manager) << "\n";);
#ifdef Z3DEBUG
if (m_check_missing_instances) {
@ -3955,13 +3955,13 @@ namespace smt {
m_context.add_instance(qa, pat, num_bindings, bindings, max_generation, min_gen, max_gen, used_enodes);
}
virtual bool is_shared(enode * n) const {
bool is_shared(enode * n) const override {
return !m_shared_enodes.empty() && m_shared_enodes.contains(n);
}
// This method is invoked when n becomes relevant.
// If lazy == true, then n is not added to the list of candidate enodes for matching. That is, the method just updates the lbls.
virtual void relevant_eh(enode * n, bool lazy) {
void relevant_eh(enode * n, bool lazy) override {
TRACE("trigger_bug", tout << "relevant_eh:\n" << mk_ismt2_pp(n->get_owner(), m_ast_manager) << "\n";
tout << "mam: " << this << "\n";);
TRACE("mam", tout << "relevant_eh: #" << n->get_owner_id() << "\n";);
@ -3984,11 +3984,11 @@ namespace smt {
}
}
virtual bool has_work() const {
bool has_work() const override {
return !m_to_match.empty() || !m_new_patterns.empty();
}
virtual void add_eq_eh(enode * r1, enode * r2) {
void add_eq_eh(enode * r1, enode * r2) override {
flet<enode *> l1(m_r1, r1);
flet<enode *> l2(m_r2, r2);

View file

@ -32,13 +32,13 @@ class array_factory : public struct_factory {
public:
array_factory(ast_manager & m, proto_model & md);
virtual ~array_factory() {}
~array_factory() override {}
virtual expr * get_some_value(sort * s);
expr * get_some_value(sort * s) override;
virtual bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2);
bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2) override;
virtual expr * get_fresh_value(sort * s);
expr * get_fresh_value(sort * s) override;
};
#endif /* ARRAY_FACTORY_H_ */

View file

@ -33,9 +33,9 @@ class datatype_factory : public struct_factory {
public:
datatype_factory(ast_manager & m, proto_model & md);
virtual ~datatype_factory() {}
virtual expr * get_some_value(sort * s);
virtual expr * get_fresh_value(sort * s);
~datatype_factory() override {}
expr * get_some_value(sort * s) override;
expr * get_fresh_value(sort * s) override;
};
#endif /* DATATYPE_FACTORY_H_ */

View file

@ -26,17 +26,17 @@ Revision History:
class numeral_factory : public simple_factory<rational> {
public:
numeral_factory(ast_manager & m, family_id fid):simple_factory<rational>(m, fid) {}
virtual ~numeral_factory() {}
~numeral_factory() override {}
};
class arith_factory : public numeral_factory {
arith_util m_util;
virtual app * mk_value_core(rational const & val, sort * s);
app * mk_value_core(rational const & val, sort * s) override;
public:
arith_factory(ast_manager & m);
virtual ~arith_factory();
~arith_factory() override;
app * mk_num_value(rational const & val, bool is_int);
};
@ -44,11 +44,11 @@ public:
class bv_factory : public numeral_factory {
bv_util m_util;
virtual app * mk_value_core(rational const & val, sort * s);
app * mk_value_core(rational const & val, sort * s) override;
public:
bv_factory(ast_manager & m);
virtual ~bv_factory();
~bv_factory() override;
app * mk_num_value(rational const & val, unsigned bv_size);
};

View file

@ -60,7 +60,7 @@ class proto_model : public model_core {
public:
proto_model(ast_manager & m, params_ref const & p = params_ref());
virtual ~proto_model() {}
~proto_model() override {}
void register_factory(value_factory * f) { m_factories.register_plugin(f); }
@ -69,7 +69,7 @@ public:
value_factory * get_factory(family_id fid);
virtual expr * get_some_value(sort * s);
expr * get_some_value(sort * s) override;
bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2);
@ -93,9 +93,9 @@ public:
void freeze_universe(sort * s);
bool is_finite(sort * s) const;
obj_hashtable<expr> const & get_known_universe(sort * s) const;
virtual ptr_vector<expr> const & get_universe(sort * s) const;
virtual unsigned get_num_uninterpreted_sorts() const;
virtual sort * get_uninterpreted_sort(unsigned idx) const;
ptr_vector<expr> const & get_universe(sort * s) const override;
unsigned get_num_uninterpreted_sorts() const override;
sort * get_uninterpreted_sort(unsigned idx) const override;
//
// Complete partial function interps

View file

@ -43,11 +43,11 @@ protected:
public:
struct_factory(ast_manager & m, family_id fid, proto_model & md);
virtual ~struct_factory();
~struct_factory() override;
virtual bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2);
bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2) override;
virtual void register_value(expr * array_value);
void register_value(expr * array_value) override;
};
#endif /* STRUCT_FACTORY_H_ */

View file

@ -64,13 +64,13 @@ class basic_factory : public value_factory {
public:
basic_factory(ast_manager & m);
virtual expr * get_some_value(sort * s);
expr * get_some_value(sort * s) override;
virtual bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2);
bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2) override;
virtual expr * get_fresh_value(sort * s);
expr * get_fresh_value(sort * s) override;
virtual void register_value(expr * n) { }
void register_value(expr * n) override { }
};
/**
@ -133,11 +133,11 @@ public:
m_sorts(m) {
}
virtual ~simple_factory() {
~simple_factory() override {
std::for_each(m_sets.begin(), m_sets.end(), delete_proc<value_set>());
}
virtual expr * get_some_value(sort * s) {
expr * get_some_value(sort * s) override {
value_set * set = 0;
expr * result = 0;
if (m_sort2value_set.find(s, set) && !set->m_values.empty())
@ -147,7 +147,7 @@ public:
return result;
}
virtual bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2) {
bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2) override {
value_set * set = 0;
if (m_sort2value_set.find(s, set)) {
switch (set->m_values.size()) {
@ -176,7 +176,7 @@ public:
return true;
}
virtual expr * get_fresh_value(sort * s) {
expr * get_fresh_value(sort * s) override {
value_set * set = get_value_set(s);
bool is_new = false;
expr * result = 0;
@ -202,7 +202,7 @@ public:
return result;
}
virtual void register_value(expr * n) {
void register_value(expr * n) override {
sort * s = this->m_manager.get_sort(n);
value_set * set = get_value_set(s);
if (!set->m_values.contains(n)) {
@ -228,10 +228,10 @@ public:
class user_sort_factory : public simple_factory<unsigned> {
obj_hashtable<sort> m_finite; //!< set of sorts that are marked as finite.
obj_hashtable<expr> m_empty_universe;
virtual app * mk_value_core(unsigned const & val, sort * s);
app * mk_value_core(unsigned const & val, sort * s) override;
public:
user_sort_factory(ast_manager & m);
virtual ~user_sort_factory() {}
~user_sort_factory() override {}
/**
\brief Make the universe of \c s finite, by preventing new
@ -257,13 +257,13 @@ public:
*/
obj_hashtable<sort> const & get_finite_sorts() const { return m_finite; }
virtual expr * get_some_value(sort * s);
expr * get_some_value(sort * s) override;
virtual bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2);
bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2) override;
virtual expr * get_fresh_value(sort * s);
expr * get_fresh_value(sort * s) override;
virtual void register_value(expr * n);
void register_value(expr * n) override;
};
#endif /* VALUE_FACTORY_H_ */

View file

@ -24,22 +24,22 @@ class include_cmd : public cmd {
char const * m_filename;
public:
include_cmd() : cmd("include"), m_filename(0) {}
virtual char const * get_usage() const { return "<string>"; }
virtual char const * get_descr(cmd_context & ctx) const { return "include a file"; }
virtual unsigned get_arity() const { return 1; }
virtual cmd_arg_kind next_arg_kind(cmd_context & ctx) const { return CPK_STRING; }
virtual void set_next_arg(cmd_context & ctx, char const * val) { m_filename = val; }
virtual void failure_cleanup(cmd_context & ctx) {}
virtual void execute(cmd_context & ctx) {
char const * get_usage() const override { return "<string>"; }
char const * get_descr(cmd_context & ctx) const override { return "include a file"; }
unsigned get_arity() const override { return 1; }
cmd_arg_kind next_arg_kind(cmd_context & ctx) const override { return CPK_STRING; }
void set_next_arg(cmd_context & ctx, char const * val) override { m_filename = val; }
void failure_cleanup(cmd_context & ctx) override {}
void execute(cmd_context & ctx) override {
std::ifstream is(m_filename);
if (is.bad() || is.fail())
throw cmd_exception(std::string("failed to open file '") + m_filename + "'");
parse_smt2_commands(ctx, is, false, params_ref(), m_filename);
is.close();
}
virtual void prepare(cmd_context & ctx) { reset(ctx); }
virtual void reset(cmd_context & ctx) { m_filename = 0; }
virtual void finalize(cmd_context & ctx) { reset(ctx); }
void prepare(cmd_context & ctx) override { reset(ctx); }
void reset(cmd_context & ctx) override { m_filename = 0; }
void finalize(cmd_context & ctx) override { reset(ctx); }
};
void install_smt2_extra_cmds(cmd_context & ctx) {

View file

@ -75,42 +75,42 @@ namespace smt {
m_queue(1024, bool_var_act_lt(ctx.get_activity_vector())) {
}
virtual void activity_increased_eh(bool_var v) {
void activity_increased_eh(bool_var v) override {
if (m_queue.contains(v))
m_queue.decreased(v);
}
virtual void mk_var_eh(bool_var v) {
void mk_var_eh(bool_var v) override {
m_queue.reserve(v+1);
SASSERT(!m_queue.contains(v));
m_queue.insert(v);
}
virtual void del_var_eh(bool_var v) {
void del_var_eh(bool_var v) override {
if (m_queue.contains(v))
m_queue.erase(v);
}
virtual void unassign_var_eh(bool_var v) {
void unassign_var_eh(bool_var v) override {
if (!m_queue.contains(v))
m_queue.insert(v);
}
virtual void relevant_eh(expr * n) {}
void relevant_eh(expr * n) override {}
virtual void init_search_eh() {}
void init_search_eh() override {}
virtual void end_search_eh() {}
void end_search_eh() override {}
virtual void reset() {
void reset() override {
m_queue.reset();
}
virtual void push_scope() {}
void push_scope() override {}
virtual void pop_scope(unsigned num_scopes) {}
void pop_scope(unsigned num_scopes) override {}
virtual void next_case_split(bool_var & next, lbool & phase) {
void next_case_split(bool_var & next, lbool & phase) override {
phase = l_undef;
if (m_context.get_random_value() < static_cast<int>(m_params.m_random_var_freq * random_gen::max_value())) {
@ -129,7 +129,7 @@ namespace smt {
next = null_bool_var;
}
virtual void display(std::ostream & out) {
void display(std::ostream & out) override {
bool first = true;
for (unsigned v : m_queue) {
if (m_context.get_assignment(v) == l_undef) {
@ -144,7 +144,7 @@ namespace smt {
out << "\n";
}
virtual ~act_case_split_queue() {};
~act_case_split_queue() override {};
};
/**
@ -159,7 +159,7 @@ namespace smt {
m_delayed_queue(1024, bool_var_act_lt(ctx.get_activity_vector())) {
}
virtual void activity_increased_eh(bool_var v) {
void activity_increased_eh(bool_var v) override {
act_case_split_queue::activity_increased_eh(v);
if (m_queue.contains(v))
m_queue.decreased(v);
@ -167,7 +167,7 @@ namespace smt {
m_delayed_queue.decreased(v);
}
virtual void mk_var_eh(bool_var v) {
void mk_var_eh(bool_var v) override {
m_queue.reserve(v+1);
m_delayed_queue.reserve(v+1);
SASSERT(!m_delayed_queue.contains(v));
@ -178,28 +178,28 @@ namespace smt {
m_queue.insert(v);
}
virtual void del_var_eh(bool_var v) {
void del_var_eh(bool_var v) override {
act_case_split_queue::del_var_eh(v);
if (m_delayed_queue.contains(v))
m_delayed_queue.erase(v);
}
virtual void relevant_eh(expr * n) {}
void relevant_eh(expr * n) override {}
virtual void init_search_eh() {}
void init_search_eh() override {}
virtual void end_search_eh() {}
void end_search_eh() override {}
virtual void reset() {
void reset() override {
act_case_split_queue::reset();
m_delayed_queue.reset();
}
virtual void push_scope() {}
void push_scope() override {}
virtual void pop_scope(unsigned num_scopes) {}
void pop_scope(unsigned num_scopes) override {}
virtual void next_case_split(bool_var & next, lbool & phase) {
void next_case_split(bool_var & next, lbool & phase) override {
act_case_split_queue::next_case_split(next, phase);
if (next != null_bool_var)
return;
@ -229,7 +229,7 @@ namespace smt {
m_cache_domain(ctx.get_manager()) {
}
virtual void mk_var_eh(bool_var v) {
void mk_var_eh(bool_var v) override {
expr * n = m_context.bool_var2expr(v);
double act;
if (m_cache.find(n, act))
@ -237,7 +237,7 @@ namespace smt {
act_case_split_queue::mk_var_eh(v);
}
virtual void del_var_eh(bool_var v) {
void del_var_eh(bool_var v) override {
if (m_context.is_searching()) {
double act = m_context.get_activity(v);
if (act > 0.0) {
@ -249,14 +249,14 @@ namespace smt {
act_case_split_queue::del_var_eh(v);
}
virtual void init_search_eh() {
void init_search_eh() override {
m_cache.reset();
m_cache_domain.reset();
}
virtual void end_search_eh() {}
void end_search_eh() override {}
virtual void reset() {
void reset() override {
init_search_eh();
}
};
@ -322,15 +322,15 @@ namespace smt {
m_head2(0) {
}
virtual void activity_increased_eh(bool_var v) {}
void activity_increased_eh(bool_var v) override {}
virtual void mk_var_eh(bool_var v) {}
void mk_var_eh(bool_var v) override {}
virtual void del_var_eh(bool_var v) {}
void del_var_eh(bool_var v) override {}
virtual void unassign_var_eh(bool_var v) {}
void unassign_var_eh(bool_var v) override {}
virtual void relevant_eh(expr * n) {
void relevant_eh(expr * n) override {
if (!m_manager.is_bool(n))
return;
bool is_or = m_manager.is_or(n);
@ -361,15 +361,15 @@ namespace smt {
m_queue2.push_back(n);
}
virtual void init_search_eh() {
void init_search_eh() override {
m_bs_num_bool_vars = m_context.get_num_bool_vars();
}
virtual void end_search_eh() {
void end_search_eh() override {
m_bs_num_bool_vars = UINT_MAX;
}
virtual void reset() {
void reset() override {
m_queue.reset();
m_head = 0;
m_queue2.reset();
@ -377,7 +377,7 @@ namespace smt {
m_scopes.reset();
}
virtual void push_scope() {
void push_scope() override {
m_scopes.push_back(scope());
scope & s = m_scopes.back();
s.m_queue_trail = m_queue.size();
@ -387,7 +387,7 @@ namespace smt {
TRACE("case_split", tout << "head: " << m_head << "\n";);
}
virtual void pop_scope(unsigned num_scopes) {
void pop_scope(unsigned num_scopes) override {
SASSERT(num_scopes <= m_scopes.size());
unsigned new_lvl = m_scopes.size() - num_scopes;
scope & s = m_scopes[new_lvl];
@ -443,7 +443,7 @@ namespace smt {
next = null_bool_var;
}
virtual void next_case_split(bool_var & next, lbool & phase) {
void next_case_split(bool_var & next, lbool & phase) override {
next_case_split_core(m_queue, m_head, next, phase);
if (next == null_bool_var)
next_case_split_core(m_queue2, m_head2, next, phase);
@ -471,7 +471,7 @@ namespace smt {
out << "\n";
}
virtual void display(std::ostream & out) {
void display(std::ostream & out) override {
if (m_queue.empty() && m_queue2.empty())
return;
out << "case-splits:\n";
@ -507,9 +507,9 @@ namespace smt {
m_delayed_queue(1024, bool_var_act_lt(ctx.get_activity_vector())) {
}
virtual void activity_increased_eh(bool_var v) {}
void activity_increased_eh(bool_var v) override {}
virtual void mk_var_eh(bool_var v) {
void mk_var_eh(bool_var v) override {
if (m_context.is_searching()) {
SASSERT(v >= m_bs_num_bool_vars);
m_delayed_queue.reserve(v+1);
@ -517,19 +517,19 @@ namespace smt {
}
}
virtual void del_var_eh(bool_var v) {
void del_var_eh(bool_var v) override {
if (v >= m_bs_num_bool_vars && m_delayed_queue.contains(v))
m_delayed_queue.erase(v);
}
virtual void unassign_var_eh(bool_var v) {
void unassign_var_eh(bool_var v) override {
if (v < m_bs_num_bool_vars)
return;
if (!m_delayed_queue.contains(v))
m_delayed_queue.insert(v);
}
virtual void relevant_eh(expr * n) {
void relevant_eh(expr * n) override {
if (!m_manager.is_bool(n))
return;
bool is_or = m_manager.is_or(n);
@ -558,22 +558,22 @@ namespace smt {
m_queue.push_back(n);
}
virtual void init_search_eh() {
void init_search_eh() override {
m_bs_num_bool_vars = m_context.get_num_bool_vars();
}
virtual void end_search_eh() {
void end_search_eh() override {
m_bs_num_bool_vars = UINT_MAX;
}
virtual void reset() {
void reset() override {
m_queue.reset();
m_head = 0;
m_delayed_queue.reset();
m_scopes.reset();
}
virtual void push_scope() {
void push_scope() override {
m_scopes.push_back(scope());
scope & s = m_scopes.back();
s.m_queue_trail = m_queue.size();
@ -581,7 +581,7 @@ namespace smt {
TRACE("case_split", tout << "head: " << m_head << "\n";);
}
virtual void pop_scope(unsigned num_scopes) {
void pop_scope(unsigned num_scopes) override {
SASSERT(num_scopes <= m_scopes.size());
unsigned new_lvl = m_scopes.size() - num_scopes;
scope & s = m_scopes[new_lvl];
@ -632,7 +632,7 @@ namespace smt {
next = null_bool_var;
}
virtual void next_case_split(bool_var & next, lbool & phase) {
void next_case_split(bool_var & next, lbool & phase) override {
if (m_context.get_random_value() < static_cast<int>(0.02 * random_gen::max_value())) {
next = m_context.get_random_value() % m_context.get_num_b_internalized();
TRACE("random_split", tout << "next: " << next << " get_assignment(next): " << m_context.get_assignment(next) << "\n";);
@ -664,7 +664,7 @@ namespace smt {
out << "\n";
}
virtual void display(std::ostream & out) {
void display(std::ostream & out) override {
if (m_queue.empty())
return;
out << "case-splits:\n";
@ -751,15 +751,15 @@ namespace smt {
set_global_generation();
}
virtual void activity_increased_eh(bool_var v) {}
void activity_increased_eh(bool_var v) override {}
virtual void mk_var_eh(bool_var v) {}
void mk_var_eh(bool_var v) override {}
virtual void del_var_eh(bool_var v) {}
void del_var_eh(bool_var v) override {}
virtual void unassign_var_eh(bool_var v) {}
void unassign_var_eh(bool_var v) override {}
virtual void relevant_eh(expr * n) {
void relevant_eh(expr * n) override {
if (get_generation(n) == 0 && m_current_generation != 0)
set_generation_rec(n, m_current_generation);
@ -793,21 +793,21 @@ namespace smt {
add_to_queue2(n);
}
virtual void internalize_instance_eh(expr * e, unsigned gen)
void internalize_instance_eh(expr * e, unsigned gen) override
{
//lower_generation(e, gen);
}
virtual void init_search_eh() {
void init_search_eh() override {
m_bs_num_bool_vars = m_context.get_num_bool_vars();
set_global_generation();
}
virtual void end_search_eh() {
void end_search_eh() override {
m_bs_num_bool_vars = UINT_MAX;
}
virtual void reset() {
void reset() override {
m_queue.reset();
m_head = 0;
m_queue2.reset();
@ -816,7 +816,7 @@ namespace smt {
set_global_generation();
}
virtual void push_scope() {
void push_scope() override {
m_scopes.push_back(scope());
scope & s = m_scopes.back();
s.m_queue_trail = m_queue.size();
@ -827,7 +827,7 @@ namespace smt {
TRACE("case_split", tout << "head: " << m_head << "\n";);
}
virtual void pop_scope(unsigned num_scopes) {
void pop_scope(unsigned num_scopes) override {
SASSERT(num_scopes <= m_scopes.size());
unsigned new_lvl = m_scopes.size() - num_scopes;
scope & s = m_scopes[new_lvl];
@ -898,7 +898,7 @@ namespace smt {
next = null_bool_var;
}
virtual void next_case_split(bool_var & next, lbool & phase) {
void next_case_split(bool_var & next, lbool & phase) override {
phase = l_undef;
next = null_bool_var;
@ -943,7 +943,7 @@ namespace smt {
out << "\n";
}
virtual void display(std::ostream & out) {
void display(std::ostream & out) override {
if (m_queue.empty() && m_queue2.empty())
return;
out << "case-splits:\n";
@ -951,7 +951,7 @@ namespace smt {
//display_core(out, m_queue2, m_head2, 2);
}
virtual void assign_lit_eh(literal l) {
void assign_lit_eh(literal l) override {
// if (m_current_generation > stop_gen)
// m_current_generation--;
@ -1128,41 +1128,41 @@ namespace smt {
m_queue(1024, theory_aware_act_lt(ctx.get_activity_vector(), m_theory_var_priority)) {
}
virtual void activity_increased_eh(bool_var v) {
void activity_increased_eh(bool_var v) override {
if (m_queue.contains(v))
m_queue.decreased(v);
}
virtual void mk_var_eh(bool_var v) {
void mk_var_eh(bool_var v) override {
m_queue.reserve(v+1);
m_queue.insert(v);
}
virtual void del_var_eh(bool_var v) {
void del_var_eh(bool_var v) override {
if (m_queue.contains(v))
m_queue.erase(v);
}
virtual void unassign_var_eh(bool_var v) {
void unassign_var_eh(bool_var v) override {
if (!m_queue.contains(v))
m_queue.insert(v);
}
virtual void relevant_eh(expr * n) {}
void relevant_eh(expr * n) override {}
virtual void init_search_eh() {}
void init_search_eh() override {}
virtual void end_search_eh() {}
void end_search_eh() override {}
virtual void reset() {
void reset() override {
m_queue.reset();
}
virtual void push_scope() {}
void push_scope() override {}
virtual void pop_scope(unsigned num_scopes) {}
void pop_scope(unsigned num_scopes) override {}
virtual void next_case_split(bool_var & next, lbool & phase) {
void next_case_split(bool_var & next, lbool & phase) override {
int threshold = static_cast<int>(m_params.m_random_var_freq * random_gen::max_value());
SASSERT(threshold >= 0);
if (m_context.get_random_value() < threshold) {
@ -1187,7 +1187,7 @@ namespace smt {
}
}
virtual void add_theory_aware_branching_info(bool_var v, double priority, lbool phase) {
void add_theory_aware_branching_info(bool_var v, double priority, lbool phase) override {
TRACE("theory_aware_branching", tout << "Add theory-aware branching information for l#" << v << ": priority=" << priority << std::endl;);
m_theory_vars.insert(v);
m_theory_var_phase.insert(v, phase);
@ -1203,7 +1203,7 @@ namespace smt {
// m_theory_queue.insert(v);
}
virtual void display(std::ostream & out) {
void display(std::ostream & out) override {
bool first = true;
bool_var_act_queue::const_iterator it = m_queue.begin();
bool_var_act_queue::const_iterator end = m_queue.end();
@ -1222,7 +1222,7 @@ namespace smt {
}
virtual ~theory_aware_branching_queue() {};
~theory_aware_branching_queue() override {};
};

View file

@ -476,7 +476,7 @@ namespace smt {
m_r2_num_parents(r2_num_parents) {
}
virtual void undo(context & ctx) {
void undo(context & ctx) override {
ctx.undo_add_eq(m_r1, m_n1, m_r2_num_parents);
}
};
@ -1451,7 +1451,7 @@ namespace smt {
bool_var m_var;
public:
set_var_theory_trail(bool_var v):m_var(v) {}
virtual void undo(context & ctx) {
void undo(context & ctx) override {
bool_var_data & d = ctx.m_bdata[m_var];
d.reset_notify_theory();
}
@ -2951,7 +2951,7 @@ namespace smt {
case_split_insert_trail(literal l):
l(l) {
}
virtual void undo(context & ctx) {
void undo(context & ctx) override {
ctx.undo_th_case_split(l);
}
};
@ -4131,7 +4131,7 @@ namespace smt {
bool_var m_var;
public:
set_true_first_trail(bool_var v):m_var(v) {}
virtual void undo(context & ctx) {
void undo(context & ctx) override {
ctx.m_bdata[m_var].reset_true_first_flag();
}
};

View file

@ -746,7 +746,7 @@ namespace smt {
friend class mk_bool_var_trail;
class mk_bool_var_trail : public trail<context> {
public:
virtual void undo(context & ctx) { ctx.undo_mk_bool_var(); }
void undo(context & ctx) override { ctx.undo_mk_bool_var(); }
};
mk_bool_var_trail m_mk_bool_var_trail;
@ -755,7 +755,7 @@ namespace smt {
friend class mk_enode_trail;
class mk_enode_trail : public trail<context> {
public:
virtual void undo(context & ctx) { ctx.undo_mk_enode(); }
void undo(context & ctx) override { ctx.undo_mk_enode(); }
};
mk_enode_trail m_mk_enode_trail;

View file

@ -93,8 +93,8 @@ namespace smt {
for_each_relevant_expr(ctx),
m_buffer(b) {
}
virtual ~collect_relevant_label_lits() {}
virtual void operator()(expr * n);
~collect_relevant_label_lits() override {}
void operator()(expr * n) override;
};
class collect_relevant_labels : public for_each_relevant_expr {
@ -104,8 +104,8 @@ namespace smt {
for_each_relevant_expr(ctx),
m_buffer(b) {
}
virtual ~collect_relevant_labels() {}
virtual void operator()(expr * n);
~collect_relevant_labels() override {}
void operator()(expr * n) override;
};
};

View file

@ -631,7 +631,7 @@ namespace smt {
set_merge_tf_trail(enode * n):
m_node(n) {
}
virtual void undo(context & ctx) {
void undo(context & ctx) override {
m_node->m_merge_tf = false;
}
};
@ -667,7 +667,7 @@ namespace smt {
set_enode_flag_trail(bool_var v):
m_var(v) {
}
virtual void undo(context & ctx) {
void undo(context & ctx) override {
bool_var_data & data = ctx.m_bdata[m_var];
data.reset_enode_flag();
}
@ -1612,7 +1612,7 @@ namespace smt {
SASSERT(m_th_var != null_theory_var);
}
virtual void undo(context & ctx) {
void undo(context & ctx) override {
theory_var v = m_enode->get_th_var(m_th_id);
SASSERT(v != null_theory_var);
SASSERT(m_th_var == v);
@ -1637,7 +1637,7 @@ namespace smt {
m_old_th_var(old_var) {
}
virtual void undo(context & ctx) {
void undo(context & ctx) override {
SASSERT(m_enode->get_th_var(m_th_id) != null_theory_var);
m_enode->replace_th_var(m_old_th_var, m_th_id);
}

View file

@ -102,15 +102,15 @@ namespace smt {
public:
justification_proof_wrapper(context & ctx, proof * pr, bool in_region = true);
virtual bool has_del_eh() const {
bool has_del_eh() const override {
return true;
}
virtual void del_eh(ast_manager & m);
void del_eh(ast_manager & m) override;
virtual proof * mk_proof(conflict_resolution & cr);
proof * mk_proof(conflict_resolution & cr) override;
virtual char const * get_name() const { return "proof-wrapper"; }
char const * get_name() const override { return "proof-wrapper"; }
};
class unit_resolution_justification : public justification {
@ -122,21 +122,21 @@ namespace smt {
unit_resolution_justification(justification * js, unsigned num_lits, literal const * lits);
~unit_resolution_justification();
~unit_resolution_justification() override;
virtual bool has_del_eh() const {
bool has_del_eh() const override {
return !in_region() && m_antecedent && m_antecedent->has_del_eh();
}
virtual void del_eh(ast_manager & m) {
void del_eh(ast_manager & m) override {
if (!in_region() && m_antecedent) m_antecedent->del_eh(m);
}
virtual void get_antecedents(conflict_resolution & cr);
void get_antecedents(conflict_resolution & cr) override;
virtual proof * mk_proof(conflict_resolution & cr);
proof * mk_proof(conflict_resolution & cr) override;
virtual char const * get_name() const { return "unit-resolution"; }
char const * get_name() const override { return "unit-resolution"; }
};
class eq_conflict_justification : public justification {
@ -150,11 +150,11 @@ namespace smt {
m_js(js) {
}
virtual void get_antecedents(conflict_resolution & cr);
void get_antecedents(conflict_resolution & cr) override;
virtual proof * mk_proof(conflict_resolution & cr);
proof * mk_proof(conflict_resolution & cr) override;
virtual char const * get_name() const { return "eq-conflict"; }
char const * get_name() const override { return "eq-conflict"; }
};
/**
@ -166,11 +166,11 @@ namespace smt {
eq_root_propagation_justification(enode * n):m_node(n) {
}
virtual void get_antecedents(conflict_resolution & cr);
void get_antecedents(conflict_resolution & cr) override;
virtual proof * mk_proof(conflict_resolution & cr);
proof * mk_proof(conflict_resolution & cr) override;
virtual char const * get_name() const { return "eq-root"; }
char const * get_name() const override { return "eq-root"; }
};
/**
@ -184,11 +184,11 @@ namespace smt {
SASSERT(n1 != n2);
}
virtual void get_antecedents(conflict_resolution & cr);
void get_antecedents(conflict_resolution & cr) override;
virtual proof * mk_proof(conflict_resolution & cr);
proof * mk_proof(conflict_resolution & cr) override;
virtual char const * get_name() const { return "eq-propagation"; }
char const * get_name() const override { return "eq-propagation"; }
};
/**
@ -201,11 +201,11 @@ namespace smt {
mp_iff_justification(enode * n1, enode * n2):m_node1(n1), m_node2(n2) {
}
virtual void get_antecedents(conflict_resolution & cr);
void get_antecedents(conflict_resolution & cr) override;
virtual proof * mk_proof(conflict_resolution & cr);
proof * mk_proof(conflict_resolution & cr) override;
virtual char const * get_name() const { return "mp-iff"; }
char const * get_name() const override { return "mp-iff"; }
};
/**
@ -221,11 +221,11 @@ namespace smt {
public:
simple_justification(region & r, unsigned num_lits, literal const * lits);
virtual void get_antecedents(conflict_resolution & cr);
void get_antecedents(conflict_resolution & cr) override;
virtual proof * mk_proof(conflict_resolution & cr) = 0;
proof * mk_proof(conflict_resolution & cr) override = 0;
virtual char const * get_name() const { return "simple"; }
char const * get_name() const override { return "simple"; }
};
@ -240,13 +240,13 @@ namespace smt {
unsigned num_params, parameter* params):
simple_justification(r, num_lits, lits),
m_th_id(fid), m_params(num_params, params) {}
virtual ~simple_theory_justification() {}
~simple_theory_justification() override {}
virtual bool has_del_eh() const { return !m_params.empty(); }
bool has_del_eh() const override { return !m_params.empty(); }
virtual void del_eh(ast_manager & m) { m_params.reset(); }
void del_eh(ast_manager & m) override { m_params.reset(); }
virtual theory_id get_from_theory() const { return m_th_id; }
theory_id get_from_theory() const override { return m_th_id; }
};
@ -258,11 +258,11 @@ namespace smt {
unsigned num_params = 0, parameter* params = 0):
simple_theory_justification(fid, r, num_lits, lits, num_params, params) {}
virtual void get_antecedents(conflict_resolution & cr) {}
void get_antecedents(conflict_resolution & cr) override {}
virtual proof * mk_proof(conflict_resolution & cr);
proof * mk_proof(conflict_resolution & cr) override;
virtual char const * get_name() const { return "theory-axiom"; }
char const * get_name() const override { return "theory-axiom"; }
};
class theory_propagation_justification : public simple_theory_justification {
@ -272,10 +272,10 @@ namespace smt {
unsigned num_params = 0, parameter* params = 0):
simple_theory_justification(fid, r, num_lits, lits, num_params, params), m_consequent(consequent) {}
virtual proof * mk_proof(conflict_resolution & cr);
proof * mk_proof(conflict_resolution & cr) override;
virtual char const * get_name() const { return "theory-propagation"; }
char const * get_name() const override { return "theory-propagation"; }
};
@ -285,9 +285,9 @@ namespace smt {
unsigned num_params = 0, parameter* params = 0):
simple_theory_justification(fid, r, num_lits, lits, num_params, params) {}
virtual proof * mk_proof(conflict_resolution & cr);
proof * mk_proof(conflict_resolution & cr) override;
virtual char const * get_name() const { return "theory-conflict"; }
char const * get_name() const override { return "theory-conflict"; }
};
/**
@ -304,11 +304,11 @@ namespace smt {
ext_simple_justification(region & r, unsigned num_lits, literal const * lits,
unsigned num_eqs, enode_pair const * eqs);
virtual void get_antecedents(conflict_resolution & cr);
void get_antecedents(conflict_resolution & cr) override;
virtual proof * mk_proof(conflict_resolution & cr) = 0;
proof * mk_proof(conflict_resolution & cr) override = 0;
virtual char const * get_name() const { return "ext-simple"; }
char const * get_name() const override { return "ext-simple"; }
};
/**
@ -325,13 +325,13 @@ namespace smt {
unsigned num_params = 0, parameter* params = 0):
ext_simple_justification(r, num_lits, lits, num_eqs, eqs), m_th_id(fid), m_params(num_params, params) {}
virtual ~ext_theory_simple_justification() {}
~ext_theory_simple_justification() override {}
virtual bool has_del_eh() const { return !m_params.empty(); }
bool has_del_eh() const override { return !m_params.empty(); }
virtual void del_eh(ast_manager & m) { m_params.reset(); }
void del_eh(ast_manager & m) override { m_params.reset(); }
virtual theory_id get_from_theory() const { return m_th_id; }
theory_id get_from_theory() const override { return m_th_id; }
};
class ext_theory_propagation_justification : public ext_theory_simple_justification {
@ -345,9 +345,9 @@ namespace smt {
ext_theory_simple_justification(fid, r, num_lits, lits, num_eqs, eqs, num_params, params),
m_consequent(consequent) {}
virtual proof * mk_proof(conflict_resolution & cr);
proof * mk_proof(conflict_resolution & cr) override;
virtual char const * get_name() const { return "ext-theory-propagation"; }
char const * get_name() const override { return "ext-theory-propagation"; }
};
class ext_theory_conflict_justification : public ext_theory_simple_justification {
@ -357,9 +357,9 @@ namespace smt {
unsigned num_params = 0, parameter* params = 0):
ext_theory_simple_justification(fid, r, num_lits, lits, num_eqs, eqs, num_params, params) {}
virtual proof * mk_proof(conflict_resolution & cr);
proof * mk_proof(conflict_resolution & cr) override;
virtual char const * get_name() const { return "ext-theory-conflict"; }
char const * get_name() const override { return "ext-theory-conflict"; }
};
class ext_theory_eq_propagation_justification : public ext_theory_simple_justification {
@ -374,9 +374,9 @@ namespace smt {
unsigned num_params = 0, parameter* params = 0):
ext_theory_simple_justification(fid, r, num_lits, lits, num_eqs, eqs, num_params, params), m_lhs(lhs), m_rhs(rhs) {}
virtual proof * mk_proof(conflict_resolution & cr);
proof * mk_proof(conflict_resolution & cr) override;
virtual char const * get_name() const { return "ext-theory-eq-propagation"; }
char const * get_name() const override { return "ext-theory-eq-propagation"; }
};
/**
@ -394,19 +394,19 @@ namespace smt {
theory_lemma_justification(family_id fid, context & ctx, unsigned num_lits, literal const * lits,
unsigned num_params = 0, parameter* params = 0);
virtual ~theory_lemma_justification();
~theory_lemma_justification() override;
virtual bool has_del_eh() const {
bool has_del_eh() const override {
return true;
}
virtual void del_eh(ast_manager & m);
void del_eh(ast_manager & m) override;
virtual void get_antecedents(conflict_resolution & cr) {}
void get_antecedents(conflict_resolution & cr) override {}
virtual proof * mk_proof(conflict_resolution & cr);
proof * mk_proof(conflict_resolution & cr) override;
virtual char const * get_name() const { return "theory-lemma"; }
char const * get_name() const override { return "theory-lemma"; }
};
};

View file

@ -566,7 +566,7 @@ namespace smt {
}
}
virtual expr * eval(expr * n, bool model_completion) {
expr * eval(expr * n, bool model_completion) override {
expr * r = 0;
if (m_eval_cache[model_completion].find(n, r)) {
return r;
@ -1134,24 +1134,24 @@ namespace smt {
unsigned m_var_j;
public:
f_var(func_decl * f, unsigned i, unsigned j):m_f(f), m_arg_i(i), m_var_j(j) {}
virtual ~f_var() {}
~f_var() override {}
virtual char const * get_kind() const {
char const * get_kind() const override {
return "f_var";
}
virtual bool is_equal(qinfo const * qi) const {
bool is_equal(qinfo const * qi) const override {
if (qi->get_kind() != get_kind())
return false;
f_var const * other = static_cast<f_var const *>(qi);
return m_f == other->m_f && m_arg_i == other->m_arg_i && m_var_j == other->m_var_j;
}
virtual void display(std::ostream & out) const {
void display(std::ostream & out) const override {
out << "(" << m_f->get_name() << ":" << m_arg_i << " -> v!" << m_var_j << ")";
}
virtual void process_auf(quantifier * q, auf_solver & s, context * ctx) {
void process_auf(quantifier * q, auf_solver & s, context * ctx) override {
node * n1 = s.get_A_f_i(m_f, m_arg_i);
node * n2 = s.get_uvar(q, m_var_j);
CTRACE("model_finder", n1->get_sort() != n2->get_sort(),
@ -1170,7 +1170,7 @@ namespace smt {
n1->merge(n2);
}
virtual void populate_inst_sets(quantifier * q, auf_solver & s, context * ctx) {
void populate_inst_sets(quantifier * q, auf_solver & s, context * ctx) override {
node * A_f_i = s.get_A_f_i(m_f, m_arg_i);
enode_vector::const_iterator it = ctx->begin_enodes_of(m_f);
enode_vector::const_iterator end = ctx->end_enodes_of(m_f);
@ -1193,7 +1193,7 @@ namespace smt {
}
}
virtual void populate_inst_sets(quantifier * q, func_decl * mhead, ptr_vector<instantiation_set> & uvar_inst_sets, context * ctx) {
void populate_inst_sets(quantifier * q, func_decl * mhead, ptr_vector<instantiation_set> & uvar_inst_sets, context * ctx) override {
if (m_f != mhead)
return;
uvar_inst_sets.reserve(m_var_j+1, 0);
@ -1222,31 +1222,31 @@ namespace smt {
f_var(f, i, j),
m_offset(offset, m) {
}
virtual ~f_var_plus_offset() {}
~f_var_plus_offset() override {}
virtual char const * get_kind() const {
char const * get_kind() const override {
return "f_var_plus_offset";
}
virtual bool is_equal(qinfo const * qi) const {
bool is_equal(qinfo const * qi) const override {
if (qi->get_kind() != get_kind())
return false;
f_var_plus_offset const * other = static_cast<f_var_plus_offset const *>(qi);
return m_f == other->m_f && m_arg_i == other->m_arg_i && m_var_j == other->m_var_j && m_offset.get() == other->m_offset.get();
}
virtual void display(std::ostream & out) const {
void display(std::ostream & out) const override {
out << "(" << m_f->get_name() << ":" << m_arg_i << " - " <<
mk_bounded_pp(m_offset.get(), m_offset.get_manager()) << " -> v!" << m_var_j << ")";
}
virtual void process_auf(quantifier * q, auf_solver & s, context * ctx) {
void process_auf(quantifier * q, auf_solver & s, context * ctx) override {
// just create the nodes
/* node * A_f_i = */ s.get_A_f_i(m_f, m_arg_i);
/* node * S_j = */ s.get_uvar(q, m_var_j);
}
virtual void populate_inst_sets(quantifier * q, auf_solver & s, context * ctx) {
void populate_inst_sets(quantifier * q, auf_solver & s, context * ctx) override {
// S_j is not necessary equal to A_f_i.
node * A_f_i = s.get_A_f_i(m_f, m_arg_i)->get_root();
node * S_j = s.get_uvar(q, m_var_j)->get_root();
@ -1318,7 +1318,7 @@ namespace smt {
}
}
virtual void populate_inst_sets2(quantifier * q, auf_solver & s, context * ctx) {
void populate_inst_sets2(quantifier * q, auf_solver & s, context * ctx) override {
node * A_f_i = s.get_A_f_i(m_f, m_arg_i)->get_root();
node * S_j = s.get_uvar(q, m_var_j)->get_root();
// If A_f_i == S_j, then there is no finite fixpoint, so we do nothing here.
@ -1331,7 +1331,7 @@ namespace smt {
}
}
virtual void populate_inst_sets(quantifier * q, func_decl * mhead, ptr_vector<instantiation_set> & uvar_inst_sets, context * ctx) {
void populate_inst_sets(quantifier * q, func_decl * mhead, ptr_vector<instantiation_set> & uvar_inst_sets, context * ctx) override {
// ignored when in macro
}
@ -1391,24 +1391,24 @@ namespace smt {
public:
select_var(ast_manager & m, app * s, unsigned i, unsigned j):m_manager(m), m_array(m), m_select(s), m_arg_i(i), m_var_j(j) {}
virtual ~select_var() {}
~select_var() override {}
virtual char const * get_kind() const {
char const * get_kind() const override {
return "select_var";
}
virtual bool is_equal(qinfo const * qi) const {
bool is_equal(qinfo const * qi) const override {
if (qi->get_kind() != get_kind())
return false;
select_var const * other = static_cast<select_var const *>(qi);
return m_select == other->m_select && m_arg_i == other->m_arg_i && m_var_j == other->m_var_j;
}
virtual void display(std::ostream & out) const {
void display(std::ostream & out) const override {
out << "(" << mk_bounded_pp(m_select, m_manager) << ":" << m_arg_i << " -> v!" << m_var_j << ")";
}
virtual void process_auf(quantifier * q, auf_solver & s, context * ctx) {
void process_auf(quantifier * q, auf_solver & s, context * ctx) override {
ptr_buffer<enode> arrays;
get_auf_arrays(get_array(), ctx, arrays);
TRACE("select_var",
@ -1428,7 +1428,7 @@ namespace smt {
}
}
virtual void populate_inst_sets(quantifier * q, auf_solver & s, context * ctx) {
void populate_inst_sets(quantifier * q, auf_solver & s, context * ctx) override {
ptr_buffer<enode> arrays;
get_auf_arrays(get_array(), ctx, arrays);
for (enode * curr : arrays) {
@ -1461,20 +1461,20 @@ namespace smt {
std::swap(m_var_i, m_var_j);
}
virtual ~var_pair() {}
~var_pair() override {}
virtual bool is_equal(qinfo const * qi) const {
bool is_equal(qinfo const * qi) const override {
if (qi->get_kind() != get_kind())
return false;
var_pair const * other = static_cast<var_pair const *>(qi);
return m_var_i == other->m_var_i && m_var_j == other->m_var_j;
}
virtual void display(std::ostream & out) const {
void display(std::ostream & out) const override {
out << "(" << get_kind() << ":v!" << m_var_i << ":v!" << m_var_j << ")";
}
virtual void populate_inst_sets(quantifier * q, auf_solver & s, context * ctx) {
void populate_inst_sets(quantifier * q, auf_solver & s, context * ctx) override {
// do nothing
}
};
@ -1482,9 +1482,9 @@ namespace smt {
class x_eq_y : public var_pair {
public:
x_eq_y(unsigned i, unsigned j):var_pair(i, j) {}
virtual char const * get_kind() const { return "x_eq_y"; }
char const * get_kind() const override { return "x_eq_y"; }
virtual void process_auf(quantifier * q, auf_solver & s, context * ctx) {
void process_auf(quantifier * q, auf_solver & s, context * ctx) override {
node * n1 = s.get_uvar(q, m_var_i);
node * n2 = s.get_uvar(q, m_var_j);
n1->insert_avoid(n2);
@ -1496,9 +1496,9 @@ namespace smt {
class x_neq_y : public var_pair {
public:
x_neq_y(unsigned i, unsigned j):var_pair(i, j) {}
virtual char const * get_kind() const { return "x_neq_y"; }
char const * get_kind() const override { return "x_neq_y"; }
virtual void process_auf(quantifier * q, auf_solver & s, context * ctx) {
void process_auf(quantifier * q, auf_solver & s, context * ctx) override {
node * n1 = s.get_uvar(q, m_var_i);
node * n2 = s.get_uvar(q, m_var_j);
n1->merge(n2);
@ -1508,9 +1508,9 @@ namespace smt {
class x_leq_y : public var_pair {
public:
x_leq_y(unsigned i, unsigned j):var_pair(i, j) {}
virtual char const * get_kind() const { return "x_leq_y"; }
char const * get_kind() const override { return "x_leq_y"; }
virtual void process_auf(quantifier * q, auf_solver & s, context * ctx) {
void process_auf(quantifier * q, auf_solver & s, context * ctx) override {
node * n1 = s.get_uvar(q, m_var_i);
node * n2 = s.get_uvar(q, m_var_j);
n1->merge(n2);
@ -1522,9 +1522,9 @@ namespace smt {
class x_sleq_y : public x_leq_y {
public:
x_sleq_y(unsigned i, unsigned j):x_leq_y(i, j) {}
virtual char const * get_kind() const { return "x_sleq_y"; }
char const * get_kind() const override { return "x_sleq_y"; }
virtual void process_auf(quantifier * q, auf_solver & s, context * ctx) {
void process_auf(quantifier * q, auf_solver & s, context * ctx) override {
node * n1 = s.get_uvar(q, m_var_i);
node * n2 = s.get_uvar(q, m_var_j);
n1->merge(n2);
@ -1540,16 +1540,16 @@ namespace smt {
public:
var_expr_pair(ast_manager & m, unsigned i, expr * t):
m_var_i(i), m_t(t, m) {}
~var_expr_pair() {}
~var_expr_pair() override {}
virtual bool is_equal(qinfo const * qi) const {
bool is_equal(qinfo const * qi) const override {
if (qi->get_kind() != get_kind())
return false;
var_expr_pair const * other = static_cast<var_expr_pair const *>(qi);
return m_var_i == other->m_var_i && m_t.get() == other->m_t.get();
}
virtual void display(std::ostream & out) const {
void display(std::ostream & out) const override {
out << "(" << get_kind() << ":v!" << m_var_i << ":" << mk_bounded_pp(m_t.get(), m_t.get_manager()) << ")";
}
};
@ -1558,14 +1558,14 @@ namespace smt {
public:
x_eq_t(ast_manager & m, unsigned i, expr * t):
var_expr_pair(m, i, t) {}
virtual char const * get_kind() const { return "x_eq_t"; }
char const * get_kind() const override { return "x_eq_t"; }
virtual void process_auf(quantifier * q, auf_solver & s, context * ctx) {
void process_auf(quantifier * q, auf_solver & s, context * ctx) override {
node * n1 = s.get_uvar(q, m_var_i);
n1->insert_exception(m_t);
}
virtual void populate_inst_sets(quantifier * q, auf_solver & slv, context * ctx) {
void populate_inst_sets(quantifier * q, auf_solver & slv, context * ctx) override {
unsigned num_vars = q->get_num_decls();
ast_manager & m = ctx->get_manager();
sort * s = q->get_decl_sort(num_vars - m_var_i - 1);
@ -1589,14 +1589,14 @@ namespace smt {
public:
x_neq_t(ast_manager & m, unsigned i, expr * t):
var_expr_pair(m, i, t) {}
virtual char const * get_kind() const { return "x_neq_t"; }
char const * get_kind() const override { return "x_neq_t"; }
virtual void process_auf(quantifier * q, auf_solver & s, context * ctx) {
void process_auf(quantifier * q, auf_solver & s, context * ctx) override {
// make sure that S_q_i is create.
s.get_uvar(q, m_var_i);
}
virtual void populate_inst_sets(quantifier * q, auf_solver & s, context * ctx) {
void populate_inst_sets(quantifier * q, auf_solver & s, context * ctx) override {
node * S_q_i = s.get_uvar(q, m_var_i);
S_q_i->insert(m_t, 0);
}
@ -1606,15 +1606,15 @@ namespace smt {
public:
x_gle_t(ast_manager & m, unsigned i, expr * t):
var_expr_pair(m, i, t) {}
virtual char const * get_kind() const { return "x_gle_t"; }
char const * get_kind() const override { return "x_gle_t"; }
virtual void process_auf(quantifier * q, auf_solver & s, context * ctx) {
void process_auf(quantifier * q, auf_solver & s, context * ctx) override {
// make sure that S_q_i is create.
node * n1 = s.get_uvar(q, m_var_i);
n1->set_mono_proj();
}
virtual void populate_inst_sets(quantifier * q, auf_solver & s, context * ctx) {
void populate_inst_sets(quantifier * q, auf_solver & s, context * ctx) override {
node * S_q_i = s.get_uvar(q, m_var_i);
S_q_i->insert(m_t, 0);
}
@ -2524,7 +2524,7 @@ namespace smt {
return false;
}
virtual bool process(ptr_vector<quantifier> const & qs, ptr_vector<quantifier> & new_qs, ptr_vector<quantifier> & residue) {
bool process(ptr_vector<quantifier> const & qs, ptr_vector<quantifier> & new_qs, ptr_vector<quantifier> & residue) override {
bool removed = false;
for (quantifier* q : qs) {
if (process(q, qs))
@ -2957,7 +2957,7 @@ namespace smt {
m_fs.reset();
}
virtual bool process(ptr_vector<quantifier> const & qs, ptr_vector<quantifier> & new_qs, ptr_vector<quantifier> & residue) {
bool process(ptr_vector<quantifier> const & qs, ptr_vector<quantifier> & new_qs, ptr_vector<quantifier> & residue) override {
reset();
ptr_vector<quantifier> qcandidates;
preprocess(qs, qcandidates, new_qs);
@ -2988,7 +2988,7 @@ namespace smt {
m_satisfied(ev_handler(this)) {
}
virtual ~hint_solver() {
~hint_solver() override {
reset();
}
@ -3115,7 +3115,7 @@ namespace smt {
}
}
virtual bool process(ptr_vector<quantifier> const & qs, ptr_vector<quantifier> & new_qs, ptr_vector<quantifier> & residue) {
bool process(ptr_vector<quantifier> const & qs, ptr_vector<quantifier> & new_qs, ptr_vector<quantifier> & residue) override {
obj_map<func_decl, mq_pair> full_macros;
func_decl_set cond_macros;
obj_hashtable<quantifier> removed;

View file

@ -159,16 +159,16 @@ namespace smt {
app * m_value;
public:
expr_wrapper_proc(app * v):m_value(v) {}
virtual app * mk_value(model_generator & m, ptr_vector<expr> & values) { return m_value; }
app * mk_value(model_generator & m, ptr_vector<expr> & values) override { return m_value; }
};
class fresh_value_proc : public model_value_proc {
extra_fresh_value * m_value;
public:
fresh_value_proc(extra_fresh_value * v):m_value(v) {}
virtual void get_dependencies(buffer<model_value_dependency> & result) { result.push_back(m_value); }
virtual app * mk_value(model_generator & m, ptr_vector<expr> & values) { return to_app(values[0]); }
virtual bool is_fresh() const { return true; }
void get_dependencies(buffer<model_value_dependency> & result) override { result.push_back(m_value); }
app * mk_value(model_generator & m, ptr_vector<expr> & values) override { return to_app(values[0]); }
bool is_fresh() const override { return true; }
};
/**

View file

@ -415,10 +415,10 @@ namespace smt {
m_active(false) {
}
virtual ~default_qm_plugin() {
~default_qm_plugin() override {
}
virtual void set_manager(quantifier_manager & qm) {
void set_manager(quantifier_manager & qm) override {
SASSERT(m_qm == 0);
m_qm = &qm;
m_context = &(qm.get_context());
@ -434,11 +434,11 @@ namespace smt {
m_model_checker->set_qm(qm);
}
virtual quantifier_manager_plugin * mk_fresh() { return alloc(default_qm_plugin); }
quantifier_manager_plugin * mk_fresh() override { return alloc(default_qm_plugin); }
virtual bool model_based() const { return m_fparams->m_mbqi; }
bool model_based() const override { return m_fparams->m_mbqi; }
virtual bool mbqi_enabled(quantifier *q) const {
bool mbqi_enabled(quantifier *q) const override {
if (!m_fparams->m_mbqi_id) return true;
const symbol &s = q->get_qid();
size_t len = strlen(m_fparams->m_mbqi_id);
@ -450,16 +450,16 @@ namespace smt {
/* Quantifier id's must begin with the prefix specified by parameter
mbqi.id to be instantiated with MBQI. The default value is the
empty string, so all quantifiers are instantiated. */
virtual void add(quantifier * q) {
void add(quantifier * q) override {
if (m_fparams->m_mbqi && mbqi_enabled(q)) {
m_active = true;
m_model_finder->register_quantifier(q);
}
}
virtual void del(quantifier * q) { }
void del(quantifier * q) override { }
virtual void push() {
void push() override {
m_mam->push_scope();
m_lazy_mam->push_scope();
if (m_fparams->m_mbqi) {
@ -467,7 +467,7 @@ namespace smt {
}
}
virtual void pop(unsigned num_scopes) {
void pop(unsigned num_scopes) override {
m_mam->pop_scope(num_scopes);
m_lazy_mam->pop_scope(num_scopes);
if (m_fparams->m_mbqi) {
@ -475,7 +475,7 @@ namespace smt {
}
}
virtual void init_search_eh() {
void init_search_eh() override {
m_lazy_matching_idx = 0;
if (m_fparams->m_mbqi) {
m_model_finder->init_search_eh();
@ -483,7 +483,7 @@ namespace smt {
}
}
virtual void assign_eh(quantifier * q) {
void assign_eh(quantifier * q) override {
m_active = true;
ast_manager& m = m_context->get_manager();
if (!m_fparams->m_ematching) {
@ -531,23 +531,23 @@ namespace smt {
return m_fparams->m_ematching && !m_qm->empty();
}
virtual void add_eq_eh(enode * e1, enode * e2) {
void add_eq_eh(enode * e1, enode * e2) override {
if (use_ematching())
m_mam->add_eq_eh(e1, e2);
}
virtual void relevant_eh(enode * e) {
void relevant_eh(enode * e) override {
if (use_ematching()) {
m_mam->relevant_eh(e, false);
m_lazy_mam->relevant_eh(e, true);
}
}
virtual bool can_propagate() const {
bool can_propagate() const override {
return m_mam->has_work();
}
virtual void restart_eh() {
void restart_eh() override {
if (m_fparams->m_mbqi) {
m_model_finder->restart_eh();
m_model_checker->restart_eh();
@ -555,17 +555,17 @@ namespace smt {
TRACE("mam_stats", m_mam->display(tout););
}
virtual bool is_shared(enode * n) const {
bool is_shared(enode * n) const override {
return m_active && (m_mam->is_shared(n) || m_lazy_mam->is_shared(n));
}
virtual void adjust_model(proto_model * m) {
void adjust_model(proto_model * m) override {
if (m_fparams->m_mbqi) {
m_model_finder->fix_model(m);
}
}
virtual void propagate() {
void propagate() override {
m_mam->match();
if (!m_context->relevancy() && use_ematching()) {
ptr_vector<enode>::const_iterator it = m_context->begin_enodes();
@ -585,8 +585,8 @@ namespace smt {
}
}
virtual quantifier_manager::check_model_result
check_model(proto_model * m, obj_map<enode, app *> const & root2value) {
quantifier_manager::check_model_result
check_model(proto_model * m, obj_map<enode, app *> const & root2value) override {
if (m_fparams->m_mbqi) {
IF_VERBOSE(10, verbose_stream() << "(smt.mbqi)\n";);
if (m_model_checker->check(m, root2value)) {
@ -599,7 +599,7 @@ namespace smt {
return quantifier_manager::UNKNOWN;
}
virtual final_check_status final_check_eh(bool full) {
final_check_status final_check_eh(bool full) override {
if (!full) {
if (m_fparams->m_qi_lazy_instantiation)
return final_check_quant();

View file

@ -52,24 +52,24 @@ namespace smt {
app * m_parent;
public:
and_relevancy_eh(app * p):m_parent(p) {}
virtual ~and_relevancy_eh() {}
virtual void operator()(relevancy_propagator & rp);
~and_relevancy_eh() override {}
void operator()(relevancy_propagator & rp) override;
};
class or_relevancy_eh : public relevancy_eh {
app * m_parent;
public:
or_relevancy_eh(app * p):m_parent(p) {}
virtual ~or_relevancy_eh() {}
virtual void operator()(relevancy_propagator & rp);
~or_relevancy_eh() override {}
void operator()(relevancy_propagator & rp) override;
};
class ite_relevancy_eh : public relevancy_eh {
app * m_parent;
public:
ite_relevancy_eh(app * p):m_parent(p) {}
virtual ~ite_relevancy_eh() {}
virtual void operator()(relevancy_propagator & rp);
~ite_relevancy_eh() override {}
void operator()(relevancy_propagator & rp) override;
};
class ite_term_relevancy_eh : public relevancy_eh {
@ -78,8 +78,8 @@ namespace smt {
app * m_else_eq;
public:
ite_term_relevancy_eh(app * p, app * then_eq, app * else_eq):m_parent(p), m_then_eq(then_eq), m_else_eq(else_eq) {}
virtual ~ite_term_relevancy_eh() {}
virtual void operator()(relevancy_propagator & rp);
~ite_term_relevancy_eh() override {}
void operator()(relevancy_propagator & rp) override;
};
relevancy_propagator::relevancy_propagator(context & ctx):
@ -154,7 +154,7 @@ namespace smt {
relevancy_propagator(ctx), m_qhead(0), m_relevant_exprs(ctx.get_manager()),
m_propagating(false) {}
virtual ~relevancy_propagator_imp() {
~relevancy_propagator_imp() override {
undo_trail(0);
}
@ -191,7 +191,7 @@ namespace smt {
m_trail.push_back(t);
}
virtual void add_handler(expr * source, relevancy_eh * eh) {
void add_handler(expr * source, relevancy_eh * eh) override {
if (!enabled())
return;
if (is_relevant_core(source)) {
@ -204,7 +204,7 @@ namespace smt {
}
}
virtual void add_watch(expr * n, bool val, relevancy_eh * eh) {
void add_watch(expr * n, bool val, relevancy_eh * eh) override {
if (!enabled())
return;
lbool lval = m_context.find_assignment(n);
@ -224,7 +224,7 @@ namespace smt {
}
}
virtual void add_watch(expr * n, bool val, expr * target) {
void add_watch(expr * n, bool val, expr * target) override {
if (!enabled())
return;
lbool lval = m_context.find_assignment(n);
@ -244,18 +244,18 @@ namespace smt {
bool is_relevant_core(expr * n) const { return m_is_relevant.contains(n); }
virtual bool is_relevant(expr * n) const {
bool is_relevant(expr * n) const override {
return !enabled() || is_relevant_core(n);
}
virtual void push() {
void push() override {
m_scopes.push_back(scope());
scope & s = m_scopes.back();
s.m_relevant_exprs_lim = m_relevant_exprs.size();
s.m_trail_lim = m_trail.size();
}
virtual void pop(unsigned num_scopes) {
void pop(unsigned num_scopes) override {
SASSERT(m_context.get_scope_level() == m_scopes.size());
unsigned lvl = m_scopes.size();
SASSERT(num_scopes <= lvl);
@ -325,7 +325,7 @@ namespace smt {
\brief Mark the given expression as relevant if it is not
already marked.
*/
virtual void mark_as_relevant(expr * n) {
void mark_as_relevant(expr * n) override {
if (!enabled())
return;
if (!is_relevant_core(n)) {
@ -450,7 +450,7 @@ namespace smt {
[m_qhead, m_relevant_exprs.size()) in the stack of
relevant expressions.
*/
virtual void propagate() {
void propagate() override {
if (m_propagating) {
return;
}
@ -494,11 +494,11 @@ namespace smt {
}
}
virtual bool can_propagate() const {
bool can_propagate() const override {
return m_qhead < m_relevant_exprs.size();
}
virtual void assign_eh(expr * n, bool val) {
void assign_eh(expr * n, bool val) override {
if (!enabled())
return;
ast_manager & m = get_manager();
@ -516,7 +516,7 @@ namespace smt {
}
}
virtual void display(std::ostream & out) const {
void display(std::ostream & out) const override {
if (enabled() && !m_relevant_exprs.empty()) {
out << "relevant exprs:\n";
for (unsigned i = 0; i < m_relevant_exprs.size(); i++) {

View file

@ -50,8 +50,8 @@ namespace smt {
expr * m_target;
public:
simple_relevancy_eh(expr * t):m_target(t) {}
virtual ~simple_relevancy_eh() {}
virtual void operator()(relevancy_propagator & rp);
~simple_relevancy_eh() override {}
void operator()(relevancy_propagator & rp) override;
};
/**
@ -63,8 +63,8 @@ namespace smt {
expr * m_target;
public:
pair_relevancy_eh(expr * s1, expr * s2, expr * t):m_source1(s1), m_source2(s2), m_target(t) {}
virtual ~pair_relevancy_eh() {}
virtual void operator()(relevancy_propagator & rp);
~pair_relevancy_eh() override {}
void operator()(relevancy_propagator & rp) override;
};
/**

View file

@ -55,7 +55,7 @@ namespace smt {
updt_params(p);
}
virtual solver * translate(ast_manager & m, params_ref const & p) {
solver * translate(ast_manager & m, params_ref const & p) override {
ast_translation translator(get_manager(), m);
smt_solver * result = alloc(smt_solver, m, p, m_logic);
@ -68,11 +68,11 @@ namespace smt {
return result;
}
virtual ~smt_solver() {
~smt_solver() override {
dec_ref_values(get_manager(), m_name2assertion);
}
virtual void updt_params(params_ref const & p) {
void updt_params(params_ref const & p) override {
solver::updt_params(p);
m_smt_params.updt_params(p);
m_context.updt_params(p);
@ -82,28 +82,28 @@ namespace smt {
m_core_extend_nonlocal_patterns = smth.core_extend_nonlocal_patterns();
}
virtual void collect_param_descrs(param_descrs & r) {
void collect_param_descrs(param_descrs & r) override {
m_context.collect_param_descrs(r);
}
virtual void collect_statistics(statistics & st) const {
void collect_statistics(statistics & st) const override {
m_context.collect_statistics(st);
}
virtual lbool get_consequences_core(expr_ref_vector const& assumptions, expr_ref_vector const& vars, expr_ref_vector& conseq) {
lbool get_consequences_core(expr_ref_vector const& assumptions, expr_ref_vector const& vars, expr_ref_vector& conseq) override {
expr_ref_vector unfixed(m_context.m());
return m_context.get_consequences(assumptions, vars, conseq, unfixed);
}
virtual lbool find_mutexes(expr_ref_vector const& vars, vector<expr_ref_vector>& mutexes) {
lbool find_mutexes(expr_ref_vector const& vars, vector<expr_ref_vector>& mutexes) override {
return m_context.find_mutexes(vars, mutexes);
}
virtual void assert_expr(expr * t) {
void assert_expr(expr * t) override {
m_context.assert_expr(t);
}
virtual void assert_expr(expr * t, expr * a) {
void assert_expr(expr * t, expr * a) override {
if (m_name2assertion.contains(a)) {
throw default_exception("named assertion defined twice");
}
@ -112,11 +112,11 @@ namespace smt {
m_name2assertion.insert(a, t);
}
virtual void push_core() {
void push_core() override {
m_context.push();
}
virtual void pop_core(unsigned n) {
void pop_core(unsigned n) override {
unsigned cur_sz = m_assumptions.size();
if (n > 0 && cur_sz > 0) {
unsigned lvl = m_scopes.size();
@ -135,7 +135,7 @@ namespace smt {
m_context.pop(n);
}
virtual lbool check_sat_core(unsigned num_assumptions, expr * const * assumptions) {
lbool check_sat_core(unsigned num_assumptions, expr * const * assumptions) override {
TRACE("solver_na2as", tout << "smt_solver::check_sat_core: " << num_assumptions << "\n";);
return m_context.check(num_assumptions, assumptions);
}
@ -154,7 +154,7 @@ namespace smt {
}
};
virtual void get_unsat_core(ptr_vector<expr> & r) {
void get_unsat_core(ptr_vector<expr> & r) override {
unsigned sz = m_context.get_unsat_core_size();
for (unsigned i = 0; i < sz; i++) {
r.push_back(m_context.get_unsat_core_expr(i));
@ -177,40 +177,40 @@ namespace smt {
add_nonlocal_pattern_literals_to_core(r);
}
virtual void get_model(model_ref & m) {
void get_model(model_ref & m) override {
m_context.get_model(m);
}
virtual proof * get_proof() {
proof * get_proof() override {
return m_context.get_proof();
}
virtual std::string reason_unknown() const {
std::string reason_unknown() const override {
return m_context.last_failure_as_string();
}
virtual void set_reason_unknown(char const* msg) {
void set_reason_unknown(char const* msg) override {
m_context.set_reason_unknown(msg);
}
virtual void get_labels(svector<symbol> & r) {
void get_labels(svector<symbol> & r) override {
buffer<symbol> tmp;
m_context.get_relevant_labels(0, tmp);
r.append(tmp.size(), tmp.c_ptr());
}
virtual ast_manager & get_manager() const { return m_context.m(); }
ast_manager & get_manager() const override { return m_context.m(); }
virtual void set_progress_callback(progress_callback * callback) {
void set_progress_callback(progress_callback * callback) override {
m_callback = callback;
m_context.set_progress_callback(callback);
}
virtual unsigned get_num_assertions() const {
unsigned get_num_assertions() const override {
return m_context.size();
}
virtual expr * get_assertion(unsigned idx) const {
expr * get_assertion(unsigned idx) const override {
SASSERT(idx < get_num_assertions());
return m_context.get_formula(idx);
}
@ -368,7 +368,7 @@ solver * mk_smt_solver(ast_manager & m, params_ref const & p, symbol const & log
class smt_solver_factory : public solver_factory {
public:
virtual solver * operator()(ast_manager & m, params_ref const & p, bool proofs_enabled, bool models_enabled, bool unsat_core_enabled, symbol const & logic) {
solver * operator()(ast_manager & m, params_ref const & p, bool proofs_enabled, bool models_enabled, bool unsat_core_enabled, symbol const & logic) override {
return mk_smt_solver(m, p, logic);
}
};

View file

@ -43,11 +43,11 @@ public:
m_fn = m.mk_func_decl(symbol(0xbeef101), i_sort, m.mk_bool_sort());
}
virtual tactic * translate(ast_manager & m) {
tactic * translate(ast_manager & m) override {
return alloc(ctx_solver_simplify_tactic, m, m_params);
}
virtual ~ctx_solver_simplify_tactic() {
~ctx_solver_simplify_tactic() override {
obj_map<sort, func_decl*>::iterator it = m_fns.begin(), end = m_fns.end();
for (; it != end; ++it) {
m.dec_ref(it->m_value);
@ -55,33 +55,32 @@ public:
m_fns.reset();
}
virtual void updt_params(params_ref const & p) {
void updt_params(params_ref const & p) override {
m_solver.updt_params(p);
}
virtual void collect_param_descrs(param_descrs & r) {
void collect_param_descrs(param_descrs & r) override {
m_solver.collect_param_descrs(r);
}
virtual void collect_statistics(statistics & st) const {
void collect_statistics(statistics & st) const override {
st.update("solver-simplify-steps", m_num_steps);
}
virtual void reset_statistics() { m_num_steps = 0; }
void reset_statistics() override { m_num_steps = 0; }
virtual void operator()(goal_ref const & in,
goal_ref_buffer & result,
model_converter_ref & mc,
proof_converter_ref & pc,
expr_dependency_ref & core) {
void operator()(goal_ref const & in,
goal_ref_buffer & result,
model_converter_ref & mc,
proof_converter_ref & pc,
expr_dependency_ref & core) override {
mc = 0; pc = 0; core = 0;
reduce(*(in.get()));
in->inc_depth();
result.push_back(in.get());
}
virtual void cleanup() {
void cleanup() override {
reset_statistics();
m_solver.reset();
}

View file

@ -53,11 +53,11 @@ public:
TRACE("smt_tactic", tout << "p: " << p << "\n";);
}
virtual tactic * translate(ast_manager & m) {
tactic * translate(ast_manager & m) override {
return alloc(smt_tactic, m_params_ref);
}
virtual ~smt_tactic() {
~smt_tactic() override {
SASSERT(m_ctx == 0);
}
@ -74,7 +74,7 @@ public:
m_fail_if_inconclusive = p.get_bool("fail_if_inconclusive", true);
}
virtual void updt_params(params_ref const & p) {
void updt_params(params_ref const & p) override {
TRACE("smt_tactic", tout << "updt_params: " << p << "\n";);
updt_params_core(p);
fparams().updt_params(p);
@ -86,7 +86,7 @@ public:
SASSERT(p.get_bool("auto_config", fparams().m_auto_config) == fparams().m_auto_config);
}
virtual void collect_param_descrs(param_descrs & r) {
void collect_param_descrs(param_descrs & r) override {
r.insert("candidate_models", CPK_BOOL, "(default: false) create candidate models even when quantifier or theory reasoning is incomplete.");
r.insert("fail_if_inconclusive", CPK_BOOL, "(default: true) fail if found unsat (sat) for under (over) approximated goal.");
smt_params_helper::collect_param_descrs(r);
@ -94,25 +94,25 @@ public:
}
virtual void collect_statistics(statistics & st) const {
void collect_statistics(statistics & st) const override {
if (m_ctx)
m_ctx->collect_statistics(st); // ctx is still running...
else
st.copy(m_stats);
}
virtual void cleanup() {
void cleanup() override {
}
virtual void reset_statistics() {
void reset_statistics() override {
m_stats.reset();
}
virtual void set_logic(symbol const & l) {
void set_logic(symbol const & l) override {
m_logic = l;
}
virtual void set_progress_callback(progress_callback * callback) {
void set_progress_callback(progress_callback * callback) override {
m_callback = callback;
}
@ -144,11 +144,11 @@ public:
};
virtual void operator()(goal_ref const & in,
goal_ref_buffer & result,
model_converter_ref & mc,
proof_converter_ref & pc,
expr_dependency_ref & core) {
void operator()(goal_ref const & in,
goal_ref_buffer & result,
model_converter_ref & mc,
proof_converter_ref & pc,
expr_dependency_ref & core) override {
try {
IF_VERBOSE(10, verbose_stream() << "(smt.tactic start)\n";);
mc = 0; pc = 0; core = 0;

View file

@ -37,22 +37,22 @@ struct unit_subsumption_tactic : public tactic {
m_clauses(m) {
}
void cleanup() {}
void cleanup() override {}
virtual void operator()(/* in */ goal_ref const & in,
/* out */ goal_ref_buffer & result,
/* out */ model_converter_ref & mc,
/* out */ proof_converter_ref & pc,
/* out */ expr_dependency_ref & core) {
void operator()(/* in */ goal_ref const & in,
/* out */ goal_ref_buffer & result,
/* out */ model_converter_ref & mc,
/* out */ proof_converter_ref & pc,
/* out */ expr_dependency_ref & core) override {
reduce_core(in, result);
}
virtual void updt_params(params_ref const& p) {
void updt_params(params_ref const& p) override {
m_params = p;
// m_context.updt_params(p); does not exist.
}
virtual tactic* translate(ast_manager& m) {
tactic* translate(ast_manager& m) override {
return alloc(unit_subsumption_tactic, m, m_params);
}

View file

@ -306,16 +306,16 @@ namespace smt {
public:
atom(bool_var bv, theory_var v, inf_numeral const & k, atom_kind kind);
atom_kind get_atom_kind() const { return static_cast<atom_kind>(m_atom_kind); }
virtual ~atom() {}
~atom() override {}
inline inf_numeral const & get_k() const { return m_k; }
bool_var get_bool_var() const { return m_bvar; }
bool is_true() const { return m_is_true; }
void assign_eh(bool is_true, inf_numeral const & epsilon);
virtual bool has_justification() const { return true; }
virtual void push_justification(antecedents& a, numeral const& coeff, bool proofs_enabled) {
bool has_justification() const override { return true; }
void push_justification(antecedents& a, numeral const& coeff, bool proofs_enabled) override {
a.push_lit(literal(get_bool_var(), !m_is_true), coeff, proofs_enabled);
}
virtual void display(theory_arith const& th, std::ostream& out) const;
void display(theory_arith const& th, std::ostream& out) const override;
};
class eq_bound : public bound {
@ -328,13 +328,13 @@ namespace smt {
m_rhs(rhs) {
SASSERT(m_lhs->get_root() == m_rhs->get_root());
}
virtual ~eq_bound() {}
virtual bool has_justification() const { return true; }
virtual void push_justification(antecedents& a, numeral const& coeff, bool proofs_enabled) {
~eq_bound() override {}
bool has_justification() const override { return true; }
void push_justification(antecedents& a, numeral const& coeff, bool proofs_enabled) override {
SASSERT(m_lhs->get_root() == m_rhs->get_root());
a.push_eq(enode_pair(m_lhs, m_rhs), coeff, proofs_enabled);
}
virtual void display(theory_arith const& th, std::ostream& out) const;
void display(theory_arith const& th, std::ostream& out) const override;
};
class derived_bound : public bound {
@ -344,14 +344,14 @@ namespace smt {
friend class theory_arith;
public:
derived_bound(theory_var v, inf_numeral const & val, bound_kind k):bound(v, val, k, false) {}
virtual ~derived_bound() {}
~derived_bound() override {}
literal_vector const& lits() const { return m_lits; }
eq_vector const& eqs() const { return m_eqs; }
virtual bool has_justification() const { return true; }
virtual void push_justification(antecedents& a, numeral const& coeff, bool proofs_enabled);
bool has_justification() const override { return true; }
void push_justification(antecedents& a, numeral const& coeff, bool proofs_enabled) override;
virtual void push_lit(literal l, numeral const&) { m_lits.push_back(l); }
virtual void push_eq(enode_pair const& p, numeral const&) { m_eqs.push_back(p); }
virtual void display(theory_arith const& th, std::ostream& out) const;
void display(theory_arith const& th, std::ostream& out) const override;
};
@ -361,12 +361,12 @@ namespace smt {
friend class theory_arith;
public:
justified_derived_bound(theory_var v, inf_numeral const & val, bound_kind k):derived_bound(v, val, k) {}
virtual ~justified_derived_bound() {}
virtual bool has_justification() const { return true; }
virtual void push_justification(antecedents& a, numeral const& coeff, bool proofs_enabled);
virtual void push_lit(literal l, numeral const& coeff);
~justified_derived_bound() override {}
bool has_justification() const override { return true; }
void push_justification(antecedents& a, numeral const& coeff, bool proofs_enabled) override;
void push_lit(literal l, numeral const& coeff) override;
virtual void push_eq(enode_pair const& p, numeral const& coeff);
void push_eq(enode_pair const& p, numeral const& coeff) override;
};
typedef int_hashtable<int_hash, default_eq<int> > literal_idx_set;
@ -510,7 +510,7 @@ namespace smt {
typedef int_hashtable<var_value_hash, var_value_eq> var_value_table;
var_value_table m_var_value_table;
virtual theory_var mk_var(enode * n);
theory_var mk_var(enode * n) override;
void found_unsupported_op(app * n);
void found_underspecified_op(app * n);
@ -635,24 +635,24 @@ namespace smt {
struct compare_atoms {
bool operator()(atom* a1, atom* a2) const { return a1->get_k() < a2->get_k(); }
};
virtual bool default_internalizer() const { return false; }
virtual bool internalize_atom(app * n, bool gate_ctx);
virtual bool internalize_term(app * term);
virtual void internalize_eq_eh(app * atom, bool_var v);
virtual void apply_sort_cnstr(enode * n, sort * s);
bool default_internalizer() const override { return false; }
bool internalize_atom(app * n, bool gate_ctx) override;
bool internalize_term(app * term) override;
void internalize_eq_eh(app * atom, bool_var v) override;
void apply_sort_cnstr(enode * n, sort * s) override;
virtual void assign_eh(bool_var v, bool is_true);
virtual void new_eq_eh(theory_var v1, theory_var v2);
virtual bool use_diseqs() const;
virtual void new_diseq_eh(theory_var v1, theory_var v2);
void assign_eh(bool_var v, bool is_true) override;
void new_eq_eh(theory_var v1, theory_var v2) override;
bool use_diseqs() const override;
void new_diseq_eh(theory_var v1, theory_var v2) override;
virtual void push_scope_eh();
virtual void pop_scope_eh(unsigned num_scopes);
void push_scope_eh() override;
void pop_scope_eh(unsigned num_scopes) override;
virtual void relevant_eh(app * n);
void relevant_eh(app * n) override;
virtual void restart_eh();
virtual void init_search_eh();
void restart_eh() override;
void init_search_eh() override;
/**
\brief True if the assignment may be changed during final
check. assume_eqs, check_int_feasibility,
@ -669,18 +669,18 @@ namespace smt {
*/
bool m_liberal_final_check;
final_check_status final_check_core();
virtual final_check_status final_check_eh();
final_check_status final_check_eh() override;
virtual bool can_propagate();
virtual void propagate();
bool can_propagate() override;
void propagate() override;
bool propagate_core();
void failed();
virtual void flush_eh();
virtual void reset_eh();
void flush_eh() override;
void reset_eh() override;
virtual bool validate_eq_in_model(theory_var v1, theory_var v2, bool is_true) const;
bool validate_eq_in_model(theory_var v1, theory_var v2, bool is_true) const override;
// -----------------------------------
//
@ -867,7 +867,7 @@ namespace smt {
void propagate_cheap_eq(unsigned rid);
void propagate_eq_to_core(theory_var x, theory_var y, antecedents& antecedents);
virtual bool is_shared(theory_var v) const;
bool is_shared(theory_var v) const override;
// -----------------------------------
//
@ -909,7 +909,7 @@ namespace smt {
/**
\brief See comment in theory::mk_eq_atom
*/
virtual app * mk_eq_atom(expr * lhs, expr * rhs) { return m_util.mk_eq(lhs, rhs); }
app * mk_eq_atom(expr * lhs, expr * rhs) override { return m_util.mk_eq(lhs, rhs); }
// -----------------------------------
//
@ -1039,13 +1039,13 @@ namespace smt {
// -----------------------------------
public:
theory_arith(ast_manager & m, theory_arith_params & params);
virtual ~theory_arith();
~theory_arith() override;
virtual theory * mk_fresh(context * new_ctx);
theory * mk_fresh(context * new_ctx) override;
virtual void setup();
void setup() override;
virtual char const * get_name() const { return "arithmetic"; }
char const * get_name() const override { return "arithmetic"; }
// -----------------------------------
//
@ -1058,15 +1058,15 @@ namespace smt {
void compute_epsilon();
void refine_epsilon();
virtual void init_model(model_generator & m);
virtual model_value_proc * mk_value(enode * n, model_generator & mg);
void init_model(model_generator & m) override;
model_value_proc * mk_value(enode * n, model_generator & mg) override;
// -----------------------------------
//
// Model checker
//
// -----------------------------------
virtual bool get_value(enode * n, expr_ref & r);
bool get_value(enode * n, expr_ref & r) override;
bool get_lower(enode* n, expr_ref& r);
bool get_upper(enode* n, expr_ref& r);
@ -1078,9 +1078,9 @@ namespace smt {
// Optimization
//
// -----------------------------------
virtual inf_eps_rational<inf_rational> maximize(theory_var v, expr_ref& blocker, bool& has_shared);
virtual inf_eps_rational<inf_rational> value(theory_var v);
virtual theory_var add_objective(app* term);
inf_eps_rational<inf_rational> maximize(theory_var v, expr_ref& blocker, bool& has_shared) override;
inf_eps_rational<inf_rational> value(theory_var v) override;
theory_var add_objective(app* term) override;
expr_ref mk_ge(filter_model_converter& fm, theory_var v, inf_numeral const& val);
void enable_record_conflict(expr* bound);
void record_conflict(unsigned num_lits, literal const * lits,
@ -1102,8 +1102,8 @@ namespace smt {
//
// -----------------------------------
public:
virtual void collect_statistics(::statistics & st) const;
virtual void display(std::ostream & out) const;
void collect_statistics(::statistics & st) const override;
void display(std::ostream & out) const override;
protected:
void display_row(std::ostream & out, unsigned r_id, bool compact = true) const;
void display_row(std::ostream & out, row const & r, bool compact = true) const;

View file

@ -472,7 +472,7 @@ namespace smt {
bounds.num_params(), bounds.params("gomory-cut")) {
}
// Remark: the assignment must be propagated back to arith
virtual theory_id get_from_theory() const { return null_theory_id; }
theory_id get_from_theory() const override { return null_theory_id; }
};
/**

View file

@ -55,19 +55,19 @@ namespace smt {
th_trail_stack m_trail_stack;
unsigned m_final_check_idx;
virtual void init(context * ctx);
virtual theory_var mk_var(enode * n);
virtual bool internalize_atom(app * atom, bool gate_ctx);
virtual bool internalize_term(app * term);
virtual void apply_sort_cnstr(enode * n, sort * s);
virtual void new_eq_eh(theory_var v1, theory_var v2);
virtual void new_diseq_eh(theory_var v1, theory_var v2);
virtual void relevant_eh(app * n);
virtual void push_scope_eh();
virtual void pop_scope_eh(unsigned num_scopes);
virtual final_check_status final_check_eh();
virtual void reset_eh();
virtual void init_search_eh() { m_final_check_idx = 0; }
void init(context * ctx) override;
theory_var mk_var(enode * n) override;
bool internalize_atom(app * atom, bool gate_ctx) override;
bool internalize_term(app * term) override;
void apply_sort_cnstr(enode * n, sort * s) override;
void new_eq_eh(theory_var v1, theory_var v2) override;
void new_diseq_eh(theory_var v1, theory_var v2) override;
void relevant_eh(app * n) override;
void push_scope_eh() override;
void pop_scope_eh(unsigned num_scopes) override;
final_check_status final_check_eh() override;
void reset_eh() override;
void init_search_eh() override { m_final_check_idx = 0; }
virtual void set_prop_upward(theory_var v);
virtual void set_prop_upward(enode* n);
@ -96,15 +96,15 @@ namespace smt {
static void display_ids(std::ostream & out, unsigned n, enode * const * v);
public:
theory_array(ast_manager & m, theory_array_params & params);
virtual ~theory_array();
~theory_array() override;
virtual theory * mk_fresh(context * new_ctx) { return alloc(theory_array, new_ctx->get_manager(), new_ctx->get_fparams()); }
theory * mk_fresh(context * new_ctx) override { return alloc(theory_array, new_ctx->get_manager(), new_ctx->get_fparams()); }
virtual char const * get_name() const { return "array"; }
char const * get_name() const override { return "array"; }
virtual void display_var(std::ostream & out, theory_var v) const;
virtual void display(std::ostream & out) const;
virtual void collect_statistics(::statistics & st) const;
void display(std::ostream & out) const override;
void collect_statistics(::statistics & st) const override;
th_trail_stack & get_trail_stack() { return m_trail_stack; }
virtual void merge_eh(theory_var v1, theory_var v2, theory_var, theory_var);
static void after_merge_eh(theory_var r1, theory_var r2, theory_var v1, theory_var v2) {}

View file

@ -828,7 +828,7 @@ namespace smt {
m_unspecified_else(true) {
}
virtual ~array_value_proc() {}
~array_value_proc() override {}
void add_entry(unsigned num_args, enode * const * args, enode * value) {
SASSERT(num_args > 0);
@ -840,11 +840,11 @@ namespace smt {
m_dependencies.push_back(model_value_dependency(value));
}
virtual void get_dependencies(buffer<model_value_dependency> & result) {
void get_dependencies(buffer<model_value_dependency> & result) override {
result.append(m_dependencies.size(), m_dependencies.c_ptr());
}
virtual app * mk_value(model_generator & mg, ptr_vector<expr> & values) {
app * mk_value(model_generator & mg, ptr_vector<expr> & values) override {
// values must have size = m_num_entries * (m_dim + 1) + ((m_else || m_unspecified_else) ? 0 : 1)
// an array value is a lookup table + else_value
// each entry has m_dim indexes that map to a value.

View file

@ -123,15 +123,15 @@ namespace smt {
//
//
// --------------------------------------------------
virtual bool is_shared(theory_var v) const;
bool is_shared(theory_var v) const override;
void collect_shared_vars(sbuffer<theory_var> & result);
unsigned mk_interface_eqs();
virtual bool can_propagate();
virtual void propagate();
virtual void push_scope_eh();
virtual void pop_scope_eh(unsigned num_scopes);
virtual void reset_eh();
bool can_propagate() override;
void propagate() override;
void push_scope_eh() override;
void pop_scope_eh(unsigned num_scopes) override;
void reset_eh() override;
void reset_queues();
// -----------------------------------
@ -177,7 +177,7 @@ namespace smt {
void set_default(theory_var v, enode* n);
enode* get_default(theory_var v);
virtual void init_model(model_generator & m);
void init_model(model_generator & m) override;
bool is_unspecified_default_ok() const;
void collect_defaults();
void collect_selects();
@ -185,12 +185,12 @@ namespace smt {
void propagate_selects_to_store_parents(enode * r, enode_pair_vector & todo);
void propagate_selects();
select_set * get_select_set(enode * n);
virtual void finalize_model(model_generator & m);
virtual model_value_proc * mk_value(enode * n, model_generator & m);
void finalize_model(model_generator & m) override;
model_value_proc * mk_value(enode * n, model_generator & m) override;
public:
theory_array_base(ast_manager & m);
virtual ~theory_array_base() { restore_sorts(0); }
~theory_array_base() override { restore_sorts(0); }
};
};

View file

@ -42,29 +42,29 @@ namespace smt {
protected:
//virtual final_check_status final_check_eh();
virtual void reset_eh();
void reset_eh() override;
virtual void set_prop_upward(theory_var v);
virtual void set_prop_upward(enode* n);
virtual void set_prop_upward(theory_var v, var_data* d);
virtual unsigned get_lambda_equiv_size(theory_var v, var_data* d);
void set_prop_upward(theory_var v) override;
void set_prop_upward(enode* n) override;
void set_prop_upward(theory_var v, var_data* d) override;
unsigned get_lambda_equiv_size(theory_var v, var_data* d) override;
virtual bool internalize_term(app * term);
virtual bool internalize_atom(app * atom, bool gate_ctx);
virtual void pop_scope_eh(unsigned num_scopes);
virtual theory_var mk_var(enode * n);
virtual void relevant_eh(app * n);
bool internalize_term(app * term) override;
bool internalize_atom(app * atom, bool gate_ctx) override;
void pop_scope_eh(unsigned num_scopes) override;
theory_var mk_var(enode * n) override;
void relevant_eh(app * n) override;
void add_const(theory_var v, enode* c);
void add_map(theory_var v, enode* s);
void add_parent_map(theory_var v, enode* s);
void add_as_array(theory_var v, enode* arr);
virtual void add_parent_select(theory_var v, enode * s);
void add_parent_select(theory_var v, enode * s) override;
void add_parent_default(theory_var v);
virtual final_check_status assert_delayed_axioms();
final_check_status assert_delayed_axioms() override;
bool instantiate_default_const_axiom(enode* cnst);
bool instantiate_default_store_axiom(enode* store);
@ -87,14 +87,14 @@ namespace smt {
public:
theory_array_full(ast_manager & m, theory_array_params & params);
virtual ~theory_array_full();
~theory_array_full() override;
virtual theory * mk_fresh(context * new_ctx);
theory * mk_fresh(context * new_ctx) override;
virtual void merge_eh(theory_var v1, theory_var v2, theory_var, theory_var);
virtual void display_var(std::ostream & out, theory_var v) const;
virtual void collect_statistics(::statistics & st) const;
virtual void init(context* ctx) {
void merge_eh(theory_var v1, theory_var v2, theory_var, theory_var) override;
void display_var(std::ostream & out, theory_var v) const override;
void collect_statistics(::statistics & st) const override;
void init(context* ctx) override {
// the parent class is theory_array.
// theory::init(ctx);
theory_array::init(ctx);

View file

@ -65,7 +65,7 @@ namespace smt {
bool_var m_var;
public:
mk_atom_trail(bool_var v):m_var(v) {}
virtual void undo(theory_bv & th) {
void undo(theory_bv & th) override {
theory_bv::atom * a = th.get_bv2a(m_var);
a->~atom();
th.erase_bv2a(m_var);
@ -213,7 +213,7 @@ namespace smt {
theory_bv::bit_atom * m_atom;
public:
add_var_pos_trail(theory_bv::bit_atom * a):m_atom(a) {}
virtual void undo(theory_bv & th) {
void undo(theory_bv & th) override {
SASSERT(m_atom->m_occs);
m_atom->m_occs = m_atom->m_occs->m_next;
}
@ -389,12 +389,12 @@ namespace smt {
m_th(th), m_var1(v1), m_var2(v2) {
}
virtual void get_antecedents(conflict_resolution & cr) {
void get_antecedents(conflict_resolution & cr) override {
mark_bits(cr, m_th.m_bits[m_var1]);
mark_bits(cr, m_th.m_bits[m_var2]);
}
virtual proof * mk_proof(conflict_resolution & cr) {
proof * mk_proof(conflict_resolution & cr) override {
ptr_buffer<proof> prs;
context & ctx = cr.get_context();
bool visited = true;
@ -414,11 +414,11 @@ namespace smt {
return m.mk_th_lemma(get_from_theory(), fact, prs.size(), prs.c_ptr());
}
virtual theory_id get_from_theory() const {
theory_id get_from_theory() const override {
return m_th.get_id();
}
virtual char const * get_name() const { return "bv-fixed-eq"; }
char const * get_name() const override { return "bv-fixed-eq"; }
};
@ -1510,13 +1510,13 @@ namespace smt {
bit_eq_justification(theory_id th_id, enode * v1, enode * v2, literal c, literal a):
m_v1(v1), m_v2(v2), m_th_id(th_id), m_consequent(c), m_antecedent(a) {}
virtual void get_antecedents(conflict_resolution & cr) {
void get_antecedents(conflict_resolution & cr) override {
cr.mark_eq(m_v1, m_v2);
if (m_antecedent.var() != true_bool_var)
cr.mark_literal(m_antecedent);
}
virtual proof * mk_proof(conflict_resolution & cr) {
proof * mk_proof(conflict_resolution & cr) override {
bool visited = true;
ptr_buffer<proof> prs;
proof * pr = cr.get_proof(m_v1, m_v2);
@ -1540,11 +1540,11 @@ namespace smt {
return m.mk_th_lemma(get_from_theory(), fact, prs.size(), prs.c_ptr());
}
virtual theory_id get_from_theory() const {
theory_id get_from_theory() const override {
return m_th_id;
}
virtual char const * get_name() const { return "bv-bit-eq"; }
char const * get_name() const override { return "bv-bit-eq"; }
};
inline justification * theory_bv::mk_bit_eq_justification(theory_var v1, theory_var v2, literal consequent, literal antecedent) {

View file

@ -58,16 +58,16 @@ namespace smt {
struct bit_atom : public atom {
var_pos_occ * m_occs;
bit_atom():m_occs(0) {}
virtual ~bit_atom() {}
virtual bool is_bit() const { return true; }
~bit_atom() override {}
bool is_bit() const override { return true; }
};
struct le_atom : public atom {
literal m_var;
literal m_def;
le_atom(literal v, literal d):m_var(v), m_def(d) {}
virtual ~le_atom() {}
virtual bool is_bit() const { return false; }
~le_atom() override {}
bool is_bit() const override { return false; }
};
/**
@ -216,21 +216,21 @@ namespace smt {
void assert_bv2int_axiom(app* n);
protected:
virtual void init(context * ctx);
virtual theory_var mk_var(enode * n);
virtual bool internalize_atom(app * atom, bool gate_ctx);
virtual bool internalize_term(app * term);
virtual void apply_sort_cnstr(enode * n, sort * s);
virtual void new_eq_eh(theory_var v1, theory_var v2);
virtual void new_diseq_eh(theory_var v1, theory_var v2);
void init(context * ctx) override;
theory_var mk_var(enode * n) override;
bool internalize_atom(app * atom, bool gate_ctx) override;
bool internalize_term(app * term) override;
void apply_sort_cnstr(enode * n, sort * s) override;
void new_eq_eh(theory_var v1, theory_var v2) override;
void new_diseq_eh(theory_var v1, theory_var v2) override;
virtual void expand_diseq(theory_var v1, theory_var v2);
virtual void assign_eh(bool_var v, bool is_true);
virtual void relevant_eh(app * n);
virtual void push_scope_eh();
virtual void pop_scope_eh(unsigned num_scopes);
virtual final_check_status final_check_eh();
virtual void reset_eh();
virtual bool include_func_interp(func_decl* f);
void assign_eh(bool_var v, bool is_true) override;
void relevant_eh(app * n) override;
void push_scope_eh() override;
void pop_scope_eh(unsigned num_scopes) override;
final_check_status final_check_eh() override;
void reset_eh() override;
bool include_func_interp(func_decl* f) override;
svector<theory_var> m_merge_aux[2]; //!< auxiliary vector used in merge_zero_one_bits
bool merge_zero_one_bits(theory_var r1, theory_var r2);
@ -240,16 +240,16 @@ namespace smt {
//
// -----------------------------------
bv_factory * m_factory;
virtual void init_model(model_generator & m);
virtual model_value_proc * mk_value(enode * n, model_generator & mg);
void init_model(model_generator & m) override;
model_value_proc * mk_value(enode * n, model_generator & mg) override;
public:
theory_bv(ast_manager & m, theory_bv_params const & params, bit_blaster_params const & bb_params);
virtual ~theory_bv();
~theory_bv() override;
virtual theory * mk_fresh(context * new_ctx);
theory * mk_fresh(context * new_ctx) override;
virtual char const * get_name() const { return "bit-vector"; }
char const * get_name() const override { return "bit-vector"; }
th_trail_stack & get_trail_stack() { return m_trail_stack; }
void merge_eh(theory_var, theory_var, theory_var v1, theory_var v2);
@ -259,8 +259,8 @@ namespace smt {
void display_var(std::ostream & out, theory_var v) const;
void display_bit_atom(std::ostream & out, bool_var v, bit_atom const * a) const;
void display_atoms(std::ostream & out) const;
virtual void display(std::ostream & out) const;
virtual void collect_statistics(::statistics & st) const;
void display(std::ostream & out) const override;
void collect_statistics(::statistics & st) const override;
bool get_fixed_value(app* x, numeral & result) const;

View file

@ -33,7 +33,7 @@ namespace smt {
ext_theory_eq_propagation_justification(fid, r, 1, &antecedent, 0, 0, lhs, rhs) {
}
// Remark: the assignment must be propagated back to the datatype theory.
virtual theory_id get_from_theory() const { return null_theory_id; }
theory_id get_from_theory() const override { return null_theory_id; }
};
@ -551,11 +551,11 @@ namespace smt {
public:
datatype_value_proc(func_decl * d):m_constructor(d) {}
void add_dependency(enode * n) { m_dependencies.push_back(model_value_dependency(n)); }
virtual ~datatype_value_proc() {}
virtual void get_dependencies(buffer<model_value_dependency> & result) {
~datatype_value_proc() override {}
void get_dependencies(buffer<model_value_dependency> & result) override {
result.append(m_dependencies.size(), m_dependencies.c_ptr());
}
virtual app * mk_value(model_generator & mg, ptr_vector<expr> & values) {
app * mk_value(model_generator & mg, ptr_vector<expr> & values) override {
SASSERT(values.size() == m_dependencies.size());
return mg.get_manager().mk_app(m_constructor, values.size(), values.c_ptr());
}

View file

@ -84,35 +84,35 @@ namespace smt {
void display_var(std::ostream & out, theory_var v) const;
protected:
virtual theory_var mk_var(enode * n);
virtual bool internalize_atom(app * atom, bool gate_ctx);
virtual bool internalize_term(app * term);
virtual void apply_sort_cnstr(enode * n, sort * s);
virtual void new_eq_eh(theory_var v1, theory_var v2);
virtual bool use_diseqs() const;
virtual void new_diseq_eh(theory_var v1, theory_var v2);
virtual void assign_eh(bool_var v, bool is_true);
virtual void relevant_eh(app * n);
virtual void push_scope_eh();
virtual void pop_scope_eh(unsigned num_scopes);
virtual final_check_status final_check_eh();
virtual void reset_eh();
virtual void restart_eh() { m_util.reset(); }
virtual bool is_shared(theory_var v) const;
theory_var mk_var(enode * n) override;
bool internalize_atom(app * atom, bool gate_ctx) override;
bool internalize_term(app * term) override;
void apply_sort_cnstr(enode * n, sort * s) override;
void new_eq_eh(theory_var v1, theory_var v2) override;
bool use_diseqs() const override;
void new_diseq_eh(theory_var v1, theory_var v2) override;
void assign_eh(bool_var v, bool is_true) override;
void relevant_eh(app * n) override;
void push_scope_eh() override;
void pop_scope_eh(unsigned num_scopes) override;
final_check_status final_check_eh() override;
void reset_eh() override;
void restart_eh() override { m_util.reset(); }
bool is_shared(theory_var v) const override;
public:
theory_datatype(ast_manager & m, theory_datatype_params & p);
virtual ~theory_datatype();
virtual theory * mk_fresh(context * new_ctx);
virtual void display(std::ostream & out) const;
virtual void collect_statistics(::statistics & st) const;
virtual void init_model(model_generator & m);
virtual model_value_proc * mk_value(enode * n, model_generator & m);
~theory_datatype() override;
theory * mk_fresh(context * new_ctx) override;
void display(std::ostream & out) const override;
void collect_statistics(::statistics & st) const override;
void init_model(model_generator & m) override;
model_value_proc * mk_value(enode * n, model_generator & m) override;
th_trail_stack & get_trail_stack() { return m_trail_stack; }
virtual void merge_eh(theory_var v1, theory_var v2, theory_var, theory_var);
static void after_merge_eh(theory_var r1, theory_var r2, theory_var v1, theory_var v2) {}
void unmerge_eh(theory_var v1, theory_var v2);
virtual char const * get_name() const { return "datatype"; }
virtual bool include_func_interp(func_decl* f);
char const * get_name() const override { return "datatype"; }
bool include_func_interp(func_decl* f) override;
};

View file

@ -182,7 +182,7 @@ namespace smt {
return false;
}
app * mk_zero_for(expr * n);
theory_var mk_var(enode * n);
theory_var mk_var(enode * n) override;
theory_var internalize_term_core(app * n);
void found_non_diff_logic_expr(expr * n);
bool is_connected(theory_var source, theory_var target) const { return m_matrix[source][target].m_edge_id != null_edge_id; }
@ -214,38 +214,38 @@ namespace smt {
// Internalization
//
// -----------------------------------
virtual bool internalize_atom(app * n, bool gate_ctx);
virtual bool internalize_term(app * term);
virtual void internalize_eq_eh(app * atom, bool_var v);
virtual void apply_sort_cnstr(enode * n, sort * s);
bool internalize_atom(app * n, bool gate_ctx) override;
bool internalize_term(app * term) override;
void internalize_eq_eh(app * atom, bool_var v) override;
void apply_sort_cnstr(enode * n, sort * s) override;
virtual void assign_eh(bool_var v, bool is_true);
virtual void new_eq_eh(theory_var v1, theory_var v2);
virtual bool use_diseqs() const;
virtual void new_diseq_eh(theory_var v1, theory_var v2);
void assign_eh(bool_var v, bool is_true) override;
void new_eq_eh(theory_var v1, theory_var v2) override;
bool use_diseqs() const override;
void new_diseq_eh(theory_var v1, theory_var v2) override;
virtual void conflict_resolution_eh(app * atom, bool_var v);
void conflict_resolution_eh(app * atom, bool_var v) override;
virtual void push_scope_eh();
virtual void pop_scope_eh(unsigned num_scopes);
void push_scope_eh() override;
void pop_scope_eh(unsigned num_scopes) override;
virtual void restart_eh();
virtual void init_search_eh();
virtual final_check_status final_check_eh();
void restart_eh() override;
void init_search_eh() override;
final_check_status final_check_eh() override;
virtual bool can_propagate();
virtual void propagate();
bool can_propagate() override;
void propagate() override;
virtual void flush_eh();
virtual void reset_eh();
void flush_eh() override;
void reset_eh() override;
bool dump_lemmas() const { return m_params.m_arith_dump_lemmas; }
virtual bool validate_eq_in_model(theory_var v1, theory_var v2, bool is_true) const;
bool validate_eq_in_model(theory_var v1, theory_var v2, bool is_true) const override;
virtual void display(std::ostream & out) const;
void display(std::ostream & out) const override;
virtual void display_atom(std::ostream & out, atom * a) const;
virtual void collect_statistics(::statistics & st) const;
void collect_statistics(::statistics & st) const override;
// -----------------------------------
//
@ -258,8 +258,8 @@ namespace smt {
void compute_epsilon();
void fix_zero();
virtual void init_model(model_generator & m);
virtual model_value_proc * mk_value(enode * n, model_generator & mg);
void init_model(model_generator & m) override;
model_value_proc * mk_value(enode * n, model_generator & mg) override;
// -----------------------------------
//
@ -267,9 +267,9 @@ namespace smt {
//
// -----------------------------------
virtual inf_eps_rational<inf_rational> maximize(theory_var v, expr_ref& blocker, bool& has_shared);
virtual inf_eps_rational<inf_rational> value(theory_var v);
virtual theory_var add_objective(app* term);
inf_eps_rational<inf_rational> maximize(theory_var v, expr_ref& blocker, bool& has_shared) override;
inf_eps_rational<inf_rational> value(theory_var v) override;
theory_var add_objective(app* term) override;
virtual expr_ref mk_gt(theory_var v, inf_eps const& val);
expr_ref mk_ge(filter_model_converter& fm, theory_var v, inf_eps const& val);
@ -280,16 +280,16 @@ namespace smt {
// -----------------------------------
public:
theory_dense_diff_logic(ast_manager & m, theory_arith_params & p);
virtual ~theory_dense_diff_logic() { reset_eh(); }
~theory_dense_diff_logic() override { reset_eh(); }
virtual theory * mk_fresh(context * new_ctx);
theory * mk_fresh(context * new_ctx) override;
virtual char const * get_name() const { return "difference-logic"; }
char const * get_name() const override { return "difference-logic"; }
/**
\brief See comment in theory::mk_eq_atom
*/
virtual app * mk_eq_atom(expr * lhs, expr * rhs) { return m_autil.mk_eq(lhs, rhs); }
app * mk_eq_atom(expr * lhs, expr * rhs) override { return m_autil.mk_eq(lhs, rhs); }
};
typedef theory_dense_diff_logic<mi_ext> theory_dense_mi;

View file

@ -208,7 +208,7 @@ namespace smt {
void new_edge(dl_var src, dl_var dst, unsigned num_edges, edge_id const* edges);
// Create a new theory variable.
virtual theory_var mk_var(enode* n);
theory_var mk_var(enode* n) override;
virtual theory_var mk_var(app* n);
@ -240,79 +240,79 @@ namespace smt {
m_num_simplex_edges(0) {
}
virtual ~theory_diff_logic() {
~theory_diff_logic() override {
reset_eh();
}
virtual theory * mk_fresh(context * new_ctx);
theory * mk_fresh(context * new_ctx) override;
virtual char const * get_name() const { return "difference-logic"; }
char const * get_name() const override { return "difference-logic"; }
/**
\brief See comment in theory::mk_eq_atom
*/
virtual app * mk_eq_atom(expr * lhs, expr * rhs) { return m_util.mk_eq(lhs, rhs); }
app * mk_eq_atom(expr * lhs, expr * rhs) override { return m_util.mk_eq(lhs, rhs); }
virtual void init(context * ctx);
void init(context * ctx) override;
virtual bool internalize_atom(app * atom, bool gate_ctx);
bool internalize_atom(app * atom, bool gate_ctx) override;
virtual bool internalize_term(app * term);
bool internalize_term(app * term) override;
virtual void internalize_eq_eh(app * atom, bool_var v);
void internalize_eq_eh(app * atom, bool_var v) override;
virtual void assign_eh(bool_var v, bool is_true);
void assign_eh(bool_var v, bool is_true) override;
virtual void new_eq_eh(theory_var v1, theory_var v2);
void new_eq_eh(theory_var v1, theory_var v2) override;
virtual bool use_diseqs() const { return true; }
bool use_diseqs() const override { return true; }
virtual void new_diseq_eh(theory_var v1, theory_var v2);
void new_diseq_eh(theory_var v1, theory_var v2) override;
virtual void push_scope_eh();
void push_scope_eh() override;
virtual void pop_scope_eh(unsigned num_scopes);
void pop_scope_eh(unsigned num_scopes) override;
virtual void restart_eh() {
void restart_eh() override {
m_arith_eq_adapter.restart_eh();
}
virtual void relevant_eh(app* e) {}
void relevant_eh(app* e) override {}
virtual void init_search_eh() {
void init_search_eh() override {
m_arith_eq_adapter.init_search_eh();
}
virtual final_check_status final_check_eh();
final_check_status final_check_eh() override;
virtual bool is_shared(theory_var v) const {
bool is_shared(theory_var v) const override {
return false;
}
virtual bool can_propagate() {
bool can_propagate() override {
return m_asserted_qhead != m_asserted_atoms.size();
}
virtual void propagate();
void propagate() override;
virtual justification * why_is_diseq(theory_var v1, theory_var v2) {
justification * why_is_diseq(theory_var v1, theory_var v2) override {
NOT_IMPLEMENTED_YET();
return 0;
}
// virtual void flush_eh();
virtual void reset_eh();
void reset_eh() override;
virtual void init_model(model_generator & m);
void init_model(model_generator & m) override;
virtual model_value_proc * mk_value(enode * n, model_generator & mg);
model_value_proc * mk_value(enode * n, model_generator & mg) override;
virtual bool validate_eq_in_model(theory_var v1, theory_var v2, bool is_true) const;
bool validate_eq_in_model(theory_var v1, theory_var v2, bool is_true) const override;
virtual void display(std::ostream & out) const;
void display(std::ostream & out) const override;
virtual void collect_statistics(::statistics & st) const;
void collect_statistics(::statistics & st) const override;
// -----------------------------------
@ -321,9 +321,9 @@ namespace smt {
//
// -----------------------------------
virtual inf_eps maximize(theory_var v, expr_ref& blocker, bool& has_shared);
virtual inf_eps value(theory_var v);
virtual theory_var add_objective(app* term);
inf_eps maximize(theory_var v, expr_ref& blocker, bool& has_shared) override;
inf_eps value(theory_var v) override;
theory_var add_objective(app* term) override;
expr_ref mk_ge(filter_model_converter& fm, theory_var v, inf_eps const& val);
bool internalize_objective(expr * n, rational const& m, rational& r, objective_term & objective);

View file

@ -51,7 +51,7 @@ namespace smt {
m_util(u)
{}
virtual app * mk_value_core(unsigned const & val, sort * s) {
app * mk_value_core(unsigned const & val, sort * s) override {
return m_util.mk_numeral(val, s);
}
};
@ -74,9 +74,9 @@ namespace smt {
dl_value_proc(theory_dl& th, smt::enode* n) : m_th(th), m_node(n) {}
virtual void get_dependencies(buffer<smt::model_value_dependency> & result) {}
void get_dependencies(buffer<smt::model_value_dependency> & result) override {}
virtual app * mk_value(smt::model_generator & mg, ptr_vector<expr> & ) {
app * mk_value(smt::model_generator & mg, ptr_vector<expr> & ) override {
smt::context& ctx = m_th.get_context();
app* result = 0;
expr* n = m_node->get_owner();
@ -111,9 +111,9 @@ namespace smt {
}
virtual char const * get_name() const { return "datalog"; }
char const * get_name() const override { return "datalog"; }
virtual bool internalize_atom(app * atom, bool gate_ctx) {
bool internalize_atom(app * atom, bool gate_ctx) override {
TRACE("theory_dl", tout << mk_pp(atom, m()) << "\n";);
context& ctx = get_context();
if (ctx.b_internalized(atom)) {
@ -136,7 +136,7 @@ namespace smt {
return false;
}
virtual bool internalize_term(app * term) {
bool internalize_term(app * term) override {
TRACE("theory_dl", tout << mk_pp(term, m()) << "\n";);
if (u().is_finite_sort(term)) {
return mk_rep(term);
@ -146,27 +146,27 @@ namespace smt {
}
}
virtual void new_eq_eh(theory_var v1, theory_var v2) {
void new_eq_eh(theory_var v1, theory_var v2) override {
}
virtual void new_diseq_eh(theory_var v1, theory_var v2) {
void new_diseq_eh(theory_var v1, theory_var v2) override {
}
virtual theory * mk_fresh(context * new_ctx) {
theory * mk_fresh(context * new_ctx) override {
return alloc(theory_dl, new_ctx->get_manager());
}
virtual void init_model(smt::model_generator & m) {
void init_model(smt::model_generator & m) override {
m.register_factory(alloc(dl_factory, m_util, m.get_model()));
}
virtual smt::model_value_proc * mk_value(smt::enode * n, smt::model_generator&) {
smt::model_value_proc * mk_value(smt::enode * n, smt::model_generator&) override {
return alloc(dl_value_proc, *this, n);
}
virtual void apply_sort_cnstr(enode * n, sort * s) {
void apply_sort_cnstr(enode * n, sort * s) override {
app* term = n->get_owner();
if (u().is_finite_sort(term)) {
mk_rep(term);
@ -174,7 +174,7 @@ namespace smt {
}
virtual void relevant_eh(app * n) {
void relevant_eh(app * n) override {
if (u().is_finite_sort(n)) {
sort* s = m().get_sort(n);
func_decl* r, *v;
@ -194,7 +194,7 @@ namespace smt {
}
}
virtual void display(std::ostream & out) const {
void display(std::ostream & out) const override {
}

View file

@ -33,25 +33,25 @@ namespace smt {
void found_theory_expr();
protected:
virtual bool internalize_atom(app * atom, bool gate_ctx);
virtual bool internalize_term(app * term);
virtual void new_eq_eh(theory_var v1, theory_var v2);
virtual bool use_diseqs() const;
virtual void new_diseq_eh(theory_var v1, theory_var v2);
virtual void reset_eh();
virtual final_check_status final_check_eh();
virtual bool build_models() const {
bool internalize_atom(app * atom, bool gate_ctx) override;
bool internalize_term(app * term) override;
void new_eq_eh(theory_var v1, theory_var v2) override;
bool use_diseqs() const override;
void new_diseq_eh(theory_var v1, theory_var v2) override;
void reset_eh() override;
final_check_status final_check_eh() override;
bool build_models() const override {
return false;
}
virtual void display(std::ostream& out) const {}
void display(std::ostream& out) const override {}
public:
theory_dummy(family_id fid, char const * name);
virtual ~theory_dummy() {}
~theory_dummy() override {}
virtual theory * mk_fresh(context * new_ctx) { return alloc(theory_dummy, get_family_id(), m_name); }
theory * mk_fresh(context * new_ctx) override { return alloc(theory_dummy, get_family_id(), m_name); }
virtual char const * get_name() const;
char const * get_name() const override;
};
};

View file

@ -32,8 +32,8 @@ namespace smt {
public:
fpa2bv_conversion_trail_elem(ast_manager & m, obj_map<expr, expr*> & map, expr * e) :
m(m), m_map(map), key(e, m) { }
virtual ~fpa2bv_conversion_trail_elem() { }
virtual void undo(theory_fpa & th) {
~fpa2bv_conversion_trail_elem() override { }
void undo(theory_fpa & th) override {
expr * val = m_map.find(key);
m_map.remove(key);
m.dec_ref(key);

View file

@ -43,16 +43,16 @@ namespace smt {
value_factory(m, fid),
m_util(m) {}
virtual ~fpa_value_factory() {}
~fpa_value_factory() override {}
virtual expr * get_some_value(sort * s) {
expr * get_some_value(sort * s) override {
mpf_manager & mpfm = m_util.fm();
scoped_mpf q(mpfm);
mpfm.set(q, m_util.get_ebits(s), m_util.get_sbits(s), 0);
return m_util.mk_value(q);
}
virtual bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2) {
bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2) override {
mpf_manager & mpfm = m_util.fm();
scoped_mpf q(mpfm);
mpfm.set(q, m_util.get_ebits(s), m_util.get_sbits(s), 0);
@ -62,8 +62,8 @@ namespace smt {
return true;
}
virtual expr * get_fresh_value(sort * s) { NOT_IMPLEMENTED_YET(); }
virtual void register_value(expr * n) { /* Ignore */ }
expr * get_fresh_value(sort * s) override { NOT_IMPLEMENTED_YET(); }
void register_value(expr * n) override { /* Ignore */ }
app * mk_value(mpf const & x) {
return m_util.mk_value(x);
@ -81,8 +81,8 @@ namespace smt {
fpa2bv_converter(m),
m_th(*th) {}
virtual ~fpa2bv_converter_wrapped() {}
virtual void mk_const(func_decl * f, expr_ref & result);
virtual void mk_rm_const(func_decl * f, expr_ref & result);
void mk_const(func_decl * f, expr_ref & result) override;
void mk_rm_const(func_decl * f, expr_ref & result) override;
};
class fpa_value_proc : public model_value_proc {
@ -100,15 +100,15 @@ namespace smt {
m_th(*th), m(th->get_manager()), m_fu(th->m_fpa_util), m_bu(th->m_bv_util),
m_ebits(ebits), m_sbits(sbits) {}
virtual ~fpa_value_proc() {}
~fpa_value_proc() override {}
void add_dependency(enode * e) { m_deps.push_back(model_value_dependency(e)); }
virtual void get_dependencies(buffer<model_value_dependency> & result) {
void get_dependencies(buffer<model_value_dependency> & result) override {
result.append(m_deps);
}
virtual app * mk_value(model_generator & mg, ptr_vector<expr> & values);
app * mk_value(model_generator & mg, ptr_vector<expr> & values) override;
};
class fpa_rm_value_proc : public model_value_proc {
@ -124,12 +124,12 @@ namespace smt {
void add_dependency(enode * e) { m_deps.push_back(model_value_dependency(e)); }
virtual void get_dependencies(buffer<model_value_dependency> & result) {
void get_dependencies(buffer<model_value_dependency> & result) override {
result.append(m_deps);
}
virtual ~fpa_rm_value_proc() {}
virtual app * mk_value(model_generator & mg, ptr_vector<expr> & values);
~fpa_rm_value_proc() override {}
app * mk_value(model_generator & mg, ptr_vector<expr> & values) override;
};
protected:
@ -145,32 +145,32 @@ namespace smt {
bool m_is_initialized;
obj_hashtable<func_decl> m_is_added_to_model;
virtual final_check_status final_check_eh();
virtual bool internalize_atom(app * atom, bool gate_ctx);
virtual bool internalize_term(app * term);
virtual void apply_sort_cnstr(enode * n, sort * s);
virtual void new_eq_eh(theory_var, theory_var);
virtual void new_diseq_eh(theory_var, theory_var);
virtual void push_scope_eh();
virtual void pop_scope_eh(unsigned num_scopes);
virtual void reset_eh();
virtual theory* mk_fresh(context* new_ctx);
virtual char const * get_name() const { return "fpa"; }
final_check_status final_check_eh() override;
bool internalize_atom(app * atom, bool gate_ctx) override;
bool internalize_term(app * term) override;
void apply_sort_cnstr(enode * n, sort * s) override;
void new_eq_eh(theory_var, theory_var) override;
void new_diseq_eh(theory_var, theory_var) override;
void push_scope_eh() override;
void pop_scope_eh(unsigned num_scopes) override;
void reset_eh() override;
theory* mk_fresh(context* new_ctx) override;
char const * get_name() const override { return "fpa"; }
virtual model_value_proc * mk_value(enode * n, model_generator & mg);
model_value_proc * mk_value(enode * n, model_generator & mg) override;
void assign_eh(bool_var v, bool is_true);
virtual void relevant_eh(app * n);
virtual void init_model(model_generator & m);
virtual void finalize_model(model_generator & mg);
void assign_eh(bool_var v, bool is_true) override;
void relevant_eh(app * n) override;
void init_model(model_generator & m) override;
void finalize_model(model_generator & mg) override;
public:
theory_fpa(ast_manager & m);
virtual ~theory_fpa();
~theory_fpa() override;
virtual void init(context * ctx);
void init(context * ctx) override;
virtual void display(std::ostream & out) const;
void display(std::ostream & out) const override;
protected:
expr_ref mk_side_conditions();

View file

@ -137,7 +137,7 @@ namespace smt {
imp& m_imp;
public:
resource_limit(imp& i): m_imp(i) { }
virtual bool get_cancel_flag() { return m_imp.m.canceled(); }
bool get_cancel_flag() override { return m_imp.m.canceled(); }
};
@ -1358,11 +1358,11 @@ namespace smt {
imp & m_imp;
local_bound_propagator(imp& i) : lp_bound_propagator(*i.m_solver), m_imp(i) {}
bool bound_is_interesting(unsigned j, lp::lconstraint_kind kind, const rational & v) {
bool bound_is_interesting(unsigned j, lp::lconstraint_kind kind, const rational & v) override {
return m_imp.bound_is_interesting(j, kind, v);
}
virtual void consume(rational const& v, unsigned j) {
void consume(rational const& v, unsigned j) override {
m_imp.set_evidence(j);
}
};

View file

@ -29,66 +29,66 @@ namespace smt {
public:
theory_lra(ast_manager& m, theory_arith_params& ap);
virtual ~theory_lra();
virtual theory* mk_fresh(context* new_ctx);
virtual char const* get_name() const { return "lra"; }
~theory_lra() override;
theory* mk_fresh(context* new_ctx) override;
char const* get_name() const override { return "lra"; }
virtual void init(context * ctx);
void init(context * ctx) override;
virtual bool internalize_atom(app * atom, bool gate_ctx);
bool internalize_atom(app * atom, bool gate_ctx) override;
virtual bool internalize_term(app * term);
bool internalize_term(app * term) override;
virtual void internalize_eq_eh(app * atom, bool_var v);
void internalize_eq_eh(app * atom, bool_var v) override;
virtual void assign_eh(bool_var v, bool is_true);
void assign_eh(bool_var v, bool is_true) override;
virtual void new_eq_eh(theory_var v1, theory_var v2);
void new_eq_eh(theory_var v1, theory_var v2) override;
virtual bool use_diseqs() const;
bool use_diseqs() const override;
virtual void new_diseq_eh(theory_var v1, theory_var v2);
void new_diseq_eh(theory_var v1, theory_var v2) override;
virtual void push_scope_eh();
void push_scope_eh() override;
virtual void pop_scope_eh(unsigned num_scopes);
void pop_scope_eh(unsigned num_scopes) override;
virtual void restart_eh();
void restart_eh() override;
virtual void relevant_eh(app* e);
void relevant_eh(app* e) override;
virtual void init_search_eh();
void init_search_eh() override;
virtual final_check_status final_check_eh();
final_check_status final_check_eh() override;
virtual bool is_shared(theory_var v) const;
bool is_shared(theory_var v) const override;
virtual bool can_propagate();
bool can_propagate() override;
virtual void propagate();
void propagate() override;
virtual justification * why_is_diseq(theory_var v1, theory_var v2);
justification * why_is_diseq(theory_var v1, theory_var v2) override;
// virtual void flush_eh();
virtual void reset_eh();
void reset_eh() override;
virtual void init_model(model_generator & m);
void init_model(model_generator & m) override;
virtual model_value_proc * mk_value(enode * n, model_generator & mg);
model_value_proc * mk_value(enode * n, model_generator & mg) override;
virtual bool get_value(enode* n, expr_ref& r);
bool get_value(enode* n, expr_ref& r) override;
virtual bool validate_eq_in_model(theory_var v1, theory_var v2, bool is_true) const;
bool validate_eq_in_model(theory_var v1, theory_var v2, bool is_true) const override;
virtual void display(std::ostream & out) const;
void display(std::ostream & out) const override;
virtual void collect_statistics(::statistics & st) const;
void collect_statistics(::statistics & st) const override;
// optimization
virtual inf_eps value(theory_var);
virtual inf_eps maximize(theory_var v, expr_ref& blocker, bool& has_shared);
virtual theory_var add_objective(app* term);
inf_eps value(theory_var) override;
inf_eps maximize(theory_var v, expr_ref& blocker, bool& has_shared) override;
theory_var add_objective(app* term) override;
virtual expr_ref mk_ge(filter_model_converter& fm, theory_var v, inf_rational const& val);

View file

@ -254,7 +254,7 @@ namespace smt {
unsigned v;
public:
remove_var(theory_pb& pb, unsigned v): pb(pb), v(v) {}
virtual void undo(context& ctx) {
void undo(context& ctx) override {
pb.m_vars.remove(v);
pb.m_simplex.unset_lower(v);
pb.m_simplex.unset_upper(v);
@ -282,7 +282,7 @@ namespace smt {
m_last_bound_valid(last_bound_valid),
m_last_explain(last_explain) {}
virtual void undo(context& ctx) {
void undo(context& ctx) override {
if (m_is_lower) {
if (m_last_bound_valid) {
pb.m_simplex.set_lower(m_v, m_last_bound);
@ -889,7 +889,7 @@ namespace smt {
ineq& c;
public:
rewatch_vars(theory_pb& p, ineq& c): pb(p), c(c) {}
virtual void undo(context& ctx) {
void undo(context& ctx) override {
for (unsigned i = 0; i < c.size(); ++i) {
pb.watch_var(c.lit(i).var(), &c);
}
@ -900,7 +900,7 @@ namespace smt {
ineq& c;
public:
negate_ineq(ineq& c): c(c) {}
virtual void undo(context& ctx) {
void undo(context& ctx) override {
c.negate();
}
};
@ -1357,7 +1357,7 @@ namespace smt {
public:
unwatch_ge(theory_pb& p, ineq& c): pb(p), c(c) {}
virtual void undo(context& ctx) {
void undo(context& ctx) override {
for (unsigned i = 0; i < c.watch_size(); ++i) {
pb.unwatch_literal(c.lit(i), &c);
}
@ -2008,11 +2008,11 @@ namespace smt {
m_dependencies.push_back(model_value_dependency(n));
}
virtual void get_dependencies(buffer<model_value_dependency> & result) {
void get_dependencies(buffer<model_value_dependency> & result) override {
result.append(m_dependencies.size(), m_dependencies.c_ptr());
}
virtual app * mk_value(model_generator & mg, ptr_vector<expr> & values) {
app * mk_value(model_generator & mg, ptr_vector<expr> & values) override {
ast_manager& m = mg.get_manager();
SASSERT(values.size() == m_dependencies.size());
SASSERT(values.size() == m_app->get_num_args());
@ -2049,18 +2049,18 @@ namespace smt {
pb_factory(ast_manager& m, family_id fid):
value_factory(m, fid) {}
virtual expr * get_some_value(sort * s) {
expr * get_some_value(sort * s) override {
return m_manager.mk_true();
}
virtual bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2) {
bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2) override {
v1 = m_manager.mk_true();
v2 = m_manager.mk_false();
return true;
}
virtual expr * get_fresh_value(sort * s) {
expr * get_fresh_value(sort * s) override {
return 0;
}
virtual void register_value(expr * n) { }
void register_value(expr * n) override { }
};
void theory_pb::init_model(model_generator & m) {

View file

@ -261,7 +261,7 @@ namespace smt {
std::ostream& display(std::ostream& out, ineq const& c, bool values = false) const;
std::ostream& display(std::ostream& out, arg_t const& c, bool values = false) const;
virtual void display(std::ostream& out) const;
void display(std::ostream& out) const override;
void display_watch(std::ostream& out, bool_var v, bool sign) const;
void display_resolved_lemma(std::ostream& out) const;
@ -317,26 +317,26 @@ namespace smt {
public:
theory_pb(ast_manager& m, theory_pb_params& p);
virtual ~theory_pb();
~theory_pb() override;
virtual theory * mk_fresh(context * new_ctx);
virtual bool internalize_atom(app * atom, bool gate_ctx);
virtual bool internalize_term(app * term) { UNREACHABLE(); return false; }
virtual void new_eq_eh(theory_var v1, theory_var v2);
virtual void new_diseq_eh(theory_var v1, theory_var v2) { }
virtual bool use_diseqs() const { return false; }
virtual bool build_models() const { return false; }
virtual final_check_status final_check_eh();
virtual void reset_eh();
virtual void assign_eh(bool_var v, bool is_true);
virtual void init_search_eh();
virtual void push_scope_eh();
virtual void pop_scope_eh(unsigned num_scopes);
virtual void restart_eh();
virtual void collect_statistics(::statistics & st) const;
virtual model_value_proc * mk_value(enode * n, model_generator & mg);
virtual void init_model(model_generator & m);
virtual bool include_func_interp(func_decl* f) { return false; }
theory * mk_fresh(context * new_ctx) override;
bool internalize_atom(app * atom, bool gate_ctx) override;
bool internalize_term(app * term) override { UNREACHABLE(); return false; }
void new_eq_eh(theory_var v1, theory_var v2) override;
void new_diseq_eh(theory_var v1, theory_var v2) override { }
bool use_diseqs() const override { return false; }
bool build_models() const override { return false; }
final_check_status final_check_eh() override;
void reset_eh() override;
void assign_eh(bool_var v, bool is_true) override;
void init_search_eh() override;
void push_scope_eh() override;
void pop_scope_eh(unsigned num_scopes) override;
void restart_eh() override;
void collect_statistics(::statistics & st) const override;
model_value_proc * mk_value(enode * n, model_generator & mg) override;
void init_model(model_generator & m) override;
bool include_func_interp(func_decl* f) override { return false; }
static literal assert_ge(context& ctx, unsigned k, unsigned n, literal const* xs);
};

View file

@ -45,7 +45,7 @@ public:
m_kernel(m, fp)
{}
virtual lbool check_sat(expr* e) {
lbool check_sat(expr* e) override {
m_kernel.push();
m_kernel.assert_expr(e);
lbool r = m_kernel.check();
@ -2691,7 +2691,7 @@ class theory_seq::seq_value_proc : public model_value_proc {
public:
seq_value_proc(theory_seq& th, sort* s): th(th), m_sort(s) {
}
virtual ~seq_value_proc() {}
~seq_value_proc() override {}
void add_unit(enode* n) {
m_dependencies.push_back(model_value_dependency(n));
m_source.push_back(unit_source);
@ -2704,7 +2704,7 @@ public:
m_strings.push_back(n);
m_source.push_back(string_source);
}
virtual void get_dependencies(buffer<model_value_dependency> & result) {
void get_dependencies(buffer<model_value_dependency> & result) override {
result.append(m_dependencies.size(), m_dependencies.c_ptr());
}
@ -2714,7 +2714,7 @@ public:
}
}
virtual app * mk_value(model_generator & mg, ptr_vector<expr> & values) {
app * mk_value(model_generator & mg, ptr_vector<expr> & values) override {
SASSERT(values.size() == m_dependencies.size());
expr_ref_vector args(th.m);
unsigned j = 0, k = 0;

View file

@ -224,8 +224,8 @@ namespace smt {
expr_ref m_e;
public:
replay_length_coherence(ast_manager& m, expr* e) : m_e(e, m) {}
virtual ~replay_length_coherence() {}
virtual void operator()(theory_seq& th) {
~replay_length_coherence() override {}
void operator()(theory_seq& th) override {
th.check_length_coherence(m_e);
m_e.reset();
}
@ -235,8 +235,8 @@ namespace smt {
expr_ref m_e;
public:
replay_fixed_length(ast_manager& m, expr* e) : m_e(e, m) {}
virtual ~replay_fixed_length() {}
virtual void operator()(theory_seq& th) {
~replay_fixed_length() override {}
void operator()(theory_seq& th) override {
th.fixed_length(m_e);
m_e.reset();
}
@ -246,8 +246,8 @@ namespace smt {
expr_ref m_e;
public:
replay_axiom(ast_manager& m, expr* e) : m_e(e, m) {}
virtual ~replay_axiom() {}
virtual void operator()(theory_seq& th) {
~replay_axiom() override {}
void operator()(theory_seq& th) override {
th.enque_axiom(m_e);
m_e.reset();
}
@ -257,7 +257,7 @@ namespace smt {
apply* m_apply;
public:
push_replay(apply* app): m_apply(app) {}
virtual void undo(theory_seq& th) {
void undo(theory_seq& th) override {
th.m_replay.push_back(m_apply);
}
};
@ -266,7 +266,7 @@ namespace smt {
unsigned k;
public:
pop_branch(unsigned k): k(k) {}
virtual void undo(theory_seq& th) {
void undo(theory_seq& th) override {
th.m_branch_start.erase(k);
}
};
@ -340,29 +340,29 @@ namespace smt {
obj_hashtable<expr> m_fixed; // string variables that are fixed length.
virtual void init(context* ctx);
virtual final_check_status final_check_eh();
virtual bool internalize_atom(app* atom, bool);
virtual bool internalize_term(app*);
virtual void internalize_eq_eh(app * atom, bool_var v);
virtual void new_eq_eh(theory_var, theory_var);
virtual void new_diseq_eh(theory_var, theory_var);
virtual void assign_eh(bool_var v, bool is_true);
virtual bool can_propagate();
virtual void propagate();
virtual void push_scope_eh();
virtual void pop_scope_eh(unsigned num_scopes);
virtual void restart_eh();
virtual void relevant_eh(app* n);
virtual theory* mk_fresh(context* new_ctx) { return alloc(theory_seq, new_ctx->get_manager()); }
virtual char const * get_name() const { return "seq"; }
virtual theory_var mk_var(enode* n);
virtual void apply_sort_cnstr(enode* n, sort* s);
virtual void display(std::ostream & out) const;
virtual void collect_statistics(::statistics & st) const;
virtual model_value_proc * mk_value(enode * n, model_generator & mg);
virtual void init_model(model_generator & mg);
virtual void init_search_eh();
void init(context* ctx) override;
final_check_status final_check_eh() override;
bool internalize_atom(app* atom, bool) override;
bool internalize_term(app*) override;
void internalize_eq_eh(app * atom, bool_var v) override;
void new_eq_eh(theory_var, theory_var) override;
void new_diseq_eh(theory_var, theory_var) override;
void assign_eh(bool_var v, bool is_true) override;
bool can_propagate() override;
void propagate() override;
void push_scope_eh() override;
void pop_scope_eh(unsigned num_scopes) override;
void restart_eh() override;
void relevant_eh(app* n) override;
theory* mk_fresh(context* new_ctx) override { return alloc(theory_seq, new_ctx->get_manager()); }
char const * get_name() const override { return "seq"; }
theory_var mk_var(enode* n) override;
void apply_sort_cnstr(enode* n, sort* s) override;
void display(std::ostream & out) const override;
void collect_statistics(::statistics & st) const override;
model_value_proc * mk_value(enode * n, model_generator & mg) override;
void init_model(model_generator & mg) override;
void init_search_eh() override;
void init_model(expr_ref_vector const& es);
// final check
@ -584,7 +584,7 @@ namespace smt {
void display_nc(std::ostream& out, nc const& nc) const;
public:
theory_seq(ast_manager& m);
virtual ~theory_seq();
~theory_seq() override;
// model building
app* mk_value(app* a);

View file

@ -63,7 +63,7 @@ namespace smt {
m_unique_sequences.insert(m.get_sort(uniq), uniq);
}
virtual expr* get_some_value(sort* s) {
expr* get_some_value(sort* s) override {
if (u.is_seq(s)) {
return u.str.mk_empty(s);
}
@ -74,7 +74,7 @@ namespace smt {
UNREACHABLE();
return 0;
}
virtual bool get_some_values(sort* s, expr_ref& v1, expr_ref& v2) {
bool get_some_values(sort* s, expr_ref& v1, expr_ref& v2) override {
if (u.is_string(s)) {
v1 = u.str.mk_string(symbol("a"));
v2 = u.str.mk_string(symbol("b"));
@ -94,7 +94,7 @@ namespace smt {
NOT_IMPLEMENTED_YET();
return false;
}
virtual expr* get_fresh_value(sort* s) {
expr* get_fresh_value(sort* s) override {
if (u.is_string(s)) {
while (true) {
std::ostringstream strm;
@ -122,7 +122,7 @@ namespace smt {
UNREACHABLE();
return 0;
}
virtual void register_value(expr* n) {
void register_value(expr* n) override {
symbol sym;
if (u.str.is_string(n, sym)) {
m_strings.insert(sym);
@ -148,17 +148,17 @@ namespace smt {
class theory_seq_empty : public theory {
bool m_used;
virtual final_check_status final_check_eh() { return m_used?FC_GIVEUP:FC_DONE; }
virtual bool internalize_atom(app*, bool) { if (!m_used) { get_context().push_trail(value_trail<context,bool>(m_used)); m_used = true; } return false; }
virtual bool internalize_term(app*) { return internalize_atom(0,false); }
virtual void new_eq_eh(theory_var, theory_var) { }
virtual void new_diseq_eh(theory_var, theory_var) {}
virtual theory* mk_fresh(context* new_ctx) { return alloc(theory_seq_empty, new_ctx->get_manager()); }
virtual char const * get_name() const { return "seq-empty"; }
virtual void display(std::ostream& out) const {}
final_check_status final_check_eh() override { return m_used?FC_GIVEUP:FC_DONE; }
bool internalize_atom(app*, bool) override { if (!m_used) { get_context().push_trail(value_trail<context,bool>(m_used)); m_used = true; } return false; }
bool internalize_term(app*) override { return internalize_atom(0,false); }
void new_eq_eh(theory_var, theory_var) override { }
void new_diseq_eh(theory_var, theory_var) override {}
theory* mk_fresh(context* new_ctx) override { return alloc(theory_seq_empty, new_ctx->get_manager()); }
char const * get_name() const override { return "seq-empty"; }
void display(std::ostream& out) const override {}
public:
theory_seq_empty(ast_manager& m):theory(m.mk_family_id("seq")), m_used(false) {}
virtual void init_model(model_generator & mg) {
void init_model(model_generator & mg) override {
mg.register_factory(alloc(seq_factory, get_manager(), get_family_id(), mg.get_model()));
}

View file

@ -46,16 +46,16 @@ public:
str_value_factory(ast_manager & m, family_id fid) :
value_factory(m, fid),
u(m), delim("!"), m_next(0) {}
virtual ~str_value_factory() {}
virtual expr * get_some_value(sort * s) {
~str_value_factory() override {}
expr * get_some_value(sort * s) override {
return u.str.mk_string(symbol("some value"));
}
virtual bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2) {
bool get_some_values(sort * s, expr_ref & v1, expr_ref & v2) override {
v1 = u.str.mk_string(symbol("value 1"));
v2 = u.str.mk_string(symbol("value 2"));
return true;
}
virtual expr * get_fresh_value(sort * s) {
expr * get_fresh_value(sort * s) override {
if (u.is_string(s)) {
while (true) {
std::ostringstream strm;
@ -74,7 +74,7 @@ public:
TRACE("t_str", tout << "unexpected sort in get_fresh_value(): " << mk_pp(s, m_manager) << std::endl;);
UNREACHABLE(); return NULL;
}
virtual void register_value(expr * n) { /* Ignore */ }
void register_value(expr * n) override { /* Ignore */ }
};
// rather than modify obj_pair_map I inherit from it and add my own helper methods
@ -104,8 +104,8 @@ class binary_search_trail : public trail<Ctx> {
public:
binary_search_trail(obj_map<expr, ptr_vector<expr> > & target, expr * entry) :
target(target), entry(entry) {}
virtual ~binary_search_trail() {}
virtual void undo(Ctx & ctx) {
~binary_search_trail() override {}
void undo(Ctx & ctx) override {
TRACE("t_str_binary_search", tout << "in binary_search_trail::undo()" << std::endl;);
if (target.contains(entry)) {
if (!target[entry].empty()) {
@ -616,10 +616,10 @@ protected:
public:
theory_str(ast_manager & m, theory_str_params const & params);
virtual ~theory_str();
~theory_str() override;
virtual char const * get_name() const { return "seq"; }
virtual void display(std::ostream & out) const;
char const * get_name() const override { return "seq"; }
void display(std::ostream & out) const override;
bool overlapping_variables_detected() const { return loopDetected; }
@ -628,33 +628,33 @@ public:
void after_merge_eh(theory_var r1, theory_var r2, theory_var v1, theory_var v2) { }
void unmerge_eh(theory_var v1, theory_var v2) {}
protected:
virtual bool internalize_atom(app * atom, bool gate_ctx);
virtual bool internalize_term(app * term);
bool internalize_atom(app * atom, bool gate_ctx) override;
bool internalize_term(app * term) override;
virtual enode* ensure_enode(expr* e);
virtual theory_var mk_var(enode * n);
theory_var mk_var(enode * n) override;
virtual void new_eq_eh(theory_var, theory_var);
virtual void new_diseq_eh(theory_var, theory_var);
void new_eq_eh(theory_var, theory_var) override;
void new_diseq_eh(theory_var, theory_var) override;
virtual theory* mk_fresh(context*) { return alloc(theory_str, get_manager(), m_params); }
virtual void init_search_eh();
virtual void add_theory_assumptions(expr_ref_vector & assumptions);
virtual lbool validate_unsat_core(expr_ref_vector & unsat_core);
virtual void relevant_eh(app * n);
virtual void assign_eh(bool_var v, bool is_true);
virtual void push_scope_eh();
virtual void pop_scope_eh(unsigned num_scopes);
virtual void reset_eh();
theory* mk_fresh(context*) override { return alloc(theory_str, get_manager(), m_params); }
void init_search_eh() override;
void add_theory_assumptions(expr_ref_vector & assumptions) override;
lbool validate_unsat_core(expr_ref_vector & unsat_core) override;
void relevant_eh(app * n) override;
void assign_eh(bool_var v, bool is_true) override;
void push_scope_eh() override;
void pop_scope_eh(unsigned num_scopes) override;
void reset_eh() override;
virtual bool can_propagate();
virtual void propagate();
bool can_propagate() override;
void propagate() override;
virtual final_check_status final_check_eh();
final_check_status final_check_eh() override;
virtual void attach_new_th_var(enode * n);
virtual void init_model(model_generator & m);
virtual model_value_proc * mk_value(enode * n, model_generator & mg);
virtual void finalize_model(model_generator & mg);
void init_model(model_generator & m) override;
model_value_proc * mk_value(enode * n, model_generator & mg) override;
void finalize_model(model_generator & mg) override;
};
};

View file

@ -166,7 +166,7 @@ namespace smt {
void new_edge(dl_var src, dl_var dst, unsigned num_edges, edge_id const* edges) {}
// Create a new theory variable.
virtual th_var mk_var(enode* n);
th_var mk_var(enode* n) override;
virtual th_var mk_var(expr* n);
@ -181,83 +181,83 @@ namespace smt {
public:
theory_utvpi(ast_manager& m);
virtual ~theory_utvpi();
~theory_utvpi() override;
virtual theory * mk_fresh(context * new_ctx);
theory * mk_fresh(context * new_ctx) override;
virtual char const * get_name() const { return "utvpi-logic"; }
char const * get_name() const override { return "utvpi-logic"; }
/**
\brief See comment in theory::mk_eq_atom
*/
virtual app * mk_eq_atom(expr * lhs, expr * rhs) { return a.mk_eq(lhs, rhs); }
app * mk_eq_atom(expr * lhs, expr * rhs) override { return a.mk_eq(lhs, rhs); }
virtual void init(context * ctx);
void init(context * ctx) override;
virtual bool internalize_atom(app * atom, bool gate_ctx);
bool internalize_atom(app * atom, bool gate_ctx) override;
virtual bool internalize_term(app * term);
bool internalize_term(app * term) override;
virtual void internalize_eq_eh(app * atom, bool_var v);
void internalize_eq_eh(app * atom, bool_var v) override;
virtual void assign_eh(bool_var v, bool is_true);
void assign_eh(bool_var v, bool is_true) override;
virtual void new_eq_eh(th_var v1, th_var v2) {
void new_eq_eh(th_var v1, th_var v2) override {
m_stats.m_num_core2th_eqs++;
m_arith_eq_adapter.new_eq_eh(v1, v2);
}
virtual bool use_diseqs() const { return true; }
bool use_diseqs() const override { return true; }
virtual void new_diseq_eh(th_var v1, th_var v2) {
void new_diseq_eh(th_var v1, th_var v2) override {
m_arith_eq_adapter.new_diseq_eh(v1, v2);
}
virtual void push_scope_eh();
void push_scope_eh() override;
virtual void pop_scope_eh(unsigned num_scopes);
void pop_scope_eh(unsigned num_scopes) override;
virtual void restart_eh() {
void restart_eh() override {
m_arith_eq_adapter.restart_eh();
}
virtual void relevant_eh(app* e) {}
void relevant_eh(app* e) override {}
virtual void init_search_eh() {
void init_search_eh() override {
m_arith_eq_adapter.init_search_eh();
}
virtual final_check_status final_check_eh();
final_check_status final_check_eh() override;
virtual bool is_shared(th_var v) const {
bool is_shared(th_var v) const override {
return false;
}
virtual bool can_propagate() {
bool can_propagate() override {
SASSERT(m_asserted_qhead <= m_asserted_atoms.size());
return m_asserted_qhead != m_asserted_atoms.size();
}
virtual void propagate();
void propagate() override;
virtual justification * why_is_diseq(th_var v1, th_var v2) {
justification * why_is_diseq(th_var v1, th_var v2) override {
UNREACHABLE();
return 0;
}
virtual void reset_eh();
void reset_eh() override;
virtual void init_model(model_generator & m);
void init_model(model_generator & m) override;
virtual model_value_proc * mk_value(enode * n, model_generator & mg);
model_value_proc * mk_value(enode * n, model_generator & mg) override;
virtual bool validate_eq_in_model(th_var v1, th_var v2, bool is_true) const {
bool validate_eq_in_model(th_var v1, th_var v2, bool is_true) const override {
return true;
}
virtual void display(std::ostream & out) const;
void display(std::ostream & out) const override;
virtual void collect_statistics(::statistics & st) const;
void collect_statistics(::statistics & st) const override;
private:

View file

@ -57,7 +57,7 @@ namespace smt {
stats m_stats;
public:
theory_wmaxsat(ast_manager& m, filter_model_converter& mc);
virtual ~theory_wmaxsat();
~theory_wmaxsat() override;
void get_assignment(svector<bool>& result);
expr* assert_weighted(expr* fml, rational const& w);
void disable_var(expr* var);
@ -76,43 +76,43 @@ namespace smt {
old.push_back(value);
}
virtual ~numeral_trail() {
~numeral_trail() override {
}
virtual void undo(context & ctx) {
void undo(context & ctx) override {
m_value = m_old_values.back();
m_old_values.shrink(m_old_values.size() - 1);
}
};
virtual void init_search_eh();
virtual void assign_eh(bool_var v, bool is_true);
virtual final_check_status final_check_eh();
virtual bool use_diseqs() const {
void init_search_eh() override;
void assign_eh(bool_var v, bool is_true) override;
final_check_status final_check_eh() override;
bool use_diseqs() const override {
return false;
}
virtual bool build_models() const {
bool build_models() const override {
return false;
}
void reset_local();
virtual void reset_eh();
virtual theory * mk_fresh(context * new_ctx) { return 0; }
virtual bool internalize_atom(app * atom, bool gate_ctx) { return false; }
virtual bool internalize_term(app * term) { return false; }
virtual void new_eq_eh(theory_var v1, theory_var v2) { }
virtual void new_diseq_eh(theory_var v1, theory_var v2) { }
virtual void display(std::ostream& out) const {}
virtual void restart_eh();
void reset_eh() override;
theory * mk_fresh(context * new_ctx) override { return 0; }
bool internalize_atom(app * atom, bool gate_ctx) override { return false; }
bool internalize_term(app * term) override { return false; }
void new_eq_eh(theory_var v1, theory_var v2) override { }
void new_diseq_eh(theory_var v1, theory_var v2) override { }
void display(std::ostream& out) const override {}
void restart_eh() override;
virtual void collect_statistics(::statistics & st) const {
void collect_statistics(::statistics & st) const override {
st.update("wmaxsat num blocks", m_stats.m_num_blocks);
st.update("wmaxsat num props", m_stats.m_num_propagations);
}
virtual bool can_propagate() {
bool can_propagate() override {
return m_propagate || m_can_propagate;
}
virtual void propagate();
void propagate() override;
bool is_optimal() const;
expr_ref mk_block();