3
0
Fork 0
mirror of https://github.com/Z3Prover/z3 synced 2025-08-21 02:30:23 +00:00

checkpoint

Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
Leonardo de Moura 2012-10-21 13:32:12 -07:00
parent 4722fdfca5
commit add684d8e9
377 changed files with 204 additions and 62 deletions

View file

@ -0,0 +1,26 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
arith_simplifier_params.cpp
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2008-05-09.
Revision History:
--*/
#include"arith_simplifier_params.h"
void arith_simplifier_params::register_params(ini_params & p) {
p.register_bool_param("ARITH_EXPAND_EQS", m_arith_expand_eqs);
p.register_bool_param("ARITH_PROCESS_ALL_EQS", m_arith_process_all_eqs);
}

View file

@ -0,0 +1,37 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
arith_simplifier_params.h
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2008-05-09.
Revision History:
--*/
#ifndef _ARITH_SIMPLIFIER_PARAMS_H_
#define _ARITH_SIMPLIFIER_PARAMS_H_
#include"ini_file.h"
struct arith_simplifier_params {
bool m_arith_expand_eqs;
bool m_arith_process_all_eqs;
arith_simplifier_params():
m_arith_expand_eqs(false),
m_arith_process_all_eqs(false) {
}
void register_params(ini_params & p);
};
#endif /* _ARITH_SIMPLIFIER_PARAMS_H_ */

421
src/simplifier/bit2int.cpp Normal file
View file

