3
0
Fork 0
mirror of https://github.com/Z3Prover/z3 synced 2025-12-01 17:49:04 +00:00

split muz_qe into two directories

Signed-off-by: Nikolaj Bjorner <nbjorner@microsoft.com>
This commit is contained in:
Nikolaj Bjorner 2013-08-28 12:08:47 -07:00
parent d795792304
commit 137339a2e1
174 changed files with 242 additions and 174 deletions

1
src/muz/README Normal file
View file

@ -0,0 +1 @@
muZ and Quantifier Elimination modules

328
src/muz/aig_exporter.cpp Normal file
View file

@ -0,0 +1,328 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
aig_exporter.cpp
Abstract:
Export AIG files from horn clauses
--*/
#include "aig_exporter.h"
#include "dl_context.h"
#include <set>
namespace datalog {
aig_exporter::aig_exporter(const rule_set& rules, context& ctx, const fact_vector *facts) :
m_rules(rules), m_facts(facts), m(ctx.get_manager()), m_rm(ctx.get_rule_manager()),
m_aigm(m), m_next_decl_id(1), m_next_aig_expr_id(2), m_num_and_gates(0),
m_latch_vars(m), m_latch_varsp(m), m_ruleid_var_set(m), m_ruleid_varp_set(m)
{
std::set<func_decl*> predicates;
for (rule_set::decl2rules::iterator I = m_rules.begin_grouped_rules(),
E = m_rules.end_grouped_rules(); I != E; ++I) {
predicates.insert(I->m_key);
}
for (fact_vector::const_iterator I = facts->begin(), E = facts->end(); I != E; ++I) {
predicates.insert(I->first);
}
// reserve pred id = 0 for initalization purposes
unsigned num_preds = (unsigned)predicates.size() + 1;
// poor's man round-up log2
unsigned preds_bitsize = log2(num_preds);
if ((1U << preds_bitsize) < num_preds)
++preds_bitsize;
SASSERT((1U << preds_bitsize) >= num_preds);
for (unsigned i = 0; i < preds_bitsize; ++i) {
m_ruleid_var_set.push_back(m.mk_fresh_const("rule_id", m.mk_bool_sort()));
m_ruleid_varp_set.push_back(m.mk_fresh_const("rule_id_p", m.mk_bool_sort()));
}
}
void aig_exporter::mk_latch_vars(unsigned n) {
for (unsigned i = m_latch_vars.size(); i <= n; ++i) {
m_latch_vars.push_back(m.mk_fresh_const("latch_var", m.mk_bool_sort()));
m_latch_varsp.push_back(m.mk_fresh_const("latch_varp", m.mk_bool_sort()));
}
SASSERT(m_latch_vars.size() > n);
}
expr* aig_exporter::get_latch_var(unsigned i, const expr_ref_vector& vars) {
mk_latch_vars(i);
return vars.get(i);
}
void aig_exporter::assert_pred_id(func_decl *decl, const expr_ref_vector& vars, expr_ref_vector& exprs) {
unsigned id = 0;
if (decl && !m_decl_id_map.find(decl, id)) {
id = m_next_decl_id++;
SASSERT(id < (1U << vars.size()));
m_decl_id_map.insert(decl, id);
}
for (unsigned i = 0; i < vars.size(); ++i) {
exprs.push_back((id & (1U << i)) ? vars[i] : m.mk_not(vars[i]));
}
}
void aig_exporter::collect_var_substs(substitution& subst, const app *h,
const expr_ref_vector& vars, expr_ref_vector& eqs) {
for (unsigned i = 0; i < h->get_num_args(); ++i) {
expr *arg = h->get_arg(i);
expr *latchvar = get_latch_var(i, vars);
if (is_var(arg)) {
var *v = to_var(arg);
expr_offset othervar;
if (subst.find(v, 0, othervar)) {
eqs.push_back(m.mk_eq(latchvar, othervar.get_expr()));
} else {
subst.insert(v, 0, expr_offset(latchvar, 0));
}
} else {
eqs.push_back(m.mk_eq(latchvar, arg));
}
}
}
void aig_exporter::operator()(std::ostream& out) {
expr_ref_vector transition_function(m), output_preds(m);
var_ref_vector input_vars(m);
rule_counter& vc = m_rm.get_counter();
expr_ref_vector exprs(m);
substitution subst(m);
for (rule_set::decl2rules::iterator I = m_rules.begin_grouped_rules(),
E = m_rules.end_grouped_rules(); I != E; ++I) {
for (rule_vector::iterator II = I->get_value()->begin(),
EE = I->get_value()->end(); II != EE; ++II) {
rule *r = *II;
unsigned numqs = r->get_positive_tail_size();
if (numqs > 1) {
std::cerr << "non-linear clauses not supported\n";
exit(-1);
}
if (numqs != r->get_uninterpreted_tail_size()) {
std::cerr << "negation of queries not supported\n";
exit(-1);
}
exprs.reset();
assert_pred_id(numqs ? r->get_tail(0)->get_decl() : 0, m_ruleid_var_set, exprs);
assert_pred_id(r->get_head()->get_decl(), m_ruleid_varp_set, exprs);
subst.reset();
subst.reserve(1, vc.get_max_rule_var(*r)+1);
if (numqs)
collect_var_substs(subst, r->get_tail(0), m_latch_vars, exprs);
collect_var_substs(subst, r->get_head(), m_latch_varsp, exprs);
for (unsigned i = numqs; i < r->get_tail_size(); ++i) {
expr_ref e(m);
subst.apply(r->get_tail(i), e);
exprs.push_back(e);
}
transition_function.push_back(m.mk_and(exprs.size(), exprs.c_ptr()));
}
}
// collect table facts
if (m_facts) {
for (fact_vector::const_iterator I = m_facts->begin(), E = m_facts->end(); I != E; ++I) {
exprs.reset();
assert_pred_id(0, m_ruleid_var_set, exprs);
assert_pred_id(I->first, m_ruleid_varp_set, exprs);
for (unsigned i = 0; i < I->second.size(); ++i) {
exprs.push_back(m.mk_eq(get_latch_var(i, m_latch_varsp), I->second[i]));
}
transition_function.push_back(m.mk_and(exprs.size(), exprs.c_ptr()));
}
}
expr *tr = m.mk_or(transition_function.size(), transition_function.c_ptr());
aig_ref aig = m_aigm.mk_aig(tr);
expr_ref aig_expr(m);
m_aigm.to_formula(aig, aig_expr);
#if 0
std::cout << mk_pp(tr, m) << "\n\n";
std::cout << mk_pp(aig_expr, m) << "\n\n";
#endif
// make rule_id vars latches
for (unsigned i = 0; i < m_ruleid_var_set.size(); ++i) {
m_latch_vars.push_back(m_ruleid_var_set.get(i));
m_latch_varsp.push_back(m_ruleid_varp_set.get(i));
}
// create vars for latches
for (unsigned i = 0; i < m_latch_vars.size(); ++i) {
mk_var(m_latch_vars.get(i));
mk_input_var(m_latch_varsp.get(i));
}
unsigned tr_id = expr_to_aig(aig_expr);
// create latch next state variables: (ite tr varp var)
unsigned_vector latch_varp_ids;
for (unsigned i = 0; i < m_latch_vars.size(); ++i) {
unsigned in_val = mk_and(tr_id, get_var(m_latch_varsp.get(i)));
unsigned latch_val = mk_and(neg(tr_id), get_var(m_latch_vars.get(i)));
latch_varp_ids.push_back(mk_or(in_val, latch_val));
}
m_latch_varsp.reset();
// create output variable (true iff an output predicate is derivable)
unsigned output_id = 0;
{
expr_ref_vector output(m);
const func_decl_set& preds = m_rules.get_output_predicates();
for (func_decl_set::iterator I = preds.begin(), E = preds.end(); I != E; ++I) {
exprs.reset();
assert_pred_id(*I, m_ruleid_var_set, exprs);
output.push_back(m.mk_and(exprs.size(), exprs.c_ptr()));
}
expr *out = m.mk_or(output.size(), output.c_ptr());
aig = m_aigm.mk_aig(out);
m_aigm.to_formula(aig, aig_expr);
output_id = expr_to_aig(aig_expr);
#if 0
std::cout << "output formula\n";
std::cout << mk_pp(out, m) << "\n\n";
std::cout << mk_pp(aig_expr, m) << "\n\n";
#endif
}
// 1) print header
// aag var_index inputs latches outputs andgates
out << "aag " << (m_next_aig_expr_id-1)/2 << ' ' << m_input_vars.size()
<< ' ' << m_latch_vars.size() << " 1 " << m_num_and_gates << '\n';
// 2) print inputs
for (unsigned i = 0; i < m_input_vars.size(); ++i) {
out << m_input_vars[i] << '\n';
}
// 3) print latches
for (unsigned i = 0; i < m_latch_vars.size(); ++i) {
out << get_var(m_latch_vars.get(i)) << ' ' << latch_varp_ids[i] << '\n';
}
// 4) print outputs (just one for now)
out << output_id << '\n';
// 5) print formula
out << m_buffer.str();
}
unsigned aig_exporter::expr_to_aig(const expr *e) {
unsigned id;
if (m_aig_expr_id_map.find(e, id))
return id;
if (is_uninterp_const(e))
return get_var(e);
switch (e->get_kind()) {
case AST_APP: {
const app *a = to_app(e);
switch (a->get_decl_kind()) {
case OP_OR:
SASSERT(a->get_num_args() > 0);
id = expr_to_aig(a->get_arg(0));
for (unsigned i = 1; i < a->get_num_args(); ++i) {
id = mk_or(id, expr_to_aig(a->get_arg(i)));
}
m_aig_expr_id_map.insert(e, id);
return id;
case OP_NOT:
return neg(expr_to_aig(a->get_arg(0)));
case OP_FALSE:
return 0;
case OP_TRUE:
return 1;
}
break;}
case AST_VAR:
return get_var(e);
default:
UNREACHABLE();
}
UNREACHABLE();
return 0;
}
unsigned aig_exporter::neg(unsigned id) const {
return (id % 2) ? (id-1) : (id+1);
}
unsigned aig_exporter::mk_and(unsigned id1, unsigned id2) {
if (id1 > id2)
std::swap(id1, id2);
std::pair<unsigned,unsigned> key(id1, id2);
and_gates_map::const_iterator I = m_and_gates_map.find(key);
if (I != m_and_gates_map.end())
return I->second;
unsigned id = mk_expr_id();
m_buffer << id << ' ' << id1 << ' ' << id2 << '\n';
m_and_gates_map[key] = id;
++m_num_and_gates;
return id;
}
unsigned aig_exporter::mk_or(unsigned id1, unsigned id2) {
return neg(mk_and(neg(id1), neg(id2)));
}
unsigned aig_exporter::get_var(const expr *e) {
unsigned id;
if (m_aig_expr_id_map.find(e, id))
return id;
return mk_input_var(e);
}
unsigned aig_exporter::mk_var(const expr *e) {
SASSERT(!m_aig_expr_id_map.contains(e));
unsigned id = mk_expr_id();
m_aig_expr_id_map.insert(e, id);
return id;
}
unsigned aig_exporter::mk_input_var(const expr *e) {
SASSERT(!m_aig_expr_id_map.contains(e));
unsigned id = mk_expr_id();
m_input_vars.push_back(id);
if (e)
m_aig_expr_id_map.insert(e, id);
return id;
}
unsigned aig_exporter::mk_expr_id() {
unsigned id = m_next_aig_expr_id;
m_next_aig_expr_id += 2;
return id;
}
}

68
src/muz/aig_exporter.h Normal file
View file

@ -0,0 +1,68 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
aig_exporter.h
Abstract:
Export AIG files from horn clauses
--*/
#ifndef _AIG_EXPORTER_H_
#define _AIG_EXPORTER_H_
#include "aig.h"
#include "dl_rule_set.h"
#include "rel_context.h"
#include <map>
#include <sstream>
namespace datalog {
class aig_exporter {
public:
aig_exporter(const rule_set& rules, context& ctx, const fact_vector *facts = 0);
void operator()(std::ostream& out);
private:
typedef obj_map<func_decl, unsigned> decl_id_map;
typedef obj_map<const expr, unsigned> aig_expr_id_map;
typedef std::map<std::pair<unsigned,unsigned>, unsigned> and_gates_map;
const rule_set& m_rules;
const fact_vector *m_facts;
ast_manager& m;
rule_manager& m_rm;
aig_manager m_aigm;
decl_id_map m_decl_id_map;
unsigned m_next_decl_id;
aig_expr_id_map m_aig_expr_id_map;
unsigned m_next_aig_expr_id;
and_gates_map m_and_gates_map;
unsigned m_num_and_gates;
expr_ref_vector m_latch_vars, m_latch_varsp;
expr_ref_vector m_ruleid_var_set, m_ruleid_varp_set;
unsigned_vector m_input_vars;
std::stringstream m_buffer;
void mk_latch_vars(unsigned n);
expr* get_latch_var(unsigned i, const expr_ref_vector& vars);
void assert_pred_id(func_decl *decl, const expr_ref_vector& vars, expr_ref_vector& exprs);
void collect_var_substs(substitution& subst, const app *h,
const expr_ref_vector& vars, expr_ref_vector& eqs);
unsigned expr_to_aig(const expr *e);
unsigned neg(unsigned id) const;
unsigned mk_and(unsigned id1, unsigned id2);
unsigned mk_or(unsigned id1, unsigned id2);
unsigned get_var(const expr *e);
unsigned mk_var(const expr *e);
unsigned mk_input_var(const expr *e = 0);
unsigned mk_expr_id();
};
}
#endif

View file

@ -0,0 +1,160 @@
#include"arith_bounds_tactic.h"
#include"arith_decl_plugin.h"
struct arith_bounds_tactic : public tactic {
ast_manager& m;
arith_util a;
volatile bool m_cancel;
arith_bounds_tactic(ast_manager& m):
m(m),
a(m),
m_cancel(false)
{
}
ast_manager& get_manager() { return m; }
void set_cancel(bool f) {
m_cancel = f;
}
virtual void cleanup() {
m_cancel = false;
}
virtual void operator()(/* in */ goal_ref const & in,
/* out */ goal_ref_buffer & result,
/* out */ model_converter_ref & mc,
/* out */ proof_converter_ref & pc,
/* out */ expr_dependency_ref & core) {
bounds_arith_subsumption(in, result);
}
virtual tactic* translate(ast_manager& m) {
return alloc(arith_bounds_tactic, m);
}
void checkpoint() {
if (m_cancel) {
throw tactic_exception(TACTIC_CANCELED_MSG);
}
}
struct info { rational r; unsigned idx; bool is_strict;};
/**
\brief Basic arithmetic subsumption simplification based on bounds.
*/
void mk_proof(proof_ref& pr, goal_ref const& s, unsigned i, unsigned j) {
if (s->proofs_enabled()) {
proof* th_lemma = m.mk_th_lemma(a.get_family_id(), m.mk_implies(s->form(i), s->form(j)), 0, 0);
pr = m.mk_modus_ponens(s->pr(i), th_lemma);
}
}
bool is_le_or_lt(expr* e, expr*& e1, expr*& e2, bool& is_strict) {
bool is_negated = m.is_not(e, e);
if ((!is_negated && (a.is_le(e, e1, e2) || a.is_ge(e, e2, e1))) ||
(is_negated && (a.is_lt(e, e2, e1) || a.is_gt(e, e1, e2)))) {
is_strict = false;
return true;
}
if ((!is_negated && (a.is_lt(e, e1, e2) || a.is_gt(e, e2, e1))) ||
(is_negated && (a.is_le(e, e2, e1) || a.is_ge(e, e1, e2)))) {
is_strict = true;
return true;
}
return false;
}
void bounds_arith_subsumption(goal_ref const& g, goal_ref_buffer& result) {
info inf;
rational r;
goal_ref s(g); // initialize result.
obj_map<expr, info> lower, upper;
expr* e1, *e2;
TRACE("arith_subsumption", s->display(tout); );
for (unsigned i = 0; i < s->size(); ++i) {
checkpoint();
expr* lemma = s->form(i);
bool is_strict = false;
bool is_lower = false;
if (!is_le_or_lt(lemma, e1, e2, is_strict)) {
continue;
}
// e1 <= e2 or e1 < e2
if (a.is_numeral(e2, r)) {
is_lower = true;
}
else if (a.is_numeral(e1, r)) {
is_lower = false;
}
else {
continue;
}
proof_ref new_pr(m);
if (is_lower && upper.find(e1, inf)) {
if (inf.r > r || (inf.r == r && is_strict && !inf.is_strict)) {
mk_proof(new_pr, s, i, inf.idx);
s->update(inf.idx, m.mk_true(), new_pr);
inf.r = r;
inf.is_strict = is_strict;
inf.idx = i;
upper.insert(e1, inf);
}
else {
mk_proof(new_pr, s, inf.idx, i);
s->update(i, m.mk_true(), new_pr);
}
}
else if (is_lower) {
inf.r = r;
inf.is_strict = is_strict;
inf.idx = i;
upper.insert(e1, inf);
}
else if (!is_lower && lower.find(e2, inf)) {
if (inf.r < r || (inf.r == r && is_strict && !inf.is_strict)) {
mk_proof(new_pr, s, i, inf.idx);
s->update(inf.idx, m.mk_true(), new_pr);
inf.r = r;
inf.is_strict = is_strict;
inf.idx = i;
lower.insert(e2, inf);
}
else {
mk_proof(new_pr, s, inf.idx, i);
s->update(i, m.mk_true());
}
}
else if (!is_lower) {
inf.r = r;
inf.is_strict = is_strict;
inf.idx = i;
lower.insert(e2, inf);
}
}
s->elim_true();
result.push_back(s.get());
TRACE("arith_subsumption", s->display(tout); );
}
};
tactic * mk_arith_bounds_tactic(ast_manager & m, params_ref const & p) {
return alloc(arith_bounds_tactic, m);
}

View file

@ -0,0 +1,37 @@
/*++
Copyright (c) 2012 Microsoft Corporation
Module Name:
arith_bounds_tactic.h
Abstract:
Fast/rudimentary arithmetic subsumption tactic.
Author:
Nikolaj Bjorner (nbjorner) 2012-9-6
Notes:
Background: The Farkas learner in PDR generates tons
of inequalities that contain redundancies.
It therefore needs a fast way to reduce these redundancies before
passing the results to routines that are more expensive.
The arith subsumption_strategy encapsulates a rudimentary
routine for simplifying inequalities. Additional simplification
routines can be added here or composed with this strategy.
Note: The bound_manager subsumes some of the collection methods used
for assembling bounds, but it does not have a way to check for
subsumption of atoms.
--*/
#ifndef _ARITH_BOUNDS_TACTIC_H_
#define _ARITH_BOUNDS_TACTIC_H_
#include "tactic.h"
tactic * mk_arith_bounds_tactic(ast_manager & m, params_ref const & p = params_ref());
#endif

240
src/muz/clp_context.cpp Normal file
View file

@ -0,0 +1,240 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
clp_context.cpp
Abstract:
Bounded CLP (symbolic simulation using Z3) context.
Author:
Nikolaj Bjorner (nbjorner) 2013-04-26
Revision History:
--*/
#include "clp_context.h"
#include "dl_context.h"
#include "unifier.h"
#include "var_subst.h"
#include "substitution.h"
namespace datalog {
class clp::imp {
struct stats {
stats() { reset(); }
void reset() { memset(this, 0, sizeof(*this)); }
unsigned m_num_unfold;
unsigned m_num_no_unfold;
unsigned m_num_subsumed;
};
context& m_ctx;
ast_manager& m;
rule_manager& rm;
smt_params m_fparams;
smt::kernel m_solver;
var_subst m_var_subst;
expr_ref_vector m_ground;
app_ref_vector m_goals;
volatile bool m_cancel;
stats m_stats;
public:
imp(context& ctx):
m_ctx(ctx),
m(ctx.get_manager()),
rm(ctx.get_rule_manager()),
m_solver(m, m_fparams), // TBD: can be replaced by efficient BV solver.
m_var_subst(m, false),
m_ground(m),
m_goals(m),
m_cancel(false)
{
// m_fparams.m_relevancy_lvl = 0;
m_fparams.m_mbqi = false;
m_fparams.m_soft_timeout = 1000;
}
~imp() {}
lbool query(expr* query) {
m_ctx.ensure_opened();
m_solver.reset();
m_goals.reset();
rm.mk_query(query, m_ctx.get_rules());
m_ctx.apply_default_transformation();
func_decl* head_decl = m_ctx.get_rules().get_output_predicate();
rule_set& rules = m_ctx.get_rules();
rule_vector const& rv = rules.get_predicate_rules(head_decl);
if (rv.empty()) {
return l_false;
}
expr_ref head(rv[0]->get_head(), m);
ground(head);
m_goals.push_back(to_app(head));
return search(20, 0);
}
void cancel() {
m_cancel = true;
m_solver.cancel();
}
void cleanup() {
m_cancel = false;
m_goals.reset();
m_solver.reset_cancel();
}
void reset_statistics() {
m_stats.reset();
}
void collect_statistics(statistics& st) const {
//st.update("tab.num_unfold", m_stats.m_num_unfold);
//st.update("tab.num_unfold_fail", m_stats.m_num_no_unfold);
//st.update("tab.num_subsumed", m_stats.m_num_subsumed);
}
void display_certificate(std::ostream& out) const {
expr_ref ans = get_answer();
out << mk_pp(ans, m) << "\n";
}
expr_ref get_answer() const {
return expr_ref(m.mk_true(), m);
}
private:
void reset_ground() {
m_ground.reset();
}
void ground(expr_ref& e) {
ptr_vector<sort> sorts;
get_free_vars(e, sorts);
if (m_ground.size() < sorts.size()) {
m_ground.resize(sorts.size());
}
for (unsigned i = 0; i < sorts.size(); ++i) {
if (sorts[i] && !m_ground[i].get()) {
m_ground[i] = m.mk_fresh_const("c",sorts[i]);
}
}
m_var_subst(e, m_ground.size(), m_ground.c_ptr(), e);
}
static bool rule_sort_fn(const rule *r1, const rule *r2) {
return r1->get_uninterpreted_tail_size() < r2->get_uninterpreted_tail_size();
}
lbool search(unsigned depth, unsigned index) {
if (index == m_goals.size()) {
return l_true;
}
if (depth == 0) {
return l_undef;
}
IF_VERBOSE(1, verbose_stream() << "search " << depth << " " << index << "\n";);
unsigned num_goals = m_goals.size();
app* head = m_goals[index].get();
rule_vector rules(m_ctx.get_rules().get_predicate_rules(head->get_decl()));
std::stable_sort(rules.begin(), rules.end(), rule_sort_fn);
lbool status = l_false;
for (unsigned i = 0; i < rules.size(); ++i) {
rule* r = rules[i];
m_solver.push();
reset_ground();
expr_ref tmp(m);
tmp = r->get_head();
IF_VERBOSE(2, verbose_stream() << index << " " << mk_pp(tmp, m) << "\n";);
ground(tmp);
for (unsigned j = 0; j < head->get_num_args(); ++j) {
expr_ref eq(m);
eq = m.mk_eq(head->get_arg(j), to_app(tmp)->get_arg(j));
m_solver.assert_expr(eq);
}
for (unsigned j = r->get_uninterpreted_tail_size(); j < r->get_tail_size(); ++j) {
tmp = r->get_tail(j);
ground(tmp);
m_solver.assert_expr(tmp);
}
lbool is_sat = m_solver.check();
switch (is_sat) {
case l_false:
break;
case l_true:
if (depth == 1 && (index+1 > m_goals.size() || r->get_uninterpreted_tail_size() > 0)) {
status = l_undef;
break;
}
for (unsigned j = 0; j < r->get_uninterpreted_tail_size(); ++j) {
tmp = r->get_tail(j);
ground(tmp);
m_goals.push_back(to_app(tmp));
}
switch(search(depth-1, index+1)) {
case l_undef:
status = l_undef;
// fallthrough
case l_false:
m_goals.resize(num_goals);
break;
case l_true:
return l_true;
}
break;
case l_undef:
status = l_undef;
throw default_exception("undef");
}
m_solver.pop(1);
}
return status;
}
proof_ref get_proof() const {
return proof_ref(0, m);
}
};
clp::clp(context& ctx):
engine_base(ctx.get_manager(), "clp"),
m_imp(alloc(imp, ctx)) {
}
clp::~clp() {
dealloc(m_imp);
}
lbool clp::query(expr* query) {
return m_imp->query(query);
}
void clp::cancel() {
m_imp->cancel();
}
void clp::cleanup() {
m_imp->cleanup();
}
void clp::reset_statistics() {
m_imp->reset_statistics();
}
void clp::collect_statistics(statistics& st) const {
m_imp->collect_statistics(st);
}
void clp::display_certificate(std::ostream& out) const {
m_imp->display_certificate(out);
}
expr_ref clp::get_answer() {
return m_imp->get_answer();
}
};

46
src/muz/clp_context.h Normal file
View file

@ -0,0 +1,46 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
clp_context.h
Abstract:
Bounded CLP (symbolic simulation using Z3) context.
Author:
Nikolaj Bjorner (nbjorner) 2013-04-26
Revision History:
--*/
#ifndef _CLP_CONTEXT_H_
#define _CLP_CONTEXT_H_
#include "ast.h"
#include "lbool.h"
#include "statistics.h"
#include "dl_util.h"
namespace datalog {
class context;
class clp : public datalog::engine_base {
class imp;
imp* m_imp;
public:
clp(context& ctx);
~clp();
virtual lbool query(expr* query);
virtual void cancel();
virtual void cleanup();
virtual void reset_statistics();
virtual void collect_statistics(statistics& st) const;
virtual void display_certificate(std::ostream& out) const;
virtual expr_ref get_answer();
};
};
#endif

1508
src/muz/datalog_parser.cpp Normal file

File diff suppressed because it is too large Load diff

48
src/muz/datalog_parser.h Normal file
View file

@ -0,0 +1,48 @@
/*++
Copyright (c) 2010 Microsoft Corporation
Module Name:
datalog_parser.h
Abstract:
Parser for Datalogish files
Author:
Nikolaj Bjorner (nbjorner) 2010-5-17
Revision History:
--*/
#ifndef _DATALOG_PARSER_H_
#define _DATALOG_PARSER_H_
#include "ast.h"
#include "dl_context.h"
namespace datalog {
class parser {
public:
static parser * create(context& ctx, ast_manager & ast_manager);
virtual ~parser() {}
virtual bool parse_file(char const * path) = 0;
virtual bool parse_string(char const * string) = 0;
};
class wpa_parser {
public:
static wpa_parser * create(context& ctx, ast_manager & ast_manager);
virtual ~wpa_parser() {}
virtual bool parse_directory(char const * path) = 0;
};
};
#endif

459
src/muz/dl_base.cpp Normal file
View file

@ -0,0 +1,459 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_base.cpp
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-09-14.
Revision History:
--*/
#include"ast_pp.h"
#include"union_find.h"
#include"vector.h"
#include"dl_context.h"
#include"dl_base.h"
#include"bool_rewriter.h"
#include<sstream>
namespace datalog {
context & get_context_from_rel_manager(const relation_manager & rm) {
return rm.get_context();
}
ast_manager & get_ast_manager_from_rel_manager(const relation_manager & rm) {
return rm.get_context().get_manager();
}
#if DL_LEAK_HUNTING
void leak_guard_check(const symbol & s) {
}
#endif
void relation_signature::output(ast_manager & m, std::ostream & out) const {
unsigned sz=size();
out<<"(";
for(unsigned i=0; i<sz; i++) {
if(i) { out<<","; }
out << mk_pp((*this)[i], m);
}
out<<")";
}
relation_fact::relation_fact(context & ctx) : app_ref_vector(ctx.get_manager()) {}
void relation_base::reset() {
ast_manager & m = get_plugin().get_ast_manager();
app_ref bottom_ref(m.mk_false(), m);
scoped_ptr<relation_mutator_fn> reset_fn =
get_manager().mk_filter_interpreted_fn(static_cast<relation_base &>(*this), bottom_ref);
if(!reset_fn) {
NOT_IMPLEMENTED_YET();
}
(*reset_fn)(*this);
}
void table_signature::from_join(const table_signature & s1, const table_signature & s2, unsigned col_cnt,
const unsigned * cols1, const unsigned * cols2, table_signature & result) {
result.reset();
unsigned s1sz=s1.size();
unsigned s2sz=s2.size();
unsigned s1first_func=s1sz-s1.functional_columns();
unsigned s2first_func=s2sz-s2.functional_columns();
for(unsigned i=0; i<s1first_func; i++) {
result.push_back(s1[i]);
}
for(unsigned i=0; i<s2first_func; i++) {
result.push_back(s2[i]);
}
for(unsigned i=s1first_func; i<s1sz; i++) {
result.push_back(s1[i]);
}
for(unsigned i=s2first_func; i<s2sz; i++) {
result.push_back(s2[i]);
}
result.set_functional_columns(s1.functional_columns()+s2.functional_columns());
}
void table_signature::from_project(const table_signature & src, unsigned col_cnt,
const unsigned * removed_cols, table_signature & result) {
signature_base::from_project(src, col_cnt, removed_cols, result);
unsigned func_cnt = src.functional_columns();
if(removed_cols==0) {
result.set_functional_columns(func_cnt);
return;
}
unsigned first_src_fun = src.first_functional();
if(removed_cols[0]<first_src_fun) {
//if we remove at least one non-functional column, all the columns in the result are non-functional
result.set_functional_columns(0);
}
else {
//all columns we are removing are functional
SASSERT(func_cnt>=col_cnt);
result.set_functional_columns(func_cnt-col_cnt);
}
}
void table_signature::from_project_with_reduce(const table_signature & src, unsigned col_cnt,
const unsigned * removed_cols, table_signature & result) {
signature_base::from_project(src, col_cnt, removed_cols, result);
unsigned remaining_fun = src.functional_columns();
unsigned first_src_fun = src.first_functional();
for(int i=col_cnt-1; i>=0; i--) {
if(removed_cols[i]<first_src_fun) {
break;
}
remaining_fun--;
}
result.set_functional_columns(remaining_fun);
}
void table_signature::from_join_project(const table_signature & s1, const table_signature & s2,
unsigned joined_col_cnt, const unsigned * cols1, const unsigned * cols2, unsigned removed_col_cnt,
const unsigned * removed_cols, table_signature & result) {
table_signature aux;
from_join(s1, s2, joined_col_cnt, cols1, cols2, aux);
//after the join the column order is
//(non-functional of s1)(non-functional of s2)(functional of s1)(functional of s2)
if(s1.functional_columns()==0 && s2.functional_columns()==0) {
from_project(aux, removed_col_cnt, removed_cols, result);
SASSERT(result.functional_columns()==0);
return;
}
unsigned join_sig_sz = s1.size()+s2.size();
unsigned s1_first_func = s1.first_functional();
unsigned s2_first_func = s2.first_functional();
unsigned second_ofs = s1_first_func;
unsigned first_func_ofs = second_ofs + s2_first_func;
unsigned second_func_ofs = second_ofs + s1.functional_columns();
svector<unsigned> remaining_in_equivalence_class;
remaining_in_equivalence_class.resize(join_sig_sz, 0);
bool merging_rows_can_happen = false;
union_find_default_ctx uf_ctx;
union_find<> uf(uf_ctx); //the numbers in uf correspond to column indexes after the join
for(unsigned i=0; i<join_sig_sz; i++) {
unsigned v = uf.mk_var();
SASSERT(v==i);
}
for(unsigned i=0; i<joined_col_cnt; i++) {
unsigned idx1 = (s1_first_func>cols1[i]) ? cols1[i] : (first_func_ofs+cols1[i]-s1_first_func);
unsigned idx2 = (s2_first_func>cols2[i]) ? (second_ofs+cols2[i]) : (second_func_ofs+cols2[i]-s2_first_func);
uf.merge(idx1, idx2);
}
for(unsigned i=0; i<first_func_ofs; i++) { //we only count the non-functional columns
remaining_in_equivalence_class[uf.find(i)]++;
}
for(unsigned i=0; i<removed_col_cnt; i++) {
unsigned rc = removed_cols[i];
if(rc>=first_func_ofs) {
//removing functional columns won't make us merge rows
continue;
}
unsigned eq_class_idx = uf.find(rc);
if(remaining_in_equivalence_class[eq_class_idx]>1) {
remaining_in_equivalence_class[eq_class_idx]--;
}
else {
merging_rows_can_happen = true;
break;
}
}
if(merging_rows_can_happen) {
//this one marks all columns as non-functional
from_project(aux, removed_col_cnt, removed_cols, result);
SASSERT(result.functional_columns()==0);
}
else {
//this one preserves columns to be functional
from_project_with_reduce(aux, removed_col_cnt, removed_cols, result);
}
}
// -----------------------------------
//
// table_base
//
// -----------------------------------
//here we give generic implementation of table operations using iterators
bool table_base::empty() const {
return begin()==end();
}
void table_base::remove_facts(unsigned fact_cnt, const table_fact * facts) {
for(unsigned i=0; i<fact_cnt; i++) {
remove_fact(facts[i]);
}
}
void table_base::remove_facts(unsigned fact_cnt, const table_element * facts) {
for(unsigned i=0; i<fact_cnt; i++) {
remove_fact(facts + i*get_signature().size());
}
}
void table_base::reset() {
vector<table_fact> to_remove;
table_base::iterator it = begin();
table_base::iterator iend = end();
table_fact row;
for(; it!=iend; ++it) {
it->get_fact(row);
to_remove.push_back(row);
}
remove_facts(to_remove.size(), to_remove.c_ptr());
}
bool table_base::contains_fact(const table_fact & f) const {
iterator it = begin();
iterator iend = end();
table_fact row;
for(; it!=iend; ++it) {
it->get_fact(row);
if(vectors_equal(row, f)) {
return true;
}
}
return false;
}
bool table_base::fetch_fact(table_fact & f) const {
if(get_signature().functional_columns()==0) {
return contains_fact(f);
}
else {
unsigned sig_sz = get_signature().size();
unsigned non_func_cnt = sig_sz-get_signature().functional_columns();
table_base::iterator it = begin();
table_base::iterator iend = end();
table_fact row;
for(; it!=iend; ++it) {
it->get_fact(row);
bool differs = false;
for(unsigned i=0; i<non_func_cnt; i++) {
if(row[i]!=f[i]) {
differs = true;
}
}
if(differs) {
continue;
}
for(unsigned i=non_func_cnt; i<sig_sz; i++) {
f[i]=row[i];
}
return true;
}
return false;
}
}
bool table_base::suggest_fact(table_fact & f) {
if(get_signature().functional_columns()==0) {
if(contains_fact(f)) {
return false;
}
add_new_fact(f);
return true;
}
else {
if(fetch_fact(f)) {
return false;
}
add_new_fact(f);
return true;
}
}
void table_base::ensure_fact(const table_fact & f) {
if(get_signature().functional_columns()==0) {
add_fact(f);
}
else {
remove_fact(f);
add_fact(f);
}
}
table_base * table_base::clone() const {
table_base * res = get_plugin().mk_empty(get_signature());
iterator it = begin();
iterator iend = end();
table_fact row;
for(; it!=iend; ++it) {
it->get_fact(row);
res->add_new_fact(row);
}
return res;
}
/**
\brief Default method for complementation.
It assumes that the compiler creates only tables with
at most one column (0 or 1 columns).
Complementation of tables with more than one columns
is transformed into a cross product of complements and/or
difference.
*/
table_base * table_base::complement(func_decl* p, const table_element * func_columns) const {
const table_signature & sig = get_signature();
SASSERT(sig.functional_columns()==0 || func_columns!=0);
SASSERT(sig.first_functional() <= 1);
table_base * res = get_plugin().mk_empty(sig);
table_fact fact;
fact.resize(sig.first_functional());
fact.append(sig.functional_columns(), func_columns);
if (sig.first_functional() == 0) {
if (empty()) {
res->add_fact(fact);
}
return res;
}
VERIFY(sig.first_functional() == 1);
uint64 upper_bound = get_signature()[0];
bool empty_table = empty();
if (upper_bound > (1 << 18)) {
std::ostringstream buffer;
buffer << "creating large table of size " << upper_bound;
if (p) buffer << " for relation " << p->get_name();
warning_msg(buffer.str().c_str());
}
for(table_element i = 0; i < upper_bound; i++) {
fact[0] = i;
if(empty_table || !contains_fact(fact)) {
res->add_fact(fact);
}
}
return res;
}
void table_base::display(std::ostream & out) const {
out << "table with signature ";
print_container(get_signature(), out);
out << ":\n";
iterator it = begin();
iterator iend = end();
for(; it!=iend; ++it) {
const row_interface & r = *it;
r.display(out);
}
out << "\n";
}
class table_base::row_interface::fact_row_iterator : public table_base::row_iterator_core {
const row_interface & m_parent;
unsigned m_index;
protected:
virtual bool is_finished() const { return m_index==m_parent.size(); }
public:
fact_row_iterator(const row_interface & row, bool finished)
: m_parent(row), m_index(finished ? row.size() : 0) {}
virtual table_element operator*() {
SASSERT(!is_finished());
return m_parent[m_index];
}
virtual void operator++() {
m_index++;
SASSERT(m_index<=m_parent.size());
}
};
table_base::row_iterator table_base::row_interface::begin() const {
return row_iterator(alloc(fact_row_iterator, *this, false));
}
table_base::row_iterator table_base::row_interface::end() const {
return row_iterator(alloc(fact_row_iterator, *this, true));
}
void table_base::row_interface::get_fact(table_fact & result) const {
result.reset();
unsigned n=size();
for(unsigned i=0; i<n; i++) {
result.push_back((*this)[i]);
}
}
void table_base::row_interface::display(std::ostream & out) const {
table_fact fact;
get_fact(fact);
print_container(fact, out);
out << "\n";
}
void table_base::to_formula(relation_signature const& sig, expr_ref& fml) const {
// iterate over rows and build disjunction
ast_manager & m = fml.get_manager();
expr_ref_vector disjs(m);
expr_ref_vector conjs(m);
dl_decl_util util(m);
bool_rewriter brw(m);
table_fact fact;
iterator it = begin();
iterator iend = end();
for(; it != iend; ++it) {
const row_interface & r = *it;
r.get_fact(fact);
conjs.reset();
for (unsigned i = 0; i < fact.size(); ++i) {
conjs.push_back(m.mk_eq(m.mk_var(i, sig[i]), util.mk_numeral(fact[i], sig[i])));
}
brw.mk_and(conjs.size(), conjs.c_ptr(), fml);
disjs.push_back(fml);
}
brw.mk_or(disjs.size(), disjs.c_ptr(), fml);
}
}

1261
src/muz/dl_base.h Normal file

File diff suppressed because it is too large Load diff

1541
src/muz/dl_bmc_engine.cpp Normal file

File diff suppressed because it is too large Load diff

82
src/muz/dl_bmc_engine.h Normal file
View file

@ -0,0 +1,82 @@
/*++
Copyright (c) 2012 Microsoft Corporation
Module Name:
dl_bmc_engine.h
Abstract:
BMC engine for fixedpoint solver.
Author:
Nikolaj Bjorner (nbjorner) 2012-9-20
Revision History:
--*/
#ifndef _DL_BMC_ENGINE_H_
#define _DL_BMC_ENGINE_H_
#include "params.h"
#include "statistics.h"
#include "smt_kernel.h"
#include "bv_decl_plugin.h"
#include "smt_params.h"
namespace datalog {
class context;
class bmc : public engine_base {
context& m_ctx;
ast_manager& m;
smt_params m_fparams;
smt::kernel m_solver;
rule_set m_rules;
func_decl_ref m_query_pred;
expr_ref m_answer;
volatile bool m_cancel;
void checkpoint();
class nonlinear_dt;
class nonlinear;
class qlinear;
class linear;
bool is_linear() const;
void assert_expr(expr* e);
public:
bmc(context& ctx);
~bmc();
lbool query(expr* query);
void cancel();
void cleanup();
void display_certificate(std::ostream& out) const;
void collect_statistics(statistics& st) const;
void reset_statistics();
expr_ref get_answer();
// direct access to (new) non-linear compiler.
void compile(rule_set const& rules, expr_ref_vector& fmls, unsigned level);
expr_ref compile_query(func_decl* query_pred, unsigned level);
};
};
#endif

297
src/muz/dl_boogie_proof.cpp Normal file
View file

@ -0,0 +1,297 @@
/**
Example from Boogie:
(derivation
(step s!4 (main_loop_LoopHead true true)
rule!3 (subst
(= assertsPassed@@1 true)
(= assertsPassed@2@@0 true)
(= main_loop_LoopHead_assertsPassed true)
)
(labels @950 +895 +668 +670 +893 +899 )
(ref true ))
(step s!3 (main true false)
rule!1 (subst
(= assertsPassed true)
(= assertsPassed@0 true)
(= assertsPassed@2 false)
(= main_assertsPassed false)
)
(labels @839 +763 +547 +546 +761 +544 +545 +si_fcall_805 +681 +768 )
(ref s!4 ))
(step s!2 (main_SeqInstr true false)
rule!2 (subst
(= assertsPassed@@0 true)
(= assertsPassed@0@@0 false)
(= main_SeqInstr_assertsPassed false)
)
(labels @890 +851 +572 +si_fcall_883 +853 )
(ref s!3 ))
(step s!1 (@Fail!0)
rule!4 (subst
(= assertsPassed@@2 true)
(= main_SeqInstr_assertsPassed@@0 false)
)
(labels )
(ref s!2 ))
)
(model
"tickleBool -> {
true -> true
false -> true
else -> true
}
")
*/
#include "dl_boogie_proof.h"
#include "model_pp.h"
#include "proof_utils.h"
#include "ast_pp.h"
#include "qe_util.h"
namespace datalog {
/**
\brief push hyper-resolution steps upwards such that every use of
hyper-resolution uses a premise that is not derived from hyper-resolution.
perform the following rewrite:
hr(hr(p1,p2,p3,..),p4,p5) => hr(p1,hr(p2,p4),p3,..,p5)
*/
void mk_input_resolution(proof_ref& pr) {
ast_manager& m = pr.get_manager();
proof_ref pr1(m);
proof_ref_vector premises1(m), premises2(m), premises(m);
expr_ref conclusion1(m), conclusion2(m), conclusion(m);
svector<std::pair<unsigned, unsigned> > positions1, positions2, positions;
vector<expr_ref_vector> substs1, substs2, substs;
if (m.is_hyper_resolve(pr, premises1, conclusion1, positions1, substs1) &&
m.is_hyper_resolve(premises1[0].get(), premises, conclusion2, positions, substs2)) {
for (unsigned i = 1; i < premises1.size(); ++i) {
pr1 = premises1[i].get();
mk_input_resolution(pr1);
premises1[i] = pr1;
}
for (unsigned i = 0; i < premises.size(); ++i) {
pr1 = premises[i].get();
mk_input_resolution(pr1);
premises[i] = pr1;
}
unsigned sz = premises.size();
for (unsigned i = 1; i < sz; ++i) {
proof* premise = premises[i].get();
expr_ref_vector literals(m);
expr* l1, *l2;
if (!m.is_implies(premise, l1, l2)) {
continue;
}
qe::flatten_and(l1, literals);
positions2.reset();
premises2.reset();
premises2.push_back(premise);
substs2.reset();
for (unsigned j = 0; j < literals.size(); ++j) {
expr* lit = literals[j].get();
for (unsigned k = 1; k < premises1.size(); ++k) {
if (m.get_fact(premises1[k].get()) == lit) {
premises2.push_back(premises1[k].get());
positions2.push_back(std::make_pair(j+1,0));
substs2.push_back(expr_ref_vector(m));
break;
}
}
}
premises[i] = m.mk_hyper_resolve(premises2.size(), premises2.c_ptr(), l2, positions2, substs2);
}
conclusion = conclusion1;
pr = m.mk_hyper_resolve(premises.size(), premises.c_ptr(), conclusion, positions, substs);
}
}
void boogie_proof::set_proof(proof* p) {
std::cout << "set proof\n";
m_proof = p;
proof_utils::push_instantiations_up(m_proof);
mk_input_resolution(m_proof);
std::cout << "proof set\n";
}
void boogie_proof::set_model(model* m) {
m_model = m;
}
void boogie_proof::pp(std::ostream& out) {
if (m_proof) {
pp_proof(out);
}
if (m_model) {
model_pp(out, *m_model);
}
}
void boogie_proof::pp_proof(std::ostream& out) {
vector<step> steps;
ptr_vector<proof> rules;
rules.push_back(m_proof);
steps.push_back(step());
obj_map<proof, unsigned> index;
index.insert(m_proof, 0);
for (unsigned j = 0; j < rules.size(); ++j) {
proof* p = rules[j];
proof_ref_vector premises(m);
expr_ref conclusion(m);
svector<std::pair<unsigned, unsigned> > positions;
vector<expr_ref_vector> substs;
expr* tmp;
steps[j].m_fact = m.get_fact(p);
m.is_implies(steps[j].m_fact, tmp, steps[j].m_fact);
get_subst(p, steps[j].m_subst);
get_labels(p, steps[j].m_labels);
if (m.is_hyper_resolve(p, premises, conclusion, positions, substs)) {
for (unsigned i = 1; i < premises.size(); ++i) {
proof* premise = premises[i].get();
unsigned position = 0;
if (!index.find(premise, position)) {
position = rules.size();
rules.push_back(premise);
steps.push_back(step());
index.insert(premise, position);
}
steps[j].m_refs.push_back(position);
}
get_rule_name(premises[0].get(), steps[j].m_rule_name);
}
}
for (unsigned j = steps.size(); j > 0; ) {
--j;
step &s = steps[j];
// TBD
// s.m_labels;
// set references, compensate for reverse ordering.
for (unsigned i = 0; i < s.m_refs.size(); ++i) {
s.m_refs[i] = rules.size()-1-s.m_refs[i];
}
}
steps.reverse();
pp_steps(out, steps);
}
/**
\brief extract the instantiation by searching for the first occurrence of a hyper-resolution
rule that produces an instance.
*/
void boogie_proof::get_subst(proof* p, subst& s) {
ptr_vector<proof> todo;
todo.push_back(p);
ast_mark visited;
std::cout << "get_subst\n" << mk_pp(p, m) << "\n";
while (!todo.empty()) {
proof* p = todo.back();
todo.pop_back();
if (visited.is_marked(p)) {
continue;
}
visited.mark(p, true);
proof_ref_vector premises(m);
expr_ref conclusion(m);
svector<std::pair<unsigned, unsigned> > positions;
vector<expr_ref_vector> substs;
if (m.is_hyper_resolve(p, premises, conclusion, positions, substs)) {
expr_ref_vector const& sub = substs[0];
if (!sub.empty()) {
quantifier* q = to_quantifier(m.get_fact(premises[0].get()));
unsigned sz = sub.size();
SASSERT(sz == q->get_num_decls());
for (unsigned i = 0; i < sz; ++i) {
s.push_back(std::make_pair(q->get_decl_name(sz-1-i), sub[i]));
}
return;
}
}
unsigned sz = m.get_num_parents(p);
for (unsigned i = 0; i < sz; ++i) {
todo.push_back(m.get_parent(p, i));
}
}
}
void boogie_proof::get_rule_name(proof* p, symbol& n) {
}
void boogie_proof::get_labels(proof* p, labels& lbls) {
}
void boogie_proof::pp_steps(std::ostream& out, vector<step>& steps) {
out << "(derivation\n";
for (unsigned i = 0; i < steps.size(); ++i) {
pp_step(out, i, steps[i]);
}
out << ")\n";
}
// step :: "(" "step" step-name fact rule-name subst labels premises ")"
void boogie_proof::pp_step(std::ostream& out, unsigned id, step& s) {
out << "(step\n";
out << " s!" << id << " ";
pp_fact(out, s.m_fact);
out << " " << s.m_rule_name << "\n";
pp_subst(out << " ", s.m_subst);
pp_labels(out << " ", s.m_labels);
pp_premises(out << " ", s.m_refs);
out << ")\n";
}
// fact :: "(" predicate theory-term* ")"
void boogie_proof::pp_fact(std::ostream& out, expr* fact) {
out << mk_pp(fact, m) << "\n";
}
// subst :: "(" "subst" assignment* ")"
void boogie_proof::pp_subst(std::ostream& out, subst& s) {
out << "(subst";
for (unsigned i = 0; i < s.size(); ++i) {
pp_assignment(out, s[i].first, s[i].second);
}
out << ")\n";
}
// assignment :: "(" "=" variable theory-term ")"
void boogie_proof::pp_assignment(std::ostream& out, symbol const& v, expr* t) {
out << "\n (= " << v << " " << mk_pp(t, m) << ")";
}
// labels :: "(" "labels" label* ")"
void boogie_proof::pp_labels(std::ostream& out, labels& lbls) {
out << "(labels";
for (unsigned i = 0; i < lbls.size(); ++i) {
out << " " << lbls[i];
}
out << ")\n";
}
// premises "(" "ref" step-name* ")"
void boogie_proof::pp_premises(std::ostream& out, refs& refs) {
out << "(ref";
for (unsigned i = 0; i < refs.size(); ++i) {
out << " s!" << refs[i];
}
out << ")\n";
}
}

87
src/muz/dl_boogie_proof.h Normal file
View file

@ -0,0 +1,87 @@
/**
output :: derivation model
derivation :: "(" "derivation" step* ")"
step :: "(" "step" step-name fact rule-name subst labels premises ")"
step-name :: identifier
rule-name :: identifier
fact :: "(" predicate theory-term* ")"
subst :: "(" "subst" assignment* ")"
assignment :: "(" "=" variable theory-term ")"
labels :: "(" "labels" label* ")"
premises :: "(" "ref" step-name* ")"
model :: "(" "model" smtlib2-model ")"
In each step the "fact" is derivable by hyper-resolution from the named
premises and the named rule, under the given substitution for the
universally quantified variables in the rule. The premises of each
step must have occurred previously in the step sequence. The last fact
is a nullary placeholder predicate representing satisfaction of the query
(its name is arbitrary).
The labels list consists of all the positively labeled sub-formulas whose
truth is used in the proof, and all the negatively labeled formulas whose
negation is used. A theory-term is a ground term using only interpreted
constants of the background theories.
The smtlib2-model gives an interpretation of the uninterpreted constants
in the background theory under which the derivation is valid. Currently
it is a quoted string in the old z3 model format, for compatibility with
Boogie, however, this should be changed to the new model format (using
define-fun) when Boogie supports this.
*/
#include "ast.h"
#include "model.h"
namespace datalog {
class boogie_proof {
typedef vector<std::pair<symbol,expr*> > subst;
typedef svector<symbol> labels;
typedef unsigned_vector refs;
struct step {
symbol m_rule_name;
expr* m_fact;
subst m_subst;
labels m_labels;
refs m_refs;
};
ast_manager& m;
proof_ref m_proof;
model_ref m_model;
void pp_proof(std::ostream& out);
void pp_steps(std::ostream& out, vector<step>& steps);
void pp_step(std::ostream& out, unsigned i, step& s);
void pp_fact(std::ostream& out, expr* fact);
void pp_subst(std::ostream& out, subst& s);
void pp_assignment(std::ostream& out, symbol const& v, expr* t);
void pp_labels(std::ostream& out, labels& lbls);
void pp_premises(std::ostream& out, refs& refs);
void get_subst(proof* p, subst& sub);
void get_rule_name(proof* p, symbol&);
void get_labels(proof* p, labels&);
public:
boogie_proof(ast_manager& m): m(m), m_proof(m), m_model(0) {}
void set_proof(proof* p);
void set_model(model* m);
void pp(std::ostream& out);
};
}

View file

@ -0,0 +1,707 @@
/*++
Copyright (c) 2010 Microsoft Corporation
Module Name:
dl_bound_relation.cpp
Abstract:
Basic (strict upper) bound relation.
Author:
Nikolaj Bjorner (nbjorner) 2010-2-11
Revision History:
--*/
#include "dl_bound_relation.h"
#include "debug.h"
#include "ast_pp.h"
namespace datalog {
bound_relation_plugin::bound_relation_plugin(relation_manager& m):
relation_plugin(bound_relation_plugin::get_name(), m),
m_arith(get_ast_manager()),
m_bsimp(get_ast_manager()) {
}
bool bound_relation_plugin::can_handle_signature(const relation_signature & sig) {
for (unsigned i = 0; i < sig.size(); ++i) {
if (!m_arith.is_int(sig[i]) && !m_arith.is_real(sig[i])) {
return false;
}
}
return true;
}
bound_relation& bound_relation_plugin::get(relation_base& r) {
return dynamic_cast<bound_relation&>(r);
}
bound_relation const & bound_relation_plugin::get(relation_base const& r) {
return dynamic_cast<bound_relation const&>(r);
}
bound_relation* bound_relation_plugin::get(relation_base* r) {
return dynamic_cast<bound_relation*>(r);
}
bool bound_relation_plugin::is_interval_relation(relation_base const& r) {
return symbol("interval_relation") == r.get_plugin().get_name();
}
interval_relation& bound_relation_plugin::get_interval_relation(relation_base& r) {
SASSERT(is_interval_relation(r));
return dynamic_cast<interval_relation&>(r);
}
interval_relation const& bound_relation_plugin::get_interval_relation(relation_base const& r) {
SASSERT(is_interval_relation(r));
return dynamic_cast<interval_relation const&>(r);
}
relation_base * bound_relation_plugin::mk_empty(const relation_signature & s) {
return alloc(bound_relation, *this, s, true);
}
relation_base * bound_relation_plugin::mk_full(func_decl* p, const relation_signature & s) {
return alloc(bound_relation, *this, s, false);
}
class bound_relation_plugin::join_fn : public convenient_relation_join_fn {
public:
join_fn(const relation_signature & o1_sig, const relation_signature & o2_sig, unsigned col_cnt,
const unsigned * cols1, const unsigned * cols2)
: convenient_relation_join_fn(o1_sig, o2_sig, col_cnt, cols1, cols2) {
}
virtual relation_base * operator()(const relation_base & _r1, const relation_base & _r2) {
bound_relation const& r1 = get(_r1);
bound_relation const& r2 = get(_r2);
bound_relation_plugin& p = r1.get_plugin();
bound_relation* result = dynamic_cast<bound_relation*>(p.mk_full(0, get_result_signature()));
result->mk_join(r1, r2, m_cols1.size(), m_cols1.c_ptr(), m_cols2.c_ptr());
return result;
}
};
relation_join_fn * bound_relation_plugin::mk_join_fn(const relation_base & r1, const relation_base & r2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2) {
if (!check_kind(r1) || !check_kind(r2)) {
return 0;
}
return alloc(join_fn, r1.get_signature(), r2.get_signature(), col_cnt, cols1, cols2);
}
class bound_relation_plugin::project_fn : public convenient_relation_project_fn {
public:
project_fn(const relation_signature & orig_sig, unsigned removed_col_cnt, const unsigned * removed_cols)
: convenient_relation_project_fn(orig_sig, removed_col_cnt, removed_cols) {
}
virtual relation_base * operator()(const relation_base & _r) {
bound_relation const& r = get(_r);
bound_relation_plugin& p = r.get_plugin();
bound_relation* result = get(p.mk_full(0, get_result_signature()));
result->mk_project(r, m_removed_cols.size(), m_removed_cols.c_ptr());
return result;
}
};
relation_transformer_fn * bound_relation_plugin::mk_project_fn(const relation_base & r,
unsigned col_cnt, const unsigned * removed_cols) {
return alloc(project_fn, r.get_signature(), col_cnt, removed_cols);
}
class bound_relation_plugin::rename_fn : public convenient_relation_rename_fn {
public:
rename_fn(const relation_signature & orig_sig, unsigned cycle_len, const unsigned * cycle)
: convenient_relation_rename_fn(orig_sig, cycle_len, cycle) {
}
virtual relation_base * operator()(const relation_base & _r) {
bound_relation const& r = get(_r);
bound_relation_plugin& p = r.get_plugin();
bound_relation* result = get(p.mk_full(0, get_result_signature()));
result->mk_rename(r, m_cycle.size(), m_cycle.c_ptr());
return result;
}
};
relation_transformer_fn * bound_relation_plugin::mk_rename_fn(const relation_base & r,
unsigned cycle_len, const unsigned * permutation_cycle) {
if(check_kind(r)) {
return alloc(rename_fn, r.get_signature(), cycle_len, permutation_cycle);
}
return 0;
}
class bound_relation_plugin::union_fn : public relation_union_fn {
bool m_is_widen;
public:
union_fn(bool is_widen) :
m_is_widen(is_widen) {
}
virtual void operator()(relation_base & _r, const relation_base & _src, relation_base * _delta) {
TRACE("bound_relation", _r.display(tout << "dst:\n"); _src.display(tout << "src:\n"););
get(_r).mk_union(get(_src), get(_delta), m_is_widen);
}
};
class bound_relation_plugin::union_fn_i : public relation_union_fn {
bool m_is_widen;
public:
union_fn_i(bool is_widen) :
m_is_widen(is_widen) {
}
virtual void operator()(relation_base & _r, const relation_base & _src, relation_base * _delta) {
TRACE("bound_relation", _r.display(tout << "dst:\n"); _src.display(tout << "src:\n"););
get(_r).mk_union_i(get_interval_relation(_src), get(_delta), m_is_widen);
TRACE("bound_relation", _r.display(tout << "dst':\n"););
}
};
relation_union_fn * bound_relation_plugin::mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta) {
if (check_kind(tgt) && is_interval_relation(src) && (!delta || check_kind(*delta))) {
return alloc(union_fn_i, false);
}
if (check_kind(tgt) && check_kind(src) && (!delta || check_kind(*delta))) {
return alloc(union_fn, false);
}
return 0;
}
relation_union_fn * bound_relation_plugin::mk_widen_fn(
const relation_base & tgt, const relation_base & src,
const relation_base * delta) {
if (check_kind(tgt) && is_interval_relation(src) && (!delta || check_kind(*delta))) {
return alloc(union_fn_i, true);
}
if (check_kind(tgt) && check_kind(src) && (!delta || check_kind(*delta))) {
return alloc(union_fn, true);
}
return 0;
}
class bound_relation_plugin::filter_identical_fn : public relation_mutator_fn {
unsigned_vector m_cols;
public:
filter_identical_fn(unsigned col_cnt, const unsigned * identical_cols)
: m_cols(col_cnt, identical_cols) {}
virtual void operator()(relation_base & r) {
for (unsigned i = 1; i < m_cols.size(); ++i) {
get(r).equate(m_cols[0], m_cols[i]);
}
}
};
relation_mutator_fn * bound_relation_plugin::mk_filter_identical_fn(
const relation_base & t, unsigned col_cnt, const unsigned * identical_cols) {
if(check_kind(t)) {
return alloc(filter_identical_fn, col_cnt, identical_cols);
}
return 0;
}
class bound_relation_plugin::filter_equal_fn : public relation_mutator_fn {
public:
filter_equal_fn(relation_element const& value, unsigned col) {}
virtual void operator()(relation_base & r) { }
};
relation_mutator_fn * bound_relation_plugin::mk_filter_equal_fn(const relation_base & r,
const relation_element & value, unsigned col) {
if (check_kind(r)) {
return alloc(filter_equal_fn, value, col);
}
return 0;
}
class bound_relation_plugin::filter_interpreted_fn : public relation_mutator_fn {
enum kind_t { NOT_APPLICABLE, EQ_VAR, EQ_SUB, LT_VAR, LE_VAR, K_FALSE };
app_ref m_cond;
app_ref m_lt;
arith_util m_arith;
interval_relation* m_interval;
unsigned_vector m_vars;
kind_t m_kind;
unsigned get_var(expr* a) {
SASSERT(is_var(a));
return to_var(a)->get_idx();
}
// x = z - y
void mk_sub_eq(expr* x, expr* z, expr* y) {
SASSERT(is_var(x));
SASSERT(is_var(z));
SASSERT(is_var(y));
m_vars.push_back(get_var(x));
m_vars.push_back(get_var(z));
m_vars.push_back(get_var(y));
m_kind = EQ_SUB;
}
void mk_lt(expr* l, expr* r) {
SASSERT(is_var(l));
SASSERT(is_var(r));
m_vars.push_back(get_var(l));
m_vars.push_back(get_var(r));
m_lt = m_arith.mk_lt(l, r);
m_kind = LT_VAR;
}
void mk_le(expr* l, expr* r) {
SASSERT(is_var(l));
SASSERT(is_var(r));
m_vars.push_back(get_var(l));
m_vars.push_back(get_var(r));
m_kind = LE_VAR;
}
void mk_eq(expr* l, expr* r) {
m_vars.push_back(get_var(l));
m_vars.push_back(get_var(r));
m_kind = EQ_VAR;
}
public:
filter_interpreted_fn(ast_manager& m, app* cond) :
m_cond(cond, m),
m_lt(m), m_arith(m), m_interval(0), m_kind(NOT_APPLICABLE) {
expr* l, *r, *r1, *r2, *c2;
rational n1;
if ((m_arith.is_lt(cond, l, r) || m_arith.is_gt(cond, r, l)) &&
is_var(l) && is_var(r)) {
mk_lt(l, r);
}
else if (m.is_not(cond, c2) &&
(m_arith.is_ge(c2, l, r) || m_arith.is_le(c2, r, l)) &&
is_var(l) && is_var(r)) {
mk_lt(l, r);
}
else if ((m_arith.is_le(cond, l, r) || m_arith.is_ge(cond, r, l)) &&
is_var(l) && is_var(r)) {
mk_le(l, r);
}
else if (m.is_not(cond, c2) &&
(m_arith.is_gt(c2, l, r) || m_arith.is_lt(c2, r, l)) &&
is_var(l) && is_var(r)) {
mk_le(l, r);
}
else if (m.is_false(cond)) {
m_kind = K_FALSE;
}
else if (m.is_eq(cond, l, r) && is_var(l) && is_var(r)) {
mk_eq(l, r);
}
else if (m.is_eq(cond, l, r) &&
m_arith.is_sub(r, r1, r2) &&
is_var(l) && is_var(r1) && is_var(r2)) {
mk_sub_eq(l, r1, r2);
}
else if (m.is_eq(cond, l, r) &&
m_arith.is_sub(l, r1, r2) &&
is_var(r) && is_var(r1) && is_var(r2)) {
mk_sub_eq(r, r1, r2);
}
else if (m.is_eq(cond, l, r) &&
m_arith.is_add(r, r1, r2) &&
m_arith.is_numeral(r1, n1) &&
n1.is_pos() && is_var(l) && is_var(r2)) {
mk_lt(r2, l);
}
else if (m.is_eq(cond, l, r) &&
m_arith.is_add(r, r1, r2) &&
m_arith.is_numeral(r2, n1) &&
n1.is_pos() && is_var(l) && is_var(r1)) {
mk_lt(r1, l);
}
else {
}
}
//
// x = z - y
// x = y
// x < y
// x <= y
// x < y + z
//
void operator()(relation_base& t) {
TRACE("dl", tout << mk_pp(m_cond, m_cond.get_manager()) << "\n"; t.display(tout););
bound_relation& r = get(t);
switch(m_kind) {
case K_FALSE:
r.set_empty();
break;
case NOT_APPLICABLE:
break;
case EQ_VAR:
r.equate(m_vars[0], m_vars[1]);
break;
case EQ_SUB:
// TBD
break;
case LT_VAR:
r.mk_lt(m_vars[0], m_vars[1]);
break;
case LE_VAR:
r.mk_le(m_vars[0], m_vars[1]);
break;
default:
UNREACHABLE();
break;
}
TRACE("dl", t.display(tout << "result\n"););
}
bool supports_attachment(relation_base& t) {
return is_interval_relation(t);
}
void attach(relation_base& t) {
SASSERT(is_interval_relation(t));
interval_relation& r = get_interval_relation(t);
m_interval = &r;
}
};
relation_mutator_fn * bound_relation_plugin::mk_filter_interpreted_fn(const relation_base & t, app * condition) {
return alloc(filter_interpreted_fn, t.get_plugin().get_ast_manager(), condition);
}
// -----------------------------
// bound_relation
void bound_relation_helper::mk_project_t(uint_set2& t, unsigned_vector const& renaming) {
if (t.lt.empty() && t.le.empty()) {
return;
}
uint_set::iterator it = t.lt.begin(), end = t.lt.end();
unsigned_vector ltv, lev;
for (; it != end; ++it) {
ltv.push_back(renaming[*it]);
}
it = t.le.begin(), end = t.le.end();
for (; it != end; ++it) {
lev.push_back(renaming[*it]);
}
TRACE("dl",
tout << "project: ";
for (unsigned i = 0; i < renaming.size(); ++i)
if (renaming[i] == UINT_MAX) tout << i << " ";
tout << ": ";
it = t.lt.begin(); end = t.lt.end();
for (; it != end; ++it) tout << *it << " ";
tout << " le ";
it = t.le.begin(); end = t.le.end();
for (; it != end; ++it) tout << *it << " ";
tout << " => ";
for (unsigned i = 0; i < ltv.size(); ++i) tout << ltv[i] << " ";
tout << " le ";
for (unsigned i = 0; i < lev.size(); ++i) tout << lev[i] << " ";
tout << "\n";);
t.lt.reset();
for (unsigned i = 0; i < ltv.size(); ++i) {
t.lt.insert(ltv[i]);
}
t.le.reset();
for (unsigned i = 0; i < lev.size(); ++i) {
t.le.insert(lev[i]);
}
}
bound_relation::bound_relation(bound_relation_plugin& p, relation_signature const& s, bool is_empty):
vector_relation<uint_set2, bound_relation_helper>(p, s, is_empty, uint_set2())
{
}
uint_set2 bound_relation::mk_intersect(uint_set2 const& t1, uint_set2 const& t2, bool& is_empty) const {
is_empty = false;
uint_set2 r(t1);
r.lt |= t2.lt;
r.le |= t2.le;
return r;
}
uint_set2 bound_relation::mk_widen(uint_set2 const& t1, uint_set2 const& t2) const {
return mk_unite(t1, t2);
}
uint_set2 bound_relation::mk_unite(uint_set2 const& t1, uint_set2 const& t2) const {
uint_set2 s1(t1);
s1.lt &= t2.lt;
s1.le &= t2.le;
return s1;
}
uint_set2 bound_relation::mk_eq(union_find<> const& old_eqs, union_find<> const& new_eqs, uint_set2 const& t) const {
unsigned sz = old_eqs.get_num_vars();
SASSERT(sz == new_eqs.get_num_vars());
uint_set2 result;
for (unsigned i = 0; i < sz; ++i) {
if (t.lt.contains(i)) {
unsigned j = i;
do {
result.lt.insert(new_eqs.find(j));
j = old_eqs.next(j);
}
while (j != i);
}
if (t.le.contains(i)) {
unsigned j = i;
do {
result.le.insert(new_eqs.find(j));
j = old_eqs.next(j);
}
while (j != i);
}
}
return result;
}
bool bound_relation::is_subset_of(uint_set2 const& t1, uint_set2 const& t2) const {
uint_set2 s1, s2;
normalize(t1, s1);
normalize(t2, s2);
return s1.lt.subset_of(s2.lt) && s1.le.subset_of(s2.le);
}
void bound_relation::mk_rename_elem(uint_set2& t, unsigned col_cnt, unsigned const* cycle) {
// [ 0 -> 2 -> 3 -> 0]
if (col_cnt == 0) return;
unsigned col1, col2;
col1 = find(cycle[0]);
col2 = find(cycle[col_cnt-1]);
bool has_col2_lt = t.lt.contains(col2);
t.lt.remove(col2);
bool has_col2_le = t.le.contains(col2);
t.le.remove(col2);
for (unsigned i = 0; i + 1 < col_cnt; ++i) {
col1 = find(cycle[i]);
col2 = find(cycle[i+1]);
if (t.lt.contains(col1)) {
t.lt.remove(col1);
t.lt.insert(col2);
}
if (t.le.contains(col1)) {
t.le.remove(col1);
t.le.insert(col2);
}
}
if (has_col2_lt) {
col1 = find(cycle[0]);
t.lt.insert(col1);
}
if (has_col2_le) {
col1 = find(cycle[0]);
t.le.insert(col1);
}
}
bool bound_relation::is_full(uint_set2 const& t) const {
return t.lt.empty() && t.le.empty();
}
bool bound_relation::is_empty(unsigned index, uint_set2 const& t) const {
return t.lt.contains(find(index)) || t.le.contains(find(index));
}
void bound_relation::normalize(uint_set const& src, uint_set& dst) const {
uint_set::iterator it = src.begin(), end = src.end();
for (; it != end; ++it) {
dst.insert(find(*it));
}
}
void bound_relation::normalize(uint_set2 const& src, uint_set2& dst) const {
normalize(src.lt, dst.lt);
normalize(src.le, dst.le);
}
void bound_relation::mk_lt(unsigned i) {
uint_set2& dst = (*this)[i];
while (!m_todo.empty()) {
unsigned j = m_todo.back().first;
bool strict = m_todo.back().second;
if (i == j && strict) {
m_todo.reset();
m_empty = true;
return;
}
m_todo.pop_back();
if (i == j) {
continue;
}
uint_set2& src = (*m_elems)[j];
uint_set::iterator it = src.lt.begin(), end = src.lt.end();
for(; it != end; ++it) {
m_todo.push_back(std::make_pair(*it, true));
}
it = src.le.begin(), end = src.le.end();
for(; it != end; ++it) {
m_todo.push_back(std::make_pair(*it, strict));
}
if (strict) {
dst.lt.insert(j);
}
else {
dst.le.insert(j);
}
}
}
void bound_relation::mk_lt(unsigned i, unsigned j) {
m_todo.reset();
i = find(i);
m_todo.push_back(std::make_pair(find(j), true));
mk_lt(i);
}
void bound_relation::mk_le(unsigned i, unsigned j) {
m_todo.reset();
i = find(i);
m_todo.push_back(std::make_pair(find(j), false));
mk_lt(i);
}
bool bound_relation::is_lt(unsigned i, unsigned j) const {
return (*this)[i].lt.contains(find(j));
}
void bound_relation::add_fact(const relation_fact & f) {
bound_relation r(get_plugin(), get_signature(), false);
for (unsigned i = 0; i < f.size(); ++i) {
scoped_ptr<relation_mutator_fn> fe = get_plugin().mk_filter_equal_fn(r, f[i], i);
(*fe)(r);
}
mk_union(r, 0, false);
}
bool bound_relation::contains_fact(const relation_fact & f) const {
if (empty()) {
return false;
}
// this is a very rough approximation.
return true;
}
bound_relation * bound_relation::clone() const {
bound_relation* result = 0;
if (empty()) {
result = bound_relation_plugin::get(get_plugin().mk_empty(get_signature()));
}
else {
result = bound_relation_plugin::get(get_plugin().mk_full(0, get_signature()));
result->copy(*this);
}
return result;
}
void bound_relation::mk_union_i(interval_relation const& src, bound_relation* delta, bool is_widen) {
unsigned size = get_signature().size();
for (unsigned i = 0; i < size; ++i) {
if (find(i) != i) {
continue;
}
uint_set2& s = (*this)[i];
ext_numeral const& lo = src[i].sup();
if (lo.is_infinite()) {
s.lt.reset();
s.le.reset();
continue;
}
uint_set::iterator it = s.lt.begin(), end = s.lt.end();
for(; it != end; ++it) {
ext_numeral const& hi = src[*it].inf();
if (hi.is_infinite() || lo.to_rational() >= hi.to_rational()) {
s.lt.remove(*it);
}
}
it = s.le.begin(), end = s.le.end();
for(; it != end; ++it) {
ext_numeral const& hi = src[*it].inf();
if (hi.is_infinite() || lo.to_rational() > hi.to_rational()) {
s.le.remove(*it);
}
}
}
}
bound_relation * bound_relation::complement(func_decl* p) const {
UNREACHABLE();
return 0;
}
void bound_relation::to_formula(expr_ref& fml) const {
ast_manager& m = get_plugin().get_ast_manager();
arith_util& arith = get_plugin().m_arith;
basic_simplifier_plugin& bsimp = get_plugin().m_bsimp;
expr_ref_vector conjs(m);
relation_signature const& sig = get_signature();
for (unsigned i = 0; i < sig.size(); ++i) {
if (i != find(i)) {
conjs.push_back(m.mk_eq(m.mk_var(i, sig[i]), m.mk_var(find(i), sig[find(i)])));
continue;
}
uint_set2 const& upper = (*this)[i];
uint_set::iterator it = upper.lt.begin(), end = upper.lt.end();
for (; it != end; ++it) {
conjs.push_back(arith.mk_lt(m.mk_var(i, sig[i]), m.mk_var(*it, sig[*it])));
}
it = upper.le.begin(), end = upper.le.end();
for (; it != end; ++it) {
conjs.push_back(arith.mk_le(m.mk_var(i, sig[i]), m.mk_var(*it, sig[*it])));
}
}
bsimp.mk_and(conjs.size(), conjs.c_ptr(), fml);
}
void bound_relation::display_index(unsigned i, uint_set2 const& src, std::ostream & out) const {
uint_set::iterator it = src.lt.begin(), end = src.lt.end();
out << "#" << i;
if (!src.lt.empty()) {
out << " < ";
for(; it != end; ++it) {
out << *it << " ";
}
}
if (!src.le.empty()) {
it = src.le.begin(), end = src.le.end();
out << " <= ";
for(; it != end; ++it) {
out << *it << " ";
}
}
if (src.lt.empty() && src.le.empty()) {
out << " < oo";
}
out << "\n";
}
bound_relation_plugin& bound_relation::get_plugin() const {
return dynamic_cast<bound_relation_plugin&>(relation_base::get_plugin());
}
};

176
src/muz/dl_bound_relation.h Normal file
View file

@ -0,0 +1,176 @@
/*++
Copyright (c) 2010 Microsoft Corporation
Module Name:
dl_bound_relation.h
Abstract:
Basic (strict upper) bound relation.
Author:
Nikolaj Bjorner (nbjorner) 2010-2-11
Revision History:
--*/
#ifndef _DL_BOUND_RELATION_H_
#define _DL_BOUND_RELATION_H_
#include "dl_context.h"
#include "uint_set.h"
#include "dl_vector_relation.h"
#include "dl_interval_relation.h"
#include "arith_decl_plugin.h"
#include "basic_simplifier_plugin.h"
namespace datalog {
class bound_relation;
class bound_relation_plugin : public relation_plugin {
friend class bound_relation;
class join_fn;
class project_fn;
class rename_fn;
class union_fn;
class union_fn_i;
class filter_equal_fn;
class filter_identical_fn;
class filter_interpreted_fn;
class filter_intersection_fn;
arith_util m_arith;
basic_simplifier_plugin m_bsimp;
public:
bound_relation_plugin(relation_manager& m);
virtual bool can_handle_signature(const relation_signature & s);
static symbol get_name() { return symbol("bound_relation"); }
virtual relation_base * mk_empty(const relation_signature & s);
virtual relation_base * mk_full(func_decl* p, const relation_signature & s);
virtual relation_join_fn * mk_join_fn(const relation_base & t1, const relation_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
virtual relation_transformer_fn * mk_project_fn(const relation_base & t, unsigned col_cnt,
const unsigned * removed_cols);
virtual relation_transformer_fn * mk_rename_fn(const relation_base & t, unsigned permutation_cycle_len,
const unsigned * permutation_cycle);
virtual relation_union_fn * mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
virtual relation_union_fn * mk_widen_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
virtual relation_mutator_fn * mk_filter_identical_fn(const relation_base & t, unsigned col_cnt,
const unsigned * identical_cols);
virtual relation_mutator_fn * mk_filter_equal_fn(const relation_base & t, const relation_element & value,
unsigned col);
virtual relation_mutator_fn * mk_filter_interpreted_fn(const relation_base & t, app * condition);
virtual relation_join_fn * mk_join_project_fn(const relation_base & t1, const relation_base & t2,
unsigned joined_col_cnt, const unsigned * cols1, const unsigned * cols2,
unsigned removed_col_cnt, const unsigned * removed_cols) { return 0; }
#if 0
virtual intersection_filter_fn * mk_filter_by_intersection_fn(
const relation_base & t,
const relation_base & src, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * src_cols) {
return 0;
}
#endif
static bound_relation* get(relation_base* r);
private:
static bound_relation& get(relation_base& r);
static bound_relation const & get(relation_base const& r);
static bool is_interval_relation(relation_base const& r);
static interval_relation& get_interval_relation(relation_base& r);
static interval_relation const& get_interval_relation(relation_base const& r);
};
struct uint_set2 {
uint_set lt;
uint_set le;
uint_set2(uint_set2 const& other):lt(other.lt), le(other.le) {}
uint_set2() {}
bool operator==(const uint_set2& other) const {
return other.lt == lt && other.le == le;
}
bool operator!=(const uint_set2& other) const {
return other.lt != lt || other.le != le;
}
};
inline std::ostream & operator<<(std::ostream & target, const uint_set2 & s) {
return target << s.lt << " " << s.le;
}
class bound_relation_helper {
public:
static void mk_project_t(uint_set2& t, unsigned_vector const& renaming);
};
class bound_relation : public vector_relation<uint_set2, bound_relation_helper> {
friend class bound_relation_plugin;
svector<std::pair<unsigned, bool> > m_todo;
public:
bound_relation(bound_relation_plugin& p, relation_signature const& s, bool is_empty);
bound_relation& operator=(bound_relation const& other);
virtual bool empty() const { return m_empty; }
virtual void add_fact(const relation_fact & f);
virtual bool contains_fact(const relation_fact & f) const;
virtual bound_relation * clone() const;
virtual bound_relation * complement(func_decl* p) const;
virtual void to_formula(expr_ref& fml) const;
bound_relation_plugin& get_plugin() const;
void mk_union_i(interval_relation const& src, bound_relation* delta, bool is_widen);
void mk_lt(unsigned i, unsigned j);
void mk_lt(unsigned i);
void mk_le(unsigned i, unsigned j);
bool is_lt(unsigned i, unsigned j) const;
virtual bool is_precise() const { return false; }
private:
typedef uint_set2 T;
virtual T mk_intersect(T const& t1, T const& t2, bool& is_empty) const;
virtual T mk_widen(T const& t1, T const& t2) const;
virtual T mk_unite(T const& t1, T const& t2) const;
virtual T mk_eq(union_find<> const& old_eqs, union_find<> const& new_eqs, T const& t) const;
virtual void mk_rename_elem(T& i, unsigned col_cnt, unsigned const* cycle);
virtual bool is_subset_of(T const& t1, T const& t2) const;
virtual bool is_full(T const& t) const;
virtual bool is_empty(unsigned idx, T const& t) const;
virtual void display_index(unsigned idx, T const& t, std::ostream& out) const;
void normalize(T const& src, T& dst) const;
void normalize(uint_set const& src, uint_set& dst) const;
};
};
#endif

439
src/muz/dl_check_table.cpp Normal file
View file

@ -0,0 +1,439 @@
/*++
Copyright (c) 2010 Microsoft Corporation
Module Name:
dl_check_table.cpp
Abstract:
<abstract>
Author:
Nikolaj Bjorner (nbjorner) 2010-11-15
Revision History:
--*/
#include "dl_check_table.h"
#include "dl_table.h"
namespace datalog {
bool check_table_plugin::can_handle_signature(table_signature const& s) {
return m_tocheck.can_handle_signature(s) && m_checker.can_handle_signature(s);
}
check_table & check_table_plugin::get(table_base& r) {
return static_cast<check_table&>(r);
}
check_table const & check_table_plugin::get(table_base const& r) {
return static_cast<check_table const &>(r);
}
table_base& check_table_plugin::checker(table_base& r) { return *get(r).m_checker; }
table_base const& check_table_plugin::checker(table_base const& r) { return *get(r).m_checker; }
table_base* check_table_plugin::checker(table_base* r) { return r?(get(*r).m_checker):0; }
table_base const* check_table_plugin::checker(table_base const* r) { return r?(get(*r).m_checker):0; }
table_base& check_table_plugin::tocheck(table_base& r) { return *get(r).m_tocheck; }
table_base const& check_table_plugin::tocheck(table_base const& r) { return *get(r).m_tocheck; }
table_base* check_table_plugin::tocheck(table_base* r) { return r?(get(*r).m_tocheck):0; }
table_base const* check_table_plugin::tocheck(table_base const* r) { return r?(get(*r).m_tocheck):0; }
table_base * check_table_plugin::mk_empty(const table_signature & s) {
IF_VERBOSE(1, verbose_stream() << __FUNCTION__ << "\n";);
table_base* checker = m_checker.mk_empty(s);
table_base* tocheck = m_tocheck.mk_empty(s);
return alloc(check_table, *this, s, tocheck, checker);
}
class check_table_plugin::join_fn : public table_join_fn {
scoped_ptr<table_join_fn> m_tocheck;
scoped_ptr<table_join_fn> m_checker;
public:
join_fn(check_table_plugin& p,
const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2) {
m_tocheck = p.get_manager().mk_join_fn(tocheck(t1), tocheck(t2), col_cnt, cols1, cols2);
m_checker = p.get_manager().mk_join_fn(checker(t1), checker(t2), col_cnt, cols1, cols2);
}
virtual table_base* operator()(const table_base & t1, const table_base & t2) {
IF_VERBOSE(1, verbose_stream() << __FUNCTION__ << "\n";);
table_base* ttocheck = (*m_tocheck)(tocheck(t1), tocheck(t2));
table_base* tchecker = (*m_checker)(checker(t1), checker(t2));
check_table* result = alloc(check_table, get(t1).get_plugin(), ttocheck->get_signature(), ttocheck, tchecker);
return result;
}
};
table_join_fn * check_table_plugin::mk_join_fn(const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2) {
if (!check_kind(t1) || !check_kind(t2)) {
return 0;
}
return alloc(join_fn, *this, t1, t2, col_cnt, cols1, cols2);
}
class check_table_plugin::join_project_fn : public table_join_fn {
scoped_ptr<table_join_fn> m_tocheck;
scoped_ptr<table_join_fn> m_checker;
public:
join_project_fn(check_table_plugin& p, const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2,
unsigned removed_col_cnt, const unsigned * removed_cols) {
m_tocheck = p.get_manager().mk_join_project_fn(tocheck(t1), tocheck(t2), col_cnt, cols1, cols2, removed_col_cnt, removed_cols);
m_checker = p.get_manager().mk_join_project_fn(checker(t1), checker(t2), col_cnt, cols1, cols2, removed_col_cnt, removed_cols);
}
virtual table_base* operator()(const table_base & t1, const table_base & t2) {
table_base* ttocheck = (*m_tocheck)(tocheck(t1), tocheck(t2));
table_base* tchecker = (*m_checker)(checker(t1), checker(t2));
check_table* result = alloc(check_table, get(t1).get_plugin(), ttocheck->get_signature(), ttocheck, tchecker);
return result;
}
};
table_join_fn * check_table_plugin::mk_join_project_fn(const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2, unsigned removed_col_cnt,
const unsigned * removed_cols) {
if (!check_kind(t1) || !check_kind(t2)) {
return 0;
}
return alloc(join_project_fn, *this, t1, t2, col_cnt, cols1, cols2, removed_col_cnt, removed_cols);
}
class check_table_plugin::union_fn : public table_union_fn {
scoped_ptr<table_union_fn> m_tocheck;
scoped_ptr<table_union_fn> m_checker;
public:
union_fn(check_table_plugin& p, table_base const& tgt, const table_base& src, table_base const* delta) {
m_tocheck = p.get_manager().mk_union_fn(tocheck(tgt), tocheck(src), tocheck(delta));
m_checker = p.get_manager().mk_union_fn(checker(tgt), checker(src), checker(delta));
}
virtual void operator()(table_base& tgt, const table_base& src, table_base* delta) {
IF_VERBOSE(1, verbose_stream() << __FUNCTION__ << "\n";);
(*m_tocheck)(tocheck(tgt), tocheck(src), tocheck(delta));
(*m_checker)(checker(tgt), checker(src), checker(delta));
get(tgt).well_formed();
if (delta) {
get(*delta).well_formed();
}
}
};
table_union_fn * check_table_plugin::mk_union_fn(const table_base & tgt, const table_base & src, const table_base * delta) {
if (!check_kind(tgt) || !check_kind(src) || (delta && !check_kind(*delta))) {
return 0;
}
return alloc(union_fn, *this, tgt, src, delta);
}
class check_table_plugin::project_fn : public table_transformer_fn {
scoped_ptr<table_transformer_fn> m_checker;
scoped_ptr<table_transformer_fn> m_tocheck;
public:
project_fn(check_table_plugin& p, const table_base & t, unsigned col_cnt, const unsigned * removed_cols) {
m_checker = p.get_manager().mk_project_fn(checker(t), col_cnt, removed_cols);
m_tocheck = p.get_manager().mk_project_fn(tocheck(t), col_cnt, removed_cols);
}
table_base* operator()(table_base const& src) {
table_base* tchecker = (*m_checker)(checker(src));
table_base* ttocheck = (*m_tocheck)(tocheck(src));
check_table* result = alloc(check_table, get(src).get_plugin(), tchecker->get_signature(), ttocheck, tchecker);
return result;
}
};
table_transformer_fn * check_table_plugin::mk_project_fn(const table_base & t, unsigned col_cnt, const unsigned * removed_cols) {
if (!check_kind(t)) {
return 0;
}
return alloc(project_fn, *this, t, col_cnt, removed_cols);
}
class check_table_plugin::select_equal_and_project_fn : public table_transformer_fn {
scoped_ptr<table_transformer_fn> m_checker;
scoped_ptr<table_transformer_fn> m_tocheck;
public:
select_equal_and_project_fn(check_table_plugin& p, const table_base & t, const table_element & value, unsigned col) {
m_checker = p.get_manager().mk_select_equal_and_project_fn(checker(t), value, col);
m_tocheck = p.get_manager().mk_select_equal_and_project_fn(tocheck(t), value, col);
}
table_base* operator()(table_base const& src) {
table_base* tchecker = (*m_checker)(checker(src));
table_base* ttocheck = (*m_tocheck)(tocheck(src));
check_table* result = alloc(check_table, get(src).get_plugin(), tchecker->get_signature(), ttocheck, tchecker);
return result;
}
};
table_transformer_fn * check_table_plugin::mk_select_equal_and_project_fn(const table_base & t,
const table_element & value, unsigned col) {
if (!check_kind(t)) {
return 0;
}
return alloc(select_equal_and_project_fn, *this, t, value, col);
}
class check_table_plugin::rename_fn : public table_transformer_fn {
scoped_ptr<table_transformer_fn> m_checker;
scoped_ptr<table_transformer_fn> m_tocheck;
public:
rename_fn(check_table_plugin& p, const table_base & t, unsigned cycle_len, unsigned const* cycle) {
m_checker = p.get_manager().mk_rename_fn(checker(t), cycle_len, cycle);
m_tocheck = p.get_manager().mk_rename_fn(tocheck(t), cycle_len, cycle);
}
table_base* operator()(table_base const& src) {
IF_VERBOSE(1, verbose_stream() << __FUNCTION__ << "\n";);
table_base* tchecker = (*m_checker)(checker(src));
table_base* ttocheck = (*m_tocheck)(tocheck(src));
check_table* result = alloc(check_table, get(src).get_plugin(), ttocheck->get_signature(), ttocheck, tchecker);
return result;
}
};
table_transformer_fn * check_table_plugin::mk_rename_fn(const table_base & t, unsigned len, const unsigned * cycle) {
if (!check_kind(t)) {
return 0;
}
return alloc(rename_fn, *this, t, len, cycle);
}
class check_table_plugin::filter_identical_fn : public table_mutator_fn {
scoped_ptr<table_mutator_fn> m_checker;
scoped_ptr<table_mutator_fn> m_tocheck;
public:
filter_identical_fn(check_table_plugin& p, const table_base & t,unsigned cnt, unsigned const* cols)
{
m_checker = p.get_manager().mk_filter_identical_fn(checker(t), cnt, cols);
m_tocheck = p.get_manager().mk_filter_identical_fn(tocheck(t), cnt, cols);
}
void operator()(table_base & t) {
(*m_checker)(checker(t));
(*m_tocheck)(tocheck(t));
get(t).well_formed();
}
};
table_mutator_fn * check_table_plugin::mk_filter_identical_fn(const table_base & t, unsigned col_cnt,
const unsigned * identical_cols) {
if (check_kind(t)) {
return alloc(filter_identical_fn, *this, t, col_cnt, identical_cols);
}
return 0;
}
class check_table_plugin::filter_equal_fn : public table_mutator_fn {
scoped_ptr<table_mutator_fn> m_checker;
scoped_ptr<table_mutator_fn> m_tocheck;
public:
filter_equal_fn(check_table_plugin& p, const table_base & t, const table_element & v, unsigned col)
{
m_checker = p.get_manager().mk_filter_equal_fn(checker(t), v, col);
m_tocheck = p.get_manager().mk_filter_equal_fn(tocheck(t), v, col);
}
virtual void operator()(table_base& src) {
(*m_checker)(checker(src));
(*m_tocheck)(tocheck(src));
get(src).well_formed();
}
};
table_mutator_fn * check_table_plugin::mk_filter_equal_fn(const table_base & t, const table_element & value, unsigned col) {
if (check_kind(t)) {
return alloc(filter_equal_fn, *this, t, value, col);
}
return 0;
}
class check_table_plugin::filter_interpreted_fn : public table_mutator_fn {
scoped_ptr<table_mutator_fn> m_checker;
scoped_ptr<table_mutator_fn> m_tocheck;
public:
filter_interpreted_fn(check_table_plugin& p, const table_base & t, app * condition)
{
m_checker = p.get_manager().mk_filter_interpreted_fn(checker(t), condition);
m_tocheck = p.get_manager().mk_filter_interpreted_fn(tocheck(t), condition);
}
virtual void operator()(table_base& src) {
(*m_checker)(checker(src));
(*m_tocheck)(tocheck(src));
get(src).well_formed();
}
};
table_mutator_fn * check_table_plugin::mk_filter_interpreted_fn(const table_base & t, app * condition) {
if (check_kind(t)) {
return alloc(filter_interpreted_fn, *this, t, condition);
}
return 0;
}
class check_table_plugin::filter_interpreted_and_project_fn : public table_transformer_fn {
scoped_ptr<table_transformer_fn> m_checker;
scoped_ptr<table_transformer_fn> m_tocheck;
public:
filter_interpreted_and_project_fn(check_table_plugin& p, const table_base & t, app * condition,
unsigned removed_col_cnt, const unsigned * removed_cols)
{
m_checker = p.get_manager().mk_filter_interpreted_and_project_fn(checker(t), condition, removed_col_cnt, removed_cols);
m_tocheck = p.get_manager().mk_filter_interpreted_and_project_fn(tocheck(t), condition, removed_col_cnt, removed_cols);
}
table_base* operator()(table_base const& src) {
table_base* tchecker = (*m_checker)(checker(src));
table_base* ttocheck = (*m_tocheck)(tocheck(src));
check_table* result = alloc(check_table, get(src).get_plugin(), ttocheck->get_signature(), ttocheck, tchecker);
return result;
}
};
table_transformer_fn * check_table_plugin::mk_filter_interpreted_and_project_fn(const table_base & t,
app * condition, unsigned removed_col_cnt, const unsigned * removed_cols) {
if (check_kind(t)) {
return alloc(filter_interpreted_and_project_fn, *this, t, condition, removed_col_cnt, removed_cols);
}
return 0;
}
class check_table_plugin::filter_by_negation_fn : public table_intersection_filter_fn {
scoped_ptr<table_intersection_filter_fn> m_checker;
scoped_ptr<table_intersection_filter_fn> m_tocheck;
public:
filter_by_negation_fn(
check_table_plugin& p,
const table_base & t,
const table_base & negated_obj, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * negated_cols) {
m_checker = p.get_manager().mk_filter_by_negation_fn(checker(t), checker(negated_obj), joined_col_cnt, t_cols, negated_cols);
m_tocheck = p.get_manager().mk_filter_by_negation_fn(tocheck(t), tocheck(negated_obj), joined_col_cnt, t_cols, negated_cols);
}
virtual void operator()(table_base& src, table_base const& negated_obj) {
IF_VERBOSE(1, verbose_stream() << __FUNCTION__ << "\n";);
(*m_checker)(checker(src), checker(negated_obj));
(*m_tocheck)(tocheck(src), tocheck(negated_obj));
get(src).well_formed();
}
};
table_intersection_filter_fn * check_table_plugin::mk_filter_by_negation_fn(const table_base & t,
const table_base & negated_obj, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * negated_cols) {
if (check_kind(t) && check_kind(negated_obj)) {
return alloc(filter_by_negation_fn, *this, t, negated_obj, joined_col_cnt, t_cols, negated_cols);
}
return 0;
}
// ------------------
// check_table
check_table::check_table(check_table_plugin & p, const table_signature & sig):
table_base(p, sig) {
(well_formed());
}
check_table::check_table(check_table_plugin & p, const table_signature & sig, table_base* tocheck, table_base* checker):
table_base(p, sig),
m_checker(checker),
m_tocheck(tocheck) {
well_formed();
}
check_table::~check_table() {
m_tocheck->deallocate();
m_checker->deallocate();
}
bool check_table::well_formed() const {
get_plugin().m_count++;
iterator it = m_tocheck->begin(), end = m_tocheck->end();
for (; it != end; ++it) {
table_fact fact;
it->get_fact(fact);
if (!m_checker->contains_fact(fact)) {
m_tocheck->display(verbose_stream());
m_checker->display(verbose_stream());
verbose_stream() << get_plugin().m_count << "\n";
UNREACHABLE();
fatal_error(0);
return false;
}
}
iterator it2 = m_checker->begin(), end2 = m_checker->end();
for (; it2 != end2; ++it2) {
table_fact fact;
it2->get_fact(fact);
if (!m_tocheck->contains_fact(fact)) {
m_tocheck->display(verbose_stream());
m_checker->display(verbose_stream());
verbose_stream() << get_plugin().m_count << "\n";
UNREACHABLE();
fatal_error(0);
return false;
}
}
return true;
}
bool check_table::empty() const {
if (m_tocheck->empty() != m_checker->empty()) {
m_tocheck->display(verbose_stream());
m_checker->display(verbose_stream());
verbose_stream() << get_plugin().m_count << "\n";
fatal_error(0);
}
return m_tocheck->empty();
}
void check_table::add_fact(const table_fact & f) {
IF_VERBOSE(1, verbose_stream() << __FUNCTION__ << "\n";);
m_tocheck->add_fact(f);
m_checker->add_fact(f);
well_formed();
}
void check_table::remove_fact(const table_element* f) {
IF_VERBOSE(1, verbose_stream() << __FUNCTION__ << "\n";);
m_tocheck->remove_fact(f);
m_checker->remove_fact(f);
well_formed();
}
bool check_table::contains_fact(const table_fact & f) const {
return m_checker->contains_fact(f);
}
table_base * check_table::clone() const {
IF_VERBOSE(1, verbose_stream() << __FUNCTION__ << "\n";);
check_table* result = alloc(check_table, get_plugin(), get_signature(), m_tocheck->clone(), m_checker->clone());
return result;
}
table_base * check_table::complement(func_decl* p, const table_element * func_columns) const {
check_table* result = alloc(check_table, get_plugin(), get_signature(), m_tocheck->complement(p, func_columns), m_checker->complement(p, func_columns));
return result;
}
};

135
src/muz/dl_check_table.h Normal file
View file

@ -0,0 +1,135 @@
/*++
Copyright (c) 2010 Microsoft Corporation
Module Name:
dl_check_table.h
Abstract:
<abstract>
Author:
Nikolaj Bjorner (nbjorner) 2010-11-15
Revision History:
--*/
#ifndef _DL_CHECK_TABLE_H_
#define _DL_CHECK_TABLE_H_
#include "dl_base.h"
#include "dl_decl_plugin.h"
#include "dl_relation_manager.h"
namespace datalog {
class check_table;
class check_table_plugin : public table_plugin {
friend class check_table;
table_plugin& m_checker;
table_plugin& m_tocheck;
unsigned m_count;
protected:
class join_fn;
class join_project_fn;
class union_fn;
class transformer_fn;
class rename_fn;
class project_fn;
class select_equal_and_project_fn;
class filter_equal_fn;
class filter_identical_fn;
class filter_interpreted_fn;
class filter_interpreted_and_project_fn;
class filter_by_negation_fn;
public:
check_table_plugin(relation_manager & manager, symbol const& checker, symbol const& tocheck)
: table_plugin(symbol("check"), manager),
m_checker(*manager.get_table_plugin(checker)),
m_tocheck(*manager.get_table_plugin(tocheck)), m_count(0) {}
virtual table_base * mk_empty(const table_signature & s);
virtual table_join_fn * mk_join_fn(const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
virtual table_join_fn * mk_join_project_fn(const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2, unsigned removed_col_cnt,
const unsigned * removed_cols);
virtual table_union_fn * mk_union_fn(const table_base & tgt, const table_base & src,
const table_base * delta);
virtual table_transformer_fn * mk_project_fn(const table_base & t, unsigned col_cnt,
const unsigned * removed_cols);
virtual table_transformer_fn * mk_select_equal_and_project_fn(const table_base & t,
const table_element & value, unsigned col);
virtual table_transformer_fn * mk_rename_fn(const table_base & t, unsigned permutation_cycle_len,
const unsigned * permutation_cycle);
virtual table_mutator_fn * mk_filter_identical_fn(const table_base & t, unsigned col_cnt,
const unsigned * identical_cols);
virtual table_mutator_fn * mk_filter_equal_fn(const table_base & t, const table_element & value,
unsigned col);
virtual table_mutator_fn * mk_filter_interpreted_fn(const table_base & t, app * condition);
virtual table_transformer_fn * mk_filter_interpreted_and_project_fn(const table_base & t,
app * condition, unsigned removed_col_cnt, const unsigned * removed_cols);
virtual table_intersection_filter_fn * mk_filter_by_negation_fn(
const table_base & t,
const table_base & negated_obj, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * negated_cols);
virtual bool can_handle_signature(table_signature const& s);
private:
static check_table& get(table_base& r);
static check_table const & get(table_base const& r);
static table_base& checker(table_base& r);
static table_base const& checker(table_base const& r);
static table_base* checker(table_base* r);
static table_base const* checker(table_base const* r);
static table_base& tocheck(table_base& r);
static table_base const& tocheck(table_base const& r);
static table_base* tocheck(table_base* r);
static table_base const* tocheck(table_base const* r);
};
class check_table : public table_base {
friend class check_table_plugin;
table_base* m_checker;
table_base* m_tocheck;
check_table(check_table_plugin & p, const table_signature & sig);
check_table(check_table_plugin & p, const table_signature & sig, table_base* tocheck, table_base* checker);
virtual ~check_table();
bool well_formed() const;
public:
check_table_plugin & get_plugin() const {
return static_cast<check_table_plugin &>(table_base::get_plugin());
}
virtual bool empty() const;
virtual void add_fact(const table_fact & f);
virtual void remove_fact(const table_element* fact);
virtual bool contains_fact(const table_fact & f) const;
virtual table_base * complement(func_decl* p, const table_element * func_columns = 0) const;
virtual table_base * clone() const;
virtual iterator begin() const { SASSERT(well_formed()); return m_tocheck->begin(); }
virtual iterator end() const { return m_tocheck->end(); }
virtual unsigned get_size_estimate_rows() const { return m_tocheck->get_size_estimate_rows(); }
virtual unsigned get_size_estimate_bytes() const { return m_tocheck->get_size_estimate_bytes(); }
};
};
#endif /* _DL_CHECK_TABLE_H_ */

467
src/muz/dl_cmds.cpp Normal file
View file

@ -0,0 +1,467 @@
/*++
Copyright (c) 2011 Microsoft Corporation
Module Name:
dl_cmds.cpp
Abstract:
Datalog commands for SMT2 front-end.
Author:
Leonardo (leonardo) 2011-03-28
Notes:
--*/
#include"cmd_context.h"
#include"dl_cmds.h"
#include"dl_external_relation.h"
#include"dl_context.h"
#include"dl_decl_plugin.h"
#include"dl_instruction.h"
#include"dl_compiler.h"
#include"dl_rule.h"
#include"ast_pp.h"
#include"parametric_cmd.h"
#include"cancel_eh.h"
#include"scoped_ctrl_c.h"
#include"scoped_timer.h"
#include"trail.h"
#include<iomanip>
struct dl_context {
smt_params m_fparams;
params_ref m_params_ref;
fixedpoint_params m_params;
cmd_context & m_cmd;
dl_collected_cmds* m_collected_cmds;
unsigned m_ref_count;
datalog::dl_decl_plugin* m_decl_plugin;
scoped_ptr<datalog::context> m_context;
trail_stack<dl_context> m_trail;
fixedpoint_params const& get_params() {
init();
return m_context->get_params();
}
dl_context(cmd_context & ctx, dl_collected_cmds* collected_cmds):
m_params(m_params_ref),
m_cmd(ctx),
m_collected_cmds(collected_cmds),
m_ref_count(0),
m_decl_plugin(0),
m_trail(*this) {}
void inc_ref() {
++m_ref_count;
}
void dec_ref() {
--m_ref_count;
if (0 == m_ref_count) {
dealloc(this);
}
}
void init() {
ast_manager& m = m_cmd.m();
if (!m_context) {
m_context = alloc(datalog::context, m, m_fparams, m_params_ref);
}
if (!m_decl_plugin) {
symbol name("datalog_relation");
if (m.has_plugin(name)) {
m_decl_plugin = static_cast<datalog::dl_decl_plugin*>(m_cmd.m().get_plugin(m.mk_family_id(name)));
}
else {
m_decl_plugin = alloc(datalog::dl_decl_plugin);
m.register_plugin(symbol("datalog_relation"), m_decl_plugin);
}
}
}
void reset() {
m_context = 0;
}
void register_predicate(func_decl* pred, unsigned num_kinds, symbol const* kinds) {
if (m_collected_cmds) {
m_collected_cmds->m_rels.push_back(pred);
m_trail.push(push_back_vector<dl_context, func_decl_ref_vector>(m_collected_cmds->m_rels));
}
dlctx().register_predicate(pred, false);
dlctx().set_predicate_representation(pred, num_kinds, kinds);
}
void add_rule(expr * rule, symbol const& name) {
init();
if (m_collected_cmds) {
expr_ref rl = m_context->bind_variables(rule, true);
m_collected_cmds->m_rules.push_back(rl);
m_collected_cmds->m_names.push_back(name);
m_trail.push(push_back_vector<dl_context, expr_ref_vector>(m_collected_cmds->m_rules));
m_trail.push(push_back_vector<dl_context, svector<symbol> >(m_collected_cmds->m_names));
}
else {
m_context->add_rule(rule, name);
}
}
bool collect_query(expr* q) {
if (m_collected_cmds) {
expr_ref qr = m_context->bind_variables(q, false);
m_collected_cmds->m_queries.push_back(qr);
m_trail.push(push_back_vector<dl_context, expr_ref_vector>(m_collected_cmds->m_queries));
return true;
}
else {
return false;
}
}
void push() {
m_trail.push_scope();
dlctx().push();
}
void pop() {
m_trail.pop_scope(1);
dlctx().pop();
}
datalog::context & dlctx() {
init();
return *m_context;
}
};
/**
\brief rule command. It is also the owner of dl_context object.
*/
class dl_rule_cmd : public cmd {
ref<dl_context> m_dl_ctx;
mutable unsigned m_arg_idx;
expr* m_t;
symbol m_name;
public:
dl_rule_cmd(dl_context * dl_ctx):
cmd("rule"),
m_dl_ctx(dl_ctx),
m_arg_idx(0),
m_t(0) {}
virtual char const * get_usage() const { return "(forall (q) (=> (and body) head)) :optional-name"; }
virtual char const * get_descr(cmd_context & ctx) const { return "add a Horn rule."; }
virtual unsigned get_arity() const { return VAR_ARITY; }
virtual cmd_arg_kind next_arg_kind(cmd_context & ctx) const {
switch(m_arg_idx) {
case 0: return CPK_EXPR;
case 1: return CPK_SYMBOL;
default: return CPK_SYMBOL;
}
}
virtual void set_next_arg(cmd_context & ctx, expr * t) {
m_t = t;
m_arg_idx++;
}
virtual void set_next_arg(cmd_context & ctx, symbol const & s) {
m_name = s;
}
virtual void reset(cmd_context & ctx) { m_dl_ctx->reset(); prepare(ctx); }
virtual void prepare(cmd_context& ctx) { m_arg_idx = 0; m_name = symbol::null; }
virtual void finalize(cmd_context & ctx) {
}
virtual void execute(cmd_context & ctx) {
m_dl_ctx->add_rule(m_t, m_name);
}
};
class dl_query_cmd : public parametric_cmd {
ref<dl_context> m_dl_ctx;
expr* m_target;
public:
dl_query_cmd(dl_context * dl_ctx):
parametric_cmd("query"),
m_dl_ctx(dl_ctx),
m_target(0) {
}
virtual char const * get_usage() const { return "(exists (q) (and body))"; }
virtual char const * get_main_descr() const {
return "pose a query based on the Horn rules.";
}
virtual cmd_arg_kind next_arg_kind(cmd_context & ctx) const {
if (m_target == 0) return CPK_EXPR;
return parametric_cmd::next_arg_kind(ctx);
}
virtual void set_next_arg(cmd_context & ctx, expr * t) {
m_target = t;
}
virtual void prepare(cmd_context & ctx) {
parametric_cmd::prepare(ctx);
m_target = 0;
}
virtual void execute(cmd_context& ctx) {
if (m_target == 0) {
throw cmd_exception("invalid query command, argument expected");
}
if (m_dl_ctx->collect_query(m_target)) {
return;
}
datalog::context& dlctx = m_dl_ctx->dlctx();
set_background(ctx);
dlctx.updt_params(m_params);
unsigned timeout = m_dl_ctx->get_params().timeout();
cancel_eh<datalog::context> eh(dlctx);
lbool status = l_undef;
{
scoped_ctrl_c ctrlc(eh);
scoped_timer timer(timeout, &eh);
cmd_context::scoped_watch sw(ctx);
try {
status = dlctx.query(m_target);
}
catch (z3_error & ex) {
ctx.regular_stream() << "(error \"query failed: " << ex.msg() << "\")" << std::endl;
throw ex;
}
catch (z3_exception& ex) {
ctx.regular_stream() << "(error \"query failed: " << ex.msg() << "\")" << std::endl;
}
}
switch (status) {
case l_false:
ctx.regular_stream() << "unsat\n";
print_certificate(ctx);
break;
case l_true:
ctx.regular_stream() << "sat\n";
print_answer(ctx);
print_certificate(ctx);
break;
case l_undef:
ctx.regular_stream() << "unknown\n";
switch(dlctx.get_status()) {
case datalog::INPUT_ERROR:
ctx.regular_stream() << "input error\n";
break;
case datalog::MEMOUT:
ctx.regular_stream() << "memory bounds exceeded\n";
break;
case datalog::TIMEOUT:
ctx.regular_stream() << "timeout\n";
break;
case datalog::APPROX:
ctx.regular_stream() << "approximated relations\n";
break;
case datalog::OK:
UNREACHABLE();
break;
case datalog::CANCELED:
ctx.regular_stream() << "canceled\n";
dlctx.display_profile(ctx.regular_stream());
break;
default:
UNREACHABLE();
break;
}
break;
}
dlctx.cleanup();
print_statistics(ctx);
m_target = 0;
}
virtual void init_pdescrs(cmd_context & ctx, param_descrs & p) {
m_dl_ctx->dlctx().collect_params(p);
}
private:
void set_background(cmd_context& ctx) {
datalog::context& dlctx = m_dl_ctx->dlctx();
ptr_vector<expr>::const_iterator it = ctx.begin_assertions();
ptr_vector<expr>::const_iterator end = ctx.end_assertions();
for (; it != end; ++it) {
dlctx.assert_expr(*it);
}
}
void print_answer(cmd_context& ctx) {
if (m_dl_ctx->get_params().print_answer()) {
datalog::context& dlctx = m_dl_ctx->dlctx();
ast_manager& m = ctx.m();
expr_ref query_result(dlctx.get_answer_as_formula(), m);
sbuffer<symbol> var_names;
unsigned num_decls = 0;
if (is_quantifier(m_target)) {
num_decls = to_quantifier(m_target)->get_num_decls();
}
ctx.display(ctx.regular_stream(), query_result, 0, num_decls, "X", var_names);
ctx.regular_stream() << std::endl;
}
}
void print_statistics(cmd_context& ctx) {
if (m_dl_ctx->get_params().print_statistics()) {
statistics st;
datalog::context& dlctx = m_dl_ctx->dlctx();
unsigned long long max_mem = memory::get_max_used_memory();
unsigned long long mem = memory::get_allocation_size();
dlctx.collect_statistics(st);
st.update("time", ctx.get_seconds());
st.update("memory", static_cast<double>(mem)/static_cast<double>(1024*1024));
st.update("max-memory", static_cast<double>(max_mem)/static_cast<double>(1024*1024));
st.display_smt2(ctx.regular_stream());
}
}
void print_certificate(cmd_context& ctx) {
if (m_dl_ctx->get_params().print_certificate()) {
datalog::context& dlctx = m_dl_ctx->dlctx();
dlctx.display_certificate(ctx.regular_stream());
ctx.regular_stream() << "\n";
}
}
};
class dl_declare_rel_cmd : public cmd {
ref<dl_context> m_dl_ctx;
unsigned m_arg_idx;
mutable unsigned m_query_arg_idx;
symbol m_rel_name;
scoped_ptr<sort_ref_vector> m_domain;
svector<symbol> m_kinds;
void ensure_domain(cmd_context& ctx) {
if (!m_domain) m_domain = alloc(sort_ref_vector, ctx.m());
}
public:
dl_declare_rel_cmd(dl_context * dl_ctx):
cmd("declare-rel"),
m_dl_ctx(dl_ctx),
m_domain(0) {}
virtual char const * get_usage() const { return "<symbol> (<arg1 sort> ...) <representation>*"; }
virtual char const * get_descr(cmd_context & ctx) const { return "declare new relation"; }
virtual unsigned get_arity() const { return VAR_ARITY; }
virtual void prepare(cmd_context & ctx) {
m_arg_idx = 0;
m_query_arg_idx = 0;
m_domain = 0;
m_kinds.reset();
}
virtual cmd_arg_kind next_arg_kind(cmd_context & ctx) const {
switch(m_query_arg_idx++) {
case 0: return CPK_SYMBOL; // relation name
case 1: return CPK_SORT_LIST; // arguments
default: return CPK_SYMBOL; // optional representation specification
}
}
virtual void set_next_arg(cmd_context & ctx, unsigned num, sort * const * slist) {
ensure_domain(ctx);
m_domain->append(num, slist);
m_arg_idx++;
}
virtual void set_next_arg(cmd_context & ctx, symbol const & s) {
if(m_arg_idx==0) {
m_rel_name = s;
}
else {
SASSERT(m_arg_idx>1);
m_kinds.push_back(s);
}
m_arg_idx++;
}
virtual void execute(cmd_context & ctx) {
if(m_arg_idx<2) {
throw cmd_exception("at least 2 arguments expected");
}
ensure_domain(ctx);
ast_manager& m = ctx.m();
func_decl_ref pred(
m.mk_func_decl(m_rel_name, m_domain->size(), m_domain->c_ptr(), m.mk_bool_sort()), m);
ctx.insert(pred);
m_dl_ctx->register_predicate(pred, m_kinds.size(), m_kinds.c_ptr());
m_domain = 0;
}
};
class dl_declare_var_cmd : public cmd {
unsigned m_arg_idx;
symbol m_var_name;
sort* m_var_sort;
ref<dl_context> m_dl_ctx;
public:
dl_declare_var_cmd(dl_context* dl_ctx):
cmd("declare-var"),
m_arg_idx(0),
m_dl_ctx(dl_ctx)
{}
virtual char const * get_usage() const { return "<symbol> <sort>"; }
virtual char const * get_descr(cmd_context & ctx) const { return "declare constant as variable"; }
virtual unsigned get_arity() const { return 2; }
virtual void prepare(cmd_context & ctx) {
m_arg_idx = 0;
}
virtual cmd_arg_kind next_arg_kind(cmd_context & ctx) const {
SASSERT(m_arg_idx <= 1);
if (m_arg_idx == 0) {
return CPK_SYMBOL;
}
return CPK_SORT;
}
virtual void set_next_arg(cmd_context & ctx, sort* s) {
m_var_sort = s;
++m_arg_idx;
}
virtual void set_next_arg(cmd_context & ctx, symbol const & s) {
m_var_name = s;
++m_arg_idx;
}
virtual void execute(cmd_context & ctx) {
ast_manager& m = ctx.m();
func_decl_ref var(m.mk_func_decl(m_var_name, 0, static_cast<sort*const*>(0), m_var_sort), m);
ctx.insert(var);
m_dl_ctx->dlctx().register_variable(var);
}
};
static void install_dl_cmds_aux(cmd_context& ctx, dl_collected_cmds* collected_cmds) {
dl_context * dl_ctx = alloc(dl_context, ctx, collected_cmds);
ctx.insert(alloc(dl_rule_cmd, dl_ctx));
ctx.insert(alloc(dl_query_cmd, dl_ctx));
ctx.insert(alloc(dl_declare_rel_cmd, dl_ctx));
ctx.insert(alloc(dl_declare_var_cmd, dl_ctx));
}
void install_dl_cmds(cmd_context & ctx) {
install_dl_cmds_aux(ctx, 0);
}
void install_dl_collect_cmds(dl_collected_cmds& collected_cmds, cmd_context & ctx) {
install_dl_cmds_aux(ctx, &collected_cmds);
}

37
src/muz/dl_cmds.h Normal file
View file

@ -0,0 +1,37 @@
/*++
Copyright (c) 2011 Microsoft Corporation
Module Name:
dl_cmds.h
Abstract:
Datalog commands for SMT2 front-end.
Author:
Nikolaj Bjorner (nbjorner) 2012-11-17
Notes:
--*/
#ifndef _DL_CMDS_H_
#define _DL_CMDS_H_
#include "ast.h"
class cmd_context;
struct dl_collected_cmds {
expr_ref_vector m_rules;
svector<symbol> m_names;
expr_ref_vector m_queries;
func_decl_ref_vector m_rels;
dl_collected_cmds(ast_manager& m) : m_rules(m), m_queries(m), m_rels(m) {}
};
void install_dl_cmds(cmd_context & ctx);
void install_dl_collect_cmds(dl_collected_cmds& collected_cmds, cmd_context& ctx);
#endif

1344
src/muz/dl_compiler.cpp Normal file

File diff suppressed because it is too large Load diff

282
src/muz/dl_compiler.h Normal file
View file

@ -0,0 +1,282 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_compiler.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-09-14.
Revision History:
--*/
#ifndef _DL_COMPILER_H_
#define _DL_COMPILER_H_
#include<iostream>
#include<list>
#include<utility>
#include "ast.h"
#include "hashtable.h"
#include "map.h"
#include "obj_pair_hashtable.h"
#include "ref_vector.h"
#include "vector.h"
#include "dl_base.h"
#include "dl_context.h"
#include "dl_instruction.h"
namespace datalog {
class compiler {
typedef instruction::reg_idx reg_idx;
typedef hashtable<unsigned, u_hash, u_eq> int_set;
typedef u_map<unsigned> int2int;
typedef u_map<unsigned_vector> int2ints;
typedef obj_map<func_decl, reg_idx> pred2idx;
typedef unsigned_vector var_vector;
typedef ptr_vector<func_decl> func_decl_vector;
enum assembling_column_kind {
ACK_BOUND_VAR,
ACK_UNBOUND_VAR,
ACK_CONSTANT
};
/**
\brief instruction for assembling head relation from a joint relation built
from rule evaluation.
ACK_BOUND_VAR(source_column) - encodes that the column contains a variable
bound in the body.
ACK_UNBOUND_VAR(var_index) - encodes that the column contains a variable that
is unbound (by the corresponding rule body),
var_index is the de-Brujin index (var->get_idx())
of the variable associated with the column.
ACK_CONSTANT(constant) - encodes that the column contains the constant.
Examples:
P(x) :- Q(x,y), Q(y,z)
The variables in the body relation are [x, y, y, z] indexed as 0, 1, 2, 3.
The variable x gets the instruction ACK_BOUND_VAR(0)
P(u,x) :- Q(x,y), Q(y,z)
The variable u gets the instruction ACK_UNBOUND_VAR(#0)
P(1, x) :- Q(x,y), Q(y,z)
The instruction for column 0 is ACK_CONSTANT(1)
*/
struct assembling_column_info {
relation_sort domain; // domain of the column
assembling_column_kind kind; // "instruction" tag
unsigned source_column; // for ACK_BOUND_VAR
unsigned var_index; // for ACK_UNBOUND_VAR
relation_element constant; // for ACK_CONSTANT
};
class instruction_observer : public instruction_block::instruction_observer {
compiler & m_parent;
rule * m_current;
public:
instruction_observer(compiler & parent) : m_parent(parent), m_current(0) {}
void start_rule(rule * r) { SASSERT(!m_current); m_current=r; }
void finish_rule() { m_current = 0; }
virtual void notify(instruction * i) {
if(m_current) {
i->set_accounting_parent_object(m_parent.m_context, m_current);
}
}
};
context & m_context;
rule_set const & m_rule_set;
/**
Invariant: the \c m_top_level_code never contains the loop that is being constructed,
so instruction that need to be executed before the loop can be pushed into it.
*/
instruction_block & m_top_level_code;
pred2idx m_pred_regs;
reg_idx m_new_reg;
vector<relation_signature> m_reg_signatures;
obj_pair_map<sort, app, reg_idx> m_constant_registers;
obj_pair_map<sort, decl, reg_idx> m_total_registers;
obj_map<decl, reg_idx> m_empty_tables_registers;
instruction_observer m_instruction_observer;
/**
If true, the union operation on the underlying structure only provides the information
whether the updated relation has changed or not. In this case we do not get anything
from using delta relations at position of input relations in the saturation loop, so we
would not do it.
*/
bool all_or_nothing_deltas() const { return m_context.all_or_nothing_deltas(); }
/**
If true, we compile the saturation loops in a way that allows us to use widening.
*/
bool compile_with_widening() const { return m_context.compile_with_widening(); }
reg_idx get_fresh_register(const relation_signature & sig);
reg_idx get_single_column_register(const relation_sort & s);
/**
\brief Allocate registers for predicates in \c pred and add them into the \c regs map.
\c regs must not already contain any predicate from \c preds.
*/
void get_fresh_registers(const func_decl_set & preds, pred2idx & regs);
void make_join(reg_idx t1, reg_idx t2, const variable_intersection & vars, reg_idx & result,
instruction_block & acc);
void make_join_project(reg_idx t1, reg_idx t2, const variable_intersection & vars,
const unsigned_vector & removed_cols, reg_idx & result, instruction_block & acc);
void make_filter_interpreted_and_project(reg_idx src, app_ref & cond,
const unsigned_vector & removed_cols, reg_idx & result, instruction_block & acc);
void make_select_equal_and_project(reg_idx src, const relation_element & val, unsigned col,
reg_idx & result, instruction_block & acc);
/**
\brief Create add an union or widen operation and put it into \c acc.
*/
void make_union(reg_idx src, reg_idx tgt, reg_idx delta, bool widening, instruction_block & acc);
void make_projection(reg_idx src, unsigned col_cnt, const unsigned * removed_cols,
reg_idx & result, instruction_block & acc);
void make_rename(reg_idx src, unsigned cycle_len, const unsigned * permutation_cycle,
reg_idx & result, instruction_block & acc);
void make_clone(reg_idx src, reg_idx & result, instruction_block & acc);
/**
\brief Into \c acc add code that will assemble columns of a relation according to description
in \c acis0. The source for bound variables is the table in register \c src.
If \c src is \c execution_context::void_register, it is assumed to be a full relation
with empty signature.
*/
void make_assembling_code(rule* compiled_rule, func_decl* head_pred, reg_idx src, const svector<assembling_column_info> & acis0,
reg_idx & result, bool & dealloc, instruction_block & acc);
void make_dealloc_non_void(reg_idx r, instruction_block & acc);
void make_add_constant_column(func_decl* pred, reg_idx src, const relation_sort & s, const relation_element & val,
reg_idx & result, bool & dealloc, instruction_block & acc);
void make_add_unbound_column(rule* compiled_rule, unsigned col_idx, func_decl* pred, reg_idx src, const relation_sort & s, reg_idx & result,
bool & dealloc, instruction_block & acc);
void make_full_relation(func_decl* pred, const relation_signature & sig, reg_idx & result,
instruction_block & acc);
void add_unbound_columns_for_negation(rule* compiled_rule, func_decl* pred, reg_idx& single_res, expr_ref_vector& single_res_expr,
bool & dealloc, instruction_block& acc);
void make_duplicate_column(reg_idx src, unsigned col, reg_idx & result, instruction_block & acc);
void ensure_predicate_loaded(func_decl * pred, instruction_block & acc);
/**
\brief For rule \c r with two positive uninterpreted predicates put into \c res indexes of
local variables in a table that results from join of the two positive predicates.
Used to get input for the "project" part of join-project.
*/
void get_local_indexes_for_projection(rule * r, unsigned_vector & res);
void get_local_indexes_for_projection(app * t, var_counter & globals, unsigned ofs,
unsigned_vector & res);
/**
\brief Into \c acc add instructions that will add new facts following from the rule into
\c head_reg, and add the facts that were not in \c head_reg before into \c delta_reg.
*/
void compile_rule_evaluation_run(rule * r, reg_idx head_reg, const reg_idx * tail_regs,
reg_idx delta_reg, bool use_widening, instruction_block & acc);
void compile_rule_evaluation(rule * r, const pred2idx * input_deltas, reg_idx output_delta,
bool use_widening, instruction_block & acc);
/**
\brief Generate code to evaluate rules corresponding to predicates in \c head_preds.
The rules are evaluated in the order their heads appear in the \c head_preds vector.
*/
void compile_preds(const func_decl_vector & head_preds, const func_decl_set & widened_preds,
const pred2idx * input_deltas, const pred2idx & output_deltas, instruction_block & acc);
/**
\brief Generate code to evaluate predicates in a stratum based on their non-recursive rules.
*/
void compile_preds_init(const func_decl_vector & head_preds, const func_decl_set & widened_preds,
const pred2idx * input_deltas, const pred2idx & output_deltas, instruction_block & acc);
void make_inloop_delta_transition(const pred2idx & global_head_deltas,
const pred2idx & global_tail_deltas, const pred2idx & local_deltas, instruction_block & acc);
void compile_loop(const func_decl_vector & head_preds, const func_decl_set & widened_preds,
const pred2idx & global_head_deltas, const pred2idx & global_tail_deltas,
const pred2idx & local_deltas, instruction_block & acc);
void compile_dependent_rules(const func_decl_set & head_preds,
const pred2idx * input_deltas, const pred2idx & output_deltas,
bool add_saturation_marks, instruction_block & acc);
void detect_chains(const func_decl_set & preds, func_decl_vector & ordered_preds,
func_decl_set & global_deltas);
/**
Return true if there is no dependency inside the \c rules stratum.
The head predicates in stratum must be strongly connected by dependency.
*/
bool is_nonrecursive_stratum(const func_decl_set & preds) const;
/**
input_deltas==0 --> we use the actual content of relations instead of deltas
*/
void compile_nonrecursive_stratum(const func_decl_set & preds,
const pred2idx * input_deltas, const pred2idx & output_deltas,
bool add_saturation_marks, instruction_block & acc);
void compile_strats(const rule_stratifier & stratifier,
const pred2idx * input_deltas, const pred2idx & output_deltas,
bool add_saturation_marks, instruction_block & acc);
bool all_saturated(const func_decl_set & preds) const;
void reset();
explicit compiler(context & ctx, rule_set const & rules, instruction_block & top_level_code)
: m_context(ctx),
m_rule_set(rules),
m_top_level_code(top_level_code),
m_instruction_observer(*this) {}
/**
\brief Compile \c rules in to pseudocode.
Instructions to load data and perform computations put into \c execution_code
*/
void do_compilation(instruction_block & execution_code,
instruction_block & termination_code);
public:
static void compile(context & ctx, rule_set const & rules, instruction_block & execution_code,
instruction_block & termination_code) {
compiler(ctx, rules, execution_code)
.do_compilation(execution_code, termination_code);
}
};
};
#endif /* _DL_COMPILER_H_ */

1366
src/muz/dl_context.cpp Normal file

File diff suppressed because it is too large Load diff

492
src/muz/dl_context.h Normal file
View file

@ -0,0 +1,492 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_context.h
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2010-05-18.
Revision History:
--*/
#ifndef _DL_CONTEXT_H_
#define _DL_CONTEXT_H_
#ifdef _CYGWIN
#undef min
#undef max
#endif
#include"arith_decl_plugin.h"
#include"map.h"
#include"th_rewriter.h"
#include"str_hashtable.h"
#include"var_subst.h"
#include"dl_base.h"
#include"dl_costs.h"
#include"dl_decl_plugin.h"
#include"dl_relation_manager.h"
#include"dl_rule_set.h"
#include"pdr_dl_interface.h"
#include"dl_bmc_engine.h"
#include"tab_context.h"
#include"rel_context.h"
#include"lbool.h"
#include"statistics.h"
#include"params.h"
#include"trail.h"
#include"model_converter.h"
#include"model2expr.h"
#include"smt_params.h"
#include"dl_rule_transformer.h"
#include"expr_abstract.h"
#include"expr_functors.h"
#include"clp_context.h"
namespace datalog {
enum execution_result {
OK,
TIMEOUT,
MEMOUT,
INPUT_ERROR,
APPROX,
CANCELED
};
class context {
public:
typedef unsigned finite_element;
enum sort_kind {
SK_UINT64,
SK_SYMBOL
};
private:
class sort_domain;
class symbol_sort_domain;
class uint64_sort_domain;
class restore_rules;
class contains_pred;
typedef hashtable<symbol, symbol_hash_proc, symbol_eq_proc> symbol_set;
typedef map<symbol, func_decl*, symbol_hash_proc, symbol_eq_proc> sym2decl;
typedef obj_map<const func_decl, svector<symbol> > pred2syms;
typedef obj_map<const sort, sort_domain*> sort_domain_map;
class contains_pred : public i_expr_pred {
context const& ctx;
public:
contains_pred(context& ctx): ctx(ctx) {}
virtual ~contains_pred() {}
virtual bool operator()(expr* e) {
return ctx.is_predicate(e);
}
};
ast_manager & m;
smt_params & m_fparams;
params_ref m_params_ref;
fixedpoint_params m_params;
dl_decl_util m_decl_util;
th_rewriter m_rewriter;
var_subst m_var_subst;
rule_manager m_rule_manager;
unused_vars_eliminator m_elim_unused_vars;
expr_abstractor m_abstractor;
contains_pred m_contains_p;
check_pred m_check_pred;
rule_transformer m_transf;
trail_stack<context> m_trail;
ast_ref_vector m_pinned;
app_ref_vector m_vars;
svector<symbol> m_names;
sort_domain_map m_sorts;
func_decl_set m_preds;
sym2decl m_preds_by_name;
pred2syms m_argument_var_names;
rule_set m_rule_set;
rule_set m_transformed_rule_set;
unsigned m_rule_fmls_head;
expr_ref_vector m_rule_fmls;
svector<symbol> m_rule_names;
expr_ref_vector m_background;
model_converter_ref m_mc;
proof_converter_ref m_pc;
rel_context* m_rel;
scoped_ptr<engine_base> m_engine;
bool m_closed;
bool m_saturation_was_run;
execution_result m_last_status;
expr_ref m_last_answer;
DL_ENGINE m_engine_type;
volatile bool m_cancel;
bool is_fact(app * head) const;
bool has_sort_domain(relation_sort s) const;
sort_domain & get_sort_domain(relation_sort s);
const sort_domain & get_sort_domain(relation_sort s) const;
class engine_type_proc;
public:
context(ast_manager & m, smt_params& fp, params_ref const& p = params_ref());
~context();
void reset();
void push();
void pop();
bool saturation_was_run() const { return m_saturation_was_run; }
void notify_saturation_was_run() { m_saturation_was_run = true; }
void configure_engine();
ast_manager & get_manager() const { return m; }
rule_manager & get_rule_manager() { return m_rule_manager; }
smt_params & get_fparams() const { return m_fparams; }
fixedpoint_params const& get_params() const { return m_params; }
DL_ENGINE get_engine() { configure_engine(); return m_engine_type; }
th_rewriter& get_rewriter() { return m_rewriter; }
var_subst & get_var_subst() { return m_var_subst; }
dl_decl_util & get_decl_util() { return m_decl_util; }
bool generate_proof_trace() const { return m_params.generate_proof_trace(); }
bool output_profile() const { return m_params.output_profile(); }
bool fix_unbound_vars() const { return m_params.fix_unbound_vars(); }
symbol default_table() const { return m_params.default_table(); }
symbol default_relation() const { return m_params.default_relation(); } // external_relation_plugin::get_name());
symbol default_table_checker() const { return m_params.default_table_checker(); }
bool default_table_checked() const { return m_params.default_table_checked(); }
bool dbg_fpr_nonempty_relation_signature() const { return m_params.dbg_fpr_nonempty_relation_signature(); }
unsigned dl_profile_milliseconds_threshold() const { return m_params.profile_timeout_milliseconds(); }
bool all_or_nothing_deltas() const { return m_params.all_or_nothing_deltas(); }
bool compile_with_widening() const { return m_params.compile_with_widening(); }
bool unbound_compressor() const { return m_params.unbound_compressor(); }
bool similarity_compressor() const { return m_params.similarity_compressor(); }
unsigned similarity_compressor_threshold() const { return m_params.similarity_compressor_threshold(); }
unsigned soft_timeout() const { return m_fparams.m_soft_timeout; }
unsigned initial_restart_timeout() const { return m_params.initial_restart_timeout(); }
bool generate_explanations() const { return m_params.generate_explanations(); }
bool explanations_on_relation_level() const { return m_params.explanations_on_relation_level(); }
bool magic_sets_for_queries() const { return m_params.magic_sets_for_queries(); }
bool eager_emptiness_checking() const { return m_params.eager_emptiness_checking(); }
void register_finite_sort(sort * s, sort_kind k);
/**
Register uninterpreted constant to be used as an implicitly quantified variable.
The variable gets quantified in the formula passed to rule::mk_rule_from_horn.
*/
void register_variable(func_decl* var);
expr_ref bind_variables(expr* fml, bool is_forall);
/**
Register datalog relation.
If names is true, we associate the predicate with its name, so that it can be
retrieved by the try_get_predicate_decl() function. Auxiliary predicates introduced
e.g. by rule transformations do not need to be named.
*/
void register_predicate(func_decl * pred, bool named);
/**
Restrict reltaions to set of predicates.
*/
void restrict_predicates(func_decl_set const& preds);
/**
\brief Retrieve predicates
*/
func_decl_set const& get_predicates() const { return m_preds; }
bool is_predicate(func_decl* pred) const { return m_preds.contains(pred); }
bool is_predicate(expr * e) const { return is_app(e) && is_predicate(to_app(e)->get_decl()); }
/**
\brief If a predicate name has a \c func_decl object assigned, return pointer to it;
otherwise return 0.
Not all \c func_decl object used as relation identifiers need to be assigned to their
names. Generally, the names coming from the parses are registered here.
*/
func_decl * try_get_predicate_decl(symbol const& pred_name) const {
func_decl * res = 0;
m_preds_by_name.find(pred_name, res);
return res;
}
/**
\brief Create a fresh head predicate declaration.
*/
func_decl * mk_fresh_head_predicate(symbol const & prefix, symbol const & suffix,
unsigned arity, sort * const * domain, func_decl* orig_pred=0);
/**
\brief Return number of a symbol in a DK_SYMBOL kind sort (\see register_sort() )
*/
finite_element get_constant_number(relation_sort sort, symbol s);
/**
\brief Return number of a symbol in a DK_UINT64 kind sort (\see register_sort() )
*/
finite_element get_constant_number(relation_sort srt, uint64 el);
/**
\brief Output name of constant with number \c num in sort \c sort.
*/
void print_constant_name(relation_sort sort, uint64 num, std::ostream & out);
bool try_get_sort_constant_count(relation_sort srt, uint64 & constant_count);
uint64 get_sort_size_estimate(relation_sort srt);
/**
\brief Assign names of variables used in the declaration of a predicate.
These names are used when printing out the relations to make the output conform
to the one of bddbddb.
*/
void set_argument_names(const func_decl * pred, svector<symbol> var_names);
symbol get_argument_name(const func_decl * pred, unsigned arg_index);
void set_predicate_representation(func_decl * pred, unsigned relation_name_cnt,
symbol const * relation_names);
void set_output_predicate(func_decl * pred) { m_rule_set.set_output_predicate(pred); }
rule_set & get_rules() { flush_add_rules(); return m_rule_set; }
void get_rules_as_formulas(expr_ref_vector& fmls, svector<symbol>& names);
void add_fact(app * head);
void add_fact(func_decl * pred, const relation_fact & fact);
bool has_facts(func_decl * pred) const;
void add_rule(rule_ref& r);
void assert_expr(expr* e);
expr_ref get_background_assertion();
unsigned get_num_assertions() { return m_background.size(); }
expr* get_assertion(unsigned i) const { return m_background[i]; }
/**
Method exposed from API for adding rules.
*/
void add_rule(expr* rl, symbol const& name);
/**
Update a named rule.
*/
void update_rule(expr* rl, symbol const& name);
/**
Retrieve the maximal number of relevant unfoldings of 'pred'
with respect to the current state.
*/
unsigned get_num_levels(func_decl* pred);
/**
Retrieve the current cover of 'pred' up to 'level' unfoldings.
Return just the delta that is known at 'level'. To
obtain the full set of properties of 'pred' one should query
at 'level+1', 'level+2' etc, and include level=-1.
*/
expr_ref get_cover_delta(int level, func_decl* pred);
/**
Add a property of predicate 'pred' at 'level'.
It gets pushed forward when possible.
*/
void add_cover(int level, func_decl* pred, expr* property);
/**
\brief Check rule subsumption.
*/
bool check_subsumes(rule const& stronger_rule, rule const& weaker_rule);
/**
\brief Check if rule is well-formed according to engine.
*/
void check_rule(rule_ref& r);
/**
\brief Return true if facts to \c pred can be added using the \c add_table_fact() function.
This function should return true if \c pred is represented by a table_relation
and there is no transformation of relation values before they are put into the
table.
*/
void add_table_fact(func_decl * pred, const table_fact & fact);
void add_table_fact(func_decl * pred, unsigned num_args, unsigned args[]);
/**
\brief To be called after all rules are added.
*/
void close();
void ensure_closed();
bool is_closed() { return m_closed; }
/**
\brief Undo the effect of the \c close operation.
*/
void reopen();
void ensure_opened();
model_converter_ref& get_model_converter() { return m_mc; }
void add_model_converter(model_converter* mc) { m_mc = concat(m_mc.get(), mc); }
proof_converter_ref& get_proof_converter() { return m_pc; }
void add_proof_converter(proof_converter* pc) { m_pc = concat(m_pc.get(), pc); }
void transform_rules();
void transform_rules(rule_transformer& transf);
void transform_rules(rule_transformer::plugin* plugin);
void replace_rules(rule_set const& rs);
void record_transformed_rules();
void apply_default_transformation();
void collect_params(param_descrs& r);
void updt_params(params_ref const& p);
void display_rules(std::ostream & out) const {
m_rule_set.display(out);
}
void display(std::ostream & out) const {
display_rules(out);
if (m_rel) m_rel->display_facts(out);
}
void display_smt2(unsigned num_queries, expr* const* queries, std::ostream& out);
void display_profile(std::ostream& out) const;
// -----------------------------------
//
// basic usage methods
//
// -----------------------------------
void cancel();
bool canceled() const { return m_cancel; }
void cleanup();
void reset_cancel() { cleanup(); }
/**
\brief check if query 'q' is satisfied under asserted rules and background.
If successful, return OK and into \c result assign a relation with all
tuples matching the query. Otherwise return reason for failure and do not modify
\c result.
The numbers of variables in the query body must form a contiguous sequence
starting from zero.
The caller becomes an owner of the relation object returned in \c result. The
relation object, however, should not outlive the datalog context since it is
linked to a relation plugin in the context.
*/
lbool query(expr* q);
/**
\brief retrieve model from inductive invariant that shows query is unsat.
\pre engine == 'pdr' - this option is only supported for PDR mode.
*/
model_ref get_model();
/**
\brief retrieve proof from derivation of the query.
\pre engine == 'pdr' - this option is only supported for PDR mode.
*/
proof_ref get_proof();
/**
Query multiple output relations.
*/
lbool rel_query(unsigned num_rels, func_decl * const* rels);
/**
\brief retrieve last proof status.
*/
execution_result get_status();
void set_status(execution_result r) { m_last_status = r; }
/**
\brief retrieve formula corresponding to query that returns l_true.
The formula describes one or more instances of the existential variables
in the query that are derivable.
*/
expr* get_answer_as_formula();
void collect_statistics(statistics& st) const;
void reset_statistics();
/**
\brief Display a certificate for reachability and/or unreachability.
*/
void display_certificate(std::ostream& out);
/**
\brief query result if it contains fact.
*/
bool result_contains_fact(relation_fact const& f);
rel_context* get_rel_context() { ensure_engine(); return m_rel; }
private:
/**
Just reset all tables.
*/
void reset_tables();
void flush_add_rules();
void ensure_engine();
void check_quantifier_free(rule_ref& r);
void check_uninterpreted_free(rule_ref& r);
void check_existential_tail(rule_ref& r);
void check_positive_predicates(rule_ref& r);
// auxilary functions for SMT2 pretty-printer.
void declare_vars(expr_ref_vector& rules, mk_fresh_name& mk_fresh, std::ostream& out);
//undefined and private copy constructor and operator=
context(const context&);
context& operator=(const context&);
};
};
#endif /* _DL_CONTEXT_H_ */

160
src/muz/dl_costs.cpp Normal file
View file

@ -0,0 +1,160 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_costs.cpp
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-09-20.
Revision History:
--*/
#include "debug.h"
#include "stopwatch.h"
#include "dl_context.h"
#include "dl_rule.h"
#include "dl_costs.h"
namespace datalog {
// -----------------------------------
//
// costs
//
// -----------------------------------
costs::costs() : milliseconds(0), instructions(0) {}
bool costs::empty() const {
return !milliseconds && !instructions;
}
void costs::reset() {
milliseconds = 0;
instructions = 0;
}
costs costs::operator-(const costs & o) const {
costs res(*this);
SASSERT(milliseconds>o.milliseconds);
res.milliseconds-=o.milliseconds;
SASSERT(instructions>o.instructions);
res.instructions-=o.instructions;
return res;
}
void costs::operator+=(const costs & o) {
milliseconds+=o.milliseconds;
instructions+=o.instructions;
}
bool costs::passes_thresholds(context & ctx) const {
return milliseconds >= ctx.dl_profile_milliseconds_threshold();
}
void costs::output(std::ostream & out) const {
out << "instr: " << instructions << " time: " << milliseconds << "ms";
}
// -----------------------------------
//
// accounted_object
//
// -----------------------------------
accounted_object::~accounted_object() {
if(m_parent_object) {
SASSERT(m_context);
m_context->get_rule_manager().dec_ref(m_parent_object);
}
}
void accounted_object::set_accounting_parent_object(context & ctx, rule * parent) {
if(m_parent_object) {
SASSERT(m_context);
SASSERT(m_context==&ctx);
m_context->get_rule_manager().dec_ref(m_parent_object);
}
m_context = &ctx;
m_parent_object = parent;
m_context->get_rule_manager().inc_ref(m_parent_object);
}
void accounted_object::process_costs() {
costs delta = get_current_costs();
if(delta.empty()) {
return;
}
get_current_costs().reset();
accounted_object * obj = this;
do {
obj->m_processed_cost+=delta;
obj=obj->m_parent_object;
} while(obj);
}
void accounted_object::get_total_cost(costs & result) const {
result.reset();
result+=m_current_cost;
result+=m_processed_cost;
}
bool accounted_object::passes_output_thresholds(context & ctx) const {
costs c;
get_total_cost(c);
return c.passes_thresholds(ctx);
}
void accounted_object::output_profile(std::ostream & out) const {
costs c;
get_total_cost(c);
c.output(out);
}
// -----------------------------------
//
// cost_recorder
//
// -----------------------------------
cost_recorder::cost_recorder() : m_obj(0) {
m_stopwatch = alloc(stopwatch);
m_stopwatch->start();
}
cost_recorder::~cost_recorder() {
if(m_obj) {
finish();
}
dealloc(m_stopwatch);
}
void cost_recorder::start(accounted_object * obj) {
uint64 curr_time = static_cast<uint64>(m_stopwatch->get_current_seconds()*1000);
if(m_obj) {
costs::time_type time_delta = static_cast<costs::time_type>(curr_time-m_last_time);
costs & c = m_obj->get_current_costs();
c.instructions++;
c.milliseconds+=time_delta;
m_obj->m_being_recorded = false;
}
m_running = obj!=0;
m_obj = obj;
m_last_time = curr_time;
if(obj) {
m_obj->m_being_recorded = true;
}
}
};

115
src/muz/dl_costs.h Normal file
View file

@ -0,0 +1,115 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_costs.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-09-20.
Revision History:
--*/
#ifndef _DL_COSTS_H_
#define _DL_COSTS_H_
#include<iosfwd>
#include "ast.h"
class stopwatch;
namespace datalog {
class context;
class rule;
class cost_recorder;
struct costs {
typedef unsigned time_type;
time_type milliseconds;
unsigned instructions;
costs();
bool empty() const;
void reset();
costs operator-(const costs & o) const;
void operator+=(const costs & o);
bool passes_thresholds(context & ctx) const;
void output(std::ostream & out) const;
};
class accounted_object {
friend class cost_recorder;
context * m_context;
rule * m_parent_object;
costs m_current_cost;
costs m_processed_cost;
bool m_being_recorded;
protected:
accounted_object() : m_context(0), m_parent_object(0), m_being_recorded(false) {}
~accounted_object();
public:
void set_accounting_parent_object(context & ctx, rule * parent);
rule * get_parent_object() const { return m_parent_object; }
costs & get_current_costs() { return m_current_cost; }
const costs & get_current_costs() const { return m_current_cost; }
const costs & get_processed_costs() const { return m_processed_cost; }
void get_total_cost(costs & result) const;
bool being_recorded() const { return m_being_recorded; }
void process_costs();
bool passes_output_thresholds(context & ctx) const;
void output_profile(std::ostream & out) const;
private:
//private and undefined copy constructor and operator= to avoid the default ones
accounted_object(const accounted_object &);
accounted_object& operator=(const accounted_object &);
};
class cost_recorder {
accounted_object * m_obj;
//it's a pointer to avoid everything depending on the stopwatch.h
//(and transitively then on windows.h) header file
stopwatch * m_stopwatch;
bool m_running;
uint64 m_last_time;
public:
cost_recorder();
~cost_recorder();
/**
\brief Start recording costs for the next object.
If the recording of the previous object did not finish, it will be finished here.
Also, it will be done more efficiently than if the \c finish() function was called
before separately.
*/
void start(accounted_object *);
void finish() { start(0); }
};
};
#endif /* _DL_COSTS_H_ */

View file

@ -0,0 +1,456 @@
/*++
Copyright (c) 2010 Microsoft Corporation
Module Name:
dl_external_relation.cpp
Abstract:
<abstract>
Author:
Nikolaj Bjorner (nbjorner) 2010-05-10
Revision History:
--*/
#include "debug.h"
#include "ast_pp.h"
#include "dl_context.h"
#include "dl_external_relation.h"
#include "dl_decl_plugin.h"
namespace datalog {
external_relation::external_relation(external_relation_plugin & p, const relation_signature & s, expr* r)
: relation_base(p, s),
m_rel(r, p.get_ast_manager()),
m_select_fn(p.get_ast_manager()),
m_store_fn(p.get_ast_manager()),
m_is_empty_fn(p.get_ast_manager())
{
}
external_relation::~external_relation() {
}
void external_relation::mk_accessor(decl_kind k, func_decl_ref& fn, const relation_fact& f, bool destructive, expr_ref& res) const {
ast_manager& m = m_rel.get_manager();
family_id fid = get_plugin().get_family_id();
ptr_vector<expr> args;
args.push_back(m_rel);
for (unsigned i = 0; i < f.size(); ++i) {
args.push_back(f[i]);
}
if (!fn.get()) {
fn = m.mk_func_decl(fid, k, 0, 0, args.size(), args.c_ptr());
}
if (destructive) {
get_plugin().reduce_assign(fn, args.size(), args.c_ptr(), 1, args.c_ptr());
res = m_rel;
}
else {
get_plugin().reduce(fn, args.size(), args.c_ptr(), res);
}
}
bool external_relation::empty() const {
ast_manager& m = m_rel.get_manager();
expr* r = m_rel.get();
expr_ref res(m);
if (!m_is_empty_fn.get()) {
family_id fid = get_plugin().get_family_id();
const_cast<func_decl_ref&>(m_is_empty_fn) = m.mk_func_decl(fid, OP_RA_IS_EMPTY, 0, 0, 1, &r);
}
get_plugin().reduce(m_is_empty_fn, 1, &r, res);
return m.is_true(res);
}
void external_relation::add_fact(const relation_fact & f) {
mk_accessor(OP_RA_STORE, m_store_fn, f, true, m_rel);
}
bool external_relation::contains_fact(const relation_fact & f) const {
ast_manager& m = get_plugin().get_ast_manager();
expr_ref res(m);
mk_accessor(OP_RA_SELECT, const_cast<func_decl_ref&>(m_select_fn), f, false, res);
return !m.is_false(res);
}
external_relation * external_relation::clone() const {
ast_manager& m = m_rel.get_manager();
family_id fid = get_plugin().get_family_id();
expr* rel = m_rel.get();
expr_ref res(m.mk_fresh_const("T", m.get_sort(rel)), m);
expr* rel_out = res.get();
func_decl_ref fn(m.mk_func_decl(fid, OP_RA_CLONE,0,0, 1, &rel), m);
get_plugin().reduce_assign(fn, 1, &rel, 1, &rel_out);
return alloc(external_relation, get_plugin(), get_signature(), res);
}
external_relation * external_relation::complement(func_decl* p) const {
ast_manager& m = m_rel.get_manager();
family_id fid = get_plugin().get_family_id();
expr_ref res(m);
expr* rel = m_rel;
func_decl_ref fn(m.mk_func_decl(fid, OP_RA_COMPLEMENT,0,0, 1, &rel), m);
get_plugin().reduce(fn, 1, &rel, res);
return alloc(external_relation, get_plugin(), get_signature(), res);
}
void external_relation::display(std::ostream & out) const {
out << mk_pp(m_rel, m_rel.get_manager()) << "\n";
}
void external_relation::display_tuples(func_decl & pred, std::ostream & out) const {
display(out);
}
external_relation_plugin & external_relation::get_plugin() const {
return static_cast<external_relation_plugin &>(relation_base::get_plugin());
}
// -----------------------------------
//
// external_relation_plugin
//
// -----------------------------------
external_relation_plugin::external_relation_plugin(external_relation_context& ctx, relation_manager & m)
: relation_plugin(external_relation_plugin::get_name(), m), m_ext(ctx) {}
external_relation const & external_relation_plugin::get(relation_base const& r) {
return dynamic_cast<external_relation const&>(r);
}
external_relation & external_relation_plugin::get(relation_base & r) {
return dynamic_cast<external_relation&>(r);
}
relation_base * external_relation_plugin::mk_empty(const relation_signature & s) {
ast_manager& m = get_ast_manager();
sort* r_sort = get_relation_sort(s);
parameter param(r_sort);
family_id fid = get_family_id();
expr_ref e(m.mk_fresh_const("T", r_sort), m);
expr* args[1] = { e.get() };
func_decl_ref empty_decl(m.mk_func_decl(fid, OP_RA_EMPTY, 1, &param, 0, (sort*const*)0), m);
reduce_assign(empty_decl, 0, 0, 1, args);
return alloc(external_relation, *this, s, e);
}
sort* external_relation_plugin::get_relation_sort(relation_signature const& sig) {
vector<parameter> sorts;
ast_manager& m = get_ast_manager();
family_id fid = get_family_id();
for (unsigned i = 0; i < sig.size(); ++i) {
sorts.push_back(parameter(sig[i]));
}
return m.mk_sort(fid, DL_RELATION_SORT, sorts.size(), sorts.c_ptr());
}
sort* external_relation_plugin::get_column_sort(unsigned col, sort* s) {
SASSERT(s->get_num_parameters() > col);
SASSERT(s->get_parameter(col).is_ast());
SASSERT(is_sort(s->get_parameter(col).get_ast()));
return to_sort(s->get_parameter(col).get_ast());
}
family_id external_relation_plugin::get_family_id() {
return m_ext.get_family_id();
}
void external_relation_plugin::mk_filter_fn(sort* s, app* condition, func_decl_ref& f) {
ast_manager& m = get_ast_manager();
family_id fid = get_family_id();
parameter param(condition);
f = m.mk_func_decl(fid, OP_RA_FILTER, 1, &param, 1, &s);
}
class external_relation_plugin::join_fn : public convenient_relation_join_fn {
external_relation_plugin& m_plugin;
func_decl_ref m_join_fn;
expr* m_args[2];
public:
join_fn(external_relation_plugin& p, const relation_signature & o1_sig, const relation_signature & o2_sig, unsigned col_cnt,
const unsigned * cols1, const unsigned * cols2)
: convenient_relation_join_fn(o1_sig, o2_sig, col_cnt, cols1, cols2),
m_plugin(p),
m_join_fn(p.get_ast_manager()) {
ast_manager& m = p.get_ast_manager();
family_id fid = p.get_family_id();
vector<parameter> params;
for (unsigned i = 0; i < col_cnt; ++i) {
params.push_back(parameter(cols1[i]));
params.push_back(parameter(cols2[i]));
}
sort* domain[2] = { p.get_relation_sort(o1_sig), p.get_relation_sort(o2_sig) };
m_join_fn = m.mk_func_decl(fid, OP_RA_JOIN, params.size(), params.c_ptr(), 2, domain);
}
virtual relation_base * operator()(const relation_base & r1, const relation_base & r2) {
expr_ref res(m_plugin.get_ast_manager());
m_args[0] = get(r1).get_relation();
m_args[1] = get(r2).get_relation();
m_plugin.reduce(m_join_fn, 2, m_args, res);
return alloc(external_relation, m_plugin, get_result_signature(), res);
}
};
relation_join_fn * external_relation_plugin::mk_join_fn(const relation_base & r1, const relation_base & r2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2) {
if (!check_kind(r1) || !check_kind(r2)) {
return 0;
}
return alloc(join_fn, *this, r1.get_signature(), r2.get_signature() , col_cnt, cols1, cols2);
}
class external_relation_plugin::project_fn : public convenient_relation_project_fn {
external_relation_plugin& m_plugin;
func_decl_ref m_project_fn;
public:
project_fn(external_relation_plugin& p, sort* relation_sort,
const relation_signature & orig_sig, unsigned removed_col_cnt, const unsigned * removed_cols)
: convenient_relation_project_fn(orig_sig, removed_col_cnt, removed_cols),
m_plugin(p),
m_project_fn(p.get_ast_manager()) {
vector<parameter> params;
ast_manager& m = p.get_ast_manager();
family_id fid = p.get_family_id();
for (unsigned i = 0; i < removed_col_cnt; ++i) {
params.push_back(parameter(removed_cols[i]));
}
m_project_fn = m.mk_func_decl(fid, OP_RA_PROJECT, params.size(), params.c_ptr(), 1, &relation_sort);
}
virtual relation_base * operator()(const relation_base & r) {
expr_ref res(m_plugin.get_ast_manager());
expr* rel = get(r).get_relation();
m_plugin.reduce(m_project_fn, 1, &rel, res);
return alloc(external_relation, m_plugin, get_result_signature(), to_app(res));
}
};
relation_transformer_fn * external_relation_plugin::mk_project_fn(const relation_base & r,
unsigned col_cnt, const unsigned * removed_cols) {
return alloc(project_fn, *this, get(r).get_sort(), r.get_signature(), col_cnt, removed_cols);
}
class external_relation_plugin::rename_fn : public convenient_relation_rename_fn {
external_relation_plugin& m_plugin;
func_decl_ref m_rename_fn;
expr* m_args[2];
public:
rename_fn(external_relation_plugin& p, sort* relation_sort, const relation_signature & orig_sig, unsigned cycle_len, const unsigned * cycle)
: convenient_relation_rename_fn(orig_sig, cycle_len, cycle),
m_plugin(p),
m_rename_fn(p.get_ast_manager()) {
ast_manager& m = p.get_ast_manager();
family_id fid = p.get_family_id();
vector<parameter> params;
for (unsigned i = 0; i < cycle_len; ++i) {
SASSERT(cycle[i] < orig_sig.size());
params.push_back(parameter(cycle[i]));
}
m_rename_fn = m.mk_func_decl(fid, OP_RA_RENAME, params.size(), params.c_ptr(), 1, &relation_sort);
}
virtual relation_base * operator()(const relation_base & r) {
expr* rel = get(r).get_relation();
expr_ref res(m_plugin.get_ast_manager());
m_args[0] = rel;
m_plugin.reduce(m_rename_fn, 1, &rel, res);
return alloc(external_relation, m_plugin, get_result_signature(), res);
}
};
relation_transformer_fn * external_relation_plugin::mk_rename_fn(const relation_base & r,
unsigned cycle_len, const unsigned * permutation_cycle) {
if(!check_kind(r)) {
return 0;
}
return alloc(rename_fn, *this, get(r).get_sort(), r.get_signature(), cycle_len, permutation_cycle);
}
class external_relation_plugin::union_fn : public relation_union_fn {
external_relation_plugin& m_plugin;
func_decl_ref m_union_fn;
expr* m_args[2];
expr* m_outs[2];
public:
union_fn(external_relation_plugin& p, decl_kind k, sort* relation_sort):
m_plugin(p),
m_union_fn(p.get_ast_manager()) {
ast_manager& m = p.get_ast_manager();
sort* domain[2] = { relation_sort, relation_sort };
m_union_fn = m.mk_func_decl(p.get_family_id(), k, 0, 0, 2, domain);
}
virtual void operator()(relation_base & r, const relation_base & src, relation_base * delta) {
ast_manager& m = m_plugin.get_ast_manager();
expr_ref_vector res(m);
m_args[0] = get(r).get_relation();
m_args[1] = get(src).get_relation();
m_outs[0] = m_args[0];
unsigned num_out = 1;
if (delta) {
m_outs[1] = get(*delta).get_relation();
++num_out;
}
m_plugin.reduce_assign(m_union_fn, 2, m_args, num_out, m_outs);
}
};
relation_union_fn * external_relation_plugin::mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta) {
if (!check_kind(tgt) || !check_kind(src) || (delta && !check_kind(*delta))) {
return 0;
}
return alloc(union_fn, *this, OP_RA_UNION, get(src).get_sort());
}
relation_union_fn * external_relation_plugin::mk_widen_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta) {
if (!check_kind(tgt) || !check_kind(src) || (delta && !check_kind(*delta))) {
return 0;
}
return alloc(union_fn, *this, OP_RA_WIDEN, get(src).get_sort());
}
class external_relation_plugin::filter_interpreted_fn : public relation_mutator_fn {
external_relation_plugin& m_plugin;
app_ref m_condition;
func_decl_ref m_filter_fn;
public:
filter_interpreted_fn(external_relation_plugin& p, sort* relation_sort, app * condition)
: m_plugin(p),
m_condition(condition, p.get_ast_manager()),
m_filter_fn(p.get_ast_manager()) {
ast_manager& m = p.get_ast_manager();
p.mk_filter_fn(relation_sort, condition, m_filter_fn);
SASSERT(m.is_bool(condition));
}
virtual void operator()(relation_base & r) {
SASSERT(m_plugin.check_kind(r));
expr* arg = get(r).get_relation();
m_plugin.reduce_assign(m_filter_fn, 1, &arg, 1, &arg);
}
};
relation_mutator_fn * external_relation_plugin::mk_filter_interpreted_fn(const relation_base & r, app * condition) {
if(!check_kind(r)) {
return 0;
}
return alloc(filter_interpreted_fn, *this, get(r).get_sort(), condition);
}
relation_mutator_fn * external_relation_plugin::mk_filter_equal_fn(const relation_base & r,
const relation_element & value, unsigned col) {
if(!check_kind(r)) {
return 0;
}
ast_manager& m = get_ast_manager();
app_ref condition(m);
expr_ref var(m);
sort* relation_sort = get(r).get_sort();
sort* column_sort = get_column_sort(col, relation_sort);
var = m.mk_var(col, column_sort);
condition = m.mk_eq(var, value);
return mk_filter_interpreted_fn(r, condition);
}
class external_relation_plugin::filter_identical_fn : public relation_mutator_fn {
external_relation_plugin& m_plugin;
func_decl_ref_vector m_filter_fn;
public:
filter_identical_fn(external_relation_plugin& p, sort* relation_sort,
unsigned col_cnt, const unsigned * identical_cols)
: m_plugin(p), m_filter_fn(p.get_ast_manager()) {
ast_manager& m = p.get_ast_manager();
func_decl_ref fn(m);
app_ref eq(m);
if (col_cnt <= 1) {
return;
}
unsigned col = identical_cols[0];
sort* s = p.get_column_sort(col, relation_sort);
var* v0 = m.mk_var(col, s);
for (unsigned i = 1; i < col_cnt; ++i) {
col = identical_cols[i];
s = p.get_column_sort(col, relation_sort);
eq = m.mk_eq(v0, m.mk_var(col, s));
p.mk_filter_fn(relation_sort, eq.get(), fn);
m_filter_fn.push_back(fn);
}
}
virtual void operator()(relation_base & r) {
expr* r0 = get(r).get_relation();
for (unsigned i = 0; i < m_filter_fn.size(); ++i) {
m_plugin.reduce_assign(m_filter_fn[i].get(), 1, &r0, 1, &r0);
}
}
};
relation_mutator_fn * external_relation_plugin::mk_filter_identical_fn(const relation_base & r,
unsigned col_cnt, const unsigned * identical_cols) {
if (!check_kind(r)) {
return 0;
}
return alloc(filter_identical_fn, *this, get(r).get_sort(), col_cnt, identical_cols);
}
class external_relation_plugin::negation_filter_fn : public convenient_relation_negation_filter_fn {
external_relation_plugin& m_plugin;
func_decl_ref m_negated_filter_fn;
expr* m_args[2];
public:
negation_filter_fn(external_relation_plugin& p,
const relation_base & tgt, const relation_base & neg_t,
unsigned joined_col_cnt, const unsigned * t_cols, const unsigned * negated_cols) :
convenient_negation_filter_fn(tgt, neg_t, joined_col_cnt, t_cols, negated_cols),
m_plugin(p),
m_negated_filter_fn(p.get_ast_manager())
{
ast_manager& m = p.get_ast_manager();
family_id fid = p.get_family_id();
vector<parameter> params;
for (unsigned i = 0; i < joined_col_cnt; ++i) {
params.push_back(parameter(t_cols[i]));
params.push_back(parameter(negated_cols[i]));
}
sort* domain[2] = { get(tgt).get_sort(), get(neg_t).get_sort() };
m_negated_filter_fn = m.mk_func_decl(fid, OP_RA_NEGATION_FILTER, params.size(), params.c_ptr(), 2, domain);
}
void operator()(relation_base & t, const relation_base & negated_obj) {
m_args[0] = get(t).get_relation();
m_args[1] = get(negated_obj).get_relation();
m_plugin.reduce_assign(m_negated_filter_fn.get(), 2, m_args, 1, m_args);
}
};
relation_intersection_filter_fn * external_relation_plugin::mk_filter_by_negation_fn(const relation_base & t,
const relation_base & negated_obj, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * negated_cols) {
if (!check_kind(t) || !check_kind(negated_obj)) {
return 0;
}
return alloc(negation_filter_fn, *this, t, negated_obj, joined_col_cnt, t_cols, negated_cols);
}
};

View file

@ -0,0 +1,154 @@
/*++
Copyright (c) 2010 Microsoft Corporation
Module Name:
dl_external_relation.h
Abstract:
<abstract>
Author:
Nikolaj Bjorner (nbjorner) 2010-05-10
Revision History:
--*/
#ifndef _DL_EXTERNAL_RELATION_H_
#define _DL_EXTERNAL_RELATION_H_
#include "dl_base.h"
namespace datalog {
class external_relation;
class external_relation_context {
public:
virtual ~external_relation_context() {}
virtual family_id get_family_id() const = 0;
// reduce arguments.
virtual void reduce(func_decl* f, unsigned num_args, expr * const* args, expr_ref& result) = 0;
// overwrite terms passed in outs vector with values computed by function.
virtual void reduce_assign(func_decl* f, unsigned num_args, expr * const* args, unsigned num_out, expr* const* outs) = 0;
};
class external_relation_plugin : public relation_plugin {
friend class external_relation;
class join_fn;
class project_fn;
class rename_fn;
class union_fn;
class filter_identical_fn;
class filter_interpreted_fn;
class negation_filter_fn;
external_relation_context& m_ext;
public:
external_relation_plugin(external_relation_context& ctx, relation_manager & m);
virtual bool can_handle_signature(const relation_signature & s) { return true; }
static symbol get_name() { return symbol("external_relation"); }
virtual relation_base * mk_empty(const relation_signature & s);
virtual relation_join_fn * mk_join_fn(const relation_base & t1, const relation_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
virtual relation_transformer_fn * mk_project_fn(const relation_base & t, unsigned col_cnt,
const unsigned * removed_cols);
virtual relation_transformer_fn * mk_rename_fn(const relation_base & t, unsigned permutation_cycle_len,
const unsigned * permutation_cycle);
virtual relation_union_fn * mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
virtual relation_union_fn * mk_widen_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
virtual relation_mutator_fn * mk_filter_identical_fn(const relation_base & t, unsigned col_cnt,
const unsigned * identical_cols);
virtual relation_mutator_fn * mk_filter_equal_fn(const relation_base & t, const relation_element & value,
unsigned col);
virtual relation_mutator_fn * mk_filter_interpreted_fn(const relation_base & t, app * condition);
virtual relation_intersection_filter_fn * mk_filter_by_negation_fn(const relation_base & t,
const relation_base & negated_obj, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * negated_cols);
private:
static external_relation& get(relation_base& r);
static external_relation const & get(relation_base const& r);
void reduce(func_decl* f, unsigned num_args, expr * const* args, expr_ref& result) {
m_ext.reduce(f, num_args, args, result);
}
void reduce_assign(func_decl* f, unsigned num_args, expr * const* args, unsigned num_out, expr* const* outs) {
m_ext.reduce_assign(f, num_args, args, num_out, outs);
}
sort* get_relation_sort(relation_signature const& sig);
sort* get_column_sort(unsigned col, sort* relation_sort);
void mk_filter_fn(sort* s, app* condition, func_decl_ref& f);
family_id get_family_id();
};
class external_relation : public relation_base {
friend class external_relation_plugin;
friend class external_relation_plugin::join_fn;
friend class external_relation_plugin::project_fn;
friend class external_relation_plugin::rename_fn;
friend class external_relation_plugin::union_fn;
friend class external_relation_plugin::filter_identical_fn;
friend class external_relation_plugin::filter_interpreted_fn;
friend class external_relation_plugin::negation_filter_fn;
expr_ref m_rel;
func_decl_ref m_select_fn;
func_decl_ref m_store_fn;
func_decl_ref m_is_empty_fn;
unsigned size() const { return get_signature().size(); }
sort* get_sort() const { return m_rel.get_manager().get_sort(m_rel); }
void mk_accessor(decl_kind k, func_decl_ref& fn, const relation_fact& f, bool destructive, expr_ref& res) const;
external_relation(external_relation_plugin & p, const relation_signature & s, expr* r);
virtual ~external_relation();
public:
external_relation_plugin & get_plugin() const;
virtual bool empty() const;
virtual void add_fact(const relation_fact & f);
virtual bool contains_fact(const relation_fact & f) const;
virtual external_relation * clone() const;
virtual external_relation * complement(func_decl*) const;
virtual void display(std::ostream & out) const;
virtual void display_tuples(func_decl & pred, std::ostream & out) const;
expr* get_relation() const { return m_rel.get(); }
virtual void to_formula(expr_ref& fml) const { fml = get_relation(); }
};
};
#endif

File diff suppressed because it is too large Load diff

View file

@ -0,0 +1,366 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_finite_product_relation.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-09-24.
Revision History:
--*/
#ifndef _DL_FINITE_PRODUCT_RELATION_H_
#define _DL_FINITE_PRODUCT_RELATION_H_
#include "dl_base.h"
#include "dl_relation_manager.h"
#include "dl_table_relation.h"
namespace datalog {
class finite_product_relation;
void universal_delete(finite_product_relation* ptr);
class finite_product_relation_plugin : public relation_plugin {
friend class finite_product_relation;
public:
struct rel_spec {
family_id m_inner_kind; //null_family_id means we don't care about the kind
svector<bool> m_table_cols;
rel_spec() : m_inner_kind(null_family_id) {}
rel_spec(const svector<bool>& table_cols)
: m_inner_kind(null_family_id), m_table_cols(table_cols) {}
bool operator==(const rel_spec & o) const {
return m_inner_kind==o.m_inner_kind && vectors_equal(m_table_cols, o.m_table_cols);
}
struct hash {
unsigned operator()(const rel_spec & o) const {
return o.m_inner_kind^svector_hash<bool_hash>()(o.m_table_cols);
}
};
};
private:
class join_fn;
class converting_join_fn;
class project_fn;
class rename_fn;
class union_fn;
class inner_singleton_union_fn;
class converting_union_fn;
class filter_identical_fn;
class filter_equal_fn;
class filter_interpreted_fn;
class negation_filter_fn;
class filter_identical_pairs_fn;
relation_plugin & m_inner_plugin;
rel_spec_store<rel_spec, rel_spec::hash, default_eq<rel_spec> > m_spec_store;
static symbol get_name(relation_plugin & inner_plugin);
family_id get_relation_kind(finite_product_relation & r, const bool * table_columns);
static void get_all_possible_table_columns(relation_manager & rmgr, const relation_signature & s,
svector<bool> & table_columns);
void get_all_possible_table_columns(const relation_signature & s, svector<bool> & table_columns) {
get_all_possible_table_columns(get_manager(), s, table_columns);
}
void split_signatures(const relation_signature & s, table_signature & table_sig,
relation_signature & remaining_sig);
void split_signatures(const relation_signature & s, const bool * table_columns,
table_signature & table_sig, relation_signature & remaining_sig);
public:
static finite_product_relation & get(relation_base & r);
static const finite_product_relation & get(const relation_base & r);
static finite_product_relation * get(relation_base * r);
static const finite_product_relation * get(const relation_base * r);
static finite_product_relation_plugin & get_plugin(relation_manager & rmgr, relation_plugin & inner);
finite_product_relation_plugin(relation_plugin & inner_plugin, relation_manager & manager);
virtual void initialize(family_id fid);
relation_plugin & get_inner_plugin() const { return m_inner_plugin; }
virtual bool can_handle_signature(const relation_signature & s);
virtual relation_base * mk_empty(const relation_signature & s);
/**
\c inner_kind==null_family_id means we don't care about the kind of the inner relation
*/
finite_product_relation * mk_empty(const relation_signature & s, const bool * table_columns,
family_id inner_kind=null_family_id);
finite_product_relation * mk_empty(const finite_product_relation & original);
virtual relation_base * mk_empty(const relation_base & original);
virtual relation_base * mk_empty(const relation_signature & s, family_id kind);
virtual relation_base * mk_full(func_decl* p, const relation_signature & s);
/**
\brief Return true if \c r can be converted to \c finite_product_relation_plugin either
by \c mk_from_table_relation or by \c mk_from_inner_relation.
*/
bool can_be_converted(const relation_base & r);
/**
If the conversion cannot be performed, 0 is returned.
*/
finite_product_relation * mk_from_table_relation(const table_relation & r);
finite_product_relation * mk_from_inner_relation(const relation_base & r);
bool can_convert_to_table_relation(const finite_product_relation & r);
table_relation * to_table_relation(const finite_product_relation & r);
protected:
virtual relation_join_fn * mk_join_fn(const relation_base & t1, const relation_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
virtual relation_transformer_fn * mk_project_fn(const relation_base & t, unsigned col_cnt,
const unsigned * removed_cols);
virtual relation_transformer_fn * mk_rename_fn(const relation_base & t, unsigned permutation_cycle_len,
const unsigned * permutation_cycle);
virtual relation_union_fn * mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
virtual relation_mutator_fn * mk_filter_identical_fn(const relation_base & t, unsigned col_cnt,
const unsigned * identical_cols);
virtual relation_mutator_fn * mk_filter_equal_fn(const relation_base & t, const relation_element & value,
unsigned col);
virtual relation_mutator_fn * mk_filter_interpreted_fn(const relation_base & t, app * condition);
virtual relation_intersection_filter_fn * mk_filter_by_negation_fn(const relation_base & t,
const relation_base & negated_obj, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * negated_cols);
private:
/**
\brief Create a filter that enforces equality between pairs of table and relation columns
The column numbers in arrays \c table_cols and \c rel_cols must be local to the table/inner relation.
*/
relation_mutator_fn * mk_filter_identical_pairs(const finite_product_relation & r, unsigned col_cnt,
const unsigned * table_cols, const unsigned * rel_cols);
/**
\brief Create a join-project operation that creates a table according to \c relation_table
but with references to relations updated and removed according to the content of \c filtered_table.
\c selected_columns contains sorted indexes of data columns in \c relation_table that are also in
the \c filtered_table (so that the first column in \c filtered_table corresponds to
\c selected_columns[0] -th column in \c relation_table etc...)
Signature of \c relation_table:
(data columns)(functional column with indexes of relation objects)
Signature of \c filtered_table:
(selected data columns)(non-functional column with original relation object indexes)
(functional column with indexes of filtered relation objects)
*/
static table_join_fn * mk_assembler_of_filter_result(const table_base & relation_table,
const table_base & filtered_table, const unsigned_vector & selected_columns);
};
class finite_product_relation : public relation_base {
friend class finite_product_relation_plugin;
friend class finite_product_relation_plugin::join_fn;
friend class finite_product_relation_plugin::project_fn;
friend class finite_product_relation_plugin::union_fn;
friend class finite_product_relation_plugin::rename_fn;
friend class finite_product_relation_plugin::inner_singleton_union_fn;
friend class finite_product_relation_plugin::filter_equal_fn;
friend class finite_product_relation_plugin::filter_identical_pairs_fn;
class live_rel_collection_reducer;
public:
/**
Size of this sort determines how many different relation objects can we refer to.
*/
static const table_sort s_rel_idx_sort;
/**
\brief The last column in the signature is a functional column with index of the
associated inner relation. The other columns correspond to the relation signature
according to \c m_table2sig.
It holds that \c m_table_sig.size()-1==m_table2sig.size()
*/
table_signature m_table_sig;
unsigned_vector m_table2sig; // (ordered list)
unsigned_vector m_sig2table; //index of corresponding table column or UINT_MAX
private:
relation_signature m_other_sig;
unsigned_vector m_other2sig; // (ordered list)
public:
unsigned_vector m_sig2other; //index of corresponding other relation column or UINT_MAX
private:
relation_plugin & m_other_plugin;
family_id m_other_kind;
mutable table_base * m_table;
public:
mutable relation_vector m_others;
private:
mutable unsigned_vector m_available_rel_indexes;
/**
\c UINT_MAX means uninitialized.
If we can get away with it, we want to have a single full relation to refer to.
*/
mutable unsigned m_full_rel_idx;
mutable idx_set m_live_rel_collection_acc;
mutable scoped_ptr<table_transformer_fn> m_live_rel_collection_project;
mutable scoped_ptr<table_intersection_filter_fn> m_empty_rel_removal_filter;
void recycle_rel_idx(unsigned idx) const;
// creates a full relation if it does not exist.
unsigned get_full_rel_idx();
public:
relation_base & get_inner_rel(table_element idx)
{ SASSERT(idx<UINT_MAX); return get_inner_rel(static_cast<unsigned>(idx)); }
relation_base & get_inner_rel(unsigned idx) { SASSERT(m_others[idx]); return *m_others[idx]; }
const relation_base & get_inner_rel(unsigned idx) const
{ return const_cast<finite_product_relation &>(*this).get_inner_rel(idx); }
unsigned get_next_rel_idx() const;
/**
The relation takes ownership of the \c inner object.
*/
void set_inner_rel(table_element idx, relation_base * inner)
{ SASSERT(idx<UINT_MAX); return set_inner_rel(static_cast<unsigned>(idx), inner); }
/**
The relation takes ownership of the \c inner object.
*/
void set_inner_rel(unsigned idx, relation_base * inner) {
SASSERT(!m_others[idx]);
SASSERT(inner);
m_others[idx] = inner;
}
table_base & get_table() { return *m_table; }
table_plugin & get_table_plugin() const { return get_table().get_plugin(); }
void garbage_collect(bool remove_empty) const;
/**
\brief Initialize an empty relation with table \c table_vals and relations in \c others.
The relation object takes ownership of relations inside the \c others vector.
If \c contiguous is true, it can be assumed that there are no zero elements in the \c others array.
*/
void init(const table_base & table_vals, const relation_vector & others, bool contiguous);
private:
/**
\brief Extract the values of table non-functional columns from the relation fact.
The value of the functional column which determines index of the inner relation is undefined.
*/
void extract_table_fact(const relation_fact rf, table_fact & tf) const;
/**
\brief Extract the values of the inner relation columns from the relation fact.
*/
void extract_other_fact(const relation_fact rf, relation_fact & of) const;
relation_base * mk_empty_inner();
relation_base * mk_full_inner(func_decl* pred);
void complement_self(func_decl* pred);
void collect_live_relation_indexes(idx_set & res) const;
/**
\brief Try to modify relations in \c rels so that they have the same columns corresponding to the table
and the inner relation (so that the union can be perofrmed on theim in a straightforward way).
Relations in \c rels must all have equal signature.
Even if the function fails and false is returned, some relations may already be modified. They are
in a valid state, but with different specification.
*/
static bool try_unify_specifications(ptr_vector<finite_product_relation> & rels);
bool try_modify_specification(const bool * table_cols);
virtual bool can_swap(const relation_base & r) const
{ return &get_plugin()==&r.get_plugin(); }
/**
\brief Swap content of the current relation with the content of \c r.
Both relations must come from the same plugin and be of the same signature.
*/
virtual void swap(relation_base & r);
/**
\brief Create a \c finite_product_relation object.
*/
finite_product_relation(finite_product_relation_plugin & p, const relation_signature & s,
const bool * table_columns, table_plugin & tplugin, relation_plugin & oplugin, family_id other_kind);
finite_product_relation(const finite_product_relation & r);
virtual ~finite_product_relation();
public:
context & get_context() const;
finite_product_relation_plugin & get_plugin() const {
return static_cast<finite_product_relation_plugin &>(relation_base::get_plugin());
}
bool is_table_column(unsigned col_idx) const { return m_sig2table[col_idx]!=UINT_MAX; }
const table_base & get_table() const { return *m_table; }
const relation_base & get_inner_rel(table_element idx) const
{ SASSERT(idx<UINT_MAX); return get_inner_rel(static_cast<unsigned>(idx)); }
/**
The function calls garbage_collect, so the internal state may change when it is called.
*/
virtual bool empty() const;
void reset() { m_table->reset(); garbage_collect(false); }
virtual void add_fact(const relation_fact & f);
virtual bool contains_fact(const relation_fact & f) const;
virtual finite_product_relation * clone() const;
virtual finite_product_relation * complement(func_decl* p) const;
virtual void display(std::ostream & out) const;
virtual void display_tuples(func_decl & pred, std::ostream & out) const;
virtual unsigned get_size_estimate_rows() const { return m_table->get_size_estimate_rows(); }
virtual unsigned get_size_estimate_bytes() const { return m_table->get_size_estimate_bytes(); }
virtual void to_formula(expr_ref& fml) const;
};
};
#endif /* _DL_FINITE_PRODUCT_RELATION_H_ */

1116
src/muz/dl_instruction.cpp Normal file

File diff suppressed because it is too large Load diff

348
src/muz/dl_instruction.h Normal file
View file

@ -0,0 +1,348 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_instruction.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-09-14.
Revision History:
--*/
#ifndef _DL_INSTRUCTION_H_
#define _DL_INSTRUCTION_H_
#include<iostream>
#include<string>
#include<utility>
#include "ast.h"
#include "vector.h"
#include "dl_base.h"
#include "dl_costs.h"
namespace datalog {
class execution_context;
class instruction_block;
class rel_context;
inline void check_overflow(unsigned i) {
if (i == UINT_MAX) {
throw out_of_memory_error();
}
}
// -----------------------------------
//
// execution_context
//
// -----------------------------------
class execution_context {
public:
typedef relation_base * reg_type;
typedef vector<reg_type> reg_vector;
typedef unsigned reg_idx;
/**
\brief A register number that should never be referenced to. Can stand e.g. for a tail
table in a rule with no tail.
*/
static const reg_idx void_register = UINT_MAX;
private:
typedef u_map<std::string> reg_annotations;
context & m_context;
reg_vector m_registers;
reg_annotations m_reg_annotation;
stopwatch * m_stopwatch;
unsigned m_timelimit_ms; //zero means no limit
/**
\brief If true, after every operation that may result in an empty relation, a check
for emptiness will be performed, and if a relation is empty, it will be deleted
and replaced by zero. This allows us to avoid performing operations that would have
no effect due to relation emptiness, but if the check for emptiness is expensive, its
cost may overcome the gains.
*/
bool m_eager_emptiness_checking;
public:
execution_context(context & context);
~execution_context();
void reset();
rel_context & get_rel_context();
void set_timelimit(unsigned time_in_ms);
void reset_timelimit();
bool should_terminate();
bool eager_emptiness_checking() const { return m_eager_emptiness_checking; }
/**
\brief Return reference to \c i -th register that contains pointer to a relation.
If register contains zero, it should be treated as if it contains an empty relation.
*/
reg_type reg(reg_idx i) const {
if (i >= m_registers.size()) {
return 0;
}
return m_registers[i];
}
/**
\brief Return value of the register and assign zero into it place.
*/
reg_type release_reg(reg_idx i) {
SASSERT(i < m_registers.size());
SASSERT(m_registers[i]);
reg_type res = m_registers[i];
m_registers[i] = 0;
return res;
}
/**
\brief Assign value to a register. If it was non-empty, deallocate its content first.
*/
void set_reg(reg_idx i, reg_type val) {
if (i >= m_registers.size()) {
check_overflow(i);
m_registers.resize(i+1,0);
}
if (m_registers[i]) {
m_registers[i]->deallocate();
}
m_registers[i] = val;
}
void make_empty(reg_idx i) {
if (reg(i)) {
set_reg(i, 0);
}
}
unsigned register_count() const {
return m_registers.size();
}
bool get_register_annotation(reg_idx reg, std::string & res) const {
return m_reg_annotation.find(reg, res);
}
void set_register_annotation(reg_idx reg, std::string str) {
m_reg_annotation.insert(reg, str);
}
void report_big_relations(unsigned threshold, std::ostream & out) const;
};
// -----------------------------------
//
// instruction
//
// -----------------------------------
/**
\brief Base class for instructions used in datalog saturation.
A relation in a register is owned by that register and is not referenced from anywhere else.
Instructions that move context of one register to another leave the source register empty
and deallocate the previous content of the target register.
*/
class instruction : public accounted_object {
typedef u_map<base_relation_fn *> fn_cache;
fn_cache m_fn_cache;
static const int rk_encode_base = 1024;
inline static unsigned encode_kind(family_id k)
{ SASSERT(k<rk_encode_base); return k; }
inline static unsigned encode_kinds(family_id k1, family_id k2)
{ SASSERT(k1<rk_encode_base && k2<rk_encode_base); return (k1+1)*rk_encode_base + k2; }
inline static unsigned encode_kinds(family_id k1, family_id k2, family_id k3) {
SASSERT(k1<rk_encode_base && k2<rk_encode_base && k3<rk_encode_base);
return ((k1+1)*rk_encode_base + k2)*rk_encode_base + k3;
}
protected:
friend class instruction_block;
template<typename T>
bool find_fn(const relation_base & r, T* & result) const
{ return m_fn_cache.find(encode_kind(r.get_kind()), reinterpret_cast<base_relation_fn*&>(result)); }
template<typename T>
bool find_fn(const relation_base & r1, const relation_base & r2, T*& result) const
{ return m_fn_cache.find(encode_kinds(r1.get_kind(), r2.get_kind()), reinterpret_cast<base_relation_fn*&>(result)); }
template<typename T>
bool find_fn(const relation_base & r1, const relation_base & r2, const relation_base & r3, T * & result) const
{ return m_fn_cache.find(encode_kinds(r1.get_kind(), r2.get_kind(), r3.get_kind()), reinterpret_cast<base_relation_fn*&>(result)); }
void store_fn(const relation_base & r, base_relation_fn * fn)
{ m_fn_cache.insert(encode_kind(r.get_kind()), fn); }
void store_fn(const relation_base & r1, const relation_base & r2, base_relation_fn * fn)
{ m_fn_cache.insert(encode_kinds(r1.get_kind(), r2.get_kind()), fn); }
void store_fn(const relation_base & r1, const relation_base & r2, const relation_base & r3,
base_relation_fn * fn)
{ m_fn_cache.insert(encode_kinds(r1.get_kind(), r2.get_kind(), r3.get_kind()), fn); }
/**
Process not only costs associated with the current instruction, but in case of
block instructions, process also costs associated with its child instructions.
*/
virtual void process_all_costs();
/**
\brief Output one line header of the current instruction.
The newline character at the end should not be printed.
*/
virtual void display_head_impl(rel_context const & ctx, std::ostream & out) const {
out << "<instruction>";
}
/**
\brief If relevant, output the body of the current instruction.
Each line must be prepended by \c indentation and ended by a newline character.
*/
virtual void display_body_impl(rel_context const & ctx, std::ostream & out, std::string indentation) const {}
public:
typedef execution_context::reg_type reg_type;
typedef execution_context::reg_idx reg_idx;
virtual ~instruction();
virtual bool perform(execution_context & ctx) = 0;
virtual void make_annotations(execution_context & ctx) = 0;
void display(rel_context const& ctx, std::ostream & out) const {
display_indented(ctx, out, "");
}
void display_indented(rel_context const & ctx, std::ostream & out, std::string indentation) const;
static instruction * mk_load(ast_manager & m, func_decl * pred, reg_idx tgt);
/**
\brief The store operation moves the relation from a register into the context. The register
is set to zero after the operation.
*/
static instruction * mk_store(ast_manager & m, func_decl * pred, reg_idx src);
static instruction * mk_dealloc(reg_idx reg); //maybe not necessary
static instruction * mk_clone(reg_idx from, reg_idx to);
static instruction * mk_move(reg_idx from, reg_idx to);
/**
\brief Return instruction that performs \c body as long as at least one register
in \c control_regs contains non-empty relation.
The instruction object takes over the ownership of the \c body object.
*/
static instruction * mk_while_loop(unsigned control_reg_cnt, const reg_idx * control_regs,
instruction_block * body);
static instruction * mk_join(reg_idx rel1, reg_idx rel2, unsigned col_cnt,
const unsigned * cols1, const unsigned * cols2, reg_idx result);
static instruction * mk_filter_equal(ast_manager & m, reg_idx reg, const relation_element & value, unsigned col);
static instruction * mk_filter_identical(reg_idx reg, unsigned col_cnt, const unsigned * identical_cols);
static instruction * mk_filter_interpreted(reg_idx reg, app_ref & condition);
static instruction * mk_filter_interpreted_and_project(reg_idx src, app_ref & condition,
unsigned col_cnt, const unsigned * removed_cols, reg_idx result);
static instruction * mk_union(reg_idx src, reg_idx tgt, reg_idx delta);
static instruction * mk_widen(reg_idx src, reg_idx tgt, reg_idx delta);
static instruction * mk_projection(reg_idx src, unsigned col_cnt, const unsigned * removed_cols,
reg_idx tgt);
static instruction * mk_join_project(reg_idx rel1, reg_idx rel2, unsigned joined_col_cnt,
const unsigned * cols1, const unsigned * cols2, unsigned removed_col_cnt,
const unsigned * removed_cols, reg_idx result);
static instruction * mk_rename(reg_idx src, unsigned cycle_len, const unsigned * permutation_cycle,
reg_idx tgt);
static instruction * mk_filter_by_negation(reg_idx tgt, reg_idx neg_rel, unsigned col_cnt,
const unsigned * cols1, const unsigned * cols2);
static instruction * mk_select_equal_and_project(ast_manager & m, reg_idx src,
const relation_element & value, unsigned col, reg_idx result);
static instruction * mk_unary_singleton(ast_manager & m, func_decl* pred, const relation_sort & s, const relation_element & val, reg_idx tgt);
static instruction * mk_total(const relation_signature & sig, func_decl* pred, reg_idx tgt);
/**
\brief The mark_saturated instruction marks a relation as saturated, so that after
next restart it does not have to be part of the saturation again.
*/
static instruction * mk_mark_saturated(ast_manager & m, func_decl * pred);
static instruction * mk_assert_signature(const relation_signature & s, reg_idx tgt);
};
// -----------------------------------
//
// instruction_block
//
// -----------------------------------
class instruction_block {
public:
struct instruction_observer {
virtual ~instruction_observer() {}
virtual void notify(instruction * i) {}
};
private:
typedef ptr_vector<instruction> instr_seq_type;
instr_seq_type m_data;
instruction_observer* m_observer;
public:
instruction_block() : m_observer(0) {}
~instruction_block();
void reset();
void push_back(instruction * i) {
m_data.push_back(i);
if(m_observer) {
m_observer->notify(i);
}
}
void set_observer(instruction_observer * o) {
SASSERT(o==0 || m_observer==0);
m_observer = o;
}
/**
\brief Perform instructions in the block. If the run was interrupted before completion,
return false; otherwise return true.
The execution can terminate before completion if the function
\c execution_context::should_terminate() returns true.
*/
bool perform(execution_context & ctx) const;
void process_all_costs();
void make_annotations(execution_context & ctx);
void display(rel_context const & ctx, std::ostream & out) const {
display_indented(ctx, out, "");
}
void display_indented(rel_context const & ctx, std::ostream & out, std::string indentation) const;
};
};
#endif /* _DL_INSTRUCTION_H_ */

View file

@ -0,0 +1,652 @@
/*++
Copyright (c) 2010 Microsoft Corporation
Module Name:
dl_interval_relation.cpp
Abstract:
Basic interval reatlion.
Author:
Nikolaj Bjorner (nbjorner) 2010-2-11
Revision History:
--*/
#include "debug.h"
#include "optional.h"
#include "ast_pp.h"
#include "dl_interval_relation.h"
#include "bool_rewriter.h"
namespace datalog {
// -------------------------
// interval_relation_plugin
interval_relation_plugin::interval_relation_plugin(relation_manager& m):
relation_plugin(interval_relation_plugin::get_name(), m),
m_empty(m_dep),
m_arith(get_ast_manager()) {
}
bool interval_relation_plugin::can_handle_signature(const relation_signature & sig) {
for (unsigned i = 0; i < sig.size(); ++i) {
if (!m_arith.is_int(sig[i]) && !m_arith.is_real(sig[i])) {
return false;
}
}
return true;
}
relation_base * interval_relation_plugin::mk_empty(const relation_signature & s) {
return alloc(interval_relation, *this, s, true);
}
relation_base * interval_relation_plugin::mk_full(func_decl* p, const relation_signature & s) {
return alloc(interval_relation, *this, s, false);
}
class interval_relation_plugin::join_fn : public convenient_relation_join_fn {
public:
join_fn(const relation_signature & o1_sig, const relation_signature & o2_sig, unsigned col_cnt,
const unsigned * cols1, const unsigned * cols2)
: convenient_relation_join_fn(o1_sig, o2_sig, col_cnt, cols1, cols2){
}
virtual relation_base * operator()(const relation_base & _r1, const relation_base & _r2) {
interval_relation const& r1 = get(_r1);
interval_relation const& r2 = get(_r2);
interval_relation_plugin& p = r1.get_plugin();
interval_relation* result = dynamic_cast<interval_relation*>(p.mk_full(0, get_result_signature()));
result->mk_join(r1, r2, m_cols1.size(), m_cols1.c_ptr(), m_cols2.c_ptr());
return result;
}
};
relation_join_fn * interval_relation_plugin::mk_join_fn(const relation_base & r1, const relation_base & r2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2) {
if (!check_kind(r1) || !check_kind(r2)) {
return 0;
}
return alloc(join_fn, r1.get_signature(), r2.get_signature(), col_cnt, cols1, cols2);
}
class interval_relation_plugin::project_fn : public convenient_relation_project_fn {
public:
project_fn(const relation_signature & orig_sig, unsigned removed_col_cnt, const unsigned * removed_cols)
: convenient_relation_project_fn(orig_sig, removed_col_cnt, removed_cols) {
}
virtual relation_base * operator()(const relation_base & _r) {
interval_relation const& r = get(_r);
interval_relation_plugin& p = r.get_plugin();
interval_relation* result = dynamic_cast<interval_relation*>(p.mk_full(0, get_result_signature()));
result->mk_project(r, m_removed_cols.size(), m_removed_cols.c_ptr());
return result;
}
};
relation_transformer_fn * interval_relation_plugin::mk_project_fn(const relation_base & r,
unsigned col_cnt, const unsigned * removed_cols) {
return alloc(project_fn, r.get_signature(), col_cnt, removed_cols);
}
class interval_relation_plugin::rename_fn : public convenient_relation_rename_fn {
public:
rename_fn(const relation_signature & orig_sig, unsigned cycle_len, const unsigned * cycle)
: convenient_relation_rename_fn(orig_sig, cycle_len, cycle) {
}
virtual relation_base * operator()(const relation_base & _r) {
interval_relation const& r = get(_r);
interval_relation_plugin& p = r.get_plugin();
interval_relation* result = dynamic_cast<interval_relation*>(p.mk_full(0, get_result_signature()));
result->mk_rename(r, m_cycle.size(), m_cycle.c_ptr());
return result;
}
};
relation_transformer_fn * interval_relation_plugin::mk_rename_fn(const relation_base & r,
unsigned cycle_len, const unsigned * permutation_cycle) {
if(!check_kind(r)) {
return 0;
}
return alloc(rename_fn, r.get_signature(), cycle_len, permutation_cycle);
}
interval interval_relation_plugin::unite(interval const& src1, interval const& src2) {
bool l_open = src1.is_lower_open();
bool r_open = src1.is_upper_open();
ext_numeral low = src1.inf();
ext_numeral high = src1.sup();
if (src2.inf() < low || (src2.inf() == low && l_open)) {
low = src2.inf();
l_open = src2.is_lower_open();
}
if (src2.sup() > high || (src2.sup() == high && r_open)) {
high = src2.sup();
r_open = src2.is_upper_open();
}
return interval(dep(), low, l_open, 0, high, r_open, 0);
}
interval interval_relation_plugin::widen(interval const& src1, interval const& src2) {
bool l_open = src1.is_lower_open();
bool r_open = src1.is_upper_open();
ext_numeral low = src1.inf();
ext_numeral high = src1.sup();
if (src2.inf() < low || (low == src2.inf() && l_open && !src2.is_lower_open())) {
low = ext_numeral(false);
l_open = true;
}
if (high < src2.sup() || (src2.sup() == high && !r_open && src2.is_upper_open())) {
high = ext_numeral(true);
r_open = true;
}
return interval(dep(), low, l_open, 0, high, r_open, 0);
}
interval interval_relation_plugin::meet(interval const& src1, interval const& src2, bool& isempty) {
isempty = false;
if (is_empty(0, src1) || is_infinite(src2)) {
return src1;
}
if (is_empty(0, src2) || is_infinite(src1)) {
return src2;
}
bool l_open = src1.is_lower_open();
bool r_open = src1.is_upper_open();
ext_numeral low = src1.inf();
ext_numeral high = src1.sup();
if (src2.inf() > low || (src2.inf() == low && !l_open)) {
low = src2.inf();
l_open = src2.is_lower_open();
}
if (src2.sup() < high || (src2.sup() == high && !r_open)) {
high = src2.sup();
r_open = src2.is_upper_open();
}
if (low > high || (low == high && (l_open || r_open))) {
isempty = true;
return interval(dep());
}
else {
return interval(dep(), low, l_open, 0, high, r_open, 0);
}
}
bool interval_relation_plugin::is_infinite(interval const& i) {
return i.plus_infinity() && i.minus_infinity();
}
bool interval_relation_plugin::is_empty(unsigned, interval const& i) {
return i.sup() < i.inf();
}
class interval_relation_plugin::union_fn : public relation_union_fn {
bool m_is_widen;
public:
union_fn(bool is_widen) :
m_is_widen(is_widen) {
}
virtual void operator()(relation_base & _r, const relation_base & _src, relation_base * _delta) {
TRACE("interval_relation", _r.display(tout << "dst:\n"); _src.display(tout << "src:\n"););
interval_relation& r = get(_r);
interval_relation const& src = get(_src);
if (_delta) {
interval_relation& d = get(*_delta);
r.mk_union(src, &d, m_is_widen);
}
else {
r.mk_union(src, 0, m_is_widen);
}
}
};
relation_union_fn * interval_relation_plugin::mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta) {
if (!check_kind(tgt) || !check_kind(src) || (delta && !check_kind(*delta))) {
return 0;
}
return alloc(union_fn, false);
}
relation_union_fn * interval_relation_plugin::mk_widen_fn(
const relation_base & tgt, const relation_base & src,
const relation_base * delta) {
if (!check_kind(tgt) || !check_kind(src) || (delta && !check_kind(*delta))) {
return 0;
}
return alloc(union_fn, true);
}
class interval_relation_plugin::filter_identical_fn : public relation_mutator_fn {
unsigned_vector m_identical_cols;
public:
filter_identical_fn(unsigned col_cnt, const unsigned * identical_cols)
: m_identical_cols(col_cnt, identical_cols) {}
virtual void operator()(relation_base & r) {
interval_relation & pr = get(r);
for (unsigned i = 1; i < m_identical_cols.size(); ++i) {
unsigned c1 = m_identical_cols[0];
unsigned c2 = m_identical_cols[i];
pr.equate(c1, c2);
}
}
};
relation_mutator_fn * interval_relation_plugin::mk_filter_identical_fn(
const relation_base & t, unsigned col_cnt, const unsigned * identical_cols) {
if(!check_kind(t)) {
return 0;
}
return alloc(filter_identical_fn, col_cnt, identical_cols);
}
class interval_relation_plugin::filter_equal_fn : public relation_mutator_fn {
unsigned m_col;
rational m_value;
public:
filter_equal_fn(relation_manager & m, const relation_element & value, unsigned col)
: m_col(col) {
arith_util arith(m.get_context().get_manager());
VERIFY(arith.is_numeral(value, m_value));
}
virtual void operator()(relation_base & _r) {
interval_relation & r = get(_r);
interval_relation_plugin & p = r.get_plugin();
r.mk_intersect(m_col, interval(p.dep(), m_value));
TRACE("interval_relation", tout << m_value << "\n"; r.display(tout););
}
};
relation_mutator_fn * interval_relation_plugin::mk_filter_equal_fn(const relation_base & r,
const relation_element & value, unsigned col) {
if(check_kind(r)) {
return alloc(filter_equal_fn, get_manager(), value, col);
}
return 0;
}
class interval_relation_plugin::filter_interpreted_fn : public relation_mutator_fn {
app_ref m_cond;
public:
filter_interpreted_fn(interval_relation const& t, app* cond):
m_cond(cond, t.get_plugin().get_ast_manager()) {
}
void operator()(relation_base& t) {
get(t).filter_interpreted(m_cond);
TRACE("interval_relation", tout << mk_pp(m_cond, m_cond.get_manager()) << "\n"; t.display(tout););
}
};
relation_mutator_fn * interval_relation_plugin::mk_filter_interpreted_fn(const relation_base & t, app * condition) {
if (check_kind(t)) {
return alloc(filter_interpreted_fn, get(t), condition);
}
return 0;
}
interval_relation& interval_relation_plugin::get(relation_base& r) {
return dynamic_cast<interval_relation&>(r);
}
interval_relation const & interval_relation_plugin::get(relation_base const& r) {
return dynamic_cast<interval_relation const&>(r);
}
// -----------------------
// interval_relation
interval_relation::interval_relation(interval_relation_plugin& p, relation_signature const& s, bool is_empty):
vector_relation<interval>(p, s, is_empty, interval(p.dep()))
{
TRACE("interval_relation", tout << s.size() << "\n";);
}
void interval_relation::add_fact(const relation_fact & f) {
interval_relation r(get_plugin(), get_signature(), false);
ast_manager& m = get_plugin().get_ast_manager();
for (unsigned i = 0; i < f.size(); ++i) {
app_ref eq(m);
expr* e = f[i];
eq = m.mk_eq(m.mk_var(i, m.get_sort(e)), e);
r.filter_interpreted(eq.get());
}
mk_union(r, 0, false);
}
bool interval_relation::contains_fact(const relation_fact & f) const {
SASSERT(f.size() == get_signature().size());
interval_relation_plugin& p = get_plugin();
for (unsigned i = 0; i < f.size(); ++i) {
if (f[i] != f[find(i)]) {
return false;
}
interval const& iv = (*this)[i];
if (p.is_infinite(iv)) {
continue;
}
rational v;
if (p.m_arith.is_numeral(f[i], v)) {
if (!iv.contains(v)) {
return false;
}
}
else {
// TBD: may or must?
}
}
return true;
}
interval_relation * interval_relation::clone() const {
interval_relation* result = alloc(interval_relation, get_plugin(), get_signature(), empty());
result->copy(*this);
return result;
}
interval_relation * interval_relation::complement(func_decl*) const {
UNREACHABLE();
return 0;
}
void interval_relation::to_formula(expr_ref& fml) const {
ast_manager& m = get_plugin().get_ast_manager();
arith_util& arith = get_plugin().m_arith;
expr_ref_vector conjs(m);
relation_signature const& sig = get_signature();
for (unsigned i = 0; i < sig.size(); ++i) {
if (i != find(i)) {
conjs.push_back(m.mk_eq(m.mk_var(i, sig[i]),
m.mk_var(find(i), sig[find(i)])));
continue;
}
interval const& iv = (*this)[i];
sort* ty = sig[i];
expr_ref var(m.mk_var(i, ty), m);
if (!iv.minus_infinity()) {
expr* lo = arith.mk_numeral(iv.get_lower_value(), ty);
if (iv.is_lower_open()) {
conjs.push_back(arith.mk_lt(lo, var));
}
else {
conjs.push_back(arith.mk_le(lo, var));
}
}
if (!iv.plus_infinity()) {
expr* hi = arith.mk_numeral(iv.get_upper_value(), ty);
if (iv.is_upper_open()) {
conjs.push_back(arith.mk_lt(var, hi));
}
else {
conjs.push_back(arith.mk_le(var, hi));
}
}
}
bool_rewriter br(m);
br.mk_and(conjs.size(), conjs.c_ptr(), fml);
}
void interval_relation::display_index(unsigned i, interval const& j, std::ostream & out) const {
out << i << " in " << j << "\n";
}
interval_relation_plugin& interval_relation::get_plugin() const {
return static_cast<interval_relation_plugin &>(relation_base::get_plugin());
}
void interval_relation::mk_intersect(unsigned idx, interval const& i) {
bool isempty;
(*this)[idx] = mk_intersect((*this)[idx], i, isempty);
if (isempty || is_empty(idx, (*this)[idx])) {
set_empty();
}
}
void interval_relation::mk_rename_elem(interval& i, unsigned, unsigned const* ) {
}
void interval_relation::filter_interpreted(app* cond) {
interval_relation_plugin& p = get_plugin();
rational k;
unsigned x, y;
if (p.is_lt(cond, x, k, y)) {
// 0 < x - y + k
if (x == UINT_MAX) {
// y < k
mk_intersect(y, interval(p.dep(), k, true, false, 0));
return;
}
if (y == UINT_MAX) {
// -k < x
mk_intersect(x, interval(p.dep(), -k, true, true, 0));
return;
}
// y < x + k
ext_numeral x_hi = (*this)[x].sup();
ext_numeral y_lo = (*this)[y].inf();
if (!x_hi.is_infinite()) {
mk_intersect(y, interval(p.dep(), k + x_hi.to_rational(), true, false, 0));
}
if (!y_lo.is_infinite()) {
mk_intersect(x, interval(p.dep(), y_lo.to_rational() - k, true, true, 0));
}
return;
}
bool is_int = false;
if (p.is_le(cond, x, k, y, is_int)) {
// 0 <= x - y + k
if (x == UINT_MAX) {
// y <= k
mk_intersect(y, interval(p.dep(), k, false, false, 0));
return;
}
if (y == UINT_MAX) {
// -k <= x
mk_intersect(x, interval(p.dep(), -k, false, true, 0));
return;
}
ext_numeral x_hi = (*this)[x].sup();
ext_numeral y_lo = (*this)[y].inf();
if (!x_hi.is_infinite()) {
mk_intersect(y, interval(p.dep(), k + x_hi.to_rational(), false, false, 0));
}
if (!y_lo.is_infinite()) {
mk_intersect(x, interval(p.dep(), y_lo.to_rational() - k, false, true, 0));
}
return;
}
if (p.is_eq(cond, x, k, y)) {
// y = x + k
if (x == UINT_MAX) {
SASSERT(y != UINT_MAX);
mk_intersect(y, interval(p.dep(), k));
return;
}
if (y == UINT_MAX) {
// x = - k
SASSERT(x != UINT_MAX);
mk_intersect(x, interval(p.dep(), -k));
return;
}
interval x_i = (*this)[x];
interval y_i = (*this)[y];
x_i += interval(p.dep(), k);
y_i -= interval(p.dep(), k);
mk_intersect(x, y_i);
mk_intersect(y, x_i);
}
if (get_plugin().get_ast_manager().is_false(cond)) {
set_empty();
}
}
bool interval_relation_plugin::is_linear(expr* e, unsigned& neg, unsigned& pos, rational& k, bool is_pos) const {
#define SET_VAR(_idx_) \
if (is_pos &&pos == UINT_MAX) { \
pos = _idx_; \
return true; \
} \
if (!is_pos && neg == UINT_MAX) { \
neg = _idx_; \
return true; \
} \
else { \
return false; \
}
if (is_var(e)) {
SET_VAR(to_var(e)->get_idx());
}
if (!is_app(e)) {
return false;
}
app* a = to_app(e);
if (m_arith.is_add(e)) {
for (unsigned i = 0; i < a->get_num_args(); ++i) {
if (!is_linear(a->get_arg(i), neg, pos, k, is_pos)) return false;
}
return true;
}
if (m_arith.is_sub(e)) {
SASSERT(a->get_num_args() == 2);
return
is_linear(a->get_arg(0), neg, pos, k, is_pos) &&
is_linear(a->get_arg(1), neg, pos, k, !is_pos);
}
rational k1;
SASSERT(!m_arith.is_mul(e) || a->get_num_args() == 2);
if (m_arith.is_mul(e) &&
m_arith.is_numeral(a->get_arg(0), k1) &&
k1.is_minus_one() &&
is_var(a->get_arg(1))) {
SET_VAR(to_var(a->get_arg(1))->get_idx());
}
if (m_arith.is_numeral(e, k1)) {
if (is_pos) {
k += k1;
}
else {
k -= k1;
}
return true;
}
return false;
}
// 0 <= x - y + k
bool interval_relation_plugin::is_le(app* cond, unsigned& x, rational& k, unsigned& y, bool& is_int) const {
ast_manager& m = get_ast_manager();
k.reset();
x = UINT_MAX;
y = UINT_MAX;
if (m_arith.is_le(cond)) {
is_int = m_arith.is_int(cond->get_arg(0));
if (!is_linear(cond->get_arg(0), y, x, k, false)) return false;
if (!is_linear(cond->get_arg(1), y, x, k, true)) return false;
return (x != UINT_MAX || y != UINT_MAX);
}
if (m_arith.is_ge(cond)) {
is_int = m_arith.is_int(cond->get_arg(0));
if (!is_linear(cond->get_arg(0), y, x, k, true)) return false;
if (!is_linear(cond->get_arg(1), y, x, k, false)) return false;
return (x != UINT_MAX || y != UINT_MAX);
}
if (m_arith.is_lt(cond) && m_arith.is_int(cond->get_arg(0))) {
is_int = true;
if (!is_linear(cond->get_arg(0), y, x, k, false)) return false;
if (!is_linear(cond->get_arg(1), y, x, k, true)) return false;
k -= rational::one();
return (x != UINT_MAX || y != UINT_MAX);
}
if (m_arith.is_gt(cond) && m_arith.is_int(cond->get_arg(0))) {
is_int = true;
if (!is_linear(cond->get_arg(0), y, x, k, true)) return false;
if (!is_linear(cond->get_arg(1), y, x, k, false)) return false;
k += rational::one();
return (x != UINT_MAX || y != UINT_MAX);
}
if (m.is_not(cond) && is_app(cond->get_arg(0))) {
// not (0 <= x - y + k)
// <=>
// 0 > x - y + k
// <=>
// 0 <= y - x - k - 1
if (is_le(to_app(cond->get_arg(0)), x, k, y, is_int) && is_int) {
k.neg();
k -= rational::one();
std::swap(x, y);
return true;
}
// not (0 < x - y + k)
// <=>
// 0 >= x - y + k
// <=>
// 0 <= y - x - k
if (is_lt(to_app(cond->get_arg(0)), x, k, y)) {
is_int = false;
k.neg();
std::swap(x, y);
return true;
}
}
return false;
}
// 0 < x - y + k
bool interval_relation_plugin::is_lt(app* cond, unsigned& x, rational& k, unsigned& y) const {
k.reset();
x = UINT_MAX;
y = UINT_MAX;
if (m_arith.is_lt(cond) && m_arith.is_real(cond->get_arg(0))) {
if (!is_linear(cond->get_arg(0), y, x, k, false)) return false;
if (!is_linear(cond->get_arg(1), y, x, k, true)) return false;
return (x != UINT_MAX || y != UINT_MAX);
}
if (m_arith.is_gt(cond) && m_arith.is_real(cond->get_arg(0))) {
if (!is_linear(cond->get_arg(0), y, x, k, true)) return false;
if (!is_linear(cond->get_arg(1), y, x, k, false)) return false;
return (x != UINT_MAX || y != UINT_MAX);
}
return false;
}
// 0 = x - y + k
bool interval_relation_plugin::is_eq(app* cond, unsigned& x, rational& k, unsigned& y) const {
ast_manager& m = get_ast_manager();
k.reset();
x = UINT_MAX;
y = UINT_MAX;
if (m.is_eq(cond)) {
if (!is_linear(cond->get_arg(0), y, x, k, false)) return false;
if (!is_linear(cond->get_arg(1), y, x, k, true)) return false;
return (x != UINT_MAX || y != UINT_MAX);
}
return false;
}
};

View file

@ -0,0 +1,140 @@
/*++
Copyright (c) 2010 Microsoft Corporation
Module Name:
dl_interval_relation.h
Abstract:
Basic interval reatlion.
Author:
Nikolaj Bjorner (nbjorner) 2010-2-11
Revision History:
--*/
#ifndef _DL_INTERVAL_RELATION_H_
#define _DL_INTERVAL_RELATION_H_
#include "dl_context.h"
#include "old_interval.h"
#include "dl_vector_relation.h"
#include "arith_decl_plugin.h"
#include "basic_simplifier_plugin.h"
namespace datalog {
class interval_relation;
class interval_relation_plugin : public relation_plugin {
v_dependency_manager m_dep;
interval m_empty;
arith_util m_arith;
class join_fn;
class project_fn;
class rename_fn;
class union_fn;
class filter_equal_fn;
class filter_identical_fn;
class filter_interpreted_fn;
friend class interval_relation;
interval unite(interval const& src1, interval const& src2);
interval widen(interval const& src1, interval const& src2);
interval meet(interval const& src1, interval const& src2, bool& is_empty);
v_dependency_manager & dep() const { return const_cast<v_dependency_manager&>(m_dep); }
public:
interval_relation_plugin(relation_manager& m);
virtual bool can_handle_signature(const relation_signature & s);
static symbol get_name() { return symbol("interval_relation"); }
virtual relation_base * mk_empty(const relation_signature & s);
virtual relation_base * mk_full(func_decl* p, const relation_signature & s);
virtual relation_join_fn * mk_join_fn(const relation_base & t1, const relation_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
virtual relation_transformer_fn * mk_project_fn(const relation_base & t, unsigned col_cnt,
const unsigned * removed_cols);
virtual relation_transformer_fn * mk_rename_fn(const relation_base & t, unsigned permutation_cycle_len,
const unsigned * permutation_cycle);
virtual relation_union_fn * mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
virtual relation_union_fn * mk_widen_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
virtual relation_mutator_fn * mk_filter_identical_fn(const relation_base & t, unsigned col_cnt,
const unsigned * identical_cols);
virtual relation_mutator_fn * mk_filter_equal_fn(const relation_base & t, const relation_element & value,
unsigned col);
virtual relation_mutator_fn * mk_filter_interpreted_fn(const relation_base & t, app * condition);
static bool is_empty(unsigned idx, interval const& i);
static bool is_infinite(interval const& i);
private:
static interval_relation& get(relation_base& r);
static interval_relation const & get(relation_base const& r);
bool is_linear(expr* e, unsigned& pos, unsigned& neg, rational& k, bool is_pos) const;
// x + k <= y
bool is_le(app* cond, unsigned& x, rational& k, unsigned& y, bool& is_int) const;
// x + k < y
bool is_lt(app* cond, unsigned& x, rational& k, unsigned& y) const;
// x + k = y
bool is_eq(app* cond, unsigned& x, rational& k, unsigned& y) const;
};
class interval_relation : public vector_relation<interval> {
friend class interval_relation_plugin;
friend class interval_relation_plugin::filter_equal_fn;
public:
interval_relation(interval_relation_plugin& p, relation_signature const& s, bool is_empty);
virtual void add_fact(const relation_fact & f);
virtual bool contains_fact(const relation_fact & f) const;
virtual interval_relation * clone() const;
virtual interval_relation * complement(func_decl*) const;
virtual void to_formula(expr_ref& fml) const;
interval_relation_plugin& get_plugin() const;
void filter_interpreted(app* cond);
virtual bool is_precise() const { return false; }
private:
virtual interval mk_intersect(interval const& t1, interval const& t2, bool& is_empty) const {
return get_plugin().meet(t1, t2, is_empty);
}
virtual interval mk_unite(interval const& t1, interval const& t2) const { return get_plugin().unite(t1,t2); }
virtual interval mk_widen(interval const& t1, interval const& t2) const { return get_plugin().widen(t1,t2); }
virtual bool is_subset_of(interval const& t1, interval const& t2) const { NOT_IMPLEMENTED_YET(); return false; }
virtual bool is_full(interval const& t) const {
return interval_relation_plugin::is_infinite(t);
}
virtual bool is_empty(unsigned idx, interval const& t) const {
return interval_relation_plugin::is_empty(idx, t);
}
virtual void mk_rename_elem(interval& i, unsigned col_cnt, unsigned const* cycle);
virtual void display_index(unsigned idx, interval const & i, std::ostream& out) const;
void mk_intersect(unsigned idx, interval const& i);
};
};
#endif

View file

@ -0,0 +1,280 @@
/*++
Copyright (c) 2012 Microsoft Corporation
Module Name:
dl_mk_array_blast.cpp
Abstract:
Remove array stores from rules.
Author:
Nikolaj Bjorner (nbjorner) 2012-11-23
Revision History:
--*/
#include "dl_mk_array_blast.h"
#include "qe_util.h"
namespace datalog {
mk_array_blast::mk_array_blast(context & ctx, unsigned priority) :
rule_transformer::plugin(priority, false),
m_ctx(ctx),
m(ctx.get_manager()),
a(m),
rm(ctx.get_rule_manager()),
m_rewriter(m, m_params),
m_simplifier(ctx),
m_sub(m),
m_next_var(0) {
m_params.set_bool("expand_select_store",true);
m_rewriter.updt_params(m_params);
}
mk_array_blast::~mk_array_blast() {
}
bool mk_array_blast::is_store_def(expr* e, expr*& x, expr*& y) {
if (m.is_iff(e, x, y) || m.is_eq(e, x, y)) {
if (!a.is_store(y)) {
std::swap(x,y);
}
if (is_var(x) && a.is_store(y)) {
return true;
}
}
return false;
}
expr* mk_array_blast::get_select(expr* e) const {
while (a.is_select(e)) {
e = to_app(e)->get_arg(0);
}
return e;
}
void mk_array_blast::get_select_args(expr* e, ptr_vector<expr>& args) const {
while (a.is_select(e)) {
app* ap = to_app(e);
for (unsigned i = 1; i < ap->get_num_args(); ++i) {
args.push_back(ap->get_arg(i));
}
e = ap->get_arg(0);
}
}
bool mk_array_blast::insert_def(rule const& r, app* e, var* v) {
//
// For the Ackermann reduction we would like the arrays
// to be variables, so that variables can be
// assumed to represent difference (alias)
// classes. Ehm., Soundness of this approach depends on
// if the arrays are finite domains...
//
if (!is_var(get_select(e))) {
return false;
}
if (v) {
m_sub.insert(e, v);
m_defs.insert(e, to_var(v));
}
else {
if (m_next_var == 0) {
ptr_vector<sort> vars;
r.get_vars(vars);
m_next_var = vars.size() + 1;
}
v = m.mk_var(m_next_var, m.get_sort(e));
m_sub.insert(e, v);
m_defs.insert(e, v);
++m_next_var;
}
return true;
}
bool mk_array_blast::ackermanize(rule const& r, expr_ref& body, expr_ref& head) {
expr_ref_vector conjs(m);
qe::flatten_and(body, conjs);
m_defs.reset();
m_sub.reset();
m_next_var = 0;
ptr_vector<expr> todo;
todo.push_back(head);
for (unsigned i = 0; i < conjs.size(); ++i) {
expr* e = conjs[i].get();
expr* x, *y;
if (m.is_eq(e, x, y) || m.is_iff(e, x, y)) {
if (a.is_select(y)) {
std::swap(x,y);
}
if (a.is_select(x) && is_var(y)) {
if (!insert_def(r, to_app(x), to_var(y))) {
return false;
}
}
}
if (a.is_select(e) && !insert_def(r, to_app(e), 0)) {
return false;
}
todo.push_back(e);
}
// now make sure to cover all occurrences.
ast_mark mark;
while (!todo.empty()) {
expr* e = todo.back();
todo.pop_back();
if (mark.is_marked(e)) {
continue;
}
mark.mark(e, true);
if (is_var(e)) {
continue;
}
if (!is_app(e)) {
return false;
}
app* ap = to_app(e);
if (a.is_select(ap) && !m_defs.contains(ap)) {
if (!insert_def(r, ap, 0)) {
return false;
}
}
if (a.is_select(e)) {
get_select_args(e, todo);
continue;
}
for (unsigned i = 0; i < ap->get_num_args(); ++i) {
todo.push_back(ap->get_arg(i));
}
}
m_sub(body);
m_sub(head);
conjs.reset();
// perform the Ackermann reduction by creating implications
// i1 = i2 => val1 = val2 for each equality pair:
// (= val1 (select a_i i1))
// (= val2 (select a_i i2))
defs_t::iterator it1 = m_defs.begin(), end = m_defs.end();
for (; it1 != end; ++it1) {
app* a1 = it1->m_key;
var* v1 = it1->m_value;
defs_t::iterator it2 = it1;
++it2;
for (; it2 != end; ++it2) {
app* a2 = it2->m_key;
var* v2 = it2->m_value;
if (get_select(a1) != get_select(a2)) {
continue;
}
expr_ref_vector eqs(m);
ptr_vector<expr> args1, args2;
get_select_args(a1, args1);
get_select_args(a2, args2);
for (unsigned j = 0; j < args1.size(); ++j) {
eqs.push_back(m.mk_eq(args1[j], args2[j]));
}
conjs.push_back(m.mk_implies(m.mk_and(eqs.size(), eqs.c_ptr()), m.mk_eq(v1, v2)));
}
}
if (!conjs.empty()) {
conjs.push_back(body);
body = m.mk_and(conjs.size(), conjs.c_ptr());
}
m_rewriter(body);
return true;
}
bool mk_array_blast::blast(rule& r, rule_set& rules) {
unsigned utsz = r.get_uninterpreted_tail_size();
unsigned tsz = r.get_tail_size();
expr_ref_vector conjs(m), new_conjs(m);
expr_ref tmp(m);
expr_safe_replace sub(m);
bool change = false;
bool inserted = false;
for (unsigned i = 0; i < utsz; ++i) {
new_conjs.push_back(r.get_tail(i));
}
for (unsigned i = utsz; i < tsz; ++i) {
conjs.push_back(r.get_tail(i));
}
qe::flatten_and(conjs);
for (unsigned i = 0; i < conjs.size(); ++i) {
expr* x, *y, *e = conjs[i].get();
if (is_store_def(e, x, y)) {
// enforce topological order consistency:
uint_set lhs = rm.collect_vars(x);
uint_set rhs_vars = rm.collect_vars(y);
lhs &= rhs_vars;
if (!lhs.empty()) {
TRACE("dl", tout << "unusable equality " << mk_pp(e, m) << "\n";);
new_conjs.push_back(e);
}
else {
sub.insert(x, y);
inserted = true;
}
}
else {
m_rewriter(e, tmp);
change = change || (tmp != e);
new_conjs.push_back(tmp);
}
}
expr_ref fml2(m), body(m), head(m);
body = m.mk_and(new_conjs.size(), new_conjs.c_ptr());
head = r.get_head();
sub(body);
m_rewriter(body);
sub(head);
m_rewriter(head);
change = ackermanize(r, body, head) || change;
if (!inserted && !change) {
rules.add_rule(&r);
return false;
}
fml2 = m.mk_implies(body, head);
proof_ref p(m);
rule_set new_rules(m_ctx);
rm.mk_rule(fml2, p, new_rules, r.name());
rule_ref new_rule(rm);
if (m_simplifier.transform_rule(new_rules.last(), new_rule)) {
rules.add_rule(new_rule.get());
rm.mk_rule_rewrite_proof(r, *new_rule.get());
TRACE("dl", new_rule->display(m_ctx, tout << "new rule\n"););
}
return true;
}
rule_set * mk_array_blast::operator()(rule_set const & source) {
rule_set* rules = alloc(rule_set, m_ctx);
rules->inherit_predicates(source);
rule_set::iterator it = source.begin(), end = source.end();
bool change = false;
for (; !m_ctx.canceled() && it != end; ++it) {
change = blast(**it, *rules) || change;
}
if (!change) {
dealloc(rules);
rules = 0;
}
return rules;
}
};

View file

@ -0,0 +1,77 @@
/*++
Copyright (c) 2012 Microsoft Corporation
Module Name:
dl_mk_array_blast.h
Abstract:
Remove array variables from rules.
Author:
Nikolaj Bjorner (nbjorner) 2012-11-23
Revision History:
--*/
#ifndef _DL_MK_ARRAY_BLAST_H_
#define _DL_MK_ARRAY_BLAST_H_
#include"dl_context.h"
#include"dl_rule_set.h"
#include"dl_rule_transformer.h"
#include"dl_mk_interp_tail_simplifier.h"
#include "equiv_proof_converter.h"
#include "array_decl_plugin.h"
#include "expr_safe_replace.h"
namespace datalog {
/**
\brief Blast occurrences of arrays in rules
*/
class mk_array_blast : public rule_transformer::plugin {
typedef obj_map<app, var*> defs_t;
context& m_ctx;
ast_manager& m;
array_util a;
rule_manager& rm;
params_ref m_params;
th_rewriter m_rewriter;
mk_interp_tail_simplifier m_simplifier;
defs_t m_defs;
expr_safe_replace m_sub;
unsigned m_next_var;
bool blast(rule& r, rule_set& new_rules);
bool is_store_def(expr* e, expr*& x, expr*& y);
bool ackermanize(rule const& r, expr_ref& body, expr_ref& head);
expr* get_select(expr* e) const;
void get_select_args(expr* e, ptr_vector<expr>& args) const;
bool insert_def(rule const& r, app* e, var* v);
public:
/**
\brief Create rule transformer that removes array stores and selects by ackermannization.
*/
mk_array_blast(context & ctx, unsigned priority);
virtual ~mk_array_blast();
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_ARRAY_BLAST_H_ */

View file

@ -0,0 +1,78 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_backwards.cpp
Abstract:
Create Horn clauses for backwards flow.
Author:
Nikolaj Bjorner (nbjorner) 2013-04-17
Revision History:
--*/
#include"dl_mk_backwards.h"
#include"dl_context.h"
namespace datalog {
mk_backwards::mk_backwards(context & ctx, unsigned priority):
plugin(priority),
m(ctx.get_manager()),
m_ctx(ctx) {
}
mk_backwards::~mk_backwards() { }
rule_set * mk_backwards::operator()(rule_set const & source) {
context& ctx = source.get_context();
rule_manager& rm = source.get_rule_manager();
rule_set * result = alloc(rule_set, ctx);
unsigned sz = source.get_num_rules();
rule_ref new_rule(rm);
app_ref_vector tail(m);
app_ref head(m);
svector<bool> neg;
app_ref query(m);
query = m.mk_fresh_const("Q", m.mk_bool_sort());
result->set_output_predicate(query->get_decl());
m_ctx.register_predicate(query->get_decl(), false);
for (unsigned i = 0; i < sz; ++i) {
tail.reset();
neg.reset();
rule & r = *source.get_rule(i);
unsigned utsz = r.get_uninterpreted_tail_size();
unsigned tsz = r.get_tail_size();
if (!source.is_output_predicate(r.get_decl())) {
tail.push_back(r.get_head());
neg.push_back(false);
}
for (unsigned j = utsz; j < tsz; ++j) {
tail.push_back(r.get_tail(j));
neg.push_back(false);
}
for (unsigned j = 0; j <= utsz; ++j) {
if (j == utsz && j > 0) {
break;
}
if (j == utsz) {
head = query;
}
else {
head = r.get_tail(j);
}
new_rule = rm.mk(head, tail.size(), tail.c_ptr(), neg.c_ptr(), r.name(), true);
result->add_rule(new_rule);
}
}
TRACE("dl", result->display(tout););
return result;
}
};

38
src/muz/dl_mk_backwards.h Normal file
View file

@ -0,0 +1,38 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_backwards.h
Abstract:
Create Horn clauses for backwards flow.
Author:
Nikolaj Bjorner (nbjorner) 2013-04-17
Revision History:
--*/
#ifndef _DL_MK_BACKWARDS_H_
#define _DL_MK_BACKWARDS_H_
#include"dl_rule_transformer.h"
namespace datalog {
class mk_backwards : public rule_transformer::plugin {
ast_manager& m;
context& m_ctx;
public:
mk_backwards(context & ctx, unsigned priority = 33000);
~mk_backwards();
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_BACKWARDS_H_ */

318
src/muz/dl_mk_bit_blast.cpp Normal file
View file

@ -0,0 +1,318 @@
/*++
Copyright (c) 2012 Microsoft Corporation
Module Name:
dl_mk_bit_blast.cpp
Abstract:
<abstract>
Author:
Nikolaj Bjorner (nbjorner) 2012-08-30
Revision History:
--*/
#include "dl_mk_bit_blast.h"
#include "bit_blaster_rewriter.h"
#include "rewriter_def.h"
#include "ast_pp.h"
#include "expr_safe_replace.h"
#include "filter_model_converter.h"
#include "dl_mk_interp_tail_simplifier.h"
namespace datalog {
//
// P(v) :- Q(extract[1:1]v ++ 0), R(1 ++ extract[0:0]v).
// ->
// P(bv(x,y)) :- Q(bv(x,0)), R(bv(1,y)) .
//
// Introduce P_bv:
// P_bv(x,y) :- Q_bv(x,0), R_bv(1,y)
// P(bv(x,y)) :- P_bv(x,y)
// Query
// this model converter should be composed with a filter converter
// that gets rid of the new functions.
class bit_blast_model_converter : public model_converter {
ast_manager& m;
bv_util m_bv;
func_decl_ref_vector m_old_funcs;
func_decl_ref_vector m_new_funcs;
public:
bit_blast_model_converter(ast_manager& m):
m(m),
m_bv(m),
m_old_funcs(m),
m_new_funcs(m) {}
void insert(func_decl* old_f, func_decl* new_f) {
m_old_funcs.push_back(old_f);
m_new_funcs.push_back(new_f);
}
virtual model_converter * translate(ast_translation & translator) {
return alloc(bit_blast_model_converter, m);
}
virtual void operator()(model_ref & model) {
for (unsigned i = 0; i < m_new_funcs.size(); ++i) {
func_decl* p = m_new_funcs[i].get();
func_decl* q = m_old_funcs[i].get();
func_interp* f = model->get_func_interp(p);
expr_ref body(m);
unsigned arity_p = p->get_arity();
unsigned arity_q = q->get_arity();
SASSERT(0 < arity_p);
model->register_decl(p, f);
func_interp* g = alloc(func_interp, m, arity_q);
if (f) {
body = f->get_interp();
SASSERT(!f->is_partial());
SASSERT(body);
}
else {
body = m.mk_false();
}
unsigned idx = 0;
expr_ref arg(m), proj(m);
expr_safe_replace sub(m);
for (unsigned j = 0; j < arity_q; ++j) {
sort* s = q->get_domain(j);
arg = m.mk_var(j, s);
if (m_bv.is_bv_sort(s)) {
expr* args[1] = { arg };
unsigned sz = m_bv.get_bv_size(s);
for (unsigned k = 0; k < sz; ++k) {
proj = m.mk_app(m_bv.get_family_id(), OP_BIT2BOOL, 1, args);
sub.insert(m.mk_var(idx++, m.mk_bool_sort()), proj);
}
}
else {
sub.insert(m.mk_var(idx++, s), arg);
}
}
sub(body);
g->set_else(body);
model->register_decl(q, g);
}
}
};
class expand_mkbv_cfg : public default_rewriter_cfg {
context& m_context;
ast_manager& m;
bv_util m_util;
expr_ref_vector m_args, m_f_vars, m_g_vars;
func_decl_ref_vector m_old_funcs;
func_decl_ref_vector m_new_funcs;
rule_set const* m_src;
rule_set* m_dst;
obj_map<func_decl,func_decl*> m_pred2blast;
public:
expand_mkbv_cfg(context& ctx):
m_context(ctx),
m(ctx.get_manager()),
m_util(m),
m_args(m),
m_f_vars(m),
m_g_vars(m),
m_old_funcs(m),
m_new_funcs(m),
m_src(0),
m_dst(0)
{}
~expand_mkbv_cfg() {}
void set_src(rule_set const* src) { m_src = src; }
void set_dst(rule_set* dst) { m_dst = dst; }
func_decl_ref_vector const& old_funcs() const { return m_old_funcs; }
func_decl_ref_vector const& new_funcs() const { return m_new_funcs; }
br_status reduce_app(func_decl * f, unsigned num, expr * const * args, expr_ref & result, proof_ref & result_pr) {
if (num == 0) {
if (m_src->is_output_predicate(f))
m_dst->set_output_predicate(f);
return BR_FAILED;
}
for (unsigned i = 0; i < num; ++i) {
if (!m_util.is_mkbv(args[i]))
return BR_FAILED;
}
//
// f(mk_bv(args),...)
//
m_args.reset();
m_g_vars.reset();
m_f_vars.reset();
expr_ref fml(m);
unsigned idx = 0;
for (unsigned j = 0; j < num; ++j) {
expr* arg = args[j];
if (m_util.is_mkbv(arg)) {
app* a = to_app(arg);
unsigned sz = a->get_num_args();
for (unsigned i = 0; i < sz; ++i) {
m_args.push_back(a->get_arg(i));
m_g_vars.push_back(m.mk_var(idx++,m.mk_bool_sort()));
}
m_f_vars.push_back(m_util.mk_bv(sz, m_g_vars.c_ptr()+m_g_vars.size()-sz));
}
else {
m_args.push_back(arg);
m_f_vars.push_back(m.mk_var(idx++, m.get_sort(arg)));
m_g_vars.push_back(m_f_vars.back());
}
}
func_decl* g = 0;
if (!m_pred2blast.find(f, g)) {
ptr_vector<sort> domain;
for (unsigned i = 0; i < m_args.size(); ++i) {
domain.push_back(m.get_sort(m_args[i].get()));
}
g = m_context.mk_fresh_head_predicate(f->get_name(), symbol("bv"), m_args.size(), domain.c_ptr(), f);
m_old_funcs.push_back(f);
m_new_funcs.push_back(g);
m_pred2blast.insert(f, g);
m_dst->inherit_predicate(*m_src, f, g);
}
result = m.mk_app(g, m_args.size(), m_args.c_ptr());
result_pr = 0;
return BR_DONE;
}
};
struct expand_mkbv : public rewriter_tpl<expand_mkbv_cfg> {
expand_mkbv_cfg m_cfg;
expand_mkbv(ast_manager& m, context& ctx):
rewriter_tpl<expand_mkbv_cfg>(m, m.proofs_enabled(), m_cfg),
m_cfg(ctx) {
}
};
class mk_bit_blast::impl {
context & m_context;
ast_manager & m;
params_ref m_params;
mk_interp_tail_simplifier m_simplifier;
bit_blaster_rewriter m_blaster;
expand_mkbv m_rewriter;
bool blast(rule *r, expr_ref& fml) {
proof_ref pr(m);
expr_ref fml1(m), fml2(m), fml3(m);
rule_ref r2(m_context.get_rule_manager());
// We need to simplify rule before bit-blasting.
if (!m_simplifier.transform_rule(r, r2)) {
r2 = r;
}
r2->to_formula(fml1);
m_blaster(fml1, fml2, pr);
m_rewriter(fml2, fml3);
TRACE("dl", tout << mk_pp(fml, m) << " -> " << mk_pp(fml2, m) << " -> " << mk_pp(fml3, m) << "\n";);
if (fml3 != fml) {
fml = fml3;
return true;
}
else {
return false;
}
}
public:
impl(context& ctx):
m_context(ctx),
m(ctx.get_manager()),
m_params(ctx.get_params().p),
m_simplifier(ctx),
m_blaster(ctx.get_manager(), m_params),
m_rewriter(ctx.get_manager(), ctx) {
m_params.set_bool("blast_full", true);
m_params.set_bool("blast_quant", true);
m_blaster.updt_params(m_params);
}
rule_set * operator()(rule_set const & source) {
// TODO pc
if (!m_context.get_params().bit_blast()) {
return 0;
}
rule_manager& rm = m_context.get_rule_manager();
unsigned sz = source.get_num_rules();
expr_ref fml(m);
rule_set * result = alloc(rule_set, m_context);
m_rewriter.m_cfg.set_src(&source);
m_rewriter.m_cfg.set_dst(result);
for (unsigned i = 0; !m_context.canceled() && i < sz; ++i) {
rule * r = source.get_rule(i);
r->to_formula(fml);
if (blast(r, fml)) {
proof_ref pr(m);
if (m_context.generate_proof_trace()) {
pr = m.mk_asserted(fml); // loses original proof of r.
}
// TODO add logic for pc:
// 1. replace fresh predicates by non-bit-blasted predicates
// 2. replace pr by the proof of r.
rm.mk_rule(fml, pr, *result, r->name());
}
else {
result->add_rule(r);
}
}
// copy output predicates without any rule (bit-blasting not really needed)
const func_decl_set& decls = source.get_output_predicates();
for (func_decl_set::iterator I = decls.begin(), E = decls.end(); I != E; ++I) {
if (!source.contains(*I))
result->set_output_predicate(*I);
}
if (m_context.get_model_converter()) {
filter_model_converter* fmc = alloc(filter_model_converter, m);
bit_blast_model_converter* bvmc = alloc(bit_blast_model_converter, m);
func_decl_ref_vector const& old_funcs = m_rewriter.m_cfg.old_funcs();
func_decl_ref_vector const& new_funcs = m_rewriter.m_cfg.new_funcs();
for (unsigned i = 0; i < old_funcs.size(); ++i) {
fmc->insert(new_funcs[i]);
bvmc->insert(old_funcs[i], new_funcs[i]);
}
m_context.add_model_converter(concat(bvmc, fmc));
}
return result;
}
};
mk_bit_blast::mk_bit_blast(context & ctx, unsigned priority) : plugin(priority) {
m_impl = alloc(impl, ctx);
}
mk_bit_blast::~mk_bit_blast() {
dealloc(m_impl);
}
rule_set * mk_bit_blast::operator()(rule_set const & source) {
return (*m_impl)(source);
}
};

38
src/muz/dl_mk_bit_blast.h Normal file
View file

@ -0,0 +1,38 @@
/*++
Copyright (c) 2012 Microsoft Corporation
Module Name:
dl_mk_bit_blast.h
Abstract:
Functor for bit-blasting a rule set
Author:
Nikolaj Bjorner (nbjorner) 2012-08-30
Revision History:
--*/
#ifndef _DL_MK_BIT_BLAST_H_
#define _DL_MK_BIT_BLAST_H_
#include"dl_rule_transformer.h"
namespace datalog {
class mk_bit_blast : public rule_transformer::plugin {
class impl;
impl* m_impl;
public:
mk_bit_blast(context & ctx, unsigned priority = 35000);
~mk_bit_blast();
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_BIT_BLAST_H_ */

199
src/muz/dl_mk_coalesce.cpp Normal file
View file

@ -0,0 +1,199 @@
/*++
Copyright (c) 2012 Microsoft Corporation
Module Name:
dl_mk_coalesce.cpp
Abstract:
Coalesce rules with shared bodies.
Author:
Nikolaj Bjorner (nbjorner) 2012-10-15
Revision History:
Notes:
Implements proof rule of the form:
a(x) & q(x) -> p(x), b(y) & q(y) -> p(y)
----------------------------------------------
(a(z) \/ b(z)) & q(z) -> p(z)
--*/
#include "dl_mk_coalesce.h"
#include "bool_rewriter.h"
namespace datalog {
mk_coalesce::mk_coalesce(context& ctx):
rule_transformer::plugin(50, false),
m_ctx(ctx),
m(ctx.get_manager()),
rm(ctx.get_rule_manager()),
m_sub1(m),
m_sub2(m),
m_idx(0)
{}
void mk_coalesce::mk_pred(app_ref& pred, app* p1, app* p2) {
SASSERT(p1->get_decl() == p2->get_decl());
unsigned sz = p1->get_num_args();
expr_ref_vector args(m);
for (unsigned i = 0; i < sz; ++i) {
expr* a = p1->get_arg(i);
expr* b = p2->get_arg(i);
SASSERT(m.get_sort(a) == m.get_sort(b));
m_sub1.push_back(a);
m_sub2.push_back(b);
args.push_back(m.mk_var(m_idx++, m.get_sort(a)));
}
pred = m.mk_app(p1->get_decl(), args.size(), args.c_ptr());
}
void mk_coalesce::extract_conjs(expr_ref_vector const& sub, rule const& rl, expr_ref& result) {
obj_map<expr, unsigned> indices;
bool_rewriter bwr(m);
rule_ref r(const_cast<rule*>(&rl), rm);
ptr_vector<sort> sorts;
expr_ref_vector revsub(m), conjs(m);
rl.get_vars(sorts);
revsub.resize(sorts.size());
svector<bool> valid(sorts.size(), true);
for (unsigned i = 0; i < sub.size(); ++i) {
expr* e = sub[i];
sort* s = m.get_sort(e);
expr_ref w(m.mk_var(i, s), m);
if (is_var(e)) {
unsigned v = to_var(e)->get_idx();
SASSERT(v < valid.size());
if (sorts[v]) {
SASSERT(s == sorts[v]);
if (valid[v]) {
revsub[v] = w;
valid[v] = false;
}
else {
SASSERT(revsub[v].get());
SASSERT(m.get_sort(revsub[v].get()) == s);
conjs.push_back(m.mk_eq(revsub[v].get(), w));
}
}
}
else {
SASSERT(m.is_value(e));
SASSERT(m.get_sort(e) == m.get_sort(w));
conjs.push_back(m.mk_eq(e, w));
}
}
for (unsigned i = 0; i < sorts.size(); ++i) {
if (valid[i] && sorts[i] && !revsub[i].get()) {
revsub[i] = m.mk_var(m_idx++, sorts[i]);
}
}
var_subst vs(m, false);
for (unsigned i = r->get_uninterpreted_tail_size(); i < r->get_tail_size(); ++i) {
vs(r->get_tail(i), revsub.size(), revsub.c_ptr(), result);
conjs.push_back(result);
}
bwr.mk_and(conjs.size(), conjs.c_ptr(), result);
}
void mk_coalesce::merge_rules(rule_ref& tgt, rule const& src) {
SASSERT(same_body(*tgt.get(), src));
m_sub1.reset();
m_sub2.reset();
m_idx = 0;
app_ref pred(m), head(m);
expr_ref fml1(m), fml2(m), fml(m);
app_ref_vector tail(m);
ptr_vector<sort> sorts1, sorts2;
expr_ref_vector conjs1(m), conjs(m);
rule_ref res(rm);
bool_rewriter bwr(m);
svector<bool> is_neg;
tgt->get_vars(sorts1);
src.get_vars(sorts2);
mk_pred(head, src.get_head(), tgt->get_head());
for (unsigned i = 0; i < src.get_uninterpreted_tail_size(); ++i) {
mk_pred(pred, src.get_tail(i), tgt->get_tail(i));
tail.push_back(pred);
is_neg.push_back(src.is_neg_tail(i));
}
extract_conjs(m_sub1, src, fml1);
extract_conjs(m_sub2, *tgt.get(), fml2);
bwr.mk_or(fml1, fml2, fml);
SASSERT(is_app(fml));
tail.push_back(to_app(fml));
is_neg.push_back(false);
res = rm.mk(head, tail.size(), tail.c_ptr(), is_neg.c_ptr(), tgt->name());
if (m_ctx.generate_proof_trace()) {
src.to_formula(fml1);
tgt->to_formula(fml2);
res->to_formula(fml);
#if 0
sort* ps = m.mk_proof_sort();
sort* domain[3] = { ps, ps, m.mk_bool_sort() };
func_decl* merge = m.mk_func_decl(symbol("merge-clauses"), 3, domain, ps); // TBD: ad-hoc proof rule
expr* args[3] = { m.mk_asserted(fml1), m.mk_asserted(fml2), fml };
// ...m_pc->insert(m.mk_app(merge, 3, args));
#else
svector<std::pair<unsigned, unsigned> > pos;
vector<expr_ref_vector> substs;
proof* p = src.get_proof();
p = m.mk_hyper_resolve(1, &p, fml, pos, substs);
res->set_proof(m, p);
#endif
}
tgt = res;
}
bool mk_coalesce::same_body(rule const& r1, rule const& r2) const {
SASSERT(r1.get_decl() == r2.get_decl());
unsigned sz = r1.get_uninterpreted_tail_size();
if (sz != r2.get_uninterpreted_tail_size()) {
return false;
}
for (unsigned i = 0; i < sz; ++i) {
if (r1.get_decl(i) != r2.get_decl(i)) {
return false;
}
if (r1.is_neg_tail(i) != r2.is_neg_tail(i)) {
return false;
}
}
return true;
}
rule_set * mk_coalesce::operator()(rule_set const & source) {
rule_set* rules = alloc(rule_set, m_ctx);
rules->inherit_predicates(source);
rule_set::decl2rules::iterator it = source.begin_grouped_rules(), end = source.end_grouped_rules();
for (; it != end; ++it) {
rule_ref_vector d_rules(rm);
d_rules.append(it->m_value->size(), it->m_value->c_ptr());
for (unsigned i = 0; i < d_rules.size(); ++i) {
rule_ref r1(d_rules[i].get(), rm);
for (unsigned j = i + 1; j < d_rules.size(); ++j) {
if (same_body(*r1.get(), *d_rules[j].get())) {
merge_rules(r1, *d_rules[j].get());
d_rules[j] = d_rules.back();
d_rules.pop_back();
--j;
}
}
rules->add_rule(r1.get());
}
}
return rules;
}
};

61
src/muz/dl_mk_coalesce.h Normal file
View file

@ -0,0 +1,61 @@
/*++
Copyright (c) 2012 Microsoft Corporation
Module Name:
dl_mk_coalesce.h
Abstract:
Coalesce rules with shared bodies.
Author:
Nikolaj Bjorner (nbjorner) 2012-10-15
Revision History:
--*/
#ifndef _DL_MK_COALESCE_H_
#define _DL_MK_COALESCE_H_
#include"dl_context.h"
#include"dl_rule_set.h"
#include"uint_set.h"
#include"dl_rule_transformer.h"
#include"dl_mk_rule_inliner.h"
namespace datalog {
/**
\brief Implements an unfolding transformation.
*/
class mk_coalesce : public rule_transformer::plugin {
context& m_ctx;
ast_manager& m;
rule_manager& rm;
expr_ref_vector m_sub1, m_sub2;
unsigned m_idx;
void mk_pred(app_ref& pred, app* p1, app* p2);
void extract_conjs(expr_ref_vector const& sub, rule const& rl, expr_ref& result);
bool same_body(rule const& r1, rule const& r2) const;
void merge_rules(rule_ref& tgt, rule const& src);
public:
/**
\brief Create coalesced rules.
*/
mk_coalesce(context & ctx);
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_COALESCE_H_ */

View file

@ -0,0 +1,201 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_coi_filter.cpp
Abstract:
Rule transformer which removes relations which are out of the cone of
influence of output relations
Author:
Krystof Hoder (t-khoder) 2011-10-01.
Revision History:
Andrey Rybalchenko (rybal) 2013-8-8
Added bottom_up pruning.
--*/
#include <sstream>
#include"ast_pp.h"
#include"dl_mk_coi_filter.h"
#include"extension_model_converter.h"
namespace datalog {
// -----------------------------------
//
// mk_coi_filter
//
// -----------------------------------
rule_set * mk_coi_filter::operator()(rule_set const & source) {
if (source.empty()) {
return 0;
}
scoped_ptr<rule_set> result1 = top_down(source);
scoped_ptr<rule_set> result2 = bottom_up(result1?*result1:source);
if (!result2) {
result2 = result1.detach();
}
return result2.detach();
}
rule_set * mk_coi_filter::bottom_up(rule_set const & source) {
decl_set all, reached;
ptr_vector<func_decl> todo;
rule_set::decl2rules body2rules;
// initialization for reachability
for (rule_set::iterator it = source.begin(); it != source.end(); ++it) {
rule * r = *it;
all.insert(r->get_decl());
if (r->get_uninterpreted_tail_size() == 0) {
if (!reached.contains(r->get_decl())) {
reached.insert(r->get_decl());
todo.insert(r->get_decl());
}
}
else {
for (unsigned i = 0; i < r->get_uninterpreted_tail_size(); ++i) {
func_decl * d = r->get_tail(i)->get_decl();
all.insert(d);
rule_set::decl2rules::obj_map_entry * e = body2rules.insert_if_not_there2(d, 0);
if (!e->get_data().m_value) {
e->get_data().m_value = alloc(ptr_vector<rule>);
}
e->get_data().m_value->push_back(r);
}
}
}
// reachability computation
while (!todo.empty()) {
func_decl * d = todo.back();
todo.pop_back();
ptr_vector<rule> * rules;
if (!body2rules.find(d, rules)) continue;
for (ptr_vector<rule>::iterator it = rules->begin(); it != rules->end(); ++it) {
rule * r = *it;
if (reached.contains(r->get_decl())) continue;
bool contained = true;
for (unsigned i = 0; contained && i < r->get_uninterpreted_tail_size(); ++i) {
contained = reached.contains(r->get_tail(i)->get_decl());
}
if (!contained) continue;
reached.insert(r->get_decl());
todo.insert(r->get_decl());
}
}
// eliminate each rule when some body predicate is not reached
scoped_ptr<rule_set> res = alloc(rule_set, m_context);
res->inherit_predicates(source);
for (rule_set::iterator it = source.begin(); it != source.end(); ++it) {
rule * r = *it;
bool contained = true;
for (unsigned i = 0; contained && i < r->get_uninterpreted_tail_size(); ++i) {
contained = reached.contains(r->get_tail(i)->get_decl());
}
if (contained) {
res->add_rule(r);
}
}
if (res->get_num_rules() == source.get_num_rules()) {
TRACE("dl", tout << "No transformation\n";);
res = 0;
}
else {
res->close();
}
// set to false each unreached predicate
if (m_context.get_model_converter()) {
extension_model_converter* mc0 = alloc(extension_model_converter, m);
for (decl_set::iterator it = all.begin(); it != all.end(); ++it) {
if (!reached.contains(*it)) {
mc0->insert(*it, m.mk_false());
}
}
m_context.add_model_converter(mc0);
}
// clean up body2rules range resources
for (rule_set::decl2rules::iterator it = body2rules.begin(); it != body2rules.end(); ++it) {
dealloc(it->m_value);
}
CTRACE("dl", 0 != res, res->display(tout););
return res.detach();
}
rule_set * mk_coi_filter::top_down(rule_set const & source) {
decl_set interesting_preds;
decl_set pruned_preds;
ptr_vector<func_decl> todo;
{
const decl_set& output_preds = source.get_output_predicates();
decl_set::iterator oend = output_preds.end();
for (decl_set::iterator it = output_preds.begin(); it!=oend; ++it) {
todo.push_back(*it);
interesting_preds.insert(*it);
}
}
const rule_dependencies& deps = source.get_dependencies();
while (!todo.empty()) {
func_decl * curr = todo.back();
todo.pop_back();
interesting_preds.insert(curr);
const rule_dependencies::item_set& cdeps = deps.get_deps(curr);
rule_dependencies::item_set::iterator dend = cdeps.end();
for (rule_dependencies::item_set::iterator it = cdeps.begin(); it != dend; ++it) {
func_decl * dep_pred = *it;
if (!interesting_preds.contains(dep_pred)) {
interesting_preds.insert(dep_pred);
todo.push_back(dep_pred);
}
}
}
scoped_ptr<rule_set> res = alloc(rule_set, m_context);
res->inherit_predicates(source);
rule_set::iterator rend = source.end();
for (rule_set::iterator rit = source.begin(); rit != rend; ++rit) {
rule * r = *rit;
func_decl * pred = r->get_decl();
if (interesting_preds.contains(pred)) {
res->add_rule(r);
}
else if (m_context.get_model_converter()) {
pruned_preds.insert(pred);
}
}
if (res->get_num_rules() == source.get_num_rules()) {
TRACE("dl", tout << "No transformation\n";);
res = 0;
}
if (res && m_context.get_model_converter()) {
decl_set::iterator end = pruned_preds.end();
decl_set::iterator it = pruned_preds.begin();
extension_model_converter* mc0 = alloc(extension_model_converter, m);
for (; it != end; ++it) {
mc0->insert(*it, m.mk_true());
}
m_context.add_model_converter(mc0);
}
CTRACE("dl", 0 != res, res->display(tout););
return res.detach();
}
};

View file

@ -0,0 +1,51 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_coi_filter.h
Abstract:
Rule transformer which removes relations which are out of the cone of
influence of output relations
Author:
Krystof Hoder (t-khoder) 2011-10-01.
Revision History:
--*/
#ifndef _DL_MK_COI_FILTER_H_
#define _DL_MK_COI_FILTER_H_
#include "dl_context.h"
#include "dl_rule_transformer.h"
namespace datalog {
class mk_coi_filter : public rule_transformer::plugin {
typedef obj_map<func_decl, func_decl *> decl_map;
ast_manager & m;
context & m_context;
rule_set * bottom_up(rule_set const & source);
rule_set * top_down(rule_set const & source);
public:
mk_coi_filter(context & ctx, unsigned priority=45000)
: plugin(priority),
m(ctx.get_manager()),
m_context(ctx) {}
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_COI_FILTER_H_ */

38
src/muz/dl_mk_different.h Normal file
View file

@ -0,0 +1,38 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_different_symbolic.h
Abstract:
Create Horn clauses for different symbolic transformation.
Author:
Nikolaj Bjorner (nbjorner) 2013-06-19
Revision History:
--*/
#ifndef _DL_MK_DIFFERENT_SYMBOLIC_H_
#define _DL_MK_DIFFERENT_SYMBOLIC_H_
#include"dl_rule_transformer.h"
namespace datalog {
class mk_different_symbolic : public rule_transformer::plugin {
ast_manager& m;
context& m_ctx;
public:
mk_different_symbolic(context & ctx, unsigned priority = 33037);
~mk_different_symbolic();
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_DIFFERENT_SYMBOLIC_H_ */

View file

@ -0,0 +1,880 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_explanations.cpp
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-11-08.
Revision History:
--*/
#include <sstream>
#include"ast_pp.h"
#include "ast_smt_pp.h"
#include"dl_finite_product_relation.h"
#include"dl_product_relation.h"
#include"dl_sieve_relation.h"
#include"dl_mk_explanations.h"
namespace datalog {
// -----------------------------------
//
// explanation_relation_plugin declaration
//
// -----------------------------------
class explanation_relation;
class explanation_relation_plugin : public relation_plugin {
friend class explanation_relation;
class join_fn;
class project_fn;
class rename_fn;
class union_fn;
class foreign_union_fn;
class assignment_filter_fn;
class negation_filter_fn;
class intersection_filter_fn;
bool m_relation_level_explanations;
func_decl_ref m_union_decl;
vector<ptr_vector<explanation_relation> > m_pool;
app * mk_union(app * a1, app * a2) {
return get_ast_manager().mk_app(m_union_decl, a1, a2);
}
public:
static symbol get_name(bool relation_level) {
return symbol(relation_level ? "relation_explanation" : "fact_explanation");
}
explanation_relation_plugin(bool relation_level, relation_manager & manager)
: relation_plugin(get_name(relation_level), manager),
m_relation_level_explanations(relation_level),
m_union_decl(mk_explanations::get_union_decl(get_context()), get_ast_manager()) {}
~explanation_relation_plugin() {
for (unsigned i = 0; i < m_pool.size(); ++i) {
for (unsigned j = 0; j < m_pool[i].size(); ++j) {
dealloc(m_pool[i][j]);
}
}
}
virtual bool can_handle_signature(const relation_signature & s) {
unsigned n=s.size();
for (unsigned i=0; i<n; i++) {
if (!get_context().get_decl_util().is_rule_sort(s[i])) {
return false;
}
}
return true;
}
virtual relation_base * mk_empty(const relation_signature & s);
void recycle(explanation_relation* r);
protected:
virtual relation_join_fn * mk_join_fn(const relation_base & t1, const relation_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
virtual relation_transformer_fn * mk_project_fn(const relation_base & t, unsigned col_cnt,
const unsigned * removed_cols);
virtual relation_transformer_fn * mk_rename_fn(const relation_base & t, unsigned permutation_cycle_len,
const unsigned * permutation_cycle);
virtual relation_union_fn * mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
virtual relation_mutator_fn * mk_filter_interpreted_fn(const relation_base & t, app * condition);
virtual relation_intersection_filter_fn * mk_filter_by_negation_fn(const relation_base & t,
const relation_base & negated_obj, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * negated_cols);
virtual relation_intersection_filter_fn * mk_filter_by_intersection_fn(const relation_base & t,
const relation_base & src, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * src_cols);
};
// -----------------------------------
//
// explanation_relation
//
// -----------------------------------
class explanation_relation : public relation_base {
friend class explanation_relation_plugin;
friend class explanation_relation_plugin::join_fn;
friend class explanation_relation_plugin::project_fn;
friend class explanation_relation_plugin::rename_fn;
friend class explanation_relation_plugin::union_fn;
friend class explanation_relation_plugin::foreign_union_fn;
friend class explanation_relation_plugin::assignment_filter_fn;
friend class explanation_relation_plugin::intersection_filter_fn;
bool m_empty;
/**
Valid only if \c !m_empty.
Zero elements mean undefined.
*/
relation_fact m_data;
explanation_relation(explanation_relation_plugin & p, const relation_signature & s)
: relation_base(p, s), m_empty(true), m_data(p.get_ast_manager()) {
DEBUG_CODE(
unsigned sz = s.size();
for (unsigned i=0;i<sz; i++) {
SASSERT( p.get_context().get_decl_util().is_rule_sort(s[i]) );
}
);
}
void assign_data(const relation_fact & f) {
m_empty = false;
unsigned n=get_signature().size();
SASSERT(f.size()==n);
m_data.reset();
m_data.append(n, f.c_ptr());
}
void set_undefined() {
m_empty = false;
m_data.reset();
m_data.resize(get_signature().size());
}
void unite_with_data(const relation_fact & f) {
if (empty()) {
assign_data(f);
return;
}
unsigned n=get_signature().size();
SASSERT(f.size()==n);
for (unsigned i=0; i<n; i++) {
SASSERT(!is_undefined(i));
m_data[i] = get_plugin().mk_union(m_data[i], f[i]);
}
}
virtual void deallocate() {
get_plugin().recycle(this);
}
public:
explanation_relation_plugin & get_plugin() const {
return static_cast<explanation_relation_plugin &>(relation_base::get_plugin());
}
virtual void to_formula(expr_ref& fml) const {
ast_manager& m = fml.get_manager();
fml = m.mk_eq(m.mk_var(0, m.get_sort(m_data[0])), m_data[0]);
}
bool is_undefined(unsigned col_idx) const {
return m_data[col_idx]==0;
}
bool no_undefined() const {
if (empty()) {
return true;
}
unsigned n = get_signature().size();
for (unsigned i=0; i<n; i++) {
if (is_undefined(i)) {
return false;
}
}
return true;
}
virtual bool empty() const { return m_empty; }
virtual void reset() {
m_empty = true;
}
virtual void add_fact(const relation_fact & f) {
SASSERT(empty());
assign_data(f);
}
virtual bool contains_fact(const relation_fact & f) const {
UNREACHABLE();
throw 0;
}
virtual explanation_relation * clone() const {
explanation_relation * res = static_cast<explanation_relation *>(get_plugin().mk_empty(get_signature()));
res->m_empty = m_empty;
SASSERT(res->m_data.empty());
res->m_data.append(m_data);
return res;
}
virtual relation_base * complement(func_decl* pred) const {
explanation_relation * res = static_cast<explanation_relation *>(get_plugin().mk_empty(get_signature()));
if (empty()) {
res->set_undefined();
}
return res;
}
void display_explanation(app * expl, std::ostream & out) const {
if (expl) {
//TODO: some nice explanation output
ast_smt_pp pp(get_plugin().get_ast_manager());
pp.display_expr_smt2(out, expl);
}
else {
out << "<undefined>";
}
}
virtual void display(std::ostream & out) const {
if (empty()) {
out << "<empty explanation relation>\n";
return;
}
unsigned sz = get_signature().size();
for (unsigned i=0; i<sz; i++) {
if (i!=0) {
out << ", ";
}
display_explanation(m_data[0], out);
}
out << "\n";
}
virtual unsigned get_size_estimate() const { return empty() ? 0 : 1; }
};
// -----------------------------------
//
// explanation_relation_plugin
//
// -----------------------------------
relation_base * explanation_relation_plugin::mk_empty(const relation_signature & s) {
if (m_pool.size() > s.size() && !m_pool[s.size()].empty()) {
explanation_relation* r = m_pool[s.size()].back();
m_pool[s.size()].pop_back();
r->m_empty = true;
r->m_data.reset();
return r;
}
return alloc(explanation_relation, *this, s);
}
void explanation_relation_plugin::recycle(explanation_relation* r) {
relation_signature const& sig = r->get_signature();
if (m_pool.size() <= sig.size()) {
m_pool.resize(sig.size()+1);
}
m_pool[sig.size()].push_back(r);
}
class explanation_relation_plugin::join_fn : public convenient_relation_join_fn {
public:
join_fn(const relation_signature & sig1, const relation_signature & sig2)
: convenient_relation_join_fn(sig1, sig2, 0, 0, 0) {}
virtual relation_base * operator()(const relation_base & r1_0, const relation_base & r2_0) {
const explanation_relation & r1 = static_cast<const explanation_relation &>(r1_0);
const explanation_relation & r2 = static_cast<const explanation_relation &>(r2_0);
explanation_relation_plugin & plugin = r1.get_plugin();
explanation_relation * res = static_cast<explanation_relation *>(plugin.mk_empty(get_result_signature()));
if (!r1.empty() && !r2.empty()) {
res->m_empty = false;
SASSERT(res->m_data.empty());
res->m_data.append(r1.m_data);
res->m_data.append(r2.m_data);
}
return res;
}
};
relation_join_fn * explanation_relation_plugin::mk_join_fn(const relation_base & r1, const relation_base & r2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2) {
if (&r1.get_plugin()!=this || &r2.get_plugin()!=this) {
return 0;
}
if (col_cnt!=0) {
return 0;
}
return alloc(join_fn, r1.get_signature(), r2.get_signature());
}
class explanation_relation_plugin::project_fn : public convenient_relation_project_fn {
public:
project_fn(const relation_signature & sig, unsigned col_cnt, const unsigned * removed_cols)
: convenient_relation_project_fn(sig, col_cnt, removed_cols) {}
virtual relation_base * operator()(const relation_base & r0) {
const explanation_relation & r = static_cast<const explanation_relation &>(r0);
explanation_relation_plugin & plugin = r.get_plugin();
explanation_relation * res = static_cast<explanation_relation *>(plugin.mk_empty(get_result_signature()));
if (!r.empty()) {
relation_fact proj_data = r.m_data;
project_out_vector_columns(proj_data, m_removed_cols);
res->assign_data(proj_data);
}
return res;
}
};
relation_transformer_fn * explanation_relation_plugin::mk_project_fn(const relation_base & r, unsigned col_cnt,
const unsigned * removed_cols) {
if (&r.get_plugin()!=this) {
return 0;
}
return alloc(project_fn, r.get_signature(), col_cnt, removed_cols);
}
class explanation_relation_plugin::rename_fn : public convenient_relation_rename_fn {
public:
rename_fn(const relation_signature & sig, unsigned permutation_cycle_len, const unsigned * permutation_cycle)
: convenient_relation_rename_fn(sig, permutation_cycle_len, permutation_cycle) {}
virtual relation_base * operator()(const relation_base & r0) {
const explanation_relation & r = static_cast<const explanation_relation &>(r0);
explanation_relation_plugin & plugin = r.get_plugin();
explanation_relation * res = static_cast<explanation_relation *>(plugin.mk_empty(get_result_signature()));
if (!r.empty()) {
relation_fact permutated_data = r.m_data;
permutate_by_cycle(permutated_data, m_cycle);
res->assign_data(permutated_data);
}
return res;
}
};
relation_transformer_fn * explanation_relation_plugin::mk_rename_fn(const relation_base & r,
unsigned permutation_cycle_len, const unsigned * permutation_cycle) {
return alloc(rename_fn, r.get_signature(), permutation_cycle_len, permutation_cycle);
}
class explanation_relation_plugin::union_fn : public relation_union_fn {
scoped_ptr<relation_union_fn> m_delta_union_fun;
public:
virtual void operator()(relation_base & tgt0, const relation_base & src0, relation_base * delta0) {
explanation_relation & tgt = static_cast<explanation_relation &>(tgt0);
const explanation_relation & src = static_cast<const explanation_relation &>(src0);
explanation_relation * delta = delta0 ? static_cast<explanation_relation *>(delta0) : 0;
explanation_relation_plugin & plugin = tgt.get_plugin();
if (!src.no_undefined() || !tgt.no_undefined() || (delta && !delta->no_undefined())) {
UNREACHABLE();
}
if (src.empty()) {
return;
}
if (plugin.m_relation_level_explanations) {
tgt.unite_with_data(src.m_data);
if (delta) {
if (!m_delta_union_fun) {
m_delta_union_fun = plugin.get_manager().mk_union_fn(*delta, src);
SASSERT(m_delta_union_fun);
}
(*m_delta_union_fun)(*delta, src);
}
}
else {
if (tgt.empty()) {
tgt.assign_data(src.m_data);
if (delta && delta->empty()) {
delta->assign_data(src.m_data);
}
}
}
}
};
class explanation_relation_plugin::foreign_union_fn : public relation_union_fn {
scoped_ptr<relation_union_fn> m_delta_union_fun;
public:
virtual void operator()(relation_base & tgt0, const relation_base & src, relation_base * delta0) {
explanation_relation & tgt = static_cast<explanation_relation &>(tgt0);
explanation_relation * delta = delta0 ? static_cast<explanation_relation *>(delta0) : 0;
if (src.empty()) {
return;
}
tgt.set_undefined();
if (delta) {
delta->set_undefined();
}
}
};
relation_union_fn * explanation_relation_plugin::mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta) {
if (!check_kind(tgt) || (delta && !check_kind(*delta))) {
return 0;
}
if (!check_kind(src)) {
//this is to handle the product relation
return alloc(foreign_union_fn);
}
return alloc(union_fn);
}
class explanation_relation_plugin::assignment_filter_fn : public relation_mutator_fn {
ast_manager & m_manager;
var_subst & m_subst;
unsigned m_col_idx;
app_ref m_new_rule;
public:
assignment_filter_fn(context & ctx, unsigned col_idx, app_ref new_rule)
: m_manager(ctx.get_manager()),
m_subst(ctx.get_var_subst()),
m_col_idx(col_idx),
m_new_rule(new_rule) {}
virtual void operator()(relation_base & r0) {
explanation_relation & r = static_cast<explanation_relation &>(r0);
if (!r.is_undefined(m_col_idx)) {
UNREACHABLE();
}
unsigned sz = r.get_signature().size();
ptr_vector<expr> subst_arg;
subst_arg.resize(sz, 0);
unsigned ofs = sz-1;
for (unsigned i=0; i<sz; i++) {
SASSERT(!r.is_undefined(i) || !contains_var(m_new_rule, i));
subst_arg[ofs-i] = r.m_data.get(i);
}
expr_ref res(m_manager);
m_subst(m_new_rule, subst_arg.size(), subst_arg.c_ptr(), res);
r.m_data[m_col_idx] = to_app(res);
}
};
relation_mutator_fn * explanation_relation_plugin::mk_filter_interpreted_fn(const relation_base & r,
app * cond) {
if (&r.get_plugin()!=this) {
return 0;
}
ast_manager & m = get_ast_manager();
if (!m.is_eq(cond)) {
return 0;
}
expr * arg1 = cond->get_arg(0);
expr * arg2 = cond->get_arg(1);
if (is_var(arg2)) {
std::swap(arg1, arg2);
}
if (!is_var(arg1) || !is_app(arg2)) {
return 0;
}
var * col_var = to_var(arg1);
app * new_rule = to_app(arg2);
if (!get_context().get_decl_util().is_rule_sort(col_var->get_sort())) {
return 0;
}
unsigned col_idx = col_var->get_idx();
return alloc(assignment_filter_fn, get_context(), col_idx, app_ref(new_rule, get_ast_manager()));
}
class explanation_relation_plugin::negation_filter_fn : public relation_intersection_filter_fn {
public:
virtual void operator()(relation_base & r, const relation_base & neg) {
if (!neg.empty()) {
r.reset();
}
}
};
relation_intersection_filter_fn * explanation_relation_plugin::mk_filter_by_negation_fn(const relation_base & r,
const relation_base & neg, unsigned joined_col_cnt, const unsigned * t_cols,
const unsigned * negated_cols) {
if (&r.get_plugin()!=this || &neg.get_plugin()!=this) {
return 0;
}
return alloc(negation_filter_fn);
}
class explanation_relation_plugin::intersection_filter_fn : public relation_intersection_filter_fn {
func_decl_ref m_union_decl;
public:
intersection_filter_fn(explanation_relation_plugin & plugin)
: m_union_decl(plugin.m_union_decl) {}
virtual void operator()(relation_base & tgt0, const relation_base & src0) {
explanation_relation & tgt = static_cast<explanation_relation &>(tgt0);
const explanation_relation & src = static_cast<const explanation_relation &>(src0);
if (src.empty()) {
tgt.reset();
return;
}
if (tgt.empty()) {
return;
}
unsigned sz = tgt.get_signature().size();
for (unsigned i=0; i<sz; i++) {
if (src.is_undefined(i)) {
continue;
}
app * curr_src = src.m_data.get(i);
if (tgt.is_undefined(i)) {
tgt.m_data.set(i, curr_src);
continue;
}
app * curr_tgt = tgt.m_data.get(i);
if (curr_tgt->get_decl()==m_union_decl.get()) {
if (curr_tgt->get_arg(0)==curr_src || curr_tgt->get_arg(1)==curr_src) {
tgt.m_data.set(i, curr_src);
continue;
}
}
//the intersection is imprecise because we do nothing here, but it is good enough for
//the purpose of explanations
}
}
};
relation_intersection_filter_fn * explanation_relation_plugin::mk_filter_by_intersection_fn(
const relation_base & tgt, const relation_base & src, unsigned joined_col_cnt,
const unsigned * tgt_cols, const unsigned * src_cols) {
if (&tgt.get_plugin()!=this || &src.get_plugin()!=this) {
return 0;
}
//this checks the join is one to one on all columns
if (tgt.get_signature()!=src.get_signature()
|| joined_col_cnt!=tgt.get_signature().size()
|| !containers_equal(tgt_cols, tgt_cols+joined_col_cnt, src_cols, src_cols+joined_col_cnt)) {
return 0;
}
counter ctr;
ctr.count(joined_col_cnt, tgt_cols);
if (ctr.get_max_counter_value()>1 || (joined_col_cnt && ctr.get_max_positive()!=joined_col_cnt-1)) {
return 0;
}
return alloc(intersection_filter_fn, *this);
}
// -----------------------------------
//
// mk_explanations
//
// -----------------------------------
mk_explanations::mk_explanations(context & ctx)
: plugin(50000),
m_manager(ctx.get_manager()),
m_context(ctx),
m_decl_util(ctx.get_decl_util()),
m_relation_level(ctx.explanations_on_relation_level()),
m_pinned(m_manager) {
m_e_sort = m_decl_util.mk_rule_sort();
m_pinned.push_back(m_e_sort);
relation_manager & rmgr = ctx.get_rel_context()->get_rmanager();
symbol er_symbol = explanation_relation_plugin::get_name(m_relation_level);
m_er_plugin = static_cast<explanation_relation_plugin *>(rmgr.get_relation_plugin(er_symbol));
if (!m_er_plugin) {
m_er_plugin = alloc(explanation_relation_plugin, m_relation_level, rmgr);
rmgr.register_plugin(m_er_plugin);
if (!m_relation_level) {
DEBUG_CODE(
finite_product_relation_plugin * dummy;
SASSERT(!rmgr.try_get_finite_product_relation_plugin(*m_er_plugin, dummy));
);
rmgr.register_plugin(alloc(finite_product_relation_plugin, *m_er_plugin, rmgr));
}
}
DEBUG_CODE(
if (!m_relation_level) {
finite_product_relation_plugin * dummy;
SASSERT(rmgr.try_get_finite_product_relation_plugin(*m_er_plugin, dummy));
}
);
}
func_decl * mk_explanations::get_union_decl(context & ctx) {
ast_manager & m = ctx.get_manager();
sort_ref s(ctx.get_decl_util().mk_rule_sort(), m);
//can it happen that the function name would collide with some other symbol?
//if functions can be overloaded by their ranges, it should be fine.
return m.mk_func_decl(symbol("e_union"), s, s, s);
}
void mk_explanations::assign_rel_level_kind(func_decl * e_decl, func_decl * orig) {
SASSERT(m_relation_level);
relation_manager & rmgr = m_context.get_rel_context()->get_rmanager();
unsigned sz = e_decl->get_arity();
relation_signature sig;
rmgr.from_predicate(e_decl, sig);
svector<bool> inner_sieve(sz-1, true);
inner_sieve.push_back(false);
svector<bool> expl_sieve(sz-1, false);
expl_sieve.push_back(true);
sieve_relation_plugin & sieve_plugin = sieve_relation_plugin::get_plugin(rmgr);
family_id inner_kind = rmgr.get_requested_predicate_kind(orig); //may be null_family_id
family_id inner_sieve_kind = sieve_plugin.get_relation_kind(sig, inner_sieve, inner_kind);
family_id expl_kind = m_er_plugin->get_kind();
family_id expl_sieve_kind = sieve_plugin.get_relation_kind(sig, expl_sieve, expl_kind);
product_relation_plugin::rel_spec product_spec;
product_spec.push_back(inner_sieve_kind);
product_spec.push_back(expl_sieve_kind);
family_id pred_kind =
product_relation_plugin::get_plugin(rmgr).get_relation_kind(sig, product_spec);
rmgr.set_predicate_kind(e_decl, pred_kind);
}
func_decl * mk_explanations::get_e_decl(func_decl * orig_decl) {
decl_map::obj_map_entry * e = m_e_decl_map.insert_if_not_there2(orig_decl, 0);
if (e->get_data().m_value==0) {
relation_signature e_domain;
e_domain.append(orig_decl->get_arity(), orig_decl->get_domain());
e_domain.push_back(m_e_sort);
func_decl * new_decl = m_context.mk_fresh_head_predicate(orig_decl->get_name(), symbol("expl"),
e_domain.size(), e_domain.c_ptr(), orig_decl);
m_pinned.push_back(new_decl);
e->get_data().m_value = new_decl;
if (m_relation_level) {
assign_rel_level_kind(new_decl, orig_decl);
}
}
return e->get_data().m_value;
}
app * mk_explanations::get_e_lit(app * lit, unsigned e_var_idx) {
expr_ref_vector args(m_manager);
func_decl * e_decl = get_e_decl(lit->get_decl());
args.append(lit->get_num_args(), lit->get_args());
args.push_back(m_manager.mk_var(e_var_idx, m_e_sort));
return m_manager.mk_app(e_decl, args.c_ptr());
}
symbol mk_explanations::get_rule_symbol(rule * r) {
if (r->name() == symbol::null) {
std::stringstream sstm;
r->display(m_context, sstm);
std::string res = sstm.str();
res = res.substr(0, res.find_last_not_of('\n')+1);
return symbol(res.c_str());
}
else {
return r->name();
}
}
rule * mk_explanations::get_e_rule(rule * r) {
rule_counter ctr;
ctr.count_rule_vars(m_manager, r);
unsigned max_var;
unsigned next_var = ctr.get_max_positive(max_var) ? (max_var+1) : 0;
unsigned head_var = next_var++;
app_ref e_head(get_e_lit(r->get_head(), head_var), m_manager);
app_ref_vector e_tail(m_manager);
svector<bool> neg_flags;
unsigned pos_tail_sz = r->get_positive_tail_size();
for (unsigned i=0; i<pos_tail_sz; i++) {
unsigned e_var = next_var++;
e_tail.push_back(get_e_lit(r->get_tail(i), e_var));
neg_flags.push_back(false);
}
unsigned tail_sz = r->get_tail_size();
for (unsigned i=pos_tail_sz; i<tail_sz; i++) {
e_tail.push_back(r->get_tail(i));
neg_flags.push_back(r->is_neg_tail(i));
}
symbol rule_repr = get_rule_symbol(r);
expr_ref_vector rule_expr_args(m_manager);
for (unsigned tail_idx=0; tail_idx<pos_tail_sz; tail_idx++) {
app * tail = e_tail.get(tail_idx);
if (true || m_relation_level) {
//this adds the explanation term of the tail
rule_expr_args.push_back(tail->get_arg(tail->get_num_args()-1));
}
else {
//this adds argument values and the explanation term
//(values will be substituted for variables at runtime by the finite_product_relation)
rule_expr_args.append(tail->get_num_args(), tail->get_args());
}
}
//rule_expr contains rule function with string representation of the rule as symbol and
//for each positive uninterpreted tail it contains its argument values and its explanation term
expr * rule_expr = m_decl_util.mk_rule(rule_repr, rule_expr_args.size(), rule_expr_args.c_ptr());
app_ref e_record(m_manager.mk_eq(m_manager.mk_var(head_var, m_e_sort), rule_expr), m_manager);
e_tail.push_back(e_record);
neg_flags.push_back(false);
SASSERT(e_tail.size()==neg_flags.size());
return m_context.get_rule_manager().mk(e_head, e_tail.size(), e_tail.c_ptr(), neg_flags.c_ptr());
}
void mk_explanations::transform_rules(const rule_set & src, rule_set & dst) {
rule_set::iterator rit = src.begin();
rule_set::iterator rend = src.end();
for (; rit!=rend; ++rit) {
rule * e_rule = get_e_rule(*rit);
dst.add_rule(e_rule);
}
//add rules that will (for output predicates) copy facts from explained relations back to
//the original ones
expr_ref_vector lit_args(m_manager);
decl_set::iterator pit = src.get_output_predicates().begin();
decl_set::iterator pend = src.get_output_predicates().end();
for (; pit != pend; ++pit) {
func_decl * orig_decl = *pit;
lit_args.reset();
unsigned arity = orig_decl->get_arity();
for (unsigned i=0; i<arity; i++) {
lit_args.push_back(m_manager.mk_var(i, orig_decl->get_domain(i)));
}
app_ref orig_lit(m_manager.mk_app(orig_decl, lit_args.c_ptr()), m_manager);
app_ref e_lit(get_e_lit(orig_lit, arity), m_manager);
app * tail[] = { e_lit.get() };
dst.add_rule(m_context.get_rule_manager().mk(orig_lit, 1, tail, 0));
}
}
void mk_explanations::translate_rel_level_relation(relation_manager & rmgr, relation_base & orig,
relation_base & e_rel) {
SASSERT(m_e_fact_relation);
SASSERT(e_rel.get_plugin().is_product_relation());
product_relation & prod_rel = static_cast<product_relation &>(e_rel);
SASSERT(prod_rel.size()==2);
SASSERT(prod_rel[0].get_plugin().is_sieve_relation());
SASSERT(prod_rel[1].get_plugin().is_sieve_relation());
sieve_relation * srels[] = {
static_cast<sieve_relation *>(&prod_rel[0]),
static_cast<sieve_relation *>(&prod_rel[1]) };
if (&srels[0]->get_inner().get_plugin()==m_er_plugin) {
std::swap(srels[0], srels[1]);
}
SASSERT(&srels[0]->get_inner().get_plugin()==&orig.get_plugin());
SASSERT(&srels[1]->get_inner().get_plugin()==m_er_plugin);
relation_base & new_orig = srels[0]->get_inner();
explanation_relation & expl_rel = static_cast<explanation_relation &>(srels[1]->get_inner());
{
scoped_ptr<relation_union_fn> orig_union_fun = rmgr.mk_union_fn(new_orig, orig);
SASSERT(orig_union_fun);
(*orig_union_fun)(new_orig, orig);
}
{
scoped_ptr<relation_union_fn> expl_union_fun = rmgr.mk_union_fn(expl_rel, *m_e_fact_relation);
SASSERT(expl_union_fun);
(*expl_union_fun)(expl_rel, *m_e_fact_relation);
}
}
void mk_explanations::transform_facts(relation_manager & rmgr, rule_set const& src, rule_set& dst) {
if (!m_e_fact_relation) {
relation_signature expl_singleton_sig;
expl_singleton_sig.push_back(m_e_sort);
relation_base * expl_singleton = rmgr.mk_empty_relation(expl_singleton_sig, m_er_plugin->get_kind());
relation_fact es_fact(m_manager);
es_fact.push_back(m_decl_util.mk_fact(symbol("fact")));
expl_singleton->add_fact(es_fact);
SASSERT(&expl_singleton->get_plugin()==m_er_plugin);
m_e_fact_relation = static_cast<explanation_relation *>(expl_singleton);
}
func_decl_set const& predicates = m_context.get_predicates();
decl_set::iterator it = predicates.begin();
decl_set::iterator end = predicates.end();
for (; it!=end; ++it) {
func_decl * orig_decl = *it;
func_decl * e_decl = get_e_decl(orig_decl);
if (!rmgr.try_get_relation(orig_decl) &&
!src.contains(orig_decl)) {
// there are no facts or rules for this predicate
continue;
}
dst.inherit_predicate(src, orig_decl, e_decl);
relation_base & orig_rel = rmgr.get_relation(orig_decl);
relation_base & e_rel = rmgr.get_relation(e_decl);
SASSERT(e_rel.empty()); //the e_rel should be a new relation
if (m_relation_level) {
translate_rel_level_relation(rmgr, orig_rel, e_rel);
}
else {
scoped_ptr<relation_join_fn> product_fun = rmgr.mk_join_fn(orig_rel, *m_e_fact_relation, 0, 0, 0);
SASSERT(product_fun);
scoped_rel<relation_base> aux_extended_rel = (*product_fun)(orig_rel, *m_e_fact_relation);
scoped_ptr<relation_union_fn> union_fun = rmgr.mk_union_fn(e_rel, *aux_extended_rel);
SASSERT(union_fun);
(*union_fun)(e_rel, *aux_extended_rel);
}
}
}
rule_set * mk_explanations::operator()(rule_set const & source) {
if (source.empty()) {
return 0;
}
if (!m_context.generate_explanations()) {
return 0;
}
rule_set * res = alloc(rule_set, m_context);
transform_facts(m_context.get_rel_context()->get_rmanager(), source, *res);
transform_rules(source, *res);
return res;
}
};

View file

@ -0,0 +1,86 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_explanations.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-11-08.
Revision History:
--*/
#ifndef _DL_MK_EXPLANATIONS_H_
#define _DL_MK_EXPLANATIONS_H_
#include "dl_context.h"
#include "dl_rule_transformer.h"
namespace datalog {
class explanation_relation;
class explanation_relation_plugin;
class mk_explanations : public rule_transformer::plugin {
typedef obj_map<func_decl, func_decl *> decl_map;
ast_manager & m_manager;
context & m_context;
dl_decl_util & m_decl_util;
bool m_relation_level;
ast_ref_vector m_pinned;
explanation_relation_plugin * m_er_plugin;
sort * m_e_sort;
scoped_rel<explanation_relation> m_e_fact_relation;
decl_map m_e_decl_map;
symbol get_rule_symbol(rule * r);
app * get_e_lit(app * lit, unsigned e_var_idx);
rule * get_e_rule(rule * r);
/**
If \c m_relation_level is true, ensure \c e_decl predicate will be represented by
the right relation object. \c orig is the predicate corresponding to \c e_decl without
the explanation column.
*/
void assign_rel_level_kind(func_decl * e_decl, func_decl * orig);
void translate_rel_level_relation(relation_manager & rmgr, relation_base & orig, relation_base & e_rel);
void transform_rules(const rule_set & src, rule_set & dst);
void transform_facts(relation_manager & rmgr, rule_set const& src, rule_set& dst);
public:
/**
If relation_level is true, the explanation will not be stored for each fact,
but we will rather store history of the whole relation.
*/
mk_explanations(context & ctx);
/**
\brief Return explanation predicate that corresponds to \c orig_decl.
*/
func_decl * get_e_decl(func_decl * orig_decl);
static func_decl * get_union_decl(context & ctx);
func_decl * get_union_decl() const {
return get_union_decl(m_context);
}
rule_set * operator()(rule_set const & source);
static expr* get_explanation(relation_base const& r);
};
};
#endif /* _DL_MK_EXPLANATIONS_H_ */

View file

@ -0,0 +1,170 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_filter_rules.cpp
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2010-05-18.
Revision History:
--*/
#include"dl_mk_filter_rules.h"
#include"dl_context.h"
#include"for_each_expr.h"
#include"ast_pp.h"
namespace datalog {
mk_filter_rules::mk_filter_rules(context & ctx):
plugin(2000),
m_context(ctx),
m(ctx.get_manager()),
rm(ctx.get_rule_manager()),
m_result(0),
m_pinned(m) {
}
mk_filter_rules::~mk_filter_rules() {
ptr_vector<filter_key> to_dealloc;
filter_cache::iterator it = m_tail2filter.begin();
filter_cache::iterator end = m_tail2filter.end();
for(; it!=end; ++it) {
to_dealloc.push_back(it->m_key);
}
m_tail2filter.reset();
ptr_vector<filter_key>::iterator dit = to_dealloc.begin();
ptr_vector<filter_key>::iterator dend = to_dealloc.end();
for(; dit!=dend; ++dit) {
dealloc(*dit);
}
}
/**
\brief Return true if \c pred is a cadidate for a "filter" rule.
*/
bool mk_filter_rules::is_candidate(app * pred) {
if (!m_context.is_predicate(pred)) {
TRACE("mk_filter_rules", tout << mk_pp(pred, m) << "\nis not a candidate because it is interpreted.\n";);
return false;
}
var_idx_set used_vars;
unsigned n = pred->get_num_args();
for (unsigned i = 0; i < n; i++) {
expr * arg = pred->get_arg(i);
if (m.is_value(arg))
return true;
SASSERT(is_var(arg));
unsigned vidx = to_var(arg)->get_idx();
if (used_vars.contains(vidx))
return true;
used_vars.insert(vidx);
}
return false;
}
/**
\brief Create a "filter" (if it doesn't exist already) for the given predicate.
*/
func_decl * mk_filter_rules::mk_filter_decl(app * pred, var_idx_set const & non_local_vars) {
sort_ref_buffer filter_domain(m);
filter_key * key = alloc(filter_key, m);
mk_new_rule_tail(m, pred, non_local_vars, filter_domain, key->filter_args, key->new_pred);
func_decl * filter_decl = 0;
if (!m_tail2filter.find(key, filter_decl)) {
filter_decl = m_context.mk_fresh_head_predicate(pred->get_decl()->get_name(), symbol("filter"),
filter_domain.size(), filter_domain.c_ptr(), pred->get_decl());
m_pinned.push_back(filter_decl);
m_tail2filter.insert(key, filter_decl);
app_ref filter_head(m);
filter_head = m.mk_app(filter_decl, key->filter_args.size(), key->filter_args.c_ptr());
app * filter_tail = key->new_pred;
rule * filter_rule = m_context.get_rule_manager().mk(filter_head, 1, &filter_tail, (const bool *)0);
filter_rule->set_accounting_parent_object(m_context, m_current);
m_result->add_rule(filter_rule);
m_context.get_rule_manager().mk_rule_asserted_proof(*filter_rule);
}
else {
dealloc(key);
}
SASSERT(filter_decl != 0);
SASSERT(filter_decl->get_arity()==filter_domain.size());
return filter_decl;
}
void mk_filter_rules::process(rule * r) {
m_current = r;
app * new_head = r->get_head();
app_ref_vector new_tail(m);
svector<bool> new_is_negated;
unsigned sz = r->get_tail_size();
bool rule_modified = false;
for (unsigned i = 0; i < sz; i++) {
app * tail = r->get_tail(i);
if (is_candidate(tail)) {
TRACE("mk_filter_rules", tout << "is_candidate: " << mk_pp(tail, m) << "\n";);
var_idx_set non_local_vars = rm.collect_rule_vars_ex(r, tail);
func_decl * filter_decl = mk_filter_decl(tail, non_local_vars);
ptr_buffer<expr> new_args;
var_idx_set used_vars;
unsigned num_args = tail->get_num_args();
for (unsigned i = 0; i < num_args; i++) {
expr * arg = tail->get_arg(i);
if (is_var(arg)) {
unsigned vidx = to_var(arg)->get_idx();
if (non_local_vars.contains(vidx) && !used_vars.contains(vidx)) {
new_args.push_back(arg);
used_vars.insert(vidx);
}
}
}
SASSERT(new_args.size() == filter_decl->get_arity());
new_tail.push_back(m.mk_app(filter_decl, new_args.size(), new_args.c_ptr()));
rule_modified = true;
}
else {
new_tail.push_back(tail);
}
new_is_negated.push_back(r->is_neg_tail(i));
}
if (rule_modified) {
remove_duplicate_tails(new_tail, new_is_negated);
SASSERT(new_tail.size() == new_is_negated.size());
rule * new_rule = m_context.get_rule_manager().mk(new_head, new_tail.size(), new_tail.c_ptr(), new_is_negated.c_ptr());
new_rule->set_accounting_parent_object(m_context, m_current);
m_result->add_rule(new_rule);
m_context.get_rule_manager().mk_rule_rewrite_proof(*r, *new_rule);
m_modified = true;
}
else {
m_result->add_rule(r);
}
}
rule_set * mk_filter_rules::operator()(rule_set const & source) {
m_tail2filter.reset();
m_result = alloc(rule_set, m_context);
m_modified = false;
unsigned num_rules = source.get_num_rules();
for (unsigned i = 0; i < num_rules; i++) {
process(source.get_rule(i));
}
if(!m_modified) {
dealloc(m_result);
return static_cast<rule_set *>(0);
}
m_result->inherit_predicates(source);
return m_result;
}
};

View file

@ -0,0 +1,87 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_filter_rules.h
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2010-05-18.
Revision History:
--*/
#ifndef _DL_MK_FILTER_RULES_H_
#define _DL_MK_FILTER_RULES_H_
#include"map.h"
#include"dl_rule_set.h"
#include"dl_rule_transformer.h"
namespace datalog {
/**
\brief Functor for applying a rule set transformation that creates "filters".
A "filter" is a rule of the form:
Head(X_1, ..., X_n) :- Tail(...)
where X_1,...,X_n are distinct, and Tail contain repeated variables and/or values.
After applying this functor only "filter" rules will contain atoms with repeated variables and/or values.
*/
class mk_filter_rules : public rule_transformer::plugin {
struct filter_key {
app_ref new_pred;
expr_ref_buffer filter_args;
filter_key(ast_manager & m) : new_pred(m), filter_args(m) {}
unsigned hash() const {
unsigned r = new_pred->hash();
for (unsigned i = 0; i < filter_args.size(); ++i) {
r ^= filter_args[i]->hash();
}
return r;
}
bool operator==(const filter_key & o) const {
return o.new_pred==new_pred && vectors_equal(o.filter_args, filter_args);
}
};
typedef obj_map<filter_key, func_decl*> filter_cache;
context & m_context;
ast_manager & m;
rule_manager & rm;
filter_cache m_tail2filter;
rule_set * m_result;
rule * m_current;
bool m_modified;
ast_ref_vector m_pinned;
bool is_candidate(app * pred);
func_decl * mk_filter_decl(app * pred, var_idx_set const & non_local_vars);
void process(rule * r);
public:
mk_filter_rules(context & ctx);
~mk_filter_rules();
/**
\brief Return a new rule set where only filter rules contain atoms with repeated variables and/or values.
*/
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_FILTER_RULES_H_ */

View file

@ -0,0 +1,618 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_interp_tail_simplifier.cpp
Abstract:
Rule transformer which simplifies interpreted tails
Author:
Krystof Hoder (t-khoder) 2011-10-01.
Revision History:
--*/
#include <sstream>
#include"ast_pp.h"
#include"bool_rewriter.h"
#include"rewriter.h"
#include"rewriter_def.h"
#include"dl_mk_rule_inliner.h"
#include"dl_mk_interp_tail_simplifier.h"
namespace datalog {
// -----------------------------------
//
// mk_interp_tail_simplifier::rule_substitution
//
// -----------------------------------
void mk_interp_tail_simplifier::rule_substitution::reset(rule * r) {
unsigned var_cnt = m_context.get_rule_manager().get_counter().get_max_rule_var(*r)+1;
m_subst.reset();
m_subst.reserve(1, var_cnt);
m_rule = r;
}
bool mk_interp_tail_simplifier::rule_substitution::unify(expr * e1, expr * e2) {
SASSERT(m_rule);
//we need to apply the current substitution in order to ensure the unifier
//works in an incremental way
expr_ref e1_s(m);
expr_ref e2_s(m);
m_subst.apply(e1,e1_s);
m_subst.apply(e2,e2_s);
//and we need to reset the cache as we're going to modify the substitution
m_subst.reset_cache();
return m_unif (e1_s, e2_s, m_subst, false);
}
void mk_interp_tail_simplifier::rule_substitution::apply(app * a, app_ref& res) {
SASSERT(m_rule);
expr_ref res_e(m);
m_subst.apply(a, res_e);
SASSERT(is_app(res_e.get()));
res = to_app(res_e.get());
}
void mk_interp_tail_simplifier::rule_substitution::get_result(rule_ref & res) {
SASSERT(m_rule);
apply(m_rule->get_head(), m_head);
m_tail.reset();
m_neg.reset();
unsigned tail_len = m_rule->get_tail_size();
for (unsigned i=0; i<tail_len; i++) {
app_ref new_tail_el(m);
apply(m_rule->get_tail(i), new_tail_el);
m_tail.push_back(new_tail_el);
m_neg.push_back(m_rule->is_neg_tail(i));
}
mk_rule_inliner::remove_duplicate_tails(m_tail, m_neg);
SASSERT(m_tail.size() == m_neg.size());
res = m_context.get_rule_manager().mk(m_head, m_tail.size(), m_tail.c_ptr(), m_neg.c_ptr());
res->set_accounting_parent_object(m_context, m_rule);
res->norm_vars(res.get_manager());
}
// -----------------------------------
//
// mk_interp_tail_simplifier
//
// -----------------------------------
class mk_interp_tail_simplifier::normalizer_cfg : public default_rewriter_cfg
{
struct expr_cmp
{
ast_manager& m;
expr_cmp(ast_manager& m) : m(m) {}
bool operator()(expr * ae, expr * be) {
return cmp_expr(ae, be, 4) == -1;
}
template<typename T>
static int cmp(T a, T b) { return (a>b) ? 1 : ((a == b) ? 0 : -1); }
int cmp_expr(expr * ae, expr * be, int depth) {
if (ae == be) { return 0; }
//remove negations
bool a_neg = m.is_not(ae, ae);
bool b_neg = m.is_not(be, be);
if (ae==be) { return cmp(a_neg, b_neg); }
if (!is_app(ae) && !is_app(be)) { return cmp(ae->get_id(), be->get_id()); }
if (!is_app(ae)) { return -1; }
if (!is_app(be)) { return 1; }
app * a = to_app(ae);
app * b = to_app(be);
if (a->get_decl()!=b->get_decl()) {
return cmp(a->get_decl()->get_id(), b->get_decl()->get_id());
}
if (a->get_num_args()!=b->get_num_args()) {
return cmp(a->get_num_args(), b->get_num_args());
}
if (depth==0) {
return cmp(a->get_id(),b->get_id());
}
unsigned arg_cnt = a->get_num_args();
unsigned neg_comparison = 0;
for (unsigned i=0; i<arg_cnt; i++) {
expr * arg_a = a->get_arg(i);
expr * arg_b = b->get_arg(i);
//we normalize away negations
bool a_is_neg = m.is_not(arg_a, arg_a);
bool b_is_neg = m.is_not(arg_b, arg_b);
if (neg_comparison==0 && a_is_neg!=b_is_neg) {
neg_comparison = a_is_neg ? -1 : 1;
}
int res = cmp_expr(arg_a, arg_b, depth-1);
if (res!=0) {
return res;
}
}
if (neg_comparison!=0) {
return neg_comparison;
}
//by normalizing away negation we may have put non-equal terms to be equal, so here we check
return cmp(a->get_id(),b->get_id());
}
};
ast_manager& m;
bool_rewriter m_brwr;
//instead of a local variable
expr_ref_vector m_app_args;
expr_cmp m_expr_cmp;
public:
normalizer_cfg(ast_manager& m)
: m(m), m_brwr(m), m_app_args(m), m_expr_cmp(m)
{
}
static void remove_duplicates(expr_ref_vector& v)
{
expr * a = v[0].get();
unsigned read_idx = 1;
unsigned write_idx = 1;
for (;;) {
while(read_idx<v.size() && a==v[read_idx].get()) {
read_idx++;
}
if (read_idx==v.size()) {
break;
}
a = v[read_idx].get();
if (write_idx!=read_idx) {
v[write_idx] = a;
}
write_idx++;
read_idx++;
}
v.shrink(write_idx);
}
typedef std::pair<expr *,expr *> arg_pair;
bool match_arg_pair(expr * e, arg_pair& pair, bool seek_conjunction)
{
if (seek_conjunction) {
return m.is_and(e, pair.first, pair.second);
}
else {
return m.is_or(e, pair.first, pair.second);
}
}
/**
If inside_disjunction is false, we're inside a conjunction (and arg pairs
represent disjunctions).
*/
app * detect_equivalence(const arg_pair& p1, const arg_pair& p2, bool inside_disjunction)
{
if (m.is_not(p1.first)==m.is_not(p2.first)) { return 0; }
if (m.is_not(p1.second)==m.is_not(p2.second)) { return 0; }
expr * first_bare = 0;
if (m.is_not(p1.first, first_bare) && p2.first!=first_bare) { return 0; }
if (m.is_not(p2.first, first_bare) && p1.first!=first_bare) { return 0; }
SASSERT(first_bare);
expr * second_bare = 0;
if (m.is_not(p1.second, second_bare) && p2.second!=second_bare) { return 0; }
if (m.is_not(p2.second, second_bare) && p1.second!=second_bare) { return 0; }
SASSERT(second_bare);
if (!m.is_bool(first_bare) || !m.is_bool(second_bare)) { return 0; }
//both negations are in the same pair
bool negs_together = m.is_not(p1.first)==m.is_not(p1.second);
if (negs_together==inside_disjunction) {
return m.mk_eq(first_bare, second_bare);
}
else {
return m.mk_eq(first_bare, m.mk_not(second_bare));
}
}
bool detect_equivalences(expr_ref_vector& v, bool inside_disjunction)
{
bool have_pair = false;
unsigned prev_pair_idx;
arg_pair ap;
unsigned read_idx = 0;
unsigned write_idx = 0;
while(read_idx<v.size()) {
expr * e = v[read_idx].get();
arg_pair new_ap;
if (match_arg_pair(e, new_ap, inside_disjunction)) {
app * neq = 0;
if (have_pair) {
neq = detect_equivalence(ap, new_ap, inside_disjunction);
}
if (neq) {
have_pair = false;
v[prev_pair_idx] = neq;
read_idx++;
continue;
}
else {
have_pair = true;
prev_pair_idx = write_idx;
ap = new_ap;
}
}
else {
have_pair = false;
}
if (write_idx!=read_idx) {
v[write_idx] = e;
}
read_idx++;
write_idx++;
}
v.shrink(write_idx);
return read_idx!=write_idx;
}
//bool detect_same_variable_conj_pairs
br_status reduce_app(func_decl * f, unsigned num, expr * const * args, expr_ref & result,
proof_ref & result_pr)
{
if (m.is_not(f) && (m.is_and(args[0]) || m.is_or(args[0]))) {
SASSERT(num==1);
expr_ref tmp(m);
app* a = to_app(args[0]);
m_app_args.reset();
for (unsigned i = 0; i < a->get_num_args(); ++i) {
m_brwr.mk_not(a->get_arg(i), tmp);
m_app_args.push_back(tmp);
}
if (m.is_and(args[0])) {
result = m.mk_or(m_app_args.size(), m_app_args.c_ptr());
}
else {
result = m.mk_and(m_app_args.size(), m_app_args.c_ptr());
}
return BR_REWRITE2;
}
if (!m.is_and(f) && !m.is_or(f)) {
return BR_FAILED;
}
if (num == 0) {
if (m.is_and(f)) {
result = m.mk_true();
}
else {
result = m.mk_false();
}
return BR_DONE;
}
if (num == 1) {
result = args[0];
return BR_DONE;
}
m_app_args.reset();
m_app_args.append(num, args);
std::sort(m_app_args.c_ptr(), m_app_args.c_ptr()+m_app_args.size(), m_expr_cmp);
remove_duplicates(m_app_args);
bool have_rewritten_args = false;
have_rewritten_args = detect_equivalences(m_app_args, m.is_or(f));
if (m_app_args.size()==1) {
result = m_app_args[0].get();
}
else {
if (m.is_and(f)) {
result = m.mk_and(m_app_args.size(), m_app_args.c_ptr());
}
else {
SASSERT(m.is_or(f));
result = m.mk_or(m_app_args.size(), m_app_args.c_ptr());
}
}
if (have_rewritten_args) {
return BR_REWRITE1;
}
return BR_DONE;
}
};
class mk_interp_tail_simplifier::normalizer_rw : public rewriter_tpl<normalizer_cfg> {
public:
normalizer_rw(ast_manager& m, normalizer_cfg& cfg): rewriter_tpl<normalizer_cfg>(m, false, cfg) {}
};
mk_interp_tail_simplifier::mk_interp_tail_simplifier(context & ctx, unsigned priority)
: plugin(priority),
m(ctx.get_manager()),
m_context(ctx),
m_simp(ctx.get_rewriter()),
a(m),
m_rule_subst(ctx),
m_tail(m),
m_itail_members(m),
m_conj(m) {
m_cfg = alloc(normalizer_cfg, m);
m_rw = alloc(normalizer_rw, m, *m_cfg);
}
mk_interp_tail_simplifier::~mk_interp_tail_simplifier() {
dealloc(m_rw);
dealloc(m_cfg);
}
void mk_interp_tail_simplifier::simplify_expr(app * a, expr_ref& res)
{
expr_ref simp1_res(m);
m_simp(a, simp1_res);
(*m_rw)(simp1_res.get(), res);
/*if (simp1_res.get()!=res.get()) {
std::cout<<"pre norm:\n"<<mk_pp(simp1_res.get(),m)<<"post norm:\n"<<mk_pp(res.get(),m)<<"\n";
}*/
m_simp(res.get(), res);
}
bool mk_interp_tail_simplifier::propagate_variable_equivalences(rule * r, rule_ref& res) {
unsigned u_len = r->get_uninterpreted_tail_size();
unsigned len = r->get_tail_size();
if (u_len == len) {
return false;
}
m_todo.reset();
m_leqs.reset();
for (unsigned i = u_len; i < len; i++) {
m_todo.push_back(r->get_tail(i));
SASSERT(!r->is_neg_tail(i));
}
m_rule_subst.reset(r);
expr_ref_vector trail(m);
expr_ref tmp1(m), tmp2(m);
bool found_something = false;
#define TRY_UNIFY(_x,_y) if (m_rule_subst.unify(_x,_y)) { found_something = true; }
#define IS_FLEX(_x) (is_var(_x) || m.is_value(_x))
while (!m_todo.empty()) {
expr * arg1, *arg2;
expr * t0 = m_todo.back();
m_todo.pop_back();
expr* t = t0;
bool neg = m.is_not(t, t);
if (is_var(t)) {
TRY_UNIFY(t, neg ? m.mk_false() : m.mk_true());
}
else if (!neg && m.is_and(t)) {
app* a = to_app(t);
m_todo.append(a->get_num_args(), a->get_args());
}
else if (!neg && m.is_eq(t, arg1, arg2) && IS_FLEX(arg1) && IS_FLEX(arg2)) {
TRY_UNIFY(arg1, arg2);
}
else if (m.is_iff(t, arg1, arg2)) {
//determine the polarity of the equivalence and remove the negations
while (m.is_not(arg1, arg1)) neg = !neg;
while (m.is_not(arg2, arg2)) neg = !neg;
if (!is_var(arg1)) {
std::swap(arg1, arg2);
}
if (!IS_FLEX(arg1) || !IS_FLEX(arg2)) {
// no-op
}
else if (is_var(arg1) && !neg) {
TRY_UNIFY(arg1, arg2);
}
else if (is_var(arg1) && neg && m.is_true(arg2)) {
TRY_UNIFY(arg1, m.mk_false());
}
else if (is_var(arg1) && neg && m.is_false(arg2)) {
TRY_UNIFY(arg1, m.mk_true());
}
}
else if (!neg && (a.is_le(t, arg1, arg2) || a.is_ge(t, arg2, arg1))) {
tmp1 = a.mk_sub(arg1, arg2);
tmp2 = a.mk_sub(arg2, arg1);
if (false && m_leqs.contains(tmp2) && IS_FLEX(arg1) && IS_FLEX(arg2)) {
TRY_UNIFY(arg1, arg2);
}
else {
trail.push_back(tmp1);
m_leqs.insert(tmp1);
}
}
}
if (!found_something) {
return false;
}
TRACE("dl_interp_tail_simplifier_propagation_pre",
tout << "will propagate rule:\n";
r->display(m_context, tout);
);
m_rule_subst.get_result(res);
TRACE("dl_interp_tail_simplifier_propagation",
tout << "propagated equivalences of:\n";
r->display(m_context, tout);
tout << "into:\n";
res->display(m_context, tout);
);
return true;
}
bool mk_interp_tail_simplifier::transform_rule(rule * r0, rule_ref & res)
{
rule_ref r(r0, m_context.get_rule_manager());
if (r->has_quantifiers()) {
res = r;
return true;
}
start:
unsigned u_len = r->get_uninterpreted_tail_size();
unsigned len = r->get_tail_size();
if (u_len == len) {
res = r;
return true;
}
app_ref head(r->get_head(), m);
m_tail.reset();
m_tail_neg.reset();
for (unsigned i=0; i<u_len; i++) {
m_tail.push_back(r->get_tail(i));
m_tail_neg.push_back(r->is_neg_tail(i));
}
bool modified = false;
app_ref itail(m);
if (u_len+1==len) {
//we have only one interpreted tail
itail = r->get_tail(u_len);
SASSERT(!r->is_neg_tail(u_len));
}
else {
m_itail_members.reset();
for (unsigned i=u_len; i<len; i++) {
m_itail_members.push_back(r->get_tail(i));
SASSERT(!r->is_neg_tail(i));
}
itail = m.mk_and(m_itail_members.size(), m_itail_members.c_ptr());
modified = true;
}
expr_ref simp_res(m);
simplify_expr(itail.get(), simp_res);
modified |= itail.get() != simp_res.get();
if (m.is_false(simp_res)) {
TRACE("dl", r->display(m_context, tout << "rule is infeasible\n"););
return false;
}
SASSERT(m.is_bool(simp_res));
if (modified) {
m_conj.reset();
qe::flatten_and(simp_res, m_conj);
for (unsigned i = 0; i < m_conj.size(); ++i) {
expr* e = m_conj[i].get();
if (is_app(e)) {
m_tail.push_back(to_app(e));
}
else {
m_tail.push_back(m.mk_eq(e, m.mk_true()));
}
m_tail_neg.push_back(false);
}
SASSERT(m_tail.size() == m_tail_neg.size());
res = m_context.get_rule_manager().mk(head, m_tail.size(), m_tail.c_ptr(), m_tail_neg.c_ptr());
res->set_accounting_parent_object(m_context, r);
}
else {
res = r;
}
rule_ref pro_var_eq_result(m_context.get_rule_manager());
if (propagate_variable_equivalences(res, pro_var_eq_result)) {
SASSERT(rule_counter().get_max_rule_var(*r.get())==0 ||
rule_counter().get_max_rule_var(*r.get()) > rule_counter().get_max_rule_var(*pro_var_eq_result.get()));
r = pro_var_eq_result;
goto start;
}
CTRACE("dl", (res != r0), r0->display(m_context, tout << "old:\n"); res->display(m_context, tout << "new:\n"););
return true;
}
bool mk_interp_tail_simplifier::transform_rules(const rule_set & orig, rule_set & tgt) {
bool modified = false;
rule_manager& rm = m_context.get_rule_manager();
rule_set::iterator rit = orig.begin();
rule_set::iterator rend = orig.end();
for (; rit!=rend; ++rit) {
rule_ref new_rule(rm);
if (transform_rule(*rit, new_rule)) {
rm.mk_rule_rewrite_proof(**rit, *new_rule.get());
bool is_modified = *rit != new_rule;
modified |= is_modified;
tgt.add_rule(new_rule);
}
else {
modified = true;
}
}
return modified;
}
rule_set * mk_interp_tail_simplifier::operator()(rule_set const & source) {
if (source.get_num_rules() == 0) {
return 0;
}
rule_set * res = alloc(rule_set, m_context);
if (transform_rules(source, *res)) {
res->inherit_predicates(source);
} else {
dealloc(res);
res = 0;
}
return res;
}
};

View file

@ -0,0 +1,109 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_interp_tail_simplifier.h
Abstract:
Rule transformer which simplifies interpreted tails
Author:
Krystof Hoder (t-khoder) 2011-10-01.
Revision History:
--*/
#ifndef _DL_MK_INTERP_TAIL_SIMPLIFIER_H_
#define _DL_MK_INTERP_TAIL_SIMPLIFIER_H_
#include "dl_context.h"
#include "dl_rule_transformer.h"
#include "unifier.h"
#include "substitution.h"
#include "arith_decl_plugin.h"
namespace datalog {
class mk_interp_tail_simplifier : public rule_transformer::plugin {
class rule_substitution
{
ast_manager& m;
context& m_context;
substitution m_subst;
unifier m_unif;
app_ref m_head;
app_ref_vector m_tail;
svector<bool> m_neg;
rule * m_rule;
void apply(app * a, app_ref& res);
public:
rule_substitution(context & ctx)
: m(ctx.get_manager()), m_context(ctx), m_subst(m), m_unif(m), m_head(m), m_tail(m), m_rule(0) {}
/**
Reset substitution and get it ready for working with rule r.
As long as this object is used without a reset, the rule r must exist.
*/
void reset(rule * r);
/** Reset subtitution and unify tail tgt_idx of the target rule and the head of the src rule */
bool unify(expr * e1, expr * e2);
void get_result(rule_ref & res);
void display(std::ostream& stm) {
m_subst.display(stm);
}
};
class normalizer_cfg;
class normalizer_rw;
ast_manager & m;
context & m_context;
th_rewriter & m_simp;
arith_util a;
rule_substitution m_rule_subst;
ptr_vector<expr> m_todo;
obj_hashtable<expr> m_leqs;
app_ref_vector m_tail;
expr_ref_vector m_itail_members;
expr_ref_vector m_conj;
svector<bool> m_tail_neg;
normalizer_cfg* m_cfg;
normalizer_rw* m_rw;
void simplify_expr(app * a, expr_ref& res);
/** return true if some propagation was done */
bool propagate_variable_equivalences(rule * r, rule_ref& res);
/** Return true if something was modified */
bool transform_rules(const rule_set & orig, rule_set & tgt);
public:
mk_interp_tail_simplifier(context & ctx, unsigned priority=40000);
virtual ~mk_interp_tail_simplifier();
/**If rule should be retained, assign transformed version to res and return true;
if rule can be deleted, return false.
This method is kind of useful, so it's public to allow other rules to use it,
e.g. on their intermediate results.
*/
bool transform_rule(rule * r, rule_ref& res);
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_INTERP_TAIL_SIMPLIFIER_H_ */

File diff suppressed because it is too large Load diff

View file

@ -0,0 +1,138 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_karr_invariants.h
Abstract:
Extract integer linear invariants.
Author:
Nikolaj Bjorner (nbjorner) 2013-03-08
Revision History:
--*/
#ifndef _DL_MK_KARR_INVARIANTS_H_
#define _DL_MK_KARR_INVARIANTS_H_
#include"dl_context.h"
#include"dl_rule_set.h"
#include"dl_rule_transformer.h"
#include"arith_decl_plugin.h"
#include"hilbert_basis.h"
namespace datalog {
/**
\brief Rule transformer that strengthens bodies with invariants.
*/
struct matrix {
vector<vector<rational> > A;
vector<rational> b;
svector<bool> eq;
unsigned size() const { return A.size(); }
void reset() { A.reset(); b.reset(); eq.reset(); }
matrix& operator=(matrix const& other);
void append(matrix const& other) { A.append(other.A); b.append(other.b); eq.append(other.eq); }
void display(std::ostream& out) const;
static void display_row(
std::ostream& out, vector<rational> const& row, rational const& b, bool is_eq);
static void display_ineq(
std::ostream& out, vector<rational> const& row, rational const& b, bool is_eq);
};
class mk_karr_invariants : public rule_transformer::plugin {
class add_invariant_model_converter;
context& m_ctx;
ast_manager& m;
rule_manager& rm;
context m_inner_ctx;
arith_util a;
obj_map<func_decl, expr*> m_fun2inv;
ast_ref_vector m_pinned;
volatile bool m_cancel;
void get_invariants(rule_set const& src);
void update_body(rule_set& result, rule& r);
rule_set* update_rules(rule_set const& src);
public:
mk_karr_invariants(context & ctx, unsigned priority);
virtual ~mk_karr_invariants();
virtual void cancel();
rule_set * operator()(rule_set const & source);
};
class karr_relation;
class karr_relation_plugin : public relation_plugin {
arith_util a;
hilbert_basis m_hb;
class join_fn;
class project_fn;
class rename_fn;
class union_fn;
class filter_equal_fn;
class filter_identical_fn;
class filter_interpreted_fn;
friend class karr_relation;
public:
karr_relation_plugin(relation_manager& rm):
relation_plugin(karr_relation_plugin::get_name(), rm),
a(get_ast_manager())
{}
virtual bool can_handle_signature(const relation_signature & sig) {
return true;
}
static symbol get_name() { return symbol("karr_relation"); }
virtual void set_cancel(bool f);
virtual relation_base * mk_empty(const relation_signature & s);
virtual relation_base * mk_full(func_decl* p, const relation_signature & s);
static karr_relation& get(relation_base& r);
static karr_relation const & get(relation_base const& r);
virtual relation_join_fn * mk_join_fn(
const relation_base & t1, const relation_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
virtual relation_transformer_fn * mk_project_fn(const relation_base & t, unsigned col_cnt,
const unsigned * removed_cols);
virtual relation_transformer_fn * mk_rename_fn(const relation_base & t, unsigned permutation_cycle_len,
const unsigned * permutation_cycle);
virtual relation_union_fn * mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
virtual relation_mutator_fn * mk_filter_identical_fn(const relation_base & t, unsigned col_cnt,
const unsigned * identical_cols);
virtual relation_mutator_fn * mk_filter_equal_fn(const relation_base & t, const relation_element & value,
unsigned col);
virtual relation_mutator_fn * mk_filter_interpreted_fn(const relation_base & t, app * condition);
private:
bool dualizeI(matrix& dst, matrix const& src);
void dualizeH(matrix& dst, matrix const& src);
};
};
#endif /* _DL_MK_KARR_INVARIANTS_H_ */

View file

@ -0,0 +1,158 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_loop_counter.cpp
Abstract:
Add loop counter argument to relations.
Author:
Nikolaj Bjorner (nbjorner) 2013-03-31
Revision History:
--*/
#include"dl_mk_loop_counter.h"
#include"dl_context.h"
namespace datalog {
mk_loop_counter::mk_loop_counter(context & ctx, unsigned priority):
plugin(priority),
m(ctx.get_manager()),
m_ctx(ctx),
a(m),
m_refs(m) {
}
mk_loop_counter::~mk_loop_counter() { }
app_ref mk_loop_counter::add_arg(rule_set const& src, rule_set& dst, app* fn, unsigned idx) {
expr_ref_vector args(m);
func_decl* new_fn, *old_fn = fn->get_decl();
args.append(fn->get_num_args(), fn->get_args());
args.push_back(m.mk_var(idx, a.mk_int()));
if (!m_old2new.find(old_fn, new_fn)) {
ptr_vector<sort> domain;
domain.append(fn->get_num_args(), old_fn->get_domain());
domain.push_back(a.mk_int());
new_fn = m.mk_func_decl(old_fn->get_name(), domain.size(), domain.c_ptr(), old_fn->get_range());
m_old2new.insert(old_fn, new_fn);
m_new2old.insert(new_fn, old_fn);
m_refs.push_back(new_fn);
m_ctx.register_predicate(new_fn, false);
if (src.is_output_predicate(old_fn)) {
dst.set_output_predicate(new_fn);
}
}
return app_ref(m.mk_app(new_fn, args.size(), args.c_ptr()), m);
}
app_ref mk_loop_counter::del_arg(app* fn) {
expr_ref_vector args(m);
func_decl* old_fn, *new_fn = fn->get_decl();
SASSERT(fn->get_num_args() > 0);
args.append(fn->get_num_args()-1, fn->get_args());
VERIFY (m_new2old.find(new_fn, old_fn));
return app_ref(m.mk_app(old_fn, args.size(), args.c_ptr()), m);
}
rule_set * mk_loop_counter::operator()(rule_set const & source) {
m_refs.reset();
m_old2new.reset();
m_new2old.reset();
rule_manager& rm = source.get_rule_manager();
rule_set * result = alloc(rule_set, m_ctx);
unsigned sz = source.get_num_rules();
rule_ref new_rule(rm);
app_ref_vector tail(m);
app_ref head(m);
svector<bool> neg;
rule_counter& vc = rm.get_counter();
for (unsigned i = 0; i < sz; ++i) {
tail.reset();
neg.reset();
rule & r = *source.get_rule(i);
unsigned cnt = vc.get_max_rule_var(r)+1;
unsigned utsz = r.get_uninterpreted_tail_size();
unsigned tsz = r.get_tail_size();
for (unsigned j = 0; j < utsz; ++j, ++cnt) {
tail.push_back(add_arg(source, *result, r.get_tail(j), cnt));
neg.push_back(r.is_neg_tail(j));
}
for (unsigned j = utsz; j < tsz; ++j) {
tail.push_back(r.get_tail(j));
neg.push_back(false);
}
head = add_arg(source, *result, r.get_head(), cnt);
// set the loop counter to be an increment of the previous
bool found = false;
unsigned last = head->get_num_args()-1;
for (unsigned j = 0; !found && j < utsz; ++j) {
if (head->get_decl() == tail[j]->get_decl()) {
tail.push_back(m.mk_eq(head->get_arg(last),
a.mk_add(tail[j]->get_arg(last),
a.mk_numeral(rational(1), true))));
neg.push_back(false);
found = true;
}
}
// initialize loop counter to 0 if none was found.
if (!found) {
expr_ref_vector args(m);
args.append(head->get_num_args(), head->get_args());
args[last] = a.mk_numeral(rational(0), true);
head = m.mk_app(head->get_decl(), args.size(), args.c_ptr());
}
new_rule = rm.mk(head, tail.size(), tail.c_ptr(), neg.c_ptr(), r.name(), true);
result->add_rule(new_rule);
}
// model converter: remove references to extra argument.
// proof converter: remove references to extra argument as well.
return result;
}
rule_set * mk_loop_counter::revert(rule_set const & source) {
context& ctx = source.get_context();
rule_manager& rm = source.get_rule_manager();
rule_set * result = alloc(rule_set, ctx);
unsigned sz = source.get_num_rules();
rule_ref new_rule(rm);
app_ref_vector tail(m);
app_ref head(m);
svector<bool> neg;
for (unsigned i = 0; i < sz; ++i) {
tail.reset();
neg.reset();
rule & r = *source.get_rule(i);
unsigned utsz = r.get_uninterpreted_tail_size();
unsigned tsz = r.get_tail_size();
for (unsigned j = 0; j < utsz; ++j) {
tail.push_back(del_arg(r.get_tail(j)));
neg.push_back(r.is_neg_tail(j));
}
for (unsigned j = utsz; j < tsz; ++j) {
tail.push_back(r.get_tail(j));
neg.push_back(false);
}
head = del_arg(r.get_head());
new_rule = rm.mk(head, tail.size(), tail.c_ptr(), neg.c_ptr(), r.name(), true);
result->add_rule(new_rule);
}
// model converter: ...
// proof converter: ...
return result;
}
};

View file

@ -0,0 +1,51 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_loop_counter.h
Abstract:
Add loop counter argument to relations.
Author:
Nikolaj Bjorner (nbjorner) 2013-03-31
Revision History:
--*/
#ifndef _DL_MK_LOOP_COUNTER_H_
#define _DL_MK_LOOP_COUNTER_H_
#include"dl_rule_transformer.h"
#include"arith_decl_plugin.h"
namespace datalog {
class mk_loop_counter : public rule_transformer::plugin {
ast_manager& m;
context& m_ctx;
arith_util a;
func_decl_ref_vector m_refs;
obj_map<func_decl, func_decl*> m_new2old;
obj_map<func_decl, func_decl*> m_old2new;
app_ref add_arg(rule_set const& src, rule_set& dst, app* fn, unsigned idx);
app_ref del_arg(app* fn);
public:
mk_loop_counter(context & ctx, unsigned priority = 33000);
~mk_loop_counter();
rule_set * operator()(rule_set const & source);
func_decl* get_old(func_decl* f) const { return m_new2old.find(f); }
rule_set * revert(rule_set const& source);
};
};
#endif /* _DL_MK_LOOP_COUNTER_H_ */

View file

@ -0,0 +1,383 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_magic_sets.cpp
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-10-04.
Revision History:
--*/
#include<utility>
#include<sstream>
#include"ast_pp.h"
#include"dl_mk_magic_sets.h"
namespace datalog {
mk_magic_sets::mk_magic_sets(context & ctx, func_decl* goal) :
plugin(10000, true),
m_context(ctx),
m(ctx.get_manager()),
rm(ctx.get_rule_manager()),
m_pinned(m),
m_goal(goal, m) {
}
void mk_magic_sets::reset() {
m_extentional.reset();
m_todo.reset();
m_adorned_preds.reset();
m_adornments.reset();
m_magic_preds.reset();
m_pinned.reset();
}
void mk_magic_sets::adornment::populate(app * lit, const var_idx_set & bound_vars) {
SASSERT(empty());
unsigned arity = lit->get_num_args();
for (unsigned i = 0; i < arity; i++) {
const expr * arg = lit->get_arg(i);
bool bound = !is_var(arg) || bound_vars.contains(to_var(arg)->get_idx());
push_back(bound ? AD_BOUND : AD_FREE);
}
}
std::string mk_magic_sets::adornment::to_string() const {
std::string res;
const_iterator eit = begin();
const_iterator eend = end();
for (; eit != eend; ++eit) {
res += (*eit == AD_BOUND)?'b':'f';
}
return res;
}
unsigned get_bound_arg_count(app * lit, const var_idx_set & bound_vars) {
unsigned res = 0;
unsigned n = lit->get_num_args();
for (unsigned i = 0; i < n; i++) {
const expr * arg = lit->get_arg(i);
if (!is_var(arg) || bound_vars.contains(to_var(arg)->get_idx())) {
SASSERT(is_var(arg) || is_app(arg));
SASSERT(!is_app(arg) || to_app(arg)->get_num_args()==0);
res++;
}
}
return res;
}
float mk_magic_sets::get_unbound_cost(app * lit, const var_idx_set & bound_vars) {
func_decl * pred = lit->get_decl();
float res = 1;
unsigned n = lit->get_num_args();
for (unsigned i = 0; i < n; i++) {
const expr * arg = lit->get_arg(i);
if (is_var(arg) && !bound_vars.contains(to_var(arg)->get_idx())) {
res *= m_context.get_sort_size_estimate(pred->get_domain(i));
}
//res-=1;
}
return res;
}
/**
\brief From \c cont which is list of indexes of tail literals of rule \c r, select
the index pointing to a literal with at least one bound variable that will be the next
bound literal in the process of creating an adorned rule. If all literals are unbound,
return -1.
*/
int mk_magic_sets::pop_bound(unsigned_vector & cont, rule * r, const var_idx_set & bound_vars) {
float best_cost;
int candidate_index = -1;
unsigned n = cont.size();
for (unsigned i=0; i<n; i++) {
app * lit = r->get_tail(cont[i]);
unsigned bound_cnt = get_bound_arg_count(lit, bound_vars);
if (bound_cnt==0) {
continue;
}
float cost = get_unbound_cost(lit, bound_vars);
if (candidate_index==-1 || cost<best_cost) {
best_cost = cost;
candidate_index = i;
}
}
if (candidate_index==-1) {
return -1;
}
if (candidate_index != static_cast<int>(n-1)) {
std::swap(cont[candidate_index], cont[n-1]);
}
unsigned res = cont.back();
cont.pop_back();
return res;
}
app * mk_magic_sets::adorn_literal(app * lit, const var_idx_set & bound_vars) {
SASSERT(!m_extentional.contains(lit->get_decl()));
func_decl * old_pred = lit->get_decl();
SASSERT(m.is_bool(old_pred->get_range()));
adornment_desc adn(old_pred);
adn.m_adornment.populate(lit, bound_vars);
adornment_map::entry * e = m_adorned_preds.insert_if_not_there2(adn, 0);
func_decl * new_pred = e->get_data().m_value;
if (new_pred==0) {
std::string suffix = "ad_"+adn.m_adornment.to_string();
new_pred = m_context.mk_fresh_head_predicate(
old_pred->get_name(), symbol(suffix.c_str()),
old_pred->get_arity(), old_pred->get_domain(), old_pred);
m_pinned.push_back(new_pred);
e->get_data().m_value = new_pred;
m_todo.push_back(adn);
m_adornments.insert(new_pred, adn.m_adornment);
}
app * res = m.mk_app(new_pred, lit->get_args());
m_pinned.push_back(res);
return res;
}
app * mk_magic_sets::create_magic_literal(app * l) {
func_decl * l_pred = l->get_decl();
SASSERT(m.is_bool(l_pred->get_range()));
pred_adornment_map::obj_map_entry * ae = m_adornments.find_core(l_pred);
SASSERT(ae);
const adornment & adn = ae->get_data().m_value;
unsigned l_arity = l->get_num_args();
ptr_vector<expr> bound_args;
for (unsigned i=0; i<l_arity; i++) {
if (adn[i]==AD_BOUND) {
bound_args.push_back(l->get_arg(i));
}
}
pred2pred::obj_map_entry * e = m_magic_preds.insert_if_not_there2(l_pred, 0);
func_decl * mag_pred = e->get_data().m_value;
if (mag_pred==0) {
unsigned mag_arity = bound_args.size();
ptr_vector<sort> mag_domain;
for (unsigned i=0; i<l_arity; i++) {
if (adn[i]==AD_BOUND) {
mag_domain.push_back(l_pred->get_domain(i));
}
}
mag_pred = m_context.mk_fresh_head_predicate(l_pred->get_name(), symbol("ms"),
mag_arity, mag_domain.c_ptr(), l_pred);
m_pinned.push_back(mag_pred);
e->get_data().m_value = mag_pred;
}
app * res = m.mk_app(mag_pred, bound_args.c_ptr());
m_pinned.push_back(res);
return res;
}
void mk_magic_sets::create_magic_rules(app * head, unsigned tail_cnt, app * const * tail, bool const* negated, rule_set& result) {
//TODO: maybe include relevant interpreted predicates from the original rule
ptr_vector<app> new_tail;
svector<bool> negations;
new_tail.push_back(create_magic_literal(head));
new_tail.append(tail_cnt, tail);
negations.push_back(false);
negations.append(tail_cnt, negated);
for (unsigned i=0; i<tail_cnt; i++) {
if (m_extentional.contains(tail[i]->get_decl())) {
continue;
}
app * mag_head = create_magic_literal(tail[i]);
rule * r = m_context.get_rule_manager().mk(mag_head, i+1, new_tail.c_ptr(), negations.c_ptr());
TRACE("dl", r->display(m_context,tout); );
result.add_rule(r);
}
}
void mk_magic_sets::transform_rule(const adornment & head_adornment, rule * r, rule_set& result) {
app * head = r->get_head();
unsigned head_len = head->get_num_args();
SASSERT(head_len==head_adornment.size());
var_idx_set bound_vars;
for (unsigned i=0; i<head_len; i++) {
expr * arg = head->get_arg(i);
if (head_adornment[i]==AD_BOUND && is_var(arg)) {
bound_vars.insert(to_var(arg)->get_idx());
}
}
unsigned processed_tail_len = r->get_uninterpreted_tail_size();
unsigned_vector exten_tails;
unsigned_vector inten_tails;
for (unsigned i=0; i<processed_tail_len; i++) {
app * t = r->get_tail(i);
if (m_extentional.contains(t->get_decl())) {
exten_tails.push_back(i);
}
else {
inten_tails.push_back(i);
}
}
ptr_vector<app> new_tail;
svector<bool> negations;
while (new_tail.size()!=processed_tail_len) {
bool intentional = false;
int curr_index = pop_bound(exten_tails, r, bound_vars);
if (curr_index==-1) {
curr_index = pop_bound(inten_tails, r,bound_vars);
if (curr_index!=-1) {
intentional = true;
}
}
if (curr_index==-1) {
if (!exten_tails.empty()) {
curr_index = exten_tails.back();
exten_tails.pop_back();
}
else {
SASSERT(!inten_tails.empty());
curr_index = inten_tails.back();
inten_tails.pop_back();
intentional = true;
}
}
SASSERT(curr_index!=-1);
app * curr = r->get_tail(curr_index);
if (intentional) {
curr = adorn_literal(curr, bound_vars);
}
new_tail.push_back(curr);
negations.push_back(r->is_neg_tail(curr_index));
bound_vars |= rm.collect_vars(curr);
}
func_decl * new_head_pred;
VERIFY( m_adorned_preds.find(adornment_desc(head->get_decl(), head_adornment), new_head_pred) );
app * new_head = m.mk_app(new_head_pred, head->get_args());
SASSERT(new_tail.size()==r->get_uninterpreted_tail_size());
create_magic_rules(new_head, new_tail.size(), new_tail.c_ptr(), negations.c_ptr(), result);
unsigned tail_len = r->get_tail_size();
for (unsigned i=processed_tail_len; i<tail_len; i++) {
new_tail.push_back(r->get_tail(i));
negations.push_back(r->is_neg_tail(i));
}
new_tail.push_back(create_magic_literal(new_head));
negations.push_back(false);
rule * nr = m_context.get_rule_manager().mk(new_head, new_tail.size(), new_tail.c_ptr(), negations.c_ptr());
result.add_rule(nr);
nr->set_accounting_parent_object(m_context, r);
}
void mk_magic_sets::create_transfer_rule(const adornment_desc & d, rule_set& result) {
func_decl * adn_pred = m_adorned_preds.find(d);
unsigned arity = adn_pred->get_arity();
SASSERT(arity == d.m_pred->get_arity());
ptr_vector<expr> args;
for (unsigned i=0; i<arity; i++) {
args.push_back(m.mk_var(i, adn_pred->get_domain(i)));
}
app * lit = m.mk_app(d.m_pred, args.c_ptr());
app * adn_lit = m.mk_app(adn_pred, args.c_ptr());
app * mag_lit = create_magic_literal(adn_lit);
app * tail[] = {lit, mag_lit};
rule * r = m_context.get_rule_manager().mk(adn_lit, 2, tail, 0);
result.add_rule(r);
}
rule_set * mk_magic_sets::operator()(rule_set const & source) {
if (!m_context.magic_sets_for_queries()) {
return 0;
}
SASSERT(source.contains(m_goal));
SASSERT(source.get_predicate_rules(m_goal).size() == 1);
app * goal_head = source.get_predicate_rules(m_goal)[0]->get_head();
unsigned init_rule_cnt = source.get_num_rules();
{
func_decl_set intentional;
for (unsigned i=0; i<init_rule_cnt; i++) {
func_decl* pred = source.get_rule(i)->get_decl();
intentional.insert(pred);
}
//now we iterate through all predicates and collect the set of extentional ones
const rule_dependencies * deps;
rule_dependencies computed_deps(m_context);
if (source.is_closed()) {
deps = &source.get_dependencies();
}
else {
computed_deps.populate(source);
deps = &computed_deps;
}
rule_dependencies::iterator it = deps->begin();
rule_dependencies::iterator end = deps->end();
for (; it!=end; ++it) {
func_decl * pred = it->m_key;
if (intentional.contains(pred)) {
continue;
}
SASSERT(it->m_value->empty());//extentional predicates have no dependency
m_extentional.insert(pred);
}
}
//adornment goal_adn;
//goal_adn.populate(goal_head, );
var_idx_set empty_var_idx_set;
adorn_literal(goal_head, empty_var_idx_set);
rule_set * result = alloc(rule_set, m_context);
result->inherit_predicates(source);
while (!m_todo.empty()) {
adornment_desc task = m_todo.back();
m_todo.pop_back();
const rule_vector & pred_rules = source.get_predicate_rules(task.m_pred);
rule_vector::const_iterator it = pred_rules.begin();
rule_vector::const_iterator end = pred_rules.end();
for (; it != end; ++it) {
rule * r = *it;
transform_rule(task.m_adornment, r, *result);
}
if (!m_context.get_rel_context()->get_relation(task.m_pred).empty()) {
//we need a rule to copy facts that are already in a relation into the adorned
//relation (since out intentional predicates can have facts, not only rules)
create_transfer_rule(task, *result);
}
}
app * adn_goal_head = adorn_literal(goal_head, empty_var_idx_set);
app * mag_goal_head = create_magic_literal(adn_goal_head);
SASSERT(mag_goal_head->is_ground());
rule * mag_goal_rule = m_context.get_rule_manager().mk(mag_goal_head, 0, 0, 0);
result->add_rule(mag_goal_rule);
rule * back_to_goal_rule = m_context.get_rule_manager().mk(goal_head, 1, &adn_goal_head, 0);
result->add_rule(back_to_goal_rule);
return result;
}
};

135
src/muz/dl_mk_magic_sets.h Normal file
View file

@ -0,0 +1,135 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_magic_sets.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-10-4.
Revision History:
--*/
#ifndef _DL_MK_MAGIC_SETS_H_
#define _DL_MK_MAGIC_SETS_H_
#include<utility>
#include"map.h"
#include"obj_pair_hashtable.h"
#include"dl_context.h"
#include"dl_rule_set.h"
#include"dl_rule_transformer.h"
namespace datalog {
/**
\brief Implements magic sets rule transformation.
According to A. Voronkov. Foundations of Deductive Databases.
The stratified negation is not in the book addressed wrt. magic sets, but it seems
that, for the purpose of magic sets, the negated literals should be treated just as
if they were non-negated (we are interested only in values of arguments, not in the
actual content of relations, at that point).
*/
class mk_magic_sets : public rule_transformer::plugin {
enum a_flag {
AD_FREE,
AD_BOUND
};
struct a_flag_hash {
typedef a_flag data;
unsigned operator()(a_flag x) const { return x; }
};
struct adornment : public svector<a_flag> {
void populate(app * lit, const var_idx_set & bound_vars);
bool operator==(const adornment & o) const {
return vectors_equal(*this, o);
}
std::string to_string() const;
};
struct adornment_desc {
func_decl * m_pred;
adornment m_adornment;
adornment_desc() {}
adornment_desc(func_decl * pred) : m_pred(pred) {}
adornment_desc(func_decl * pred, const adornment & a)
: m_pred(pred), m_adornment(a) {}
bool operator==(const adornment_desc & o) const {
//m_tail_adornment value is implied by the rule and the head adornment
return m_pred==o.m_pred && m_adornment==o.m_adornment;
}
unsigned hash() const {
return m_pred->hash()^svector_hash<a_flag_hash>()(m_adornment);
}
};
struct adorned_rule {
app * m_head;
adornment m_head_adornment;
ptr_vector<app> m_tail;
};
typedef hashtable<adornment_desc, obj_hash<adornment_desc>,
default_eq<adornment_desc> > adornment_set;
typedef map<adornment_desc, func_decl *, obj_hash<adornment_desc>,
default_eq<adornment_desc> > adornment_map;
typedef obj_map<func_decl, adornment> pred_adornment_map;
typedef obj_map<func_decl, func_decl *> pred2pred;
context & m_context;
ast_manager & m;
rule_manager& rm;
ast_ref_vector m_pinned;
/**
\brief Predicates from the original set that appear in a head of a rule
*/
func_decl_set m_extentional;
//adornment_set m_processed;
vector<adornment_desc> m_todo;
adornment_map m_adorned_preds;
pred_adornment_map m_adornments;
pred2pred m_magic_preds;
func_decl_ref m_goal;
void reset();
float get_unbound_cost(app * lit, const var_idx_set & bound_vars);
int pop_bound(unsigned_vector & cont, rule * r, const var_idx_set & bound_vars);
app * create_magic_literal(app * l);
void create_magic_rules(app * head, unsigned tail_cnt, app * const * tail, bool const* negated, rule_set& result);
app * adorn_literal(app * lit, const var_idx_set & bound_vars);
void transform_rule(const adornment & head_adornment, rule * r, rule_set& result);
void create_transfer_rule(const adornment_desc & d, rule_set& result);
public:
/**
\brief Create magic sets rule transformer for \c goal_rule. When applying the transformer,
the \c goal_rule must be present in the \c rule_set that is being transformed.
*/
mk_magic_sets(context & ctx, func_decl* goal);
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_MAGIC_SETS_H_ */

View file

@ -0,0 +1,135 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_magic_symbolic.cpp
Abstract:
Create Horn clauses for magic symbolic flow.
Q(x) :- A(y), B(z), phi1(x,y,z).
Q(x) :- C(y), phi2(x,y).
A(x) :- C(y), phi3(x,y).
A(x) :- A(y), phi3(x,y).
B(x) :- C(y), A(z), phi4(x,y,z).
C(x) :- phi5(x).
Transformed clauses:
Q_ans(x) :- Q_query(x), A_ans(y), B_ans(z), phi1(x,y,z).
Q_ans(x) :- Q_query(x), C_ans(y), phi2(x,y).
Q_query(x) :- true.
A_ans(x) :- A_query(x), C_ans(y), phi2(x,y)
A_ans(x) :- A_query(x), A_ans(y), phi3(x,y).
A_query(y) :- Q_query(x), phi1(x,y,z).
A_query(y) :- A_query(x), phi3(x,y).
A_query(z) :- B_query(x), C_ans(y), phi4(x,y,z).
B_ans(x) :- B_query(x), C_ans(y), A_ans(z), phi4(x,y,z).
B_query(z) :- Q_query(x), A_ans(y), phi1(x,y,z).
C_ans(x) :- C_query(x), phi5(x).
C_query(y) :- Q_query(x), phi2(x,y).
C_query(y) :- Q_query(x), phi3(x,y).
C_query(y) :- B_query(x), phi4(x,y,z).
General scheme:
A(x) :- P1(x_1), ..., Pn(x_n), phi(x,x1,..,x_n).
P(x) :- Prefix(x,y,z), A(z) ...
A_ans(x) :- A_query(x), P_i_ans(x_i), phi(x,..).
A_query(z) :- P_query(x), Prefix_ans(x,y,z).
Author:
Nikolaj Bjorner (nbjorner) 2013-06-19
Revision History:
--*/
#include"dl_mk_magic_symbolic.h"
#include"dl_context.h"
namespace datalog {
mk_magic_symbolic::mk_magic_symbolic(context & ctx, unsigned priority):
plugin(priority),
m(ctx.get_manager()),
m_ctx(ctx) {
}
mk_magic_symbolic::~mk_magic_symbolic() { }
rule_set * mk_magic_symbolic::operator()(rule_set const & source) {
if (!m_ctx.get_params().magic()) {
return 0;
}
context& ctx = source.get_context();
rule_manager& rm = source.get_rule_manager();
rule_set * result = alloc(rule_set, ctx);
unsigned sz = source.get_num_rules();
rule_ref new_rule(rm);
app_ref_vector tail(m);
app_ref head(m);
svector<bool> neg;
for (unsigned i = 0; i < sz; ++i) {
rule & r = *source.get_rule(i);
unsigned utsz = r.get_uninterpreted_tail_size();
unsigned tsz = r.get_tail_size();
tail.reset();
neg.reset();
for (unsigned j = utsz; j < tsz; ++j) {
tail.push_back(r.get_tail(j));
neg.push_back(false);
}
tail.push_back(mk_query(r.get_head()));
neg.push_back(false);
for (unsigned j = 0; j < utsz; ++j) {
tail.push_back(mk_ans(r.get_tail(j)));
neg.push_back(false);
}
new_rule = rm.mk(mk_ans(r.get_head()), tail.size(), tail.c_ptr(), neg.c_ptr(), r.name(), true);
result->add_rule(new_rule);
if (source.is_output_predicate(r.get_decl())) {
result->set_output_predicate(new_rule->get_decl());
new_rule = rm.mk(mk_query(r.get_head()), 0, 0, 0, r.name(), true);
result->add_rule(new_rule);
}
for (unsigned j = 0; j < utsz; ++j) {
new_rule = rm.mk(mk_query(r.get_tail(j)), tail.size()-utsz+j, tail.c_ptr(), neg.c_ptr(), r.name(), true);
result->add_rule(new_rule);
}
}
TRACE("dl", result->display(tout););
return result;
}
app_ref mk_magic_symbolic::mk_query(app* q) {
string_buffer<64> name;
func_decl* f = q->get_decl();
name << f->get_name() << "!query";
func_decl_ref g(m);
g = m.mk_func_decl(symbol(name.c_str()), f->get_arity(), f->get_domain(), f->get_range());
m_ctx.register_predicate(g, false);
return app_ref(m.mk_app(g, q->get_num_args(), q->get_args()), m);
}
app_ref mk_magic_symbolic::mk_ans(app* q) {
string_buffer<64> name;
func_decl* f = q->get_decl();
func_decl_ref g(m);
name << f->get_name() << "!ans";
g = m.mk_func_decl(symbol(name.c_str()), f->get_arity(), f->get_domain(), f->get_range());
m_ctx.register_predicate(g, false);
return app_ref(m.mk_app(g, q->get_num_args(), q->get_args()), m);
}
};

View file

@ -0,0 +1,40 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_magic_symbolic.h
Abstract:
Create Horn clauses for magic symbolic transformation.
Author:
Nikolaj Bjorner (nbjorner) 2013-06-19
Revision History:
--*/
#ifndef _DL_MK_MAGIC_SYMBOLIC_H_
#define _DL_MK_MAGIC_SYMBOLIC_H_
#include"dl_rule_transformer.h"
namespace datalog {
class mk_magic_symbolic : public rule_transformer::plugin {
ast_manager& m;
context& m_ctx;
app_ref mk_ans(app* q);
app_ref mk_query(app* q);
public:
mk_magic_symbolic(context & ctx, unsigned priority = 33037);
~mk_magic_symbolic();
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_MAGIC_SYMBOLIC_H_ */

View file

@ -0,0 +1,154 @@
/*++
Copyright (c) 2012 Microsoft Corporation
Module Name:
dl_mk_partial_equiv.cpp
Abstract:
Rule transformer which identifies predicates that are partial equivalence relations.
Author:
Nikolaj Bjorner (nbjorner) 2012-05-14
Revision History:
--*/
#include "dl_mk_partial_equiv.h"
#include "ast_pp.h"
namespace datalog {
bool mk_partial_equivalence_transformer::is_symmetry(rule const* r) {
func_decl* p = r->get_decl();
return
p->get_arity() == 2 &&
p->get_domain(0) == p->get_domain(1) &&
r->get_tail_size() == 1 &&
r->get_tail(0)->get_decl() == p &&
r->get_head()->get_arg(0) == r->get_tail(0)->get_arg(1) &&
r->get_head()->get_arg(1) == r->get_tail(0)->get_arg(0) &&
is_var(r->get_head()->get_arg(0)) &&
is_var(r->get_head()->get_arg(1)) &&
r->get_head()->get_arg(0) != r->get_head()->get_arg(1);
}
bool mk_partial_equivalence_transformer::is_transitivity(rule const* r) {
func_decl* p = r->get_decl();
if (p->get_arity() != 2 ||
p->get_domain(0) != p->get_domain(1) ||
r->get_tail_size() != 2 ||
r->get_tail(0)->get_decl() != p ||
r->get_tail(1)->get_decl() != p) {
return false;
}
app* h = r->get_head();
app* a = r->get_tail(0);
app* b = r->get_tail(1);
expr* x1 = h->get_arg(0);
expr* x2 = h->get_arg(1);
expr* a1 = a->get_arg(0);
expr* a2 = a->get_arg(1);
expr* b1 = b->get_arg(0);
expr* b2 = b->get_arg(1);
if (!(is_var(x1) && is_var(x2) && is_var(a1) && is_var(a2) && is_var(b1) && is_var(b2))) {
return false;
}
if (x1 == x2 || a1 == a2 || b1 == b2) {
return false;
}
if (a2 == b1) {
if (x1 == b2 && x2 == a1) {
return true;
}
if (x1 == a1 && x2 == b2) {
return true;
}
return false;
}
if (a1 == b2) {
if (x1 == b1 && x2 == a2) {
return true;
}
if (x1 == a2 && x2 == b1) {
return true;
}
return false;
}
return false;
;
}
rule_set * mk_partial_equivalence_transformer::operator()(rule_set const & source) {
// TODO mc
if (source.get_num_rules() == 0) {
return 0;
}
if (m_context.get_engine() != DATALOG_ENGINE) {
return 0;
}
relation_manager & rm = m_context.get_rel_context()->get_rmanager();
rule_set::decl2rules::iterator it = source.begin_grouped_rules();
rule_set::decl2rules::iterator end = source.end_grouped_rules();
rule_set* res = alloc(rule_set, m_context);
for (; it != end; ++it) {
func_decl* p = it->m_key;
rule_vector const& rv = *(it->m_value);
bool has_symmetry = false;
bool has_transitivity = false;
unsigned i_symmetry, i_transitivity;
family_id kind = rm.get_requested_predicate_kind(p);
for (unsigned i = 0; i < rv.size(); ++i) {
if (kind != null_family_id) {
res->add_rule(rv[i]);
}
else if (is_symmetry(rv[i])) {
i_symmetry = i;
has_symmetry = true;
}
else if (is_transitivity(rv[i])) {
i_transitivity = i;
has_transitivity = true;
}
else {
res->add_rule(rv[i]);
}
}
if (has_symmetry && !has_transitivity) {
res->add_rule(rv[i_symmetry]);
}
else if (!has_symmetry && has_transitivity) {
res->add_rule(rv[i_transitivity]);
}
else if (has_symmetry && has_transitivity) {
TRACE("dl", tout << "updating predicate " << mk_pp(p, m) << " to partial equivalence\n";);
SASSERT(kind == null_family_id);
rm.set_predicate_kind(p, rm.get_table_plugin(symbol("equivalence"))->get_kind());
}
}
if (res->get_num_rules() == source.get_num_rules()) {
dealloc(res);
return 0;
}
res->inherit_predicates(source);
return res;
}
};

View file

@ -0,0 +1,50 @@
/*++
Copyright (c) 2012 Microsoft Corporation
Module Name:
dl_mk_partial_equiv.h
Abstract:
Rule transformer which identifies predicates that are partial equivalence relations.
Author:
Nikolaj Bjorner (nbjorner) 2012-05-14
Revision History:
--*/
#ifndef _DL_MK_PARTIAL_EQUIVALENCE_TRANSFORMER_H_
#define _DL_MK_PARTIAL_EQUIVALENCE_TRANSFORMER_H_
#include "dl_context.h"
#include "dl_rule_transformer.h"
namespace datalog {
class mk_partial_equivalence_transformer : public rule_transformer::plugin {
ast_manager & m;
context & m_context;
public:
mk_partial_equivalence_transformer(context & ctx, unsigned priority=45000)
: plugin(priority),
m(ctx.get_manager()),
m_context(ctx) {}
rule_set * operator()(rule_set const & source);
private:
bool is_symmetry(rule const* r);
bool is_transitivity(rule const* r);
};
};
#endif /* _DL_MK_PARTIAL_EQUIV_TRANSFORMER_H_ */

View file

@ -0,0 +1,367 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_quantifier_abstraction.cpp
Abstract:
Create quantified Horn clauses from benchmarks with arrays.
Author:
Ken McMillan
Andrey Rybalchenko
Nikolaj Bjorner (nbjorner) 2013-04-02
Revision History:
--*/
#include "dl_mk_quantifier_abstraction.h"
#include "dl_context.h"
#include "expr_safe_replace.h"
#include "expr_abstract.h"
namespace datalog {
// model converter:
// Given model for P^(x, y, i, a[i])
// create model: P(x,y,a) == forall i . P^(x,y,i,a[i])
// requires substitution and list of bound variables.
class mk_quantifier_abstraction::qa_model_converter : public model_converter {
ast_manager& m;
func_decl_ref_vector m_old_funcs;
func_decl_ref_vector m_new_funcs;
vector<expr_ref_vector> m_subst;
vector<sort_ref_vector> m_sorts;
vector<svector<bool> > m_bound;
public:
qa_model_converter(ast_manager& m):
m(m), m_old_funcs(m), m_new_funcs(m) {}
virtual ~qa_model_converter() {}
virtual model_converter * translate(ast_translation & translator) {
return alloc(qa_model_converter, m);
}
void insert(func_decl* old_p, func_decl* new_p, expr_ref_vector& sub, sort_ref_vector& sorts, svector<bool> const& bound) {
m_old_funcs.push_back(old_p);
m_new_funcs.push_back(new_p);
m_subst.push_back(sub);
m_bound.push_back(bound);
m_sorts.push_back(sorts);
}
virtual void operator()(model_ref & old_model) {
model_ref new_model = alloc(model, m);
for (unsigned i = 0; i < m_new_funcs.size(); ++i) {
func_decl* p = m_new_funcs[i].get();
func_decl* q = m_old_funcs[i].get();
expr_ref_vector const& sub = m_subst[i];
sort_ref_vector const& sorts = m_sorts[i];
svector<bool> const& is_bound = m_bound[i];
func_interp* f = old_model->get_func_interp(p);
expr_ref body(m);
unsigned arity_p = p->get_arity();
unsigned arity_q = q->get_arity();
SASSERT(0 < arity_p);
func_interp* g = alloc(func_interp, m, arity_q);
if (f) {
body = f->get_interp();
SASSERT(!f->is_partial());
SASSERT(body);
}
else {
body = m.mk_false();
}
// Create quantifier wrapper around body.
TRACE("dl", tout << mk_pp(body, m) << "\n";);
// 1. replace variables by the compound terms from
// the original predicate.
expr_safe_replace rep(m);
for (unsigned i = 0; i < sub.size(); ++i) {
rep.insert(m.mk_var(i, m.get_sort(sub[i])), sub[i]);
}
rep(body);
rep.reset();
TRACE("dl", tout << mk_pp(body, m) << "\n";);
// 2. replace bound variables by constants.
expr_ref_vector consts(m), bound(m), free(m);
svector<symbol> names;
ptr_vector<sort> bound_sorts;
for (unsigned i = 0; i < sorts.size(); ++i) {
sort* s = sorts[i];
consts.push_back(m.mk_fresh_const("C", s));
rep.insert(m.mk_var(i, s), consts.back());
if (is_bound[i]) {
bound.push_back(consts.back());
names.push_back(symbol(i));
bound_sorts.push_back(s);
}
else {
free.push_back(consts.back());
}
}
rep(body);
rep.reset();
TRACE("dl", tout << mk_pp(body, m) << "\n";);
// 3. abstract and quantify those variables that should be bound.
expr_abstract(m, 0, bound.size(), bound.c_ptr(), body, body);
body = m.mk_forall(names.size(), bound_sorts.c_ptr(), names.c_ptr(), body);
TRACE("dl", tout << mk_pp(body, m) << "\n";);
// 4. replace remaining constants by variables.
for (unsigned i = 0; i < free.size(); ++i) {
rep.insert(free[i].get(), m.mk_var(i, m.get_sort(free[i].get())));
}
rep(body);
g->set_else(body);
TRACE("dl", tout << mk_pp(body, m) << "\n";);
new_model->register_decl(q, g);
}
old_model = new_model;
}
};
mk_quantifier_abstraction::mk_quantifier_abstraction(
context & ctx, unsigned priority):
plugin(priority),
m(ctx.get_manager()),
m_ctx(ctx),
a(m),
m_refs(m) {
}
mk_quantifier_abstraction::~mk_quantifier_abstraction() {
}
func_decl* mk_quantifier_abstraction::declare_pred(rule_set const& rules, rule_set& dst, func_decl* old_p) {
if (rules.is_output_predicate(old_p)) {
dst.inherit_predicate(rules, old_p, old_p);
return 0;
}
unsigned sz = old_p->get_arity();
unsigned num_arrays = 0;
for (unsigned i = 0; i < sz; ++i) {
if (a.is_array(old_p->get_domain(i))) {
num_arrays++;
}
}
if (num_arrays == 0) {
return 0;
}
func_decl* new_p = 0;
if (!m_old2new.find(old_p, new_p)) {
expr_ref_vector sub(m), vars(m);
svector<bool> bound;
sort_ref_vector domain(m), sorts(m);
expr_ref arg(m);
for (unsigned i = 0; i < sz; ++i) {
sort* s0 = old_p->get_domain(i);
unsigned lookahead = 0;
sort* s = s0;
while (a.is_array(s)) {
lookahead += get_array_arity(s);
s = get_array_range(s);
}
arg = m.mk_var(bound.size() + lookahead, s0);
s = s0;
while (a.is_array(s)) {
unsigned arity = get_array_arity(s);
expr_ref_vector args(m);
for (unsigned j = 0; j < arity; ++j) {
sort* s1 = get_array_domain(s, j);
domain.push_back(s1);
args.push_back(m.mk_var(bound.size(), s1));
bound.push_back(true);
sorts.push_back(s1);
}
arg = mk_select(arg, args.size(), args.c_ptr());
s = get_array_range(s);
}
domain.push_back(s);
bound.push_back(false);
sub.push_back(arg);
sorts.push_back(s0);
}
SASSERT(old_p->get_range() == m.mk_bool_sort());
new_p = m.mk_func_decl(old_p->get_name(), domain.size(), domain.c_ptr(), old_p->get_range());
m_refs.push_back(new_p);
m_ctx.register_predicate(new_p, false);
if (m_mc) {
m_mc->insert(old_p, new_p, sub, sorts, bound);
}
m_old2new.insert(old_p, new_p);
}
return new_p;
}
app_ref mk_quantifier_abstraction::mk_head(rule_set const& rules, rule_set& dst, app* p, unsigned idx) {
func_decl* new_p = declare_pred(rules, dst, p->get_decl());
if (!new_p) {
return app_ref(p, m);
}
expr_ref_vector args(m);
expr_ref arg(m);
unsigned sz = p->get_num_args();
for (unsigned i = 0; i < sz; ++i) {
arg = p->get_arg(i);
sort* s = m.get_sort(arg);
while (a.is_array(s)) {
unsigned arity = get_array_arity(s);
for (unsigned j = 0; j < arity; ++j) {
args.push_back(m.mk_var(idx++, get_array_domain(s, j)));
}
arg = mk_select(arg, arity, args.c_ptr()+args.size()-arity);
s = get_array_range(s);
}
args.push_back(arg);
}
TRACE("dl",
tout << mk_pp(new_p, m) << "\n";
for (unsigned i = 0; i < args.size(); ++i) {
tout << mk_pp(args[i].get(), m) << "\n";
});
return app_ref(m.mk_app(new_p, args.size(), args.c_ptr()), m);
}
app_ref mk_quantifier_abstraction::mk_tail(rule_set const& rules, rule_set& dst, app* p) {
func_decl* old_p = p->get_decl();
func_decl* new_p = declare_pred(rules, dst, old_p);
if (!new_p) {
return app_ref(p, m);
}
SASSERT(new_p->get_arity() > old_p->get_arity());
unsigned num_extra_args = new_p->get_arity() - old_p->get_arity();
var_shifter shift(m);
expr_ref p_shifted(m);
shift(p, num_extra_args, p_shifted);
app* ps = to_app(p_shifted);
expr_ref_vector args(m);
app_ref_vector pats(m);
sort_ref_vector vars(m);
svector<symbol> names;
expr_ref arg(m);
unsigned idx = 0;
unsigned sz = p->get_num_args();
for (unsigned i = 0; i < sz; ++i) {
arg = ps->get_arg(i);
sort* s = m.get_sort(arg);
bool is_pattern = false;
while (a.is_array(s)) {
is_pattern = true;
unsigned arity = get_array_arity(s);
for (unsigned j = 0; j < arity; ++j) {
vars.push_back(get_array_domain(s, j));
names.push_back(symbol(idx));
args.push_back(m.mk_var(idx++, vars.back()));
}
arg = mk_select(arg, arity, args.c_ptr()+args.size()-arity);
s = get_array_range(s);
}
if (is_pattern) {
pats.push_back(to_app(arg));
}
args.push_back(arg);
}
expr* pat = 0;
expr_ref pattern(m);
pattern = m.mk_pattern(pats.size(), pats.c_ptr());
pat = pattern.get();
app_ref result(m);
symbol qid, skid;
result = m.mk_app(new_p, args.size(), args.c_ptr());
result = m.mk_eq(m.mk_forall(vars.size(), vars.c_ptr(), names.c_ptr(), result, 1, qid, skid, 1, &pat), m.mk_true());
return result;
}
expr * mk_quantifier_abstraction::mk_select(expr* arg, unsigned num_args, expr* const* args) {
ptr_vector<expr> args2;
args2.push_back(arg);
args2.append(num_args, args);
return a.mk_select(args2.size(), args2.c_ptr());
}
rule_set * mk_quantifier_abstraction::operator()(rule_set const & source) {
TRACE("dl", tout << "quantify " << source.get_num_rules() << " " << m_ctx.get_params().quantify_arrays() << "\n";);
if (!m_ctx.get_params().quantify_arrays()) {
return 0;
}
unsigned sz = source.get_num_rules();
for (unsigned i = 0; i < sz; ++i) {
rule& r = *source.get_rule(i);
if (r.has_negation()) {
return 0;
}
}
m_refs.reset();
m_old2new.reset();
m_new2old.reset();
rule_manager& rm = source.get_rule_manager();
rule_ref new_rule(rm);
expr_ref_vector tail(m);
app_ref head(m);
expr_ref fml(m);
rule_counter& vc = rm.get_counter();
if (m_ctx.get_model_converter()) {
m_mc = alloc(qa_model_converter, m);
}
rule_set * result = alloc(rule_set, m_ctx);
for (unsigned i = 0; i < sz; ++i) {
tail.reset();
rule & r = *source.get_rule(i);
TRACE("dl", r.display(m_ctx, tout); );
unsigned cnt = vc.get_max_rule_var(r)+1;
unsigned utsz = r.get_uninterpreted_tail_size();
unsigned tsz = r.get_tail_size();
for (unsigned j = 0; j < utsz; ++j) {
tail.push_back(mk_tail(source, *result, r.get_tail(j)));
}
for (unsigned j = utsz; j < tsz; ++j) {
tail.push_back(r.get_tail(j));
}
head = mk_head(source, *result, r.get_head(), cnt);
fml = m.mk_implies(m.mk_and(tail.size(), tail.c_ptr()), head);
rule_ref_vector added_rules(rm);
proof_ref pr(m);
rm.mk_rule(fml, pr, *result);
TRACE("dl", result->last()->display(m_ctx, tout););
}
// proof converter: proofs are not necessarily preserved using this transformation.
if (m_old2new.empty()) {
dealloc(result);
dealloc(m_mc);
result = 0;
}
else {
m_ctx.add_model_converter(m_mc);
}
m_mc = 0;
return result;
}
};

View file

@ -0,0 +1,64 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_quantifier_abstraction.h
Abstract:
Convert clauses with array arguments to predicates
into Quantified Horn clauses.
Author:
Ken McMillan
Andrey Rybalchenko
Nikolaj Bjorner (nbjorner) 2013-04-02
Revision History:
Based on approach suggested in SAS 2013 paper
"On Solving Universally Quantified Horn Clauses"
--*/
#ifndef _DL_MK_QUANTIFIER_ABSTRACTION_H_
#define _DL_MK_QUANTIFIER_ABSTRACTION_H_
#include"dl_rule_transformer.h"
#include"array_decl_plugin.h"
namespace datalog {
class context;
class mk_quantifier_abstraction : public rule_transformer::plugin {
class qa_model_converter;
ast_manager& m;
context& m_ctx;
array_util a;
func_decl_ref_vector m_refs;
obj_map<func_decl, func_decl*> m_new2old;
obj_map<func_decl, func_decl*> m_old2new;
qa_model_converter* m_mc;
func_decl* declare_pred(rule_set const& rules, rule_set& dst, func_decl* old_p);
app_ref mk_head(rule_set const& rules, rule_set& dst, app* p, unsigned idx);
app_ref mk_tail(rule_set const& rules, rule_set& dst, app* p);
expr* mk_select(expr* a, unsigned num_args, expr* const* args);
public:
mk_quantifier_abstraction(context & ctx, unsigned priority);
virtual ~mk_quantifier_abstraction();
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_QUANTIFIER_ABSTRACTION_H_ */

View file

@ -0,0 +1,302 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_quantifier_instantiation.cpp
Abstract:
Convert Quantified Horn clauses into non-quantified clauses using
instantiation.
Author:
Ken McMillan
Andrey Rybalchenko
Nikolaj Bjorner (nbjorner) 2013-04-02
Revision History:
Based on approach suggested in the SAS 2013 paper
"On Solving Universally Quantified Horn Clauses"
--*/
#include "dl_mk_quantifier_instantiation.h"
#include "dl_context.h"
#include "pattern_inference.h"
namespace datalog {
mk_quantifier_instantiation::mk_quantifier_instantiation(
context & ctx, unsigned priority):
plugin(priority),
m(ctx.get_manager()),
m_ctx(ctx),
m_var2cnst(m),
m_cnst2var(m) {
}
mk_quantifier_instantiation::~mk_quantifier_instantiation() {
}
void mk_quantifier_instantiation::extract_quantifiers(rule& r, expr_ref_vector& conjs, quantifier_ref_vector& qs) {
conjs.reset();
qs.reset();
unsigned tsz = r.get_tail_size();
for (unsigned j = 0; j < tsz; ++j) {
conjs.push_back(r.get_tail(j));
}
qe::flatten_and(conjs);
for (unsigned j = 0; j < conjs.size(); ++j) {
expr* e = conjs[j].get();
quantifier* q;
if (rule_manager::is_forall(m, e, q)) {
qs.push_back(q);
conjs[j] = conjs.back();
conjs.pop_back();
--j;
}
}
}
void mk_quantifier_instantiation::instantiate_quantifier(quantifier* q, expr_ref_vector & conjs) {
expr_ref qe(m);
qe = q;
m_var2cnst(qe);
q = to_quantifier(qe);
if (q->get_num_patterns() == 0) {
proof_ref new_pr(m);
pattern_inference_params params;
pattern_inference infer(m, params);
infer(q, qe, new_pr);
q = to_quantifier(qe);
}
unsigned num_patterns = q->get_num_patterns();
for (unsigned i = 0; i < num_patterns; ++i) {
expr * pat = q->get_pattern(i);
SASSERT(m.is_pattern(pat));
instantiate_quantifier(q, to_app(pat), conjs);
}
}
void mk_quantifier_instantiation::instantiate_quantifier(quantifier* q, app* pat, expr_ref_vector & conjs) {
m_binding.reset();
m_binding.resize(q->get_num_decls());
term_pairs todo;
match(0, pat, 0, todo, q, conjs);
}
void mk_quantifier_instantiation::match(unsigned i, app* pat, unsigned j, term_pairs& todo, quantifier* q, expr_ref_vector& conjs) {
TRACE("dl", tout << "match" << mk_pp(pat, m) << "\n";);
while (j < todo.size()) {
expr* p = todo[j].first;
expr* t = todo[j].second;
if (is_var(p)) {
unsigned idx = to_var(p)->get_idx();
if (!m_binding[idx]) {
m_binding[idx] = t;
match(i, pat, j + 1, todo, q, conjs);
m_binding[idx] = 0;
return;
}
++j;
continue;
}
if (!is_app(p)) {
return;
}
app* a1 = to_app(p);
unsigned id = t->get_id();
unsigned next_id = id;
unsigned sz = todo.size();
do {
expr* t2 = m_terms[next_id];
if (is_app(t2)) {
app* a2 = to_app(t2);
if (a1->get_decl() == a2->get_decl() &&
a1->get_num_args() == a2->get_num_args()) {
for (unsigned k = 0; k < a1->get_num_args(); ++k) {
todo.push_back(std::make_pair(a1->get_arg(k), a2->get_arg(k)));
}
match(i, pat, j + 1, todo, q, conjs);
todo.resize(sz);
}
}
next_id = m_uf.next(next_id);
}
while (next_id != id);
return;
}
if (i == pat->get_num_args()) {
yield_binding(q, conjs);
return;
}
expr* arg = pat->get_arg(i);
ptr_vector<expr>* terms = 0;
if (m_funs.find(to_app(arg)->get_decl(), terms)) {
for (unsigned k = 0; k < terms->size(); ++k) {
todo.push_back(std::make_pair(arg, (*terms)[k]));
match(i + 1, pat, j, todo, q, conjs);
todo.pop_back();
}
}
}
void mk_quantifier_instantiation::yield_binding(quantifier* q, expr_ref_vector& conjs) {
DEBUG_CODE(
for (unsigned i = 0; i < m_binding.size(); ++i) {
SASSERT(m_binding[i]);
});
m_binding.reverse();
expr_ref res(m);
instantiate(m, q, m_binding.c_ptr(), res);
m_binding.reverse();
m_cnst2var(res);
conjs.push_back(res);
TRACE("dl", tout << mk_pp(q, m) << "\n==>\n" << mk_pp(res, m) << "\n";);
}
void mk_quantifier_instantiation::collect_egraph(expr* e) {
expr* e1, *e2;
m_todo.push_back(e);
expr_fast_mark1 visited;
while (!m_todo.empty()) {
e = m_todo.back();
m_todo.pop_back();
if (visited.is_marked(e)) {
continue;
}
unsigned n = e->get_id();
if (n >= m_terms.size()) {
m_terms.resize(n+1);
}
m_terms[n] = e;
visited.mark(e);
if (m.is_eq(e, e1, e2) || m.is_iff(e, e1, e2)) {
m_uf.merge(e1->get_id(), e2->get_id());
}
if (is_app(e)) {
app* ap = to_app(e);
ptr_vector<expr>* terms = 0;
if (!m_funs.find(ap->get_decl(), terms)) {
terms = alloc(ptr_vector<expr>);
m_funs.insert(ap->get_decl(), terms);
}
terms->push_back(e);
m_todo.append(ap->get_num_args(), ap->get_args());
}
}
}
void mk_quantifier_instantiation::instantiate_rule(rule& r, expr_ref_vector& conjs, quantifier_ref_vector& qs, rule_set& rules) {
rule_manager& rm = m_ctx.get_rule_manager();
expr_ref fml(m), cnst(m);
var_ref var(m);
ptr_vector<sort> sorts;
r.get_vars(sorts);
m_uf.reset();
m_terms.reset();
m_var2cnst.reset();
m_cnst2var.reset();
fml = m.mk_and(conjs.size(), conjs.c_ptr());
for (unsigned i = 0; i < sorts.size(); ++i) {
if (!sorts[i]) {
sorts[i] = m.mk_bool_sort();
}
var = m.mk_var(i, sorts[i]);
cnst = m.mk_fresh_const("C", sorts[i]);
m_var2cnst.insert(var, cnst);
m_cnst2var.insert(cnst, var);
}
fml = m.mk_and(conjs.size(), conjs.c_ptr());
m_var2cnst(fml);
collect_egraph(fml);
for (unsigned i = 0; i < qs.size(); ++i) {
instantiate_quantifier(qs[i].get(), conjs);
}
obj_map<func_decl, ptr_vector<expr>*>::iterator it = m_funs.begin(), end = m_funs.end();
for (; it != end; ++it) {
dealloc(it->m_value);
}
m_funs.reset();
fml = m.mk_and(conjs.size(), conjs.c_ptr());
fml = m.mk_implies(fml, r.get_head());
TRACE("dl", r.display(m_ctx, tout); tout << mk_pp(fml, m) << "\n";);
rule_set added_rules(m_ctx);
proof_ref pr(m);
rm.mk_rule(fml, pr, added_rules);
if (r.get_proof()) {
// use def-axiom to encode that new rule is a weakening of the original.
proof* p1 = r.get_proof();
for (unsigned i = 0; i < added_rules.get_num_rules(); ++i) {
rule* r2 = added_rules.get_rule(i);
r2->to_formula(fml);
pr = m.mk_modus_ponens(m.mk_def_axiom(m.mk_implies(m.get_fact(p1), fml)), p1);
r2->set_proof(m, pr);
}
}
rules.add_rules(added_rules);
}
rule_set * mk_quantifier_instantiation::operator()(rule_set const & source) {
TRACE("dl", tout << m_ctx.get_params().instantiate_quantifiers() << "\n";);
if (!m_ctx.get_params().instantiate_quantifiers()) {
return 0;
}
bool has_quantifiers = false;
unsigned sz = source.get_num_rules();
for (unsigned i = 0; !has_quantifiers && i < sz; ++i) {
rule& r = *source.get_rule(i);
has_quantifiers = has_quantifiers || r.has_quantifiers();
if (r.has_negation()) {
return 0;
}
}
if (!has_quantifiers) {
return 0;
}
expr_ref_vector conjs(m);
quantifier_ref_vector qs(m);
rule_set * result = alloc(rule_set, m_ctx);
bool instantiated = false;
for (unsigned i = 0; i < sz; ++i) {
rule * r = source.get_rule(i);
extract_quantifiers(*r, conjs, qs);
if (qs.empty()) {
result->add_rule(r);
}
else {
instantiate_rule(*r, conjs, qs, *result);
instantiated = true;
}
}
// model convertion: identity function.
if (instantiated) {
result->inherit_predicates(source);
}
else {
dealloc(result);
result = 0;
}
return result;
}
};

View file

@ -0,0 +1,136 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_quantifier_instantiation.h
Abstract:
Convert Quantified Horn clauses into non-quantified clauses using
instantiation.
Author:
Ken McMillan
Andrey Rybalchenko
Nikolaj Bjorner (nbjorner) 2013-04-02
Revision History:
Based on approach suggested in the SAS 2013 paper
"On Solving Universally Quantified Horn Clauses"
--*/
#ifndef _DL_MK_QUANTIFIER_INSTANTIATION_H_
#define _DL_MK_QUANTIFIER_INSTANTIATION_H_
#include"dl_rule_transformer.h"
#include"expr_safe_replace.h"
namespace datalog {
class context;
class mk_quantifier_instantiation : public rule_transformer::plugin {
typedef svector<std::pair<expr*,expr*> > term_pairs;
class union_find {
unsigned_vector m_find;
unsigned_vector m_size;
unsigned_vector m_next;
void ensure_size(unsigned v) {
while (v >= get_num_vars()) {
mk_var();
}
}
public:
unsigned mk_var() {
unsigned r = m_find.size();
m_find.push_back(r);
m_size.push_back(1);
m_next.push_back(r);
return r;
}
unsigned get_num_vars() const { return m_find.size(); }
unsigned find(unsigned v) const {
if (v >= get_num_vars()) {
return v;
}
while (true) {
unsigned new_v = m_find[v];
if (new_v == v)
return v;
v = new_v;
}
}
unsigned next(unsigned v) const {
if (v >= get_num_vars()) {
return v;
}
return m_next[v];
}
bool is_root(unsigned v) const {
return v >= get_num_vars() || m_find[v] == v;
}
void merge(unsigned v1, unsigned v2) {
unsigned r1 = find(v1);
unsigned r2 = find(v2);
if (r1 == r2)
return;
ensure_size(v1);
ensure_size(v2);
if (m_size[r1] > m_size[r2])
std::swap(r1, r2);
m_find[r1] = r2;
m_size[r2] += m_size[r1];
std::swap(m_next[r1], m_next[r2]);
}
void reset() {
m_find.reset();
m_next.reset();
m_size.reset();
}
};
ast_manager& m;
context& m_ctx;
expr_safe_replace m_var2cnst;
expr_safe_replace m_cnst2var;
union_find m_uf;
ptr_vector<expr> m_todo;
ptr_vector<expr> m_terms;
ptr_vector<expr> m_binding;
obj_map<func_decl, ptr_vector<expr>*> m_funs;
void extract_quantifiers(rule& r, expr_ref_vector& conjs, quantifier_ref_vector& qs);
void collect_egraph(expr* e);
void instantiate_rule(rule& r, expr_ref_vector& conjs, quantifier_ref_vector& qs, rule_set& rules);
void instantiate_quantifier(quantifier* q, expr_ref_vector & conjs);
void instantiate_quantifier(quantifier* q, app* pat, expr_ref_vector & conjs);
void match(unsigned i, app* pat, unsigned j, term_pairs& todo, quantifier* q, expr_ref_vector& conjs);
void yield_binding(quantifier* q, expr_ref_vector& conjs);
public:
mk_quantifier_instantiation(context & ctx, unsigned priority);
virtual ~mk_quantifier_instantiation();
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_QUANTIFIER_INSTANTIATION_H_ */

View file

@ -0,0 +1,899 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_rule_inliner.cpp
Abstract:
Rule transformer which simplifies interpreted tails
Author:
Krystof Hoder (t-khoder) 2011-10-01.
Revision History:
Added linear_inline 2012-9-10 (nbjorner)
Disable inliner for quantified rules 2012-10-31 (nbjorner)
Notes:
Resolution transformation (resolve):
P(x) :- Q(y), phi(x,y) Q(y) :- R(z), psi(y,z)
--------------------------------------------------
P(x) :- R(z), phi(x,y), psi(y,z)
Proof converter:
replace assumption (*) by rule and upper assumptions.
Subsumption transformation (remove rule):
P(x) :- Q(y), phi(x,y) Rules
---------------------------------
Rules
Model converter:
P(x) := P(x) or (exists y . Q(y) & phi(x,y))
--*/
#include <sstream>
#include "ast_pp.h"
#include "dl_finite_product_relation.h"
#include "dl_product_relation.h"
#include "dl_sieve_relation.h"
#include "rewriter.h"
#include "rewriter_def.h"
#include "dl_mk_rule_inliner.h"
namespace datalog {
// -----------------------------------
//
// mk_rule_inliner::rule_unifier
//
// -----------------------------------
bool rule_unifier::unify_rules(const rule& tgt, unsigned tgt_idx, const rule& src) {
rule_counter& vc = m_rm.get_counter();
unsigned var_cnt = std::max(vc.get_max_rule_var(tgt), vc.get_max_rule_var(src))+1;
m_subst.reset();
m_subst.reserve(2, var_cnt);
m_ready = m_unif(tgt.get_tail(tgt_idx), src.get_head(), m_subst);
if (m_ready) {
m_deltas[0] = 0;
m_deltas[1] = var_cnt;
TRACE("dl",
output_predicate(m_context, src.get_head(), tout << "unify rules ");
output_predicate(m_context, tgt.get_head(), tout << "\n");
tout << "\n";);
}
return m_ready;
}
void rule_unifier::apply(app * a, bool is_tgt, app_ref& res) {
expr_ref res_e(m);
TRACE("dl", output_predicate(m_context, a, tout); tout << "\n";);
m_subst.apply(2, m_deltas, expr_offset(a, is_tgt ? 0 : 1), res_e);
SASSERT(is_app(res_e.get()));
res = to_app(res_e.get());
}
void rule_unifier::apply(
rule const& r, bool is_tgt, unsigned skipped_index,
app_ref_vector& res, svector<bool>& res_neg) {
unsigned rule_len = r.get_tail_size();
for (unsigned i = 0; i < rule_len; i++) {
if (i != skipped_index) { //i can never be UINT_MAX, so we'll never skip if we're not supposed to
app_ref new_tail_el(m);
apply(r.get_tail(i), is_tgt, new_tail_el);
res.push_back(new_tail_el);
res_neg.push_back(r.is_neg_tail(i));
}
}
}
bool rule_unifier::apply(rule const& tgt, unsigned tail_index, rule const& src, rule_ref& res) {
SASSERT(m_ready);
app_ref new_head(m);
app_ref_vector tail(m);
svector<bool> tail_neg;
rule_ref simpl_rule(m_rm);
apply(tgt.get_head(), true, new_head);
apply(tgt, true, tail_index, tail, tail_neg);
apply(src, false, UINT_MAX, tail, tail_neg);
mk_rule_inliner::remove_duplicate_tails(tail, tail_neg);
SASSERT(tail.size()==tail_neg.size());
res = m_rm.mk(new_head, tail.size(), tail.c_ptr(), tail_neg.c_ptr(), tgt.name(), m_normalize);
res->set_accounting_parent_object(m_context, const_cast<rule*>(&tgt));
TRACE("dl",
tgt.display(m_context, tout << "tgt (" << tail_index << "): \n");
src.display(m_context, tout << "src:\n");
res->display(m_context, tout << "res\n");
// m_unif.display(tout << "subst:\n");
);
if (m_normalize) {
m_rm.fix_unbound_vars(res, true);
if (m_interp_simplifier.transform_rule(res.get(), simpl_rule)) {
res = simpl_rule;
return true;
}
else {
return false;
}
}
else {
return true;
}
}
expr_ref_vector rule_unifier::get_rule_subst(const rule& r, bool is_tgt) {
SASSERT(m_ready);
expr_ref_vector result(m);
ptr_vector<sort> sorts;
expr_ref v(m), w(m);
r.get_vars(sorts);
for (unsigned i = 0; i < sorts.size(); ++i) {
if (!sorts[i]) {
sorts[i] = m.mk_bool_sort();
}
v = m.mk_var(i, sorts[i]);
m_subst.apply(2, m_deltas, expr_offset(v, is_tgt?0:1), w);
result.push_back(w);
}
return result;
}
// -----------------------------------
//
// mk_rule_inliner
//
// -----------------------------------
/**
Inline occurrences of rule src at tail_index in tgt and return the result in res.
*/
bool mk_rule_inliner::try_to_inline_rule(rule& tgt, rule& src, unsigned tail_index, rule_ref& res)
{
SASSERT(tail_index<tgt.get_positive_tail_size());
SASSERT(!tgt.is_neg_tail(tail_index));
tgt.norm_vars(m_context.get_rule_manager());
SASSERT(!has_quantifier(src));
if (!m_unifier.unify_rules(tgt, tail_index, src)) {
return false;
}
if (m_unifier.apply(tgt, tail_index, src, res)) {
if (m_context.generate_proof_trace()) {
expr_ref_vector s1 = m_unifier.get_rule_subst(tgt, true);
expr_ref_vector s2 = m_unifier.get_rule_subst(src, false);
datalog::resolve_rule(tgt, src, tail_index, s1, s2, *res.get());
}
return true;
}
else {
TRACE("dl", res->display(m_context, tout << "interpreted tail is unsat\n"););
//the interpreted part is unsatisfiable
return false;
}
}
// TBD: replace by r.has_quantifiers() and test
bool mk_rule_inliner::has_quantifier(rule const& r) const {
unsigned utsz = r.get_uninterpreted_tail_size();
for (unsigned i = utsz; i < r.get_tail_size(); ++i) {
if (r.get_tail(i)->has_quantifiers()) return true;
}
return false;
}
void mk_rule_inliner::count_pred_occurrences(rule_set const & orig)
{
rel_context* rel = m_context.get_rel_context();
if (rel) {
rel->get_rmanager().collect_non_empty_predicates(m_preds_with_facts);
}
rule_set::iterator rend = orig.end();
for (rule_set::iterator rit = orig.begin(); rit!=rend; ++rit) {
rule * r = *rit;
func_decl * head_pred = r->get_decl();
m_head_pred_ctr.inc(head_pred);
if (r->get_tail_size()>0) {
m_head_pred_non_empty_tails_ctr.inc(head_pred);
}
unsigned ut_len = r->get_uninterpreted_tail_size();
for (unsigned i=0; i<ut_len; i++) {
func_decl * pred = r->get_decl(i);
m_tail_pred_ctr.inc(pred);
if (r->is_neg_tail(i)) {
m_preds_with_neg_occurrence.insert(pred);
}
}
}
}
bool mk_rule_inliner::inlining_allowed(rule_set const& source, func_decl * pred)
{
if (//these three conditions are important for soundness
source.is_output_predicate(pred) ||
m_preds_with_facts.contains(pred) ||
m_preds_with_neg_occurrence.contains(pred) ||
//this condition is used for breaking of cycles among inlined rules
m_forbidden_preds.contains(pred)) {
return false;
}
//
// these conditions are optional, they avoid possible exponential increase
// in the size of the problem
//
return
//m_head_pred_non_empty_tails_ctr.get(pred)<=1
m_head_pred_ctr.get(pred) <= 1
|| (m_tail_pred_ctr.get(pred) <= 1 && m_head_pred_ctr.get(pred) <= 4)
;
}
/** Caller has to dealloc the returned object */
rule_set * mk_rule_inliner::create_allowed_rule_set(rule_set const & orig)
{
rule_set * res = alloc(rule_set, m_context);
unsigned rcnt = orig.get_num_rules();
for (unsigned i=0; i<rcnt; i++) {
rule * r = orig.get_rule(i);
if (inlining_allowed(orig, r->get_decl())) {
res->add_rule(r);
}
}
//the rule set should be stratified, since orig (which is its superset) is as well
VERIFY(res->close());
return res;
}
/**
Try to make the set of inlined predicates acyclic by forbidding inlining of one
predicate from each strongly connected component. Return true if we did forbide some
predicate, and false if the set of rules is already acyclic.
*/
bool mk_rule_inliner::forbid_preds_from_cycles(rule_set const & r)
{
SASSERT(r.is_closed());
bool something_forbidden = false;
const rule_stratifier::comp_vector& comps = r.get_stratifier().get_strats();
rule_stratifier::comp_vector::const_iterator cend = comps.end();
for (rule_stratifier::comp_vector::const_iterator it = comps.begin(); it!=cend; ++it) {
rule_stratifier::item_set * stratum = *it;
if (stratum->size()==1) {
continue;
}
SASSERT(stratum->size()>1);
func_decl * first_stratum_pred = *stratum->begin();
//we're trying to break cycles by removing one predicate from each of them
m_forbidden_preds.insert(first_stratum_pred);
something_forbidden = true;
}
return something_forbidden;
}
bool mk_rule_inliner::forbid_multiple_multipliers(const rule_set & orig,
rule_set const & proposed_inlined_rules) {
bool something_forbidden = false;
const rule_stratifier::comp_vector& comps =
proposed_inlined_rules.get_stratifier().get_strats();
rule_stratifier::comp_vector::const_iterator cend = comps.end();
for (rule_stratifier::comp_vector::const_iterator it = comps.begin(); it!=cend; ++it) {
rule_stratifier::item_set * stratum = *it;
SASSERT(stratum->size()==1);
func_decl * head_pred = *stratum->begin();
bool is_multi_head_pred = m_head_pred_ctr.get(head_pred)>1;
bool is_multi_occurrence_pred = m_tail_pred_ctr.get(head_pred)>1;
const rule_vector& pred_rules = proposed_inlined_rules.get_predicate_rules(head_pred);
rule_vector::const_iterator iend = pred_rules.end();
for (rule_vector::const_iterator iit = pred_rules.begin(); iit!=iend; ++iit) {
rule * r = *iit;
unsigned pt_len = r->get_positive_tail_size();
for (unsigned ti = 0; ti<pt_len; ++ti) {
func_decl * tail_pred = r->get_decl(ti);
if (!inlining_allowed(orig, tail_pred)) {
continue;
}
unsigned tail_pred_head_cnt = m_head_pred_ctr.get(tail_pred);
if (tail_pred_head_cnt<=1) {
continue;
}
if (is_multi_head_pred) {
m_forbidden_preds.insert(head_pred);
something_forbidden = true;
goto process_next_pred;
}
if (is_multi_occurrence_pred) {
m_forbidden_preds.insert(tail_pred);
something_forbidden = true;
}
else {
is_multi_head_pred = true;
m_head_pred_ctr.get(head_pred) =
m_head_pred_ctr.get(head_pred)*tail_pred_head_cnt;
}
}
}
process_next_pred:;
}
unsigned rule_cnt = orig.get_num_rules();
for (unsigned ri=0; ri<rule_cnt; ri++) {
rule * r = orig.get_rule(ri);
func_decl * head_pred = r->get_decl();
if (inlining_allowed(orig, head_pred)) {
//we have already processed inlined rules
continue;
}
bool has_multi_head_pred = false;
unsigned pt_len = r->get_positive_tail_size();
for (unsigned ti = 0; ti<pt_len; ++ti) {
func_decl * pred = r->get_decl(ti);
if (!inlining_allowed(orig, pred)) {
continue;
}
if (m_head_pred_ctr.get(pred)<=1) {
continue;
}
if (has_multi_head_pred) {
m_forbidden_preds.insert(pred);
something_forbidden = true;
}
else {
has_multi_head_pred = true;
}
}
}
return something_forbidden;
}
void mk_rule_inliner::plan_inlining(rule_set const & orig)
{
count_pred_occurrences(orig);
scoped_ptr<rule_set> candidate_inlined_set = create_allowed_rule_set(orig);
while (forbid_preds_from_cycles(*candidate_inlined_set)) {
candidate_inlined_set = create_allowed_rule_set(orig);
}
if (forbid_multiple_multipliers(orig, *candidate_inlined_set)) {
candidate_inlined_set = create_allowed_rule_set(orig);
}
TRACE("dl", tout<<"rules to be inlined:\n" << (*candidate_inlined_set); );
// now we start filling in the set of the inlined rules in a topological order,
// so that we inline rules into other rules
SASSERT(m_inlined_rules.get_num_rules()==0);
const rule_stratifier::comp_vector& comps = candidate_inlined_set->get_stratifier().get_strats();
rule_stratifier::comp_vector::const_iterator cend = comps.end();
for (rule_stratifier::comp_vector::const_iterator it = comps.begin(); it!=cend; ++it) {
rule_stratifier::item_set * stratum = *it;
SASSERT(stratum->size()==1);
func_decl * pred = *stratum->begin();
const rule_vector& pred_rules = candidate_inlined_set->get_predicate_rules(pred);
rule_vector::const_iterator iend = pred_rules.end();
for (rule_vector::const_iterator iit = pred_rules.begin(); iit!=iend; ++iit) {
transform_rule(orig, *iit, m_inlined_rules);
}
}
TRACE("dl", tout << "inlined rules after mutual inlining:\n" << m_inlined_rules; );
}
bool mk_rule_inliner::transform_rule(rule_set const& orig, rule * r0, rule_set& tgt) {
bool modified = false;
rule_ref_vector todo(m_rm);
todo.push_back(r0);
while (!todo.empty()) {
rule_ref r(todo.back(), m_rm);
todo.pop_back();
unsigned pt_len = r->get_positive_tail_size();
unsigned i = 0;
for (; i < pt_len && !inlining_allowed(orig, r->get_decl(i)); ++i) {};
SASSERT(!has_quantifier(*r.get()));
if (i == pt_len) {
//there's nothing we can inline in this rule
tgt.add_rule(r);
continue;
}
modified = true;
func_decl * pred = r->get_decl(i);
const rule_vector& pred_rules = m_inlined_rules.get_predicate_rules(pred);
rule_vector::const_iterator iend = pred_rules.end();
for (rule_vector::const_iterator iit = pred_rules.begin(); iit!=iend; ++iit) {
rule * inl_rule = *iit;
rule_ref inl_result(m_rm);
if (try_to_inline_rule(*r.get(), *inl_rule, i, inl_result)) {
todo.push_back(inl_result);
}
}
}
if (modified) {
datalog::del_rule(m_mc, *r0);
}
return modified;
}
bool mk_rule_inliner::transform_rules(const rule_set & orig, rule_set & tgt) {
bool something_done = false;
rule_set::iterator rend = orig.end();
for (rule_set::iterator rit = orig.begin(); rit!=rend; ++rit) {
rule_ref r(*rit, m_rm);
func_decl * pred = r->get_decl();
// if inlining is allowed, then we are eliminating
// this relation through inlining,
// so we don't add its rules to the result
something_done |= !inlining_allowed(orig, pred) && transform_rule(orig, r, tgt);
}
if (something_done && m_mc) {
for (rule_set::iterator rit = orig.begin(); rit!=rend; ++rit) {
if (inlining_allowed(orig, (*rit)->get_decl())) {
datalog::del_rule(m_mc, **rit);
}
}
}
return something_done;
}
/**
Check whether rule r is oriented in a particular ordering.
This is to avoid infinite cycle of inlining in the eager inliner.
Out ordering is lexicographic, comparing atoms first on stratum they are in,
then on arity and then on ast ID of their func_decl.
*/
bool mk_rule_inliner::is_oriented_rewriter(rule * r, rule_stratifier const& strat) {
func_decl * head_pred = r->get_decl();
unsigned head_strat = strat.get_predicate_strat(head_pred);
unsigned head_arity = head_pred->get_arity();
unsigned pt_len = r->get_positive_tail_size();
for (unsigned ti=0; ti<pt_len; ++ti) {
func_decl * pred = r->get_decl(ti);
unsigned pred_strat = strat.get_predicate_strat(pred);
SASSERT(pred_strat<=head_strat);
if (pred_strat==head_strat) {
if (pred->get_arity()>head_arity
|| (pred->get_arity()==head_arity && pred->get_id()>=head_pred->get_id()) ) {
return false;
}
}
}
return true;
}
bool mk_rule_inliner::do_eager_inlining(rule * r, rule_set const& rules, rule_ref& res) {
SASSERT(rules.is_closed());
const rule_stratifier& strat = rules.get_stratifier();
func_decl * head_pred = r->get_decl();
unsigned pt_len = r->get_positive_tail_size();
for (unsigned ti = 0; ti < pt_len; ++ti) {
func_decl * pred = r->get_decl(ti);
if (pred == head_pred || m_preds_with_facts.contains(pred)) { continue; }
const rule_vector& pred_rules = rules.get_predicate_rules(pred);
rule * inlining_candidate = 0;
unsigned rule_cnt = pred_rules.size();
if (rule_cnt == 0) {
inlining_candidate = 0;
}
else if (rule_cnt == 1) {
inlining_candidate = pred_rules[0];
}
else {
inlining_candidate = 0;
for (unsigned ri = 0; ri < rule_cnt; ++ri) {
rule * pred_rule = pred_rules[ri];
if (!m_unifier.unify_rules(*r, ti, *pred_rule)) {
//we skip rules which don't unify with the tail atom
continue;
}
if (inlining_candidate != 0) {
// We have two rules that can be inlined into the current
// tail predicate. In this situation we don't do inlinning
// on this tail atom, as we don't want the overall number
// of rules to increase.
goto process_next_tail;
}
inlining_candidate = pred_rule;
}
}
if (inlining_candidate == 0) {
// nothing unifies with the tail atom, therefore the rule is unsatisfiable
// (we can say this because relation pred doesn't have any ground facts either)
res = 0;
datalog::del_rule(m_mc, *r);
return true;
}
if (!is_oriented_rewriter(inlining_candidate, strat)) {
// The rule which should be used for inlining isn't oriented
// in a simplifying direction. Inlining with such rule might lead to
// infinite loops, so we don't do it.
goto process_next_tail;
}
if (!try_to_inline_rule(*r, *inlining_candidate, ti, res)) {
datalog::del_rule(m_mc, *r);
res = 0;
}
return true;
process_next_tail:;
}
return false;
}
bool mk_rule_inliner::do_eager_inlining(scoped_ptr<rule_set> & rules) {
scoped_ptr<rule_set> res = alloc(rule_set, m_context);
bool done_something = false;
rule_set::iterator rend = rules->end();
for (rule_set::iterator rit = rules->begin(); rit!=rend; ++rit) {
rule_ref r(*rit, m_rm);
rule_ref replacement(m_rm);
while (r && do_eager_inlining(r, *rules, replacement)) {
r = replacement;
done_something = true;
}
if (!r) {
continue;
}
res->add_rule(r);
}
if (done_something) {
rules = res.detach();
}
return done_something;
}
/**
Inline predicates that are known to not be join-points.
P(1,x) :- P(0,y), phi(x,y)
P(0,x) :- P(1,z), psi(x,z)
->
P(1,x) :- P(1,z), phi(x,y), psi(y,z)
whenever P(0,x) is not unifiable with the
body of the rule where it appears (P(1,z))
and P(0,x) is unifiable with at most one (?)
other rule (and it does not occur negatively).
*/
bool mk_rule_inliner::visitor::operator()(expr* e) {
m_unifiers.append(m_positions.find(e));
TRACE("dl",
tout << "unifier: " << (m_unifiers.empty()?0:m_unifiers.back());
tout << " num unifiers: " << m_unifiers.size();
tout << " num positions: " << m_positions.find(e).size() << "\n";
output_predicate(m_context, to_app(e), tout); tout << "\n";);
return true;
}
void mk_rule_inliner::visitor::reset(unsigned sz) {
m_unifiers.reset();
m_can_remove.reset();
m_can_remove.resize(sz, true);
m_can_expand.reset();
m_can_expand.resize(sz, true);
m_positions.reset();
}
unsigned_vector const& mk_rule_inliner::visitor::add_position(expr* e, unsigned j) {
obj_map<expr, unsigned_vector>::obj_map_entry * et = m_positions.insert_if_not_there2(e, unsigned_vector());
et->get_data().m_value.push_back(j);
return et->get_data().m_value;
}
unsigned_vector const& mk_rule_inliner::visitor::del_position(expr* e, unsigned j) {
obj_map<expr, unsigned_vector>::obj_map_entry * et = m_positions.find_core(e);
SASSERT(et && et->get_data().m_value.contains(j));
et->get_data().m_value.erase(j);
return et->get_data().m_value;
}
void mk_rule_inliner::add_rule(rule_set const& source, rule* r, unsigned i) {
svector<bool>& can_remove = m_head_visitor.can_remove();
svector<bool>& can_expand = m_head_visitor.can_expand();
app* head = r->get_head();
func_decl* headd = head->get_decl();
m_head_visitor.add_position(head, i);
m_head_index.insert(head);
m_pinned.push_back(r);
if (source.is_output_predicate(headd) ||
m_preds_with_facts.contains(headd)) {
can_remove.set(i, false);
TRACE("dl", output_predicate(m_context, head, tout << "cannot remove: " << i << " "); tout << "\n";);
}
unsigned tl_sz = r->get_uninterpreted_tail_size();
for (unsigned j = 0; j < tl_sz; ++j) {
app* tail = r->get_tail(j);
m_tail_visitor.add_position(tail, i);
m_tail_index.insert(tail);
}
bool can_exp =
tl_sz == 1
&& r->get_positive_tail_size() == 1
&& !m_preds_with_facts.contains(r->get_decl(0))
&& !source.is_output_predicate(r->get_decl(0));
can_expand.set(i, can_exp);
}
void mk_rule_inliner::del_rule(rule* r, unsigned i) {
app* head = r->get_head();
m_head_visitor.del_position(head, i);
unsigned tl_sz = r->get_uninterpreted_tail_size();
for (unsigned j = 0; j < tl_sz; ++j) {
app* tail = r->get_tail(j);
m_tail_visitor.del_position(tail, i);
}
}
#define PRT(_x_) ((_x_)?"T":"F")
bool mk_rule_inliner::inline_linear(scoped_ptr<rule_set>& rules) {
bool done_something = false;
unsigned sz = rules->get_num_rules();
m_head_visitor.reset(sz);
m_tail_visitor.reset(sz);
m_head_index.reset();
m_tail_index.reset();
TRACE("dl", rules->display(tout););
rule_ref_vector acc(m_rm);
for (unsigned i = 0; i < sz; ++i) {
acc.push_back(rules->get_rule(i));
}
// set up unification index.
svector<bool>& can_remove = m_head_visitor.can_remove();
svector<bool>& can_expand = m_head_visitor.can_expand();
for (unsigned i = 0; i < sz; ++i) {
add_rule(*rules, acc[i].get(), i);
}
// initialize substitution.
rule_counter& vc = m_rm.get_counter();
unsigned max_var = 0;
for (unsigned i = 0; i < sz; ++i) {
rule* r = acc[i].get();
max_var = std::max(max_var, vc.get_max_var(r->get_head()));
unsigned tl_sz = r->get_uninterpreted_tail_size();
for (unsigned j = 0; j < tl_sz; ++j) {
max_var = std::max(max_var, vc.get_max_var(r->get_tail(j)));
}
}
m_subst.reset();
m_subst.reserve_vars(max_var+1);
m_subst.reserve_offsets(std::max(m_tail_index.get_approx_num_regs(), 2+m_head_index.get_approx_num_regs()));
svector<bool> valid;
valid.reset();
valid.resize(sz, true);
bool allow_branching = m_context.get_params().inline_linear_branch();
for (unsigned i = 0; i < sz; ++i) {
while (true) {
rule_ref r(acc[i].get(), m_rm);
TRACE("dl", r->display(m_context, tout << "processing: " << i << "\n"););
if (!valid.get(i)) {
TRACE("dl", tout << "invalid: " << i << "\n";);
break;
}
if (!can_expand.get(i)) {
TRACE("dl", tout << "cannot expand: " << i << "\n";);
break;
}
m_head_visitor.reset();
m_head_index.unify(r->get_tail(0), m_head_visitor);
unsigned num_head_unifiers = m_head_visitor.get_unifiers().size();
if (num_head_unifiers != 1) {
TRACE("dl", tout << "no unique unifier " << num_head_unifiers << "\n";);
break;
}
unsigned j = m_head_visitor.get_unifiers()[0];
if (!can_remove.get(j) || !valid.get(j) || i == j) {
TRACE("dl", tout << PRT(can_remove.get(j)) << " " << PRT(valid.get(j)) << " " << PRT(i != j) << "\n";);
break;
}
rule* r2 = acc[j].get();
// check that the head of r2 only unifies with this single body position.
TRACE("dl", output_predicate(m_context, r2->get_head(), tout << "unify head: "); tout << "\n";);
m_tail_visitor.reset();
m_tail_index.unify(r2->get_head(), m_tail_visitor);
unsigned_vector const& tail_unifiers = m_tail_visitor.get_unifiers();
unsigned num_tail_unifiers = tail_unifiers.size();
SASSERT(!tail_unifiers.empty());
if (!allow_branching && num_tail_unifiers != 1) {
TRACE("dl", tout << "too many tails " << num_tail_unifiers << "\n";);
break;
}
rule_ref rl_res(m_rm);
if (!try_to_inline_rule(*r.get(), *r2, 0, rl_res)) {
TRACE("dl", r->display(m_context, tout << "inlining failed\n"); r2->display(m_context, tout); );
break;
}
done_something = true;
TRACE("dl", r->display(m_context, tout); r2->display(m_context, tout); rl_res->display(m_context, tout); );
del_rule(r, i);
add_rule(*rules, rl_res.get(), i);
r = rl_res;
acc[i] = r.get();
can_expand.set(i, can_expand.get(j));
if (num_tail_unifiers == 1) {
TRACE("dl", tout << "setting invalid: " << j << "\n";);
valid.set(j, false);
datalog::del_rule(m_mc, *r2);
del_rule(r2, j);
}
max_var = std::max(max_var, vc.get_max_rule_var(*r.get()));
m_subst.reserve_vars(max_var+1);
}
}
if (done_something) {
scoped_ptr<rule_set> res = alloc(rule_set, m_context);
for (unsigned i = 0; i < sz; ++i) {
if (valid.get(i)) {
res->add_rule(acc[i].get());
}
}
res->inherit_predicates(*rules);
TRACE("dl", res->display(tout););
rules = res.detach();
}
return done_something;
}
rule_set * mk_rule_inliner::operator()(rule_set const & source) {
bool something_done = false;
ref<horn_subsume_model_converter> hsmc;
if (source.get_num_rules() == 0) {
return 0;
}
rule_set::iterator end = source.end();
for (rule_set::iterator it = source.begin(); it != end; ++ it) {
if (has_quantifier(**it)) {
return 0;
}
}
if (m_context.get_model_converter()) {
hsmc = alloc(horn_subsume_model_converter, m);
}
m_mc = hsmc.get();
scoped_ptr<rule_set> res = alloc(rule_set, m_context);
if (m_context.get_params().inline_eager()) {
TRACE("dl", source.display(tout << "before eager inlining\n"););
plan_inlining(source);
something_done = transform_rules(source, *res);
VERIFY(res->close()); //this transformation doesn't break the negation stratification
// try eager inlining
if (do_eager_inlining(res)) {
something_done = true;
}
TRACE("dl", res->display(tout << "after eager inlining\n"););
}
if (something_done) {
res->inherit_predicates(source);
}
else {
res = alloc(rule_set, source);
}
if (m_context.get_params().inline_linear() && inline_linear(res)) {
something_done = true;
}
if (!something_done) {
res = 0;
}
else {
m_context.add_model_converter(hsmc.get());
}
return res.detach();
}
};

View file

@ -0,0 +1,205 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_interp_tail_simplifier.h
Abstract:
Rule transformer which inlines some of the rules
Author:
Krystof Hoder (t-khoder) 2011-10-02.
Revision History:
--*/
#ifndef _DL_MK_RULE_INLINER_H_
#define _DL_MK_RULE_INLINER_H_
#include "dl_context.h"
#include "dl_rule_transformer.h"
#include "dl_mk_interp_tail_simplifier.h"
#include "unifier.h"
#include "substitution.h"
#include "substitution_tree.h"
namespace datalog {
class rule_unifier {
ast_manager& m;
rule_manager& m_rm;
context& m_context;
/** We use this simplifier after inlining to get nicer intermediate rules */
mk_interp_tail_simplifier m_interp_simplifier;
substitution m_subst;
unifier m_unif;
bool m_ready;
bool m_normalize;
unsigned m_deltas[2];
public:
rule_unifier(context& ctx)
: m(ctx.get_manager()), m_rm(ctx.get_rule_manager()), m_context(ctx),
m_interp_simplifier(ctx), m_subst(m), m_unif(m), m_ready(false), m_normalize(true) {}
/** Reset subtitution and unify tail tgt_idx of the target rule and the head of the src rule */
bool unify_rules(rule const& tgt, unsigned tgt_idx, rule const& src);
/**
\brief Apply unifier to rules.
Return false if the resulting rule is a tautology (the interpreted tail is unsat).
*/
bool apply(rule const& tgt, unsigned tgt_idx, rule const& src, rule_ref& result);
void display(std::ostream& stm) { m_subst.display(stm, 2, m_deltas); }
/**
Retrieve substitutions for src/tgt. (second argument of unify_rules).
*/
expr_ref_vector get_rule_subst(rule const& r, bool is_tgt);
/**
Control if bound variables are normalized after unification.
The default is 'true': bound variables are re-mapped to an
initial segment of de-Bruijn indices.
*/
void set_normalize(bool n) { m_normalize = n; }
private:
void apply(app * a, bool is_tgt, app_ref& res);
/**
Apply substitution to a rule tail. Tail with skipped_index is skipped,
unless skipped_index is equal to UINT_MAX
*/
void apply(rule const& r, bool is_tgt, unsigned skipped_index, app_ref_vector& res,
svector<bool>& res_neg);
};
class mk_rule_inliner : public rule_transformer::plugin {
class visitor : public st_visitor {
context& m_context;
unsigned_vector m_unifiers;
svector<bool> m_can_remove, m_can_expand;
obj_map<expr, unsigned_vector> m_positions;
public:
visitor(context& c, substitution & s): st_visitor(s), m_context(c) {}
virtual bool operator()(expr* e);
void reset() { m_unifiers.reset(); }
void reset(unsigned sz);
svector<bool>& can_remove() { return m_can_remove; }
svector<bool>& can_expand() { return m_can_expand; }
unsigned_vector const& add_position(expr* e, unsigned j);
unsigned_vector const& del_position(expr* e, unsigned j);
unsigned_vector const& get_unifiers() { return m_unifiers; }
};
typedef obj_map<func_decl, func_decl *> decl_map;
ast_manager & m;
rule_manager & m_rm;
context & m_context;
th_rewriter& m_simp;
rule_ref_vector m_pinned;
decl_set m_forbidden_preds;
decl_set m_preds_with_facts;
decl_set m_preds_with_neg_occurrence;
ast_counter m_head_pred_ctr;
ast_counter m_head_pred_non_empty_tails_ctr;
ast_counter m_tail_pred_ctr;
rule_set m_inlined_rules;
horn_subsume_model_converter* m_mc;
//used in try_to_inline_rule and do_eager_inlining
rule_unifier m_unifier;
substitution_tree m_head_index; // for straight-line relation inlining.
substitution_tree m_tail_index;
substitution m_subst;
visitor m_head_visitor;
visitor m_tail_visitor;
bool tail_matches_head(app * tail, app* head);
bool try_to_inline_rule(rule& tgt, rule& src, unsigned tail_index, rule_ref& res);
bool inlining_allowed(rule_set const& orig, func_decl * pred);
void count_pred_occurrences(rule_set const & orig);
void plan_inlining(rule_set const & orig);
rule_set * create_allowed_rule_set(rule_set const & orig);
bool forbid_preds_from_cycles(rule_set const & r);
/** Ensure we don't inline two multi-head rules that would appear together in some tail */
bool forbid_multiple_multipliers(const rule_set & orig, rule_set const & proposed_inlined_rules);
/** Return true if the rule was modified */
bool transform_rule(rule_set const& orig, rule * r, rule_set& tgt);
/** Return true if some transformation was performed */
bool transform_rules(const rule_set & orig, rule_set & tgt);
bool is_oriented_rewriter(rule * r, rule_stratifier const& strat);
/**
Return false if nothing was done with the rule.
res may be set to zero if we managed to prove the rule unsatisfiable.
*/
bool do_eager_inlining(rule * r, rule_set const& rules, rule_ref& res);
/**
Inline rules even if it doesn't lead to elimination of the whole predicate.
The inlining is done as long as it doesn't increase the number of rules
(i.e. when only one rule defining a predicate can replace tail atom).
The original rule-set must be closed before passing t this function
*/
bool do_eager_inlining(scoped_ptr<rule_set> & rules);
bool has_quantifier(rule const& r) const;
/**
Inline predicates that are known to not be join-points.
*/
bool inline_linear(scoped_ptr<rule_set>& rules);
void add_rule(rule_set const& rule_set, rule* r, unsigned i);
void del_rule(rule* r, unsigned i);
public:
mk_rule_inliner(context & ctx, unsigned priority=35000)
: plugin(priority),
m(ctx.get_manager()),
m_rm(ctx.get_rule_manager()),
m_context(ctx),
m_simp(m_context.get_rewriter()),
m_pinned(m_rm),
m_inlined_rules(m_context),
m_mc(0),
m_unifier(ctx),
m_head_index(m),
m_tail_index(m),
m_subst(m),
m_head_visitor(ctx, m_subst),
m_tail_visitor(ctx, m_subst)
{}
virtual ~mk_rule_inliner() { }
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_INTERP_TAIL_SIMPLIFIER_H_ */

241
src/muz/dl_mk_scale.cpp Normal file
View file

@ -0,0 +1,241 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_scale.cpp
Abstract:
Author:
Nikolaj Bjorner (nbjorner) 2013-08-19
Revision History:
--*/
#include"dl_mk_scale.h"
#include"dl_context.h"
namespace datalog {
class mk_scale::scale_model_converter : public model_converter {
ast_manager& m;
func_decl_ref_vector m_trail;
arith_util a;
obj_map<func_decl, func_decl*> m_new2old;
public:
scale_model_converter(ast_manager& m): m(m), m_trail(m), a(m) {}
virtual ~scale_model_converter() {}
void add_new2old(func_decl* new_f, func_decl* old_f) {
m_trail.push_back(old_f);
m_trail.push_back(new_f);
m_new2old.insert(new_f, old_f);
}
virtual void operator()(model_ref& md) {
model_ref old_model = alloc(model, m);
obj_map<func_decl, func_decl*>::iterator it = m_new2old.begin();
obj_map<func_decl, func_decl*>::iterator end = m_new2old.end();
for (; it != end; ++it) {
func_decl* old_p = it->m_value;
func_decl* new_p = it->m_key;
func_interp* old_fi = alloc(func_interp, m, old_p->get_arity());
if (new_p->get_arity() == 0) {
old_fi->set_else(md->get_const_interp(new_p));
}
else {
func_interp* new_fi = md->get_func_interp(new_p);
expr_ref_vector subst(m);
var_subst vs(m, false);
expr_ref tmp(m);
if (!new_fi) {
TRACE("dl", tout << new_p->get_name() << " has no value in the current model\n";);
dealloc(old_fi);
continue;
}
for (unsigned i = 0; i < old_p->get_arity(); ++i) {
subst.push_back(m.mk_var(i, old_p->get_domain(i)));
}
subst.push_back(a.mk_numeral(rational(1), a.mk_real()));
// Hedge that we don't have to handle the general case for models produced
// by Horn clause solvers.
SASSERT(!new_fi->is_partial() && new_fi->num_entries() == 0);
vs(new_fi->get_else(), subst.size(), subst.c_ptr(), tmp);
old_fi->set_else(tmp);
old_model->register_decl(old_p, old_fi);
}
}
// register values that have not been scaled.
unsigned sz = md->get_num_constants();
for (unsigned i = 0; i < sz; ++i) {
func_decl* c = md->get_constant(i);
if (!m_new2old.contains(c)) {
old_model->register_decl(c, md->get_const_interp(c));
}
}
sz = md->get_num_functions();
for (unsigned i = 0; i < sz; ++i) {
func_decl* f = md->get_function(i);
if (!m_new2old.contains(f)) {
func_interp* fi = md->get_func_interp(f);
old_model->register_decl(f, fi->copy());
}
}
md = old_model;
//TRACE("dl", model_smt2_pp(tout, m, *md, 0); );
}
virtual model_converter * translate(ast_translation & translator) {
UNREACHABLE();
return 0;
}
};
mk_scale::mk_scale(context & ctx, unsigned priority):
plugin(priority),
m(ctx.get_manager()),
m_ctx(ctx),
a(m),
m_trail(m),
m_eqs(m) {
}
mk_scale::~mk_scale() {
}
rule_set * mk_scale::operator()(rule_set const & source) {
if (!m_ctx.get_params().scale()) {
return 0;
}
rule_manager& rm = source.get_rule_manager();
rule_set * result = alloc(rule_set, m_ctx);
unsigned sz = source.get_num_rules();
rule_ref new_rule(rm);
app_ref_vector tail(m);
app_ref head(m);
svector<bool> neg;
ptr_vector<sort> vars;
ref<scale_model_converter> smc;
if (m_ctx.get_model_converter()) {
smc = alloc(scale_model_converter, m);
}
m_mc = smc.get();
for (unsigned i = 0; i < sz; ++i) {
rule & r = *source.get_rule(i);
unsigned utsz = r.get_uninterpreted_tail_size();
unsigned tsz = r.get_tail_size();
tail.reset();
vars.reset();
m_cache.reset();
m_trail.reset();
m_eqs.reset();
r.get_vars(vars);
unsigned num_vars = vars.size();
for (unsigned j = 0; j < utsz; ++j) {
tail.push_back(mk_pred(num_vars, r.get_tail(j)));
}
for (unsigned j = utsz; j < tsz; ++j) {
tail.push_back(mk_constraint(num_vars, r.get_tail(j)));
}
app_ref new_pred = mk_pred(num_vars, r.get_head());
tail.append(m_eqs);
tail.push_back(a.mk_gt(m.mk_var(num_vars, a.mk_real()), a.mk_numeral(rational(0), false)));
neg.resize(tail.size(), false);
new_rule = rm.mk(new_pred, tail.size(), tail.c_ptr(), neg.c_ptr(), r.name(), true);
result->add_rule(new_rule);
if (source.is_output_predicate(r.get_decl())) {
result->set_output_predicate(new_rule->get_decl());
}
}
TRACE("dl", result->display(tout););
if (m_mc) {
m_ctx.add_model_converter(m_mc);
}
m_trail.reset();
m_cache.reset();
return result;
}
app_ref mk_scale::mk_pred(unsigned sigma_idx, app* q) {
func_decl* f = q->get_decl();
ptr_vector<sort> domain(f->get_arity(), f->get_domain());
domain.push_back(a.mk_real());
func_decl_ref g(m);
g = m.mk_func_decl(f->get_name(), f->get_arity() + 1, domain.c_ptr(), f->get_range());
expr_ref_vector args(m);
for (unsigned i = 0; i < q->get_num_args(); ++i) {
expr* arg = q->get_arg(i);
rational val;
if (a.is_numeral(arg, val)) {
if (val.is_zero()) {
// arg is unchanged.
}
else if (val.is_one()) {
arg = m.mk_var(sigma_idx, a.mk_real());
}
else {
// create a fresh variable 'v', add 'v == sigma*arg'
expr* v = m.mk_var(sigma_idx + 1 + m_eqs.size(), a.mk_real());
m_eqs.push_back(m.mk_eq(v, a.mk_mul(arg, m.mk_var(sigma_idx, a.mk_real()))));
arg = v;
}
}
args.push_back(arg);
}
args.push_back(m.mk_var(sigma_idx, a.mk_real()));
m_ctx.register_predicate(g, false);
if (m_mc) {
m_mc->add_new2old(g, f);
}
return app_ref(m.mk_app(g, q->get_num_args() + 1, args.c_ptr()), m);
}
app_ref mk_scale::mk_constraint(unsigned sigma_idx, app* q) {
expr* r = linearize(sigma_idx, q);
SASSERT(is_app(r));
return app_ref(to_app(r), m);
}
expr* mk_scale::linearize(unsigned sigma_idx, expr* e) {
expr* r;
if (m_cache.find(e, r)) {
return r;
}
if (!is_app(e)) {
return e;
}
expr_ref result(m);
app* ap = to_app(e);
if (ap->get_family_id() == m.get_basic_family_id() ||
a.is_add(e) || a.is_sub(e) ||
a.is_le(e) || a.is_ge(e) ||
a.is_lt(e) || a.is_gt(e)) {
expr_ref_vector args(m);
for (unsigned i = 0; i < ap->get_num_args(); ++i) {
args.push_back(linearize(sigma_idx, ap->get_arg(i)));
}
result = m.mk_app(ap->get_decl(), args.size(), args.c_ptr());
}
else if (a.is_numeral(e)) {
result = a.mk_mul(m.mk_var(sigma_idx, a.mk_real()), e);
}
else {
result = e;
}
m_trail.push_back(result);
m_cache.insert(e, result);
return result;
}
};

53
src/muz/dl_mk_scale.h Normal file
View file

@ -0,0 +1,53 @@
/*++
Copyright (c) 2013 Microsoft Corporation
Module Name:
dl_mk_scale.h
Abstract:
Add scale factor to linear (Real) arithemetic Horn clauses.
The transformation replaces occurrences of isolated constants by
a scale multiplied to each constant.
Author:
Nikolaj Bjorner (nbjorner) 2013-08-19
Revision History:
--*/
#ifndef _DL_MK_SCALE_H_
#define _DL_MK_SCALE_H_
#include"dl_rule_transformer.h"
#include"arith_decl_plugin.h"
namespace datalog {
class mk_scale : public rule_transformer::plugin {
class scale_model_converter;
ast_manager& m;
context& m_ctx;
arith_util a;
expr_ref_vector m_trail;
app_ref_vector m_eqs;
obj_map<expr, expr*> m_cache;
scale_model_converter* m_mc;
expr* linearize(unsigned num_vars, expr* e);
app_ref mk_pred(unsigned num_vars, app* q);
app_ref mk_constraint(unsigned num_vars, app* q);
public:
mk_scale(context & ctx, unsigned priority = 33039);
virtual ~mk_scale();
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_SCALE_H_ */

View file

@ -0,0 +1,545 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_similarity_compressor.cpp
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-10-22.
Revision History:
--*/
#include<utility>
#include<sstream>
#include"dl_mk_similarity_compressor.h"
namespace datalog {
mk_similarity_compressor::mk_similarity_compressor(context & ctx) :
plugin(5000),
m_context(ctx),
m_manager(ctx.get_manager()),
m_threshold_count(ctx.similarity_compressor_threshold()),
m_result_rules(ctx.get_rule_manager()),
m_modified(false),
m_pinned(m_manager) {
SASSERT(m_threshold_count>1);
}
void mk_similarity_compressor::reset() {
m_rules.reset();
m_result_rules.reset();
m_pinned.reset();
}
/**
Allows to traverse head and positive tails in a single for loop starting from -1
*/
static app * get_by_tail_index(rule * r, int idx) {
if (idx < 0) {
return r->get_head();
}
SASSERT(idx < static_cast<int>(r->get_positive_tail_size()));
return r->get_tail(idx);
}
template<typename T>
static int aux_compare(T a, T b) {
return (a>b) ? 1 : ( (a==b) ? 0 : -1);
}
template<typename T>
static int aux_compare(T* a, T* b);
static int compare_var_args(app* t1, app* t2) {
SASSERT(t1->get_num_args()==t2->get_num_args());
int res;
unsigned n = t1->get_num_args();
for (unsigned i = 0; i < n; i++) {
expr * a1 = t1->get_arg(i);
expr * a2 = t2->get_arg(i);
res = aux_compare(is_var(a1), is_var(a2));
if (res != 0) {
return res;
}
if (is_var(a1)) {
res = aux_compare(to_var(a1)->get_idx(), to_var(a2)->get_idx());
if (res != 0) {
return res;
}
}
}
return 0;
}
static int compare_args(app* t1, app* t2, int & skip_countdown) {
SASSERT(t1->get_num_args()==t2->get_num_args());
int res;
unsigned n = t1->get_num_args();
for (unsigned i=0; i<n; i++) {
if (is_var(t1->get_arg(i))) {
SASSERT(t1->get_arg(i) == t2->get_arg(i));
continue;
}
if ((skip_countdown--) == 0) {
continue;
}
res = aux_compare(t1->get_arg(i)->get_id(), t2->get_arg(i)->get_id());
if (res!=0) { return res; }
}
return 0;
}
/**
\brief Return 0 if r1 and r2 could be similar. If the rough similarity
equaivelance class of r1 is greater than the one of r2, return 1; otherwise return -1.
Two rules are in the same rough similarity class if they differ only in constant arguments
of positive uninterpreted predicates.
*/
static int rough_compare(rule * r1, rule * r2) {
int res = aux_compare(r1->get_tail_size(), r2->get_tail_size());
if (res!=0) { return res; }
res = aux_compare(r1->get_uninterpreted_tail_size(), r2->get_uninterpreted_tail_size());
if (res!=0) { return res; }
res = aux_compare(r1->get_positive_tail_size(), r2->get_positive_tail_size());
if (res!=0) { return res; }
int pos_tail_sz = r1->get_positive_tail_size();
for (int i=-1; i<pos_tail_sz; i++) {
app * t1 = get_by_tail_index(r1, i);
app * t2 = get_by_tail_index(r2, i);
res = aux_compare(t1->get_decl()->get_id(), t2->get_decl()->get_id());
if (res!=0) { return res; }
res = compare_var_args(t1, t2);
if (res!=0) { return res; }
}
unsigned tail_sz = r1->get_tail_size();
for (unsigned i=pos_tail_sz; i<tail_sz; i++) {
res = aux_compare(r1->get_tail(i)->get_id(), r2->get_tail(i)->get_id());
if (res!=0) { return res; }
}
return 0;
}
/**
\c r1 and \c r2 must be equal according to the \c rough_compare function for this function
to be called.
*/
static int total_compare(rule * r1, rule * r2, int skipped_arg_index = INT_MAX) {
SASSERT(rough_compare(r1, r2)==0);
int pos_tail_sz = r1->get_positive_tail_size();
for (int i=-1; i<pos_tail_sz; i++) {
int res = compare_args(get_by_tail_index(r1, i), get_by_tail_index(r2, i), skipped_arg_index);
if (res!=0) { return res; }
}
return 0;
}
class const_info {
int m_tail_index;
unsigned m_arg_index;
bool m_has_parent;
/** Parent is a constant that appears earlier in the rule and has always the same value
as this constant. */
unsigned m_parent_index;
public:
const_info(int tail_index, unsigned arg_index)
: m_tail_index(tail_index), m_arg_index(arg_index), m_has_parent(false) {}
int tail_index() const { return m_tail_index; }
unsigned arg_index() const { return m_arg_index; }
bool has_parent() const { return m_has_parent; }
unsigned parent_index() const { SASSERT(has_parent()); return m_parent_index; }
void set_parent_index(unsigned idx) {
SASSERT(!m_has_parent);
m_has_parent = true;
m_parent_index = idx;
}
};
typedef svector<const_info> info_vector;
static void collect_const_indexes(app * t, int tail_index, info_vector & res) {
unsigned n = t->get_num_args();
for (unsigned i=0; i<n; i++) {
if (is_var(t->get_arg(i))) {
continue;
}
res.push_back(const_info(tail_index, i));
}
}
static void collect_const_indexes(rule * r, info_vector & res) {
collect_const_indexes(r->get_head(), -1, res);
unsigned pos_tail_sz = r->get_positive_tail_size();
for (unsigned i=0; i<pos_tail_sz; i++) {
collect_const_indexes(r->get_tail(i), i, res);
}
}
template<class T>
static void collect_orphan_consts(rule * r, const info_vector & const_infos, T & tgt) {
unsigned const_cnt = const_infos.size();
tgt.reset();
for (unsigned i=0; i<const_cnt; i++) {
const_info inf = const_infos[i];
if (inf.has_parent()) {
continue;
}
app * pred = get_by_tail_index(r, inf.tail_index());
tgt.push_back(to_app(pred->get_arg(inf.arg_index())));
SASSERT(tgt.back()->get_num_args()==0);
}
}
template<class T>
static void collect_orphan_sorts(rule * r, const info_vector & const_infos, T & tgt) {
unsigned const_cnt = const_infos.size();
tgt.reset();
for (unsigned i=0; i<const_cnt; i++) {
const_info inf = const_infos[i];
if (inf.has_parent()) {
continue;
}
app * pred = get_by_tail_index(r, inf.tail_index());
tgt.push_back(pred->get_decl()->get_domain(inf.arg_index()));
}
}
/**
\brief From the \c tail_indexes and \c arg_indexes remove elements corresponding to constants
that are the same in rules \c *first ... \c *(after_last-1).
*/
static void remove_stable_constants(rule_vector::iterator first, rule_vector::iterator after_last,
info_vector & const_infos) {
SASSERT(after_last-first>1);
unsigned const_cnt = const_infos.size();
ptr_vector<app> vals;
rule * r = *(first++);
collect_orphan_consts(r, const_infos, vals);
SASSERT(vals.size()==const_cnt);
rule_vector::iterator it = first;
for (; it!=after_last; ++it) {
for (unsigned i=0; i<const_cnt; i++) {
app * pred = get_by_tail_index(*it, const_infos[i].tail_index());
app * val = to_app(pred->get_arg(const_infos[i].arg_index()));
if (vals[i]!=val) {
vals[i] = 0;
}
}
}
unsigned removed_cnt = 0;
for (unsigned i=0; i<const_cnt; i++) {
if (vals[i]!=0) {
removed_cnt++;
}
else if (removed_cnt!=0) {
const_infos[i-removed_cnt] = const_infos[i];
}
}
if (removed_cnt!=0) {
const_infos.shrink(const_cnt-removed_cnt);
}
}
/**
\brief When function returns, \c parents will contain for each constant the index of the
first constant that is equal to it in all the rules. If there is no such, it will contain
its own index.
*/
static void detect_equal_constants(rule_vector::iterator first, rule_vector::iterator after_last,
info_vector & const_infos) {
SASSERT(first!=after_last);
unsigned const_cnt = const_infos.size();
ptr_vector<app> vals;
ptr_vector<sort> sorts;
rule * r = *(first++);
collect_orphan_consts(r, const_infos, vals);
collect_orphan_sorts(r, const_infos, sorts);
SASSERT(vals.size()==const_cnt);
vector<unsigned_vector> possible_parents(const_cnt);
for (unsigned i=1; i<const_cnt; i++) {
for (unsigned j=0; j<i; j++) {
if (vals[i]==vals[j] && sorts[i]==sorts[j]) {
possible_parents[i].push_back(j);
}
}
}
rule_vector::iterator it = first;
for (; it!=after_last; ++it) {
collect_orphan_consts(*it, const_infos, vals);
for (unsigned i=1; i<const_cnt; i++) {
unsigned_vector & ppars = possible_parents[i];
unsigned j=0;
while(j<ppars.size()) {
if (vals[i]!=vals[ppars[j]]) {
ppars[j] = ppars.back();
ppars.pop_back();
}
else {
j++;
}
}
}
}
for (unsigned i=0; i<const_cnt; i++) {
unsigned parent = i;
unsigned_vector & ppars = possible_parents[i];
unsigned ppars_sz = ppars.size();
for (unsigned j=0; j<ppars_sz; j++) {
if (ppars[j]<parent) {
parent = ppars[j];
}
}
if (parent!=i) {
const_infos[i].set_parent_index(parent);
}
}
}
static unsigned get_constant_count(rule * r) {
unsigned res = r->get_head()->get_num_args() - count_variable_arguments(r->get_head());
unsigned pos_tail_sz = r->get_positive_tail_size();
for (unsigned i=0; i<pos_tail_sz; i++) {
res+= r->get_tail(i)->get_num_args() - count_variable_arguments(r->get_tail(i));
}
return res;
}
static bool initial_comparator(rule * r1, rule * r2) {
int res = rough_compare(r1, r2);
if (res!=0) { return res>0; }
return total_compare(r1, r2)>0;
}
class arg_ignoring_comparator {
unsigned m_ignored_index;
public:
arg_ignoring_comparator(unsigned ignored_index) : m_ignored_index(ignored_index) {}
bool operator()(rule * r1, rule * r2) const {
return total_compare(r1, r2, m_ignored_index)>0;
}
bool eq(rule * r1, rule * r2) const {
return total_compare(r1, r2, m_ignored_index)==0;
}
};
void mk_similarity_compressor::merge_class(rule_vector::iterator first,
rule_vector::iterator after_last) {
SASSERT(after_last-first>1);
info_vector const_infos;
rule * r = *first; //an arbitrary representative of the class
collect_const_indexes(r, const_infos);
remove_stable_constants(first, after_last, const_infos);
unsigned const_cnt = const_infos.size();
SASSERT(const_cnt>0);
detect_equal_constants(first, after_last, const_infos);
//The aux relation contains column for each constant which does not have an earlier constant
//that it is equal to (i.e. only has no parent)
ptr_vector<sort> aux_domain;
collect_orphan_sorts(r, const_infos, aux_domain);
func_decl* head_pred = r->get_decl();
symbol const& name_prefix = head_pred->get_name();
std::string name_suffix = "sc_" + to_string(const_cnt);
func_decl * aux_pred = m_context.mk_fresh_head_predicate(name_prefix, symbol(name_suffix.c_str()),
aux_domain.size(), aux_domain.c_ptr(), head_pred);
m_pinned.push_back(aux_pred);
relation_fact val_fact(m_manager, const_cnt);
rule_vector::iterator it = first;
for (; it!=after_last; ++it) {
collect_orphan_consts(*it, const_infos, val_fact);
m_context.add_fact(aux_pred, val_fact);
}
m_context.get_rel_context()->get_rmanager().mark_saturated(aux_pred);
app * new_head = r->get_head();
ptr_vector<app> new_tail;
svector<bool> new_negs;
unsigned tail_sz = r->get_tail_size();
for (unsigned i=0; i<tail_sz; i++) {
new_tail.push_back(r->get_tail(i));
new_negs.push_back(r->is_neg_tail(i));
}
rule_counter ctr;
ctr.count_rule_vars(m_manager, r);
unsigned max_var_idx, new_var_idx_base;
if (ctr.get_max_positive(max_var_idx)) {
new_var_idx_base = max_var_idx+1;
}
else {
new_var_idx_base = 0;
}
ptr_vector<var> const_vars; //variables at indexes of their corresponding constants
expr_ref_vector aux_vars(m_manager); //variables as arguments for the auxiliary predicate
unsigned aux_column_index = 0;
for (unsigned i=0; i<const_cnt; ) {
int tail_idx = const_infos[i].tail_index();
app * & mod_tail = (tail_idx==-1) ? new_head : new_tail[tail_idx];
ptr_vector<expr> mod_args(mod_tail->get_num_args(), mod_tail->get_args());
for (; i<const_cnt && const_infos[i].tail_index()==tail_idx; i++) { //here the outer loop counter is modified
const_info & inf = const_infos[i];
var * mod_var;
if (!inf.has_parent()) {
mod_var = m_manager.mk_var(new_var_idx_base+aux_column_index,
aux_domain[aux_column_index]);
aux_column_index++;
aux_vars.push_back(mod_var);
}
else {
mod_var = const_vars[inf.parent_index()];
}
const_vars.push_back(mod_var);
mod_args[inf.arg_index()] = mod_var;
}
app * upd_tail = m_manager.mk_app(mod_tail->get_decl(), mod_args.c_ptr());
m_pinned.push_back(upd_tail);
mod_tail = upd_tail;
}
app_ref aux_tail(m_manager.mk_app(aux_pred, aux_vars.c_ptr()), m_manager);
new_tail.push_back(aux_tail);
new_negs.push_back(false);
rule * new_rule = m_context.get_rule_manager().mk(new_head, new_tail.size(), new_tail.c_ptr(),
new_negs.c_ptr());
m_result_rules.push_back(new_rule);
//TODO: allow for a rule to have multiple parent objects
new_rule->set_accounting_parent_object(m_context, r);
m_modified = true;
}
void mk_similarity_compressor::process_class(rule_set const& source, rule_vector::iterator first,
rule_vector::iterator after_last) {
SASSERT(first!=after_last);
//remove duplicates
{
rule_vector::iterator it = first;
rule_vector::iterator prev = it;
++it;
while(it!=after_last) {
if (it!=after_last && total_compare(*prev, *it)==0) {
--after_last;
std::swap(*it, *after_last);
m_modified = true;
}
else {
prev = it;
++it;
}
}
}
SASSERT(first!=after_last);
unsigned const_cnt = get_constant_count(*first);
#if 0
for (unsigned ignored_index=0; ignored_index<const_cnt; ignored_index++) {
arg_ignoring_comparator comparator(ignored_index);
std::sort(first, after_last, comparator);
rule_vector::iterator it = first;
rule_vector::iterator grp_begin = it;
unsigned grp_size=0;
while(it!=after_last) {
rule_vector::iterator prev = it;
++it;
grp_size++;
if (it==after_last || !comparator.eq(*prev, *it)) {
if (grp_size>m_threshold_count) {
merge_class(grp_begin, it);
//group was processed, so we remove it from the class
if (it==after_last) {
after_last=grp_begin;
it=after_last;
}
else {
while(it!=grp_begin) {
std::swap(*--it, *--after_last);
}
}
}
grp_begin = it;
grp_size = 0;
}
}
}
#endif
//TODO: compress also rules with pairs (or tuples) of equal constants
#if 1
if (const_cnt>0 && !source.is_output_predicate((*first)->get_decl())) {
unsigned rule_cnt = static_cast<unsigned>(after_last-first);
if (rule_cnt>m_threshold_count) {
merge_class(first, after_last);
return;
}
}
#endif
//put rules which weren't merged into result
rule_vector::iterator it = first;
for (; it!=after_last; ++it) {
m_result_rules.push_back(*it);
}
}
rule_set * mk_similarity_compressor::operator()(rule_set const & source) {
// TODO mc
m_modified = false;
unsigned init_rule_cnt = source.get_num_rules();
SASSERT(m_rules.empty());
for (unsigned i=0; i<init_rule_cnt; i++) {
m_rules.push_back(source.get_rule(i));
}
std::sort(m_rules.begin(), m_rules.end(), initial_comparator);
rule_vector::iterator it = m_rules.begin();
rule_vector::iterator end = m_rules.end();
rule_vector::iterator cl_begin = it;
while(it!=end) {
rule_vector::iterator prev = it;
++it;
if (it==end || rough_compare(*prev, *it)!=0) {
process_class(source, cl_begin, it);
cl_begin = it;
}
}
rule_set * result = static_cast<rule_set *>(0);
if (m_modified) {
result = alloc(rule_set, m_context);
unsigned fin_rule_cnt = m_result_rules.size();
for (unsigned i=0; i<fin_rule_cnt; i++) {
result->add_rule(m_result_rules.get(i));
}
result->inherit_predicates(source);
}
reset();
return result;
}
};

View file

@ -0,0 +1,78 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_similarity_compressor.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-10-22.
Revision History:
--*/
#ifndef _DL_MK_SIMILARITY_COMPRESSOR_H_
#define _DL_MK_SIMILARITY_COMPRESSOR_H_
#include<utility>
#include"map.h"
#include"obj_pair_hashtable.h"
#include"dl_context.h"
#include"dl_rule_set.h"
#include"dl_rule_transformer.h"
namespace datalog {
/**
\brief Functor for merging groups of similar rules.
A rule sequence
P("1",x):-Q(x).
...
P("N",x):-Q(x).
will be replaced by
P(y,x):-Q(x), Aux(y).
and a set of facts
Aux("1").
...
Aux("N").
Similar transformation is performed when the varying constant appears in the positive tail.
*/
class mk_similarity_compressor : public rule_transformer::plugin {
context & m_context;
ast_manager & m_manager;
/** number of similar rules necessary for a group to be introduced */
unsigned m_threshold_count;
rule_vector m_rules;
rule_ref_vector m_result_rules;
bool m_modified;
ast_ref_vector m_pinned;
void merge_class(rule_vector::iterator first, rule_vector::iterator after_last);
void process_class(rule_set const& source, rule_vector::iterator first, rule_vector::iterator after_last);
void reset();
public:
mk_similarity_compressor(context & ctx);
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_SIMILARITY_COMPRESSOR_H_ */

View file

@ -0,0 +1,741 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_simple_joins.cpp
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2010-05-20.
Revision History:
--*/
#include<utility>
#include<sstream>
#include<limits>
#include"dl_mk_simple_joins.h"
#include"ast_pp.h"
#include"trace.h"
namespace datalog {
mk_simple_joins::mk_simple_joins(context & ctx):
plugin(1000),
m_context(ctx),
rm(ctx.get_rule_manager()) {
}
class join_planner {
typedef float cost;
class pair_info {
cost m_total_cost;
/**
\brief Number of rules longer than two that contain this pair.
This number is being updated by \c add_rule and \remove rule. Even though between
adding a rule and removing it, the length of a rule can decrease without this pair
being notified about it, it will surely see the decrease from length 3 to 2 which
the threshold for rule being counted in this counter.
*/
unsigned m_consumers;
bool m_stratified;
unsigned m_src_stratum;
public:
var_idx_set m_all_nonlocal_vars;
rule_vector m_rules;
pair_info() : m_consumers(0), m_stratified(true), m_src_stratum(0) {}
bool can_be_joined() const {
return m_consumers>0;
}
cost get_cost() const {
/*if(m_instantiated) {
return std::numeric_limits<cost>::min();
}*/
SASSERT(m_consumers>0);
cost amortized = m_total_cost/m_consumers;
if(m_stratified) {
return amortized * ( (amortized>0) ? (1/16.0f) : 16.0f);
}
else {
return amortized;
}
}
/**
\brief Add rule \c r among rules interested in current predicate pair.
The \c pl.m_rule_content entry of the rule has to be properly filled in
by the time of a call to this function
*/
void add_rule(join_planner & pl, app * t1, app * t2, rule * r,
const var_idx_set & non_local_vars_normalized) {
if(m_rules.empty()) {
m_total_cost = pl.compute_cost(t1, t2);
m_src_stratum = std::max(pl.get_stratum(t1->get_decl()), pl.get_stratum(t2->get_decl()));
}
m_rules.push_back(r);
if(pl.m_rules_content.find_core(r)->get_data().m_value.size()>2) {
m_consumers++;
}
if(m_stratified) {
unsigned head_stratum = pl.get_stratum(r->get_decl());
SASSERT(head_stratum>=m_src_stratum);
if(head_stratum==m_src_stratum) {
m_stratified = false;
}
}
idx_set_union(m_all_nonlocal_vars, non_local_vars_normalized);
}
/**
\brief Remove rule from the pair record. Return true if no rules remain
in the pair, and so it should be removed.
*/
bool remove_rule(rule * r, unsigned original_length) {
TRUSTME( remove_from_vector(m_rules, r) );
if(original_length>2) {
SASSERT(m_consumers>0);
m_consumers--;
}
SASSERT(!m_rules.empty() || m_consumers==0);
return m_rules.empty();
}
private:
pair_info & operator=(const pair_info &); //to avoid the implicit one
};
typedef std::pair<app*, app*> app_pair;
typedef map<app_pair, pair_info *,
pair_hash<obj_ptr_hash<app>, obj_ptr_hash<app> >, default_eq<app_pair> > cost_map;
typedef map<rule *, ptr_vector<app>, ptr_hash<rule>, ptr_eq<rule> > rule_pred_map;
context & m_context;
ast_manager & m;
rule_manager & rm;
var_subst & m_var_subst;
rule_set & m_rs_aux_copy; //reference to a rule_set that will allow to ask for stratum levels
cost_map m_costs;
ptr_vector<app> m_interpreted;
rule_pred_map m_rules_content;
rule_ref_vector m_introduced_rules;
ptr_hashtable<rule, ptr_hash<rule>, ptr_eq<rule> > m_modified_rules;
ast_ref_vector m_pinned;
mutable ptr_vector<sort> m_vars;
public:
join_planner(context & ctx, rule_set & rs_aux_copy)
: m_context(ctx), m(ctx.get_manager()),
rm(ctx.get_rule_manager()),
m_var_subst(ctx.get_var_subst()),
m_rs_aux_copy(rs_aux_copy),
m_introduced_rules(ctx.get_rule_manager()),
m_pinned(ctx.get_manager())
{
}
~join_planner()
{
cost_map::iterator it = m_costs.begin();
cost_map::iterator end = m_costs.end();
for (; it != end; ++it) {
dealloc(it->m_value);
}
m_costs.reset();
}
private:
void get_normalizer(app * t, unsigned & next_var, expr_ref_vector & result) const {
SASSERT(result.size()>0);
unsigned res_ofs = result.size()-1;
unsigned n=t->get_num_args();
for(unsigned i=0; i<n; i++) {
SASSERT(is_var(t->get_arg(i)));
var * v = to_var(t->get_arg(i));
unsigned var_idx = v->get_idx();
if(result[res_ofs-var_idx]==0) {
result[res_ofs-var_idx]=m.mk_var(next_var, v->get_sort());
next_var++;
}
}
}
void get_normalizer(app * t1, app * t2, expr_ref_vector & result) const {
SASSERT(result.empty());
if(t1->get_num_args()==0 && t2->get_num_args()==0) {
return; //nothing to normalize
}
SASSERT(!t1->is_ground() || !t2->is_ground());
unsigned max_var_idx = 0;
{
var_idx_set& orig_var_set = rm.collect_vars(t1, t2);
var_idx_set::iterator ovit = orig_var_set.begin();
var_idx_set::iterator ovend = orig_var_set.end();
for(; ovit!=ovend; ++ovit) {
unsigned var_idx = *ovit;
if(var_idx>max_var_idx) {
max_var_idx = var_idx;
}
}
}
if(t1->get_decl()!=t2->get_decl()) {
if(t1->get_decl()->get_id()<t2->get_decl()->get_id()) {
std::swap(t1, t2);
}
}
else {
int_vector norm1(max_var_idx+1, -1);
int_vector norm2(max_var_idx+1, -1);
unsigned n=t1->get_num_args();
SASSERT(n==t2->get_num_args());
for(unsigned i=0; i<n; i++) {
//We assume that the mk_simple_joins transformer is applied after mk_filter_rules,
//so the only literals which appear in pairs are the ones that contain only variables.
var * v1 = to_var(t1->get_arg(i));
var * v2 = to_var(t2->get_arg(i));
if(v1->get_sort()!=v2->get_sort()) {
//different sorts mean we can distinguish the two terms
if(v1->get_sort()->get_id()<v2->get_sort()->get_id()) {
std::swap(t1, t2);
}
break;
}
unsigned v1_idx = v1->get_idx();
unsigned v2_idx = v2->get_idx();
//since the rules already went through the mk_filter_rules transformer,
//variables must be linear
SASSERT(norm1[v1_idx]==-1);
SASSERT(norm2[v2_idx]==-1);
if(norm2[v1_idx]!=norm1[v2_idx]) {
//now we can distinguish the two terms
if(norm2[v1_idx]<norm1[v2_idx]) {
std::swap(t1, t2);
}
break;
}
norm1[v1_idx]=i;
norm2[v2_idx]=i;
}
//if we did not exit the loop prematurely, the two terms are indistinguishable,
//so the order should not matter
}
result.resize(max_var_idx+1, static_cast<expr *>(0));
unsigned next_var = 0;
get_normalizer(t1, next_var, result);
get_normalizer(t2, next_var, result);
}
app_pair get_key(app * t1, app * t2) {
expr_ref_vector norm_subst(m);
get_normalizer(t1, t2, norm_subst);
expr_ref t1n_ref(m);
expr_ref t2n_ref(m);
m_var_subst(t1, norm_subst.size(), norm_subst.c_ptr(), t1n_ref);
m_var_subst(t2, norm_subst.size(), norm_subst.c_ptr(), t2n_ref);
app * t1n = to_app(t1n_ref);
app * t2n = to_app(t2n_ref);
if(t1n>t2n) {
std::swap(t1n, t2n);
}
m_pinned.push_back(t1n);
m_pinned.push_back(t2n);
/*
IF_VERBOSE(0,
print_renaming(norm_subst, verbose_stream());
display_predicate(m_context, t1, verbose_stream());
display_predicate(m_context, t2, verbose_stream());
display_predicate(m_context, t1n, verbose_stream());
display_predicate(m_context, t2n, verbose_stream()););
*/
return app_pair(t1n, t2n);
}
/**
\brief Add rule \c r among rules interested in predicate pair \c t1, \c t2.
The \c m_rule_content entry of the rule \c r has to be properly filled in
by the time of a call to this function
*/
void register_pair(app * t1, app * t2, rule * r, const var_idx_set & non_local_vars) {
TRACE("dl", tout << mk_pp(t1, m) << " " << mk_pp(t2, m) << "\n";
r->display(m_context, tout); tout << "\n";);
SASSERT(t1!=t2);
cost_map::entry * e = m_costs.insert_if_not_there2(get_key(t1, t2), 0);
pair_info * & ptr_inf = e->get_data().m_value;
if(ptr_inf==0) {
ptr_inf = alloc(pair_info);
}
pair_info & inf = *ptr_inf;
expr_ref_vector normalizer(m);
get_normalizer(t1, t2, normalizer);
unsigned norm_ofs = normalizer.size()-1;
var_idx_set normalized_vars;
var_idx_set::iterator vit = non_local_vars.begin();
var_idx_set::iterator vend = non_local_vars.end();
for(; vit!=vend; ++vit) {
unsigned norm_var = to_var(normalizer.get(norm_ofs-*vit))->get_idx();
normalized_vars.insert(norm_var);
}
inf.add_rule(*this, t1, t2, r, normalized_vars);
}
pair_info & get_pair(app_pair key) const {
cost_map::entry * e = m_costs.find_core(key);
SASSERT(e);
return *e->get_data().m_value;
}
void remove_rule_from_pair(app_pair key, rule * r, unsigned original_len) {
pair_info * ptr = &get_pair(key);
if(ptr->remove_rule(r, original_len)) {
SASSERT(ptr->m_rules.empty());
m_costs.remove(key);
dealloc(ptr);
}
}
void register_rule(rule * r) {
rule_counter counter;
counter.count_rule_vars(m, r, 1);
ptr_vector<app> & rule_content =
m_rules_content.insert_if_not_there2(r, ptr_vector<app>())->get_data().m_value;
SASSERT(rule_content.empty());
unsigned pos_tail_size=r->get_positive_tail_size();
for(unsigned i=0; i<pos_tail_size; i++) {
rule_content.push_back(r->get_tail(i));
}
for(unsigned i=0; i<pos_tail_size; i++) {
app * t1 = r->get_tail(i);
var_idx_set t1_vars = rm.collect_vars(t1);
counter.count_vars(m, t1, -1); //temporarily remove t1 variables from counter
for(unsigned j=i+1; j<pos_tail_size; j++) {
app * t2 = r->get_tail(j);
counter.count_vars(m, t2, -1); //temporarily remove t2 variables from counter
var_idx_set scope_vars = rm.collect_vars(t2);
scope_vars |= t1_vars;
var_idx_set non_local_vars;
counter.collect_positive(non_local_vars);
counter.count_vars(m, t2, 1); //restore t2 variables in counter
set_intersection(non_local_vars, scope_vars);
register_pair(t1, t2, r, non_local_vars);
}
counter.count_vars(m, t1, 1); //restore t1 variables in counter
}
}
bool extract_argument_info(unsigned var_idx, app * t, expr_ref_vector & args,
ptr_vector<sort> & domain) {
unsigned n=t->get_num_args();
for(unsigned i=0; i<n; i++) {
var * v=to_var(t->get_arg(i));
if(v->get_idx()==var_idx) {
args.push_back(v);
domain.push_back(m.get_sort(v));
return true;
}
}
return false;
}
void join_pair(app_pair pair_key) {
app * t1 = pair_key.first;
app * t2 = pair_key.second;
pair_info & inf = get_pair(pair_key);
SASSERT(!inf.m_rules.empty());
var_idx_set & output_vars = inf.m_all_nonlocal_vars;
expr_ref_vector args(m);
ptr_vector<sort> domain;
unsigned arity = output_vars.num_elems();
idx_set::iterator ovit=output_vars.begin();
idx_set::iterator ovend=output_vars.end();
//TODO: improve quadratic complexity
for(;ovit!=ovend;++ovit) {
unsigned var_idx=*ovit;
bool found=extract_argument_info(var_idx, t1, args, domain);
if(!found) {
found=extract_argument_info(var_idx, t2, args, domain);
}
SASSERT(found);
}
SASSERT(args.size()==arity);
SASSERT(domain.size()==arity);
rule * one_parent = inf.m_rules.back();
func_decl* parent_head = one_parent->get_decl();
const char * one_parent_name = parent_head->get_name().bare_str();
std::string parent_name;
if(inf.m_rules.size()>1) {
parent_name = one_parent_name + std::string("_and_") + to_string(inf.m_rules.size()-1);
}
else {
parent_name = one_parent_name;
}
func_decl * decl = m_context.mk_fresh_head_predicate(
symbol(parent_name.c_str()), symbol("split"),
arity, domain.c_ptr(), parent_head);
app_ref head(m.mk_app(decl, arity, args.c_ptr()), m);
app * tail[] = {t1, t2};
rule * new_rule = m_context.get_rule_manager().mk(head, 2, tail, 0);
//TODO: update accounting so that it can handle multiple parents
new_rule->set_accounting_parent_object(m_context, one_parent);
m_introduced_rules.push_back(new_rule);
//here we copy the inf.m_rules vector because inf.m_rules will get changed
//in the iteration. Also we use hashtable instead of vector because we do
//not want to process one rule twice.
typedef ptr_hashtable<rule, ptr_hash<rule>, default_eq<rule *> > rule_hashtable;
rule_hashtable relevant_rules;
insert_into_set(relevant_rules, inf.m_rules);
rule_hashtable::iterator rit = relevant_rules.begin();
rule_hashtable::iterator rend = relevant_rules.end();
for(; rit!=rend; ++rit) {
apply_binary_rule(*rit, pair_key, head);
}
// SASSERT(!m_costs.contains(pair_key));
}
void replace_edges(rule * r, const ptr_vector<app> & removed_tails,
const ptr_vector<app> & added_tails0, const ptr_vector<app> & rule_content) {
SASSERT(removed_tails.size()>=added_tails0.size());
unsigned len = rule_content.size();
unsigned original_len = len+removed_tails.size()-added_tails0.size();
ptr_vector<app> added_tails(added_tails0); //we need a copy since we'll be modifying it
unsigned rt_sz = removed_tails.size();
//remove edges between removed tails
for(unsigned i=0; i<rt_sz; i++) {
for(unsigned j=i+1; j<rt_sz; j++) {
app_pair pair_key = get_key(removed_tails[i], removed_tails[j]);
remove_rule_from_pair(pair_key, r, original_len);
}
}
//remove edges between surviving tails and removed tails
for(unsigned i=0; i<len; i++) {
if(added_tails.contains(rule_content[i])) {
continue;
}
for(unsigned ri=0; ri<rt_sz; ri++) {
app_pair pair_key = get_key(rule_content[i], removed_tails[ri]);
remove_rule_from_pair(pair_key, r, original_len);
}
}
if(len==1) {
return;
}
app * head = r->get_head();
var_counter counter;
counter.count_vars(m, head, 1);
unsigned tail_size=r->get_tail_size();
unsigned pos_tail_size=r->get_positive_tail_size();
for(unsigned i=pos_tail_size; i<tail_size; i++) {
counter.count_vars(m, r->get_tail(i), 1);
}
for(unsigned i=0; i<len; i++) {
counter.count_vars(m, rule_content[i], 1);
}
//add edges that contain added tails
while(!added_tails.empty()) {
app * a_tail = added_tails.back(); //added tail
var_idx_set a_tail_vars = rm.collect_vars(a_tail);
counter.count_vars(m, a_tail, -1); //temporarily remove a_tail variables from counter
for(unsigned i=0; i<len; i++) {
app * o_tail = rule_content[i]; //other tail
if(added_tails.contains(o_tail)) {
//this avoids adding edges between new tails twice
continue;
}
counter.count_vars(m, o_tail, -1); //temporarily remove o_tail variables from counter
var_idx_set scope_vars = rm.collect_vars(o_tail);
scope_vars |= a_tail_vars;
var_idx_set non_local_vars;
counter.collect_positive(non_local_vars);
counter.count_vars(m, o_tail, 1); //restore o_tail variables in counter
set_intersection(non_local_vars, scope_vars);
register_pair(o_tail, a_tail, r, non_local_vars);
}
counter.count_vars(m, a_tail, 1); //restore t1 variables in counter
added_tails.pop_back();
}
}
void apply_binary_rule(rule * r, app_pair pair_key, app * t_new) {
app * t1 = pair_key.first;
app * t2 = pair_key.second;
ptr_vector<app> & rule_content = m_rules_content.find_core(r)->get_data().m_value;
unsigned len = rule_content.size();
if(len==1) {
return;
}
func_decl * t1_pred = t1->get_decl();
func_decl * t2_pred = t2->get_decl();
ptr_vector<app> removed_tails;
ptr_vector<app> added_tails;
for(unsigned i1=0; i1<len; i1++) {
app * rt1 = rule_content[i1];
if(rt1->get_decl()!=t1_pred) {
continue;
}
unsigned i2start = (t1_pred==t2_pred) ? (i1+1) : 0;
for(unsigned i2=i2start; i2<len; i2++) {
app * rt2 = rule_content[i2];
if(i1==i2 || rt2->get_decl()!=t2_pred) {
continue;
}
if(get_key(rt1, rt2)!=pair_key) {
continue;
}
expr_ref_vector normalizer(m);
get_normalizer(rt1, rt2, normalizer);
expr_ref_vector denormalizer(m);
reverse_renaming(m, normalizer, denormalizer);
expr_ref new_transf(m);
m_var_subst(t_new, denormalizer.size(), denormalizer.c_ptr(), new_transf);
app * new_lit = to_app(new_transf);
m_pinned.push_back(new_lit);
rule_content[i1]=new_lit;
rule_content[i2]=rule_content.back();
rule_content.pop_back();
len--; //here the bound of both loops changes!!!
removed_tails.push_back(rt1);
removed_tails.push_back(rt2);
added_tails.push_back(new_lit);
//this exits the inner loop, the outer one continues in case there will
//be other matches
break;
}
}
SASSERT(!removed_tails.empty());
SASSERT(!added_tails.empty());
m_modified_rules.insert(r);
replace_edges(r, removed_tails, added_tails, rule_content);
}
cost get_domain_size(func_decl * pred, unsigned arg_index) const {
relation_sort sort = pred->get_domain(arg_index);
return static_cast<cost>(m_context.get_sort_size_estimate(sort));
//unsigned sz;
//if(!m_context.get_sort_size(sort, sz)) {
// sz=UINT_MAX;
//}
//return static_cast<cost>(sz);
}
unsigned get_stratum(func_decl * pred) const {
return m_rs_aux_copy.get_predicate_strat(pred);
}
cost estimate_size(app * t) const {
func_decl * pred = t->get_decl();
unsigned n=pred->get_arity();
rel_context* rel = m_context.get_rel_context();
if (!rel) {
return cost(1);
}
relation_manager& rm = rel->get_rmanager();
if( (m_context.saturation_was_run() && rm.try_get_relation(pred))
|| rm.is_saturated(pred)) {
SASSERT(rm.try_get_relation(pred)); //if it is saturated, it should exist
unsigned rel_size_int = rel->get_relation(pred).get_size_estimate_rows();
if(rel_size_int!=0) {
cost rel_size = static_cast<cost>(rel_size_int);
cost curr_size = rel_size;
for(unsigned i=0; i<n; i++) {
if(!is_var(t->get_arg(i))) {
curr_size /= get_domain_size(pred, i);
}
}
return curr_size;
}
}
cost res = 1;
for(unsigned i=0; i<n; i++) {
if(is_var(t->get_arg(i))) {
res *= get_domain_size(pred, i);
}
}
return res;
}
cost compute_cost(app * t1, app * t2) const {
func_decl * t1_pred = t1->get_decl();
func_decl * t2_pred = t2->get_decl();
cost inters_size = 1;
variable_intersection vi(m_context.get_manager());
vi.populate(t1, t2);
unsigned n = vi.size();
for(unsigned i=0; i<n; i++) {
unsigned arg_index1, arg_index2;
vi.get(i, arg_index1, arg_index2);
inters_size *= get_domain_size(t1_pred, arg_index1);
//joined arguments must have the same domain
SASSERT(get_domain_size(t1_pred, arg_index1)==get_domain_size(t2_pred, arg_index2));
}
cost res = estimate_size(t1)*estimate_size(t2)/(inters_size*inters_size);
//cost res = -inters_size;
/*unsigned t1_strat = get_stratum(t1_pred);
SASSERT(t1_strat<=m_head_stratum);
if(t1_strat<m_head_stratum) {
unsigned t2_strat = get_stratum(t2_pred);
SASSERT(t2_strat<=m_head_stratum);
if(t2_strat<m_head_stratum) {
//the rule of this predicates would depend on predicates
//in lower stratum than the head, which is a good thing, since
//then the rule code will not need to appear in a loop
if(res>0) {
res /= 2;
}
else {
res *= 2;
}
}
}*/
TRACE("report_costs",
display_predicate(m_context, t1, tout);
display_predicate(m_context, t2, tout);
tout << res << "\n";);
return res;
}
bool pick_best_pair(app_pair & p) {
app_pair best;
bool found = false;
cost best_cost;
cost_map::iterator it = m_costs.begin();
cost_map::iterator end = m_costs.end();
for(; it!=end; ++it) {
app_pair key = it->m_key;
pair_info & inf = *it->m_value;
if(!inf.can_be_joined()) {
continue;
}
cost c = inf.get_cost();
if(!found || c<best_cost) {
found = true;
best_cost = c;
best = key;
}
}
if(!found) {
return false;
}
p=best;
return true;
}
public:
rule_set * run(rule_set const & source) {
unsigned num_rules = source.get_num_rules();
for (unsigned i = 0; i < num_rules; i++) {
register_rule(source.get_rule(i));
}
app_pair selected;
while(pick_best_pair(selected)) {
join_pair(selected);
}
if(m_modified_rules.empty()) {
return 0;
}
rule_set * result = alloc(rule_set, m_context);
rule_pred_map::iterator rcit = m_rules_content.begin();
rule_pred_map::iterator rcend = m_rules_content.end();
for(; rcit!=rcend; ++rcit) {
rule * orig_r = rcit->m_key;
ptr_vector<app> content = rcit->m_value;
SASSERT(content.size()<=2);
if(content.size()==orig_r->get_positive_tail_size()) {
//rule did not change
result->add_rule(orig_r);
continue;
}
ptr_vector<app> tail(content);
svector<bool> negs(tail.size(), false);
unsigned or_len = orig_r->get_tail_size();
for(unsigned i=orig_r->get_positive_tail_size(); i<or_len; i++) {
tail.push_back(orig_r->get_tail(i));
negs.push_back(orig_r->is_neg_tail(i));
}
rule * new_rule = m_context.get_rule_manager().mk(orig_r->get_head(), tail.size(), tail.c_ptr(),
negs.c_ptr());
new_rule->set_accounting_parent_object(m_context, orig_r);
m_context.get_rule_manager().mk_rule_rewrite_proof(*orig_r, *new_rule);
result->add_rule(new_rule);
}
while (!m_introduced_rules.empty()) {
result->add_rule(m_introduced_rules.back());
m_context.get_rule_manager().mk_rule_asserted_proof(*m_introduced_rules.back());
m_introduced_rules.pop_back();
}
result->inherit_predicates(source);
return result;
}
};
rule_set * mk_simple_joins::operator()(rule_set const & source) {
rule_set rs_aux_copy(m_context);
rs_aux_copy.replace_rules(source);
if(!rs_aux_copy.is_closed()) {
rs_aux_copy.close();
}
join_planner planner(m_context, rs_aux_copy);
return planner.run(source);
}
};

View file

@ -0,0 +1,63 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_simple_joins.h
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2010-05-20.
Revision History:
--*/
#ifndef _DL_MK_SIMPLE_JOINS_H_
#define _DL_MK_SIMPLE_JOINS_H_
#include"map.h"
#include"obj_pair_hashtable.h"
#include"dl_context.h"
#include"dl_rule_set.h"
#include"dl_rule_transformer.h"
namespace datalog {
/**
\brief Functor for creating rules that contain simple joins.
A simple join is the join of two tables.
After applying this transformation, every rule has at most one join.
So, the rules will have the form
HEAD :- TAIL.
HEAD :- TAIL_1, TAIL_2.
We also assume a rule may contain interpreted expressions that work as filtering conditions.
So, we may also have:
HEAD :- TAIL, C_1, ..., C_n.
HEAD :- TAIL_1, TAIL_2, C_1, ..., C_n.
Where the C_i's are interpreted expressions.
We say that a rule containing C_i's is a rule with a "big tail".
*/
class mk_simple_joins : public rule_transformer::plugin {
context & m_context;
rule_manager & rm;
public:
mk_simple_joins(context & ctx);
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_SIMPLE_JOINS_H_ */

859
src/muz/dl_mk_slice.cpp Normal file
View file

@ -0,0 +1,859 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_slice.cpp
Abstract:
<abstract>
Author:
Nikolaj Bjorner (nbjorner) 2012-9-12.
Revision History:
Consider a rule:
P(x,y) :- R(x,z), phi(x,y,z)
input: x, z
output: x, y
Let x_i, y_i, z_i be incides into the vectors x, y, z.
Suppose that positions in P and R are annotated with what is
slicable.
Sufficient conditions for sliceability:
x_i is sliceable if x_i does not appear in phi(x,y,z)
and the positions where x_i is used in P and R are sliceable
y_i is sliceable if y_i does not occur in phi(x,y,z), or
if it occurs in phi(x,y,z) it is only in one conjunct of the form
y_i = t[x_j,y_j,z_j]
and the positions where y_i is used in P and R are sliceable
z_i is sliceable if z_i does not occur in phi(x,y,z), or
if it occurs in phi(x,y,z) it is only in one conjunct of the form
y_i = t[x_j,y_j,z_i] where y_i is sliceable
and the positions where z_i is used in P and R are sliceable
A more refined approach may be using Gaussean elimination based
on x,z and eliminating variables from x,y (expressing them in terms
of a disjoint subeset of x,z).
--*/
#include "dl_mk_slice.h"
#include "ast_pp.h"
#include "expr_functors.h"
#include "dl_mk_rule_inliner.h"
#include "model_smt2_pp.h"
namespace datalog {
/**
Convert from sliced proofs to original proofs.
Given sliced rule
fml0: forall x y z u. p(x,y) & z = f(x,y) & phi(x,u) => p(u, z)
into
fml1: forall a b . q(a) & phi(a,b) => q(b)
It induces mappings:
theta: a |-> x, b |-> u
vars: x y z u.
predicates: -q(a) |-> p(x,y)
+q(b) |-> z = f(x,y) => p(u,z)
fml1 |-> fml0
The mapping theta is an injective function from variable indices
to variable indices. We can apply it as a substitution on expressions,
but we can also apply it as a transformation on substitutions. We
write theta[subst] when applying theta on substitution 'subst' such
that if [x |-> t] is in subst, then [theta(x) |-> theta(t)] is in
the result.
Given hyper-resolvent: fml1 subst1 fml2 subst2 |- fml3
where fml1 |-> fml1' with theta1
fml2 |-> fml2' with theta2
Perform the following steps:
1. [Convert fml1' fml2' to datalog rules because we have resolution routines]
2. Create subst1' := theta1[subst1]
subst2' := theta2[subst2]
3. Set fml1'' := subst1'(fml1')
fml2'' := subst2'(fml2')
4. Resolve fml1'' and fml2''
extract subst1'', subst2'' from resolvents.
extract goal fml3'
5. Create subst1''' := subst1'' o subst1'
subst2''' := subst2'' o subst2'
6. Return fml1'' subst1''' fml2'' subst2''' |- fml3'
7. Attach to fml3' the transformation ...?
*/
class mk_slice::slice_proof_converter : public proof_converter {
context& m_ctx;
ast_manager& m;
rule_manager& rm;
rule_ref_vector m_pinned_rules;
expr_ref_vector m_pinned_exprs;
obj_map<rule, rule*> m_rule2slice; // rule to sliced rule
obj_map<rule, unsigned_vector> m_renaming; // rule to renaming
obj_map<expr, rule*> m_sliceform2rule; // sliced formula to rule.
ptr_vector<proof> m_todo;
obj_map<proof, proof*> m_new_proof;
rule_unifier m_unifier;
slice_proof_converter(slice_proof_converter const& other);
void init_form2rule() {
if (!m_sliceform2rule.empty()) {
return;
}
obj_map<rule, rule*>::iterator it = m_rule2slice.begin();
obj_map<rule, rule*>::iterator end = m_rule2slice.end();
expr_ref fml(m);
for (; it != end; ++it) {
it->m_value->to_formula(fml);
m_pinned_exprs.push_back(fml);
TRACE("dl",
tout << "orig: " << mk_pp(fml, m) << "\n";
it->m_value->display(m_ctx, tout << "new:\n"););
m_sliceform2rule.insert(fml, it->m_key);
}
}
void translate_proof(proof_ref& pr) {
m_todo.reset();
m_new_proof.reset();
m_todo.push_back(pr);
while (!m_todo.empty()) {
proof* p = m_todo.back();
if (m_new_proof.contains(p)) {
m_todo.pop_back();
}
else if (translate_asserted(p)) {
// done
}
else if (translate_hyper_res(p)) {
// done
}
else {
m_new_proof.insert(p, p);
m_todo.pop_back();
TRACE("dl", tout << "unhandled proof term\n" << mk_pp(p, m) << "\n";);
}
}
pr = m_new_proof.find(pr);
}
bool translate_asserted(proof* p) {
expr* fact = 0;
rule* r = 0;
if (!m.is_asserted(p, fact)) {
return false;
}
if (!m_sliceform2rule.find(fact, r)) {
TRACE("dl", tout << "does not have fact\n" << mk_pp(fact, m) << "\n";);
return false;
}
proof_ref new_p(m);
new_p = r->get_proof();
m_pinned_exprs.push_back(new_p);
m_todo.pop_back();
m_new_proof.insert(p, new_p);
return true;
}
bool translate_hyper_res(proof* p) {
dl_decl_util util(m);
svector<std::pair<unsigned, unsigned> > positions;
expr_ref concl(m), slice_concl(m);
proof_ref_vector premises0(m);
vector<expr_ref_vector> substs, substs0;
if (!m.is_hyper_resolve(p, premises0, slice_concl, positions, substs0)) {
return false;
}
unsigned num_args = p->get_num_args();
SASSERT(num_args >= 2);
bool all_found = true;
for (unsigned i = 0; i < num_args-1; ++i) {
proof* arg = to_app(p->get_arg(i));
SASSERT(m.is_proof(arg));
if (!m_new_proof.contains(arg)) {
m_todo.push_back(arg);
all_found = false;
}
}
if (!all_found) {
return true;
}
ptr_vector<proof> premises;
proof* p0 = to_app(p->get_arg(0));
proof* p0_new = m_new_proof.find(p0);
expr* fact0 = m.get_fact(p0);
TRACE("dl", tout << "fact0: " << mk_pp(fact0, m) << "\n";);
rule* orig0;
if (!m_sliceform2rule.find(fact0, orig0)) {
return false;
}
premises.push_back(p0_new);
rule_ref r1(rm), r2(rm), r3(rm);
r1 = orig0;
substs.push_back(expr_ref_vector(m));
for (unsigned i = 1; i < num_args-1; ++i) {
proof* p1 = to_app(p->get_arg(i));
proof* p1_new = m_new_proof.find(p1);
expr* fact1 = m.get_fact(p1);
TRACE("dl", tout << "fact1: " << mk_pp(fact1, m) << "\n";);
rule* orig1 = 0;
if (!m_sliceform2rule.find(fact1, orig1)) {
return false;
}
premises.push_back(p1_new);
// TODO: work with substitutions.
r2 = orig1;
unsigned idx = 0; // brittle. TBD get index from positions.
VERIFY(m_unifier.unify_rules(*r1, idx, *r2));
m_unifier.apply(*r1.get(), idx, *r2.get(), r3);
expr_ref_vector const sub1 = m_unifier.get_rule_subst(*r1.get(), true);
for (unsigned j = 0; j < substs.size(); ++j) {
apply_subst(substs[j], sub1);
// size of substitutions may have grown...substs[j].resize(num_args[j]);
}
substs.push_back(m_unifier.get_rule_subst(*r2.get(), false));
TRACE("dl",
r1->display(m_ctx, tout << "rule1:");
r2->display(m_ctx, tout << "rule2:");
r3->display(m_ctx, tout << "res:"););
r1 = r3;
}
r1->to_formula(concl);
proof* new_p = m.mk_hyper_resolve(premises.size(), premises.c_ptr(), concl, positions, substs);
m_pinned_exprs.push_back(new_p);
m_pinned_rules.push_back(r1.get());
TRACE("dl",
tout << "orig: " << mk_pp(slice_concl, m) << "\n";
r1->display(m_ctx, tout << "new:"););
m_sliceform2rule.insert(slice_concl, r1.get());
m_rule2slice.insert(r1.get(), 0);
m_renaming.insert(r1.get(), unsigned_vector());
m_new_proof.insert(p, new_p);
m_todo.pop_back();
TRACE("dl", tout << "translated:\n" << mk_pp(p, m) << "\nto\n" << mk_pp(new_p, m) << "\n";);
return true;
}
public:
slice_proof_converter(context& ctx):
m_ctx(ctx),
m(ctx.get_manager()),
rm(ctx.get_rule_manager()),
m_pinned_rules(rm),
m_pinned_exprs(m),
m_unifier(ctx) {}
void insert(rule* orig_rule, rule* slice_rule, unsigned sz, unsigned const* renaming) {
m_rule2slice.insert(orig_rule, slice_rule);
m_pinned_rules.push_back(orig_rule);
m_pinned_rules.push_back(slice_rule);
m_renaming.insert(orig_rule, unsigned_vector(sz, renaming));
}
virtual void operator()(ast_manager& m, unsigned num_source, proof * const * source, proof_ref & result) {
SASSERT(num_source == 1);
result = source[0];
init_form2rule();
translate_proof(result);
}
virtual proof_converter * translate(ast_translation & translator) {
UNREACHABLE();
// this would require implementing translation for the dl_context.
return 0;
}
};
class mk_slice::slice_model_converter : public model_converter {
ast_manager& m;
obj_map<func_decl, func_decl*> m_slice2old;
obj_map<func_decl, bit_vector> m_sliceable;
ast_ref_vector m_pinned;
public:
slice_model_converter(mk_slice& parent, ast_manager& m): m(m), m_pinned(m) {}
void add_predicate(func_decl* old_f, func_decl* slice_f) {
m_pinned.push_back(old_f);
m_pinned.push_back(slice_f);
m_slice2old.insert(slice_f, old_f);
}
void add_sliceable(func_decl* f, bit_vector const& bv) {
m_pinned.push_back(f);
m_sliceable.insert(f, bv);
}
virtual void operator()(model_ref & md) {
if (m_slice2old.empty()) {
return;
}
TRACE("dl", model_smt2_pp(tout, m, *md, 0); );
model_ref old_model = alloc(model, m);
obj_map<func_decl, func_decl*>::iterator it = m_slice2old.begin();
obj_map<func_decl, func_decl*>::iterator end = m_slice2old.end();
for (; it != end; ++it) {
func_decl* old_p = it->m_value;
func_decl* new_p = it->m_key;
bit_vector const& is_sliced = m_sliceable.find(old_p);
SASSERT(is_sliced.size() == old_p->get_arity());
SASSERT(is_sliced.size() > new_p->get_arity());
func_interp* old_fi = alloc(func_interp, m, is_sliced.size());
TRACE("dl", tout << mk_pp(old_p, m) << " " << mk_pp(new_p, m) << "\n";
for (unsigned j = 0; j < is_sliced.size(); ++j) {
tout << (is_sliced.get(j)?"1":"0");
}
tout << "\n";);
if (new_p->get_arity() == 0) {
old_fi->set_else(md->get_const_interp(new_p));
}
else {
expr_ref_vector subst(m);
expr_ref tmp(m);
var_subst vs(m, false);
for (unsigned i = 0; i < is_sliced.size(); ++i) {
if (!is_sliced.get(i)) {
subst.push_back(m.mk_var(i, old_p->get_domain(i)));
}
}
func_interp* new_fi = md->get_func_interp(new_p);
if (!new_fi) {
TRACE("dl", tout << new_p->get_name() << " has no value in the current model\n";);
dealloc(old_fi);
continue;
}
if (!new_fi->is_partial()) {
TRACE("dl", tout << mk_pp(new_fi->get_else(), m) << "\n";);
vs(new_fi->get_else(), subst.size(), subst.c_ptr(), tmp);
old_fi->set_else(tmp);
}
unsigned num_entries = new_fi->num_entries();
for (unsigned j = 0; j < num_entries; ++j) {
expr_ref res(m);
expr_ref_vector args(m);
func_entry const* e = new_fi->get_entry(j);
for (unsigned k = 0, l = 0; k < old_p->get_arity(); ++k) {
if (!is_sliced.get(k)) {
vs(e->get_arg(l++), subst.size(), subst.c_ptr(), tmp);
args.push_back(tmp);
}
else {
args.push_back(m.mk_var(k, old_p->get_domain(k)));
}
SASSERT(l <= new_p->get_arity());
}
vs(e->get_result(), subst.size(), subst.c_ptr(), res);
old_fi->insert_entry(args.c_ptr(), res.get());
}
old_model->register_decl(old_p, old_fi);
}
}
// register values that have not been sliced.
unsigned sz = md->get_num_constants();
for (unsigned i = 0; i < sz; ++i) {
func_decl* c = md->get_constant(i);
if (!m_slice2old.contains(c)) {
old_model->register_decl(c, md->get_const_interp(c));
}
}
sz = md->get_num_functions();
for (unsigned i = 0; i < sz; ++i) {
func_decl* f = md->get_function(i);
if (!m_slice2old.contains(f)) {
func_interp* fi = md->get_func_interp(f);
old_model->register_decl(f, fi->copy());
}
}
md = old_model;
TRACE("dl", model_smt2_pp(tout, m, *md, 0); );
}
virtual model_converter * translate(ast_translation & translator) {
UNREACHABLE();
return 0;
}
};
mk_slice::mk_slice(context & ctx):
plugin(1),
m_ctx(ctx),
m(ctx.get_manager()),
rm(ctx.get_rule_manager()),
m_solved_vars(m),
m_pinned(m),
m_pc(0),
m_mc(0)
{}
bit_vector& mk_slice::get_predicate_slice(func_decl* h) {
if (!m_sliceable.contains(h)) {
bit_vector bv;
bv.resize(h->get_arity(), true);
m_sliceable.insert(h, bv);
}
return m_sliceable.find(h);
}
/**
\brief Saturate set of rules with respect to slicing criteria.
*/
void mk_slice::saturate(rule_set const& src) {
bool change = true;
while (change) {
change = false;
for (unsigned i = 0; i < src.get_num_rules(); ++i) {
change = prune_rule(*src.get_rule(i)) || change;
}
}
}
void mk_slice::filter_unique_vars(rule& r) {
//
// Variables that occur in multiple uinterpreted predicates are not sliceable.
//
uint_set used_vars;
for (unsigned j = 0; j < r.get_uninterpreted_tail_size(); ++j) {
app* p = r.get_tail(j);
for (unsigned i = 0; i < p->get_num_args(); ++i) {
expr* v = p->get_arg(i);
if (is_var(v)) {
unsigned vi = to_var(v)->get_idx();
add_var(vi);
if (used_vars.contains(vi)) {
m_var_is_sliceable[vi] = false;
}
else {
used_vars.insert(vi);
}
}
}
}
}
void mk_slice::solve_vars(rule& r, uint_set& used_vars, uint_set& parameter_vars) {
expr_ref_vector conjs = get_tail_conjs(r);
for (unsigned j = 0; j < conjs.size(); ++j) {
expr* e = conjs[j].get();
expr_ref r(m);
unsigned v;
if (is_eq(e, v, r) && is_output(v) && m_var_is_sliceable[v]) {
TRACE("dl", tout << "is_eq: " << mk_pp(e, m) << " " << (m_solved_vars[v].get()?"solved":"new") << "\n";);
add_var(v);
if (!m_solved_vars[v].get()) {
add_free_vars(parameter_vars, r);
m_solved_vars[v] = r;
}
else {
// variables can only be solved once.
add_free_vars(used_vars, e);
add_free_vars(used_vars, m_solved_vars[v].get());
used_vars.insert(v);
}
}
else {
add_free_vars(used_vars, e);
}
}
}
bool mk_slice::prune_rule(rule& r) {
TRACE("dl", r.display(m_ctx, tout << "prune:\n"); );
bool change = false;
init_vars(r);
//
// if a predicate in the body takes a constant as argument,
// the corresponding position is not sliceable.
//
for (unsigned j = 0; j < r.get_uninterpreted_tail_size(); ++j) {
app* p = r.get_tail(j);
bit_vector& bv = get_predicate_slice(p);
for (unsigned i = 0; i < p->get_num_args(); ++i) {
if (!is_var(p->get_arg(i)) && bv.get(i)) {
bv.unset(i);
change = true;
TRACE("dl", tout << "argument " << i << " is not a variable " << p->get_decl()->get_name() << "\n";);
}
}
}
filter_unique_vars(r);
//
// Collect the set of variables that are solved.
// Collect the occurrence count of the variables per conjunct.
//
uint_set used_vars, parameter_vars;
solve_vars(r, used_vars, parameter_vars);
uint_set::iterator it = used_vars.begin(), end = used_vars.end();
for (; it != end; ++it) {
if (*it < m_var_is_sliceable.size()) {
m_var_is_sliceable[*it] = false;
}
}
//
// Check if sliceable variables are either solved
// or are used to solve output sliceable variables, or
// don't occur in interpreted tail.
//
for (unsigned i = 0; i < num_vars(); ++i) {
if (!m_var_is_sliceable[i]) {
continue;
}
if (used_vars.contains(i)) {
m_var_is_sliceable[i] = false;
continue;
}
bool is_input = m_input[i];
bool is_output = m_output[i];
if (is_input && is_output) {
if (m_solved_vars[i].get()) {
m_var_is_sliceable[i] = false;
}
}
else if (is_output) {
if (parameter_vars.contains(i)) {
m_var_is_sliceable[i] = false;
}
}
else if (is_input) {
// I can be a parameter var, but not in used_vars.
}
else {
// variable does not correspond to
// any position in predicates.
}
}
//
// Update sliceable predicates based on slicing information of variables.
//
change = finalize_vars(r.get_head()) || change;
for (unsigned j = 0; j < r.get_uninterpreted_tail_size(); ++j) {
change = finalize_vars(r.get_tail(j)) || change;
}
return change;
}
bool mk_slice::is_eq(expr* e, unsigned& v, expr_ref& r) {
expr* c, *th, *el, *e1, *e2;
unsigned v1, v2;
expr_ref r1(m), r2(m);
if (m.is_ite(e, c, th, el)) {
if (is_eq(th, v1, r1) && is_eq(el, v2, r2) && v1 == v2) {
v = v1;
r = m.mk_ite(c, r1, r2);
return true;
}
}
if (is_var(e)) {
v = to_var(e)->get_idx();
r = m.mk_true();
return true;
}
if (m.is_not(e,e) && is_var(e)) {
v = to_var(e)->get_idx();
r = m.mk_false();
return true;
}
if (m.is_eq(e, e1, e2) && is_var(e1)) {
v = to_var(e1)->get_idx();
r = e2;
return true;
}
if (m.is_eq(e, e1, e2) && is_var(e2)) {
v = to_var(e2)->get_idx();
r = e1;
return true;
}
return false;
}
bool mk_slice::is_output(unsigned idx) {
return idx < m_output.size() && m_output[idx] && !m_input[idx];
}
bool mk_slice::is_output(expr* e) {
if (is_var(e)) {
return is_output(to_var(e)->get_idx());
}
else {
return false;
}
}
void mk_slice::init_vars(rule& r) {
m_input.reset();
m_output.reset();
m_var_is_sliceable.reset();
m_solved_vars.reset();
init_vars(r.get_head(), true, false);
for (unsigned j = 0; j < r.get_uninterpreted_tail_size(); ++j) {
init_vars(r.get_tail(j), false, r.is_neg_tail(j));
}
}
expr_ref_vector mk_slice::get_tail_conjs(rule const& r) {
expr_ref_vector conjs(m);
for (unsigned j = r.get_uninterpreted_tail_size(); j < r.get_tail_size(); ++j) {
conjs.push_back(r.get_tail(j));
}
qe::flatten_and(conjs);
return conjs;
}
void mk_slice::add_var(unsigned idx) {
if (idx >= m_input.size()) {
m_input.resize(idx+1, false);
m_output.resize(idx+1, false);
m_var_is_sliceable.resize(idx+1, true);
m_solved_vars.resize(idx+1);
}
}
void mk_slice::init_vars(app* p, bool is_output, bool is_neg_tail) {
bit_vector& bv = get_predicate_slice(p);
for (unsigned i = 0; i < p->get_num_args(); ++i) {
if (is_neg_tail) {
TRACE("dl", tout << "negated " << i << " in " << p->get_decl()->get_name() << "\n";);
bv.unset(i);
}
expr* arg = p->get_arg(i);
if (is_var(arg)) {
unsigned idx = to_var(arg)->get_idx();
add_var(idx);
if (is_output) {
m_output[idx] = true;
}
else {
m_input[idx] = true;
}
m_var_is_sliceable[idx] &= bv.get(i);
}
else {
SASSERT(m.is_value(arg));
if (!is_output) {
TRACE("dl", tout << "input " << i << " in " << p->get_decl()->get_name() << "\n";);
bv.unset(i);
}
}
}
}
bool mk_slice::finalize_vars(app* p) {
bool change = false;
bit_vector& bv = get_predicate_slice(p);
for (unsigned i = 0; i < p->get_num_args(); ++i) {
expr* arg = p->get_arg(i);
if (is_var(arg) && !m_var_is_sliceable[to_var(arg)->get_idx()] && bv.get(i)) {
bv.unset(i);
change = true;
TRACE("dl", tout << "variable is unslicable " << mk_pp(arg, m) << " for index " << i << " in " << p->get_decl()->get_name() << "\n";);
}
}
return change;
}
void mk_slice::add_free_vars(uint_set& result, expr* e) {
ptr_vector<sort> sorts;
get_free_vars(e, sorts);
for (unsigned i = 0; i < sorts.size(); ++i) {
if (sorts[i]) {
result.insert(i);
}
}
}
void mk_slice::display(std::ostream& out) {
obj_map<func_decl, bit_vector>::iterator it = m_sliceable.begin();
obj_map<func_decl, bit_vector>::iterator end = m_sliceable.end();
for (; it != end; ++it) {
out << it->m_key->get_name() << " ";
bit_vector const& bv = it->m_value;
for (unsigned i = 0; i < bv.size(); ++i) {
out << (bv.get(i)?"1":"0");
}
out << "\n";
}
}
void mk_slice::reset() {
m_input.reset();
m_output.reset();
m_var_is_sliceable.reset();
m_solved_vars.reset();
m_predicates.reset();
m_pinned.reset();
}
void mk_slice::declare_predicates(rule_set const& src, rule_set& dst) {
obj_map<func_decl, bit_vector>::iterator it = m_sliceable.begin(), end = m_sliceable.end();
ptr_vector<sort> domain;
bool has_output = false;
func_decl* f;
for (; it != end; ++it) {
domain.reset();
func_decl* p = it->m_key;
bit_vector const& bv = it->m_value;
for (unsigned i = 0; i < bv.size(); ++i) {
if (!bv.get(i)) {
domain.push_back(p->get_domain(i));
}
}
if (domain.size() < bv.size()) {
f = m_ctx.mk_fresh_head_predicate(p->get_name(), symbol("slice"), domain.size(), domain.c_ptr(), p);
m_pinned.push_back(f);
m_predicates.insert(p, f);
dst.inherit_predicate(src, p, f);
if (m_mc) {
m_mc->add_predicate(p, f);
}
}
else if (src.is_output_predicate(p)) {
dst.set_output_predicate(p);
has_output = true;
}
}
// disable slicing if the output predicates don't occur in rules.
if (!has_output) {
m_predicates.reset();
}
}
bool mk_slice::rule_updated(rule const& r) {
if (m_predicates.contains(r.get_decl())) return true;
for (unsigned i = 0; i < r.get_uninterpreted_tail_size(); ++i) {
if (m_predicates.contains(r.get_decl(i))) return true;
}
return false;
}
void mk_slice::update_predicate(app* p, app_ref& q) {
func_decl* qd;
if (m_predicates.find(p->get_decl(), qd)) {
bit_vector const& bv = get_predicate_slice(p->get_decl());
ptr_vector<expr> args;
for (unsigned i = 0; i < bv.size(); ++i) {
if (!bv.get(i)) {
args.push_back(p->get_arg(i));
}
}
q = m.mk_app(qd, args.size(), args.c_ptr());
}
else {
q = p;
}
}
void mk_slice::update_rule(rule& r, rule_set& dst) {
rule_ref new_rule(rm);
if (rule_updated(r)) {
init_vars(r);
app_ref_vector tail(m);
app_ref head(m);
ptr_vector<sort> sorts;
update_predicate(r.get_head(), head);
get_free_vars(head.get(), sorts);
for (unsigned i = 0; i < r.get_uninterpreted_tail_size(); ++i) {
app_ref t(m);
update_predicate(r.get_tail(i), t);
tail.push_back(t);
get_free_vars(t, sorts);
}
expr_ref_vector conjs = get_tail_conjs(r);
m_solved_vars.reset();
for (unsigned i = 0; i < conjs.size(); ++i) {
expr* e = conjs[i].get();
tail.push_back(to_app(e));
}
new_rule = rm.mk(head.get(), tail.size(), tail.c_ptr(), (const bool*) 0);
rm.fix_unbound_vars(new_rule, false);
TRACE("dl", r.display(m_ctx, tout << "replacing:\n"); new_rule->display(m_ctx, tout << "by:\n"););
if (m_ctx.generate_proof_trace()) {
rm.mk_rule_asserted_proof(*new_rule.get());
}
}
else {
new_rule = &r;
}
dst.add_rule(new_rule.get());
if (m_pc) {
m_pc->insert(&r, new_rule.get(), 0, 0);
}
}
void mk_slice::update_rules(rule_set const& src, rule_set& dst) {
for (unsigned i = 0; i < src.get_num_rules(); ++i) {
update_rule(*src.get_rule(i), dst);
}
}
rule_set * mk_slice::operator()(rule_set const & src) {
for (unsigned i = 0; i < src.get_num_rules(); ++i) {
if (src.get_rule(i)->has_quantifiers()) {
return 0;
}
}
ref<slice_proof_converter> spc;
ref<slice_model_converter> smc;
if (m_ctx.generate_proof_trace()) {
spc = alloc(slice_proof_converter, m_ctx);
}
if (m_ctx.get_model_converter()) {
smc = alloc(slice_model_converter, *this, m);
}
m_pc = spc.get();
m_mc = smc.get();
reset();
saturate(src);
rule_set* result = alloc(rule_set, m_ctx);
declare_predicates(src, *result);
if (m_predicates.empty()) {
// nothing could be sliced.
dealloc(result);
return 0;
}
TRACE("dl", display(tout););
update_rules(src, *result);
TRACE("dl", result->display(tout););
if (m_mc) {
obj_map<func_decl, bit_vector>::iterator it = m_sliceable.begin(), end = m_sliceable.end();
for (; it != end; ++it) {
m_mc->add_sliceable(it->m_key, it->m_value);
}
}
m_ctx.add_proof_converter(spc.get());
m_ctx.add_model_converter(smc.get());
return result;
}
};

115
src/muz/dl_mk_slice.h Normal file
View file

@ -0,0 +1,115 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_slice.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-10-4.
Revision History:
--*/
#ifndef _DL_MK_SLICE_H_
#define _DL_MK_SLICE_H_
#include"dl_context.h"
#include"dl_rule_set.h"
#include"uint_set.h"
#include"dl_rule_transformer.h"
namespace datalog {
/**
\brief Implements a slicing rule transformation.
*/
class mk_slice : public rule_transformer::plugin {
class slice_proof_converter;
class slice_model_converter;
context& m_ctx;
ast_manager& m;
rule_manager& rm;
svector<bool> m_input;
svector<bool> m_output;
expr_ref_vector m_solved_vars;
svector<bool> m_var_is_sliceable;
obj_map<func_decl, func_decl*> m_predicates;
obj_map<func_decl, bit_vector> m_sliceable;
ast_ref_vector m_pinned;
slice_proof_converter* m_pc;
slice_model_converter* m_mc;
void reset();
void init(rule_set const& source);
void saturate(rule_set const& source);
void display(std::ostream& out);
bool prune_rule(rule& r);
void init_vars(rule& r);
void init_vars(app* p, bool is_output, bool is_neg_tail);
bool finalize_vars(app* p);
unsigned num_vars() const { return m_input.size(); }
bit_vector& get_predicate_slice(func_decl* p);
bit_vector& get_predicate_slice(app* p) { return get_predicate_slice(p->get_decl()); }
bool is_eq(expr* e, unsigned& v, expr_ref& r);
void add_free_vars(uint_set& s, expr* e);
void add_var(unsigned idx);
bool is_output(expr* e);
bool is_output(unsigned idx);
void update_rules(rule_set const& src, rule_set& dst);
void update_rule(rule& r, rule_set& dst);
expr_ref_vector get_tail_conjs(rule const& r);
void declare_predicates(rule_set const& src, rule_set& dst);
bool rule_updated(rule const& r);
void update_predicate(app* p, app_ref& q);
void filter_unique_vars(rule& r);
void solve_vars(rule& r, uint_set& used_vars, uint_set& parameter_vars);
public:
/**
\brief Create slice rule transformer for \c goal predicate. When applying the transformer,
the \c goal must be present in the \c rule_set that is being transformed.
*/
mk_slice(context & ctx);
virtual ~mk_slice() { }
rule_set * operator()(rule_set const & source);
func_decl* get_predicate(func_decl* p) { func_decl* q = p; m_predicates.find(p, q); return q; }
obj_map<func_decl, func_decl*> const& get_predicates() { return m_predicates; }
};
};
#endif /* _DL_MK_SLICE_H_ */

View file

@ -0,0 +1,368 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_subsumption_checker.cpp
Abstract:
Rule transformer which checks for subsumption
(currently just for subsumption with total relations)
Author:
Krystof Hoder (t-khoder) 2011-10-01.
Revision History:
--*/
#include <sstream>
#include"ast_pp.h"
#include "rewriter.h"
#include "rewriter_def.h"
#include"dl_mk_subsumption_checker.h"
#include"dl_table_relation.h"
namespace datalog {
// -----------------------------------
//
// mk_subsumption_checker
//
// -----------------------------------
bool mk_subsumption_checker::is_total_rule(const rule * r) {
if(r->get_tail_size()!=0) { return false; }
unsigned pt_len = r->get_positive_tail_size();
if(pt_len!=r->get_uninterpreted_tail_size()) {
//we dont' expect rules with negative tails to be total
return false;
}
for(unsigned i=0; i<pt_len; i++) {
func_decl * tail_pred = r->get_tail(i)->get_decl();
if(!m_total_relations.contains(tail_pred)) {
//this rule has a non-total predicate in the tail
return false;
}
}
unsigned t_len = r->get_positive_tail_size();
for(unsigned i=pt_len; i<t_len; i++) {
SASSERT(!r->is_neg_tail(i)); //we assume interpreted tail not to be negated
if(!m.is_true(r->get_tail(i))) {
//this rule has an interpreted tail which is not constant true
return false;
}
}
var_idx_set head_vars;
app * head = r->get_head();
unsigned arity = head->get_num_args();
for(unsigned i=0; i<arity; i++) {
expr * arg = head->get_arg(i);
if(!is_var(arg)) { return false; }
unsigned idx = to_var(arg)->get_idx();
if(head_vars.contains(idx)) { return false; }
head_vars.insert(idx);
}
SASSERT(head_vars.num_elems()==arity);
return true;
}
void mk_subsumption_checker::on_discovered_total_relation(func_decl * pred, rule * r) {
//this should be rule marking a new relation as total
SASSERT(!m_total_relations.contains(pred));
SASSERT(!r || pred==r->get_decl());
SASSERT(!r || is_total_rule(r));
m_total_relations.insert(pred);
m_total_relation_defining_rules.insert(pred, r);
m_have_new_total_rule = true;
if(r) {
m_ref_holder.push_back(r);
}
}
void mk_subsumption_checker::scan_for_total_rules(const rule_set & rules) {
bool new_discovered;
//we cycle through the rules until we keep discovering new total relations
//(discovering a total relation migh reveal other total relations)
do {
new_discovered = false;
rule_set::iterator rend = rules.end();
for(rule_set::iterator rit = rules.begin(); rit!=rend; ++rit) {
rule * r = *rit;
func_decl * head_pred = r->get_decl();
if(is_total_rule(r) && !m_total_relations.contains(head_pred)) {
on_discovered_total_relation(head_pred, r);
new_discovered = true;
}
}
} while(new_discovered);
}
bool mk_subsumption_checker::transform_rule(rule * r,
rule_subsumption_index& subs_index, rule_ref & res)
{
unsigned u_len = r->get_uninterpreted_tail_size();
unsigned len = r->get_tail_size();
if(u_len==0) {
res = r;
return true;
}
app_ref head(r->get_head(), m);
app_ref_vector tail(m);
svector<bool> tail_neg;
for(unsigned i=0; i<u_len; i++) {
app * tail_atom = r->get_tail(i);
bool neg = r->is_neg_tail(i);
if(m_total_relations.contains(tail_atom->get_decl())
|| subs_index.is_subsumed(tail_atom)) {
if(neg) {
//rule contains negated total relation, this means that it is unsatisfiable
//and can be removed
return false;
}
else {
//we remove total relations from the tail
continue;
}
}
if(!neg && head.get()==tail_atom) {
//rule contains its head positively in the tail, therefore
//it will never add any new facts to the relation, so it
//can be removed
return false;
}
tail.push_back(tail_atom);
tail_neg.push_back(neg);
}
if(tail.size()==u_len) {
res = r;
return true;
}
//we just copy the interpreted part of the tail
for(unsigned i=u_len; i<len; i++) {
tail.push_back(r->get_tail(i));
tail_neg.push_back(r->is_neg_tail(i));
}
SASSERT(tail.size()==tail_neg.size());
res = m_context.get_rule_manager().mk(head, tail.size(), tail.c_ptr(), tail_neg.c_ptr());
res->set_accounting_parent_object(m_context, r);
m_context.get_rule_manager().fix_unbound_vars(res, true);
m_context.get_rule_manager().mk_rule_rewrite_proof(*r, *res.get());
return true;
}
bool rule_size_comparator(rule * r1, rule * r2) {
return r1->get_tail_size() < r2->get_tail_size();
}
bool mk_subsumption_checker::transform_rules(const rule_set & orig, rule_set & tgt) {
bool modified = false;
func_decl_set total_relations_with_included_rules;
rule_subsumption_index subs_index(m_context);
rule_ref_vector orig_rules(m_context.get_rule_manager());
orig_rules.append(orig.get_num_rules(), orig.begin());
rule * * rbegin = orig_rules.c_ptr();
rule * * rend = rbegin + orig_rules.size();
//before traversing we sort rules so that the shortest are in the beginning.
//this will help make subsumption checks more efficient
std::sort(rbegin, rend, rule_size_comparator);
for(rule_set::iterator rit = rbegin; rit!=rend; ++rit) {
rule * r = *rit;
func_decl * head_pred = r->get_decl();
if(m_total_relations.contains(head_pred)) {
if(!orig.is_output_predicate(head_pred) ||
total_relations_with_included_rules.contains(head_pred)) {
//We just skip definitions of total non-output relations as
//we'll eliminate them from the problem.
//We also skip rules of total output relations for which we have
//already output the rule which implies their totality.
modified = true;
continue;
}
rule * defining_rule;
VERIFY(m_total_relation_defining_rules.find(head_pred, defining_rule));
if (defining_rule) {
rule_ref totality_rule(m_context.get_rule_manager());
VERIFY(transform_rule(defining_rule, subs_index, totality_rule));
if(defining_rule!=totality_rule) {
modified = true;
}
tgt.add_rule(totality_rule);
SASSERT(totality_rule->get_decl()==head_pred);
}
else {
modified = true;
}
total_relations_with_included_rules.insert(head_pred);
continue;
}
rule_ref new_rule(m_context.get_rule_manager());
if(!transform_rule(r, subs_index, new_rule)) {
modified = true;
continue;
}
if(m_new_total_relation_discovery_during_transformation && is_total_rule(new_rule)) {
on_discovered_total_relation(head_pred, new_rule.get());
}
if(subs_index.is_subsumed(new_rule)) {
modified = true;
continue;
}
if(new_rule.get()!=r) {
modified = true;
}
tgt.add_rule(new_rule);
subs_index.add(new_rule);
}
tgt.inherit_predicates(orig);
TRACE("dl",
tout << "original set size: "<<orig.get_num_rules()<<"\n"
<< "reduced set size: "<<tgt.get_num_rules()<<"\n"; );
return modified;
}
void mk_subsumption_checker::scan_for_relations_total_due_to_facts(rule_set const& source) {
rel_context* rel = m_context.get_rel_context();
if (!rel) {
return;
}
relation_manager& rm = rel->get_rmanager();
decl_set const& candidate_preds = m_context.get_predicates();
decl_set::iterator end = candidate_preds.end();
for(decl_set::iterator it = candidate_preds.begin(); it!=end; ++it) {
func_decl * pred = *it;
if (m_total_relations.contains(pred)) { continue; } //already total
relation_base * rel = rm.try_get_relation(pred);
if (!rel || !rel->knows_exact_size()) { continue; }
unsigned arity = pred->get_arity();
if (arity > 30) { continue; }
//for now we only check booleans domains
for(unsigned i=0; i<arity; i++) {
if(!m.is_bool(pred->get_domain(i))) {
goto next_pred;
}
}
{
unsigned total_size = 1<<arity;
//by calling rel.knows_exact_size() we got assured that the estimate is exact
unsigned rel_sz = rel->get_size_estimate_rows();
obj_hashtable<app> * head_store;
if(m_ground_unconditional_rule_heads.find(pred, head_store)) {
//Some relations may receive facts by ground unconditioned rules.
//We scanned for those earlier, so now we check whether we cannot get a
//better estimate of relation size from these.
unsigned gnd_rule_cnt = head_store->size();
if(gnd_rule_cnt>rel_sz) {
rel_sz = gnd_rule_cnt;
}
}
SASSERT(total_size>=rel_sz);
if(total_size==rel_sz) {
on_discovered_total_relation(pred, 0);
}
}
next_pred:;
}
}
void mk_subsumption_checker::collect_ground_unconditional_rule_heads(const rule_set & rules)
{
rule_set::iterator rend = rules.end();
for(rule_set::iterator rit = rules.begin(); rit!=rend; ++rit) {
rule * r = *rit;
func_decl * pred = r->get_decl();
if(r->get_tail_size()!=0) { continue; }
app * head = r->get_head();
unsigned arity = pred->get_arity();
for(unsigned i=0; i<arity; i++) {
expr * arg = head->get_arg(i);
if(!is_app(arg)) {
goto next_rule;
}
}
if(!m_ground_unconditional_rule_heads.contains(pred)) {
m_ground_unconditional_rule_heads.insert(pred, alloc(obj_hashtable<app>));
}
m_ground_unconditional_rule_heads.find(pred)->insert(head);
next_rule:;
}
}
rule_set * mk_subsumption_checker::operator()(rule_set const & source) {
// TODO mc
m_have_new_total_rule = false;
collect_ground_unconditional_rule_heads(source);
scan_for_relations_total_due_to_facts(source);
scan_for_total_rules(source);
m_have_new_total_rule = false;
rule_set * res = alloc(rule_set, m_context);
bool modified = transform_rules(source, *res);
if (!m_have_new_total_rule && !modified) {
dealloc(res);
return 0;
}
//During the construction of the new set we may discover new total relations
//(by quantifier elimination on the uninterpreted tails).
SASSERT(m_new_total_relation_discovery_during_transformation || !m_have_new_total_rule);
while (m_have_new_total_rule) {
m_have_new_total_rule = false;
rule_set * old = res;
res = alloc(rule_set, m_context);
transform_rules(*old, *res);
dealloc(old);
}
return res;
}
};

View file

@ -0,0 +1,93 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
mk_subsumption_checker.h
Abstract:
Rule transformer which checks for subsumption
(currently just for subsumption with total relations)
Author:
Krystof Hoder (t-khoder) 2011-10-01.
Revision History:
--*/
#ifndef _DL_MK_SUBSUMPTION_CHECKER_H_
#define _DL_MK_SUBSUMPTION_CHECKER_H_
#include "dl_context.h"
#include "dl_rule_transformer.h"
#include "dl_rule_subsumption_index.h"
namespace datalog {
class mk_subsumption_checker : public rule_transformer::plugin {
ast_manager & m;
context & m_context;
rule_ref_vector m_ref_holder;
func_decl_set m_total_relations;
/** Map that for each relation contains the rule which implies its totality.
If the totality is due to the relation containing all facts, the rule stored
here is zero*/
obj_map<func_decl, rule *> m_total_relation_defining_rules;
/**
Contains heads of rules of shape
R(c1,c2,...cN).
grouped by their predicate.
This information helps to improve the results of the
scan_for_relations_total_due_to_facts() function.
*/
obj_map<func_decl, obj_hashtable<app> *> m_ground_unconditional_rule_heads;
bool m_have_new_total_rule;
bool m_new_total_relation_discovery_during_transformation;
bool is_total_rule(const rule * r);
/** Function to be called when a new total relation is discovered */
void on_discovered_total_relation(func_decl * pred, rule * r);
void scan_for_total_rules(rule_set const& rules);
void scan_for_relations_total_due_to_facts(rule_set const& rules);
void collect_ground_unconditional_rule_heads(const rule_set & rules);
/** Return false if rule is unsatisfiable */
bool transform_rule(rule * r, rule_subsumption_index& subs_index, rule_ref & res);
/** Return false if the rule set hasn't changed */
bool transform_rules(const rule_set & orig, rule_set & tgt);
public:
mk_subsumption_checker(context & ctx, unsigned priority=31000)
: plugin(priority),
m(ctx.get_manager()),
m_context(ctx),
m_ref_holder(ctx.get_rule_manager()),
m_new_total_relation_discovery_during_transformation(true) {}
~mk_subsumption_checker() {
reset_dealloc_values(m_ground_unconditional_rule_heads);
}
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_SUBSUMPTION_CHECKER_H_ */

View file

@ -0,0 +1,383 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_unbound_compressor.cpp
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-10-04.
Revision History:
--*/
#include<utility>
#include<sstream>
#include"dl_mk_unbound_compressor.h"
namespace datalog {
mk_unbound_compressor::mk_unbound_compressor(context & ctx) :
plugin(500),
m_context(ctx),
m(ctx.get_manager()),
rm(ctx.get_rule_manager()),
m_rules(rm),
m_pinned(m) {
}
void mk_unbound_compressor::reset() {
m_rules.reset();
m_todo.reset();
m_in_progress.reset();
m_map.reset();
m_pinned.reset();
}
bool mk_unbound_compressor::is_unbound_argument(rule * r, unsigned head_index) {
app * head = r->get_head();
expr * head_arg = head->get_arg(head_index);
if (!is_var(head_arg)) {
return false;
}
unsigned var_idx = to_var(head_arg)->get_idx();
return rm.collect_tail_vars(r).contains(var_idx);
}
void mk_unbound_compressor::add_task(func_decl * pred, unsigned arg_index) {
c_info ci = c_info(pred, arg_index);
if (m_map.contains(ci)) {
return; //this task was already added
}
unsigned parent_arity = pred->get_arity();
sort * const * parent_domain = pred->get_domain();
symbol const& parent_name = pred->get_name();
unsigned arity = parent_arity-1;
ptr_vector<sort> domain;
for (unsigned i = 0; i < parent_arity; i++) {
if (i != arg_index) {
domain.push_back(parent_domain[i]);
}
}
std::stringstream name_suffix;
name_suffix << "compr_arg_" << arg_index;
func_decl * cpred = m_context.mk_fresh_head_predicate(parent_name, symbol(name_suffix.str().c_str()),
arity, domain.c_ptr(), pred);
m_pinned.push_back(cpred);
m_todo.push_back(ci);
m_map.insert(ci, cpred);
}
void mk_unbound_compressor::detect_tasks(rule_set const& source, unsigned rule_index) {
rule * r = m_rules.get(rule_index);
var_idx_set& tail_vars = rm.collect_tail_vars(r);
app * head = r->get_head();
func_decl * head_pred = head->get_decl();
if (source.is_output_predicate(head_pred)) {
//we don't compress output predicates
return;
}
unsigned n = head_pred->get_arity();
rm.get_counter().reset();
rm.get_counter().count_vars(m, head, 1);
for (unsigned i=0; i<n; i++) {
expr * arg = head->get_arg(i);
if (!is_var(arg)) {
continue;
}
unsigned var_idx = to_var(arg)->get_idx();
if (!tail_vars.contains(var_idx)) {
//unbound
unsigned occurence_cnt = rm.get_counter().get(var_idx);
SASSERT(occurence_cnt>0);
if (occurence_cnt == 1) {
TRACE("dl", r->display(m_context, tout << "Compress: "););
add_task(head_pred, i);
return; //we compress out the unbound arguments one by one
}
}
}
}
void mk_unbound_compressor::try_compress(rule_set const& source, unsigned rule_index) {
start:
rule * r = m_rules.get(rule_index);
var_idx_set& tail_vars = rm.collect_tail_vars(r);
app * head = r->get_head();
func_decl * head_pred = head->get_decl();
unsigned head_arity = head_pred->get_arity();
rm.get_counter().reset();
rm.get_counter().count_vars(m, head);
unsigned arg_index;
for (arg_index = 0; arg_index < head_arity; arg_index++) {
expr * arg = head->get_arg(arg_index);
if (!is_var(arg)) {
continue;
}
unsigned var_idx = to_var(arg)->get_idx();
if (!tail_vars.contains(var_idx)) {
//unbound
unsigned occurence_cnt = rm.get_counter().get(var_idx);
SASSERT(occurence_cnt>0);
if ( occurence_cnt==1 && m_in_progress.contains(c_info(head_pred, arg_index)) ) {
//we have found what to compress
break;
}
}
}
if (arg_index == head_arity) {
//we didn't find anything to compress
return;
}
SASSERT(arg_index<head_arity);
c_info ci(head_pred, arg_index);
func_decl * cpred;
TRUSTME( m_map.find(ci, cpred) );
ptr_vector<expr> cargs;
for (unsigned i=0; i<head_arity; i++) {
if (i != arg_index) {
cargs.push_back(head->get_arg(i));
}
}
app_ref chead(m.mk_app(cpred, head_arity-1, cargs.c_ptr()), m);
if (r->get_tail_size()==0 && m_context.get_rule_manager().is_fact(chead)) {
m_non_empty_rels.insert(cpred);
m_context.add_fact(chead);
//remove the rule that became fact by placing the last rule on its place
m_head_occurrence_ctr.dec(m_rules.get(rule_index)->get_decl());
m_rules.set(rule_index, m_rules.get(m_rules.size()-1));
m_rules.shrink(m_rules.size()-1);
//since we moved the last rule to rule_index, we have to try to compress it as well
if (rule_index<m_rules.size()) {
goto start;
}
}
else {
rule_ref new_rule(m_context.get_rule_manager().mk(r, chead), m_context.get_rule_manager());
new_rule->set_accounting_parent_object(m_context, r);
m_head_occurrence_ctr.dec(m_rules.get(rule_index)->get_decl());
m_rules.set(rule_index, new_rule);
m_head_occurrence_ctr.inc(m_rules.get(rule_index)->get_decl());
detect_tasks(source, rule_index);
}
m_modified = true;
}
void mk_unbound_compressor::mk_decompression_rule(rule * r, unsigned tail_index, unsigned arg_index,
rule_ref& res)
{
app * orig_dtail = r->get_tail(tail_index); //dtail ~ decompressed tail
c_info ci(orig_dtail->get_decl(), arg_index);
func_decl * dtail_pred;
TRUSTME( m_map.find(ci, dtail_pred) );
ptr_vector<expr> dtail_args;
unsigned orig_dtail_arity = orig_dtail->get_num_args();
for (unsigned i=0;i<orig_dtail_arity;i++) {
if (i != arg_index) {
dtail_args.push_back(orig_dtail->get_arg(i));
}
}
SASSERT(dtail_args.size()==dtail_pred->get_arity());
app_ref dtail(m.mk_app(dtail_pred, dtail_args.size(), dtail_args.c_ptr()), m);
svector<bool> tails_negated;
app_ref_vector tails(m);
unsigned tail_len = r->get_tail_size();
for (unsigned i=0; i<tail_len; i++) {
tails_negated.push_back(r->is_neg_tail(i));
if (i==tail_index && !r->is_neg_tail(i)) {
tails.push_back(dtail);
}
else {
tails.push_back(r->get_tail(i));
}
}
// Accumulate negated filtered rule instead
// of replacing the original predicate.
if (r->is_neg_tail(tail_index)) {
tails_negated.push_back(true);
tails.push_back(dtail);
}
res = m_context.get_rule_manager().mk( r->get_head(), tails.size(), tails.c_ptr(), tails_negated.c_ptr());
res->set_accounting_parent_object(m_context, r);
m_context.get_rule_manager().fix_unbound_vars(res, true);
}
void mk_unbound_compressor::add_decompression_rule(rule_set const& source, rule * r, unsigned tail_index, unsigned arg_index) {
rule_ref new_rule(m_context.get_rule_manager());
mk_decompression_rule(r, tail_index, arg_index, new_rule);
unsigned new_rule_index = m_rules.size();
m_rules.push_back(new_rule);
m_context.get_rule_manager().mk_rule_rewrite_proof(*r, *new_rule.get());
m_head_occurrence_ctr.inc(new_rule->get_decl());
detect_tasks(source, new_rule_index);
m_modified = true;
//TODO: avoid rule duplicity
//If two predicates are compressed in a rule, applying decompression
//to the results can cause a rule being added multiple times:
//P:- R(x,y), S(x,y)
//is decompressed into rules
//P:- R1(x), S(x,y)
//P:- R(x,y), S1(x)
//and each of these rules is again decompressed giving the same rule
//P:- R1(x), S1(x)
//P:- R1(x), S1(x)
}
void mk_unbound_compressor::replace_by_decompression_rule(rule_set const& source, unsigned rule_index, unsigned tail_index, unsigned arg_index)
{
rule * r = m_rules.get(rule_index);
rule_ref new_rule(m_context.get_rule_manager());
mk_decompression_rule(r, tail_index, arg_index, new_rule);
m_rules.set(rule_index, new_rule);
//we don't update the m_head_occurrence_ctr because the head predicate doesn't change
detect_tasks(source, rule_index);
m_modified = true;
}
void mk_unbound_compressor::add_decompression_rules(rule_set const& source, unsigned rule_index) {
unsigned_vector compressed_tail_pred_arg_indexes;
//this value is updated inside the loop if replace_by_decompression_rule is called
rule_ref r(m_rules.get(rule_index), m_context.get_rule_manager());
unsigned utail_len = r->get_uninterpreted_tail_size();
unsigned tail_index=0;
while (tail_index<utail_len) {
app * t = r->get_tail(tail_index);
func_decl * t_pred = t->get_decl();
unsigned t_arity = t_pred->get_arity();
bool is_negated_predicate = r->is_neg_tail(tail_index);
compressed_tail_pred_arg_indexes.reset();
for (unsigned arg_index=0; arg_index<t_arity; arg_index++) {
c_info ci(t_pred, arg_index);
if (m_in_progress.contains(ci)) {
compressed_tail_pred_arg_indexes.push_back(arg_index);
}
}
bool orig_rule_replaced = false;
while (!compressed_tail_pred_arg_indexes.empty()) {
unsigned arg_index = compressed_tail_pred_arg_indexes.back();
compressed_tail_pred_arg_indexes.pop_back();
bool can_remove_orig_rule =
compressed_tail_pred_arg_indexes.empty() &&
!m_non_empty_rels.contains(t_pred) &&
m_head_occurrence_ctr.get(t_pred)==0;
if (can_remove_orig_rule || is_negated_predicate) {
replace_by_decompression_rule(source, rule_index, tail_index, arg_index);
orig_rule_replaced = true;
}
else {
add_decompression_rule(source, r, tail_index, arg_index);
}
}
if (orig_rule_replaced) {
//update r with the new rule
rule * new_rule = m_rules.get(rule_index);
SASSERT(new_rule->get_uninterpreted_tail_size() >= utail_len);
//here we check that the rule replacement didn't affect other uninterpreted literals
//in the tail (aside of variable renaming)
SASSERT(tail_index==0 ||
new_rule->get_tail(tail_index-1)->get_decl()==r->get_tail(tail_index-1)->get_decl());
r = new_rule;
//we have replaced the original rule, with one that has different
//content of the tail_index -th tail. we will therefore not do
//tail_index++, so that we examine the new tail literal as well
}
else {
tail_index++;
}
}
}
rule_set * mk_unbound_compressor::operator()(rule_set const & source) {
// TODO mc
m_modified = false;
rel_context* rel = m_context.get_rel_context();
if (rel) {
rel->get_rmanager().collect_non_empty_predicates(m_non_empty_rels);
}
unsigned init_rule_cnt = source.get_num_rules();
SASSERT(m_rules.empty());
for (unsigned i=0; i<init_rule_cnt; i++) {
rule * r = source.get_rule(i);
m_rules.push_back(r);
m_head_occurrence_ctr.inc(r->get_decl());
}
for (unsigned i=0; i<init_rule_cnt; i++) {
detect_tasks(source, i);
}
while (!m_todo.empty()) {
m_in_progress.reset();
while (!m_todo.empty()) {
m_in_progress.insert(m_todo.back());
m_todo.pop_back();
}
unsigned rule_index = 0;
while (rule_index<m_rules.size()) {
try_compress(source, rule_index); //m_rules.size() can change here
if (rule_index<m_rules.size()) {
add_decompression_rules(source, rule_index); //m_rules.size() can change here
}
rule_index++;
}
}
rule_set * result = static_cast<rule_set *>(0);
if (m_modified) {
result = alloc(rule_set, m_context);
unsigned fin_rule_cnt = m_rules.size();
for (unsigned i=0; i<fin_rule_cnt; i++) {
result->add_rule(m_rules.get(i));
}
result->inherit_predicates(source);
}
reset();
return result;
}
};

View file

@ -0,0 +1,92 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_unbound_compressor.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-10-4.
Revision History:
--*/
#ifndef _DL_MK_UNBOUND_COMPRESSOR_H_
#define _DL_MK_UNBOUND_COMPRESSOR_H_
#include<utility>
#include"map.h"
#include"obj_pair_hashtable.h"
#include"dl_context.h"
#include"dl_rule_set.h"
#include"dl_rule_transformer.h"
namespace datalog {
/**
\brief Functor for introducing auxiliary predicates to avoid unbound variables in
rule heads.
A rule
P(x,_) :- T(x).
is replaced by
P1(x) :- T(x).
and for each occurrence of P in a tail of a rule, a new rule is added with P1 in
its place.
*/
class mk_unbound_compressor : public rule_transformer::plugin {
/** predicate and index of compressed argument */
typedef std::pair<func_decl*,unsigned> c_info;
typedef pair_hash<ptr_hash<func_decl>,unsigned_hash> c_info_hash;
/** predicates that are results of compression */
typedef map<c_info, func_decl*, c_info_hash, default_eq<c_info> > c_map;
typedef hashtable<c_info, c_info_hash, default_eq<c_info> > in_progress_table;
typedef svector<c_info> todo_stack;
context & m_context;
ast_manager & m;
rule_manager & rm;
rule_ref_vector m_rules;
bool m_modified;
todo_stack m_todo;
in_progress_table m_in_progress;
c_map m_map;
/**
Relations that contain facts
*/
decl_set m_non_empty_rels;
ast_counter m_head_occurrence_ctr;
ast_ref_vector m_pinned;
bool is_unbound_argument(rule * r, unsigned head_index);
bool has_unbound_head_var(rule * r);
void detect_tasks(rule_set const& source, unsigned rule_index);
void add_task(func_decl * pred, unsigned arg_index);
void try_compress(rule_set const& source, unsigned rule_index);
void add_decompression_rules(rule_set const& source, unsigned rule_index);
void mk_decompression_rule(rule * r, unsigned tail_index, unsigned arg_index, rule_ref& res);
void add_decompression_rule(rule_set const& source, rule * r, unsigned tail_index, unsigned arg_index);
void replace_by_decompression_rule(rule_set const& source, unsigned rule_index, unsigned tail_index, unsigned arg_index);
void reset();
public:
mk_unbound_compressor(context & ctx);
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_UNBOUND_COMPRESSOR_H_ */

64
src/muz/dl_mk_unfold.cpp Normal file
View file

@ -0,0 +1,64 @@
/*++
Copyright (c) 2012 Microsoft Corporation
Module Name:
dl_mk_unfold.cpp
Abstract:
Unfold rules once, return the unfolded set of rules.
Author:
Nikolaj Bjorner (nbjorner) 2012-10-15
Revision History:
--*/
#include "dl_mk_unfold.h"
namespace datalog {
mk_unfold::mk_unfold(context& ctx):
rule_transformer::plugin(100, false),
m_ctx(ctx),
m(ctx.get_manager()),
rm(ctx.get_rule_manager()),
m_unify(ctx)
{}
void mk_unfold::expand_tail(rule& r, unsigned tail_idx, rule_set const& src, rule_set& dst) {
SASSERT(tail_idx <= r.get_uninterpreted_tail_size());
if (tail_idx == r.get_uninterpreted_tail_size()) {
dst.add_rule(&r);
}
else {
func_decl* p = r.get_decl(tail_idx);
rule_vector const& p_rules = src.get_predicate_rules(p);
rule_ref new_rule(rm);
for (unsigned i = 0; i < p_rules.size(); ++i) {
rule const& r2 = *p_rules[i];
if (m_unify.unify_rules(r, tail_idx, r2) &&
m_unify.apply(r, tail_idx, r2, new_rule)) {
expr_ref_vector s1 = m_unify.get_rule_subst(r, true);
expr_ref_vector s2 = m_unify.get_rule_subst(r2, false);
resolve_rule(r, r2, tail_idx, s1, s2, *new_rule.get());
expand_tail(*new_rule.get(), tail_idx+r2.get_uninterpreted_tail_size(), src, dst);
}
}
}
}
rule_set * mk_unfold::operator()(rule_set const & source) {
rule_set* rules = alloc(rule_set, m_ctx);
rule_set::iterator it = source.begin(), end = source.end();
for (; it != end; ++it) {
expand_tail(**it, 0, source, *rules);
}
rules->inherit_predicates(source);
return rules;
}
};

53
src/muz/dl_mk_unfold.h Normal file
View file

@ -0,0 +1,53 @@
/*++
Copyright (c) 2012 Microsoft Corporation
Module Name:
dl_mk_unfold.h
Abstract:
Unfold rules once, return the unfolded set of rules.
Author:
Nikolaj Bjorner (nbjorner) 2012-10-15
Revision History:
--*/
#ifndef _DL_MK_UNFOLD_H_
#define _DL_MK_UNFOLD_H_
#include"dl_context.h"
#include"dl_rule_set.h"
#include"uint_set.h"
#include"dl_rule_transformer.h"
#include"dl_mk_rule_inliner.h"
namespace datalog {
/**
\brief Implements an unfolding transformation.
*/
class mk_unfold : public rule_transformer::plugin {
context& m_ctx;
ast_manager& m;
rule_manager& rm;
rule_unifier m_unify;
void expand_tail(rule& r, unsigned tail_idx, rule_set const& src, rule_set& dst);
public:
/**
\brief Create unfold rule transformer.
*/
mk_unfold(context & ctx);
rule_set * operator()(rule_set const & source);
};
};
#endif /* _DL_MK_UNFOLD_H_ */

File diff suppressed because it is too large Load diff

View file

@ -0,0 +1,190 @@
/*++
Copyright (c) 2010 Microsoft Corporation
Module Name:
dl_product_relation.h
Abstract:
A Relation relation combinator.
Author:
Nikolaj Bjorner (nbjorner) 2010-4-11
Revision History:
--*/
#ifndef _DL_PRODUCT_RELATION_H_
#define _DL_PRODUCT_RELATION_H_
#include "dl_context.h"
namespace datalog {
class product_relation;
class product_relation_plugin : public relation_plugin {
friend class product_relation;
public:
typedef svector<family_id> rel_spec;
private:
class join_fn;
class transform_fn;
class mutator_fn;
class aligned_union_fn;
class unaligned_union_fn;
class single_non_transparent_src_union_fn;
class filter_equal_fn;
class filter_identical_fn;
class filter_interpreted_fn;
struct fid_hash {
typedef family_id data;
unsigned operator()(data x) const { return static_cast<unsigned>(x); }
};
rel_spec_store<rel_spec, svector_hash<fid_hash> > m_spec_store;
family_id get_relation_kind(const product_relation & r);
bool is_product_relation(relation_base * r) { return r->get_plugin().is_product_relation(); }
public:
static product_relation_plugin& get_plugin(relation_manager & rmgr);
product_relation_plugin(relation_manager& m);
virtual void initialize(family_id fid);
virtual bool can_handle_signature(const relation_signature & s);
virtual bool can_handle_signature(const relation_signature & s, family_id kind);
static symbol get_name() { return symbol("product_relation"); }
family_id get_relation_kind(const relation_signature & sig, const rel_spec & spec);
virtual relation_base * mk_empty(const relation_signature & s);
virtual relation_base * mk_empty(const relation_signature & s, family_id kind);
virtual relation_base * mk_full(func_decl* p, const relation_signature & s);
virtual relation_base * mk_full(func_decl* p, const relation_signature & s, family_id kind);
protected:
virtual relation_join_fn * mk_join_fn(const relation_base & t1, const relation_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
virtual relation_transformer_fn * mk_project_fn(const relation_base & t, unsigned col_cnt,
const unsigned * removed_cols);
virtual relation_transformer_fn * mk_rename_fn(const relation_base & t, unsigned permutation_cycle_len,
const unsigned * permutation_cycle);
virtual relation_union_fn * mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
virtual relation_union_fn * mk_widen_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
virtual relation_mutator_fn * mk_filter_identical_fn(const relation_base & t, unsigned col_cnt,
const unsigned * identical_cols);
virtual relation_mutator_fn * mk_filter_equal_fn(const relation_base & t, const relation_element & value,
unsigned col);
virtual relation_mutator_fn * mk_filter_interpreted_fn(const relation_base & t, app * condition);
static bool is_product_relation(relation_base const& r);
private:
static product_relation& get(relation_base& r);
static product_relation const & get(relation_base const& r);
static product_relation* get(relation_base* r);
static product_relation const* get(relation_base const* r);
relation_union_fn * mk_union_w_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta, bool is_widen);
bool are_aligned(const product_relation& r1, const product_relation& r2);
static void get_common_spec(const ptr_vector<const product_relation> & rels, rel_spec & res);
};
class product_relation : public relation_base {
friend class product_relation_plugin;
friend class product_relation_plugin::join_fn;
friend class product_relation_plugin::transform_fn;
friend class product_relation_plugin::mutator_fn;
friend class product_relation_plugin::aligned_union_fn;
friend class product_relation_plugin::unaligned_union_fn;
friend class product_relation_plugin::single_non_transparent_src_union_fn;
friend class product_relation_plugin::filter_equal_fn;
friend class product_relation_plugin::filter_identical_fn;
friend class product_relation_plugin::filter_interpreted_fn;
typedef product_relation_plugin::rel_spec rel_spec;
/**
If m_relations is empty, value of this determines whether the relation is empty or full.
*/
bool m_default_empty;
/**
There must not be two relations of the same kind
*/
ptr_vector<relation_base> m_relations;
/**
Array of kinds of inner relations.
If two product relations have equal signature and specification, their
m_relations arrays contain corresponding relations at the same indexes.
The value returned by get_kind() depends uniquely on the specification.
*/
rel_spec m_spec;
/**
\brief Ensure the kind assigned to this relation reflects the types of inner relations.
*/
void ensure_correct_kind();
/**
The current specification must be a subset of the new one.
*/
void convert_spec(const rel_spec & spec);
public:
product_relation(product_relation_plugin& p, relation_signature const& s);
product_relation(product_relation_plugin& p, relation_signature const& s, unsigned num_relations, relation_base** relations);
~product_relation();
virtual bool empty() const;
virtual void add_fact(const relation_fact & f);
virtual bool contains_fact(const relation_fact & f) const;
virtual product_relation * clone() const;
virtual product_relation * complement(func_decl* p) const;
virtual void display(std::ostream & out) const;
virtual void to_formula(expr_ref& fml) const;
product_relation_plugin& get_plugin() const;
unsigned size() const { return m_relations.size(); }
relation_base& operator[](unsigned i) const { return *m_relations[i]; }
/**
If all relations except one are sieve_relations with no inner columns,
return true and into \c idx assign index of that relation. Otherwise return
false.
*/
bool try_get_single_non_transparent(unsigned & idx) const;
virtual bool is_precise() const {
for (unsigned i = 0; i < m_relations.size(); ++i) {
if (!m_relations[i]->is_precise()) {
return false;
}
}
return true;
}
};
};
#endif

File diff suppressed because it is too large Load diff

View file

@ -0,0 +1,688 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_relation_manager.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-09-24.
Revision History:
--*/
#ifndef _DL_RELATION_MANAGER_H_
#define _DL_RELATION_MANAGER_H_
#include"map.h"
#include"vector.h"
#include"dl_base.h"
namespace datalog {
class context;
class dl_decl_util;
class table_relation;
class table_relation_plugin;
class finite_product_relation;
class finite_product_relation_plugin;
class sieve_relation;
class sieve_relation_plugin;
class rule_set;
class relation_manager {
class empty_signature_relation_join_fn;
class default_relation_join_project_fn;
class default_relation_select_equal_and_project_fn;
class default_relation_intersection_filter_fn;
class default_relation_filter_interpreted_and_project_fn;
class auxiliary_table_transformer_fn;
class auxiliary_table_filter_fn;
class default_table_join_fn;
class default_table_project_fn;
class null_signature_table_project_fn;
class default_table_join_project_fn;
class default_table_rename_fn;
class default_table_union_fn;
class default_table_filter_equal_fn;
class default_table_filter_identical_fn;
class default_table_filter_interpreted_fn;
class default_table_filter_interpreted_and_project_fn;
class default_table_negation_filter_fn;
class default_table_filter_not_equal_fn;
class default_table_select_equal_and_project_fn;
class default_table_map_fn;
class default_table_project_with_reduce_fn;
typedef obj_map<func_decl, family_id> decl2kind_map;
typedef u_map<relation_plugin *> kind2plugin_map;
typedef map<const table_plugin *, table_relation_plugin *, ptr_hash<const table_plugin>,
ptr_eq<const table_plugin> > tp2trp_map;
typedef map<const relation_plugin *, finite_product_relation_plugin *, ptr_hash<const relation_plugin>,
ptr_eq<const relation_plugin> > rp2fprp_map;
typedef map<func_decl *, relation_base *, ptr_hash<func_decl>, ptr_eq<func_decl> > relation_map;
typedef ptr_vector<table_plugin> table_plugin_vector;
typedef ptr_vector<relation_plugin> relation_plugin_vector;
context & m_context;
table_plugin_vector m_table_plugins;
relation_plugin_vector m_relation_plugins;
//table_relation_plugins corresponding to table_plugins
tp2trp_map m_table_relation_plugins;
rp2fprp_map m_finite_product_relation_plugins;
kind2plugin_map m_kind2plugin;
table_plugin * m_favourite_table_plugin;
relation_plugin * m_favourite_relation_plugin;
relation_map m_relations;
decl_set m_saturated_rels;
family_id m_next_table_fid;
family_id m_next_relation_fid;
/**
Map specifying what kind of relation should be used to represent particular predicate.
*/
decl2kind_map m_pred_kinds;
void register_relation_plugin_impl(relation_plugin * plugin);
relation_manager(const relation_manager &); //private and undefined copy constructor
relation_manager & operator=(const relation_manager &); //private and undefined operator=
public:
relation_manager(context & ctx) :
m_context(ctx),
m_favourite_table_plugin(0),
m_favourite_relation_plugin(0),
m_next_table_fid(0),
m_next_relation_fid(0) {}
virtual ~relation_manager();
void reset();
void reset_relations();
context & get_context() const { return m_context; }
dl_decl_util & get_decl_util() const;
family_id get_next_table_fid() { return m_next_table_fid++; }
family_id get_next_relation_fid(relation_plugin & claimer);
/**
Set what kind of relation is going to be used to represent the predicate \c pred.
This function can be called only before the relation object for \c pred is created
(i.e. before the \c get_relation function is called with \c pred as argument for the
first time).
*/
void set_predicate_kind(func_decl * pred, family_id kind);
/**
Return the relation kind that was requested to represent the predicate \c pred by
\c set_predicate_kind. If there was no such request, return \c null_family_id.
*/
family_id get_requested_predicate_kind(func_decl * pred);
relation_base & get_relation(func_decl * pred);
relation_base * try_get_relation(func_decl * pred) const;
/**
\brief Store the relation \c rel under the predicate \c pred. The \c relation_manager
takes over the relation object.
*/
void store_relation(func_decl * pred, relation_base * rel);
bool is_saturated(func_decl * pred) const { return m_saturated_rels.contains(pred); }
void mark_saturated(func_decl * pred) { m_saturated_rels.insert(pred); }
void reset_saturated_marks() {
if(!m_saturated_rels.empty()) {
m_saturated_rels.reset();
}
}
void collect_non_empty_predicates(decl_set & res) const;
void restrict_predicates(const decl_set & preds);
void register_plugin(table_plugin * plugin);
/**
table_relation_plugins should not be passed to this function since they are
created automatically when registering a table plugin.
*/
void register_plugin(relation_plugin * plugin) {
SASSERT(!plugin->from_table());
register_relation_plugin_impl(plugin);
}
table_plugin & get_appropriate_plugin(const table_signature & t);
relation_plugin & get_appropriate_plugin(const relation_signature & t);
table_plugin * try_get_appropriate_plugin(const table_signature & t);
relation_plugin * try_get_appropriate_plugin(const relation_signature & t);
table_plugin * get_table_plugin(symbol const& s);
relation_plugin * get_relation_plugin(symbol const& s);
relation_plugin & get_relation_plugin(family_id kind);
table_relation_plugin & get_table_relation_plugin(table_plugin & tp);
bool try_get_finite_product_relation_plugin(const relation_plugin & inner,
finite_product_relation_plugin * & res);
table_base * mk_empty_table(const table_signature & s);
relation_base * mk_implicit_relation(const relation_signature & s, app * expr);
relation_base * mk_empty_relation(const relation_signature & s, family_id kind);
relation_base * mk_empty_relation(const relation_signature & s, func_decl* pred);
relation_base * mk_full_relation(const relation_signature & s, func_decl* pred, family_id kind);
relation_base * mk_full_relation(const relation_signature & s, func_decl* pred);
relation_base * mk_table_relation(const relation_signature & s, table_base * table);
bool mk_empty_table_relation(const relation_signature & s, relation_base * & result);
bool is_non_explanation(relation_signature const& s) const;
/**
\brief Convert relation value to table one.
This function can be called only for the relation sorts that have a table counterpart.
*/
void relation_to_table(const relation_sort & sort, const relation_element & from, table_element & to);
void table_to_relation(const relation_sort & sort, const table_element & from, relation_element & to);
void table_to_relation(const relation_sort & sort, const table_element & from,
const relation_fact::el_proxy & to);
void table_to_relation(const relation_sort & sort, const table_element & from,
relation_element_ref & to);
bool relation_sort_to_table(const relation_sort & from, table_sort & to);
void from_predicate(func_decl * pred, unsigned arg_index, relation_sort & result);
void from_predicate(func_decl * pred, relation_signature & result);
/**
\brief Convert relation signature to table signature and return true if successful. If false
is returned, the value of \c to is undefined.
*/
bool relation_signature_to_table(const relation_signature & from, table_signature & to);
void relation_fact_to_table(const relation_signature & s, const relation_fact & from,
table_fact & to);
void table_fact_to_relation(const relation_signature & s, const table_fact & from,
relation_fact & to);
void set_cancel(bool f);
// -----------------------------------
//
// relation operations
//
// -----------------------------------
//TODO: If multiple operation implementations are available, we may want to do something to
//select the best one here.
/**
If \c allow_product_relation is true, we will create a join that builds a product relation,
if there is no other way to do the join. If \c allow_product_relation is false, we will return
zero in that case.
*/
relation_join_fn * mk_join_fn(const relation_base & t1, const relation_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2, bool allow_product_relation=true);
relation_join_fn * mk_join_fn(const relation_base & t1, const relation_base & t2,
const unsigned_vector & cols1, const unsigned_vector & cols2, bool allow_product_relation=true) {
SASSERT(cols1.size()==cols2.size());
return mk_join_fn(t1, t2, cols1.size(), cols1.c_ptr(), cols2.c_ptr(), allow_product_relation);
}
/**
\brief Return functor that transforms a table into one that lacks columns listed in
\c removed_cols array.
The \c removed_cols cotains columns of table \c t in strictly ascending order.
*/
relation_transformer_fn * mk_project_fn(const relation_base & t, unsigned col_cnt,
const unsigned * removed_cols);
relation_transformer_fn * mk_project_fn(const relation_base & t, const unsigned_vector & removed_cols) {
return mk_project_fn(t, removed_cols.size(), removed_cols.c_ptr());
}
/**
\brief Return an operation that is a composition of a join an a project operation.
*/
relation_join_fn * mk_join_project_fn(const relation_base & t1, const relation_base & t2,
unsigned joined_col_cnt, const unsigned * cols1, const unsigned * cols2,
unsigned removed_col_cnt, const unsigned * removed_cols, bool allow_product_relation_join=true);
relation_join_fn * mk_join_project_fn(const relation_base & t1, const relation_base & t2,
const unsigned_vector & cols1, const unsigned_vector & cols2,
const unsigned_vector & removed_cols, bool allow_product_relation_join=true) {
return mk_join_project_fn(t1, t2, cols1.size(), cols1.c_ptr(), cols2.c_ptr(), removed_cols.size(),
removed_cols.c_ptr(), allow_product_relation_join);
}
relation_transformer_fn * mk_rename_fn(const relation_base & t, unsigned permutation_cycle_len,
const unsigned * permutation_cycle);
relation_transformer_fn * mk_rename_fn(const relation_base & t, const unsigned_vector & permutation_cycle) {
return mk_rename_fn(t, permutation_cycle.size(), permutation_cycle.c_ptr());
}
/**
Like \c mk_rename_fn, only the permutation is not specified by cycle, but by a permutated array
of column number.
*/
relation_transformer_fn * mk_permutation_rename_fn(const relation_base & t,
const unsigned * permutation);
relation_transformer_fn * mk_permutation_rename_fn(const relation_base & t,
const unsigned_vector permutation) {
SASSERT(t.get_signature().size()==permutation.size());
return mk_permutation_rename_fn(t, permutation.c_ptr());
}
/**
The post-condition for an ideal union operation is be
Union(tgt, src, delta):
tgt_1==tgt_0 \union src
delta_1== delta_0 \union ( tgt_1 \setminus tgt_0 )
A required post-condition is
Union(tgt, src, delta):
tgt_1==tgt_0 \union src
tgt_1==tgt_0 => delta_1==delta_0
delta_0 \subset delta_1
delta_1 \subset (delta_0 \union tgt_1)
( tgt_1 \setminus tgt_0 ) \subset delta_1
So that a sufficient implementation is
Union(tgt, src, delta) {
oldTgt:=tgt.clone();
tgt:=tgt \union src
if(tgt!=oldTgt) {
delta:=delta \union src //also ?delta \union tgt? would work
}
}
If rules are compiled with all_or_nothing_deltas parameter set to true, a sufficient
post-condition is
Union(tgt, src, delta):
tgt_1==tgt_0 \union src
(tgt_1==tgt_0 || delta_0 is non-empty) <=> delta_1 is non-empty
*/
relation_union_fn * mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
relation_union_fn * mk_union_fn(const relation_base & tgt, const relation_base & src) {
return mk_union_fn(tgt, src, static_cast<relation_base *>(0));
}
/**
Similar to union, but this one should be used inside loops to allow for abstract
domain convergence.
*/
relation_union_fn * mk_widen_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
relation_mutator_fn * mk_filter_identical_fn(const relation_base & t, unsigned col_cnt,
const unsigned * identical_cols);
relation_mutator_fn * mk_filter_identical_fn(const relation_base & t, const unsigned_vector identical_cols) {
return mk_filter_identical_fn(t, identical_cols.size(), identical_cols.c_ptr());
}
relation_mutator_fn * mk_filter_equal_fn(const relation_base & t, const relation_element & value,
unsigned col);
relation_mutator_fn * mk_filter_interpreted_fn(const relation_base & t, app * condition);
relation_transformer_fn * mk_filter_interpreted_and_project_fn(const relation_base & t, app * condition,
unsigned removed_col_cnt, const unsigned * removed_cols);
/**
\brief Operations that returns all rows of \c t for which is column \c col equal to \c value
with the column \c col removed.
This operation can often be efficiently implemented and is useful for evaluating rules
of the form
F(x):-P("c",x).
*/
relation_transformer_fn * mk_select_equal_and_project_fn(const relation_base & t,
const relation_element & value, unsigned col);
relation_intersection_filter_fn * mk_filter_by_intersection_fn(const relation_base & tgt,
const relation_base & src, unsigned joined_col_cnt,
const unsigned * tgt_cols, const unsigned * src_cols);
relation_intersection_filter_fn * mk_filter_by_intersection_fn(const relation_base & tgt,
const relation_base & src, const unsigned_vector & tgt_cols, const unsigned_vector & src_cols) {
SASSERT(tgt_cols.size()==src_cols.size());
return mk_filter_by_intersection_fn(tgt, src, tgt_cols.size(), tgt_cols.c_ptr(), src_cols.c_ptr());
}
relation_intersection_filter_fn * mk_filter_by_intersection_fn(const relation_base & tgt,
const relation_base & src);
/**
The filter_by_negation postcondition:
filter_by_negation(tgt, neg, columns in tgt: c1,...,cN,
corresponding columns in neg: d1,...,dN):
tgt_1:={x: x\in tgt_0 && ! \exists y: ( y \in neg & pi_c1(x)= pi_d1(y) & ... & pi_cN(x)= pi_dN(y) ) }
*/
relation_intersection_filter_fn * mk_filter_by_negation_fn(const relation_base & t,
const relation_base & negated_obj, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * negated_cols);
relation_intersection_filter_fn * mk_filter_by_negation_fn(const relation_base & t,
const relation_base & negated_obj, const unsigned_vector & t_cols,
const unsigned_vector & negated_cols) {
SASSERT(t_cols.size()==negated_cols.size());
return mk_filter_by_negation_fn(t, negated_obj, t_cols.size(), t_cols.c_ptr(), negated_cols.c_ptr());
}
// -----------------------------------
//
// table operations
//
// -----------------------------------
table_join_fn * mk_join_fn(const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
table_join_fn * mk_join_fn(const table_base & t1, const table_base & t2,
const unsigned_vector & cols1, const unsigned_vector & cols2) {
SASSERT(cols1.size()==cols2.size());
return mk_join_fn(t1, t2, cols1.size(), cols1.c_ptr(), cols2.c_ptr());
}
/**
\brief Return functor that transforms a table into one that lacks columns listed in
\c removed_cols array.
The \c removed_cols cotains columns of table \c t in strictly ascending order.
If a project operation removes a non-functional column, all functional columns become
non-functional (so that none of the values in functional columns are lost)
*/
table_transformer_fn * mk_project_fn(const table_base & t, unsigned col_cnt,
const unsigned * removed_cols);
table_transformer_fn * mk_project_fn(const table_base & t, const unsigned_vector & removed_cols) {
return mk_project_fn(t, removed_cols.size(), removed_cols.c_ptr());
}
/**
\brief Return an operation that is a composition of a join an a project operation.
This operation is equivalent to the two operations performed separately, unless functional
columns are involved.
The ordinary project would make all of the functional columns into non-functional if any
non-functional column was removed. In function, however, we group columns into equivalence
classes (according to the equalities in \c cols1 and \c cols2) and make everything non-functional
only if some equivalence class of non-functional columns would have no non-functional columns
remain after the removal.
This behavior is implemented in the \c table_signature::from_join_project function.
*/
table_join_fn * mk_join_project_fn(const table_base & t1, const table_base & t2,
unsigned joined_col_cnt, const unsigned * cols1, const unsigned * cols2,
unsigned removed_col_cnt, const unsigned * removed_cols);
table_join_fn * mk_join_project_fn(const table_base & t1, const table_base & t2,
const unsigned_vector & cols1, const unsigned_vector & cols2,
const unsigned_vector & removed_cols) {
return mk_join_project_fn(t1, t2, cols1.size(), cols1.c_ptr(), cols2.c_ptr(), removed_cols.size(),
removed_cols.c_ptr());
}
table_transformer_fn * mk_rename_fn(const table_base & t, unsigned permutation_cycle_len,
const unsigned * permutation_cycle);
table_transformer_fn * mk_rename_fn(const table_base & t, const unsigned_vector & permutation_cycle) {
return mk_rename_fn(t, permutation_cycle.size(), permutation_cycle.c_ptr());
}
/**
Like \c mk_rename_fn, only the permutation is not specified by cycle, but by a permutated array
of column number.
*/
table_transformer_fn * mk_permutation_rename_fn(const table_base & t, const unsigned * permutation);
table_transformer_fn * mk_permutation_rename_fn(const table_base & t, const unsigned_vector permutation) {
SASSERT(t.get_signature().size()==permutation.size());
return mk_permutation_rename_fn(t, permutation.c_ptr());
}
/**
The post-condition for an ideal union operation is be
Union(tgt, src, delta):
tgt_1==tgt_0 \union src
delta_1== delta_0 \union ( tgt_1 \setminus tgt_0 )
A required post-condition is
Union(tgt, src, delta):
tgt_1==tgt_0 \union src
tgt_1==tgt_0 => delta_1==delta_0
delta_0 \subset delta_1
delta_1 \subset (delta_0 \union tgt_1)
( tgt_1 \setminus tgt_0 ) \subset delta_1
So that a sufficient implementation is
Union(tgt, src, delta) {
oldTgt:=tgt.clone();
tgt:=tgt \union src
if(tgt!=oldTgt) {
delta:=delta \union src //also ?delta \union tgt? would work
}
}
If rules are compiled with all_or_nothing_deltas parameter set to true, a sufficient
post-condition is
Union(tgt, src, delta):
tgt_1==tgt_0 \union src
(tgt_1==tgt_0 || delta_0 is non-empty) <=> delta_1 is non-empty
*/
table_union_fn * mk_union_fn(const table_base & tgt, const table_base & src,
const table_base * delta);
table_union_fn * mk_union_fn(const table_base & tgt, const table_base & src) {
return mk_union_fn(tgt, src, static_cast<table_base *>(0));
}
/**
Similar to union, but this one should be used inside loops to allow for abstract
domain convergence.
*/
table_union_fn * mk_widen_fn(const table_base & tgt, const table_base & src,
const table_base * delta);
table_mutator_fn * mk_filter_identical_fn(const table_base & t, unsigned col_cnt,
const unsigned * identical_cols);
table_mutator_fn * mk_filter_identical_fn(const table_base & t, const unsigned_vector identical_cols) {
return mk_filter_identical_fn(t, identical_cols.size(), identical_cols.c_ptr());
}
table_mutator_fn * mk_filter_equal_fn(const table_base & t, const table_element & value,
unsigned col);
table_mutator_fn * mk_filter_interpreted_fn(const table_base & t, app * condition);
table_transformer_fn * mk_filter_interpreted_and_project_fn(const table_base & t, app * condition,
unsigned removed_col_cnt, const unsigned * removed_cols);
/**
\brief Operations that returns all rows of \c t for which is column \c col equal to \c value
with the column \c col removed.
This operation can often be efficiently implemented and is useful for evaluating rules
of the form
F(x):-P("c",x).
*/
table_transformer_fn * mk_select_equal_and_project_fn(const table_base & t,
const table_element & value, unsigned col);
table_intersection_filter_fn * mk_filter_by_intersection_fn(const table_base & t,
const table_base & src, unsigned joined_col_cnt, const unsigned * t_cols, const unsigned * src_cols);
table_intersection_filter_fn * mk_filter_by_intersection_fn(const table_base & t,
const table_base & src, const unsigned_vector & t_cols, const unsigned_vector & src_cols) {
SASSERT(t_cols.size()==src_cols.size());
return mk_filter_by_intersection_fn(t, src, t_cols.size(), t_cols.c_ptr(), src_cols.c_ptr());
}
/**
The filter_by_negation postcondition:
filter_by_negation(tgt, neg, columns in tgt: c1,...,cN,
corresponding columns in neg: d1,...,dN):
tgt_1:={x: x\in tgt_0 && ! \exists y: ( y \in neg & pi_c1(x)= pi_d1(y) & ... & pi_cN(x)= pi_dN(y) ) }
*/
table_intersection_filter_fn * mk_filter_by_negation_fn(const table_base & t, const table_base & negated_obj,
unsigned joined_col_cnt, const unsigned * t_cols, const unsigned * negated_cols);
table_intersection_filter_fn * mk_filter_by_negation_fn(const table_base & t, const table_base & negated_obj,
const unsigned_vector & t_cols, const unsigned_vector & negated_cols) {
SASSERT(t_cols.size()==negated_cols.size());
return mk_filter_by_negation_fn(t, negated_obj, t_cols.size(), t_cols.c_ptr(), negated_cols.c_ptr());
}
/**
\c t must contain at least one functional column.
Created object takes ownership of the \c mapper object.
*/
virtual table_mutator_fn * mk_map_fn(const table_base & t, table_row_mutator_fn * mapper);
/**
\c t must contain at least one functional column.
Created object takes ownership of the \c mapper object.
*/
virtual table_transformer_fn * mk_project_with_reduce_fn(const table_base & t, unsigned col_cnt,
const unsigned * removed_cols, table_row_pair_reduce_fn * reducer);
// -----------------------------------
//
// output functions
//
// -----------------------------------
std::string to_nice_string(const relation_element & el) const;
/**
This one may give a nicer representation of \c el than the
\c to_nice_string(const relation_element & el) function, by unsing the information about the sort
of the element.
*/
std::string to_nice_string(const relation_sort & s, const relation_element & el) const;
std::string to_nice_string(const relation_sort & s) const;
std::string to_nice_string(const relation_signature & s) const;
void display(std::ostream & out) const;
void display_relation_sizes(std::ostream & out) const;
void display_output_tables(rule_set const& rules, std::ostream & out) const;
private:
relation_intersection_filter_fn * try_mk_default_filter_by_intersection_fn(const relation_base & t,
const relation_base & src, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * src_cols);
};
/**
This is a helper class for relation_plugins whose relations can be of various kinds.
*/
template<class Spec, class Hash, class Eq=vector_eq_proc<Spec> >
class rel_spec_store {
typedef relation_signature::hash r_hash;
typedef relation_signature::eq r_eq;
typedef map<Spec, unsigned, Hash, Eq > family_id_idx_store;
typedef map<relation_signature, family_id_idx_store *, r_hash, r_eq> sig2store;
typedef u_map<Spec> family_id2spec;
typedef map<relation_signature, family_id2spec *, r_hash, r_eq> sig2spec_store;
relation_plugin & m_parent;
svector<family_id> m_allocated_kinds;
sig2store m_kind_assignment;
sig2spec_store m_kind_specs;
relation_manager & get_manager() { return m_parent.get_manager(); }
void add_new_kind() {
add_available_kind(get_manager().get_next_relation_fid(m_parent));
}
public:
rel_spec_store(relation_plugin & parent) : m_parent(parent) {}
~rel_spec_store() {
reset_dealloc_values(m_kind_assignment);
reset_dealloc_values(m_kind_specs);
}
void add_available_kind(family_id k) {
m_allocated_kinds.push_back(k);
}
bool contains_signature(relation_signature const& sig) const {
return m_kind_assignment.contains(sig);
}
family_id get_relation_kind(const relation_signature & sig, const Spec & spec) {
typename sig2store::entry * e = m_kind_assignment.find_core(sig);
if(!e) {
e = m_kind_assignment.insert_if_not_there2(sig, alloc(family_id_idx_store));
m_kind_specs.insert(sig, alloc(family_id2spec));
}
family_id_idx_store & ids = *e->get_data().m_value;
unsigned res_idx;
if(!ids.find(spec, res_idx)) {
res_idx = ids.size();
if(res_idx==m_allocated_kinds.size()) {
add_new_kind();
}
SASSERT(res_idx<m_allocated_kinds.size());
ids.insert(spec, res_idx);
family_id2spec * idspecs = m_kind_specs.find(sig);
idspecs->insert(m_allocated_kinds[res_idx], spec);
}
return m_allocated_kinds[res_idx];
}
void get_relation_spec(const relation_signature & sig, family_id kind, Spec & spec) {
family_id2spec * idspecs = m_kind_specs.find(sig);
spec = idspecs->find(kind);
}
};
};
#endif /* _DL_RELATION_MANAGER_H_ */

1141
src/muz/dl_rule.cpp Normal file

File diff suppressed because it is too large Load diff

331
src/muz/dl_rule.h Normal file
View file

@ -0,0 +1,331 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_rule.h
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2010-05-17.
Revision History:
--*/
#ifndef _DL_RULE_H_
#define _DL_RULE_H_
#include"ast.h"
#include"dl_costs.h"
#include"dl_util.h"
#include"used_vars.h"
#include"proof_converter.h"
#include"model_converter.h"
#include"ast_counter.h"
#include"rewriter.h"
#include"hnf.h"
#include"qe_lite.h"
namespace datalog {
class rule;
class rule_manager;
class rule_set;
class table;
class context;
typedef obj_ref<rule, rule_manager> rule_ref;
typedef ref_vector<rule, rule_manager> rule_ref_vector;
typedef ptr_vector<rule> rule_vector;
/**
\brief Manager for the \c rule class
\remark \c rule_manager objects are interchangable as long as they
contain the same \c ast_manager object.
*/
class rule_manager
{
class remove_label_cfg : public default_rewriter_cfg {
family_id m_label_fid;
public:
remove_label_cfg(ast_manager& m): m_label_fid(m.get_label_family_id()) {}
virtual ~remove_label_cfg();
br_status reduce_app(func_decl * f, unsigned num, expr * const * args, expr_ref & result,
proof_ref & result_pr);
};
ast_manager& m;
context& m_ctx;
rule_counter m_counter;
used_vars m_used;
ptr_vector<sort> m_vars;
var_idx_set m_var_idx;
ptr_vector<expr> m_todo;
ast_mark m_mark;
app_ref_vector m_body;
app_ref m_head;
expr_ref_vector m_args;
svector<bool> m_neg;
hnf m_hnf;
qe_lite m_qe;
remove_label_cfg m_cfg;
rewriter_tpl<remove_label_cfg> m_rwr;
// only the context can create a rule_manager
friend class context;
explicit rule_manager(context& ctx);
/**
\brief Move functions from predicate tails into the interpreted tail by introducing new variables.
*/
void hoist_compound_predicates(unsigned num_bound, app_ref& head, app_ref_vector& body);
void hoist_compound(unsigned& num_bound, app_ref& fml, app_ref_vector& body);
void flatten_body(app_ref_vector& body);
void remove_labels(expr_ref& fml, proof_ref& pr);
app_ref ensure_app(expr* e);
void check_app(expr* e);
bool contains_predicate(expr* fml) const;
void bind_variables(expr* fml, bool is_forall, expr_ref& result);
void mk_negations(app_ref_vector& body, svector<bool>& is_negated);
void mk_rule_core(expr* fml, proof* p, rule_set& rules, symbol const& name);
void mk_horn_rule(expr* fml, proof* p, rule_set& rules, symbol const& name);
static expr_ref mk_implies(app_ref_vector const& body, expr* head);
unsigned extract_horn(expr* fml, app_ref_vector& body, app_ref& head);
/**
\brief Perform cheap quantifier elimination to reduce the number of variables in the interpreted tail.
*/
void reduce_unbound_vars(rule_ref& r);
void reset_collect_vars();
var_idx_set& finalize_collect_vars();
public:
ast_manager& get_manager() const { return m; }
void inc_ref(rule * r);
void dec_ref(rule * r);
used_vars& reset_used() { m_used.reset(); return m_used; }
var_idx_set& collect_vars(expr * pred);
var_idx_set& collect_vars(expr * e1, expr* e2);
var_idx_set& collect_rule_vars(rule * r);
var_idx_set& collect_rule_vars_ex(rule * r, app* t);
var_idx_set& collect_tail_vars(rule * r);
void accumulate_vars(expr* pred);
ptr_vector<sort>& get_var_sorts() { return m_vars; }
var_idx_set& get_var_idx() { return m_var_idx; }
/**
\brief Create a Datalog rule from a Horn formula.
The formula is of the form (forall (...) (forall (...) (=> (and ...) head)))
*/
void mk_rule(expr* fml, proof* p, rule_set& rules, symbol const& name = symbol::null);
/**
\brief Create a Datalog query from an expression.
The formula is of the form (exists (...) (exists (...) (and ...))
*/
func_decl* mk_query(expr* query, rule_set& rules);
/**
\brief Create a Datalog rule head :- tail[0], ..., tail[n-1].
Return 0 if it is not a valid rule.
\remark A tail may contain negation. tail[i] is assumed to be negated if is_neg != 0 && is_neg[i] == true
*/
rule * mk(app * head, unsigned n, app * const * tail, bool const * is_neg = 0,
symbol const& name = symbol::null, bool normalize = true);
/**
\brief Create a rule with the same tail as \c source and with a specified head.
*/
rule * mk(rule const * source, app * new_head, symbol const& name = symbol::null);
/**
\brief Create a copy of the given rule.
*/
rule * mk(rule const * source, symbol const& name = symbol::null);
/** make sure there are not non-quantified variables that occur only in interpreted predicates */
void fix_unbound_vars(rule_ref& r, bool try_quantifier_elimination);
/**
\brief add proof that new rule is obtained by rewriting old rule.
*/
void mk_rule_rewrite_proof(rule& old_rule, rule& new_rule);
/**
\brief tag rule as asserted.
*/
void mk_rule_asserted_proof(rule& r);
/**
\brief apply substitution to variables of rule.
*/
void substitute(rule_ref& r, unsigned sz, expr*const* es);
/**
\brief Check that head :- tail[0], ..., tail[n-1]
is a valid Datalog rule.
*/
void check_valid_rule(app * head, unsigned n, app * const * tail) const;
/**
\brief Check that \c head may occur as a Datalog rule head.
*/
void check_valid_head(expr * head) const;
/**
\brief Return true if \c head may occur as a fact.
*/
bool is_fact(app * head) const;
static bool is_forall(ast_manager& m, expr* e, quantifier*& q);
rule_counter& get_counter() { return m_counter; }
};
class rule : public accounted_object {
friend class rule_manager;
app * m_head;
proof* m_proof;
unsigned m_tail_size:20;
// unsigned m_reserve:12;
unsigned m_ref_cnt;
unsigned m_positive_cnt;
unsigned m_uninterp_cnt;
symbol m_name;
/**
The following field is an array of tagged pointers.
- Tag 0: the atom is not negated
- Tag 1: the atom is negated.
The order of tail formulas is the following:
uninterpreted positive,
uninterpreted negative,
interpreted.
The negated flag is never set for interpreted tails.
*/
app * m_tail[0];
static unsigned get_obj_size(unsigned n) { return sizeof(rule) + n * sizeof(app *); }
rule() : m_ref_cnt(0) {}
~rule() {}
void deallocate(ast_manager & m);
void get_used_vars(used_vars& uv) const;
public:
proof * get_proof() const { return m_proof; }
void set_proof(ast_manager& m, proof* p);
app * get_head() const { return m_head; }
func_decl* get_decl() const { return get_head()->get_decl(); }
unsigned get_tail_size() const { return m_tail_size; }
/**
\brief Return number of positive uninterpreted predicates in the tail.
These predicates are the first in the tail.
*/
unsigned get_positive_tail_size() const { return m_positive_cnt; }
unsigned get_uninterpreted_tail_size() const { return m_uninterp_cnt; }
/**
\brief Return i-th tail atom. The first \c get_uninterpreted_tail_size()
atoms are uninterpreted and the first \c get_positive_tail_size() are
uninterpreted and non-negated.
*/
app * get_tail(unsigned i) const { SASSERT(i < m_tail_size); return UNTAG(app *, m_tail[i]); }
func_decl* get_decl(unsigned i) const { SASSERT(i < get_uninterpreted_tail_size()); return get_tail(i)->get_decl(); }
bool is_neg_tail(unsigned i) const { SASSERT(i < m_tail_size); return GET_TAG(m_tail[i]) == 1; }
/**
Check whether predicate p is in the interpreted tail.
If only_positive is true, only the positive predicate tail atoms are checked.
*/
bool is_in_tail(const func_decl * p, bool only_positive=false) const;
bool has_uninterpreted_non_predicates(func_decl*& f) const;
void has_quantifiers(bool& existential, bool& universal) const;
bool has_quantifiers() const;
bool has_negation() const;
/**
\brief Store in d the (direct) dependencies of the given rule.
*/
void norm_vars(rule_manager & rm);
void get_vars(ptr_vector<sort>& sorts) const;
void to_formula(expr_ref& result) const;
void display(context & ctx, std::ostream & out) const;
std::ostream& display_smt2(ast_manager& m, std::ostream & out) const;
symbol const& name() const { return m_name; }
unsigned hash() const;
};
struct rule_eq_proc {
bool operator()(const rule * r1, const rule * r2) const;
};
struct rule_hash_proc {
unsigned operator()(const rule * r) const;
};
};
#endif /* _DL_RULE_H_ */

773
src/muz/dl_rule_set.cpp Normal file
View file

@ -0,0 +1,773 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_rule_set.cpp
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2010-05-17.
Revision History:
--*/
#include<algorithm>
#include<functional>
#include"dl_context.h"
#include"dl_rule_set.h"
#include"ast_pp.h"
namespace datalog {
rule_dependencies::rule_dependencies(context& ctx): m_context(ctx) {
}
rule_dependencies::rule_dependencies(const rule_dependencies & o, bool reversed):
m_context(o.m_context) {
if (reversed) {
iterator oit = o.begin();
iterator oend = o.end();
for (; oit!=oend; ++oit) {
func_decl * pred = oit->m_key;
item_set & orig_items = *oit->get_value();
ensure_key(pred);
item_set::iterator dit = orig_items.begin();
item_set::iterator dend = orig_items.end();
for (; dit!=dend; ++dit) {
func_decl * master_pred = *dit;
insert(master_pred, pred);
}
}
}
else {
iterator oit = o.begin();
iterator oend = o.end();
for (; oit!=oend; ++oit) {
func_decl * pred = oit->m_key;
item_set & orig_items = *oit->get_value();
m_data.insert(pred, alloc(item_set, orig_items));
}
}
}
rule_dependencies::~rule_dependencies() {
reset();
}
void rule_dependencies::reset() {
reset_dealloc_values(m_data);
}
void rule_dependencies::remove_m_data_entry(func_decl * key) {
item_set * itm_set = m_data.find(key);
dealloc(itm_set);
m_data.remove(key);
}
rule_dependencies::item_set & rule_dependencies::ensure_key(func_decl * pred) {
deps_type::obj_map_entry * e = m_data.insert_if_not_there2(pred, 0);
if (!e->get_data().m_value) {
e->get_data().m_value = alloc(item_set);
}
return *e->get_data().m_value;
}
void rule_dependencies::insert(func_decl * depending, func_decl * master) {
SASSERT(m_data.contains(master)); //see m_data documentation
item_set & s = ensure_key(depending);
s.insert(master);
}
void rule_dependencies::populate(const rule_set & rules) {
SASSERT(m_data.empty());
rule_set::decl2rules::iterator it = rules.m_head2rules.begin();
rule_set::decl2rules::iterator end = rules.m_head2rules.end();
for (; it != end; ++it) {
ptr_vector<rule> * rules = it->m_value;
ptr_vector<rule>::iterator it2 = rules->begin();
ptr_vector<rule>::iterator end2 = rules->end();
for (; it2 != end2; ++it2) {
populate(*it2);
}
}
}
void rule_dependencies::populate(unsigned n, rule * const * rules) {
SASSERT(m_data.empty());
for (unsigned i=0; i<n; i++) {
populate(rules[i]);
}
}
void rule_dependencies::populate(rule const* r) {
TRACE("dl_verbose", tout << r->get_decl()->get_name() << "\n";);
m_visited.reset();
func_decl * d = r->get_decl();
func_decl_set & s = ensure_key(d);
for (unsigned i = 0; i < r->get_tail_size(); ++i) {
m_todo.push_back(r->get_tail(i));
}
while (!m_todo.empty()) {
expr* e = m_todo.back();
m_todo.pop_back();
if (m_visited.is_marked(e)) {
continue;
}
m_visited.mark(e, true);
if (is_app(e)) {
app* a = to_app(e);
d = a->get_decl();
if (m_context.is_predicate(d)) {
// insert d and ensure the invariant
// that every predicate is present as
// a key in m_data
s.insert(d);
ensure_key(d);
}
m_todo.append(a->get_num_args(), a->get_args());
}
else if (is_quantifier(e)) {
m_todo.push_back(to_quantifier(e)->get_expr());
}
}
}
const rule_dependencies::item_set & rule_dependencies::get_deps(func_decl * f) const {
deps_type::obj_map_entry * e = m_data.find_core(f);
if (!e) {
return m_empty_item_set;
}
SASSERT(e->get_data().get_value());
return *e->get_data().get_value();
}
void rule_dependencies::restrict(const item_set & allowed) {
ptr_vector<func_decl> to_remove;
iterator pit = begin();
iterator pend = end();
for (; pit!=pend; ++pit) {
func_decl * pred = pit->m_key;
if (!allowed.contains(pred)) {
to_remove.insert(pred);
continue;
}
item_set& itms = *pit->get_value();
set_intersection(itms, allowed);
}
ptr_vector<func_decl>::iterator rit = to_remove.begin();
ptr_vector<func_decl>::iterator rend = to_remove.end();
for (; rit != rend; ++rit) {
remove_m_data_entry(*rit);
}
}
void rule_dependencies::remove(func_decl * itm) {
remove_m_data_entry(itm);
iterator pit = begin();
iterator pend = end();
for (; pit != pend; ++pit) {
item_set & itms = *pit->get_value();
itms.remove(itm);
}
}
void rule_dependencies::remove(const item_set & to_remove) {
item_set::iterator rit = to_remove.begin();
item_set::iterator rend = to_remove.end();
for (; rit!=rend; ++rit) {
remove_m_data_entry(*rit);
}
iterator pit = begin();
iterator pend = end();
for (; pit!=pend; ++pit) {
item_set * itms = pit->get_value();
set_difference(*itms, to_remove);
}
}
unsigned rule_dependencies::out_degree(func_decl * f) const {
unsigned res = 0;
iterator pit = begin();
iterator pend = end();
for (; pit!=pend; ++pit) {
item_set & itms = *pit->get_value();
if (itms.contains(f)) {
res++;
}
}
return res;
}
bool rule_dependencies::sort_deps(ptr_vector<func_decl> & res) {
typedef obj_map<func_decl, unsigned> deg_map;
unsigned init_len = res.size();
deg_map degs;
unsigned curr_index = init_len;
rule_dependencies reversed(*this, true);
iterator pit = begin();
iterator pend = end();
for (; pit!=pend; ++pit) {
func_decl * pred = pit->m_key;
unsigned deg = in_degree(pred);
if (deg==0) {
res.push_back(pred);
}
else {
degs.insert(pred, deg);
}
}
while (curr_index<res.size()) { //res.size() can change in the loop iteration
func_decl * curr = res[curr_index];
const item_set & children = reversed.get_deps(curr);
item_set::iterator cit = children.begin();
item_set::iterator cend = children.end();
for (; cit!=cend; ++cit) {
func_decl * child = *cit;
deg_map::obj_map_entry * e = degs.find_core(child);
SASSERT(e);
unsigned & child_deg = e->get_data().m_value;
SASSERT(child_deg>0);
child_deg--;
if (child_deg==0) {
res.push_back(child);
}
}
curr_index++;
}
if (res.size() < init_len + m_data.size()) {
res.shrink(init_len);
return false;
}
SASSERT(res.size()==init_len+m_data.size());
return true;
}
void rule_dependencies::display(std::ostream & out ) const
{
iterator pit = begin();
iterator pend = end();
for (; pit != pend; ++pit) {
func_decl * pred = pit->m_key;
const item_set & deps = *pit->m_value;
item_set::iterator dit=deps.begin();
item_set::iterator dend=deps.end();
if (dit == dend) {
out<<pred->get_name()<<" - <none>\n";
}
for (; dit != dend; ++dit) {
func_decl * dep = *dit;
out << pred->get_name() << " -> " << dep->get_name() << "\n";
}
}
}
// -----------------------------------
//
// rule_set
//
// -----------------------------------
rule_set::rule_set(context & ctx)
: m_context(ctx),
m_rule_manager(ctx.get_rule_manager()),
m_rules(m_rule_manager),
m_deps(ctx),
m_stratifier(0),
m_refs(ctx.get_manager()) {
}
rule_set::rule_set(const rule_set & other)
: m_context(other.m_context),
m_rule_manager(other.m_rule_manager),
m_rules(m_rule_manager),
m_deps(other.m_context),
m_stratifier(0),
m_refs(m_context.get_manager()) {
add_rules(other);
if (other.m_stratifier) {
VERIFY(close());
}
}
rule_set::~rule_set() {
reset();
}
void rule_set::reset() {
m_rules.reset();
reset_dealloc_values(m_head2rules);
m_deps.reset();
m_stratifier = 0;
m_output_preds.reset();
m_orig2pred.reset();
m_pred2orig.reset();
m_refs.reset();
}
ast_manager & rule_set::get_manager() const {
return m_context.get_manager();
}
func_decl* rule_set::get_orig(func_decl* pred) const {
func_decl* orig = pred;
m_pred2orig.find(pred, orig);
return orig;
}
func_decl* rule_set::get_pred(func_decl* orig) const {
func_decl* pred = orig;
m_orig2pred.find(orig, pred);
return pred;
}
void rule_set::inherit_predicates(rule_set const& other) {
m_refs.append(other.m_refs);
set_union(m_output_preds, other.m_output_preds);
{
obj_map<func_decl, func_decl*>::iterator it = other.m_orig2pred.begin();
obj_map<func_decl, func_decl*>::iterator end = other.m_orig2pred.end();
for (; it != end; ++it) {
m_orig2pred.insert(it->m_key, it->m_value);
}
}
{
obj_map<func_decl, func_decl*>::iterator it = other.m_pred2orig.begin();
obj_map<func_decl, func_decl*>::iterator end = other.m_pred2orig.end();
for (; it != end; ++it) {
m_pred2orig.insert(it->m_key, it->m_value);
}
}
}
void rule_set::inherit_predicate(rule_set const& other, func_decl* orig, func_decl* pred) {
if (other.is_output_predicate(orig)) {
set_output_predicate(pred);
}
orig = other.get_orig(orig);
m_refs.push_back(pred);
m_refs.push_back(orig);
m_orig2pred.insert(orig, pred);
m_pred2orig.insert(pred, orig);
}
void rule_set::add_rule(rule * r) {
TRACE("dl_verbose", r->display(m_context, tout << "add:"););
SASSERT(!is_closed());
m_rules.push_back(r);
app * head = r->get_head();
SASSERT(head != 0);
func_decl * d = head->get_decl();
decl2rules::obj_map_entry* e = m_head2rules.insert_if_not_there2(d, 0);
if (!e->get_data().m_value) e->get_data().m_value = alloc(ptr_vector<rule>);
e->get_data().m_value->push_back(r);
}
void rule_set::del_rule(rule * r) {
TRACE("dl", r->display(m_context, tout << "del:"););
func_decl* d = r->get_decl();
rule_vector* rules = m_head2rules.find(d);
#define DEL_VECTOR(_v) \
for (unsigned i = (_v).size(); i > 0; ) { \
--i; \
if ((_v)[i] == r) { \
(_v)[i] = (_v).back(); \
(_v).pop_back(); \
break; \
} \
} \
DEL_VECTOR(*rules);
DEL_VECTOR(m_rules);
}
void rule_set::ensure_closed() {
if (!is_closed()) {
VERIFY(close());
}
}
bool rule_set::close() {
SASSERT(!is_closed()); //the rule_set is not already closed
m_deps.populate(*this);
m_stratifier = alloc(rule_stratifier, m_deps);
if (!stratified_negation()) {
m_stratifier = 0;
m_deps.reset();
return false;
}
return true;
}
void rule_set::reopen() {
if (is_closed()) {
m_stratifier = 0;
m_deps.reset();
}
}
/**
\brief Return true if the negation is indeed stratified.
*/
bool rule_set::stratified_negation() {
ptr_vector<rule>::const_iterator it = m_rules.c_ptr();
ptr_vector<rule>::const_iterator end = m_rules.c_ptr()+m_rules.size();
for (; it != end; it++) {
rule * r = *it;
app * head = r->get_head();
func_decl * head_decl = head->get_decl();
unsigned n = r->get_uninterpreted_tail_size();
for (unsigned i = r->get_positive_tail_size(); i < n; i++) {
SASSERT(r->is_neg_tail(i));
func_decl * tail_decl = r->get_tail(i)->get_decl();
unsigned neg_strat = get_predicate_strat(tail_decl);
unsigned head_strat = get_predicate_strat(head_decl);
SASSERT(head_strat>=neg_strat); //head strat can never be lower than that of a tail
if (head_strat == neg_strat) {
return false;
}
}
}
return true;
}
void rule_set::replace_rules(const rule_set & src) {
if (this != &src) {
reset();
add_rules(src);
}
}
void rule_set::add_rules(const rule_set & src) {
SASSERT(!is_closed());
unsigned n = src.get_num_rules();
for (unsigned i = 0; i < n; i++) {
add_rule(src.get_rule(i));
}
inherit_predicates(src);
}
const rule_vector & rule_set::get_predicate_rules(func_decl * pred) const {
decl2rules::obj_map_entry * e = m_head2rules.find_core(pred);
if (!e) {
return m_empty_rule_vector;
}
return *e->get_data().m_value;
}
const rule_set::pred_set_vector & rule_set::get_strats() const {
SASSERT(m_stratifier);
return m_stratifier->get_strats();
}
unsigned rule_set::get_predicate_strat(func_decl * pred) const {
SASSERT(m_stratifier);
return m_stratifier->get_predicate_strat(pred);
}
void rule_set::split_founded_rules(func_decl_set& founded, func_decl_set& non_founded) {
founded.reset();
non_founded.reset();
{
decl2rules::iterator it = begin_grouped_rules(), end = end_grouped_rules();
for (; it != end; ++it) {
non_founded.insert(it->m_key);
}
}
bool change = true;
while (change) {
change = false;
func_decl_set::iterator it = non_founded.begin(), end = non_founded.end();
for (; it != end; ++it) {
rule_vector const& rv = get_predicate_rules(*it);
bool found = false;
for (unsigned i = 0; !found && i < rv.size(); ++i) {
rule const& r = *rv[i];
bool is_founded = true;
for (unsigned j = 0; is_founded && j < r.get_uninterpreted_tail_size(); ++j) {
is_founded = founded.contains(r.get_decl(j));
}
if (is_founded) {
founded.insert(*it);
non_founded.remove(*it);
change = true;
found = true;
}
}
}
}
}
void rule_set::display(std::ostream & out) const {
out << "; rule count: " << get_num_rules() << "\n";
out << "; predicate count: " << m_head2rules.size() << "\n";
func_decl_set::iterator pit = m_output_preds.begin();
func_decl_set::iterator pend = m_output_preds.end();
for (; pit != pend; ++pit) {
out << "; output: " << (*pit)->get_name() << '\n';
}
decl2rules::iterator it = m_head2rules.begin();
decl2rules::iterator end = m_head2rules.end();
for (; it != end; ++it) {
ptr_vector<rule> * rules = it->m_value;
ptr_vector<rule>::iterator it2 = rules->begin();
ptr_vector<rule>::iterator end2 = rules->end();
for (; it2 != end2; ++it2) {
rule * r = *it2;
if (!r->passes_output_thresholds(m_context)) {
continue;
}
r->display(m_context, out);
}
}
}
void rule_set::display_deps( std::ostream & out ) const
{
const pred_set_vector & strats = get_strats();
pred_set_vector::const_iterator sit = strats.begin();
pred_set_vector::const_iterator send = strats.end();
for (; sit!=send; ++sit) {
func_decl_set & strat = **sit;
func_decl_set::iterator fit=strat.begin();
func_decl_set::iterator fend=strat.end();
bool non_empty = false;
for (; fit!=fend; ++fit) {
func_decl * first = *fit;
const func_decl_set & deps = m_deps.get_deps(first);
func_decl_set::iterator dit=deps.begin();
func_decl_set::iterator dend=deps.end();
for (; dit!=dend; ++dit) {
non_empty = true;
func_decl * dep = *dit;
out<<first->get_name()<<" -> "<<dep->get_name()<<"\n";
}
}
if (non_empty && sit!=send) {
out << "\n";
}
}
}
// -----------------------------------
//
// rule_stratifier
//
// -----------------------------------
rule_stratifier::~rule_stratifier() {
comp_vector::iterator it = m_strats.begin();
comp_vector::iterator end = m_strats.end();
for (; it!=end; ++it) {
SASSERT(*it);
dealloc(*it);
}
}
unsigned rule_stratifier::get_predicate_strat(func_decl * pred) const {
unsigned num;
if (!m_pred_strat_nums.find(pred, num)) {
//the number of the predicate is not stored, therefore it did not appear
//in the algorithm and therefore it does not depend on anything and nothing
//depends on it. So it is safe to assign zero strate to it, although it is
//not strictly true.
num = 0;
}
return num;
}
void rule_stratifier::traverse(T* el) {
unsigned p_num;
if (m_preorder_nums.find(el, p_num)) {
if (p_num < m_first_preorder) {
//traversed in a previous sweep
return;
}
if (m_component_nums.contains(el)) {
//we already assigned a component for el
return;
}
while (!m_stack_P.empty()) {
unsigned on_stack_num;
VERIFY( m_preorder_nums.find(m_stack_P.back(), on_stack_num) );
if (on_stack_num <= p_num) {
break;
}
m_stack_P.pop_back();
}
}
else {
p_num=m_next_preorder++;
m_preorder_nums.insert(el, p_num);
m_stack_S.push_back(el);
m_stack_P.push_back(el);
const item_set & children = m_deps.get_deps(el);
item_set::iterator cit=children.begin();
item_set::iterator cend=children.end();
for (; cit!=cend; ++cit) {
traverse(*cit);
}
if (el == m_stack_P.back()) {
unsigned comp_num = m_components.size();
item_set * new_comp = alloc(item_set);
m_components.push_back(new_comp);
T* s_el;
do {
s_el=m_stack_S.back();
m_stack_S.pop_back();
new_comp->insert(s_el);
m_component_nums.insert(s_el, comp_num);
} while (s_el!=el);
m_stack_P.pop_back();
}
}
}
void rule_stratifier::process() {
if (m_deps.empty()) {
return;
}
//detect strong components
rule_dependencies::iterator it = m_deps.begin();
rule_dependencies::iterator end = m_deps.end();
for (; it!=end; ++it) {
T * el = it->m_key;
//we take a note of the preorder number with which this sweep started
m_first_preorder = m_next_preorder;
traverse(el);
}
//do topological sorting
//degres of components (number of inter-component edges ending up in the component)
svector<unsigned> in_degrees;
in_degrees.resize(m_components.size());
//init in_degrees
it = m_deps.begin();
end = m_deps.end();
for (; it != end; ++it) {
T * el = it->m_key;
item_set * out_edges = it->m_value;
unsigned el_comp;
VERIFY( m_component_nums.find(el, el_comp) );
item_set::iterator eit = out_edges->begin();
item_set::iterator eend = out_edges->end();
for (; eit!=eend; ++eit) {
T * tgt = *eit;
unsigned tgt_comp = m_component_nums.find(tgt);
if (el_comp != tgt_comp) {
in_degrees[tgt_comp]++;
}
}
}
// We put components whose indegree is zero to m_strats and assign its
// m_components entry to zero.
unsigned comp_cnt = m_components.size();
for (unsigned i = 0; i < comp_cnt; i++) {
if (in_degrees[i] == 0) {
m_strats.push_back(m_components[i]);
m_components[i] = 0;
}
}
SASSERT(!m_strats.empty()); //the component graph is acyclic and non-empty
//we remove edges from components with zero indegre building the topological ordering
unsigned strats_index = 0;
while (strats_index < m_strats.size()) { //m_strats.size() changes inside the loop!
item_set * comp = m_strats[strats_index];
item_set::iterator cit=comp->begin();
item_set::iterator cend=comp->end();
for (; cit!=cend; ++cit) {
T * el = *cit;
const item_set & deps = m_deps.get_deps(el);
item_set::iterator eit=deps.begin();
item_set::iterator eend=deps.end();
for (; eit!=eend; ++eit) {
T * tgt = *eit;
unsigned tgt_comp;
VERIFY( m_component_nums.find(tgt, tgt_comp) );
//m_components[tgt_comp]==0 means the edge is intra-component.
//Otherwise it would go to another component, but it is not possible, since
//as m_components[tgt_comp]==0, its indegree has already reached zero.
if (m_components[tgt_comp]) {
SASSERT(in_degrees[tgt_comp]>0);
in_degrees[tgt_comp]--;
if (in_degrees[tgt_comp]==0) {
m_strats.push_back(m_components[tgt_comp]);
m_components[tgt_comp] = 0;
}
}
traverse(*cit);
}
}
strats_index++;
}
//we have managed to topologicaly order all the components
SASSERT(std::find_if(m_components.begin(), m_components.end(),
std::bind1st(std::not_equal_to<item_set*>(), (item_set*)0)) == m_components.end());
//reverse the strats array, so that the only the later components would depend on earlier ones
std::reverse(m_strats.begin(), m_strats.end());
SASSERT(m_pred_strat_nums.empty());
unsigned strat_cnt = m_strats.size();
for (unsigned strat_index=0; strat_index<strat_cnt; strat_index++) {
item_set * comp = m_strats[strat_index];
item_set::iterator cit=comp->begin();
item_set::iterator cend=comp->end();
for (; cit != cend; ++cit) {
T * el = *cit;
m_pred_strat_nums.insert(el, strat_index);
}
}
//finalize structures that are not needed anymore
m_preorder_nums.finalize();
m_stack_S.finalize();
m_stack_P.finalize();
m_component_nums.finalize();
m_components.finalize();
}
void rule_stratifier::display(std::ostream& out) const {
m_deps.display(out << "dependencies\n");
out << "strata\n";
for (unsigned i = 0; i < m_strats.size(); ++i) {
item_set::iterator it = m_strats[i]->begin();
item_set::iterator end = m_strats[i]->end();
for (; it != end; ++it) {
out << (*it)->get_name() << " ";
}
out << "\n";
}
}
};

282
src/muz/dl_rule_set.h Normal file
View file

@ -0,0 +1,282 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_rule_set.h
Abstract:
<abstract>
Author:
Leonardo de Moura (leonardo) 2010-05-17.
Revision History:
--*/
#ifndef _DL_RULE_SET_H_
#define _DL_RULE_SET_H_
#include"obj_hashtable.h"
#include"dl_rule.h"
namespace datalog {
class rule_set;
class rule_dependencies {
public:
typedef obj_hashtable<func_decl> item_set;
typedef obj_map<func_decl, item_set * > deps_type;
typedef deps_type::iterator iterator;
private:
/**
Map (dependent -> set of master objects)
Each master object is also present as a key of the map, even if its master set
is empty.
*/
deps_type m_data;
context & m_context;
ptr_vector<expr> m_todo;
ast_mark m_visited;
//we need to take care with removing to avoid memory leaks
void remove_m_data_entry(func_decl * key);
//sometimes we need to return reference to an empty set,
//so we return reference to this one.
item_set m_empty_item_set;
item_set & ensure_key(func_decl * pred);
void insert(func_decl * depending, func_decl * master);
void populate(rule const* r);
public:
rule_dependencies(context& ctx);
rule_dependencies(const rule_dependencies & o, bool reversed = false);
~rule_dependencies();
void reset();
void populate(const rule_set & rules);
void populate(unsigned n, rule * const * rules);
void restrict(const item_set & allowed);
void remove(func_decl * itm);
void remove(const item_set & to_remove);
bool empty() const { return m_data.empty(); }
const item_set & get_deps(func_decl * f) const;
/**
\brief Number of predicates \c f depends on.
*/
unsigned in_degree(func_decl * f) const { return get_deps(f).size(); }
/**
\brief Number of predicates that depend on \c f.
*/
unsigned out_degree(func_decl * f) const;
/**
\brief If the rependency graph is acyclic, put all elements into \c res
ordered so that elements can depend only on elements that are before them.
If the graph is not acyclic, return false.
*/
bool sort_deps(ptr_vector<func_decl> & res);
iterator begin() const { return m_data.begin(); }
iterator end() const { return m_data.end(); }
void display( std::ostream & out ) const;
};
class rule_stratifier {
public:
typedef func_decl T;
typedef obj_hashtable<T> item_set;
typedef ptr_vector<item_set> comp_vector;
typedef obj_map<T, item_set *> deps_type;
private:
const rule_dependencies & m_deps;
comp_vector m_strats;
obj_map<T, unsigned> m_preorder_nums;
ptr_vector<T> m_stack_S;
ptr_vector<T> m_stack_P;
obj_map<T, unsigned> m_component_nums;
comp_vector m_components;
obj_map<T, unsigned> m_pred_strat_nums;
unsigned m_next_preorder;
unsigned m_first_preorder;
/**
Finds strongly connected components using the Gabow's algorithm.
*/
void traverse(T* el);
/**
Calls \c traverse to identify strognly connected components and then
orders them using topological sorting.
*/
void process();
public:
/**
\remark The \c stratifier object keeps a reference to the \c deps object, so
it must exist for the whole lifetime of the \c stratifier object.
*/
rule_stratifier(const rule_dependencies & deps)
: m_deps(deps), m_next_preorder(0)
{
process();
}
~rule_stratifier();
/**
Return vector of components ordered so that the only dependencies are from
later components to earlier.
*/
const comp_vector & get_strats() const { return m_strats; }
unsigned get_predicate_strat(func_decl * pred) const;
void display( std::ostream & out ) const;
};
/**
\brief Datalog "Program" (aka rule set).
*/
class rule_set {
friend class rule_dependencies;
public:
typedef ptr_vector<func_decl_set> pred_set_vector;
typedef obj_map<func_decl, rule_vector*> decl2rules;
private:
typedef obj_map<func_decl, func_decl_set*> decl2deps;
context & m_context;
rule_manager & m_rule_manager;
rule_ref_vector m_rules; //!< all rules
decl2rules m_head2rules; //!< mapping from head symbol to rules.
rule_dependencies m_deps; //!< dependencies
scoped_ptr<rule_stratifier> m_stratifier; //!< contains stratifier object iff the rule_set is closed
func_decl_set m_output_preds; //!< output predicates
obj_map<func_decl,func_decl*> m_orig2pred;
obj_map<func_decl,func_decl*> m_pred2orig;
func_decl_ref_vector m_refs;
//sometimes we need to return reference to an empty rule_vector,
//so we return reference to this one.
rule_vector m_empty_rule_vector;
void compute_deps();
void compute_tc_deps();
bool stratified_negation();
public:
rule_set(context & ctx);
rule_set(const rule_set & rs);
~rule_set();
ast_manager & get_manager() const;
rule_manager & get_rule_manager() const { return const_cast<rule_manager&>(m_rule_manager); }
context& get_context() const { return m_context; }
void inherit_predicates(rule_set const& other);
void inherit_predicate(rule_set const& other, func_decl* orig, func_decl* pred);
func_decl* get_orig(func_decl* pred) const;
func_decl* get_pred(func_decl* orig) const;
/**
\brief Add rule \c r to the rule set.
*/
void add_rule(rule * r);
/**
\brief Remove rule \c r from the rule set.
*/
void del_rule(rule * r);
/**
\brief Add all rules from a different rule_set.
*/
void add_rules(const rule_set& src);
void replace_rules(const rule_set& other);
/**
\brief This method should be invoked after all rules are added to the rule set.
It will check if the negation is indeed stratified.
Return true if succeeded.
\remark If new rules are added, the rule_set will be "reopen".
*/
bool close();
void ensure_closed();
/**
\brief Undo the effect of the \c close() operation.
*/
void reopen();
bool is_closed() const { return m_stratifier != 0; }
unsigned get_num_rules() const { return m_rules.size(); }
bool empty() const { return m_rules.size() == 0; }
rule * get_rule(unsigned i) const { return m_rules[i]; }
rule * last() const { return m_rules[m_rules.size()-1]; }
rule_ref_vector const& get_rules() const { return m_rules; }
const rule_vector & get_predicate_rules(func_decl * pred) const;
bool contains(func_decl* pred) const { return m_head2rules.contains(pred); }
const rule_stratifier & get_stratifier() const {
SASSERT(m_stratifier);
return *m_stratifier;
}
const pred_set_vector & get_strats() const;
unsigned get_predicate_strat(func_decl * pred) const;
const rule_dependencies & get_dependencies() const { SASSERT(is_closed()); return m_deps; }
// split predicats into founded and non-founded.
void split_founded_rules(func_decl_set& founded, func_decl_set& non_founded);
void reset();
void set_output_predicate(func_decl * pred) { m_refs.push_back(pred); m_output_preds.insert(pred); }
bool is_output_predicate(func_decl * pred) const { return m_output_preds.contains(pred); }
const func_decl_set & get_output_predicates() const { return m_output_preds; }
func_decl* get_output_predicate() const { SASSERT(m_output_preds.size() == 1); return *m_output_preds.begin(); }
void display(std::ostream & out) const;
/**
\brief Output rule dependencies.
The rule set must be closed before calling this function.
*/
void display_deps(std::ostream & out) const;
typedef rule * const * iterator;
iterator begin() const { return m_rules.c_ptr(); }
iterator end() const { return m_rules.c_ptr()+m_rules.size(); }
decl2rules::iterator begin_grouped_rules() const { return m_head2rules.begin(); }
decl2rules::iterator end_grouped_rules() const { return m_head2rules.end(); }
};
inline std::ostream& operator<<(std::ostream& out, rule_set const& r) { r.display(out); return out; }
};
#endif /* _DL_RULE_SET_H_ */

View file

@ -0,0 +1,84 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_rule_subsumption_index.cpp
Abstract:
Subsumption index for rules.
Currently an underapproximation (fails to identify some subsumptions).
Author:
Krystof Hoder (t-khoder) 2011-10-10.
Revision History:
--*/
#include <sstream>
#include "ast_pp.h"
#include "dl_rule_subsumption_index.h"
namespace datalog {
// -----------------------------------
//
// rule_subsumption_index
//
// -----------------------------------
void rule_subsumption_index::handle_unconditioned_rule(rule * r) {
SASSERT(r->get_tail_size()==0);
app * head = r->get_head();
func_decl * pred = head->get_decl();
app_set * head_set;
if(!m_unconditioned_heads.find(pred, head_set)) {
head_set = alloc(app_set);
m_unconditioned_heads.insert(pred, head_set);
}
head_set->insert(head);
}
void rule_subsumption_index::add(rule * r) {
m_ref_holder.push_back(r);
if(r->get_tail_size()==0) {
handle_unconditioned_rule(r);
}
m_rule_set.insert(r);
}
bool rule_subsumption_index::is_subsumed(app * query) {
func_decl * pred = query->get_decl();
app_set * head_set;
if(m_unconditioned_heads.find(pred, head_set)) {
if(head_set->contains(query)) {
return true;
}
}
return false;
}
bool rule_subsumption_index::is_subsumed(rule * r) {
app * head = r->get_head();
if(is_subsumed(head)) {
return true;
}
if(m_rule_set.contains(r)) {
return true;
}
return false;
}
};

View file

@ -0,0 +1,65 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_rule_subsumption_index.h
Abstract:
Subsumption index for rules.
Currently an underapproximation (fails to identify some subsumptions).
Author:
Krystof Hoder (t-khoder) 2011-10-10.
Revision History:
--*/
#ifndef _DL_RULE_SUBSUMPTION_INDEX_H_
#define _DL_RULE_SUBSUMPTION_INDEX_H_
#include "dl_context.h"
namespace datalog {
class rule_subsumption_index {
//private and undefined copy constroctor
rule_subsumption_index(rule_subsumption_index const&);
//private and undefined operator=
rule_subsumption_index& operator=(rule_subsumption_index const&);
typedef obj_hashtable<app> app_set;
ast_manager & m;
context & m_context;
rule_ref_vector m_ref_holder;
obj_map<func_decl, app_set *> m_unconditioned_heads;
hashtable<rule *, rule_hash_proc, rule_eq_proc> m_rule_set;
void handle_unconditioned_rule(rule * r);
public:
rule_subsumption_index(context & ctx) :
m(ctx.get_manager()),
m_context(ctx),
m_ref_holder(ctx.get_rule_manager()) {}
~rule_subsumption_index() {
reset_dealloc_values(m_unconditioned_heads);
}
void add(rule * r);
bool is_subsumed(rule * r);
bool is_subsumed(app * query);
};
};
#endif /* _DL_RULE_SUBSUMPTION_INDEX_H_ */

View file

@ -0,0 +1,159 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_rule_transformer.cpp
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-09-20.
Revision History:
--*/
#include <algorithm>
#include<typeinfo>
#include"dl_context.h"
#include"dl_rule_transformer.h"
#include"stopwatch.h"
namespace datalog {
rule_transformer::rule_transformer(context & ctx)
: m_context(ctx), m_rule_manager(m_context.get_rule_manager()), m_dirty(false) {
}
rule_transformer::~rule_transformer() {
reset();
}
void rule_transformer::reset() {
plugin_vector::iterator it = m_plugins.begin();
plugin_vector::iterator end = m_plugins.end();
for(; it!=end; ++it) {
dealloc(*it);
}
m_plugins.reset();
m_dirty = false;
}
void rule_transformer::cancel() {
plugin_vector::iterator it = m_plugins.begin();
plugin_vector::iterator end = m_plugins.end();
for(; it!=end; ++it) {
(*it)->cancel();
}
}
struct rule_transformer::plugin_comparator {
bool operator()(rule_transformer::plugin * p1, rule_transformer::plugin * p2) {
return p1->get_priority() > p2->get_priority();
}
};
void rule_transformer::ensure_ordered() {
if (m_dirty) {
std::sort(m_plugins.begin(), m_plugins.end(), plugin_comparator());
m_dirty = false;
}
}
void rule_transformer::register_plugin(plugin * p) {
m_plugins.push_back(p);
p->attach(*this);
m_dirty=true;
}
bool rule_transformer::operator()(rule_set & rules) {
ensure_ordered();
bool modified = false;
TRACE("dl_rule_transf",
tout<<"init:\n";
rules.display(tout);
);
rule_set* new_rules = alloc(rule_set, rules);
plugin_vector::iterator it = m_plugins.begin();
plugin_vector::iterator end = m_plugins.end();
for(; it!=end && !m_context.canceled(); ++it) {
plugin & p = **it;
IF_VERBOSE(1, verbose_stream() << "(transform " << typeid(p).name() << "...";);
stopwatch sw;
sw.start();
rule_set * new_rules1 = p(*new_rules);
sw.stop();
double sec = sw.get_seconds();
if (sec < 0.001) sec = 0.0;
if (!new_rules1) {
IF_VERBOSE(1, verbose_stream() << "no-op " << sec << "s)\n";);
continue;
}
if (p.can_destratify_negation() &&
!new_rules1->is_closed() &&
!new_rules1->close()) {
warning_msg("a rule transformation skipped "
"because it destratified negation");
dealloc(new_rules1);
IF_VERBOSE(1, verbose_stream() << "no-op " << sec << "s)\n";);
continue;
}
modified = true;
dealloc(new_rules);
new_rules = new_rules1;
new_rules->ensure_closed();
IF_VERBOSE(1, verbose_stream() << new_rules->get_num_rules() << " rules " << sec << "s)\n";);
TRACE("dl_rule_transf",
tout << typeid(p).name()<<":\n";
new_rules->display(tout);
);
}
if (modified) {
rules.replace_rules(*new_rules);
}
dealloc(new_rules);
return modified;
}
//------------------------------
//
//rule_transformer::plugin
//
//------------------------------
void rule_transformer::plugin::remove_duplicate_tails(app_ref_vector& tail, svector<bool>& tail_neg)
{
//one set for positive and one for negative
obj_hashtable<app> tail_apps[2];
unsigned i=0;
while(i<tail.size()) {
unsigned set_idx = tail_neg[i] ? 1 : 0;
if (tail_apps[set_idx].contains(tail[i].get())) {
if (i!=tail.size()-1) {
tail[i] = tail.back();
tail_neg[i] = tail_neg.back();
}
tail.pop_back();
tail_neg.pop_back();
}
else {
tail_apps[set_idx].insert(tail[i].get());
i++;
}
}
}
};

View file

@ -0,0 +1,117 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_rule_transformer.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-09-20.
Revision History:
--*/
#ifndef _DL_RULE_TRANSFORMER_H_
#define _DL_RULE_TRANSFORMER_H_
#include"map.h"
#include"vector.h"
#include"dl_rule.h"
#include"dl_rule_set.h"
namespace datalog {
class context;
class rule_transformer {
public:
class plugin;
private:
friend class plugin;
typedef svector<plugin*> plugin_vector;
struct plugin_comparator;
context & m_context;
rule_manager & m_rule_manager;
bool m_dirty;
svector<plugin*> m_plugins;
void ensure_ordered();
public:
rule_transformer(context & ctx);
~rule_transformer();
/**
\brief Reset all registered transformers.
*/
void reset();
void cancel();
/**
\brief Add a plugin for rule transformation.
The rule_transformer object takes ownership of the plugin object.
*/
void register_plugin(plugin * p);
/**
\brief Transform the rule set using the registered transformation plugins. If the rule
set has changed, return true; otherwise return false.
*/
bool operator()(rule_set & rules);
};
class rule_transformer::plugin {
friend class rule_transformer;
unsigned m_priority;
bool m_can_destratify_negation;
rule_transformer * m_transformer;
void attach(rule_transformer & transformer) { m_transformer = &transformer; }
protected:
/**
\brief Create a plugin object for rule_transformer.
The priority argument determines in what order will rules be applied to the rules
(higher priority plugins will be applied first).
*/
plugin(unsigned priority, bool can_destratify_negation = false) : m_priority(priority),
m_can_destratify_negation(can_destratify_negation), m_transformer(0) {}
public:
virtual ~plugin() {}
unsigned get_priority() { return m_priority; }
bool can_destratify_negation() const { return m_can_destratify_negation; }
virtual void cancel() {}
/**
\brief Return \c rule_set object with containing transformed rules or 0 if no
transformation was done.
The caller takes ownership of the returned \c rule_set object.
*/
virtual rule_set * operator()(rule_set const & source) = 0;
/**
Removes duplicate tails.
*/
static void remove_duplicate_tails(app_ref_vector& tail, svector<bool>& tail_neg);
};
};
#endif /* _DL_RULE_TRANSFORMER_H_ */

View file

@ -0,0 +1,666 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_mk_explanations.cpp
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-11-08.
Revision History:
--*/
#include <sstream>
#include"ast_pp.h"
#include"dl_sieve_relation.h"
namespace datalog {
// -----------------------------------
//
// sieve_relation
//
// -----------------------------------
sieve_relation::sieve_relation(sieve_relation_plugin & p, const relation_signature & s,
const bool * inner_columns, relation_base * inner)
: relation_base(p, s), m_inner_cols(s.size(), inner_columns), m_inner(inner) {
unsigned n = s.size();
for(unsigned i=0; i<n; i++) {
if(inner_columns && inner_columns[i]) {
unsigned inner_idx = m_inner2sig.size();
SASSERT(get_inner().get_signature()[inner_idx]==s[i]);
m_sig2inner.push_back(inner_idx);
m_inner2sig.push_back(i);
}
else {
m_sig2inner.push_back(UINT_MAX);
m_ignored_cols.push_back(i);
}
}
set_kind(p.get_relation_kind(*this, inner_columns));
}
void sieve_relation::add_fact(const relation_fact & f) {
relation_fact inner_f = f;
project_out_vector_columns(inner_f, m_ignored_cols);
get_inner().add_fact(inner_f);
}
bool sieve_relation::contains_fact(const relation_fact & f) const {
relation_fact inner_f = f;
project_out_vector_columns(inner_f, m_ignored_cols);
return get_inner().contains_fact(inner_f);
}
sieve_relation * sieve_relation::clone() const {
relation_base * new_inner = get_inner().clone();
return get_plugin().mk_from_inner(get_signature(), m_inner_cols.c_ptr(), new_inner);
}
relation_base * sieve_relation::complement(func_decl* p) const {
//this is not precisely a complement, because we still treat the ignored collumns as
//full, but it should give reasonable results inside the product relation
relation_base * new_inner = get_inner().complement(p);
return get_plugin().mk_from_inner(get_signature(), m_inner_cols.c_ptr(), new_inner);
}
void sieve_relation::to_formula(expr_ref& fml) const {
ast_manager& m = fml.get_manager();
expr_ref_vector s(m);
expr_ref tmp(m);
relation_signature const& sig = get_inner().get_signature();
unsigned sz = sig.size();
for (unsigned i = sz ; i > 0; ) {
--i;
unsigned idx = m_inner2sig[i];
s.push_back(m.mk_var(idx, sig[i]));
}
get_inner().to_formula(tmp);
get_plugin().get_context().get_var_subst()(tmp, sz, s.c_ptr(), fml);
}
void sieve_relation::display(std::ostream & out) const {
out << "Sieve relation ";
print_container(m_inner_cols, out);
out <<"\n";
get_inner().display(out);
}
// -----------------------------------
//
// sieve_relation_plugin
//
// -----------------------------------
sieve_relation_plugin & sieve_relation_plugin::get_plugin(relation_manager & rmgr) {
sieve_relation_plugin * res = static_cast<sieve_relation_plugin *>(rmgr.get_relation_plugin(get_name()));
if(!res) {
res = alloc(sieve_relation_plugin, rmgr);
rmgr.register_plugin(res);
}
return *res;
}
sieve_relation& sieve_relation_plugin::get(relation_base& r) {
return dynamic_cast<sieve_relation&>(r);
}
sieve_relation const & sieve_relation_plugin::get(relation_base const& r) {
return dynamic_cast<sieve_relation const&>(r);
}
sieve_relation* sieve_relation_plugin::get(relation_base* r) {
return dynamic_cast<sieve_relation*>(r);
}
sieve_relation const* sieve_relation_plugin::get(relation_base const* r) {
return dynamic_cast<sieve_relation const*>(r);
}
sieve_relation_plugin::sieve_relation_plugin(relation_manager & manager)
: relation_plugin(get_name(), manager, ST_SIEVE_RELATION),
m_spec_store(*this) {}
void sieve_relation_plugin::initialize(family_id fid) {
relation_plugin::initialize(fid);
m_spec_store.add_available_kind(get_kind());
}
family_id sieve_relation_plugin::get_relation_kind(const relation_signature & sig,
const bool * inner_columns, family_id inner_kind) {
rel_spec spec(sig.size(), inner_columns, inner_kind);
return m_spec_store.get_relation_kind(sig, spec);
}
family_id sieve_relation_plugin::get_relation_kind(sieve_relation & r, const bool * inner_columns) {
const relation_signature & sig = r.get_signature();
return get_relation_kind(sig, inner_columns, r.get_inner().get_kind());
}
void sieve_relation_plugin::extract_inner_columns(const relation_signature & s, relation_plugin & inner,
svector<bool> & inner_columns) {
SASSERT(inner_columns.size()==s.size());
unsigned n = s.size();
relation_signature inner_sig_singleton;
for(unsigned i=0; i<n; i++) {
inner_sig_singleton.reset();
inner_sig_singleton.push_back(s[i]);
inner_columns[i] = inner.can_handle_signature(inner_sig_singleton);
}
#if Z3DEBUG
//we assume that if a relation plugin can handle two sets of columns separetely,
//it can also handle them in one relation
relation_signature inner_sig;
collect_inner_signature(s, inner_columns, inner_sig);
SASSERT(inner.can_handle_signature(inner_sig));
#endif
}
void sieve_relation_plugin::collect_inner_signature(const relation_signature & s,
const svector<bool> & inner_columns, relation_signature & inner_sig) {
SASSERT(inner_columns.size()==s.size());
inner_sig.reset();
unsigned n = s.size();
for(unsigned i=0; i<n; i++) {
if(inner_columns[i]) {
inner_sig.push_back(s[i]);
}
}
}
void sieve_relation_plugin::extract_inner_signature(const relation_signature & s,
relation_signature & inner_sig) {
UNREACHABLE();
#if 0
svector<bool> inner_cols(s.size());
extract_inner_columns(s, inner_cols.c_ptr());
collect_inner_signature(s, inner_cols, inner_sig);
#endif
}
bool sieve_relation_plugin::can_handle_signature(const relation_signature & s) {
//we do not want this plugin to handle anything by default
return false;
#if 0
relation_signature inner_sig;
extract_inner_signature(s, inner_sig);
SASSERT(inner_sig.size()<=s.size());
return !inner_sig.empty() && inner_sig.size()!=s.size();
#endif
}
sieve_relation * sieve_relation_plugin::mk_from_inner(const relation_signature & s, const bool * inner_columns,
relation_base * inner_rel) {
SASSERT(!inner_rel->get_plugin().is_sieve_relation()); //it does not make sense to make a sieve of a sieve
return alloc(sieve_relation, *this, s, inner_columns, inner_rel);
}
sieve_relation * sieve_relation_plugin::mk_empty(const sieve_relation & original) {
return static_cast<sieve_relation *>(mk_empty(original.get_signature(), original.get_kind()));
}
relation_base * sieve_relation_plugin::mk_empty(const relation_base & original) {
return mk_empty(static_cast<const sieve_relation &>(original));
}
relation_base * sieve_relation_plugin::mk_empty(const relation_signature & s, family_id kind) {
rel_spec spec;
m_spec_store.get_relation_spec(s, kind, spec);
relation_signature inner_sig;
collect_inner_signature(s, spec.m_inner_cols, inner_sig);
relation_base * inner = get_manager().mk_empty_relation(inner_sig, spec.m_inner_kind);
return mk_from_inner(s, spec.m_inner_cols.c_ptr(), inner);
}
relation_base * sieve_relation_plugin::mk_empty(const relation_signature & s) {
UNREACHABLE();
return 0;
#if 0
svector<bool> inner_cols(s.size());
extract_inner_columns(s, inner_cols.c_ptr());
return mk_empty(s, inner_cols.c_ptr());
#endif
}
sieve_relation * sieve_relation_plugin::mk_empty(const relation_signature & s, relation_plugin & inner_plugin) {
SASSERT(!inner_plugin.is_sieve_relation()); //it does not make sense to make a sieve of a sieve
svector<bool> inner_cols(s.size());
extract_inner_columns(s, inner_plugin, inner_cols);
relation_signature inner_sig;
collect_inner_signature(s, inner_cols, inner_sig);
relation_base * inner_rel = inner_plugin.mk_empty(inner_sig);
return mk_from_inner(s, inner_cols, inner_rel);
}
relation_base * sieve_relation_plugin::mk_full(func_decl* p, const relation_signature & s) {
relation_signature empty_sig;
relation_plugin& plugin = get_manager().get_appropriate_plugin(s);
relation_base * inner = plugin.mk_full(p, empty_sig, null_family_id);
svector<bool> inner_cols;
inner_cols.resize(s.size(), false);
return mk_from_inner(s, inner_cols, inner);
}
sieve_relation * sieve_relation_plugin::mk_full(func_decl* p, const relation_signature & s, relation_plugin & inner_plugin) {
SASSERT(!inner_plugin.is_sieve_relation()); //it does not make sense to make a sieve of a sieve
svector<bool> inner_cols(s.size());
extract_inner_columns(s, inner_plugin, inner_cols);
relation_signature inner_sig;
collect_inner_signature(s, inner_cols, inner_sig);
relation_base * inner_rel = inner_plugin.mk_full(p, inner_sig, null_family_id);
return mk_from_inner(s, inner_cols, inner_rel);
}
class sieve_relation_plugin::join_fn : public convenient_relation_join_fn {
sieve_relation_plugin & m_plugin;
unsigned_vector m_inner_cols_1;
unsigned_vector m_inner_cols_2;
svector<bool> m_result_inner_cols;
scoped_ptr<relation_join_fn> m_inner_join_fun;
public:
join_fn(sieve_relation_plugin & p, const relation_base & r1, const relation_base & r2, unsigned col_cnt,
const unsigned * cols1, const unsigned * cols2, relation_join_fn * inner_join_fun)
: convenient_relation_join_fn(r1.get_signature(), r2.get_signature(), col_cnt, cols1, cols2),
m_plugin(p),
m_inner_join_fun(inner_join_fun) {
bool r1_sieved = r1.get_plugin().is_sieve_relation();
bool r2_sieved = r2.get_plugin().is_sieve_relation();
const sieve_relation * sr1 = r1_sieved ? static_cast<const sieve_relation *>(&r1) : 0;
const sieve_relation * sr2 = r2_sieved ? static_cast<const sieve_relation *>(&r2) : 0;
if(r1_sieved) {
m_result_inner_cols.append(sr1->m_inner_cols);
}
else {
m_result_inner_cols.resize(r1.get_signature().size(), true);
}
if(r2_sieved) {
m_result_inner_cols.append(sr2->m_inner_cols);
}
else {
m_result_inner_cols.resize(m_result_inner_cols.size() + r2.get_signature().size(), true);
}
}
virtual relation_base * operator()(const relation_base & r1, const relation_base & r2) {
bool r1_sieved = r1.get_plugin().is_sieve_relation();
bool r2_sieved = r2.get_plugin().is_sieve_relation();
SASSERT(r1_sieved || r2_sieved);
const sieve_relation * sr1 = r1_sieved ? static_cast<const sieve_relation *>(&r1) : 0;
const sieve_relation * sr2 = r2_sieved ? static_cast<const sieve_relation *>(&r2) : 0;
const relation_base & inner1 = r1_sieved ? sr1->get_inner() : r1;
const relation_base & inner2 = r2_sieved ? sr2->get_inner() : r2;
relation_base * inner_res = (*m_inner_join_fun)(inner1, inner2);
return m_plugin.mk_from_inner(get_result_signature(), m_result_inner_cols.c_ptr(), inner_res);
}
};
relation_join_fn * sieve_relation_plugin::mk_join_fn(const relation_base & r1, const relation_base & r2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2) {
if( &r1.get_plugin()!=this && &r2.get_plugin()!=this ) {
//we create just operations that involve the current plugin
return 0;
}
bool r1_sieved = r1.get_plugin().is_sieve_relation();
bool r2_sieved = r2.get_plugin().is_sieve_relation();
const sieve_relation * sr1 = r1_sieved ? static_cast<const sieve_relation *>(&r1) : 0;
const sieve_relation * sr2 = r2_sieved ? static_cast<const sieve_relation *>(&r2) : 0;
const relation_base & inner1 = r1_sieved ? sr1->get_inner() : r1;
const relation_base & inner2 = r2_sieved ? sr2->get_inner() : r2;
unsigned_vector inner_cols1;
unsigned_vector inner_cols2;
for(unsigned i=0; i<col_cnt; i++) {
//if at least one end of an equality is not an inner column, we ignore that equality
//(which introduces imprecision)
if(r1_sieved && !sr1->is_inner_col(cols1[i])) {
continue;
}
if(r2_sieved && !sr2->is_inner_col(cols2[i])) {
continue;
}
inner_cols1.push_back( r1_sieved ? sr1->get_inner_col(cols1[i]) : cols1[i] );
inner_cols2.push_back( r2_sieved ? sr2->get_inner_col(cols2[i]) : cols2[i] );
}
relation_join_fn * inner_join_fun = get_manager().mk_join_fn(inner1, inner2, inner_cols1, inner_cols2, false);
if(!inner_join_fun) {
return 0;
}
return alloc(join_fn, *this, r1, r2, col_cnt, cols1, cols2, inner_join_fun);
}
class sieve_relation_plugin::transformer_fn : public convenient_relation_transformer_fn {
svector<bool> m_result_inner_cols;
scoped_ptr<relation_transformer_fn> m_inner_fun;
public:
transformer_fn(relation_transformer_fn * inner_fun, const relation_signature & result_sig,
const bool * result_inner_cols)
: m_result_inner_cols(result_sig.size(), result_inner_cols), m_inner_fun(inner_fun) {
get_result_signature() = result_sig;
}
virtual relation_base * operator()(const relation_base & r0) {
SASSERT(r0.get_plugin().is_sieve_relation());
const sieve_relation & r = static_cast<const sieve_relation &>(r0);
sieve_relation_plugin & plugin = r.get_plugin();
relation_base * inner_res = (*m_inner_fun)(r.get_inner());
return plugin.mk_from_inner(get_result_signature(), m_result_inner_cols.c_ptr(), inner_res);
}
};
relation_transformer_fn * sieve_relation_plugin::mk_project_fn(const relation_base & r0, unsigned col_cnt,
const unsigned * removed_cols) {
if(&r0.get_plugin()!=this) {
return 0;
}
const sieve_relation & r = static_cast<const sieve_relation &>(r0);
unsigned_vector inner_removed_cols;
for(unsigned i=0; i<col_cnt; i++) {
unsigned col = removed_cols[i];
if(r.is_inner_col(col)) {
inner_removed_cols.push_back(r.get_inner_col(col));
}
}
svector<bool> result_inner_cols = r.m_inner_cols;
project_out_vector_columns(result_inner_cols, col_cnt, removed_cols);
relation_signature result_sig;
relation_signature::from_project(r.get_signature(), col_cnt, removed_cols, result_sig);
relation_transformer_fn * inner_fun;
if(inner_removed_cols.empty()) {
inner_fun = alloc(identity_relation_transformer_fn);
}
else {
inner_fun = get_manager().mk_project_fn(r.get_inner(), inner_removed_cols);
}
if(!inner_fun) {
return 0;
}
return alloc(transformer_fn, inner_fun, result_sig, result_inner_cols.c_ptr());
}
relation_transformer_fn * sieve_relation_plugin::mk_rename_fn(const relation_base & r0,
unsigned cycle_len, const unsigned * permutation_cycle) {
if(&r0.get_plugin()!=this) {
return 0;
}
const sieve_relation & r = static_cast<const sieve_relation &>(r0);
unsigned sig_sz = r.get_signature().size();
unsigned_vector permutation;
add_sequence(0, sig_sz, permutation);
permutate_by_cycle(permutation, cycle_len, permutation_cycle);
bool inner_identity;
unsigned_vector inner_permutation;
collect_sub_permutation(permutation, r.m_sig2inner, inner_permutation, inner_identity);
svector<bool> result_inner_cols = r.m_inner_cols;
permutate_by_cycle(result_inner_cols, cycle_len, permutation_cycle);
relation_signature result_sig;
relation_signature::from_rename(r.get_signature(), cycle_len, permutation_cycle, result_sig);
relation_transformer_fn * inner_fun =
get_manager().mk_permutation_rename_fn(r.get_inner(), inner_permutation);
if(!inner_fun) {
return 0;
}
return alloc(transformer_fn, inner_fun, result_sig, result_inner_cols.c_ptr());
}
class sieve_relation_plugin::union_fn : public relation_union_fn {
scoped_ptr<relation_union_fn> m_union_fun;
public:
union_fn(relation_union_fn * union_fun) : m_union_fun(union_fun) {}
virtual void operator()(relation_base & tgt, const relation_base & src, relation_base * delta) {
bool tgt_sieved = tgt.get_plugin().is_sieve_relation();
bool src_sieved = src.get_plugin().is_sieve_relation();
bool delta_sieved = delta && delta->get_plugin().is_sieve_relation();
sieve_relation * stgt = tgt_sieved ? static_cast<sieve_relation *>(&tgt) : 0;
const sieve_relation * ssrc = src_sieved ? static_cast<const sieve_relation *>(&src) : 0;
sieve_relation * sdelta = delta_sieved ? static_cast<sieve_relation *>(delta) : 0;
relation_base & itgt = tgt_sieved ? stgt->get_inner() : tgt;
const relation_base & isrc = src_sieved ? ssrc->get_inner() : src;
relation_base * idelta = delta_sieved ? &sdelta->get_inner() : delta;
(*m_union_fun)(itgt, isrc, idelta);
}
};
relation_union_fn * sieve_relation_plugin::mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta) {
if(&tgt.get_plugin()!=this && &src.get_plugin()!=this && (delta && &delta->get_plugin()!=this)) {
//we create the operation only if it involves this plugin
return 0;
}
bool tgt_sieved = tgt.get_plugin().is_sieve_relation();
bool src_sieved = src.get_plugin().is_sieve_relation();
bool delta_sieved = delta && delta->get_plugin().is_sieve_relation();
const sieve_relation * stgt = tgt_sieved ? static_cast<const sieve_relation *>(&tgt) : 0;
const sieve_relation * ssrc = src_sieved ? static_cast<const sieve_relation *>(&src) : 0;
const sieve_relation * sdelta = delta_sieved ? static_cast<const sieve_relation *>(delta) : 0;
const relation_base & itgt = tgt_sieved ? stgt->get_inner() : tgt;
const relation_base & isrc = src_sieved ? ssrc->get_inner() : src;
const relation_base * idelta = delta_sieved ? &sdelta->get_inner() : delta;
//Now we require that the sieved and inner columns must match on all relations.
//We may want to allow for some cases of misalignment even though it could introcude imprecision
if( tgt_sieved && src_sieved && (!delta || delta_sieved) ) {
if( !vectors_equal(stgt->m_inner_cols, ssrc->m_inner_cols)
|| (delta && !vectors_equal(stgt->m_inner_cols, sdelta->m_inner_cols)) ) {
return 0;
}
}
else {
if( (stgt && !stgt->no_sieved_columns())
|| (ssrc && !ssrc->no_sieved_columns())
|| (sdelta && !sdelta->no_sieved_columns()) ) {
//We have an unsieved relation and then some relation with some sieved columns,
//which means there is an misalignment.
return 0;
}
}
relation_union_fn * union_fun = get_manager().mk_union_fn(itgt, isrc, idelta);
if(!union_fun) {
return 0;
}
return alloc(union_fn, union_fun);
}
class sieve_relation_plugin::filter_fn : public relation_mutator_fn {
scoped_ptr<relation_mutator_fn> m_inner_fun;
public:
filter_fn(relation_mutator_fn * inner_fun)
: m_inner_fun(inner_fun) {}
virtual void operator()(relation_base & r0) {
SASSERT(r0.get_plugin().is_sieve_relation());
sieve_relation & r = static_cast<sieve_relation &>(r0);
(*m_inner_fun)(r.get_inner());
}
};
relation_mutator_fn * sieve_relation_plugin::mk_filter_identical_fn(const relation_base & r0,
unsigned col_cnt, const unsigned * identical_cols) {
if(&r0.get_plugin()!=this) {
return 0;
}
const sieve_relation & r = static_cast<const sieve_relation &>(r0);
unsigned_vector inner_icols;
//we ignore the columns which do not belong to the inner relation (which introduces imprecision)
for(unsigned i=0; i<col_cnt; i++) {
unsigned col = identical_cols[i];
if(r.is_inner_col(col)) {
inner_icols.push_back(r.get_inner_col(col));
}
}
if(inner_icols.size()<2) {
return alloc(identity_relation_mutator_fn);
}
relation_mutator_fn * inner_fun = get_manager().mk_filter_identical_fn(r.get_inner(), inner_icols);
if(!inner_fun) {
return 0;
}
return alloc(filter_fn, inner_fun);
}
relation_mutator_fn * sieve_relation_plugin::mk_filter_equal_fn(const relation_base & r0,
const relation_element & value, unsigned col) {
if(&r0.get_plugin()!=this) {
return 0;
}
const sieve_relation & r = static_cast<const sieve_relation &>(r0);
if(!r.is_inner_col(col)) {
//if the column which do not belong to the inner relation, we do nothing (which introduces imprecision)
return alloc(identity_relation_mutator_fn);
}
unsigned inner_col = r.get_inner_col(col);
relation_mutator_fn * inner_fun = get_manager().mk_filter_equal_fn(r.get_inner(), value, inner_col);
if(!inner_fun) {
return 0;
}
return alloc(filter_fn, inner_fun);
}
relation_mutator_fn * sieve_relation_plugin::mk_filter_interpreted_fn(const relation_base & rb,
app * condition) {
if(&rb.get_plugin()!=this) {
return 0;
}
ast_manager & m = get_ast_manager();
const sieve_relation & r = static_cast<const sieve_relation &>(rb);
const relation_signature sig = r.get_signature();
unsigned sz = sig.size();
var_idx_set& cond_vars = get_context().get_rule_manager().collect_vars(condition);
expr_ref_vector subst_vect(m);
subst_vect.resize(sz);
unsigned subst_ofs = sz-1;
for(unsigned i=0; i<sz; i++) {
if(!cond_vars.contains(i)) {
continue;
}
if(!r.is_inner_col(i)) {
//If the condition involves columns which do not belong to the inner relation,
//we do nothing (which introduces imprecision).
//Maybe we might try to do some quantifier elimination...
return alloc(identity_relation_mutator_fn);
}
subst_vect[subst_ofs-i] = m.mk_var(r.m_sig2inner[i], sig[i]);
}
expr_ref inner_cond(m);
get_context().get_var_subst()(condition, subst_vect.size(), subst_vect.c_ptr(), inner_cond);
relation_mutator_fn * inner_fun = get_manager().mk_filter_interpreted_fn(r.get_inner(), to_app(inner_cond));
if(!inner_fun) {
return 0;
}
return alloc(filter_fn, inner_fun);
}
class sieve_relation_plugin::negation_filter_fn : public relation_intersection_filter_fn {
scoped_ptr<relation_intersection_filter_fn> m_inner_fun;
public:
negation_filter_fn(relation_intersection_filter_fn * inner_fun)
: m_inner_fun(inner_fun) {}
virtual void operator()(relation_base & r, const relation_base & neg) {
bool r_sieved = r.get_plugin().is_sieve_relation();
bool neg_sieved = neg.get_plugin().is_sieve_relation();
SASSERT(r_sieved || neg_sieved);
sieve_relation * sr = r_sieved ? static_cast<sieve_relation *>(&r) : 0;
const sieve_relation * sneg = neg_sieved ? static_cast<const sieve_relation *>(&neg) : 0;
relation_base & inner_r = r_sieved ? sr->get_inner() : r;
const relation_base & inner_neg = neg_sieved ? sneg->get_inner() : neg;
(*m_inner_fun)(inner_r, inner_neg);
}
};
relation_intersection_filter_fn * sieve_relation_plugin::mk_filter_by_negation_fn(const relation_base & r,
const relation_base & neg, unsigned col_cnt, const unsigned * r_cols,
const unsigned * neg_cols) {
if(&r.get_plugin()!=this && &neg.get_plugin()!=this) {
//we create just operations that involve the current plugin
return 0;
}
bool r_sieved = r.get_plugin().is_sieve_relation();
bool neg_sieved = neg.get_plugin().is_sieve_relation();
SASSERT(r_sieved || neg_sieved);
const sieve_relation * sr = r_sieved ? static_cast<const sieve_relation *>(&r) : 0;
const sieve_relation * sneg = neg_sieved ? static_cast<const sieve_relation *>(&neg) : 0;
const relation_base & inner_r = r_sieved ? sr->get_inner() : r;
const relation_base & inner_neg = neg_sieved ? sneg->get_inner() : neg;
unsigned_vector ir_cols;
unsigned_vector ineg_cols;
for(unsigned i=0; i<col_cnt; i++) {
//if at least one end of an equality is not an inner column, we ignore that equality
//(which introduces imprecision)
bool r_col_inner = r_sieved && !sr->is_inner_col(r_cols[i]);
bool neg_col_inner = neg_sieved && !sneg->is_inner_col(neg_cols[i]);
if(r_col_inner && neg_col_inner) {
ir_cols.push_back( r_sieved ? sr->get_inner_col(i) : i );
ineg_cols.push_back( neg_sieved ? sneg->get_inner_col(i) : i );
}
else if(!r_col_inner && neg_col_inner) {
//Sieved (i.e. full) column in r is matched on an inner column in neg.
//If we assume the column in neg is not full, no rows from the inner relation of
//r would be removed. So in this case we perform no operation at cost of a little
//impresicion.
return alloc(identity_relation_intersection_filter_fn);
}
else {
//Inner or sieved column in r must match a sieved column in neg.
//Since sieved columns are full, this is always true so we can skip the equality.
continue;
}
}
relation_intersection_filter_fn * inner_fun =
get_manager().mk_filter_by_negation_fn(inner_r, inner_neg, ir_cols, ineg_cols);
if(!inner_fun) {
return 0;
}
return alloc(negation_filter_fn, inner_fun);
}
};

197
src/muz/dl_sieve_relation.h Normal file
View file

@ -0,0 +1,197 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_sieve_relation.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-11-11.
Revision History:
--*/
#ifndef _DL_SIEVE_RELATION_H_
#define _DL_SIEVE_RELATION_H_
#include "dl_context.h"
namespace datalog {
class sieve_relation;
class sieve_relation_plugin : public relation_plugin {
friend class sieve_relation;
public:
struct rel_spec {
svector<bool> m_inner_cols;
family_id m_inner_kind;
/**
Create uninitialized rel_spec.
*/
rel_spec() {}
/**
\c inner_kind==null_family_id means we will not specify a relation kind when requesting
the relation object from the relation_manager.
\c inner_kind==null_family_id cannot hold in a specification of existing relation object.
*/
rel_spec(unsigned sig_sz, const bool * inner_cols, family_id inner_kind=null_family_id)
: m_inner_cols(sig_sz, inner_cols), m_inner_kind(inner_kind) {}
bool operator==(const rel_spec & o) const {
return m_inner_kind==o.m_inner_kind && vectors_equal(m_inner_cols, o.m_inner_cols);
}
struct hash {
unsigned operator()(const rel_spec & s) const {
return svector_hash<bool_hash>()(s.m_inner_cols)^s.m_inner_kind;
}
};
};
private:
class join_fn;
class transformer_fn;
class union_fn;
class filter_fn;
class negation_filter_fn;
rel_spec_store<rel_spec, rel_spec::hash, default_eq<rel_spec> > m_spec_store;
family_id get_relation_kind(sieve_relation & r, const bool * inner_columns);
void extract_inner_columns(const relation_signature & s, relation_plugin & inner,
svector<bool> & inner_columns);
void extract_inner_signature(const relation_signature & s, relation_signature & inner_sig);
void collect_inner_signature(const relation_signature & s, const svector<bool> & inner_columns,
relation_signature & inner_sig);
public:
static symbol get_name() { return symbol("sieve_relation"); }
static sieve_relation_plugin& get_plugin(relation_manager & rmgr);
static sieve_relation& get(relation_base& r);
static sieve_relation const & get(relation_base const& r);
static sieve_relation* get(relation_base* r);
static sieve_relation const* get(relation_base const* r);
sieve_relation_plugin(relation_manager & manager);
virtual void initialize(family_id fid);
family_id get_relation_kind(const relation_signature & sig, const bool * inner_columns,
family_id inner_kind);
family_id get_relation_kind(const relation_signature & sig, const svector<bool> & inner_columns,
family_id inner_kind) {
SASSERT(sig.size()==inner_columns.size());
return get_relation_kind(sig, inner_columns.c_ptr(), inner_kind);
}
virtual bool can_handle_signature(const relation_signature & s);
virtual relation_base * mk_empty(const relation_signature & s);
sieve_relation * mk_empty(const sieve_relation & original);
virtual relation_base * mk_empty(const relation_base & original);
virtual relation_base * mk_empty(const relation_signature & s, family_id kind);
sieve_relation * mk_empty(const relation_signature & s, relation_plugin & inner_plugin);
virtual relation_base * mk_full(func_decl* p, const relation_signature & s);
sieve_relation * mk_full(func_decl* p, const relation_signature & s, relation_plugin & inner_plugin);
sieve_relation * mk_from_inner(const relation_signature & s, const bool * inner_columns,
relation_base * inner_rel);
sieve_relation * mk_from_inner(const relation_signature & s, const svector<bool> inner_columns,
relation_base * inner_rel) {
SASSERT(inner_columns.size()==s.size());
return mk_from_inner(s, inner_columns.c_ptr(), inner_rel);
}
protected:
virtual relation_join_fn * mk_join_fn(const relation_base & t1, const relation_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
virtual relation_transformer_fn * mk_project_fn(const relation_base & t, unsigned col_cnt,
const unsigned * removed_cols);
virtual relation_transformer_fn * mk_rename_fn(const relation_base & t, unsigned permutation_cycle_len,
const unsigned * permutation_cycle);
virtual relation_union_fn * mk_union_fn(const relation_base & tgt, const relation_base & src,
const relation_base * delta);
virtual relation_mutator_fn * mk_filter_identical_fn(const relation_base & t, unsigned col_cnt,
const unsigned * identical_cols);
virtual relation_mutator_fn * mk_filter_equal_fn(const relation_base & t, const relation_element & value,
unsigned col);
virtual relation_mutator_fn * mk_filter_interpreted_fn(const relation_base & t, app * condition);
virtual relation_intersection_filter_fn * mk_filter_by_negation_fn(const relation_base & t,
const relation_base & negated_obj, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * negated_cols);
};
// -----------------------------------
//
// sieve_relation
//
// -----------------------------------
class sieve_relation : public relation_base {
friend class sieve_relation_plugin;
friend class sieve_relation_plugin::join_fn;
friend class sieve_relation_plugin::transformer_fn;
friend class sieve_relation_plugin::union_fn;
friend class sieve_relation_plugin::filter_fn;
svector<bool> m_inner_cols;
unsigned_vector m_sig2inner;
unsigned_vector m_inner2sig;
unsigned_vector m_ignored_cols; //in ascending order, so that it can be used in project-like functions
scoped_rel<relation_base> m_inner;
sieve_relation(sieve_relation_plugin & p, const relation_signature & s,
const bool * inner_columns, relation_base * inner);
public:
sieve_relation_plugin & get_plugin() const {
return static_cast<sieve_relation_plugin &>(relation_base::get_plugin());
}
bool is_inner_col(unsigned idx) const { return m_sig2inner[idx]!=UINT_MAX; }
unsigned get_inner_col(unsigned idx) const {
SASSERT(is_inner_col(idx));
return m_sig2inner[idx];
}
bool no_sieved_columns() const { return m_ignored_cols.size()==0; }
bool no_inner_columns() const { return m_ignored_cols.size()==get_signature().size(); }
relation_base & get_inner() { return *m_inner; }
const relation_base & get_inner() const { return *m_inner; }
virtual void add_fact(const relation_fact & f);
virtual bool contains_fact(const relation_fact & f) const;
virtual sieve_relation * clone() const;
virtual relation_base * complement(func_decl*p) const;
virtual void to_formula(expr_ref& fml) const;
virtual bool empty() const { return get_inner().empty(); }
virtual void reset() { get_inner().reset(); }
virtual unsigned get_size_estimate_rows() const { return get_inner().get_size_estimate_rows(); }
virtual unsigned get_size_estimate_bytes() const { return get_inner().get_size_estimate_bytes(); }
virtual void display(std::ostream & out) const;
};
};
#endif /* _DL_SIEVE_RELATION_H_ */

1246
src/muz/dl_sparse_table.cpp Normal file

File diff suppressed because it is too large Load diff

480
src/muz/dl_sparse_table.h Normal file
View file

@ -0,0 +1,480 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_table.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-09-01.
Revision History:
--*/
#ifndef _DL_SPARSE_TABLE_H_
#define _DL_SPARSE_TABLE_H_
#include<iostream>
#include<list>
#include<utility>
#include "ast.h"
#include "bit_vector.h"
#include "buffer.h"
#include "hashtable.h"
#include "map.h"
#include "ref_vector.h"
#include "vector.h"
#include "dl_base.h"
namespace datalog {
class sparse_table;
class sparse_table_plugin : public table_plugin {
friend class sparse_table;
protected:
class join_project_fn;
class union_fn;
class transformer_fn;
class rename_fn;
class project_fn;
class negation_filter_fn;
class select_equal_and_project_fn;
typedef ptr_vector<sparse_table> sp_table_vector;
typedef map<table_signature, sp_table_vector *,
table_signature::hash, table_signature::eq > table_pool;
table_pool m_pool;
void recycle(sparse_table * t);
void garbage_collect();
void reset();
static bool join_involves_functional(const table_signature & s1, const table_signature & s2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
public:
typedef sparse_table table;
sparse_table_plugin(relation_manager & manager);
~sparse_table_plugin();
virtual bool can_handle_signature(const table_signature & s)
{ return s.size()>0; }
virtual table_base * mk_empty(const table_signature & s);
sparse_table * mk_clone(const sparse_table & t);
protected:
virtual table_join_fn * mk_join_fn(const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
virtual table_join_fn * mk_join_project_fn(const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2, unsigned removed_col_cnt,
const unsigned * removed_cols);
virtual table_union_fn * mk_union_fn(const table_base & tgt, const table_base & src,
const table_base * delta);
virtual table_transformer_fn * mk_project_fn(const table_base & t, unsigned col_cnt,
const unsigned * removed_cols);
virtual table_transformer_fn * mk_rename_fn(const table_base & t, unsigned permutation_cycle_len,
const unsigned * permutation_cycle);
virtual table_transformer_fn * mk_select_equal_and_project_fn(const table_base & t,
const table_element & value, unsigned col);
virtual table_intersection_filter_fn * mk_filter_by_negation_fn(const table_base & t,
const table_base & negated_obj, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * negated_cols);
};
class entry_storage {
public:
typedef unsigned store_offset;
private:
typedef svector<char> storage;
class offset_hash_proc {
storage & m_storage;
unsigned m_unique_entry_size;
public:
offset_hash_proc(storage & s, unsigned unique_entry_sz)
: m_storage(s), m_unique_entry_size(unique_entry_sz) {}
unsigned operator()(store_offset ofs) const {
return string_hash(m_storage.c_ptr()+ofs, m_unique_entry_size, 0);
}
};
class offset_eq_proc {
storage & m_storage;
unsigned m_unique_entry_size;
public:
offset_eq_proc(storage & s, unsigned unique_entry_sz)
: m_storage(s), m_unique_entry_size(unique_entry_sz) {}
bool operator()(store_offset o1, store_offset o2) const {
const char * base = m_storage.c_ptr();
return memcmp(base+o1, base+o2, m_unique_entry_size)==0;
}
};
typedef hashtable<store_offset, offset_hash_proc, offset_eq_proc> storage_indexer;
static const store_offset NO_RESERVE = UINT_MAX;
unsigned m_entry_size;
unsigned m_unique_part_size;
unsigned m_data_size;
/**
Invariant: Every or all but one blocks of length \c m_entry_size in the \c m_data vector
are unique sequences of bytes and have their offset stored in the \c m_data_indexer hashtable.
If the offset of the last block is not stored in the hashtable, it is stored in the \c m_reserve
variable. Otherwise \c m_reserve==NO_RESERVE.
The size of m_data is actually 8 bytes larger than stated in m_data_size, so that we may
deref an uint64 pointer at the end of the array.
*/
storage m_data;
storage_indexer m_data_indexer;
store_offset m_reserve;
public:
entry_storage(unsigned entry_size, unsigned functional_size = 0, unsigned init_size = 0)
: m_entry_size(entry_size),
m_unique_part_size(entry_size-functional_size),
m_data_indexer(next_power_of_two(std::max(8u,init_size)),
offset_hash_proc(m_data, m_unique_part_size), offset_eq_proc(m_data, m_unique_part_size)),
m_reserve(NO_RESERVE) {
SASSERT(entry_size>0);
SASSERT(functional_size<=entry_size);
resize_data(init_size);
resize_data(0);
}
entry_storage(const entry_storage &s)
: m_entry_size(s.m_entry_size),
m_unique_part_size(s.m_unique_part_size),
m_data_size(s.m_data_size),
m_data(s.m_data),
m_data_indexer(next_power_of_two(std::max(8u,s.entry_count())),
offset_hash_proc(m_data, m_unique_part_size), offset_eq_proc(m_data, m_unique_part_size)),
m_reserve(s.m_reserve) {
store_offset after_last=after_last_offset();
for(store_offset i=0; i<after_last; i+=m_entry_size) {
m_data_indexer.insert(i);
}
}
entry_storage & operator=(const entry_storage & o) {
m_data_indexer.reset();
m_entry_size = o.m_entry_size;
m_unique_part_size = o.m_unique_part_size;
m_data_size = o.m_data_size;
m_data = o.m_data;
m_reserve = o.m_reserve;
store_offset after_last=after_last_offset();
for(store_offset i=0; i<after_last; i+=m_entry_size) {
m_data_indexer.insert(i);
}
return *this;
}
void reset() {
resize_data(0);
m_data_indexer.reset();
m_reserve = NO_RESERVE;
}
unsigned entry_size() const { return m_entry_size; }
unsigned get_size_estimate_bytes() const;
char * get(store_offset ofs) { return m_data.begin()+ofs; }
const char * get(store_offset ofs) const
{ return const_cast<entry_storage *>(this)->get(ofs); }
unsigned entry_count() const { return m_data_indexer.size(); }
store_offset after_last_offset() const {
return (m_reserve==NO_RESERVE) ? m_data_size : m_reserve;
}
char * begin() { return get(0); }
const char * begin() const { return get(0); }
const char * after_last() const { return get(after_last_offset()); }
bool has_reserve() const { return m_reserve!=NO_RESERVE; }
store_offset reserve() const { SASSERT(has_reserve()); return m_reserve; }
void ensure_reserve() {
if(has_reserve()) {
SASSERT(m_reserve==m_data_size-m_entry_size);
return;
}
m_reserve=m_data_size;
resize_data(m_data_size+m_entry_size);
}
/**
\brief Return pointer to the reserve.
The reserve must exist when the function is called.
*/
char * get_reserve_ptr() {
SASSERT(has_reserve());
return &m_data.get(reserve());
}
bool reserve_content_already_present() const {
SASSERT(has_reserve());
return m_data_indexer.contains(reserve());
}
bool find_reserve_content(store_offset & result) const {
SASSERT(has_reserve());
storage_indexer::entry * indexer_entry = m_data_indexer.find_core(reserve());
if(!indexer_entry) {
return false;
}
result = indexer_entry->get_data();
return true;
}
/**
\brief Write fact \c f into the reserve at the end of the \c m_data storage.
If the reserve does not exist, this function creates it.
*/
void write_into_reserve(const char * data) {
ensure_reserve();
memcpy(get_reserve_ptr(), data, m_entry_size);
}
/**
\brief If the fact in reserve is not in the table, insert it there and return true;
otherwise return false.
When a fact is inserted into the table, the reserve becomes part of the table and
is no longer a reserve.
*/
bool insert_reserve_content();
store_offset insert_or_get_reserve_content();
bool remove_reserve_content();
/**
Remove data at the offset \c ofs.
Data with offset lower than \c ofs are not be modified by this function, data with
higher offset may be moved.
*/
void remove_offset(store_offset ofs);
//the following two operations allow breaking of the object invariant!
void resize_data(unsigned sz) {
m_data_size = sz;
m_data.resize(sz + sizeof(uint64));
}
bool insert_offset(store_offset ofs) {
return m_data_indexer.insert_if_not_there(ofs)==ofs;
}
};
class sparse_table : public table_base {
friend class sparse_table_plugin;
friend class sparse_table_plugin::join_project_fn;
friend class sparse_table_plugin::union_fn;
friend class sparse_table_plugin::transformer_fn;
friend class sparse_table_plugin::rename_fn;
friend class sparse_table_plugin::project_fn;
friend class sparse_table_plugin::negation_filter_fn;
friend class sparse_table_plugin::select_equal_and_project_fn;
class our_iterator_core;
class key_indexer;
class general_key_indexer;
class full_signature_key_indexer;
typedef entry_storage::store_offset store_offset;
class column_info {
unsigned m_big_offset;
unsigned m_small_offset;
uint64 m_mask;
uint64 m_write_mask;
public:
unsigned m_offset; //!< in bits
unsigned m_length; //!< in bits
column_info(unsigned offset, unsigned length) \
: m_big_offset(offset/8),
m_small_offset(offset%8),
m_mask( length==64 ? ULLONG_MAX : (static_cast<uint64>(1)<<length)-1 ),
m_write_mask( ~(m_mask<<m_small_offset) ),
m_offset(offset),
m_length(length) {
SASSERT(length<=64);
SASSERT(length+m_small_offset<=64);
}
table_element get(const char * rec) const {
const uint64 * ptr = reinterpret_cast<const uint64*>(rec+m_big_offset);
uint64 res = *ptr;
res>>=m_small_offset;
res&=m_mask;
return res;
}
void set(char * rec, table_element val) const {
SASSERT( (val&~m_mask)==0 ); //the value fits into the column
uint64 * ptr = reinterpret_cast<uint64*>(rec+m_big_offset);
*ptr&=m_write_mask;
*ptr|=val<<m_small_offset;
}
unsigned const next_ofs() const { return m_offset+m_length; }
};
class column_layout : public svector<column_info> {
void make_byte_aligned_end(unsigned col_index);
public:
unsigned m_entry_size;
/**
Number of last bytes which correspond to functional columns in the signature.
*/
unsigned m_functional_part_size;
unsigned m_functional_col_cnt;
column_layout(const table_signature & sig);
table_element get(const char * rec, unsigned col) const {
return (*this)[col].get(rec);
}
void set(char * rec, unsigned col, table_element val) const {
return (*this)[col].set(rec, val);
}
};
typedef svector<unsigned> key_spec; //sequence of columns in a key
typedef svector<table_element> key_value; //values of key columns
typedef map<key_spec, key_indexer*, svector_hash_proc<unsigned_hash>,
vector_eq_proc<key_spec> > key_index_map;
static const store_offset NO_RESERVE = UINT_MAX;
column_layout m_column_layout;
unsigned m_fact_size;
entry_storage m_data;
mutable key_index_map m_key_indexes;
const char * get_at_offset(store_offset i) const {
return m_data.get(i);
}
table_element get_cell(store_offset ofs, unsigned column) const {
return m_column_layout.get(m_data.get(ofs), column);
}
void set_cell(store_offset ofs, unsigned column, table_element val) {
m_column_layout.set(m_data.get(ofs), column, val);
}
void write_into_reserve(const table_element* f);
/**
\brief Return reference to an indexer over columns in \c key_cols.
An indexer can retrieve a sequence of offsets that with \c key_cols columns equal to
the specified key. Indexers are populated lazily -- they remember the position of the
last fact they contain, and when an indexer is retrieved by the \c get_key_indexer function,
all the new facts are added into the indexer.
When a fact is removed from the table, all indexers are destroyed. This is not an extra
expense in the current use scenario, because we first perform all fact removals and do the
joins only after that (joins are the only operations that lead to index construction).
*/
key_indexer& get_key_indexer(unsigned key_len, const unsigned * key_cols) const;
void reset_indexes();
static void copy_columns(const column_layout & src_layout, const column_layout & dest_layout,
unsigned start_index, unsigned after_last, const char * src, char * dest,
unsigned & dest_idx, unsigned & pre_projection_idx, const unsigned * & next_removed);
/**
\c array \c removed_cols contains column indexes to be removed in ascending order and
is terminated by a number greated than the highest column index of a join the the two tables.
This is to simplify the traversal of the array when building facts.
*/
static void concatenate_rows(const column_layout & layout1, const column_layout & layout2,
const column_layout & layout_res, const char * ptr1, const char * ptr2, char * res,
const unsigned * removed_cols);
/**
\brief Perform join-project between t1 and t2 iterating through t1 and retrieving relevant
columns from t2 using indexing.
\c array \c removed_cols contains column indexes to be removed in ascending order and
is terminated by a number greated than the highest column index of a join the the two tables.
This is to simplify the traversal of the array when building facts.
\c tables_swapped value means that the resulting facts should contain facts from t2 first,
instead of the default behavior that would concatenate the two facts as \c (t1,t2).
\remark The function is called \c self_agnostic_join since, unlike the virtual method
\c join, it is static and therefore allows to easily swap the roles of the two joined
tables (the indexed and iterated one) in a way that is expected to give better performance.
*/
static void self_agnostic_join_project(const sparse_table & t1, const sparse_table & t2,
unsigned joined_col_cnt, const unsigned * t1_joined_cols, const unsigned * t2_joined_cols,
const unsigned * removed_cols, bool tables_swapped, sparse_table & result);
/**
If the fact at \c data (in table's native representation) is not in the table,
add it and return true. Otherwise return false.
*/
bool add_fact(const char * data);
bool add_reserve_content();
void garbage_collect();
sparse_table(sparse_table_plugin & p, const table_signature & sig, unsigned init_capacity=0);
sparse_table(const sparse_table & t);
virtual ~sparse_table();
public:
virtual void deallocate() {
get_plugin().recycle(this);
}
unsigned row_count() const { return m_data.entry_count(); }
sparse_table_plugin & get_plugin() const
{ return static_cast<sparse_table_plugin &>(table_base::get_plugin()); }
virtual bool empty() const { return row_count()==0; }
virtual void add_fact(const table_fact & f);
virtual bool contains_fact(const table_fact & f) const;
virtual bool fetch_fact(table_fact & f) const;
virtual void ensure_fact(const table_fact & f);
virtual void remove_fact(const table_element* fact);
virtual void reset();
virtual table_base * clone() const;
virtual table_base::iterator begin() const;
virtual table_base::iterator end() const;
virtual unsigned get_size_estimate_rows() const { return row_count(); }
virtual unsigned get_size_estimate_bytes() const;
virtual bool knows_exact_size() const { return true; }
};
};
#endif /* _DL_SPARSE_TABLE_H_ */

772
src/muz/dl_table.cpp Normal file
View file

@ -0,0 +1,772 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_table.cpp
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-09-01.
Revision History:
--*/
#include"dl_context.h"
#include"dl_util.h"
#include"dl_table.h"
namespace datalog {
// -----------------------------------
//
// hashtable_table
//
// -----------------------------------
table_base * hashtable_table_plugin::mk_empty(const table_signature & s) {
SASSERT(can_handle_signature(s));
return alloc(hashtable_table, *this, s);
}
class hashtable_table_plugin::join_fn : public convenient_table_join_fn {
unsigned m_joined_col_cnt;
public:
join_fn(const table_signature & t1_sig, const table_signature & t2_sig, unsigned col_cnt, const unsigned * cols1, const unsigned * cols2)
: convenient_table_join_fn(t1_sig, t2_sig, col_cnt, cols1, cols2),
m_joined_col_cnt(col_cnt) {}
virtual table_base * operator()(const table_base & t1, const table_base & t2) {
const hashtable_table & ht1 = static_cast<const hashtable_table &>(t1);
const hashtable_table & ht2 = static_cast<const hashtable_table &>(t2);
hashtable_table_plugin & plugin = ht1.get_plugin();
hashtable_table * res = static_cast<hashtable_table *>(plugin.mk_empty(get_result_signature()));
hashtable_table::storage::iterator els1it = ht1.m_data.begin();
hashtable_table::storage::iterator els1end = ht1.m_data.end();
hashtable_table::storage::iterator els2end = ht2.m_data.end();
table_fact acc;
for(; els1it!=els1end; ++els1it) {
const table_fact & row1 = *els1it;
hashtable_table::storage::iterator els2it = ht2.m_data.begin();
for(; els2it!=els2end; ++els2it) {
const table_fact & row2 = *els2it;
bool match=true;
for(unsigned i=0; i<m_joined_col_cnt; i++) {
if(row1[m_cols1[i]]!=row2[m_cols2[i]]) {
match=false;
break;
}
}
if(!match) {
continue;
}
acc.reset();
acc.append(row1);
acc.append(row2);
res->m_data.insert(acc);
}
}
return res;
}
};
table_join_fn * hashtable_table_plugin::mk_join_fn(const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2) {
if(t1.get_kind()!=get_kind() || t2.get_kind()!=get_kind()) {
return 0;
}
return alloc(join_fn, t1.get_signature(), t2.get_signature(), col_cnt, cols1, cols2);
}
class hashtable_table::our_iterator_core : public iterator_core {
const hashtable_table & m_parent;
storage::iterator m_inner;
storage::iterator m_end;
class our_row : public row_interface {
const our_iterator_core & m_parent;
public:
our_row(const our_iterator_core & parent) : row_interface(parent.m_parent), m_parent(parent) {}
virtual void get_fact(table_fact & result) const {
result = *m_parent.m_inner;
}
virtual table_element operator[](unsigned col) const {
return (*m_parent.m_inner)[col];
}
};
our_row m_row_obj;
public:
our_iterator_core(const hashtable_table & t, bool finished) :
m_parent(t), m_inner(finished ? t.m_data.end() : t.m_data.begin()),
m_end(t.m_data.end()), m_row_obj(*this) {}
virtual bool is_finished() const {
return m_inner==m_end;
}
virtual row_interface & operator*() {
SASSERT(!is_finished());
return m_row_obj;
}
virtual void operator++() {
SASSERT(!is_finished());
++m_inner;
}
};
table_base::iterator hashtable_table::begin() const {
return mk_iterator(alloc(our_iterator_core, *this, false));
}
table_base::iterator hashtable_table::end() const {
return mk_iterator(alloc(our_iterator_core, *this, true));
}
// -----------------------------------
//
// bitvector_table
//
// -----------------------------------
bool bitvector_table_plugin::can_handle_signature(const table_signature & sig) {
if(sig.functional_columns()!=0) {
return false;
}
unsigned cols = sig.size();
unsigned shift = 0;
for (unsigned i = 0; i < cols; ++i) {
unsigned s = static_cast<unsigned>(sig[i]);
if (s != sig[i] || !is_power_of_two(s)) {
return false;
}
unsigned num_bits = 0;
unsigned bit_pos = 1;
for (num_bits = 1; num_bits < 32; ++num_bits) {
if (bit_pos & s) {
break;
}
bit_pos <<= 1;
}
shift += num_bits;
if (shift >= 32) {
return false;
}
}
return true;
}
table_base * bitvector_table_plugin::mk_empty(const table_signature & s) {
SASSERT(can_handle_signature(s));
return alloc(bitvector_table, *this, s);
}
class bitvector_table::bv_iterator : public iterator_core {
bitvector_table const& m_bv;
unsigned m_offset;
class our_row : public caching_row_interface {
const bv_iterator& m_parent;
public:
our_row(const bv_iterator & p) : caching_row_interface(p.m_bv), m_parent(p) {}
virtual void get_fact(table_fact& result) const {
if (result.size() < size()) {
result.resize(size(), 0);
}
m_parent.m_bv.offset2fact(m_parent.m_offset, result);
}
};
our_row m_row_obj;
public:
bv_iterator(const bitvector_table& bv, bool end):
m_bv(bv), m_offset(end?m_bv.m_bv.size():0), m_row_obj(*this)
{
if (!is_finished() && !m_bv.m_bv.get(m_offset)) {
++(*this);
}
}
virtual bool is_finished() const {
return m_offset == m_bv.m_bv.size();
}
virtual row_interface & operator*() {
SASSERT(!is_finished());
return m_row_obj;
}
virtual void operator++() {
SASSERT(!is_finished());
++m_offset;
while (!is_finished() && !m_bv.m_bv.get(m_offset)) {
++m_offset;
}
m_row_obj.reset();
}
};
bitvector_table::bitvector_table(bitvector_table_plugin & plugin, const table_signature & sig)
: table_base(plugin, sig) {
SASSERT(plugin.can_handle_signature(sig));
m_num_cols = sig.size();
unsigned shift = 0;
for (unsigned i = 0; i < m_num_cols; ++i) {
unsigned s = static_cast<unsigned>(sig[i]);
if (s != sig[i] || !is_power_of_two(s)) {
throw default_exception("bit-vector table is specialized to small domains that are powers of two");
}
m_shift.push_back(shift);
m_mask.push_back(s - 1);
unsigned num_bits = 0;
unsigned bit_pos = 1;
for (num_bits = 1; num_bits < 32; ++num_bits) {
if (bit_pos & s) {
break;
}
bit_pos <<= 1;
}
shift += num_bits;
if (shift >= 32) {
throw default_exception("bit-vector table is specialized to small domains that are powers of two");
}
m_bv.reserve(1 << shift);
}
}
unsigned bitvector_table::fact2offset(const table_element* f) const {
unsigned result = 0;
for (unsigned i = 0; i < m_num_cols; ++i) {
SASSERT(f[i]<get_signature()[i]);
result += ((unsigned)f[i]) << m_shift[i];
}
return result;
}
void bitvector_table::offset2fact(unsigned offset, table_fact& f) const {
SASSERT(m_num_cols == f.size());
for (unsigned i = 0; i < m_num_cols; ++i) {
f[i] = m_mask[i] & (offset >> m_shift[i]);
}
}
void bitvector_table::add_fact(const table_fact & f) {
m_bv.set(fact2offset(f.c_ptr()));
}
void bitvector_table::remove_fact(const table_element* fact) {
m_bv.unset(fact2offset(fact));
}
bool bitvector_table::contains_fact(const table_fact & f) const {
return m_bv.get(fact2offset(f.c_ptr()));
}
table_base::iterator bitvector_table::begin() const {
return mk_iterator(alloc(bv_iterator, *this, false));
}
table_base::iterator bitvector_table::end() const {
return mk_iterator(alloc(bv_iterator, *this, true));
}
// -----------------------------------
//
// equivalence_table
//
// -----------------------------------
bool equivalence_table_plugin::can_handle_signature(const table_signature & sig) {
return sig.functional_columns() == 0 && sig.size() == 2 && sig[0] < UINT_MAX && sig[0] == sig[1];
}
bool equivalence_table_plugin::is_equivalence_table(table_base const& tbl) const {
if (tbl.get_kind() != get_kind()) return false;
equivalence_table const& t = static_cast<equivalence_table const&>(tbl);
return !t.is_sparse();
}
table_base * equivalence_table_plugin::mk_empty(const table_signature & s) {
TRACE("dl", for (unsigned i = 0; i < s.size(); ++i) tout << s[i] << " "; tout << "\n";);
SASSERT(can_handle_signature(s));
return alloc(equivalence_table, *this, s);
}
class equivalence_table_plugin::select_equal_and_project_fn : public table_transformer_fn {
unsigned m_val;
table_sort m_sort;
public:
select_equal_and_project_fn(const table_signature & sig, table_element val, unsigned col)
: m_val(static_cast<unsigned>(val)),
m_sort(sig[0]) {
SASSERT(val <= UINT_MAX);
SASSERT(col == 0 || col == 1);
SASSERT(sig.functional_columns() == 0);
SASSERT(sig.size() == 2);
SASSERT(sig[0] < UINT_MAX && sig[0] == sig[1]);
}
virtual table_base* operator()(const table_base& tb) {
TRACE("dl", tout << "\n";);
table_plugin & plugin = tb.get_plugin();
table_plugin* rp = plugin.get_manager().get_table_plugin(symbol("sparse"));
SASSERT(rp);
table_signature sig;
sig.push_back(m_sort);
table_base* result = rp->mk_empty(sig);
equivalence_table const& eq_table = static_cast<equivalence_table const&>(tb);
if (eq_table.is_valid(m_val)) {
table_fact fact;
fact.resize(1);
unsigned r = m_val;
do {
fact[0] = r;
result->add_fact(fact);
r = eq_table.m_uf.next(r);
}
while (r != m_val);
}
TRACE("dl", tb.display(tout << "src:\n"); result->display(tout << "result\n"););
return result;
}
};
table_transformer_fn * equivalence_table_plugin::mk_select_equal_and_project_fn(
const table_base & t, const table_element & value, unsigned col) {
return alloc(select_equal_and_project_fn, t.get_signature(), value, col);
}
class equivalence_table_plugin::union_fn : public table_union_fn {
equivalence_table_plugin& m_plugin;
void mk_union1(equivalence_table & tgt, const equivalence_table & src, table_base * delta) {
unsigned num_vars = src.m_uf.get_num_vars();
table_fact fact;
fact.resize(2);
for (unsigned i = 0; i < num_vars; ++i) {
if (src.is_valid(i) && src.m_uf.find(i) == i) {
fact[0] = i;
equivalence_table::class_iterator it = src.class_begin(i);
equivalence_table::class_iterator end = src.class_end(i);
for (; it != end; ++it) {
fact[1] = *it;
if (!tgt.contains_fact(fact)) {
tgt.add_fact(fact);
if (delta) {
delta->add_fact(fact);
}
}
}
}
}
}
void mk_union2(equivalence_table & tgt, const table_base & src, table_base * delta) {
table_fact fact;
table_base::iterator it = src.begin(), end = src.end();
for (; it != end; ++it) {
it->get_fact(fact);
if (!tgt.contains_fact(fact)) {
tgt.add_fact(fact);
if (delta) {
delta->add_fact(fact);
TRACE("dl",
tout << "Add: ";
for (unsigned i = 0; i < fact.size(); ++i) tout << fact[i] << " ";
tout << "\n";);
}
}
}
}
public:
union_fn(equivalence_table_plugin& p) : m_plugin(p) {}
virtual void operator()(table_base & tgt0, const table_base & src, table_base * delta) {
TRACE("dl", tout << "union\n";);
equivalence_table & tgt = static_cast<equivalence_table &>(tgt0);
if (m_plugin.is_equivalence_table(src)) {
mk_union1(tgt, static_cast<equivalence_table const&>(src), delta);
}
else {
mk_union2(tgt, src, delta);
}
TRACE("dl", src.display(tout << "src\n"); tgt.display(tout << "tgt\n");
if (delta) delta->display(tout << "delta\n"););
}
};
table_union_fn * equivalence_table_plugin::mk_union_fn(
const table_base & tgt, const table_base & src, const table_base * delta) {
if (!is_equivalence_table(tgt) ||
tgt.get_signature() != src.get_signature() ||
(delta && delta->get_signature() != tgt.get_signature())) {
return 0;
}
return alloc(union_fn,*this);
}
class equivalence_table_plugin::join_project_fn : public convenient_table_join_project_fn {
equivalence_table_plugin& m_plugin;
public:
join_project_fn(
equivalence_table_plugin& plugin, const table_signature & t1_sig, const table_signature & t2_sig, unsigned col_cnt,
const unsigned * cols1, const unsigned * cols2, unsigned removed_col_cnt,
const unsigned * removed_cols)
: convenient_table_join_project_fn(t1_sig, t2_sig, col_cnt, cols1, cols2, removed_col_cnt, removed_cols),
m_plugin(plugin) {
m_removed_cols.push_back(UINT_MAX);
}
virtual table_base * operator()(const table_base & tb1, const table_base & tb2) {
SASSERT(m_cols1.size() == 1);
const table_signature & res_sign = get_result_signature();
table_plugin * plugin = &tb1.get_plugin();
if (!plugin->can_handle_signature(res_sign)) {
plugin = &tb2.get_plugin();
if (!plugin->can_handle_signature(res_sign)) {
plugin = &tb1.get_manager().get_appropriate_plugin(res_sign);
}
}
SASSERT(plugin->can_handle_signature(res_sign));
table_base * result = plugin->mk_empty(res_sign);
if (m_plugin.is_equivalence_table(tb1)) {
mk_join(0, m_cols1[0], static_cast<const equivalence_table&>(tb1),
2, m_cols2[0], tb2, result);
}
else if (m_plugin.is_equivalence_table(tb2)) {
mk_join(tb1.get_signature().size(), m_cols2[0], static_cast<const equivalence_table&>(tb2),
0, m_cols1[0], tb1, result);
}
else {
UNREACHABLE();
}
TRACE("dl", tb1.display(tout << "tb1\n"); tb2.display(tout << "tb2\n"); result->display(tout << "result\n"););
return result;
}
private:
table_base * mk_join(unsigned offs1, unsigned col1, equivalence_table const & t1,
unsigned offs2, unsigned col2, table_base const& t2, table_base* res) {
table_base::iterator els2it = t2.begin();
table_base::iterator els2end = t2.end();
table_fact acc, proj;
acc.resize(t1.get_signature().size() + t2.get_signature().size());
for(; els2it != els2end; ++els2it) {
const table_base::row_interface & row2 = *els2it;
table_element const& e2 = row2[col2];
equivalence_table::class_iterator it = t1.class_begin(e2);
equivalence_table::class_iterator end = t1.class_end(e2);
if (it != end) {
for (unsigned i = 0; i < row2.size(); ++i) {
acc[i+offs2] = row2[i];
}
}
for (; it != end; ++it) {
acc[offs1+col1] = e2;
acc[offs1+1-col1] = *it;
mk_project(acc, proj);
TRACE("dl", for (unsigned i = 0; i < proj.size(); ++i) tout << proj[i] << " "; tout << "\n";);
res->add_fact(proj);
}
}
return res;
}
virtual void mk_project(table_fact const & f, table_fact & p) const {
unsigned sz = f.size();
p.reset();
for (unsigned i = 0, r = 0; i < sz; ++i) {
if (r < m_removed_cols.size() && m_removed_cols[r] == i) {
++r;
}
else {
p.push_back(f[i]);
}
}
}
};
table_join_fn * equivalence_table_plugin::mk_join_project_fn(
const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2, unsigned removed_col_cnt,
const unsigned * removed_cols) {
if (col_cnt != 1) {
TRACE("dl", tout << "WARNING: join_project on multiple columns is not implemented\n";);
return 0;
}
if (is_equivalence_table(t1) || is_equivalence_table(t2)) {
return alloc(join_project_fn, *this, t1.get_signature(), t2.get_signature(), col_cnt, cols1, cols2,
removed_col_cnt, removed_cols);
}
return 0;
}
class equivalence_table::eq_iterator : public iterator_core {
equivalence_table const& m_eq;
unsigned m_last;
unsigned m_current;
unsigned m_next;
class our_row : public caching_row_interface {
const eq_iterator& m_parent;
public:
our_row(const eq_iterator & p) : caching_row_interface(p.m_eq), m_parent(p) {}
virtual void get_fact(table_fact& result) const {
if (result.size() < size()) {
result.resize(size(), 0);
}
result[0] = m_parent.m_current;
result[1] = m_parent.m_next;
}
virtual table_element operator[](unsigned col) const {
if (col == 0) return m_parent.m_current;
if (col == 1) return m_parent.m_next;
UNREACHABLE();
return 0;
}
};
our_row m_row_obj;
public:
eq_iterator(const equivalence_table& eq, bool end):
m_eq(eq),
m_last(eq.m_uf.get_num_vars()),
m_current(end?m_last:0),
m_next(0),
m_row_obj(*this)
{
while (m_current < m_last && !m_eq.is_valid(m_current)) {
m_current++;
m_next = m_current;
}
}
virtual bool is_finished() const {
return m_current == m_last;
}
virtual row_interface & operator*() {
SASSERT(!is_finished());
return m_row_obj;
}
virtual void operator++() {
SASSERT(!is_finished());
m_next = m_eq.m_uf.next(m_next);
if (m_next == m_current) {
do {
m_current++;
m_next = m_current;
}
while (m_current < m_last && !m_eq.is_valid(m_current));
}
}
};
equivalence_table::equivalence_table(equivalence_table_plugin & plugin, const table_signature & sig)
: table_base(plugin, sig), m_uf(m_ctx), m_sparse(0) {
SASSERT(plugin.can_handle_signature(sig));
}
equivalence_table::~equivalence_table() {
if (is_sparse()) {
m_sparse->deallocate();
}
}
void equivalence_table::add_fact(const table_fact & f) {
if (is_sparse()) {
add_fact_sparse(f);
}
else {
TRACE("dl_verbose", for (unsigned i = 0; i < f.size(); ++i) tout << f[i] << " "; tout << "\n";);
while (first(f) >= m_uf.get_num_vars()) m_uf.mk_var();
while (second(f) >= m_uf.get_num_vars()) m_uf.mk_var();
m_uf.merge(first(f), second(f));
m_valid.reserve(m_uf.get_num_vars());
m_valid.set(first(f));
m_valid.set(second(f));
}
}
void equivalence_table::remove_fact(const table_element* fact) {
mk_sparse();
m_sparse->remove_fact(fact);
}
void equivalence_table::mk_sparse() {
if (m_sparse) return;
TRACE("dl",tout << "\n";);
table_plugin & plugin = get_plugin();
table_plugin* rp = plugin.get_manager().get_table_plugin(symbol("sparse"));
SASSERT(rp);
table_base* result = rp->mk_empty(get_signature());
table_base::iterator it = begin(), e = end();
table_fact fact;
for (; it != e; ++it) {
it->get_fact(fact);
result->add_fact(fact);
}
m_sparse = result;
}
void equivalence_table::add_fact_sparse(table_fact const& f) {
table_base::iterator it = m_sparse->begin(), end = m_sparse->end();
vector<table_fact> to_add;
to_add.push_back(f);
table_fact f1(f);
f1[0] = f[1];
f1[1] = f[0];
to_add.push_back(f1);
f1[0] = f[1];
f1[1] = f[1];
to_add.push_back(f1);
f1[0] = f[0];
f1[1] = f[0];
to_add.push_back(f1);
for (; it != end; ++it) {
if ((*it)[0] == f[0]) {
f1[0] = f[1];
f1[1] = (*it)[1];
to_add.push_back(f1);
std::swap(f1[0],f1[1]);
to_add.push_back(f1);
}
}
for (unsigned i = 0; i < to_add.size(); ++i) {
m_sparse->add_fact(to_add[i]);
}
}
bool equivalence_table::contains_fact(const table_fact & f) const {
TRACE("dl_verbose", for (unsigned i = 0; i < f.size(); ++i) tout << f[i] << " "; tout << "\n";);
if (is_sparse()) {
return m_sparse->contains_fact(f);
}
return
is_valid(first(f)) &&
is_valid(second(f)) &&
m_uf.find(first(f)) == m_uf.find(second(f));
}
table_base* equivalence_table::clone() const {
if (is_sparse()) {
return m_sparse->clone();
}
TRACE("dl",tout << "\n";);
table_plugin & plugin = get_plugin();
table_base* result = plugin.mk_empty(get_signature());
table_fact fact;
fact.resize(2);
for (unsigned i = 0; i < m_uf.get_num_vars(); ++i) {
if (m_valid.get(i) && m_uf.find(i) == i) {
unsigned n = m_uf.next(i);
fact[0] = i;
while (n != i) {
fact[1] = n;
result->add_fact(fact);
n = m_uf.next(n);
}
}
}
return result;
}
table_base::iterator equivalence_table::begin() const {
if (is_sparse()) return m_sparse->begin();
return mk_iterator(alloc(eq_iterator, *this, false));
}
table_base::iterator equivalence_table::end() const {
if (is_sparse()) return m_sparse->end();
return mk_iterator(alloc(eq_iterator, *this, true));
}
equivalence_table::class_iterator equivalence_table::class_begin(table_element const& _e) const {
SASSERT(!is_sparse());
unsigned e = static_cast<unsigned>(_e);
return class_iterator(*this, e, !is_valid(e));
}
equivalence_table::class_iterator equivalence_table::class_end(table_element const& _e) const {
SASSERT(!is_sparse());
unsigned e = static_cast<unsigned>(_e);
return class_iterator(*this, e, true);
}
void equivalence_table::display(std::ostream& out) const {
if (is_sparse()) {
m_sparse->display(out);
return;
}
for (unsigned i = 0; i < m_uf.get_num_vars(); ++i) {
if (is_valid(i) && m_uf.find(i) == i) {
unsigned j = i, last = i;
do {
out << "<" << i << " " << j << ">\n";
j = m_uf.next(j);
}
while (last != j);
}
}
}
unsigned equivalence_table::get_size_estimate_rows() const {
if (is_sparse()) return m_sparse->get_size_estimate_rows();
return static_cast<unsigned>(get_signature()[0]);
}
unsigned equivalence_table::get_size_estimate_bytes() const {
if (is_sparse()) return m_sparse->get_size_estimate_bytes();
return static_cast<unsigned>(get_signature()[0]);
}
bool equivalence_table::knows_exact_size() const {
return (!is_sparse() || m_sparse->knows_exact_size());
}
};

265
src/muz/dl_table.h Normal file
View file

@ -0,0 +1,265 @@
/*++
Copyright (c) 2006 Microsoft Corporation
Module Name:
dl_table.h
Abstract:
<abstract>
Author:
Krystof Hoder (t-khoder) 2010-09-01.
Revision History:
--*/
#ifndef _DL_TABLE_H_
#define _DL_TABLE_H_
#include<iostream>
#include<list>
#include<utility>
#include "ast.h"
#include "bit_vector.h"
#include "buffer.h"
#include "hashtable.h"
#include "map.h"
#include "ref_vector.h"
#include "vector.h"
#include "union_find.h"
#include "dl_base.h"
#include "dl_util.h"
#include "bit_vector.h"
namespace datalog {
class context;
class variable_intersection;
// -----------------------------------
//
// hashtable_table
//
// -----------------------------------
class hashtable_table;
class hashtable_table_plugin : public table_plugin {
friend class hashtable_table;
protected:
class join_fn;
public:
typedef hashtable_table table;
hashtable_table_plugin(relation_manager & manager)
: table_plugin(symbol("hashtable"), manager) {}
virtual table_base * mk_empty(const table_signature & s);
virtual table_join_fn * mk_join_fn(const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
};
class hashtable_table : public table_base {
friend class hashtable_table_plugin;
friend class hashtable_table_plugin::join_fn;
class our_iterator_core;
typedef hashtable<table_fact, svector_hash_proc<table_element_hash>,
vector_eq_proc<table_fact> > storage;
storage m_data;
hashtable_table(hashtable_table_plugin & plugin, const table_signature & sig)
: table_base(plugin, sig) {}
public:
hashtable_table_plugin & get_plugin() const
{ return static_cast<hashtable_table_plugin &>(table_base::get_plugin()); }
virtual void add_fact(const table_fact & f) {
m_data.insert(f);
}
virtual void remove_fact(const table_element* fact) {
table_fact f(get_signature().size(), fact);
m_data.remove(f);
}
virtual bool contains_fact(const table_fact & f) const {
return m_data.contains(f);
}
virtual iterator begin() const;
virtual iterator end() const;
virtual unsigned get_size_estimate_rows() const { return m_data.size(); }
virtual unsigned get_size_estimate_bytes() const { return m_data.size()*get_signature().size()*8; }
virtual bool knows_exact_size() const { return true; }
};
// -----------------------------------
//
// bitvector_table
//
// -----------------------------------
class bitvector_table;
class bitvector_table_plugin : public table_plugin {
public:
typedef bitvector_table table;
bitvector_table_plugin(relation_manager & manager)
: table_plugin(symbol("bitvector"), manager) {}
virtual bool can_handle_signature(const table_signature & s);
virtual table_base * mk_empty(const table_signature & s);
};
class bitvector_table : public table_base {
friend class bitvector_table_plugin;
class bv_iterator;
bit_vector m_bv;
unsigned m_num_cols;
unsigned_vector m_shift;
unsigned_vector m_mask;
unsigned fact2offset(const table_element* f) const;
void offset2fact(unsigned offset, table_fact& f) const;
bitvector_table(bitvector_table_plugin & plugin, const table_signature & sig);
public:
virtual void add_fact(const table_fact & f);
virtual void remove_fact(const table_element* fact);
virtual bool contains_fact(const table_fact & f) const;
virtual iterator begin() const;
virtual iterator end() const;
};
// -------------------------------------------
// Equivalence table.
// Really: partial equivalence relation table.
// -------------------------------------------
class equivalence_table;
class equivalence_table_plugin : public table_plugin {
class union_fn;
class select_equal_and_project_fn;
class join_project_fn;
bool is_equivalence_table(table_base const& tbl) const;
public:
typedef equivalence_table table;
equivalence_table_plugin(relation_manager & manager)
: table_plugin(symbol("equivalence"), manager) {}
virtual bool can_handle_signature(const table_signature & s);
virtual table_base * mk_empty(const table_signature & s);
protected:
virtual table_union_fn * mk_union_fn(const table_base & tgt, const table_base & src,
const table_base * delta);
virtual table_transformer_fn * mk_select_equal_and_project_fn(
const table_base & t,
const table_element & value, unsigned col);
virtual table_join_fn * mk_join_project_fn(const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2, unsigned removed_col_cnt,
const unsigned * removed_cols);
#if 0
virtual table_join_fn * mk_join_fn(const table_base & t1, const table_base & t2,
unsigned col_cnt, const unsigned * cols1, const unsigned * cols2);
virtual table_transformer_fn * mk_project_fn(const table_base & t, unsigned col_cnt,
const unsigned * removed_cols);
virtual table_transformer_fn * mk_rename_fn(const table_base & t, unsigned permutation_cycle_len,
const unsigned * permutation_cycle);
const table_element & value, unsigned col);
virtual table_intersection_filter_fn * mk_filter_by_negation_fn(const table_base & t,
const table_base & negated_obj, unsigned joined_col_cnt,
const unsigned * t_cols, const unsigned * negated_cols);
#endif
};
class equivalence_table : public table_base {
friend class equivalence_table_plugin;
class eq_iterator;
union_find_default_ctx m_ctx;
bit_vector m_valid;
union_find<> m_uf;
table_base* m_sparse;
equivalence_table(equivalence_table_plugin & plugin, const table_signature & sig);
virtual ~equivalence_table();
unsigned first(table_fact const& f) const { return static_cast<unsigned>(f[0]); }
unsigned second(table_fact const& f) const { return static_cast<unsigned>(f[1]); }
bool is_valid(unsigned entry) const { return entry < m_valid.size() && m_valid.get(entry); }
bool is_sparse() const { return m_sparse != 0; }
// iterator over equivalence class of 'n'.
class class_iterator {
equivalence_table const& m_parent;
unsigned m_current;
unsigned m_last;
bool m_end;
public:
class_iterator(equivalence_table const& s, unsigned n, bool end):
m_parent(s), m_current(n), m_last(n), m_end(end) {}
unsigned operator*() { return m_current; }
class_iterator& operator++() {
m_current = m_parent.m_uf.next(m_current);
m_end = (m_current == m_last);
return *this;
}
bool operator==(const class_iterator & it) const {
return
(m_end && it.m_end) ||
(!m_end && !it.m_end && m_current == it.m_current);
}
bool operator!=(const class_iterator & it) const { return !operator==(it); }
};
class_iterator class_begin(table_element const& e) const;
class_iterator class_end(table_element const& e) const;
void add_fact_sparse(table_fact const& f);
void mk_sparse();
public:
virtual void add_fact(const table_fact & f);
virtual void remove_fact(const table_element* fact);
virtual bool contains_fact(const table_fact & f) const;
virtual table_base* clone() const;
virtual iterator begin() const;
virtual iterator end() const;
virtual unsigned get_size_estimate_rows() const;
virtual unsigned get_size_estimate_bytes() const;
virtual bool knows_exact_size() const;
virtual void display(std::ostream & out) const;
};
};
#endif /* _DL_TABLE_H_ */

Some files were not shown because too many files have changed in this diff Show more