mirror of
https://github.com/Z3Prover/z3
synced 2025-04-28 03:15:50 +00:00
Merge branch 'master' into polysat
This commit is contained in:
commit
f54f33551e
308 changed files with 6606 additions and 18485 deletions
|
@ -103,6 +103,7 @@ public:
|
|||
new_f = conj;
|
||||
g->assert_expr(new_f);
|
||||
}
|
||||
g->elim_true();
|
||||
}
|
||||
|
||||
void operator()(goal_ref const & g, goal_ref_buffer & result) override {
|
||||
|
|
|
@ -1,42 +0,0 @@
|
|||
/*++
|
||||
Copyright (c) 2023 Microsoft Corporation
|
||||
|
||||
Module Name:
|
||||
|
||||
bound_simplifier_tactic.h
|
||||
|
||||
Author:
|
||||
|
||||
Nikolaj Bjorner (nbjorner) 2023-01-22
|
||||
|
||||
Tactic Documentation:
|
||||
|
||||
## Tactic bound-simplifier
|
||||
|
||||
### Short Description
|
||||
|
||||
Tactic for simplifying arithmetical expressions modulo bounds
|
||||
|
||||
### Long Description
|
||||
|
||||
The tactic is used to eliminate occurrences of modulus expressions when it is known that terms are within the bounds
|
||||
of the modulus.
|
||||
|
||||
|
||||
--*/
|
||||
#pragma once
|
||||
|
||||
#include "util/params.h"
|
||||
#include "tactic/tactic.h"
|
||||
#include "tactic/dependent_expr_state_tactic.h"
|
||||
#include "ast/simplifiers/bound_simplifier.h"
|
||||
|
||||
inline tactic* mk_bound_simplifier_tactic(ast_manager& m, params_ref const& p = params_ref()) {
|
||||
return alloc(dependent_expr_state_tactic, m, p,
|
||||
[](auto& m, auto& p, auto &s) -> dependent_expr_simplifier* { return alloc(bound_simplifier, m, p, s); });
|
||||
}
|
||||
|
||||
/*
|
||||
ADD_TACTIC("bound-simplifier", "Simplify arithmetical expressions modulo bounds.", "mk_bound_simplifier_tactic(m, p)")
|
||||
ADD_SIMPLIFIER("bound-simplifier", "Simplify arithmetical expressions modulo bounds.", "alloc(bound_simplifier, m, p, s)")
|
||||
*/
|
|
@ -24,6 +24,7 @@ Notes:
|
|||
#include "ast/recfun_decl_plugin.h"
|
||||
#include "ast/array_decl_plugin.h"
|
||||
#include "ast/datatype_decl_plugin.h"
|
||||
#include "ast/seq_decl_plugin.h"
|
||||
#include "tactic/core/collect_occs.h"
|
||||
#include "ast/ast_smt2_pp.h"
|
||||
#include "ast/ast_ll_pp.h"
|
||||
|
@ -44,6 +45,7 @@ class elim_uncnstr_tactic : public tactic {
|
|||
bv_util m_bv_util;
|
||||
array_util m_ar_util;
|
||||
datatype_util m_dt_util;
|
||||
seq_util m_seq_util;
|
||||
app_ref_vector m_fresh_vars;
|
||||
obj_map<app, app*> m_cache;
|
||||
app_ref_vector m_cache_domain;
|
||||
|
@ -60,6 +62,7 @@ class elim_uncnstr_tactic : public tactic {
|
|||
m_bv_util(m),
|
||||
m_ar_util(m),
|
||||
m_dt_util(m),
|
||||
m_seq_util(m),
|
||||
m_fresh_vars(m),
|
||||
m_cache_domain(m),
|
||||
m_max_memory(max_memory),
|
||||
|
@ -792,6 +795,44 @@ class elim_uncnstr_tactic : public tactic {
|
|||
}
|
||||
return nullptr;
|
||||
}
|
||||
|
||||
// x ++ y -> z, x -> z, y -> eps
|
||||
app * process_seq_app(func_decl * f, unsigned num, expr * const * args) {
|
||||
switch (f->get_decl_kind()) {
|
||||
case _OP_STRING_CONCAT:
|
||||
case OP_SEQ_CONCAT: {
|
||||
app * r = nullptr;
|
||||
expr* x, *y;
|
||||
if (uncnstr(args[0]) && num == 2 &&
|
||||
args[1]->get_ref_count() == 1 &&
|
||||
m_seq_util.str.is_concat(args[1], x, y) &&
|
||||
uncnstr(x)) {
|
||||
if (!mk_fresh_uncnstr_var_for(f, num, args, r))
|
||||
return r;
|
||||
|
||||
if (m_mc) {
|
||||
add_def(args[0], r);
|
||||
add_def(x, m_seq_util.str.mk_empty(args[0]->get_sort()));
|
||||
}
|
||||
r = m_seq_util.str.mk_concat(r, y);
|
||||
return r;
|
||||
|
||||
}
|
||||
if (!uncnstr(num, args))
|
||||
return nullptr;
|
||||
if (!mk_fresh_uncnstr_var_for(f, num, args, r))
|
||||
return r;
|
||||
|
||||
expr_ref id(m_seq_util.str.mk_empty(args[0]->get_sort()), m());
|
||||
add_defs(num, args, r, id);
|
||||
|
||||
return r;
|
||||
}
|
||||
default:
|
||||
return nullptr;
|
||||
}
|
||||
}
|
||||
|
||||
|
||||
br_status reduce_app(func_decl * f, unsigned num, expr * const * args, expr_ref & result, proof_ref & result_pr) {
|
||||
if (num == 0)
|
||||
|
@ -817,7 +858,8 @@ class elim_uncnstr_tactic : public tactic {
|
|||
u = process_array_app(f, num, args);
|
||||
else if (fid == m_dt_util.get_family_id())
|
||||
u = process_datatype_app(f, num, args);
|
||||
|
||||
else if (fid == m_seq_util.get_family_id())
|
||||
u = process_seq_app(f, num, args);
|
||||
if (u == nullptr)
|
||||
return BR_FAILED;
|
||||
|
||||
|
|
|
@ -30,7 +30,7 @@ resolution.
|
|||
the predicate `p` occurs once positively. All negative occurrences of `p` are resolved against this positive occurrence.
|
||||
The result of resolution is a set of equalities between arguments to `p`. The function `f` is replaced by a partial solution.
|
||||
|
||||
```
|
||||
```z3
|
||||
(declare-fun f (Int Int Int) Int)
|
||||
(declare-fun p (Int) Bool)
|
||||
(declare-const a Int)
|
||||
|
|
|
@ -47,13 +47,13 @@ public:
|
|||
};
|
||||
|
||||
class ac_rewriter {
|
||||
ast_manager& m_manager;
|
||||
ast_manager& m;
|
||||
public:
|
||||
ac_rewriter(ast_manager& m): m_manager(m) {}
|
||||
ac_rewriter(ast_manager& m): m(m) {}
|
||||
|
||||
br_status mk_app_core(func_decl * f, unsigned num_args, expr * const * args, expr_ref & result) {
|
||||
if ((f->is_associative() && f->is_commutative()) ||
|
||||
m_manager.is_distinct(f)) {
|
||||
m.is_distinct(f)) {
|
||||
ptr_buffer<expr> buffer;
|
||||
buffer.append(num_args, args);
|
||||
std::sort(buffer.begin(), buffer.end(), ast_lt_proc());
|
||||
|
@ -62,13 +62,13 @@ public:
|
|||
change = (args[i] != buffer[i]);
|
||||
}
|
||||
if (change) {
|
||||
result = m().mk_app(f, num_args, buffer.begin());
|
||||
result = m.mk_app(f, num_args, buffer.begin());
|
||||
return BR_DONE;
|
||||
}
|
||||
}
|
||||
else if (f->is_commutative() && num_args == 2 && args[0]->get_id() > args[1]->get_id()) {
|
||||
expr* args2[2] = { args[1], args[0] };
|
||||
result = m().mk_app(f, num_args, args2);
|
||||
result = m.mk_app(f, num_args, args2);
|
||||
return BR_DONE;
|
||||
}
|
||||
return BR_FAILED;
|
||||
|
@ -76,10 +76,8 @@ public:
|
|||
|
||||
void mk_app(func_decl * f, unsigned num_args, expr * const * args, expr_ref & result) {
|
||||
if (mk_app_core(f, num_args, args, result) == BR_FAILED)
|
||||
result = m().mk_app(f, num_args, args);
|
||||
result = m.mk_app(f, num_args, args);
|
||||
}
|
||||
private:
|
||||
ast_manager& m() const { return m_manager; }
|
||||
};
|
||||
|
||||
|
||||
|
@ -110,13 +108,12 @@ class symmetry_reduce_tactic::imp {
|
|||
typedef ptr_vector<app> term_set;
|
||||
typedef obj_map<app, unsigned> app_map;
|
||||
typedef u_map<ptr_vector<app> > inv_app_map;
|
||||
ast_manager& m_manager;
|
||||
ast_manager& m;
|
||||
ac_rewriter_star m_rewriter;
|
||||
scoped_ptr<expr_replacer> m_replace;
|
||||
|
||||
ast_manager& m() const { return m_manager; }
|
||||
public:
|
||||
imp(ast_manager& m) : m_manager(m), m_rewriter(m) {
|
||||
imp(ast_manager& m) : m(m), m_rewriter(m) {
|
||||
m_replace = mk_default_expr_replacer(m, false);
|
||||
}
|
||||
|
||||
|
@ -124,10 +121,10 @@ public:
|
|||
|
||||
void operator()(goal & g) {
|
||||
if (g.inconsistent())
|
||||
return;
|
||||
return;
|
||||
tactic_report report("symmetry-reduce", g);
|
||||
vector<ptr_vector<app> > P;
|
||||
expr_ref fml(m());
|
||||
expr_ref fml(m);
|
||||
to_formula(g, fml);
|
||||
app_map occs;
|
||||
compute_occurrences(fml, occs);
|
||||
|
@ -149,11 +146,13 @@ public:
|
|||
app* c = select_const(consts, cts);
|
||||
if (!c) break;
|
||||
cts.push_back(c);
|
||||
expr* mem = mk_member(t, cts);
|
||||
expr_ref mem = mk_member(t, cts);
|
||||
g.assert_expr(mem);
|
||||
num_sym_break_preds++;
|
||||
TRACE("symmetry_reduce", tout << "member predicate: " << mk_pp(mem, m()) << "\n";);
|
||||
fml = m().mk_and(fml.get(), mem);
|
||||
|
||||
TRACE("symmetry_reduce", tout << "member predicate: " << mem << "\n";);
|
||||
|
||||
fml = m.mk_and(fml.get(), mem);
|
||||
normalize(fml);
|
||||
}
|
||||
}
|
||||
|
@ -167,7 +166,7 @@ private:
|
|||
for (unsigned i = 0; i < g.size(); ++i) {
|
||||
conjs.push_back(g.form(i));
|
||||
}
|
||||
fml = m().mk_and(conjs.size(), conjs.data());
|
||||
fml = m.mk_and(conjs.size(), conjs.data());
|
||||
normalize(fml);
|
||||
}
|
||||
|
||||
|
@ -184,28 +183,17 @@ private:
|
|||
// compute_siblings(fml, coloring);
|
||||
compute_inv_app(coloring, inv_color);
|
||||
|
||||
inv_app_map::iterator it = inv_color.begin(), end = inv_color.end();
|
||||
for (; it != end; ++it) {
|
||||
if (it->m_value.size() < 2) {
|
||||
for (auto const& [k, v] : inv_color) {
|
||||
if (v.size() < 2)
|
||||
continue;
|
||||
}
|
||||
VERIFY(occs.find(it->m_value[0], num_occs));
|
||||
if (num_occs < 2) {
|
||||
VERIFY(occs.find(v[0], num_occs));
|
||||
if (num_occs < 2)
|
||||
continue;
|
||||
}
|
||||
bool is_const = true;
|
||||
for (unsigned j = 0; is_const && j < it->m_value.size(); ++j) {
|
||||
is_const = it->m_value[j]->get_num_args() == 0;
|
||||
}
|
||||
if (!is_const) {
|
||||
bool is_const = all_of(v, [&](app* a) { return a->get_num_args() == 0; });
|
||||
if (!is_const)
|
||||
continue;
|
||||
}
|
||||
P.push_back(it->m_value);
|
||||
TRACE("symmetry_reduce",
|
||||
for (unsigned i = 0; i < it->m_value.size(); ++i) {
|
||||
tout << mk_pp(it->m_value[i], m()) << " ";
|
||||
}
|
||||
tout << "\n";);
|
||||
P.push_back(v);
|
||||
TRACE("symmetry_reduce", for (app * a : v) tout << mk_pp(a, m) << " "; tout << "\n";);
|
||||
}
|
||||
}
|
||||
|
||||
|
@ -426,14 +414,14 @@ private:
|
|||
tout << "Not symmetric: ";
|
||||
}
|
||||
for (unsigned i = 0; i < p.size(); ++i) {
|
||||
tout << mk_pp(p[i], m()) << " ";
|
||||
tout << mk_pp(p[i], m) << " ";
|
||||
}
|
||||
tout << "\n";);
|
||||
return result;
|
||||
}
|
||||
|
||||
bool check_swap(expr* fml, app* t1, app* t2) {
|
||||
expr_substitution sub(m());
|
||||
expr_substitution sub(m);
|
||||
sub.insert(t1, t2);
|
||||
sub.insert(t2, t1);
|
||||
m_replace->set_substitution(&sub);
|
||||
|
@ -441,7 +429,7 @@ private:
|
|||
}
|
||||
|
||||
bool check_cycle(expr* fml, permutation& p) {
|
||||
expr_substitution sub(m());
|
||||
expr_substitution sub(m);
|
||||
for (unsigned i = 0; i + 1 < p.size(); ++i) {
|
||||
sub.insert(p[i], p[i+1]);
|
||||
}
|
||||
|
@ -451,15 +439,15 @@ private:
|
|||
}
|
||||
|
||||
bool check_substitution(expr* t) {
|
||||
expr_ref r(m());
|
||||
expr_ref r(m);
|
||||
(*m_replace)(t, r);
|
||||
normalize(r);
|
||||
return t == r.get();
|
||||
}
|
||||
|
||||
void normalize(expr_ref& r) {
|
||||
proof_ref pr(m());
|
||||
expr_ref result(m());
|
||||
proof_ref pr(m);
|
||||
expr_ref result(m);
|
||||
m_rewriter(r.get(), result, pr);
|
||||
r = result;
|
||||
}
|
||||
|
@ -473,7 +461,7 @@ private:
|
|||
while (!todo.empty()) {
|
||||
fml = todo.back();
|
||||
todo.pop_back();
|
||||
if (m().is_and(fml)) {
|
||||
if (m.is_and(fml)) {
|
||||
todo.append(to_app(fml)->get_num_args(), to_app(fml)->get_args());
|
||||
}
|
||||
else if (is_range_restriction(fml, p, t)) {
|
||||
|
@ -482,13 +470,13 @@ private:
|
|||
}
|
||||
}
|
||||
bool is_range_restriction(expr* form, term_set const& C, app*& t) {
|
||||
if (!m().is_or(form)) return false;
|
||||
if (!m.is_or(form)) return false;
|
||||
unsigned sz = to_app(form)->get_num_args();
|
||||
t = nullptr;
|
||||
for (unsigned i = 0; i < sz; ++i) {
|
||||
expr* e = to_app(form)->get_arg(i);
|
||||
expr* e1, *e2;
|
||||
if (!m().is_eq(e, e1, e2)) return false;
|
||||
if (!m.is_eq(e, e1, e2)) return false;
|
||||
if (!is_app(e1) || !is_app(e2)) return false;
|
||||
app* a1 = to_app(e1), *a2 = to_app(e2);
|
||||
if (C.contains(a1) && (t == nullptr || t == a2)) {
|
||||
|
@ -515,13 +503,9 @@ private:
|
|||
num_occurrences(app_map& occs): m_occs(occs) {}
|
||||
void operator()(app* n) {
|
||||
m_occs.insert_if_not_there(n, 0);
|
||||
unsigned sz = n->get_num_args();
|
||||
for (unsigned i = 0; i < sz; ++i) {
|
||||
expr* arg = n->get_arg(i);
|
||||
if (is_app(arg)) {
|
||||
m_occs.insert_if_not_there(to_app(arg), 0)++;
|
||||
}
|
||||
}
|
||||
for (expr* arg : *n)
|
||||
if (is_app(arg))
|
||||
m_occs.insert_if_not_there(to_app(arg), 0)++;
|
||||
}
|
||||
void operator()(quantifier * n) {}
|
||||
void operator()(var * n) {}
|
||||
|
@ -540,7 +524,7 @@ private:
|
|||
unsigned weight = 0, weight1 = 0;
|
||||
VERIFY(occs.find(t, weight));
|
||||
unsigned cts_delta = compute_cts_delta(t, cts, consts);
|
||||
TRACE("symmetry_reduce", tout << mk_pp(t, m()) << " " << weight << " " << cts_delta << "\n";);
|
||||
TRACE("symmetry_reduce", tout << mk_pp(t, m) << " " << weight << " " << cts_delta << "\n";);
|
||||
for (unsigned i = 1; i < T.size(); ++i) {
|
||||
app* t1 = T[i];
|
||||
VERIFY(occs.find(t1, weight1));
|
||||
|
@ -548,7 +532,7 @@ private:
|
|||
continue;
|
||||
}
|
||||
unsigned cts_delta1 = compute_cts_delta(t1, cts, consts);
|
||||
TRACE("symmetry_reduce", tout << mk_pp(t1, m()) << " " << weight1 << " " << cts_delta1 << "\n";);
|
||||
TRACE("symmetry_reduce", tout << mk_pp(t1, m) << " " << weight1 << " " << cts_delta1 << "\n";);
|
||||
if ((t->get_num_args() == t1->get_num_args() && (weight1 > weight || cts_delta1 < cts_delta)) ||
|
||||
t->get_num_args() > t1->get_num_args()) {
|
||||
cts_delta = cts_delta1;
|
||||
|
@ -577,15 +561,15 @@ private:
|
|||
member_of mem(consts, cts);
|
||||
for_each_expr(mem, t);
|
||||
TRACE("symmetry_reduce",
|
||||
tout << "Term: " << mk_pp(t, m()) << "\n";
|
||||
tout << "Term: " << mk_pp(t, m) << "\n";
|
||||
tout << "Support set: ";
|
||||
for (unsigned i = 0; i < consts.size(); ++i) {
|
||||
tout << mk_pp(consts[i], m()) << " ";
|
||||
tout << mk_pp(consts[i], m) << " ";
|
||||
}
|
||||
tout << "\n";
|
||||
tout << "Constants: ";
|
||||
for (unsigned i = 0; i < cts.size(); ++i) {
|
||||
tout << mk_pp(cts[i], m()) << " ";
|
||||
tout << mk_pp(cts[i], m) << " ";
|
||||
}
|
||||
tout << "\n";
|
||||
);
|
||||
|
@ -606,15 +590,14 @@ private:
|
|||
app* select_const(term_set const& A, term_set const& B) {
|
||||
unsigned j;
|
||||
for (j = 0; j < A.size() && B.contains(A[j]); ++j);
|
||||
return (j == A.size())?0:A[j];
|
||||
return (j == A.size())? nullptr:A[j];
|
||||
}
|
||||
|
||||
expr* mk_member(app* t, term_set const& C) {
|
||||
expr_ref_vector eqs(m());
|
||||
for (unsigned i = 0; i < C.size(); ++i) {
|
||||
eqs.push_back(m().mk_eq(t, C[i]));
|
||||
}
|
||||
return mk_or(m(), eqs.size(), eqs.data());
|
||||
expr_ref mk_member(app* t, term_set const& C) {
|
||||
expr_ref_vector eqs(m);
|
||||
for (expr* e : C)
|
||||
eqs.push_back(m.mk_eq(t, e));
|
||||
return mk_or(eqs);
|
||||
}
|
||||
};
|
||||
|
||||
|
|
|
@ -141,6 +141,7 @@ class tseitin_cnf_tactic : public tactic {
|
|||
sign = !sign;
|
||||
goto start;
|
||||
case OP_OR:
|
||||
// case OP_AND:
|
||||
l = nullptr;
|
||||
m_cache.find(to_app(n), l);
|
||||
SASSERT(l != 0);
|
||||
|
@ -187,6 +188,7 @@ class tseitin_cnf_tactic : public tactic {
|
|||
goto start;
|
||||
}
|
||||
case OP_OR:
|
||||
// case OP_AND:
|
||||
visited = false;
|
||||
push_frame(to_app(n));
|
||||
return;
|
||||
|
@ -197,10 +199,10 @@ class tseitin_cnf_tactic : public tactic {
|
|||
push_frame(to_app(n));
|
||||
}
|
||||
return;
|
||||
case OP_AND:
|
||||
case OP_XOR:
|
||||
case OP_IMPLIES:
|
||||
case OP_DISTINCT:
|
||||
case OP_AND:
|
||||
throw_op_not_handled();
|
||||
default:
|
||||
return;
|
||||
|
@ -617,6 +619,43 @@ class tseitin_cnf_tactic : public tactic {
|
|||
}
|
||||
return DONE;
|
||||
}
|
||||
|
||||
mres match_and(app * t, bool first, bool root) {
|
||||
if (!m.is_and(t))
|
||||
return NO;
|
||||
if (first) {
|
||||
bool visited = true;
|
||||
for (expr* a : *t)
|
||||
visit(a, visited);
|
||||
if (!visited)
|
||||
return CONT;
|
||||
}
|
||||
expr_ref_buffer lits(m);
|
||||
expr_ref l(m), nl(m);
|
||||
app_ref k(m), nk(m);
|
||||
if (root) {
|
||||
for (expr* arg : *t) {
|
||||
get_lit(arg, false, l);
|
||||
expr* lits[1] = { l };
|
||||
mk_clause(1, lits);
|
||||
}
|
||||
}
|
||||
else {
|
||||
k = mk_fresh();
|
||||
nk = m.mk_not(k);
|
||||
cache_result(t, k);
|
||||
|
||||
for (expr* arg : *t) {
|
||||
get_lit(arg, false, l);
|
||||
mk_clause(nk, l);
|
||||
inv(l, nl);
|
||||
lits.push_back(nl);
|
||||
}
|
||||
lits.push_back(k);
|
||||
mk_clause(lits.size(), lits.data());
|
||||
}
|
||||
return DONE;
|
||||
}
|
||||
|
||||
mres match_or(app * t, bool first, bool root) {
|
||||
if (!m.is_or(t))
|
||||
|
@ -778,6 +817,7 @@ class tseitin_cnf_tactic : public tactic {
|
|||
fr.m_first = false;
|
||||
TRY(match_or_3and);
|
||||
TRY(match_or);
|
||||
TRY(match_and);
|
||||
TRY(match_iff3);
|
||||
// TRY(match_iff_or);
|
||||
TRY(match_iff);
|
||||
|
|
|
@ -692,28 +692,23 @@ bool goal::is_cnf() const {
|
|||
for (unsigned i = 0; i < size(); i++) {
|
||||
expr * f = form(i);
|
||||
if (m_manager.is_or(f)) {
|
||||
for (expr* lit : *to_app(f)) {
|
||||
if (!is_literal(lit)) {
|
||||
for (expr* lit : *to_app(f))
|
||||
if (!is_literal(lit))
|
||||
return false;
|
||||
}
|
||||
}
|
||||
return true;
|
||||
}
|
||||
if (!is_literal(f)) {
|
||||
if (!is_literal(f))
|
||||
return false;
|
||||
}
|
||||
}
|
||||
return true;
|
||||
}
|
||||
|
||||
bool goal::is_literal(expr* f) const {
|
||||
m_manager.is_not(f, f);
|
||||
if (!is_app(f)) return false;
|
||||
if (to_app(f)->get_family_id() == m_manager.get_basic_family_id()) {
|
||||
if (!is_app(f))
|
||||
return false;
|
||||
if (to_app(f)->get_family_id() == m_manager.get_basic_family_id())
|
||||
for (expr* arg : *to_app(f))
|
||||
if (m_manager.is_bool(arg)) {
|
||||
if (m_manager.is_bool(arg))
|
||||
return false;
|
||||
}
|
||||
}
|
||||
return true;
|
||||
}
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue