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

prepare polysat

Signed-off-by: Nikolaj Bjorner <nbjorner@microsoft.com>
This commit is contained in:
Nikolaj Bjorner 2022-01-26 06:19:24 +01:00
commit 3f5df04dc4
252 changed files with 5792 additions and 2553 deletions

View file

@ -273,6 +273,17 @@ public:
bool is_rem0(func_decl const * n) const { return is_decl_of(n, arith_family_id, OP_REM0); }
bool is_mod0(func_decl const * n) const { return is_decl_of(n, arith_family_id, OP_MOD0); }
bool is_power0(func_decl const * n) const { return is_decl_of(n, arith_family_id, OP_POWER0); }
bool is_power(func_decl const * n) const { return is_decl_of(n, arith_family_id, OP_POWER); }
bool is_add(func_decl const* f) const { return is_decl_of(f, arith_family_id, OP_ADD); }
bool is_mul(func_decl const* f) const { return is_decl_of(f, arith_family_id, OP_MUL); }
bool is_sub(func_decl const* f) const { return is_decl_of(f, arith_family_id, OP_SUB); }
bool is_uminus(func_decl const* f) const { return is_decl_of(f, arith_family_id, OP_UMINUS); }
bool is_div(func_decl const* f) const { return is_decl_of(f, arith_family_id, OP_DIV); }
bool is_rem(func_decl const* f) const { return is_decl_of(f, arith_family_id, OP_REM); }
bool is_mod(func_decl const* f) const { return is_decl_of(f, arith_family_id, OP_MOD); }
bool is_to_real(func_decl const* f) const { return is_decl_of(f, arith_family_id, OP_TO_REAL); }
bool is_to_int(func_decl const* f) const { return is_decl_of(f, arith_family_id, OP_TO_INT); }
bool is_is_int(func_decl const* f) const { return is_decl_of(f, arith_family_id, OP_IS_INT); }
bool is_add(expr const * n) const { return is_app_of(n, arith_family_id, OP_ADD); }
bool is_sub(expr const * n) const { return is_app_of(n, arith_family_id, OP_SUB); }
@ -291,6 +302,7 @@ public:
bool is_is_int(expr const * n) const { return is_app_of(n, arith_family_id, OP_IS_INT); }
bool is_power(expr const * n) const { return is_app_of(n, arith_family_id, OP_POWER); }
bool is_power0(expr const * n) const { return is_app_of(n, arith_family_id, OP_POWER0); }
bool is_abs(expr const* n) const { return is_app_of(n, arith_family_id, OP_ABS); }
bool is_int(sort const * s) const { return is_sort_of(s, arith_family_id, INT_SORT); }
bool is_int(expr const * n) const { return is_int(n->get_sort()); }
@ -330,6 +342,7 @@ public:
MATCH_UNARY(is_to_real);
MATCH_UNARY(is_to_int);
MATCH_UNARY(is_is_int);
MATCH_UNARY(is_abs);
MATCH_BINARY(is_sub);
MATCH_BINARY(is_add);
MATCH_BINARY(is_mul);

View file