@ -0,0 +1,421 @@
/*++
Copyright (c) 2009 Microsoft Corporation
Module Name:
bit2cpp.cpp
Abstract:
Routines for simplifying bit2int expressions.
This propagates bv2int over arithmetical symbols as much as possible,
converting arithmetic operations into bit-vector operations.
Author:
Nikolaj Bjorner (nbjorner) 2009-08-28
Revision History:
--*/
#include "bit2int.h"
#include "ast_pp.h"
#include "ast_ll_pp.h"
#include "for_each_ast.h"
#define CHECK(_x_) if (!(_x_)) { UNREACHABLE(); }
bit2int::bit2int(ast_manager & m) :
m_manager(m), m_bv_util(m), m_arith_util(m), m_cache(m), m_bit0(m) {
m_bit0 = m_bv_util.mk_numeral(0,1);
}
void bit2int::operator()(expr * m, expr_ref & result, proof_ref& p) {
flush_cache();
expr_reduce emap(*this);
for_each_ast(emap, m);
result = get_cached(m);
if (m_manager.proofs_enabled() && m != result.get()) {
// TBD: rough
p = m_manager.mk_rewrite(m, result);
}
TRACE("bit2int",
tout << mk_pp(m, m_manager) << "======>\n"
<< mk_pp(result, m_manager) << "\n";);
}
unsigned bit2int::get_b2i_size(expr* n) {
SASSERT(m_bv_util.is_bv2int(n));
return m_bv_util.get_bv_size(to_app(n)->get_arg(0));
}
unsigned bit2int::get_numeral_bits(numeral const& k) {
numeral two(2);
numeral n(abs(k));
unsigned num_bits = 1;
n = div(n, two);
while (n.is_pos()) {
++num_bits;
n = div(n, two);
}
return num_bits;
}
void bit2int::align_size(expr* e, unsigned sz, expr_ref& result) {
unsigned sz1 = m_bv_util.get_bv_size(e);
SASSERT(sz1 <= sz);
m_bv_simplifier->mk_zeroext(sz-sz1, e, result);
}
void bit2int::align_sizes(expr_ref& a, expr_ref& b) {
unsigned sz1 = m_bv_util.get_bv_size(a);
unsigned sz2 = m_bv_util.get_bv_size(b);
expr_ref tmp(m_manager);
if (sz1 > sz2) {
m_bv_simplifier->mk_zeroext(sz1-sz2, b, tmp);
b = tmp;
}
else if (sz2 > sz1) {
m_bv_simplifier->mk_zeroext(sz2-sz1, a, tmp);
a = tmp;
}
}
bool bit2int::extract_bv(expr* n, unsigned& sz, bool& sign, expr_ref& bv) {
numeral k;
bool is_int;
if (m_bv_util.is_bv2int(n)) {
bv = to_app(n)->get_arg(0);
sz = m_bv_util.get_bv_size(bv);
sign = false;
return true;
}
else if (m_arith_util.is_numeral(n, k, is_int) && is_int) {
sz = get_numeral_bits(k);
bv = m_bv_util.mk_numeral(k, m_bv_util.mk_sort(sz));
sign = k.is_neg();
return true;
}
else {
return false;
}
}
bool bit2int::mk_add(expr* e1, expr* e2, expr_ref& result) {
unsigned sz1, sz2;
bool sign1, sign2;
expr_ref tmp1(m_manager), tmp2(m_manager), tmp3(m_manager);
if (extract_bv(e1, sz1, sign1, tmp1) && !sign1 &&
extract_bv(e2, sz2, sign2, tmp2) && !sign2) {
unsigned sz;
numeral k;
if (m_bv_util.is_numeral(tmp1, k, sz) && k.is_zero()) {
result = e2;
return true;
}
if (m_bv_util.is_numeral(tmp2, k, sz) && k.is_zero()) {
result = e1;
return true;
}
align_sizes(tmp1, tmp2);
m_bv_simplifier->mk_zeroext(1, tmp1, tmp1);
m_bv_simplifier->mk_zeroext(1, tmp2, tmp2);
SASSERT(m_bv_util.get_bv_size(tmp1) == m_bv_util.get_bv_size(tmp2));
m_bv_simplifier->mk_add(tmp1, tmp2, tmp3);
m_bv_simplifier->mk_bv2int(tmp3, m_arith_util.mk_int(), result);
return true;
}
return false;
}
bool bit2int::mk_comp(eq_type ty, expr* e1, expr* e2, expr_ref& result) {
unsigned sz1, sz2;
bool sign1, sign2;
expr_ref tmp1(m_manager), tmp2(m_manager), tmp3(m_manager);
if (extract_bv(e1, sz1, sign1, tmp1) && !sign1 &&
extract_bv(e2, sz2, sign2, tmp2) && !sign2) {
align_sizes(tmp1, tmp2);
SASSERT(m_bv_util.get_bv_size(tmp1) == m_bv_util.get_bv_size(tmp2));
switch(ty) {
case lt:
m_bv_simplifier->mk_leq_core(false, tmp2, tmp1, tmp3);
result = m_manager.mk_not(tmp3);
break;
case le:
m_bv_simplifier->mk_leq_core(false,tmp1, tmp2, result);
break;
case eq:
result = m_manager.mk_eq(tmp1,tmp2);
break;
}
return true;
}
return false;
}
bool bit2int::mk_mul(expr* e1, expr* e2, expr_ref& result) {
unsigned sz1, sz2;
bool sign1, sign2;
expr_ref tmp1(m_manager), tmp2(m_manager);
expr_ref tmp3(m_manager);
if (extract_bv(e1, sz1, sign1, tmp1) &&
extract_bv(e2, sz2, sign2, tmp2)) {
align_sizes(tmp1, tmp2);
m_bv_simplifier->mk_zeroext(m_bv_util.get_bv_size(tmp1), tmp1, tmp1);
m_bv_simplifier->mk_zeroext(m_bv_util.get_bv_size(tmp2), tmp2, tmp2);
SASSERT(m_bv_util.get_bv_size(tmp1) == m_bv_util.get_bv_size(tmp2));
m_bv_simplifier->mk_mul(tmp1, tmp2, tmp3);
m_bv_simplifier->mk_bv2int(tmp3, m_arith_util.mk_int(), result);
if (sign1 != sign2) {
result = m_arith_util.mk_uminus(result);
}
return true;
}
return false;
}
bool bit2int::is_bv_poly(expr* n, expr_ref& pos, expr_ref& neg) {
ptr_vector<expr> todo;
expr_ref tmp(m_manager);
numeral k;
bool is_int;
todo.push_back(n);
m_bv_simplifier->mk_bv2int(m_bit0, m_arith_util.mk_int(), pos);
m_bv_simplifier->mk_bv2int(m_bit0, m_arith_util.mk_int(), neg);
while (!todo.empty()) {
n = todo.back();
todo.pop_back();
if (m_bv_util.is_bv2int(n)) {
CHECK(mk_add(n, pos, pos));
}
else if (m_arith_util.is_numeral(n, k, is_int) && is_int) {
if (k.is_nonneg()) {
CHECK(mk_add(n, pos, pos));
}
else {
tmp = m_arith_util.mk_numeral(-k, true);
CHECK(mk_add(tmp, neg, neg));
}
}
else if (m_arith_util.is_add(n)) {
for (unsigned i = 0; i < to_app(n)->get_num_args(); ++i) {
todo.push_back(to_app(n)->get_arg(i));
}
}
else if (m_arith_util.is_mul(n) &&
to_app(n)->get_num_args() == 2 &&
m_arith_util.is_numeral(to_app(n)->get_arg(0), k, is_int) && is_int && k.is_minus_one() &&
m_bv_util.is_bv2int(to_app(n)->get_arg(1))) {
CHECK(mk_add(to_app(n)->get_arg(1), neg, neg));
}
else if (m_arith_util.is_mul(n) &&
to_app(n)->get_num_args() == 2 &&
m_arith_util.is_numeral(to_app(n)->get_arg(1), k, is_int) && is_int && k.is_minus_one() &&
m_bv_util.is_bv2int(to_app(n)->get_arg(0))) {
CHECK(mk_add(to_app(n)->get_arg(0), neg, neg));
}
else if (m_arith_util.is_uminus(n) &&
m_bv_util.is_bv2int(to_app(n)->get_arg(0))) {
CHECK(mk_add(to_app(n)->get_arg(0), neg, neg));
}
else {
TRACE("bit2int", tout << "Not a poly: " << mk_pp(n, m_manager) << "\n";);
return false;
}
}
return true;
}
void bit2int::visit(quantifier* q) {
expr_ref result(m_manager);
result = get_cached(q->get_expr());
result = m_manager.update_quantifier(q, result);
cache_result(q, result);
}
void bit2int::visit(app* n) {
func_decl* f = n->get_decl();
unsigned num_args = n->get_num_args();
m_args.reset();
for (unsigned i = 0; i < num_args; ++i) {
m_args.push_back(get_cached(n->get_arg(i)));
}
expr* const* args = m_args.c_ptr();
bool has_b2i =
m_arith_util.is_le(n) || m_arith_util.is_ge(n) || m_arith_util.is_gt(n) ||
m_arith_util.is_lt(n) || m_manager.is_eq(n);
expr_ref result(m_manager);
for (unsigned i = 0; !has_b2i && i < num_args; ++i) {
has_b2i = m_bv_util.is_bv2int(args[i]);
}
if (!has_b2i) {
result = m_manager.mk_app(f, num_args, args);
cache_result(n, result);
return;
}
//
// bv2int(x) + bv2int(y) -> bv2int(pad(x) + pad(y))
// bv2int(x) + k -> bv2int(pad(x) + pad(k))
// bv2int(x) * bv2int(y) -> bv2int(pad(x) * pad(y))
// bv2int(x) * k -> sign(k)*bv2int(pad(x) * pad(k))
// bv2int(x) - bv2int(y) <= z -> bv2int(x) <= bv2int(y) + z
// bv2int(x) <= z - bv2int(y) -> bv2int(x) + bv2int(y) <= z
//
expr* e1, *e2;
expr_ref tmp1(m_manager), tmp2(m_manager);
expr_ref tmp3(m_manager);
expr_ref pos1(m_manager), neg1(m_manager);
expr_ref pos2(m_manager), neg2(m_manager);
expr_ref e2bv(m_manager);
bool sign2;
numeral k;
unsigned sz2;
if (num_args >= 2) {
e1 = args[0];
e2 = args[1];
}
if (m_arith_util.is_add(n) && num_args >= 1) {
result = e1;
for (unsigned i = 1; i < num_args; ++i) {
e1 = result;
e2 = args[i];
if (!mk_add(e1, e2, result)) {
result = m_manager.mk_app(f, num_args, args);
cache_result(n, result);
return;
}
}
cache_result(n, result);
}
else if (m_arith_util.is_mul(n) && num_args >= 1) {
result = e1;
for (unsigned i = 1; i < num_args; ++i) {
e1 = result;
e2 = args[i];
if (!mk_mul(e1, e2, result)) {
result = m_manager.mk_app(f, num_args, args);
cache_result(n, result);
return;
}
}
cache_result(n, result);
}
else if (m_manager.is_eq(n) &&
is_bv_poly(e1, pos1, neg1) &&
is_bv_poly(e2, pos2, neg2) &&
mk_add(pos1, neg2, tmp1) &&
mk_add(neg1, pos2, tmp2) &&
mk_comp(eq, tmp1, tmp2, result)) {
cache_result(n, result);
}
else if (m_arith_util.is_le(n) &&
is_bv_poly(e1, pos1, neg1) &&
is_bv_poly(e2, pos2, neg2) &&
mk_add(pos1, neg2, tmp1) &&
mk_add(neg1, pos2, tmp2) &&
mk_comp(le, tmp1, tmp2, result)) {
cache_result(n, result);
}
else if (m_arith_util.is_lt(n) &&
is_bv_poly(e1, pos1, neg1) &&
is_bv_poly(e2, pos2, neg2) &&
mk_add(pos1, neg2, tmp1) &&
mk_add(neg1, pos2, tmp2) &&
mk_comp(lt, tmp1, tmp2, result)) {
cache_result(n, result);
}
else if (m_arith_util.is_ge(n) &&
is_bv_poly(e1, pos1, neg1) &&
is_bv_poly(e2, pos2, neg2) &&
mk_add(pos1, neg2, tmp1) &&
mk_add(neg1, pos2, tmp2) &&
mk_comp(le, tmp2, tmp1, result)) {
cache_result(n, result);
}
else if (m_arith_util.is_gt(n) &&
is_bv_poly(e1, pos1, neg1) &&
is_bv_poly(e2, pos2, neg2) &&
mk_add(pos1, neg2, tmp1) &&
mk_add(neg1, pos2, tmp2) &&
mk_comp(lt, tmp2, tmp1, result)) {
cache_result(n, result);
}
else if (m_arith_util.is_mod(n) &&
is_bv_poly(e1, pos1, neg1) &&
extract_bv(e2, sz2, sign2, e2bv) && !sign2) {
//
// (pos1 - neg1) mod e2 = (pos1 + (e2 - (neg1 mod e2))) mod e2
//
unsigned sz_p, sz_n, sz;
bool sign_p, sign_n;
expr_ref tmp_p(m_manager), tmp_n(m_manager);
CHECK(extract_bv(pos1, sz_p, sign_p, tmp_p));
CHECK(extract_bv(neg1, sz_n, sign_n, tmp_n));
SASSERT(!sign_p && !sign_n);
// pos1 mod e2
if (m_bv_util.is_numeral(tmp_n, k, sz) && k.is_zero()) {
tmp1 = tmp_p;
tmp2 = e2bv;
align_sizes(tmp1, tmp2);
m_bv_simplifier->mk_bv_urem(tmp1, tmp2, tmp3);
m_bv_simplifier->mk_bv2int(tmp3, m_arith_util.mk_int(), result);
cache_result(n, result);
return;
}
// neg1 mod e2;
tmp1 = tmp_n;
tmp2 = e2bv;
align_sizes(tmp1, tmp2);
m_bv_simplifier->mk_bv_urem(tmp1, tmp2, tmp3);
// e2 - (neg1 mod e2)
tmp1 = e2bv;
tmp2 = tmp3;
align_sizes(tmp1, tmp2);
m_bv_simplifier->mk_sub(tmp1, tmp2, tmp3);
// pos1 + (e2 - (neg1 mod e2))
tmp1 = tmp_p;
tmp2 = tmp3;
align_sizes(tmp1, tmp2);
m_bv_simplifier->mk_zeroext(1, tmp1, tmp_p);
m_bv_simplifier->mk_zeroext(1, tmp2, tmp_n);
m_bv_simplifier->mk_add(tmp_p, tmp_n, tmp1);
// (pos1 + (e2 - (neg1 mod e2))) mod e2
tmp2 = e2bv;
align_sizes(tmp1, tmp2);
m_bv_simplifier->mk_bv_urem(tmp1, tmp2, tmp3);
m_bv_simplifier->mk_bv2int(tmp3, m_arith_util.mk_int(), result);
cache_result(n, result);
}
else {
result = m_manager.mk_app(f, num_args, args);
cache_result(n, result);
}
}
expr * bit2int::get_cached(expr * n) const {
return const_cast<bit2int*>(this)->m_cache.find(n);
}
void bit2int::cache_result(expr * n, expr * r) {
TRACE("bit2int_verbose", tout << "caching:\n" << mk_ll_pp(n, m_manager) <<
"======>\n" << mk_ll_pp(r, m_manager) << "\n";);
m_cache.insert(n, r);
}