@ -149,7 +149,12 @@ public:
bool is_store(expr* n) const { return is_app_of(n, m_fid, OP_STORE); }
bool is_const(expr* n) const { return is_app_of(n, m_fid, OP_CONST_ARRAY); }
bool is_ext(expr* n) const { return is_app_of(n, m_fid, OP_ARRAY_EXT); }
bool is_ext(func_decl const* f) const { return is_decl_of(f, m_fid, OP_ARRAY_EXT); }
bool is_map(expr* n) const { return is_app_of(n, m_fid, OP_ARRAY_MAP); }
bool is_union(expr* n) const { return is_app_of(n, m_fid, OP_SET_UNION); }
bool is_intersect(expr* n) const { return is_app_of(n, m_fid, OP_SET_INTERSECT); }
bool is_difference(expr* n) const { return is_app_of(n, m_fid, OP_SET_DIFFERENCE); }
bool is_complement(expr* n) const { return is_app_of(n, m_fid, OP_SET_COMPLEMENT); }
bool is_as_array(expr * n) const { return is_app_of(n, m_fid, OP_AS_ARRAY); }
bool is_as_array(expr * n, func_decl*& f) const { return is_as_array(n) && (f = get_as_array_func_decl(n), true); }
bool is_set_has_size(expr* e) const { return is_app_of(e, m_fid, OP_SET_HAS_SIZE); }
@ -158,11 +163,14 @@ public:
bool is_store(func_decl* f) const { return is_decl_of(f, m_fid, OP_STORE); }
bool is_const(func_decl* f) const { return is_decl_of(f, m_fid, OP_CONST_ARRAY); }
bool is_map(func_decl* f) const { return is_decl_of(f, m_fid, OP_ARRAY_MAP); }
bool is_union(func_decl* f) const { return is_decl_of(f, m_fid, OP_SET_UNION); }
bool is_intersect(func_decl* f) const { return is_decl_of(f, m_fid, OP_SET_INTERSECT); }
bool is_as_array(func_decl* f) const { return is_decl_of(f, m_fid, OP_AS_ARRAY); }
bool is_set_has_size(func_decl* f) const { return is_decl_of(f, m_fid, OP_SET_HAS_SIZE); }
bool is_set_card(func_decl* f) const { return is_decl_of(f, m_fid, OP_SET_CARD); }
bool is_default(func_decl* f) const { return is_decl_of(f, m_fid, OP_ARRAY_DEFAULT); }
bool is_default(expr* n) const { return is_app_of(n, m_fid, OP_ARRAY_DEFAULT); }
bool is_subset(expr const* n) const { return is_app_of(n, m_fid, OP_SET_SUBSET); }
bool is_as_array(func_decl* f, func_decl*& g) const { return is_decl_of(f, m_fid, OP_AS_ARRAY) && (g = get_as_array_func_decl(f), true); }
func_decl * get_as_array_func_decl(expr * n) const;
func_decl * get_as_array_func_decl(func_decl* f) const;
@ -172,6 +180,8 @@ public:
bool is_const(expr* e, expr*& v) const;
bool is_store_ext(expr* e, expr_ref& a, expr_ref_vector& args, expr_ref& value);
MATCH_BINARY(is_subset);
};
class array_util : public array_recognizers {

View file

@ -845,7 +845,7 @@ class quantifier : public expr {
char m_patterns_decls[0];
static unsigned get_obj_size(unsigned num_decls, unsigned num_patterns, unsigned num_no_patterns) {
return sizeof(quantifier) + num_decls * (sizeof(sort *) + sizeof(symbol)) + (num_patterns + num_no_patterns) * sizeof(expr*);
return (unsigned)(sizeof(quantifier) + num_decls * (sizeof(sort *) + sizeof(symbol)) + (num_patterns + num_no_patterns) * sizeof(expr*));
}
quantifier(quantifier_kind k, unsigned num_decls, sort * const * decl_sorts, symbol const * decl_names, expr * body, sort* s,

View file

@ -260,39 +260,52 @@ expr_ref mk_distinct(expr_ref_vector const& args) {
void flatten_and(expr_ref_vector& result) {
ast_manager& m = result.get_manager();
expr* e1, *e2, *e3;
expr_ref_vector pin(m);
expr_fast_mark1 seen;
for (unsigned i = 0; i < result.size(); ++i) {
if (m.is_and(result.get(i))) {
app* a = to_app(result.get(i));
for (expr* arg : *a) result.push_back(arg);
expr* e = result.get(i);
if (seen.is_marked(e)) {
result[i] = result.back();
result.pop_back();
--i;
continue;
}
seen.mark(e);
pin.push_back(e);
if (m.is_and(e)) {
app* a = to_app(e);
for (expr* arg : *a)
result.push_back(arg);
result[i] = result.back();
result.pop_back();
--i;
}
else if (m.is_not(result.get(i), e1) && m.is_not(e1, e2)) {
else if (m.is_not(e, e1) && m.is_not(e1, e2)) {
result[i] = e2;
--i;
}
else if (m.is_not(result.get(i), e1) && m.is_or(e1)) {
else if (m.is_not(e, e1) && m.is_or(e1)) {
app* a = to_app(e1);
for (expr* arg : *a) result.push_back(m.mk_not(arg));
for (expr* arg : *a)
result.push_back(mk_not(m, arg));
result[i] = result.back();
result.pop_back();
--i;
}
else if (m.is_not(result.get(i), e1) && m.is_implies(e1,e2,e3)) {
else if (m.is_not(e, e1) && m.is_implies(e1, e2, e3)) {
result.push_back(e2);
result[i] = m.mk_not(e3);
result[i] = mk_not(m, e3);
--i;
}
else if (m.is_true(result.get(i)) ||
(m.is_not(result.get(i), e1) &&
else if (m.is_true(e) ||
(m.is_not(e, e1) &&
m.is_false(e1))) {
result[i] = result.back();
result.pop_back();
--i;
}
else if (m.is_false(result.get(i)) ||
(m.is_not(result.get(i), e1) &&
else if (m.is_false(e) ||
(m.is_not(e, e1) &&
m.is_true(e1))) {
result.reset();
result.push_back(m.mk_false());
@ -317,39 +330,52 @@ void flatten_and(expr_ref& fml) {
void flatten_or(expr_ref_vector& result) {
ast_manager& m = result.get_manager();
expr* e1, *e2, *e3;
expr_ref_vector pin(m);
expr_fast_mark1 seen;
for (unsigned i = 0; i < result.size(); ++i) {
if (m.is_or(result.get(i))) {
app* a = to_app(result.get(i));
for (expr* arg : *a) result.push_back(arg);
expr* e = result.get(i);
if (seen.is_marked(e)) {
result[i] = result.back();
result.pop_back();
--i;
continue;
}
seen.mark(e);
pin.push_back(e);
if (m.is_or(e)) {
app* a = to_app(e);
for (expr* arg : *a)
result.push_back(arg);
result[i] = result.back();
result.pop_back();
--i;
}
else if (m.is_not(result.get(i), e1) && m.is_not(e1, e2)) {
else if (m.is_not(e, e1) && m.is_not(e1, e2)) {
result[i] = e2;
--i;
}
else if (m.is_not(result.get(i), e1) && m.is_and(e1)) {
else if (m.is_not(e, e1) && m.is_and(e1)) {
app* a = to_app(e1);
for (expr* arg : *a) result.push_back(m.mk_not(arg));
for (expr* arg : *a)
result.push_back(mk_not(m, arg));
result[i] = result.back();
result.pop_back();
--i;
}
else if (m.is_implies(result.get(i),e2,e3)) {
else if (m.is_implies(e,e2,e3)) {
result.push_back(e3);
result[i] = m.mk_not(e2);
result[i] = mk_not(m, e2);
--i;
}
else if (m.is_false(result.get(i)) ||
(m.is_not(result.get(i), e1) &&
else if (m.is_false(e) ||
(m.is_not(e, e1) &&
m.is_true(e1))) {
result[i] = result.back();
result.pop_back();
--i;
}
else if (m.is_true(result.get(i)) ||
(m.is_not(result.get(i), e1) &&
else if (m.is_true(e) ||
(m.is_not(e, e1) &&
m.is_false(e1))) {
result.reset();
result.push_back(m.mk_true());

View file

@ -580,7 +580,7 @@ namespace datatype {
m_defs.remove(s);
}
bool plugin::is_value_visit(expr * arg, ptr_buffer<app> & todo) const {
bool plugin::is_value_visit(bool unique, expr * arg, ptr_buffer<app> & todo) const {
if (!is_app(arg))
return false;
family_id fid = to_app(arg)->get_family_id();
@ -592,12 +592,13 @@ namespace datatype {
todo.push_back(to_app(arg));
return true;
}
else {
else if (unique)
return m_manager->is_unique_value(arg);
else
return m_manager->is_value(arg);
}
}
bool plugin::is_value(app * e) const {
bool plugin::is_value_aux(bool unique, app * e) const {
TRACE("dt_is_value", tout << "checking\n" << mk_ismt2_pp(e, *m_manager) << "\n";);
if (!u().is_constructor(e))
return false;
@ -608,7 +609,7 @@ namespace datatype {
ptr_buffer<app> todo;
// potentially expensive check for common sub-expressions.
for (expr* arg : *e) {
if (!is_value_visit(arg, todo)) {
if (!is_value_visit(unique, arg, todo)) {
TRACE("dt_is_value", tout << "not-value:\n" << mk_ismt2_pp(arg, *m_manager) << "\n";);
return false;
}
@ -618,7 +619,7 @@ namespace datatype {
SASSERT(u().is_constructor(curr));
todo.pop_back();
for (expr* arg : *curr) {
if (!is_value_visit(arg, todo)) {
if (!is_value_visit(unique, arg, todo)) {
TRACE("dt_is_value", tout << "not-value:\n" << mk_ismt2_pp(arg, *m_manager) << "\n";);
return false;
}

View file

@ -230,9 +230,9 @@ namespace datatype {
bool is_fully_interp(sort * s) const override;
bool is_value(app* e) const override;
bool is_value(app* e) const override { return is_value_aux(false, e); }
bool is_unique_value(app * e) const override { return is_value(e); }
bool is_unique_value(app * e) const override { return is_value_aux(true, e); }
void get_op_names(svector<builtin_name> & op_names, symbol const & logic) override;
@ -257,7 +257,8 @@ namespace datatype {
bool has_nested_arrays() const { return m_has_nested_arrays; }
private:
bool is_value_visit(expr * arg, ptr_buffer<app> & todo) const;
bool is_value_visit(bool unique, expr * arg, ptr_buffer<app> & todo) const;
bool is_value_aux(bool unique, app * arg) const;
func_decl * mk_update_field(
unsigned num_parameters, parameter const * parameters,

View file

@ -25,6 +25,8 @@ namespace euf {
enode* egraph::mk_enode(expr* f, unsigned generation, unsigned num_args, enode * const* args) {
enode* n = enode::mk(m_region, f, generation, num_args, args);
if (m_default_relevant)
n->set_relevant(true);
m_nodes.push_back(n);
m_exprs.push_back(f);
if (is_app(f) && num_args > 0) {
@ -187,10 +189,10 @@ namespace euf {
add_th_diseq(id, v1, v2, n->get_expr());
return;
}
for (auto p : euf::enode_th_vars(r1)) {
for (auto const& p : euf::enode_th_vars(r1)) {
if (!th_propagates_diseqs(p.get_id()))
continue;
for (auto q : euf::enode_th_vars(r2))
for (auto const& q : euf::enode_th_vars(r2))
if (p.get_id() == q.get_id())
add_th_diseq(p.get_id(), p.get_var(), q.get_var(), n->get_expr());
}
@ -269,6 +271,13 @@ namespace euf {
}
}
void egraph::set_relevant(enode* n) {
if (n->is_relevant())
return;
n->set_relevant(true);
m_updates.push_back(update_record(n, update_record::set_relevant()));
}
void egraph::toggle_merge_enabled(enode* n, bool backtracking) {
bool enable_merge = !n->merge_enabled();
n->set_merge_enabled(enable_merge);
@ -380,6 +389,10 @@ namespace euf {
case update_record::tag_t::is_lbl_set:
p.r1->m_lbls.set(p.m_lbls);
break;
case update_record::tag_t::is_set_relevant:
SASSERT(p.r1->is_relevant());
p.r1->set_relevant(false);
break;
case update_record::tag_t::is_update_children:
for (unsigned i = 0; i < p.r1->num_args(); ++i) {
SASSERT(p.r1->m_args[i]->get_root()->m_parents.back() == p.r1);
@ -446,8 +459,8 @@ namespace euf {
r2->inc_class_size(r1->class_size());
merge_th_eq(r1, r2);
reinsert_parents(r1, r2);
if (m_on_merge)
m_on_merge(r2, r1);
for (auto& cb : m_on_merge)
cb(r2, r1);
}
void egraph::remove_parents(enode* r1, enode* r2) {
@ -493,7 +506,7 @@ namespace euf {
void egraph::merge_th_eq(enode* n, enode* root) {
SASSERT(n != root);
for (auto iv : enode_th_vars(n)) {
for (auto const& iv : enode_th_vars(n)) {
theory_id id = iv.get_id();
theory_var v = root->get_th_var(id);
if (v == null_theory_var) {
@ -754,6 +767,8 @@ namespace euf {
}
std::ostream& egraph::display(std::ostream& out, unsigned max_args, enode* n) const {
if (!n->is_relevant())
out << "n";
out << "#" << n->get_expr_id() << " := ";
expr* f = n->get_expr();
if (is_app(f))
@ -770,11 +785,18 @@ namespace euf {
out << " " << p->get_expr_id();
out << "] ";
}
if (n->value() != l_undef)
out << "[b" << n->bool_var() << " := " << (n->value() == l_true ? "T":"F") << (n->merge_tf()?"":" no merge") << "] ";
auto value_of = [&]() {
switch (n->value()) {
case l_true: return "T";
case l_false: return "F";
default: return "?";
}
};
if (n->bool_var() != sat::null_bool_var)
out << "[b" << n->bool_var() << " := " << value_of() << (n->merge_tf() ? "" : " no merge") << "] ";
if (n->has_th_vars()) {
out << "[t";
for (auto v : enode_th_vars(n))
for (auto const& v : enode_th_vars(n))
out << " " << v.get_id() << ":" << v.get_var();
out << "] ";
}

View file

@ -32,6 +32,7 @@ Notes:
#include "ast/euf/euf_enode.h"
#include "ast/euf/euf_etable.h"
#include "ast/ast_ll_pp.h"
#include <vector>
namespace euf {
@ -105,10 +106,11 @@ namespace euf {
struct lbl_hash {};
struct lbl_set {};
struct update_children {};
struct set_relevant {};
enum class tag_t { is_set_parent, is_add_node, is_toggle_merge, is_update_children,
is_add_th_var, is_replace_th_var, is_new_lit, is_new_th_eq,
is_lbl_hash, is_new_th_eq_qhead, is_new_lits_qhead,
is_inconsistent, is_value_assignment, is_lbl_set };
is_inconsistent, is_value_assignment, is_lbl_set, is_set_relevant };
tag_t tag;
enode* r1;
enode* n1;
@ -151,6 +153,8 @@ namespace euf {
tag(tag_t::is_lbl_set), r1(n), n1(nullptr), m_lbls(n->m_lbls.get()) {}
update_record(enode* n, update_children) :
tag(tag_t::is_update_children), r1(n), n1(nullptr), r2_num_parents(UINT_MAX) {}
update_record(enode* n, set_relevant) :
tag(tag_t::is_set_relevant), r1(n), n1(nullptr), r2_num_parents(UINT_MAX) {}
};
ast_manager& m;
svector<to_merge> m_to_merge;
@ -181,7 +185,8 @@ namespace euf {
enode_vector m_todo;
stats m_stats;
bool m_uses_congruence = false;
std::function<void(enode*,enode*)> m_on_merge;
bool m_default_relevant = true;
std::vector<std::function<void(enode*,enode*)>> m_on_merge;
std::function<void(enode*)> m_on_make;
std::function<void(expr*,expr*,expr*)> m_used_eq;
std::function<void(app*,app*)> m_used_cc;
@ -292,8 +297,10 @@ namespace euf {
void set_merge_enabled(enode* n, bool enable_merge);
void set_value(enode* n, lbool value);
void set_bool_var(enode* n, unsigned v) { n->set_bool_var(v); }
void set_relevant(enode* n);
void set_default_relevant(bool b) { m_default_relevant = b; }
void set_on_merge(std::function<void(enode* root,enode* other)>& on_merge) { m_on_merge = on_merge; }
void set_on_merge(std::function<void(enode* root,enode* other)>& on_merge) { m_on_merge.push_back(on_merge); }
void set_on_make(std::function<void(enode* n)>& on_make) { m_on_make = on_make; }
void set_used_eq(std::function<void(expr*,expr*,expr*)>& used_eq) { m_used_eq = used_eq; }
void set_used_cc(std::function<void(app*,app*)>& used_cc) { m_used_cc = used_cc; }

View file

@ -45,10 +45,12 @@ namespace euf {
expr* m_expr = nullptr;
bool m_mark1 = false;
bool m_mark2 = false;
bool m_mark3 = false;
bool m_commutative = false;
bool m_interpreted = false;
bool m_merge_enabled = true;
bool m_is_equality = false; // Does the expression represent an equality
bool m_is_relevant = false;
lbool m_value = l_undef; // Assignment by SAT solver for Boolean node
sat::bool_var m_bool_var = sat::null_bool_var; // SAT solver variable associated with Boolean node
unsigned m_class_size = 1; // Size of the equivalence class if the enode is the root.
@ -145,6 +147,8 @@ namespace euf {
unsigned num_parents() const { return m_parents.size(); }
bool interpreted() const { return m_interpreted; }
bool is_equality() const { return m_is_equality; }
bool is_relevant() const { return m_is_relevant; }
void set_relevant(bool b) { m_is_relevant = b; }
lbool value() const { return m_value; }
bool value_conflict() const { return value() != l_undef && get_root()->value() != l_undef && value() != get_root()->value(); }
sat::bool_var bool_var() const { return m_bool_var; }
@ -170,6 +174,9 @@ namespace euf {
void mark2() { m_mark2 = true; }
void unmark2() { m_mark2 = false; }
bool is_marked2() { return m_mark2; }
void mark3() { m_mark3 = true; }
void unmark3() { m_mark3 = false; }
bool is_marked3() { return m_mark3; }
template<bool m> void mark1_targets() {
enode* n = this;

View file

@ -410,10 +410,7 @@ namespace recfun {
promise_def plugin::ensure_def(symbol const& name, unsigned n, sort *const * params, sort * range, bool is_generated) {
def* d = u().decl_fun(name, n, params, range, is_generated);
def* d2 = nullptr;
if (m_defs.find(d->get_decl(), d2)) {
dealloc(d2);
}
erase_def(d->get_decl());
m_defs.insert(d->get_decl(), d);
return promise_def(&u(), d);
}

View file

@ -305,7 +305,7 @@ namespace recfun {
m_pred(pred), m_cdef(&d), m_args(args) {}
body_expansion(body_expansion const & from):
m_pred(from.m_pred), m_cdef(from.m_cdef), m_args(from.m_args) {}
body_expansion(body_expansion && from) :
body_expansion(body_expansion && from) noexcept :
m_pred(from.m_pred), m_cdef(from.m_cdef), m_args(std::move(from.m_args)) {}
std::ostream& display(std::ostream& out) const;

View file

@ -722,12 +722,12 @@ br_status bool_rewriter::mk_eq_core(expr * lhs, expr * rhs, expr_ref & result) {
return BR_DONE;
}
if (m().is_not(rhs))
if (m().is_not(rhs))
std::swap(lhs, rhs);
if (m().is_not(lhs, lhs)) {
result = m().mk_not(m().mk_eq(lhs, rhs));
return BR_REWRITE2;
if (m().is_not(lhs, lhs)) {
result = m().mk_not(m().mk_eq(lhs, rhs));
return BR_REWRITE2;
}
if (unfolded) {

View file

@ -2818,14 +2818,15 @@ br_status bv_rewriter::mk_bvsmul_no_overflow(unsigned num, expr * const * args,
return BR_DONE;
}
if (is_num1 && is_num2) {
rational mr = a0_val * a1_val;
rational lim = rational::power_of_two(bv_sz-1);
result = m().mk_bool_val(mr < lim);
return BR_DONE;
}
return BR_FAILED;
if (!is_num1 || !is_num2)
return BR_FAILED;
rational lim = rational::power_of_two(bv_sz);
rational r = a0_val * a1_val;
bool sign1 = m_util.has_sign_bit(a0_val, bv_sz);
bool sign2 = m_util.has_sign_bit(a1_val, bv_sz);
result = m().mk_bool_val((sign1 != sign2) || r < lim);
return BR_DONE;
}
br_status bv_rewriter::mk_bvumul_no_overflow(unsigned num, expr * const * args, expr_ref & result) {

View file

@ -46,7 +46,7 @@ class maximize_ac_sharing : public default_rewriter_cfg {
entry(func_decl * d = nullptr, expr * arg1 = nullptr, expr * arg2 = nullptr):m_decl(d), m_arg1(arg1), m_arg2(arg2) {
SASSERT((d == 0 && arg1 == 0 && arg2 == 0) || (d != 0 && arg1 != 0 && arg2 != 0));
if (arg1->get_id() > arg2->get_id())
if (arg1 && arg2 && arg1->get_id() > arg2->get_id())
std::swap(m_arg1, m_arg2);
}

View file

@ -21,8 +21,11 @@ Revision History:
#include "ast/ast_pp.h"
#include "ast/ast_ll_pp.h"
#include "ast/recfun_decl_plugin.h"
#include "ast/rewriter/recfun_replace.h"
#include "ast/rewriter/seq_axioms.h"
namespace seq {
axioms::axioms(th_rewriter& r):
@ -1052,6 +1055,81 @@ namespace seq {
NOT_IMPLEMENTED_YET();
}
// A basic strategy for supporting replace_all and other
// similar functions is to use recursive relations.
// Then the features for recursive functions that allow expanding definitions
// using iterative deepening can be re-used.
//
// create recursive relation 'ra' with properties:
// ra(i, j, s, p, t, r) <- len(s) = i && len(r) = j
// ra(i, j, s, p, t, r) <- len(s) > i = 0 && p = "" && r = t + s
// ra(i, j, s, p, t, r) <- len(s) > i && p != "" && s = extract(s, 0, i) + p + extract(s, i + len(p), len(s)) && r = extract(r, 0, i) + t + extract(r, i + len(p), len(r)) && ra(i + len(p), j + len(t), s, p, t, r)
// ra(i, s, p, t, r) <- ~prefix(p, extract(s, i, len(s)) && at(s,i) = at(r,j) && ra(i + 1, j + 1, s, p, t, r)
// which amounts to:
//
//
// Then assert
// ra(s, p, t, replace_all(s, p, t))
//
void axioms::replace_all_axiom(expr* r) {
expr* s = nullptr, *p = nullptr, *t = nullptr;
VERIFY(seq.str.is_replace_all(r, s, p, t));
recfun::util rec(m);
recfun::decl::plugin& plugin = rec.get_plugin();
recfun_replace replace(m);
sort* srt = s->get_sort();
sort* domain[4] = { srt, srt, srt, srt };
auto d = plugin.ensure_def(symbol("ra"), 4, domain, m.mk_bool_sort(), true);
func_decl* ra = d.get_def()->get_decl();
(void)ra;
sort* isrt = a.mk_int();
var_ref vi(m.mk_var(5, isrt), m);
var_ref vj(m.mk_var(4, isrt), m);
var_ref vs(m.mk_var(3, srt), m);
var_ref vp(m.mk_var(2, srt), m);
var_ref vt(m.mk_var(1, srt), m);
var_ref vr(m.mk_var(0, srt), m);
var* vars[6] = { vi, vj, vs, vp, vt, vr };
(void)vars;
expr_ref len_s(seq.str.mk_length(vs), m);
expr_ref len_r(seq.str.mk_length(vr), m);
expr_ref test1(m.mk_eq(len_s, vi), m);
expr_ref branch1(m.mk_eq(len_r, vj), m);
expr_ref test2(m.mk_and(a.mk_gt(len_s, vi), m.mk_eq(vi, a.mk_int(0)), seq.str.mk_is_empty(vp)), m);
expr_ref branch2(m.mk_eq(vr, seq.str.mk_concat(vt, vs)), m);
NOT_IMPLEMENTED_YET();
#if 0
expr_ref test3(, m);
expr_ref s1(m_sk.mk_prefix_inv(vp, vs), m);
expr_ref r1(m_sk.mk_prefix_inv(vp, vr), m);
expr* args1[4] = { s1, vp, vt, r1 };
expr_ref branch3(m.mk_and(m.mk_eq(seq.str.mk_concat(vp, s1), vs),
m.mk_eq(seq.str.mk_concat(vr, r1), vr),
m.mk_app(ra, 4, args1)
), m);
expr_ref s0(m), r0(m);
m_sk.decompose(vs, s0, s1);
m_sk.decompose(vr, r0, r1);
expr* args2[4] = { s1, vp, vt, r1 };
expr_ref branch4(m.mk_and(m.mk_eq(vs, seq.str.mk_concat(s0, s1)),
m.mk_eq(vr, seq.str.mk_concat(s0, r1)),
m.mk_app(ra, 4, args2)), m);
// s = [s0] + s' && r = [s0] + r' && ra(s', p, t, r')
expr_ref body(m.mk_ite(test1, branch1, m.mk_ite(test2, branch2, m.mk_ite(test3, branch3, branch4))), m);
plugin.set_definition(replace, d, true, 4, vars, body);
expr* args3[4] = { s, p, t, r };
expr_ref lit(m.mk_app(ra, 4, args3), m);
add_clause(lit);
#endif
}
void axioms::replace_re_all_axiom(expr* e) {
expr* s = nullptr, *p = nullptr, *t = nullptr;
VERIFY(seq.str.is_replace_re_all(e, s, p, t));
NOT_IMPLEMENTED_YET();
}
/**

View file

@ -107,6 +107,8 @@ namespace seq {
void length_axiom(expr* n);
void unroll_not_contains(expr* e);
void replace_re_axiom(expr* e);
void replace_all_axiom(expr* e);
void replace_re_all_axiom(expr* e);
expr_ref length_limit(expr* s, unsigned k);
expr_ref is_digit(expr* ch);

File diff suppressed because it is too large Load diff

View file

@ -200,7 +200,7 @@ class seq_rewriter {
expr_ref mk_der_inter(expr* a, expr* b);
expr_ref mk_der_compl(expr* a);
expr_ref mk_der_cond(expr* cond, expr* ele, sort* seq_sort);
expr_ref mk_der_antimorov_union(expr* r1, expr* r2);
expr_ref mk_der_antimirov_union(expr* r1, expr* r2);
bool ite_bdds_compatabile(expr* a, expr* b);
/* if r has the form deriv(en..deriv(e1,to_re(s))..) returns 's = [e1..en]' else returns '() in r'*/
expr_ref is_nullable_symbolic_regex(expr* r, sort* seq_sort);
@ -214,14 +214,19 @@ class seq_rewriter {
expr_ref mk_in_antimirov_rec(expr* s, expr* d);
expr_ref mk_in_antimirov(expr* s, expr* d);
expr_ref mk_antimirov_deriv_intersection(expr* d1, expr* d2, expr* path);
expr_ref mk_antimirov_deriv_intersection(expr* elem, expr* d1, expr* d2, expr* path);
expr_ref mk_antimirov_deriv_concat(expr* d, expr* r);
expr_ref mk_antimirov_deriv_negate(expr* d);
expr_ref mk_antimirov_deriv_negate(expr* elem, expr* d);
expr_ref mk_antimirov_deriv_union(expr* d1, expr* d2);
expr_ref mk_antimirov_deriv_restrict(expr* elem, expr* d1, expr* cond);
expr_ref mk_regex_reverse(expr* r);
expr_ref mk_regex_concat(expr* r1, expr* r2);
expr_ref simplify_path(expr* path);
expr_ref merge_regex_sets(expr* r1, expr* r2, expr* unit, std::function<bool(expr*, expr*&, expr*&)>& decompose, std::function<expr* (expr*, expr*)>& compose);
// elem is (:var 0) and path a condition that may have (:var 0) as a free variable
// simplify path, e.g., (:var 0) = 'a' & (:var 0) = 'b' is simplified to false
expr_ref simplify_path(expr* elem, expr* path);
bool lt_char(expr* ch1, expr* ch2);
bool eq_char(expr* ch1, expr* ch2);
@ -414,5 +419,10 @@ public:
// heuristic elimination of element from condition that comes form a derivative.
// special case optimization for conjunctions of equalities, disequalities and ranges.
void elim_condition(expr* elem, expr_ref& cond);
/* Apply simplifications to the union to keep the union normalized (r1 and r2 are not normalized)*/
expr_ref mk_regex_union_normalize(expr* r1, expr* r2);
/* Apply simplifications to the intersection to keep it normalized (r1 and r2 are not normalized)*/
expr_ref mk_regex_inter_normalize(expr* r1, expr* r2);
};

View file

@ -36,7 +36,8 @@ namespace seq {
symbol m_indexof_left, m_indexof_right; // inverse of indexof: (indexof_left s t) + s + (indexof_right s t) = t, for s in t.
symbol m_aut_step; // regex unfolding state
symbol m_accept; // regex
symbol m_is_empty, m_is_non_empty; // regex emptiness check
symbol m_is_empty; // regex emptiness check
symbol m_is_non_empty;
symbol m_pre, m_post; // inverse of at: (pre s i) + (at s i) + (post s i) = s if 0 <= i < (len s)
symbol m_postp;
symbol m_eq; // equality atom
@ -73,8 +74,8 @@ namespace seq {
}
expr_ref mk_accept(expr_ref_vector const& args) { return expr_ref(seq.mk_skolem(m_accept, args.size(), args.data(), m.mk_bool_sort()), m); }
expr_ref mk_accept(expr* s, expr* i, expr* r) { return mk(m_accept, s, i, r, nullptr, m.mk_bool_sort()); }
expr_ref mk_is_non_empty(expr* r, expr* u, expr* n) { return mk(m_is_non_empty, r, u, n, m.mk_bool_sort(), false); }
expr_ref mk_is_empty(expr* r, expr* u, expr* n) { return mk(m_is_empty, r, u, n, m.mk_bool_sort(), false); }
expr_ref mk_is_non_empty(expr* r, expr* u, expr* n) { return mk(m_is_non_empty, r, u, n, m.mk_bool_sort(), false); }
expr_ref mk_indexof_left(expr* t, expr* s, expr* offset = nullptr) { return mk(m_indexof_left, t, s, offset); }
expr_ref mk_indexof_right(expr* t, expr* s, expr* offset = nullptr) { return mk(m_indexof_right, t, s, offset); }
@ -153,12 +154,12 @@ namespace seq {
bool is_length_limit(expr* e) const { return is_skolem(m_length_limit, e); }
bool is_length_limit(expr* p, unsigned& lim, expr*& s) const;
bool is_is_empty(expr* e) const { return is_skolem(m_is_empty, e); }
bool is_is_non_empty(expr* e) const { return is_skolem(m_is_non_empty, e); }
bool is_is_empty(expr* e, expr*& r, expr*& u, expr*& n) const {
return is_skolem(m_is_empty, e) && (r = to_app(e)->get_arg(0), u = to_app(e)->get_arg(1), n = to_app(e)->get_arg(2), true);
}
bool is_is_non_empty(expr* e, expr*& r, expr*& u, expr*& n) const {
return is_skolem(m_is_non_empty, e) && (r = to_app(e)->get_arg(0), u = to_app(e)->get_arg(1), n = to_app(e)->get_arg(2), true);
bool is_is_non_empty(expr* e) const { return is_skolem(m_is_non_empty, e); }
bool is_is_non_empty(expr* e, expr*& r, expr*& u, expr*& n) const {
return is_skolem(m_is_non_empty, e) && (r = to_app(e)->get_arg(0), u = to_app(e)->get_arg(1), n = to_app(e)->get_arg(2), true);
}
void decompose(expr* e, expr_ref& head, expr_ref& tail);

View file

@ -627,7 +627,7 @@ struct th_rewriter_cfg : public default_rewriter_cfg {
count_down_subterm_references(result, reference_map);
// Any term that was newly introduced by the rewrite step is only referenced within / reachable from the result term.
for (auto kv : reference_map) {
for (auto const& kv : reference_map) {
if (kv.m_value == 0) {
m().trace_stream() << "[attach-enode] #" << kv.m_key->get_id() << " 0\n";
}
@ -691,7 +691,7 @@ struct th_rewriter_cfg : public default_rewriter_cfg {
return;
if (m_new_subst) {
m_rep.reset();
for (auto kv : m_subst->sub())
for (auto const& kv : m_subst->sub())
m_rep.insert(kv.m_key, kv.m_value);
m_new_subst = false;
}
@ -859,8 +859,8 @@ ast_manager & th_rewriter::m() const {
}
void th_rewriter::updt_params(params_ref const & p) {
m_params = p;
m_imp->cfg().updt_params(p);
m_params.append(p);
m_imp->cfg().updt_params(m_params);
}
void th_rewriter::get_param_descrs(param_descrs & r) {

View file

@ -48,6 +48,7 @@ public:
Otherwise, (VAR 0) is stored in the first position, (VAR 1) in the second, and so on.
*/
expr_ref operator()(expr * n, unsigned num_args, expr * const * args);
inline expr_ref operator()(expr* n, expr* arg) { return (*this)(n, 1, &arg); }
inline expr_ref operator()(expr * n, expr_ref_vector const& args) { return (*this)(n, args.size(), args.data()); }
inline expr_ref operator()(expr * n, var_ref_vector const& args) { return (*this)(n, args.size(), (expr*const*)args.data()); }
inline expr_ref operator()(expr * n, app_ref_vector const& args) { return (*this)(n, args.size(), (expr*const*)args.data()); }

View file

@ -243,7 +243,7 @@ void seq_decl_plugin::init() {
m_sigs[OP_RE_OF_PRED] = alloc(psig, m, "re.of.pred", 1, 1, &predA, reA);
m_sigs[OP_RE_REVERSE] = alloc(psig, m, "re.reverse", 1, 1, &reA, reA);
m_sigs[OP_RE_DERIVATIVE] = alloc(psig, m, "re.derivative", 1, 2, AreA, reA);
m_sigs[_OP_RE_ANTIMOROV_UNION] = alloc(psig, m, "re.union", 1, 2, reAreA, reA);
m_sigs[_OP_RE_ANTIMIROV_UNION] = alloc(psig, m, "re.union", 1, 2, reAreA, reA);
m_sigs[OP_SEQ_TO_RE] = alloc(psig, m, "seq.to.re", 1, 1, &seqA, reA);
m_sigs[OP_SEQ_IN_RE] = alloc(psig, m, "seq.in.re", 1, 2, seqAreA, boolT);
m_sigs[OP_SEQ_REPLACE_RE_ALL] = alloc(psig, m, "str.replace_re_all", 1, 3, seqAreAseqA, seqA);
@ -414,9 +414,9 @@ func_decl* seq_decl_plugin::mk_func_decl(decl_kind k, unsigned num_parameters, p
case OP_RE_COMPLEMENT:
case OP_RE_REVERSE:
case OP_RE_DERIVATIVE:
case _OP_RE_ANTIMOROV_UNION:
case _OP_RE_ANTIMIROV_UNION:
m_has_re = true;
// fall-through
Z3_fallthrough;
case OP_SEQ_UNIT:
case OP_STRING_ITOS:
case OP_STRING_STOI:
@ -516,6 +516,7 @@ func_decl* seq_decl_plugin::mk_func_decl(decl_kind k, unsigned num_parameters, p
case OP_SEQ_REPLACE_RE_ALL:
case OP_SEQ_REPLACE_RE:
m_has_re = true;
Z3_fallthrough;
case OP_SEQ_REPLACE_ALL:
return mk_str_fun(k, arity, domain, range, k);
@ -830,6 +831,39 @@ app* seq_util::mk_lt(expr* ch1, expr* ch2) const {
return m.mk_not(mk_le(ch2, ch1));
}
bool seq_util::is_char_const_range(expr const* x, expr* e, unsigned& l, unsigned& u, bool& negated) const {
expr* a, * b, * e0, * e1, * e2, * lb, * ub;
e1 = e;
negated = (m.is_not(e, e1)) ? true : false;
if (m.is_eq(e1, a, b) && (a == x && is_const_char(b, l))) {
u = l;
return true;
}
if (is_char_le(e1, a, b) && a == x && is_const_char(b, u)) {
// (x <= u)
l = 0;
return true;
}
if (is_char_le(e1, a, b) && b == x && is_const_char(a, l)) {
// (l <= x)
u = max_char();
return true;
}
if (m.is_and(e1, e0, e2) && is_char_le(e0, lb, a) && a == x && is_const_char(lb, l) &&
is_char_le(e2, b, ub) && b == x && is_const_char(ub, u))
// (l <= x) & (x <= u)
return true;
if (m.is_eq(e1, a, b) && b == x && is_const_char(a, l)) {
u = l;
return true;
}
if (m.is_and(e1, e0, e2) && is_char_le(e0, a, ub) && a == x && is_const_char(ub, u) &&
is_char_le(e2, lb, b) && b == x && is_const_char(lb, l))
// (x <= u) & (l <= x)
return true;
return false;
}
bool seq_util::str::is_string(func_decl const* f, zstring& s) const {
if (is_string(f)) {
s = f->get_parameter(0).get_zstring();
@ -1029,6 +1063,23 @@ app* seq_util::rex::mk_loop(expr* r, unsigned lo, unsigned hi) {
return m.mk_app(m_fid, OP_RE_LOOP, 2, params, 1, &r);
}
expr* seq_util::rex::mk_loop_proper(expr* r, unsigned lo, unsigned hi)
{
if (lo == 0 && hi == 0) {
sort* seq_sort = nullptr;
VERIFY(u.is_re(r, seq_sort));
// avoid creating a loop with both bounds 0
// such an expression is invalid as a loop
// it is BY DEFINITION = epsilon
return mk_epsilon(seq_sort);
}
if (lo == 1 && hi == 1)
// do not create a loop unless it actually is a loop
return r;
parameter params[2] = { parameter(lo), parameter(hi) };
return m.mk_app(m_fid, OP_RE_LOOP, 2, params, 1, &r);
}
app* seq_util::rex::mk_loop(expr* r, expr* lo) {
expr* rs[2] = { r, lo };
return m.mk_app(m_fid, OP_RE_LOOP, 0, nullptr, 2, rs);
@ -1283,7 +1334,7 @@ std::ostream& seq_util::rex::pp::print(std::ostream& out, expr* e) const {
print(out, r1);
print(out, r2);
}
else if (re.is_antimorov_union(e, r1, r2) || re.is_union(e, r1, r2)) {
else if (re.is_antimirov_union(e, r1, r2) || re.is_union(e, r1, r2)) {
out << "(";
print(out, r1);
out << (html_encode ? "&#x22C3;" : "|");
@ -1502,9 +1553,9 @@ seq_util::rex::info seq_util::rex::mk_info_rec(app* e) const {
if (e->get_family_id() == u.get_family_id()) {
switch (e->get_decl()->get_decl_kind()) {
case OP_RE_EMPTY_SET:
return info(true, true, true, true, true, true, false, l_false, UINT_MAX, 0);
return info(true, l_false, UINT_MAX);
case OP_RE_FULL_SEQ_SET:
return info(true, true, true, true, true, true, false, l_true, 0, 1);
return info(true, l_true, 0);
case OP_RE_STAR:
i1 = get_info_rec(e->get_arg(0));
return i1.star();
@ -1516,7 +1567,7 @@ seq_util::rex::info seq_util::rex::mk_info_rec(app* e) const {
case OP_RE_OF_PRED:
//TBD: check if the character predicate contains uninterpreted symbols or is nonground or is unsat
//TBD: check if the range is unsat
return info(true, true, true, true, true, true, true, l_false, 1, 0);
return info(true, l_false, 1);
case OP_RE_CONCAT:
i1 = get_info_rec(e->get_arg(0));
i2 = get_info_rec(e->get_arg(1));
@ -1533,7 +1584,7 @@ seq_util::rex::info seq_util::rex::mk_info_rec(app* e) const {
min_length = u.str.min_length(e->get_arg(0));
is_value = m.is_value(e->get_arg(0));
nullable = (is_value && min_length == 0 ? l_true : (min_length > 0 ? l_false : l_undef));
return info(true, true, is_value, true, true, true, (min_length == 1 && u.str.max_length(e->get_arg(0)) == 1), nullable, min_length, 0);
return info(is_value, nullable, min_length);
case OP_RE_REVERSE:
return get_info_rec(e->get_arg(0));
case OP_RE_PLUS:
@ -1569,14 +1620,7 @@ std::ostream& seq_util::rex::info::display(std::ostream& out) const {
if (is_known()) {
out << "info("
<< "nullable=" << (nullable == l_true ? "T" : (nullable == l_false ? "F" : "U")) << ", "
<< "classical=" << (classical ? "T" : "F") << ", "
<< "standard=" << (standard ? "T" : "F") << ", "
<< "nonbranching=" << (nonbranching ? "T" : "F") << ", "
<< "normalized=" << (normalized ? "T" : "F") << ", "
<< "monadic=" << (monadic ? "T" : "F") << ", "
<< "singleton=" << (singleton ? "T" : "F") << ", "
<< "min_length=" << min_length << ", "
<< "star_height=" << star_height << ")";
<< "min_length=" << min_length << ")";
}
else if (is_valid())
out << "UNKNOWN";
@ -1596,13 +1640,13 @@ std::string seq_util::rex::info::str() const {
seq_util::rex::info seq_util::rex::info::star() const {
//if is_known() is false then all mentioned properties will remain false
return seq_util::rex::info(classical, classical, interpreted, nonbranching, normalized, monadic, false, l_true, 0, star_height + 1);
return seq_util::rex::info(interpreted, l_true, 0);
}
seq_util::rex::info seq_util::rex::info::plus() const {
if (is_known()) {
//plus never occurs in a normalized regex
return info(classical, classical, interpreted, nonbranching, false, monadic, false, nullable, min_length, star_height + 1);
return info(interpreted, nullable, min_length);
}
else
return *this;
@ -1611,14 +1655,14 @@ seq_util::rex::info seq_util::rex::info::plus() const {
seq_util::rex::info seq_util::rex::info::opt() const {
// if is_known() is false then all mentioned properties will remain false
// optional construct never occurs in a normalized regex
return seq_util::rex::info(classical, classical, interpreted, nonbranching, false, monadic, false, l_true, 0, star_height);
return seq_util::rex::info(interpreted, l_true, 0);
}
seq_util::rex::info seq_util::rex::info::complement() const {
if (is_known()) {
lbool compl_nullable = (nullable == l_true ? l_false : (nullable == l_false ? l_true : l_undef));
unsigned compl_min_length = (compl_nullable == l_false ? 1 : 0);
return info(false, standard, interpreted, nonbranching, normalized, monadic, false, compl_nullable, compl_min_length, star_height);
return info(interpreted, compl_nullable, compl_min_length);
}
else
return *this;
@ -1630,16 +1674,9 @@ seq_util::rex::info seq_util::rex::info::concat(seq_util::rex::info const& rhs,
unsigned m = min_length + rhs.min_length;
if (m < min_length || m < rhs.min_length)
m = UINT_MAX;
return info(classical & rhs.classical,
classical && rhs.classical, //both args of concat must be classical for it to be standard
interpreted && rhs.interpreted,
nonbranching && rhs.nonbranching,
(normalized && !lhs_is_concat && rhs.normalized),
monadic && rhs.monadic,
false,
return info(interpreted && rhs.interpreted,
((nullable == l_false || rhs.nullable == l_false) ? l_false : ((nullable == l_true && rhs.nullable == l_true) ? l_true : l_undef)),
m,
std::max(star_height, rhs.star_height));
m);
}
else
return rhs;
@ -1651,16 +1688,9 @@ seq_util::rex::info seq_util::rex::info::concat(seq_util::rex::info const& rhs,
seq_util::rex::info seq_util::rex::info::disj(seq_util::rex::info const& rhs) const {
if (is_known() || rhs.is_known()) {
//works correctly if one of the arguments is unknown
return info(classical & rhs.classical,
standard && rhs.standard,
interpreted && rhs.interpreted,
nonbranching && rhs.nonbranching,
normalized && rhs.normalized,
monadic && rhs.monadic,
singleton && rhs.singleton,
return info(interpreted && rhs.interpreted,
((nullable == l_true || rhs.nullable == l_true) ? l_true : ((nullable == l_false && rhs.nullable == l_false) ? l_false : l_undef)),
std::min(min_length, rhs.min_length),
std::max(star_height, rhs.star_height));
std::min(min_length, rhs.min_length));
}
else
return rhs;
@ -1669,16 +1699,9 @@ seq_util::rex::info seq_util::rex::info::disj(seq_util::rex::info const& rhs) co
seq_util::rex::info seq_util::rex::info::conj(seq_util::rex::info const& rhs) const {
if (is_known()) {
if (rhs.is_known()) {
return info(false,
standard && rhs.standard,
interpreted && rhs.interpreted,
nonbranching && rhs.nonbranching,
normalized && rhs.normalized,
monadic && rhs.monadic,
singleton && rhs.singleton,
return info(interpreted && rhs.interpreted,
((nullable == l_true && rhs.nullable == l_true) ? l_true : ((nullable == l_false || rhs.nullable == l_false) ? l_false : l_undef)),
std::max(min_length, rhs.min_length),
std::max(star_height, rhs.star_height));
std::max(min_length, rhs.min_length));
}
else
return rhs;
@ -1690,16 +1713,9 @@ seq_util::rex::info seq_util::rex::info::conj(seq_util::rex::info const& rhs) co
seq_util::rex::info seq_util::rex::info::diff(seq_util::rex::info const& rhs) const {
if (is_known()) {
if (rhs.is_known()) {
return info(false,
standard & rhs.standard,
interpreted & rhs.interpreted,
nonbranching & rhs.nonbranching,
normalized & rhs.normalized,
monadic & rhs.monadic,
false,
return info(interpreted & rhs.interpreted,
((nullable == l_true && rhs.nullable == l_false) ? l_true : ((nullable == l_false || rhs.nullable == l_false) ? l_false : l_undef)),
std::max(min_length, rhs.min_length),
std::max(star_height, rhs.star_height));
std::max(min_length, rhs.min_length));
}
else
return rhs;
@ -1714,13 +1730,9 @@ seq_util::rex::info seq_util::rex::info::orelse(seq_util::rex::info const& i) co
// unsigned ite_min_length = std::min(min_length, i.min_length);
// lbool ite_nullable = (nullable == i.nullable ? nullable : l_undef);
// TBD: whether ite is interpreted or not depends on whether the condition is interpreted and both branches are interpreted
return info(false, false, false, false,
normalized && i.normalized,
monadic && i.monadic,
singleton && i.singleton,
return info(false,
((nullable == l_true && i.nullable == l_true) ? l_true : ((nullable == l_false && i.nullable == l_false) ? l_false : l_undef)),
std::min(min_length, i.min_length),
std::max(star_height, i.star_height));
std::min(min_length, i.min_length));
}
else
return i;
@ -1736,24 +1748,22 @@ seq_util::rex::info seq_util::rex::info::loop(unsigned lower, unsigned upper) co
if (m > 0 && (m < min_length || m < lower))
m = UINT_MAX;
lbool loop_nullable = (nullable == l_true || lower == 0 ? l_true : nullable);
if (upper == UINT_MAX) {
// this means the loop is r{lower,*} and is therefore not normalized
// normalized regex would be r{lower,lower}r* and would in particular not use r{0,} for r*
return info(classical, classical, interpreted, nonbranching, false, singleton, false, loop_nullable, m, star_height + 1);
}
else {
bool loop_normalized = normalized;
// r{lower,upper} is not normalized if r is nullable but lower > 0
// r{0,1} is not normalized: it should be ()|r
// r{1,1} is not normalized: it should be r
// r{lower,upper} is not normalized if lower > upper it should then be [] (empty)
if ((nullable == l_true && lower > 0) || upper == 1 || lower > upper)
loop_normalized = false;
return info(classical, classical, interpreted, nonbranching, loop_normalized, singleton, false, loop_nullable, m, star_height);
}
return info(interpreted, loop_nullable, m);
}
else
return *this;
}
seq_util::rex::info& seq_util::rex::info::operator=(info const& other) {
if (this == &other) {
return *this;
}
known = other.known;
interpreted = other.interpreted;
nullable = other.nullable;
min_length = other.min_length;
return *this;
}

View file

@ -103,7 +103,7 @@ enum seq_op_kind {
_OP_REGEXP_EMPTY,
_OP_REGEXP_FULL_CHAR,
_OP_RE_IS_NULLABLE,
_OP_RE_ANTIMOROV_UNION, // Lifted union for antimorov-style derivatives
_OP_RE_ANTIMIROV_UNION, // Lifted union for antimirov-style derivatives
_OP_SEQ_SKOLEM,
LAST_SEQ_OP
};
@ -252,6 +252,12 @@ public:
unsigned max_char() const { return seq.max_char(); }
unsigned num_bits() const { return seq.num_bits(); }
/*
e has a form that is equivalent to l <= x <= u (then negated = false)
or e is equivalent to !(l <= x <= u) (then negated = true)
*/
bool is_char_const_range(expr const* x, expr * e, unsigned& l, unsigned& u, bool& negated) const;
app* mk_skolem(symbol const& name, unsigned n, expr* const* args, sort* range);
bool is_skolem(expr const* e) const { return is_app_of(e, m_fid, _OP_SEQ_SKOLEM); }
@ -411,26 +417,11 @@ public:
struct info {
/* Value is either undefined (known=l_undef) or defined and known (l_true) or defined but unknown (l_false)*/
lbool known { l_undef };
/* No complement, no intersection, no difference, and no if-then-else is used. Reverse is allowed. */
bool classical { false };
/* Boolean-reverse combination of classical regexes (using reverse, union, complement, intersection or difference). */
bool standard { false };
/* There are no uninterpreted symbols. */
bool interpreted { false };
/* No if-then-else is used. */
bool nonbranching { false };
/* Concatenations are right associative and if a loop body is nullable then the lower bound is zero. */
bool normalized { false };
/* All bounded loops have a body that is a singleton. */
bool monadic { false };
/* Positive Boolean combination of ranges or predicates or singleton sequences. */
bool singleton { false };
/* If l_true then empty word is accepted, if l_false then empty word is not accepted. */
lbool nullable { l_undef };
/* Lower bound on the length of all accepted words. */
unsigned min_length { 0 };
/* Maximum nesting depth of Kleene stars. */
unsigned star_height { 0 };
/*
Default constructor of invalid info.
@ -445,19 +436,13 @@ public:
/*
General info constructor.
*/
info(bool is_classical,
bool is_standard,
bool is_interpreted,
bool is_nonbranching,
bool is_normalized,
bool is_monadic,
bool is_singleton,
info(bool is_interpreted,
lbool is_nullable,
unsigned min_l,
unsigned star_h) :
known(l_true), classical(is_classical), standard(is_standard), interpreted(is_interpreted), nonbranching(is_nonbranching),
normalized(is_normalized), monadic(is_monadic), singleton(is_singleton), nullable(is_nullable),
min_length(min_l), star_height(star_h) {}
unsigned min_l) :
known(l_true),
interpreted(is_interpreted),
nullable(is_nullable),
min_length(min_l) {}
/*
Appends a string representation of the info into the stream.
@ -483,6 +468,8 @@ public:
info diff(info const& rhs) const;
info orelse(info const& rhs) const;
info loop(unsigned lower, unsigned upper) const;
info& operator=(info const& other);
};
private:
seq_util& u;
@ -517,6 +504,7 @@ public:
app* mk_opt(expr* r) { return m.mk_app(m_fid, OP_RE_OPTION, r); }
app* mk_loop(expr* r, unsigned lo);
app* mk_loop(expr* r, unsigned lo, unsigned hi);
expr* mk_loop_proper(expr* r, unsigned lo, unsigned hi);
app* mk_loop(expr* r, expr* lo);
app* mk_loop(expr* r, expr* lo, expr* hi);
app* mk_full_char(sort* s);
@ -525,7 +513,7 @@ public:
app* mk_of_pred(expr* p);
app* mk_reverse(expr* r) { return m.mk_app(m_fid, OP_RE_REVERSE, r); }
app* mk_derivative(expr* ele, expr* r) { return m.mk_app(m_fid, OP_RE_DERIVATIVE, ele, r); }
app* mk_antimorov_union(expr* r1, expr* r2) { return m.mk_app(m_fid, _OP_RE_ANTIMOROV_UNION, r1, r2); }
app* mk_antimirov_union(expr* r1, expr* r2) { return m.mk_app(m_fid, _OP_RE_ANTIMIROV_UNION, r1, r2); }
bool is_to_re(expr const* n) const { return is_app_of(n, m_fid, OP_SEQ_TO_RE); }
bool is_concat(expr const* n) const { return is_app_of(n, m_fid, OP_RE_CONCAT); }
@ -557,7 +545,7 @@ public:
bool is_of_pred(expr const* n) const { return is_app_of(n, m_fid, OP_RE_OF_PRED); }
bool is_reverse(expr const* n) const { return is_app_of(n, m_fid, OP_RE_REVERSE); }
bool is_derivative(expr const* n) const { return is_app_of(n, m_fid, OP_RE_DERIVATIVE); }
bool is_antimorov_union(expr const* n) const { return is_app_of(n, m_fid, _OP_RE_ANTIMOROV_UNION); }
bool is_antimirov_union(expr const* n) const { return is_app_of(n, m_fid, _OP_RE_ANTIMIROV_UNION); }
MATCH_UNARY(is_to_re);
MATCH_BINARY(is_concat);
MATCH_BINARY(is_union);
@ -571,7 +559,7 @@ public:
MATCH_UNARY(is_of_pred);
MATCH_UNARY(is_reverse);
MATCH_BINARY(is_derivative);
MATCH_BINARY(is_antimorov_union);
MATCH_BINARY(is_antimirov_union);
bool is_loop(expr const* n, expr*& body, unsigned& lo, unsigned& hi) const;
bool is_loop(expr const* n, expr*& body, unsigned& lo) const;
bool is_loop(expr const* n, expr*& body, expr*& lo, expr*& hi) const;

View file

@ -18,6 +18,7 @@ Revision History:
--*/
#include "ast/static_features.h"
#include "ast/ast_pp.h"
#include "ast/ast_ll_pp.h"
#include "ast/for_each_expr.h"
static_features::static_features(ast_manager & m):
@ -39,7 +40,8 @@ static_features::static_features(ast_manager & m):
}
void static_features::reset() {
m_already_visited .reset();
m_pre_processed .reset();
m_post_processed.reset();
m_cnf = true;
m_num_exprs = 0;
m_num_roots = 0;
@ -343,8 +345,8 @@ void static_features::update_core(expr * e) {
}
func_decl * d = to_app(e)->get_decl();
inc_num_apps(d);
if (d->get_arity() > 0 && !is_marked(d)) {
mark(d);
if (d->get_arity() > 0 && !is_marked_pre(d)) {
mark_pre(d);
if (fid == null_family_id)
m_num_uninterpreted_functions++;
}
@ -396,73 +398,68 @@ void static_features::update_core(sort * s) {
check_array(s);
}
void static_features::process(expr * e, bool form_ctx, bool or_and_ctx, bool ite_ctx, unsigned stack_depth) {
TRACE("static_features", tout << "processing\n" << mk_pp(e, m) << "\n";);
if (is_var(e))
return;
if (is_marked(e)) {
m_num_sharing++;
return;
}
bool static_features::pre_process(expr * e, bool form_ctx, bool or_and_ctx, bool ite_ctx) {
if (is_marked_post(e))
return true;
if (stack_depth > m_max_stack_depth) {
ptr_vector<expr> todo;
todo.push_back(e);
for (unsigned i = 0; i < todo.size(); ++i) {
e = todo[i];
if (is_marked(e))
continue;
if (is_basic_expr(e)) {
mark(e);
todo.append(to_app(e)->get_num_args(), to_app(e)->get_args());
}
else {
process(e, form_ctx, or_and_ctx, ite_ctx, stack_depth - 10);
}
}
return;
if (is_marked_pre(e))
return true;
if (is_var(e)) {
mark_pre(e);
mark_post(e);
return true;
}
mark(e);
mark_pre(e);
update_core(e);
if (is_quantifier(e)) {
expr * body = to_quantifier(e)->get_expr();
process(body, false, false, false, stack_depth+1);
if (is_marked_post(body))
return true;
add_process(body, false, false, false);
return false;
}
auto [form_ctx_new, or_and_ctx_new, ite_ctx_new] = new_ctx(e);
bool all_processed = true;
for (expr* arg : *to_app(e)) {
m.is_not(arg, arg);
if (is_marked_post(arg))
++m_num_sharing;
else {
add_process(arg, form_ctx_new, or_and_ctx_new, ite_ctx_new);
all_processed = false;
}
}
return all_processed;
}
void static_features::post_process(expr * e, bool form_ctx, bool or_and_ctx, bool ite_ctx) {
if (is_marked_post(e))
return;
mark_post(e);
if (is_quantifier(e)) {
expr * body = to_quantifier(e)->get_expr();
set_depth(e, get_depth(body)+1);
return;
}
bool form_ctx_new = false;
bool or_and_ctx_new = false;
bool ite_ctx_new = false;
if (is_basic_expr(e)) {
switch (to_app(e)->get_decl_kind()) {
case OP_ITE:
form_ctx_new = m.is_bool(e);
ite_ctx_new = true;
break;
case OP_AND:
case OP_OR:
form_ctx_new = true;
or_and_ctx_new = true;
break;
case OP_EQ:
form_ctx_new = true;
break;
}
}
unsigned depth = 0;
unsigned form_depth = 0;
unsigned or_and_depth = 0;
unsigned ite_depth = 0;
auto [form_ctx_new, or_and_ctx_new, ite_ctx_new] = new_ctx(e);
for (expr* arg : *to_app(e)) {
m.is_not(arg, arg);
process(arg, form_ctx_new, or_and_ctx_new, ite_ctx_new, stack_depth+1);
SASSERT(is_marked_post(arg));
depth = std::max(depth, get_depth(arg));
if (form_ctx_new)
form_depth = std::max(form_depth, get_form_depth(arg));
@ -507,16 +504,58 @@ void static_features::process(expr * e, bool form_ctx, bool or_and_ctx, bool ite
}
set_ite_depth(e, ite_depth);
}
}
std::tuple<bool, bool, bool> static_features::new_ctx(expr* e) {
bool form_ctx_new = false;
bool or_and_ctx_new = false;
bool ite_ctx_new = false;
if (is_basic_expr(e)) {
switch (to_app(e)->get_decl_kind()) {
case OP_ITE:
form_ctx_new = m.is_bool(e);
ite_ctx_new = true;
break;
case OP_AND:
case OP_OR:
form_ctx_new = true;
or_and_ctx_new = true;
break;
case OP_EQ:
form_ctx_new = true;
break;
}
}
return std::tuple(form_ctx_new, or_and_ctx_new, ite_ctx_new);
}
void static_features::process_all() {
while (!m_to_process.empty()) {
auto const& [e, form, or_and, ite] = m_to_process.back();
if (is_marked_post(e)) {
m_to_process.pop_back();
++m_num_sharing;
continue;
}
if (!pre_process(e, form, or_and, ite))
continue;
post_process(e, form, or_and, ite);
m_to_process.pop_back();
}
}
void static_features::process_root(expr * e) {
if (is_marked(e)) {
if (is_marked_post(e)) {
m_num_sharing++;
return;
}
m_num_roots++;
if (m.is_or(e)) {
mark(e);
mark_post(e);
m_num_clauses++;
m_num_bool_exprs++;
unsigned num_args = to_app(e)->get_num_args();
@ -530,7 +569,8 @@ void static_features::process_root(expr * e) {
expr * arg = to_app(e)->get_arg(i);
if (m.is_not(arg))
arg = to_app(arg)->get_arg(0);
process(arg, true, true, false, 0);
add_process(arg, true, true, false);
process_all();
depth = std::max(depth, get_depth(arg));
form_depth = std::max(form_depth, get_form_depth(arg));
or_and_depth = std::max(or_and_depth, get_or_and_depth(arg));
@ -558,7 +598,8 @@ void static_features::process_root(expr * e) {
m_num_units++;
m_num_clauses++;
}
process(e, false, false, false, 0);
add_process(e, false, false, false);
process_all();
}
void static_features::collect(unsigned num_formulas, expr * const * formulas) {

View file

@ -28,6 +28,12 @@ Revision History:
#include "util/map.h"
struct static_features {
struct to_process {
expr* e;
bool form_ctx;
bool and_or_ctx;
bool ite_ctx;
};
ast_manager & m;
arith_util m_autil;
bv_util m_bvutil;
@ -39,7 +45,7 @@ struct static_features {
family_id m_lfid;
family_id m_arrfid;
family_id m_srfid;
ast_mark m_already_visited;
ast_mark m_pre_processed, m_post_processed;
bool m_cnf;
unsigned m_num_exprs; //
unsigned m_num_roots; //
@ -111,14 +117,17 @@ struct static_features {
u_map<unsigned> m_expr2formula_depth;
unsigned m_num_theories;
bool_vector m_theories; // mapping family_id -> bool
bool_vector m_theories; // mapping family_id -> bool
symbol m_label_sym;
symbol m_pattern_sym;
symbol m_expr_list_sym;
svector<to_process> m_to_process;
bool is_marked(ast * e) const { return m_already_visited.is_marked(e); }
void mark(ast * e) { m_already_visited.mark(e, true); }
bool is_marked_pre(ast * e) const { return m_pre_processed.is_marked(e); }
void mark_pre(ast * e) { m_pre_processed.mark(e, true); }
bool is_marked_post(ast * e) const { return m_post_processed.is_marked(e); }
void mark_post(ast * e) { m_post_processed.mark(e, true); }
bool is_bool(expr const * e) const { return m.is_bool(e); }
bool is_basic_expr(expr const * e) const { return is_app(e) && to_app(e)->get_family_id() == m_bfid; }
bool is_arith_expr(expr const * e) const { return is_app(e) && to_app(e)->get_family_id() == m_afid; }
@ -161,7 +170,12 @@ struct static_features {
void inc_num_aliens(family_id fid) { m_num_aliens_per_family.reserve(fid+1, 0); m_num_aliens_per_family[fid]++; }
void update_core(expr * e);
void update_core(sort * s);
void process(expr * e, bool form_ctx, bool or_and_ctx, bool ite_ctx, unsigned stack_depth);
// void process(expr * e, bool form_ctx, bool or_and_ctx, bool ite_ctx, unsigned stack_depth);
bool pre_process(expr * e, bool form_ctx, bool or_and_ctx, bool ite_ctx);
void post_process(expr * e, bool form_ctx, bool or_and_ctx, bool ite_ctx);
void add_process(expr * e, bool form_ctx, bool or_and_ctx, bool ite_ctx) { m_to_process.push_back({e, form_ctx, or_and_ctx, ite_ctx}); }
void process_all();
std::tuple<bool, bool, bool> new_ctx(expr* );
void process_root(expr * e);
unsigned get_depth(expr const * e) const { return m_expr2depth.get(e->get_id(), 1); }
void set_depth(expr const * e, unsigned d) { m_expr2depth.setx(e->get_id(), d, 1); }