96
src/simplifier/bit2int.h Normal file
View file

@ -0,0 +1,96 @@
/*++
Copyright (c) 2009 Microsoft Corporation
Module Name:
bit2int.h
Abstract:
Routines for simplifying bit2int expressions.
Author:
Nikolaj Bjorner (nbjorner) 2009-08-28
Revision History:
--*/
#ifndef _BIT2INT_H_
#define _BIT2INT_H_
#include"bv_decl_plugin.h"
#include"arith_decl_plugin.h"
#include"act_cache.h"
#include"basic_simplifier_plugin.h"
#include"bv_simplifier_plugin.h"
class bit2int {
protected:
typedef rational numeral;
enum eq_type {
lt,
le,
eq
};
class expr_reduce {
bit2int& m_super;
public:
expr_reduce(bit2int& s) : m_super(s) {}
void operator()(var* v) {
m_super.cache_result(v, v);
}
void operator()(quantifier* q) {
m_super.visit(q);
}
void operator()(app* a) {
m_super.visit(a);
}
void operator()(ast* a) {}
};
typedef act_cache expr_map;
ast_manager & m_manager;
bv_util m_bv_util;
arith_util m_arith_util;
bv_simplifier_plugin * m_bv_simplifier;
expr_map m_cache; // map: ast -> ast ref. counters are incremented when inserted here.
expr_ref m_bit0;
ptr_vector<expr> m_args;
void visit(app* n);
void visit(quantifier* q);
unsigned get_b2i_size(expr * n);
bool extract_bv(expr* n, unsigned& sz, bool& sign, expr_ref& bv);
unsigned get_numeral_bits(numeral const& k);
bool is_bv_poly(expr* n, expr_ref& pos, expr_ref& neg);
bool mk_mul(expr* a, expr* b, expr_ref& result);
bool mk_comp(eq_type ty, expr* e1, expr* e2, expr_ref& result);
bool mk_add(expr* e1, expr* e2, expr_ref& result);
expr * get_cached(expr * n) const;
bool is_cached(expr * n) const { return get_cached(n) != 0; }
void cache_result(expr * n, expr * r);
void reset_cache() { m_cache.reset(); }
void flush_cache() { m_cache.cleanup(); }
void align_size(expr* e, unsigned sz, expr_ref& result);
void align_sizes(expr_ref& a, expr_ref& b);
public:
bit2int(ast_manager & m);
void set_bv_simplifier(bv_simplifier_plugin * p) { m_bv_simplifier = p; }
void operator()(expr * m, expr_ref & result, proof_ref& p);
};
#endif /* _BIT2INT_H_ */

113
src/simplifier/bv_elim.cpp Normal file
View file

@ -0,0 +1,113 @@
#include "bv_elim.h"
#include "bv_decl_plugin.h"
#include "var_subst.h"
#include <sstream>
void bv_elim::elim(quantifier* q, quantifier_ref& r) {
svector<symbol> names, _names;
sort_ref_buffer sorts(m_manager), _sorts(m_manager);
expr_ref_buffer pats(m_manager);
expr_ref_buffer no_pats(m_manager);
expr_ref_buffer subst_map(m_manager), _subst_map(m_manager);
var_subst subst(m_manager);
bv_util bv(m_manager);
expr_ref new_body(m_manager);
expr* old_body = q->get_expr();
unsigned num_decls = q->get_num_decls();
family_id bfid = m_manager.get_family_id("bv");
//
// Traverse sequence of bound variables to eliminate
// bit-vecctor variables and replace them by
// Booleans.
//
unsigned var_idx = 0;
for (unsigned i = num_decls; i > 0; ) {
--i;
sort* s = q->get_decl_sort(i);
symbol nm = q->get_decl_name(i);
if (bv.is_bv_sort(s)) {
// convert n-bit bit-vector variable into sequence of n-Booleans.
unsigned num_bits = bv.get_bv_size(s);
expr_ref_buffer args(m_manager);
expr_ref bv(m_manager);
for (unsigned j = 0; j < num_bits; ++j) {
std::ostringstream new_name;
new_name << nm.str();
new_name << "_";
new_name << j;
var* v = m_manager.mk_var(var_idx++, m_manager.mk_bool_sort());
args.push_back(v);
_sorts.push_back(m_manager.mk_bool_sort());
_names.push_back(symbol(new_name.str().c_str()));
}
bv = m_manager.mk_app(bfid, OP_MKBV, 0, 0, args.size(), args.c_ptr());
_subst_map.push_back(bv.get());
}
else {
_subst_map.push_back(m_manager.mk_var(var_idx++, s));
_sorts.push_back(s);
_names.push_back(nm);
}
}
//
// reverse the vectors.
//
SASSERT(_names.size() == _sorts.size());
for (unsigned i = _names.size(); i > 0; ) {
--i;
names.push_back(_names[i]);
sorts.push_back(_sorts[i]);
}
for (unsigned i = _subst_map.size(); i > 0; ) {
--i;
subst_map.push_back(_subst_map[i]);
}
expr* const* sub = subst_map.c_ptr();
unsigned sub_size = subst_map.size();
subst(old_body, sub_size, sub, new_body);
for (unsigned j = 0; j < q->get_num_patterns(); j++) {
expr_ref pat(m_manager);
subst(q->get_pattern(j), sub_size, sub, pat);
pats.push_back(pat);
}
for (unsigned j = 0; j < q->get_num_no_patterns(); j++) {
expr_ref nopat(m_manager);
subst(q->get_no_pattern(j), sub_size, sub, nopat);
no_pats.push_back(nopat);
}
r = m_manager.mk_quantifier(true,
names.size(),
sorts.c_ptr(),
names.c_ptr(),
new_body.get(),
q->get_weight(),
q->get_qid(),
q->get_skid(),
pats.size(), pats.c_ptr(),
no_pats.size(), no_pats.c_ptr());
}
bool bv_elim_star::visit_quantifier(quantifier* q) {
// behave like destructive resolution, do not recurse.
return true;
}
void bv_elim_star::reduce1_quantifier(quantifier* q) {
quantifier_ref r(m_manager);
proof_ref pr(m_manager);
m_bv_elim.elim(q, r);
if (m_manager.fine_grain_proofs()) {
pr = m_manager.mk_rewrite(q, r.get());
}
else {
pr = 0;
}
cache_result(q, r, pr);
}

45
src/simplifier/bv_elim.h Normal file
View file

@ -0,0 +1,45 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
bv_elim.h
Abstract:
Eliminate bit-vectors variables from clauses, by
replacing them by bound Boolean variables.
Author:
Nikolaj Bjorner (nbjorner) 2008-12-16.
Revision History:
--*/
#ifndef _BV_ELIM_H_
#define _BV_ELIM_H_
#include "ast.h"
#include "simplifier.h"
class bv_elim {
ast_manager& m_manager;
public:
bv_elim(ast_manager& m) : m_manager(m) {};
void elim(quantifier* q, quantifier_ref& r);
};
class bv_elim_star : public simplifier {
protected:
bv_elim m_bv_elim;
virtual bool visit_quantifier(quantifier* q);
virtual void reduce1_quantifier(quantifier* q);
public:
bv_elim_star(ast_manager& m) : simplifier(m), m_bv_elim(m) { enable_ac_support(false); }
virtual ~bv_elim_star() {}
};
#endif /* _BV_ELIM_H_ */

View file

@ -0,0 +1,39 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
bv_simplifier_params.h
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2008-10-10.
Revision History:
--*/
#ifndef _BV_SIMPLIFIER_PARAMS_H_
#define _BV_SIMPLIFIER_PARAMS_H_
#include"ini_file.h"
struct bv_simplifier_params {
bool m_hi_div0; //!< if true, uses the hardware interpretation for div0, mod0, ... if false, div0, mod0, ... are considered uninterpreted.
bool m_bv2int_distribute; //!< if true allows downward propagation of bv2int.
bv_simplifier_params():
m_hi_div0(true),
m_bv2int_distribute(true) {
}
void register_params(ini_params & p) {
p.register_bool_param("HI_DIV0", m_hi_div0, "if true, then Z3 uses the usual hardware interpretation for division (rem, mod) by zero. Otherwise, these operations are considered uninterpreted.");
p.register_bool_param("BV2INT_DISTRIBUTE", m_bv2int_distribute, "if true, then int2bv is distributed over arithmetical operators.");
}
};
#endif /* _BV_SIMPLIFIER_PARAMS_H_ */

View file

@ -0,0 +1,170 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
distribute_forall.cpp
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2010-04-02.
Revision History:
Christoph Wintersteiger 2010-04-06: Added implementation.
--*/
#include"var_subst.h"
#include"ast_ll_pp.h"
#include"distribute_forall.h"
distribute_forall::distribute_forall(ast_manager & m, basic_simplifier_plugin & p) :
m_manager(m),
m_bsimp(p),
m_cache(m) {
}
void distribute_forall::visit(expr * n, bool & visited) {
if (!is_cached(n)) {
m_todo.push_back(n);
visited = false;
}
}
bool distribute_forall::visit_children(expr * n) {
bool visited = true;
unsigned j;
switch(n->get_kind()) {
case AST_VAR:
break;
case AST_APP:
j = to_app(n)->get_num_args();
while (j > 0) {
--j;
visit(to_app(n)->get_arg(j), visited);
}
break;
case AST_QUANTIFIER:
visit(to_quantifier(n)->get_expr(), visited);
break;
default:
UNREACHABLE();
}
return visited;
}
void distribute_forall::reduce1(expr * n) {
switch (n->get_kind()) {
case AST_VAR:
cache_result(n, n);
break;
case AST_APP:
reduce1_app(to_app(n));
break;
case AST_QUANTIFIER:
reduce1_quantifier(to_quantifier(n));
break;
default: UNREACHABLE();
}
}
void distribute_forall::reduce1_app(app * a) {
SASSERT(a);
unsigned num_args = a->get_num_args();
unsigned j = num_args;
bool reduced = false;
m_new_args.reserve(num_args);
app * na = a;
while(j > 0) {
--j;
SASSERT(is_cached(a->get_arg(j)));
expr * c = get_cached(a->get_arg(j));
SASSERT(c!=0);
if (c != a->get_arg(j))
reduced = true;
m_new_args[j] = c;
}
if (reduced) {
na = m_manager.mk_app(a->get_decl(), num_args, m_new_args.c_ptr());
}
cache_result(a, na);
}
void distribute_forall::reduce1_quantifier(quantifier * q) {
// This transformation is applied after skolemization/quantifier elimination. So, all quantifiers are universal.
SASSERT(q->is_forall());
// This transformation is applied after basic pre-processing steps.
// So, we can assume that
// 1) All (and f1 ... fn) are already encoded as (not (or (not f1 ... fn)))
// 2) All or-formulas are flat (or f1 (or f2 f3)) is encoded as (or f1 f2 f3)
expr * e = get_cached(q->get_expr());
if (m_manager.is_not(e) && m_manager.is_or(to_app(e)->get_arg(0))) {
// found target for simplification
// (forall X (not (or F1 ... Fn)))
// -->
// (and (forall X (not F1))
// ...
// (forall X (not Fn)))
app * or_e = to_app(to_app(e)->get_arg(0));
unsigned num_args = or_e->get_num_args();
expr_ref_buffer new_args(m_manager);
for (unsigned i = 0; i < num_args; i++) {
expr * arg = or_e->get_arg(i);
expr_ref not_arg(m_manager);
// m_bsimp.mk_not applies basic simplifications. For example, if arg is of the form (not a), then it will return a.
m_bsimp.mk_not(arg, not_arg);
quantifier_ref tmp_q(m_manager);
tmp_q = m_manager.update_quantifier(q, not_arg);
expr_ref new_q(m_manager);
elim_unused_vars(m_manager, tmp_q, new_q);
new_args.push_back(new_q);
}
expr_ref result(m_manager);
// m_bsimp.mk_and actually constructs a (not (or ...)) formula,
// it will also apply basic simplifications.
m_bsimp.mk_and(new_args.size(), new_args.c_ptr(), result);
cache_result(q, result);
}
else {
cache_result(q, m_manager.update_quantifier(q, e));
}
}
void distribute_forall::operator()(expr * f, expr_ref & result) {
m_todo.reset();
flush_cache();
m_todo.push_back(f);
while (!m_todo.empty()) {
expr * e = m_todo.back();
if (visit_children(e)) {
m_todo.pop_back();
reduce1(e);
}
}
result = get_cached(f);
SASSERT(result!=0);
TRACE("distribute_forall", tout << mk_ll_pp(f, m_manager) << "======>\n"
<< mk_ll_pp(result, m_manager););
}
expr * distribute_forall::get_cached(expr * n) const {
return const_cast<distribute_forall*>(this)->m_cache.find(n);
}
void distribute_forall::cache_result(expr * n, expr * r) {
SASSERT(r != 0);
m_cache.insert(n, r);
}

View file

@ -0,0 +1,82 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
distribute_forall.h
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2010-04-02.
Revision History:
Christoph Wintersteiger 2010-04-06: Added implementation
--*/
#ifndef _DISTRIBUTE_FORALL_H_
#define _DISTRIBUTE_FORALL_H_
#include"ast.h"
#include"basic_simplifier_plugin.h"
#include"act_cache.h"
/**
\brief Apply the following transformation
(forall X (and F1 ... Fn))
-->
(and (forall X F1) ... (forall X Fn))
The actual transformation is slightly different since the "and" connective is eliminated and
replaced with a "not or".
So, the actual transformation is:
(forall X (not (or F1 ... Fn)))
-->
(not (or (not (forall X (not F1)))
...
(not (forall X (not Fn)))))
The implementation uses the visit_children/reduce1 idiom. A cache is used as usual.
*/
class distribute_forall {
typedef act_cache expr_map;
ast_manager & m_manager;
basic_simplifier_plugin & m_bsimp; // useful for constructing formulas and/or/not in simplified form.
ptr_vector<expr> m_todo;
expr_map m_cache;
ptr_vector<expr> m_new_args;
// The new expressions are stored in a mapping that increments their reference counter. So, we do not need to store them in
// m_new_exprs
// expr_ref_vector m_new_exprs;
public:
distribute_forall(ast_manager & m, basic_simplifier_plugin & p);
/**
\brief Apply the distribute_forall transformation (when possible) to all universal quantifiers in \c f.
Store the result in \c result.
*/
void operator()(expr * f, expr_ref & result);
protected:
inline void visit(expr * n, bool & visited);
bool visit_children(expr * n);
void reduce1(expr * n);
void reduce1_quantifier(quantifier * q);
void reduce1_app(app * a);
expr * get_cached(expr * n) const;
bool is_cached(expr * n) const { return get_cached(n) != 0; }
void cache_result(expr * n, expr * r);
void reset_cache() { m_cache.reset(); }
void flush_cache() { m_cache.cleanup(); }
};
#endif /* _DISTRIBUTE_FORALL_H_ */

View file

@ -0,0 +1,222 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
elim_bounds.cpp
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2008-06-28.
Revision History:
--*/
#include"elim_bounds.h"
#include"used_vars.h"
#include"obj_hashtable.h"
#include"var_subst.h"
#include"ast_pp.h"
elim_bounds::elim_bounds(ast_manager & m):
m_manager(m),
m_util(m) {
}
/**
\brief Find bounds of the form
(<= x k)
(<= (+ x (* -1 y)) k)
(<= (+ x (* -1 t)) k)
(<= (+ t (* -1 x)) k)
x and y are a bound variables, t is a ground term and k is a numeral
It also detects >=, and the atom can be negated.
*/
bool elim_bounds::is_bound(expr * n, var * & lower, var * & upper) {
upper = 0;
lower = 0;
bool neg = false;
if (m_manager.is_not(n)) {
n = to_app(n)->get_arg(0);
neg = true;
}
bool le = false;
if (m_util.is_le(n)) {
SASSERT(m_util.is_numeral(to_app(n)->get_arg(1)));
n = to_app(n)->get_arg(0);
le = true;
}
else if (m_util.is_ge(n)) {
SASSERT(m_util.is_numeral(to_app(n)->get_arg(1)));
n = to_app(n)->get_arg(0);
le = false;
}
else {
return false;
}
if (neg)
le = !le;
if (is_var(n)) {
upper = to_var(n);
}
else if (m_util.is_add(n) && to_app(n)->get_num_args() == 2) {
expr * arg1 = to_app(n)->get_arg(0);
expr * arg2 = to_app(n)->get_arg(1);
if (is_var(arg1))
upper = to_var(arg1);
else if (!is_ground(arg1))
return false;
rational k;
bool is_int;
if (m_util.is_mul(arg2) && m_util.is_numeral(to_app(arg2)->get_arg(0), k, is_int) && k.is_minus_one()) {
arg2 = to_app(arg2)->get_arg(1);
if (is_var(arg2))
lower = to_var(arg2);
else if (!is_ground(arg2))
return false; // not supported
}
else {
return false; // not supported
}
}
else {
return false;
}
if (!le)
std::swap(upper, lower);
return true;
}
bool elim_bounds::is_bound(expr * n) {
var * lower, * upper;
return is_bound(n, lower, upper);
}
void elim_bounds::operator()(quantifier * q, expr_ref & r) {
if (!q->is_forall()) {
r = q;
return;
}
expr * n = q->get_expr();
ptr_buffer<expr> atoms;
if (m_manager.is_or(n))
atoms.append(to_app(n)->get_num_args(), to_app(n)->get_args());
else
atoms.push_back(n);
used_vars m_used_vars;
// collect non-candidates
unsigned sz = atoms.size();
for (unsigned i = 0; i < sz; i++) {
expr * a = atoms[i];
if (!is_bound(a))
m_used_vars.process(a);
}
if (m_used_vars.uses_all_vars(q->get_num_decls())) {
r = q;
return;
}
// collect candidates
obj_hashtable<var> m_lowers;
obj_hashtable<var> m_uppers;
obj_hashtable<var> m_candidate_set;
ptr_buffer<var> m_candidates;
#define ADD_CANDIDATE(V) if (!m_lowers.contains(V) && !m_uppers.contains(V)) { m_candidate_set.insert(V); m_candidates.push_back(V); }
for (unsigned i = 0; i < sz; i++) {
expr * a = atoms[i];
var * lower = 0;
var * upper = 0;
if (is_bound(a, lower, upper)) {
if (lower != 0 && !m_used_vars.contains(lower->get_idx())) {
ADD_CANDIDATE(lower);
m_lowers.insert(lower);
}
if (upper != 0 && !m_used_vars.contains(upper->get_idx())) {
ADD_CANDIDATE(upper);
m_uppers.insert(upper);
}
}
}
TRACE("elim_bounds", tout << "candidates:\n"; for (unsigned i = 0; i < m_candidates.size(); i++) tout << mk_pp(m_candidates[i], m_manager) << "\n";);
// remove candidates that have lower and upper bounds
for (unsigned i = 0; i < m_candidates.size(); i++) {
var * v = m_candidates[i];
if (m_lowers.contains(v) && m_uppers.contains(v))
m_candidate_set.erase(v);
}
TRACE("elim_bounds", tout << "candidates after filter:\n"; for (unsigned i = 0; i < m_candidates.size(); i++) tout << mk_pp(m_candidates[i], m_manager) << "\n";);
if (m_candidate_set.empty()) {
r = q;
return;
}
// remove bounds that contain variables in m_candidate_set
unsigned j = 0;
for (unsigned i = 0; i < sz; i++) {
expr * a = atoms[i];
var * lower = 0;
var * upper = 0;
if (is_bound(a, lower, upper) && ((lower != 0 && m_candidate_set.contains(lower)) || (upper != 0 && m_candidate_set.contains(upper))))
continue;
atoms[j] = a;
j++;
}
atoms.resize(j);
expr * new_body = 0;
switch (atoms.size()) {
case 0:
r = m_manager.mk_false();
return;
case 1:
new_body = atoms[0];
break;
default:
new_body = m_manager.mk_or(atoms.size(), atoms.c_ptr());
break;
}
quantifier_ref new_q(m_manager);
new_q = m_manager.update_quantifier(q, new_body);
elim_unused_vars(m_manager, new_q, r);
TRACE("elim_bounds", tout << mk_pp(q, m_manager) << "\n" << mk_pp(r, m_manager) << "\n";);
}
bool elim_bounds_star::visit_quantifier(quantifier * q) {
if (!q->is_forall() || q->get_num_patterns() != 0)
return true;
bool visited = true;
visit(q->get_expr(), visited);
return visited;
}
void elim_bounds_star::reduce1_quantifier(quantifier * q) {
if (!q->is_forall() || q->get_num_patterns() != 0) {
cache_result(q, q, 0);
return;
}
quantifier_ref new_q(m_manager);
expr * new_body = 0;
proof * new_pr;
get_cached(q->get_expr(), new_body, new_pr);
new_q = m_manager.update_quantifier(q, new_body);
expr_ref r(m_manager);
m_elim(new_q, r);
if (q == r.get()) {
cache_result(q, q, 0);
return;
}
proof_ref pr(m_manager);
if (m_manager.fine_grain_proofs())
pr = m_manager.mk_rewrite(q, r); // TODO: improve justification
cache_result(q, r, pr);
}

View file

@ -0,0 +1,69 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
elim_bounds.h
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2008-06-28.
Revision History:
--*/
#ifndef _ELIM_BOUNDS_H_
#define _ELIM_BOUNDS_H_
#include"ast.h"
#include"arith_decl_plugin.h"
#include"simplifier.h"
/**
\brief Functor for eliminating irrelevant bounds in quantified formulas.
Example:
(forall (x Int) (y Int) (or (not (>= y x) (not (>= x 0)) (= (select a x) 1))))
The bound (>= y x) is irrelevant and can be eliminated.
This can be easily proved by using Fourier-Motzkin elimination.
Limitations & Assumptions:
- It assumes the input formula was already simplified.
- It can only handle bounds in the diff-logic fragment.
\remark This operation is subsumed by Fourier-Motzkin elimination.
*/
class elim_bounds {
ast_manager & m_manager;
arith_util m_util;
bool is_bound(expr * n, var * & lower, var * & upper);
bool is_bound(expr * n);
public:
elim_bounds(ast_manager & m);
void operator()(quantifier * q, expr_ref & r);
};
/**
\brief Functor for applying elim_bounds in all
universal quantifiers in an expression.
Assumption: the formula was already skolemized.
*/
class elim_bounds_star : public simplifier {
protected:
elim_bounds m_elim;
virtual bool visit_quantifier(quantifier * q);
virtual void reduce1_quantifier(quantifier * q);
public:
elim_bounds_star(ast_manager & m):simplifier(m), m_elim(m) { enable_ac_support(false); }
virtual ~elim_bounds_star() {}
};
#endif /* _ELIM_BOUNDS_H_ */

View file

@ -0,0 +1,142 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
inj_axiom.cpp
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2008-06-23.
Revision History:
--*/
#include"inj_axiom.h"
#include"ast_pp.h"
#include"ast_ll_pp.h"
#include"has_free_vars.h"
#include"well_sorted.h"
/**
\brief Little HACK for simplifying injectivity axioms
\remark It is not covering all possible cases.
*/
bool simplify_inj_axiom(ast_manager & m, quantifier * q, expr_ref & result) {
expr * n = q->get_expr();
if (q->is_forall() && m.is_or(n) && to_app(n)->get_num_args() == 2) {
expr * arg1 = to_app(n)->get_arg(0);
expr * arg2 = to_app(n)->get_arg(1);
if (m.is_not(arg2))
std::swap(arg1, arg2);
if (m.is_not(arg1) &&
m.is_eq(to_app(arg1)->get_arg(0)) &&
m.is_eq(arg2)) {
expr * app1 = to_app(to_app(arg1)->get_arg(0))->get_arg(0);
expr * app2 = to_app(to_app(arg1)->get_arg(0))->get_arg(1);
expr * var1 = to_app(arg2)->get_arg(0);
expr * var2 = to_app(arg2)->get_arg(1);
if (is_app(app1) &&
is_app(app2) &&
to_app(app1)->get_decl() == to_app(app2)->get_decl() &&
to_app(app1)->get_num_args() == to_app(app2)->get_num_args() &&
to_app(app1)->get_family_id() == null_family_id &&
to_app(app1)->get_num_args() > 0 &&
is_var(var1) &&
is_var(var2) &&
var1 != var2) {
app * f1 = to_app(app1);
app * f2 = to_app(app2);
bool found_vars = false;
unsigned num = f1->get_num_args();
unsigned idx = UINT_MAX;
unsigned num_vars = 1;
for (unsigned i = 0; i < num; i++) {
expr * c1 = f1->get_arg(i);
expr * c2 = f2->get_arg(i);
if (!is_var(c1) && !is_uninterp_const(c1))
return false;
if ((c1 == var1 && c2 == var2) || (c1 == var2 && c2 == var1)) {
if (found_vars)
return false;
found_vars = true;
idx = i;
}
else if (c1 == c2 && c1 != var1 && c1 != var2) {
if (is_var(c1)) {
++num_vars;
}
}
else {
return false;
}
}
if (found_vars && !has_free_vars(q)) {
TRACE("inj_axiom",
tout << "Cadidate for simplification:\n" << mk_ll_pp(q, m) << mk_pp(app1, m) << "\n" << mk_pp(app2, m) << "\n" <<
mk_pp(var1, m) << "\n" << mk_pp(var2, m) << "\nnum_vars: " << num_vars << "\n";);
// Building new (optimized) axiom
func_decl * decl = f1->get_decl();
unsigned var_idx = 0;
ptr_buffer<expr> f_args, inv_vars;
ptr_buffer<sort> decls;
buffer<symbol> names;
expr * var = 0;
for (unsigned i = 0; i < num; i++) {
expr * c = f1->get_arg(i);
if (is_var(c)) {
names.push_back(symbol(i));
sort * s = decl->get_domain(i);
decls.push_back(s);
expr * new_c = m.mk_var(var_idx, s);
var_idx++;
f_args.push_back(new_c);
if (i == idx) {
var = new_c;
}
else {
inv_vars.push_back(new_c);
}
}
else {
SASSERT(is_uninterp_const(c));
f_args.push_back(c);
}
}
SASSERT(var != 0);
app * f = m.mk_app(decl, f_args.size(), f_args.c_ptr());
ptr_vector<sort> domain;
inv_vars.push_back(f);
for (unsigned i = 0; i < inv_vars.size(); ++i) {
domain.push_back(m.get_sort(inv_vars[i]));
}
sort * d = decl->get_domain(idx);
func_decl * inv_decl = m.mk_fresh_func_decl("inj", domain.size(), domain.c_ptr(), d);
expr * proj = m.mk_app(inv_decl, inv_vars.size(), inv_vars.c_ptr());
expr * eq = m.mk_eq(proj, var);
expr * p = m.mk_pattern(f);
// decls are in the wrong order...
// Remark: the sort of the var 0 must be in the last position.
std::reverse(decls.begin(), decls.end());
result = m.mk_forall(decls.size(), decls.c_ptr(), names.c_ptr(), eq,
0, symbol(), symbol(), 1, &p);
TRACE("inj_axiom", tout << "new axiom:\n" << mk_pp(result, m) << "\n";);
SASSERT(is_well_sorted(m, result));
return true;
}
}
}
}
return false;
}

View file

@ -0,0 +1,27 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
inj_axiom.h
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2008-06-23.
Revision History:
--*/
#ifndef _INJ_AXIOM_H_
#define _INJ_AXIOM_H_
#include"ast.h"
bool simplify_inj_axiom(ast_manager & m, quantifier * q, expr_ref & result);
#endif /* _INJ_AXIOM_H_ */

View file

@ -0,0 +1,192 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
maximise_ac_sharing.cpp
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2008-10-22.
Revision History:
--*/
#include"maximise_ac_sharing.h"
#include"ast_pp.h"
#include"basic_simplifier_plugin.h"
maximise_ac_sharing::ac_plugin::ac_plugin(symbol const & fname, ast_manager & m, maximise_ac_sharing & owner):
simplifier_plugin(fname, m),
m_owner(owner) {
}
void maximise_ac_sharing::ac_plugin::register_kind(decl_kind k) {
m_kinds.push_back(k);
}
maximise_ac_sharing::ac_plugin::~ac_plugin() {
}
bool maximise_ac_sharing::ac_plugin::reduce(func_decl * f, unsigned num_args, expr * const * args, expr_ref & result) {
decl_kind k = f->get_kind();
if (!f->is_associative())
return false;
if (num_args <= 2)
return false;
if (std::find(m_kinds.begin(), m_kinds.end(), k) == m_kinds.end())
return false;
ptr_buffer<expr, 128> _args;
expr * numeral = 0;
if (m_owner.is_numeral(args[0])) {
numeral = args[0];
for (unsigned i = 1; i < num_args; i++)
_args.push_back(args[i]);
num_args--;
}
else {
_args.append(num_args, args);
}
TRACE("ac_sharing_detail", tout << "before-reuse: num_args: " << num_args << "\n";);
#define MAX_NUM_ARGS_FOR_OPT 128
// Try to reuse already created circuits.
TRACE("ac_sharing_detail", tout << "args: "; for (unsigned i = 0; i < num_args; i++) tout << mk_pp(_args[i], m_manager) << "\n";);
try_to_reuse:
if (num_args > 1 && num_args < MAX_NUM_ARGS_FOR_OPT) {
for (unsigned i = 0; i < num_args - 1; i++) {
for (unsigned j = i + 1; j < num_args; j++) {
if (m_owner.contains(f, _args[i], _args[j])) {
TRACE("ac_sharing_detail", tout << "reusing args: " << i << " " << j << "\n";);
_args[i] = m_manager.mk_app(f, _args[i], _args[j]);
SASSERT(num_args > 1);
for (unsigned w = j; w < num_args - 1; w++) {
_args[w] = _args[w+1];
}
num_args--;
goto try_to_reuse;
}
}
}
}
// Create "tree-like circuit"
while (true) {
TRACE("ac_sharing_detail", tout << "tree-loop: num_args: " << num_args << "\n";);
unsigned j = 0;
for (unsigned i = 0; i < num_args; i += 2, j++) {
if (i == num_args - 1) {
_args[j] = _args[i];
}
else {
m_owner.insert(f, _args[i], _args[i+1]);
_args[j] = m_manager.mk_app(f, _args[i], _args[i+1]);
}
}
num_args = j;
if (num_args == 1) {
if (numeral == 0) {
result = _args[0];
}
else {
result = m_manager.mk_app(f, numeral, _args[0]);
}
TRACE("ac_sharing_detail", tout << "result: " << mk_pp(result, m_manager) << "\n";);
return true;
}
}
UNREACHABLE();
return false;
}
bool maximise_ac_sharing::contains(func_decl * f, expr * arg1, expr * arg2) {
entry e(f, arg1, arg2);
return m_cache.contains(&e);
}
void maximise_ac_sharing::insert(func_decl * f, expr * arg1, expr * arg2) {
entry * e = new (m_region) entry(f, arg1, arg2);
m_entries.push_back(e);
m_cache.insert(e);
m_manager.inc_ref(arg1);
m_manager.inc_ref(arg2);
}
maximise_ac_sharing::maximise_ac_sharing(ast_manager & m):
m_manager(m),
m_simplifier(m),
m_init(false) {
basic_simplifier_plugin* basic_simp = alloc(basic_simplifier_plugin,m);
m_simplifier.register_plugin(basic_simp);
}
maximise_ac_sharing::~maximise_ac_sharing() {
restore_entries(0);
}
void maximise_ac_sharing::operator()(expr * s, expr_ref & r, proof_ref & p) {
init();
m_simplifier.operator()(s, r, p);
}
void maximise_ac_sharing::push_scope() {
init();
m_scopes.push_back(m_entries.size());
m_region.push_scope();
}
void maximise_ac_sharing::pop_scope(unsigned num_scopes) {
SASSERT(num_scopes <= m_scopes.size());
unsigned new_lvl = m_scopes.size() - num_scopes;
unsigned old_lim = m_scopes[new_lvl];
restore_entries(old_lim);
m_region.pop_scope(num_scopes);
m_scopes.shrink(new_lvl);
}
void maximise_ac_sharing::restore_entries(unsigned old_lim) {
unsigned i = m_entries.size();
while (i != old_lim) {
--i;
entry * e = m_entries[i];
m_manager.dec_ref(e->m_arg1);
m_manager.dec_ref(e->m_arg2);
}
m_entries.shrink(old_lim);
}
void maximise_ac_sharing::reset() {
restore_entries(0);
m_entries.reset();
m_cache.reset();
m_simplifier.reset();
m_region.reset();
m_scopes.reset();
}
void maximise_bv_sharing::init_core() {
maximise_ac_sharing::ac_plugin * p = alloc(maximise_ac_sharing::ac_plugin, symbol("bv"), get_manager(), *this);
p->register_kind(OP_BADD);
p->register_kind(OP_BMUL);
p->register_kind(OP_BOR);
p->register_kind(OP_BAND);
register_plugin(p);
}
bool maximise_bv_sharing::is_numeral(expr * n) const {
return m_util.is_numeral(n);
}
maximise_bv_sharing::maximise_bv_sharing(ast_manager & m):
maximise_ac_sharing(m),
m_util(m) {
}

View file

@ -0,0 +1,124 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
maximise_ac_sharing.h
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2008-10-22.
Revision History:
--*/
#ifndef _MAXIMISE_AC_SHARING_H_
#define _MAXIMISE_AC_SHARING_H_
#include"simplifier.h"
#include"hashtable.h"
#include"bv_decl_plugin.h"
#include"region.h"
/**
\brief Functor used to maximise the amount of shared terms in an expression.
The idea is to rewrite AC terms to maximise sharing.
Example:
(f (bvadd a (bvadd b c)) (bvadd a (bvadd b d)))
is rewritten to:
(f (bvadd (bvadd a b) c) (bvadd (bvadd a b) d))
\warning This class uses an opportunistic heuristic to maximise sharing.
There is no guarantee that the optimal expression will be produced.
*/
class maximise_ac_sharing {
struct entry {
func_decl * m_decl;
expr * m_arg1;
expr * m_arg2;
entry(func_decl * d = 0, expr * arg1 = 0, expr * arg2 = 0):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())
std::swap(m_arg1, m_arg2);
}
unsigned hash() const {
unsigned a = m_decl->get_id();
unsigned b = m_arg1->get_id();
unsigned c = m_arg2->get_id();
mix(a,b,c);
return c;
}
bool operator==(entry const & e) const {
return m_decl == e.m_decl && m_arg1 == e.m_arg1 && m_arg2 == e.m_arg2;
}
};
typedef ptr_hashtable<entry, obj_ptr_hash<entry>, deref_eq<entry> > cache;
protected:
class ac_plugin : public simplifier_plugin {
maximise_ac_sharing & m_owner;
svector<decl_kind> m_kinds; //!< kinds to be processed
public:
ac_plugin(symbol const & fname, ast_manager & m, maximise_ac_sharing & owner);
void register_kind(decl_kind k);
virtual ~ac_plugin();
virtual bool reduce(func_decl * f, unsigned num_args, expr * const * args, expr_ref & result);
};
friend class ac_plugin;
private:
ast_manager & m_manager;
simplifier m_simplifier;
bool m_init;
region m_region;
cache m_cache;
ptr_vector<entry> m_entries;
unsigned_vector m_scopes;
bool contains(func_decl * f, expr * arg1, expr * arg2);
void insert(func_decl * f, expr * arg1, expr * arg2);
void restore_entries(unsigned old_lim);
void init() {
if (!m_init) {
init_core();
m_init = true;
}
}
protected:
virtual void init_core() = 0;
virtual bool is_numeral(expr * n) const = 0;
void register_plugin(ac_plugin * p) { m_simplifier.register_plugin(p); }
public:
maximise_ac_sharing(ast_manager & m);
virtual ~maximise_ac_sharing();
void operator()(expr * s, expr_ref & r, proof_ref & p);
void push_scope();
void pop_scope(unsigned num_scopes);
void reset();
ast_manager & get_manager() { return m_manager; }
};
class maximise_bv_sharing : public maximise_ac_sharing {
bv_util m_util;
protected:
virtual void init_core();
virtual bool is_numeral(expr * n) const;
public:
maximise_bv_sharing(ast_manager & m);
};
#endif /* _MAXIMISE_AC_SHARING_H_ */

View file

@ -0,0 +1,213 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
pull_ite_tree.cpp
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2008-06-22.
Revision History:
--*/
#include"pull_ite_tree.h"
#include"recurse_expr_def.h"
#include"for_each_expr.h"
#include"ast_pp.h"
pull_ite_tree::pull_ite_tree(ast_manager & m, simplifier & s):
m_manager(m),
m_simplifier(s),
m_cache(m) {
}
void pull_ite_tree::cache_result(expr * n, expr * r, proof * pr) {
m_cache.insert(n, r, pr);
}
void pull_ite_tree::visit(expr * n, bool & visited) {
if (!is_cached(n)) {
m_todo.push_back(n);
visited = false;
}
}
bool pull_ite_tree::visit_children(expr * n) {
if (m_manager.is_ite(n)) {
bool visited = true;
visit(to_app(n)->get_arg(1), visited);
visit(to_app(n)->get_arg(2), visited);
return visited;
}
else {
return true;
}
}
void pull_ite_tree::reduce(expr * n) {
// Remark: invoking the simplifier to build the new expression saves a lot of memory.
if (m_manager.is_ite(n)) {
expr * c = to_app(n)->get_arg(0);
expr * t_old = to_app(n)->get_arg(1);
expr * e_old = to_app(n)->get_arg(2);
expr * t = 0;
proof * t_pr = 0;
expr * e = 0;
proof * e_pr = 0;
get_cached(t_old, t, t_pr);
get_cached(e_old, e, e_pr);
expr_ref r(m_manager);
expr * args[3] = {c, t, e};
m_simplifier.mk_app(to_app(n)->get_decl(), 3, args, r);
if (!m_manager.proofs_enabled()) {
// expr * r = m_manager.mk_ite(c, t, e);
cache_result(n, r, 0);
}
else {
// t_pr is a proof for (m_p ... t_old ...) == t
// e_pr is a proof for (m_p ... e_old ...) == e
expr_ref old(m_manager);
expr_ref p_t_old(m_manager);
expr_ref p_e_old(m_manager);
old = mk_p_arg(n); // (m_p ... n ...) where n is (ite c t_old e_old)
p_t_old = mk_p_arg(t_old); // (m_p ... t_old ...)
p_e_old = mk_p_arg(e_old); // (m_p ... e_old ...)
expr_ref tmp1(m_manager);
tmp1 = m_manager.mk_ite(c, p_t_old, p_e_old); // (ite c (m_p ... t_old ...) (m_p ... e_old ...))
proof * pr1 = m_manager.mk_rewrite(old, tmp1); // proof for (m_p ... (ite c t_old e_old) ...) = (ite c (m_p ... t_old ...) (m_p ... e_old ...))
expr_ref tmp2(m_manager);
tmp2 = m_manager.mk_ite(c, t, e); // (ite c t e)
proof * pr2 = 0; // it will contain a proof for (ite c (m_p ... t_old ...) (m_p ... e_old ...)) = (ite c t e)
proof * pr3 = 0; // it will contain a proof for (m_p ... (ite c t_old e_old) ...) = (ite c t e)
proof * proofs[2];
unsigned num_proofs = 0;
if (t_pr != 0) {
proofs[num_proofs] = t_pr;
num_proofs++;
}
if (e_pr != 0) {
proofs[num_proofs] = e_pr;
num_proofs++;
}
if (num_proofs > 0) {
pr2 = m_manager.mk_congruence(to_app(tmp1), to_app(tmp2), num_proofs, proofs);
pr3 = m_manager.mk_transitivity(pr1, pr2);
}
else {
pr3 = pr1;
}
proof * pr4 = 0; // it will contain a proof for (ite c t e) = r
proof * pr5 = 0; // it will contain a proof for (m_p ... (ite c t_old e_old) ...) = r
if (tmp2 != r) {
pr4 = m_manager.mk_rewrite(tmp2, r);
pr5 = m_manager.mk_transitivity(pr3, pr4);
}
else {
pr5 = pr3;
}
cache_result(n, r, pr5);
}
}
else {
expr_ref r(m_manager);
m_args[m_arg_idx] = n;
m_simplifier.mk_app(m_p, m_args.size(), m_args.c_ptr(), r);
if (!m_manager.proofs_enabled()) {
// expr * r = m_manager.mk_app(m_p, m_args.size(), m_args.c_ptr());
cache_result(n, r, 0);
}
else {
expr_ref old(m_manager);
proof * p;
old = mk_p_arg(n);
if (old == r)
p = 0;
else
p = m_manager.mk_rewrite(old, r);
cache_result(n, r, p);
}
}
}
void pull_ite_tree::operator()(app * n, app_ref & r, proof_ref & pr) {
unsigned num_args = n->get_num_args();
m_args.resize(num_args);
m_p = n->get_decl();
expr * ite = 0;
for (unsigned i = 0; i < num_args; i++) {
expr * arg = n->get_arg(i);
if (ite) {
m_args[i] = arg;
}
else if (m_manager.is_ite(arg)) {
m_arg_idx = i;
m_args[i] = 0;
ite = arg;
}
else {
m_args[i] = arg;
}
}
if (!ite) {
r = n;
pr = 0;
return;
}
m_todo.push_back(ite);
while (!m_todo.empty()) {
expr * n = m_todo.back();
if (is_cached(n))
m_todo.pop_back();
else if (visit_children(n)) {
m_todo.pop_back();
reduce(n);
}
}
SASSERT(is_cached(ite));
expr * _r = 0;
proof * _pr = 0;
get_cached(ite, _r, _pr);
r = to_app(_r);
pr = _pr;
m_cache.reset();
m_todo.reset();
}
pull_ite_tree_star::pull_ite_tree_star(ast_manager & m, simplifier & s):
simplifier(m),
m_proc(m, s) {
borrow_plugins(s);
}
bool pull_ite_tree_star::get_subst(expr * n, expr_ref & r, proof_ref & p) {
if (is_app(n) && is_target(to_app(n))) {
app_ref tmp(m_manager);
m_proc(to_app(n), tmp, p);
r = tmp;
return true;
}
return false;
}
bool pull_cheap_ite_tree_star::is_target(app * n) const {
bool r =
n->get_num_args() == 2 &&
n->get_family_id() != null_family_id &&
m_manager.is_bool(n) &&
(m_manager.is_value(n->get_arg(0)) || m_manager.is_value(n->get_arg(1))) &&
(m_manager.is_term_ite(n->get_arg(0)) || m_manager.is_term_ite(n->get_arg(1)));
TRACE("pull_ite_target", tout << mk_pp(n, m_manager) << "\nresult: " << r << "\n";);
return r;
}

View file

@ -0,0 +1,101 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
pull_ite_tree.h
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2008-06-22.
Revision History:
--*/
#ifndef _PULL_ITE_TREE_H_
#define _PULL_ITE_TREE_H_
#include"ast.h"
#include"simplifier.h"
#include"recurse_expr.h"
#include"obj_hashtable.h"
#include"expr_map.h"
/**
\brief Functor for applying the following transformation
F[(p (ite c t1 t2) args)] = F'[(ite c t1 t2), p, args]
F'[(ite c t1 t2), p, args] = (ite c F'[t1, p, args] F'[t2, p, args])
F'[t, p, args] = (p t args)
*/
class pull_ite_tree {
ast_manager & m_manager;
simplifier & m_simplifier;
func_decl * m_p;
ptr_vector<expr> m_args;
unsigned m_arg_idx; //!< position of the ite argument
expr_map m_cache;
ptr_vector<expr> m_todo;
bool is_cached(expr * n) const { return m_cache.contains(n); }
void get_cached(expr * n, expr * & r, proof * & p) const { m_cache.get(n, r, p); }
void cache_result(expr * n, expr * r, proof * pr);
void visit(expr * n, bool & visited);
bool visit_children(expr * n);
void reduce(expr * n);
/**
\brief Creante an application (m_p ... n ...) where n is the argument m_arg_idx and the other arguments
are in m_args.
*/
expr * mk_p_arg(expr * n) {
m_args[m_arg_idx] = n;
return m_manager.mk_app(m_p, m_args.size(), m_args.c_ptr());
}
public:
pull_ite_tree(ast_manager & m, simplifier & s);
/**
\brief Apply the transformation above if n contains an ite-expression.
Store the result in r. If n does not contain an ite-expression, then
store n in r.
When proof generation is enabled, pr is a proof for n = r.
*/
void operator()(app * n, app_ref & r, proof_ref & pr);
};
/**
\brief Functor for applying the pull_ite_tree on subexpressions n that
satisfy the is_target virtual predicate.
*/
class pull_ite_tree_star : public simplifier {
protected:
pull_ite_tree m_proc;
virtual bool get_subst(expr * n, expr_ref & r, proof_ref & p);
public:
pull_ite_tree_star(ast_manager & m, simplifier & s);
virtual ~pull_ite_tree_star() { release_plugins(); }
virtual bool is_target(app * n) const = 0;
};
/**
\brief Apply pull_ite_tree on predicates of the form
(p ite v) and (p v ite)
where:
- p is an interpreted predicate
- ite is an ite-term expression
- v is a value
*/
class pull_cheap_ite_tree_star : public pull_ite_tree_star {
public:
pull_cheap_ite_tree_star(ast_manager & m, simplifier & s):pull_ite_tree_star(m, s) {}
virtual ~pull_cheap_ite_tree_star() {}
virtual bool is_target(app * n) const;
};
#endif /* _PULL_ITE_TREE_H_ */

View file

@ -0,0 +1,76 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
theory_array_params.h
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2008-06-01.
Revision History:
--*/
#ifndef _THEORY_ARRAY_PARAMS_H_
#define _THEORY_ARRAY_PARAMS_H_
#include"ini_file.h"
enum array_solver_id {
AR_NO_ARRAY,
AR_SIMPLE,
AR_MODEL_BASED,
AR_FULL
};
struct theory_array_params {
array_solver_id m_array_mode;
bool m_array_weak;
bool m_array_extensional;
unsigned m_array_laziness;
bool m_array_delay_exp_axiom;
bool m_array_cg;
bool m_array_always_prop_upward;
bool m_array_lazy_ieq;
unsigned m_array_lazy_ieq_delay;
bool m_array_canonize_simplify;
bool m_array_simplify; // temporary hack for disabling array simplifier plugin.
theory_array_params():
m_array_mode(AR_FULL),
m_array_weak(false),
m_array_extensional(true),
m_array_laziness(1),
m_array_delay_exp_axiom(true),
m_array_cg(false),
m_array_always_prop_upward(true), // UPWARDs filter is broken... TODO: fix it
m_array_lazy_ieq(false),
m_array_lazy_ieq_delay(10),
m_array_canonize_simplify(false),
m_array_simplify(true) {
}
void register_params(ini_params & p) {
p.register_int_param("ARRAY_SOLVER", 0, 3, reinterpret_cast<int&>(m_array_mode), "0 - no array, 1 - simple, 2 - model based, 3 - full");
p.register_bool_param("ARRAY_WEAK", m_array_weak);
p.register_bool_param("ARRAY_EXTENSIONAL", m_array_extensional);
p.register_unsigned_param("ARRAY_LAZINESS", m_array_laziness);
p.register_bool_param("ARRAY_DELAY_EXP_AXIOM", m_array_delay_exp_axiom);
p.register_bool_param("ARRAY_CG", m_array_cg);
p.register_bool_param("ARRAY_ALWAYS_PROP_UPWARD", m_array_always_prop_upward,
"Disable the built-in filter upwards propagation");
p.register_bool_param("ARRAY_LAZY_IEQ", m_array_lazy_ieq);
p.register_unsigned_param("ARRAY_LAZY_IEQ_DELAY", m_array_lazy_ieq_delay);
p.register_bool_param("ARRAY_CANONIZE", m_array_canonize_simplify,
"Normalize arrays into normal form during simplification");
}
};
#endif /* _THEORY_ARRAY_PARAMS_H_ */