3
0
Fork 0
mirror of https://github.com/Z3Prover/z3 synced 2025-04-30 04:15:51 +00:00

overhaul of proof format for new solver

This commit overhauls the proof format (in development) for the new core.

NOTE: this functionality is work in progress with a long way to go.
It is shielded by the sat.euf option, which is off by default and in pre-release state.
It is too early to fuzz or use it. It is pushed into master to shed light on road-map for certifying inferences of sat.euf.

It retires the ad-hoc extension of DRUP used by the SAT solver.
Instead it relies on SMT with ad-hoc extensions for proof terms.
It adds the following commands (consumed by proof_cmds.cpp):

- assume  - for input clauses
- learn   - when a clause is learned (or redundant clause is added)
- del     - when a clause is deleted.

The commands take a list of expressions of type Bool and the
last argument can optionally be of type Proof.
When the last argument is of type Proof it is provided as a hint
to justify the learned clause.

Proof hints can be checked using a self-contained proof
checker. The sat/smt/euf_proof_checker.h class provides
a plugin dispatcher for checkers.
It is instantiated with a checker for arithmetic lemmas,
so far for Farkas proofs.

Use example:
```
(set-option :sat.euf true)
(set-option :tactic.default_tactic smt)
(set-option :sat.smt.proof f.proof)
(declare-const x Int)
(declare-const y Int)
(declare-const z Int)
(declare-const u Int)
(assert (< x y))
(assert (< y z))
(assert (< z x))
(check-sat)
```

Run z3 on a file with above content.
Then run z3 on f.proof

```
(verified-smt)
(verified-smt)
(verified-smt)
(verified-farkas)
(verified-smt)
```
This commit is contained in:
Nikolaj Bjorner 2022-08-28 17:44:33 -07:00
parent 9922c766b9
commit e2f4fc2307
37 changed files with 809 additions and 1078 deletions

View file

@ -21,6 +21,7 @@ z3_add_component(sat_smt
euf_invariant.cpp
euf_model.cpp
euf_proof.cpp
euf_proof_checker.cpp
euf_relevancy.cpp
euf_solver.cpp
fpa_solver.cpp

View file

@ -264,11 +264,12 @@ namespace arith {
SASSERT(k1 != k2 || kind1 != kind2);
auto bin_clause = [&](sat::literal l1, sat::literal l2) {
sat::proof_hint* bound_params = nullptr;
arith_proof_hint* bound_params = nullptr;
if (ctx.use_drat()) {
bound_params = &m_farkas2;
m_farkas2.m_literals[0] = std::make_pair(rational(1), ~l1);
m_farkas2.m_literals[1] = std::make_pair(rational(1), ~l2);
m_arith_hint.set_type(ctx, hint_type::farkas_h);
m_arith_hint.add_lit(rational(1), ~l1);
m_arith_hint.add_lit(rational(1), ~l2);
bound_params = m_arith_hint.mk(ctx);
}
add_clause(l1, l2, bound_params);
};

View file

@ -15,6 +15,8 @@ Author:
--*/
#include "ast/ast_util.h"
#include "ast/scoped_proof.h"
#include "sat/smt/euf_solver.h"
#include "sat/smt/arith_solver.h"
@ -81,7 +83,6 @@ namespace arith {
}
void solver::explain_assumptions() {
m_arith_hint.reset();
unsigned i = 0;
for (auto const & ev : m_explanation) {
++i;
@ -91,14 +92,12 @@ namespace arith {
switch (m_constraint_sources[idx]) {
case inequality_source: {
literal lit = m_inequalities[idx];
m_arith_hint.m_literals.push_back({ev.coeff(), lit});
m_arith_hint.add_lit(ev.coeff(), lit);
break;
}
case equality_source: {
auto [u, v] = m_equalities[idx];
ctx.drat_log_expr(u->get_expr());
ctx.drat_log_expr(v->get_expr());
m_arith_hint.m_eqs.push_back({u->get_expr_id(), v->get_expr_id()});
m_arith_hint.add_eq(u, v);
break;
}
default:
@ -115,22 +114,65 @@ namespace arith {
* such that there is a r >= 1
* (r1*a1+..+r_k*a_k) = r*a, (r1*b1+..+r_k*b_k) <= r*b
*/
sat::proof_hint const* solver::explain(sat::hint_type ty, sat::literal lit) {
arith_proof_hint const* solver::explain(hint_type ty, sat::literal lit) {
if (!ctx.use_drat())
return nullptr;
m_arith_hint.m_ty = ty;
m_arith_hint.set_type(ctx, ty);
explain_assumptions();
if (lit != sat::null_literal)
m_arith_hint.m_literals.push_back({rational(1), ~lit});
return &m_arith_hint;
m_arith_hint.add_lit(rational(1), ~lit);
return m_arith_hint.mk(ctx);
}
sat::proof_hint const* solver::explain_implied_eq(euf::enode* a, euf::enode* b) {
arith_proof_hint const* solver::explain_implied_eq(euf::enode* a, euf::enode* b) {
if (!ctx.use_drat())
return nullptr;
m_arith_hint.m_ty = sat::hint_type::implied_eq_h;
m_arith_hint.set_type(ctx, hint_type::implied_eq_h);
explain_assumptions();
m_arith_hint.m_diseqs.push_back({a->get_expr_id(), b->get_expr_id()});
return &m_arith_hint;
m_arith_hint.add_diseq(a, b);
return m_arith_hint.mk(ctx);
}
expr* arith_proof_hint::get_hint(euf::solver& s) const {
ast_manager& m = s.get_manager();
family_id fid = m.get_family_id("arith");
arith_util arith(m);
solver& a = dynamic_cast<solver&>(*s.fid2solver(fid));
char const* name;
switch (m_ty) {
case hint_type::farkas_h:
name = "farkas";
break;
case hint_type::bound_h:
name = "bound";
break;
case hint_type::implied_eq_h:
name = "implied-eq";
break;
}
rational lc(1);
for (unsigned i = m_lit_head; i < m_lit_tail; ++i)
lc = lcm(lc, denominator(a.m_arith_hint.lit(i).first));
expr_ref_vector args(m);
sort_ref_vector sorts(m);
for (unsigned i = m_lit_head; i < m_lit_tail; ++i) {
auto const& [coeff, lit] = a.m_arith_hint.lit(i);
args.push_back(arith.mk_int(coeff*lc));
args.push_back(s.literal2expr(lit));
}
for (unsigned i = m_eq_head; i < m_eq_tail; ++i) {
auto const& [a, b, is_eq] = a.m_arith_hint.eq(i);
expr_ref eq(m.mk_eq(a->get_expr(), b->get_expr()), m);
if (!is_eq) eq = m.mk_not(eq);
args.push_back(arith.mk_int(lc));
args.push_back(eq);
}
for (expr* a : args)
sorts.push_back(a->get_sort());
sort* range = m.mk_proof_sort();
func_decl* d = m.mk_func_decl(symbol(name), args.size(), sorts.data(), range);
expr* r = m.mk_app(d, args);
return r;
}
}

View file

@ -1,5 +1,5 @@
/*++
Copyright (c) 2020 Microsoft Corporation
Copyright (c) 2022 Microsoft Corporation
Module Name:
@ -11,7 +11,15 @@ Abstract:
Author:
Nikolaj Bjorner (nbjorner) 2020-09-08
Nikolaj Bjorner (nbjorner) 2022-08-28
Notes:
The module assumes a limited repertoire of arithmetic proof rules.
- farkas - inequalities, equalities and disequalities with coefficients
- implied-eq - last literal is a disequality. The literals before imply the corresponding equality.
- bound - last literal is a bound. It is implied by prior literals.
--*/
#pragma once
@ -19,11 +27,12 @@ Author:
#include "util/obj_pair_set.h"
#include "ast/ast_trail.h"
#include "ast/arith_decl_plugin.h"
#include "sat/smt/euf_proof_checker.h"
namespace arith {
class proof_checker {
class proof_checker : public euf::proof_checker_plugin {
struct row {
obj_map<expr, rational> m_coeffs;
rational m_coeff;
@ -300,6 +309,8 @@ namespace arith {
public:
proof_checker(ast_manager& m): m(m), a(m) {}
~proof_checker() override {}
void reset() {
m_ineq.reset();
@ -350,6 +361,47 @@ namespace arith {
return out;
}
bool check(expr_ref_vector const& clause, app* jst) override {
reset();
if (jst->get_name() == symbol("farkas")) {
bool even = true;
rational coeff;
expr* x, *y;
for (expr* arg : *jst) {
if (even) {
VERIFY(a.is_numeral(arg, coeff));
}
else {
bool sign = m.is_not(arg, arg);
if (a.is_le(arg) || a.is_lt(arg) || a.is_ge(arg) || a.is_gt(arg))
add_ineq(coeff, arg, sign);
else if (m.is_eq(arg, x, y)) {
if (sign)
add_diseq(x, y);
else
add_eq(x, y);
}
else
return false;
}
even = !even;
}
// display(verbose_stream());
// todo: correlate with literals in clause, literals that are not in clause should have RUP property.
return check_farkas();
}
// todo: rules for bounds and implied-by
return false;
}
void register_plugins(euf::proof_checker& pc) {
pc.register_plugin(symbol("farkas"), this);
pc.register_plugin(symbol("bound"), this);
pc.register_plugin(symbol("implied-eq"), this);
}
};

View file

@ -39,8 +39,6 @@ namespace arith {
lp().settings().set_random_seed(get_config().m_random_seed);
m_lia = alloc(lp::int_solver, *m_solver.get());
m_farkas2.m_ty = sat::hint_type::farkas_h;
m_farkas2.m_literals.resize(2);
}
solver::~solver() {
@ -197,11 +195,12 @@ namespace arith {
reset_evidence();
m_core.push_back(lit1);
TRACE("arith", tout << lit2 << " <- " << m_core << "\n";);
sat::proof_hint* ph = nullptr;
arith_proof_hint* ph = nullptr;
if (ctx.use_drat()) {
ph = &m_farkas2;
m_farkas2.m_literals[0] = std::make_pair(rational(1), lit1);
m_farkas2.m_literals[1] = std::make_pair(rational(1), ~lit2);
m_arith_hint.set_type(ctx, hint_type::farkas_h);
m_arith_hint.add_lit(rational(1), lit1);
m_arith_hint.add_lit(rational(1), ~lit2);
ph = m_arith_hint.mk(ctx);
}
assign(lit2, m_core, m_eqs, ph);
++m_stats.m_bounds_propagations;
@ -262,7 +261,7 @@ namespace arith {
TRACE("arith", for (auto lit : m_core) tout << lit << ": " << s().value(lit) << "\n";);
DEBUG_CODE(for (auto lit : m_core) { VERIFY(s().value(lit) == l_true); });
++m_stats.m_bound_propagations1;
assign(lit, m_core, m_eqs, explain(sat::hint_type::bound_h, lit));
assign(lit, m_core, m_eqs, explain(hint_type::bound_h, lit));
}
if (should_refine_bounds() && first)
@ -378,7 +377,7 @@ namespace arith {
reset_evidence();
m_explanation.clear();
lp().explain_implied_bound(be, m_bp);
assign(bound, m_core, m_eqs, explain(sat::hint_type::farkas_h, bound));
assign(bound, m_core, m_eqs, explain(hint_type::farkas_h, bound));
}
@ -1178,7 +1177,7 @@ namespace arith {
app_ref b = mk_bound(m_lia->get_term(), m_lia->get_offset(), !m_lia->is_upper());
IF_VERBOSE(4, verbose_stream() << "cut " << b << "\n");
literal lit = expr2literal(b);
assign(lit, m_core, m_eqs, explain(sat::hint_type::bound_h, lit));
assign(lit, m_core, m_eqs, explain(hint_type::bound_h, lit));
lia_check = l_false;
break;
}
@ -1200,7 +1199,7 @@ namespace arith {
return lia_check;
}
void solver::assign(literal lit, literal_vector const& core, svector<enode_pair> const& eqs, sat::proof_hint const* pma) {
void solver::assign(literal lit, literal_vector const& core, svector<enode_pair> const& eqs, euf::th_proof_hint const* pma) {
if (core.size() < small_lemma_size() && eqs.empty()) {
m_core2.reset();
for (auto const& c : core)
@ -1247,7 +1246,7 @@ namespace arith {
for (literal& c : m_core)
c.neg();
add_clause(m_core, explain(sat::hint_type::farkas_h));
add_clause(m_core, explain(hint_type::farkas_h));
}
bool solver::is_infeasible() const {

View file

@ -48,8 +48,61 @@ namespace arith {
typedef sat::literal_vector literal_vector;
typedef lp_api::bound<sat::literal> api_bound;
enum class hint_type {
farkas_h,
bound_h,
implied_eq_h
};
struct arith_proof_hint : public euf::th_proof_hint {
hint_type m_ty;
unsigned m_lit_head, m_lit_tail, m_eq_head, m_eq_tail;
arith_proof_hint(hint_type t, unsigned lh, unsigned lt, unsigned eh, unsigned et):
m_ty(t), m_lit_head(lh), m_lit_tail(lt), m_eq_head(eh), m_eq_tail(et) {}
expr* get_hint(euf::solver& s) const override;
};
class arith_proof_hint_builder {
vector<std::pair<rational, literal>> m_literals;
svector<std::tuple<euf::enode*,euf::enode*,bool>> m_eqs;
hint_type m_ty;
unsigned m_lit_head = 0, m_lit_tail = 0, m_eq_head = 0, m_eq_tail;
void reset() { m_lit_head = m_lit_tail; m_eq_head = m_eq_tail; }
void add(euf::enode* a, euf::enode* b, bool is_eq) {
if (m_eq_tail < m_eqs.size())
m_eqs[m_eq_tail] = std::tuple(a, b, is_eq);
else
m_eqs.push_back(std::tuple(a, b, is_eq));
m_eq_tail++;
}
public:
void set_type(euf::solver& ctx, hint_type ty) {
ctx.push(value_trail<unsigned>(m_eq_tail));
ctx.push(value_trail<unsigned>(m_lit_tail));
m_ty = ty;
reset();
}
void add_eq(euf::enode* a, euf::enode* b) { add(a, b, true); }
void add_diseq(euf::enode* a, euf::enode* b) { add(a, b, false); }
void add_lit(rational const& coeff, literal lit) {
if (m_lit_tail < m_literals.size())
m_literals[m_lit_tail] = {coeff, lit};
else
m_literals.push_back({coeff, lit});
m_lit_tail++;
}
std::pair<rational, literal> const& lit(unsigned i) const { return m_literals[i]; }
std::tuple<enode*, enode*, bool> const& eq(unsigned i) const { return m_eqs[i]; }
arith_proof_hint* mk(euf::solver& s) {
return new (s.get_region()) arith_proof_hint(m_ty, m_lit_head, m_lit_tail, m_eq_head, m_eq_tail);
}
};
class solver : public euf::th_euf_solver {
friend struct arith_proof_hint;
struct scope {
unsigned m_bounds_lim;
unsigned m_idiv_lim;
@ -414,15 +467,15 @@ namespace arith {
void set_conflict();
void set_conflict_or_lemma(literal_vector const& core, bool is_conflict);
void set_evidence(lp::constraint_index idx);
void assign(literal lit, literal_vector const& core, svector<enode_pair> const& eqs, sat::proof_hint const* pma);
void assign(literal lit, literal_vector const& core, svector<enode_pair> const& eqs, euf::th_proof_hint const* pma);
void false_case_of_check_nla(const nla::lemma& l);
void dbg_finalize_model(model& mdl);
sat::proof_hint m_arith_hint;
sat::proof_hint m_farkas2;
sat::proof_hint const* explain(sat::hint_type ty, sat::literal lit = sat::null_literal);
sat::proof_hint const* explain_implied_eq(euf::enode* a, euf::enode* b);
arith_proof_hint_builder m_arith_hint;
arith_proof_hint const* explain(hint_type ty, sat::literal lit = sat::null_literal);
arith_proof_hint const* explain_implied_eq(euf::enode* a, euf::enode* b);
void explain_assumptions();

View file

@ -391,8 +391,8 @@ namespace bv {
if (c.m_kind != bv_justification::kind_t::bit2ne) {
expr* e1 = var2expr(c.m_v1);
expr* e2 = var2expr(c.m_v2);
eq = m.mk_eq(e1, e2);
ctx.drat_eq_def(leq, eq);
eq = m.mk_eq(e1, e2);
ctx.set_tmp_bool_var(leq.var(), eq);
}
sat::literal_vector lits;

View file

@ -21,130 +21,21 @@ Author:
namespace euf {
void solver::init_drat() {
if (!m_drat_initialized) {
void solver::init_proof() {
if (!m_proof_initialized) {
get_drat().add_theory(get_id(), symbol("euf"));
get_drat().add_theory(m.get_basic_family_id(), symbol("bool"));
}
m_drat_initialized = true;
}
void solver::def_add_arg(unsigned arg) {
auto* out = get_drat().out();
if (out)
(*out) << " " << arg;
}
void solver::def_end() {
auto* out = get_drat().out();
if (out)
(*out) << " 0\n";
}
void solver::def_begin(char id, unsigned n, std::string const& name) {
auto* out = get_drat().out();
if (out)
(*out) << id << " " << n << " " << name;
}
void solver::bool_def(bool_var v, unsigned n) {
auto* out = get_drat().out();
if (out)
(*out) << "b " << v << " " << n << " 0\n";
}
void solver::drat_log_params(func_decl* f) {
for (unsigned i = f->get_num_parameters(); i-- > 0; ) {
auto const& p = f->get_parameter(i);
if (!p.is_ast())
continue;
ast* a = p.get_ast();
if (is_func_decl(a))
drat_log_decl(to_func_decl(a));
if (!m_proof_out && s().get_config().m_drat &&
(get_config().m_lemmas2console || s().get_config().m_smt_proof.is_non_empty_string())) {
TRACE("euf", tout << "init-proof\n");
m_proof_out = alloc(std::ofstream, s().get_config().m_smt_proof.str(), std::ios_base::out);
if (get_config().m_lemmas2console)
get_drat().set_clause_eh(*this);
if (s().get_config().m_smt_proof.is_non_empty_string())
get_drat().set_clause_eh(*this);
}
}
void solver::drat_log_expr1(expr* e) {
if (is_app(e)) {
app* a = to_app(e);
drat_log_params(a->get_decl());
drat_log_decl(a->get_decl());
std::stringstream strm;
strm << mk_ismt2_func(a->get_decl(), m);
def_begin('e', e->get_id(), strm.str());
for (expr* arg : *a)
def_add_arg(arg->get_id());
def_end();
}
else if (is_var(e)) {
var* v = to_var(e);
def_begin('v', v->get_id(), "" + mk_pp(e->get_sort(), m));
def_add_arg(v->get_idx());
def_end();
}
else if (is_quantifier(e)) {
quantifier* q = to_quantifier(e);
std::stringstream strm;
strm << "(" << (is_forall(q) ? "forall" : (is_exists(q) ? "exists" : "lambda"));
for (unsigned i = 0; i < q->get_num_decls(); ++i)
strm << " (" << q->get_decl_name(i) << " " << mk_pp(q->get_decl_sort(i), m) << ")";
strm << ")";
def_begin('q', q->get_id(), strm.str());
def_add_arg(q->get_expr()->get_id());
def_end();
}
else
UNREACHABLE();
m_drat_asts.insert(e);
push(insert_obj_trail<ast>(m_drat_asts, e));
}
void solver::drat_log_expr(expr* e) {
if (m_drat_asts.contains(e))
return;
ptr_vector<expr>::scoped_stack _sc(m_drat_todo);
m_drat_todo.push_back(e);
while (!m_drat_todo.empty()) {
e = m_drat_todo.back();
unsigned sz = m_drat_todo.size();
if (is_app(e))
for (expr* arg : *to_app(e))
if (!m_drat_asts.contains(arg))
m_drat_todo.push_back(arg);
if (is_quantifier(e)) {
expr* arg = to_quantifier(e)->get_expr();
if (!m_drat_asts.contains(arg))
m_drat_todo.push_back(arg);
}
if (m_drat_todo.size() != sz)
continue;
if (!m_drat_asts.contains(e))
drat_log_expr1(e);
m_drat_todo.pop_back();
}
}
void solver::drat_bool_def(sat::bool_var v, expr* e) {
if (!use_drat())
return;
drat_log_expr(e);
bool_def(v, e->get_id());
}
void solver::drat_log_decl(func_decl* f) {
if (f->get_family_id() != null_family_id)
return;
if (m_drat_asts.contains(f))
return;
m_drat_asts.insert(f);
push(insert_obj_trail< ast>(m_drat_asts, f));
std::ostringstream strm;
smt2_pp_environment_dbg env(m);
ast_smt2_pp(strm, f, env);
def_begin('f', f->get_small_id(), strm.str());
def_end();
m_proof_initialized = true;
}
/**
@ -183,16 +74,19 @@ namespace euf {
}
}
void solver::set_tmp_bool_var(bool_var b, expr* e) {
m_bool_var2expr.setx(b, e, nullptr);
}
void solver::log_justification(literal l, th_explain const& jst) {
literal_vector lits;
unsigned nv = s().num_vars();
expr_ref_vector eqs(m);
unsigned nv = s().num_vars();
auto add_lit = [&](enode_pair const& eq) {
++nv;
literal lit(nv, false);
eqs.push_back(m.mk_eq(eq.first->get_expr(), eq.second->get_expr()));
drat_eq_def(lit, eqs.back());
return lit;
set_tmp_bool_var(nv, eqs.back());
return literal(nv, false);
};
for (auto lit : euf::th_explain::lits(jst))
@ -208,68 +102,133 @@ namespace euf {
get_drat().add(lits, sat::status::th(m_is_redundant, jst.ext().get_id(), jst.get_pragma()));
}
void solver::drat_eq_def(literal lit, expr* eq) {
expr *a = nullptr, *b = nullptr;
VERIFY(m.is_eq(eq, a, b));
drat_log_expr(a);
drat_log_expr(b);
def_begin('e', eq->get_id(), std::string("="));
def_add_arg(a->get_id());
def_add_arg(b->get_id());
def_end();
bool_def(lit.var(), eq->get_id());
void solver::on_clause(unsigned n, literal const* lits, sat::status st) {
TRACE("euf", tout << "on-clause " << n << "\n");
on_lemma(n, lits, st);
on_proof(n, lits, st);
}
void solver::on_clause(unsigned n, literal const* lits, sat::status st) {
void solver::on_proof(unsigned n, literal const* lits, sat::status st) {
if (!m_proof_out)
return;
flet<bool> _display_all_decls(m_display_all_decls, true);
std::ostream& out = *m_proof_out;
if (!visit_clause(out, n, lits))
return;
if (st.is_asserted())
display_redundant(out, n, lits, status2proof_hint(st));
else if (st.is_deleted())
display_deleted(out, n, lits);
else if (st.is_redundant())
display_redundant(out, n, lits, status2proof_hint(st));
else if (st.is_input())
display_assume(out, n, lits);
else
UNREACHABLE();
out.flush();
}
void solver::on_lemma(unsigned n, literal const* lits, sat::status st) {
if (!get_config().m_lemmas2console)
return;
if (!st.is_redundant() && !st.is_asserted())
return;
if (!visit_clause(n, lits))
std::ostream& out = std::cout;
if (!visit_clause(out, n, lits))
return;
std::function<symbol(int)> ppth = [&](int th) {
return m.get_family_name(th);
};
if (!st.is_sat())
std::cout << "; " << sat::status_pp(st, ppth) << "\n";
out << "; " << sat::status_pp(st, ppth) << "\n";
display_clause(n, lits);
display_assert(out, n, lits);
}
void solver::on_instantiation(unsigned n, sat::literal const* lits, unsigned k, euf::enode* const* bindings) {
std::ostream& out = std::cout;
for (unsigned i = 0; i < k; ++i)
visit_expr(out, bindings[i]->get_expr());
VERIFY(visit_clause(out, n, lits));
out << "(instantiate";
display_literals(out, n, lits);
for (unsigned i = 0; i < k; ++i)
display_expr(out << " :binding ", bindings[i]->get_expr());
out << ")\n";
}
bool solver::visit_clause(unsigned n, literal const* lits) {
bool solver::visit_clause(std::ostream& out, unsigned n, literal const* lits) {
for (unsigned i = 0; i < n; ++i) {
expr* e = bool_var2expr(lits[i].var());
if (!e)
return false;
visit_expr(e);
visit_expr(out, e);
}
return true;
}
void solver::display_clause(unsigned n, literal const* lits) {
std::cout << "(assert (or";
void solver::display_assert(std::ostream& out, unsigned n, literal const* lits) {
display_literals(out << "(assert (or", n, lits) << "))\n";
}
void solver::display_assume(std::ostream& out, unsigned n, literal const* lits) {
display_literals(out << "(assume", n, lits) << ")\n";
}
void solver::display_redundant(std::ostream& out, unsigned n, literal const* lits, expr_ref& proof_hint) {
if (proof_hint)
visit_expr(out, proof_hint);
display_hint(display_literals(out << "(learn", n, lits), proof_hint) << ")\n";
}
void solver::display_deleted(std::ostream& out, unsigned n, literal const* lits) {
display_literals(out << "(del", n, lits) << ")\n";
}
std::ostream& solver::display_hint(std::ostream& out, expr* proof_hint) {
if (proof_hint)
return display_expr(out << " ", proof_hint);
else
return out;
}
expr_ref solver::status2proof_hint(sat::status st) {
if (st.is_sat())
return expr_ref(m.mk_const("rup", m.mk_proof_sort()), m); // provable by reverse unit propagation
auto* h = reinterpret_cast<euf::th_proof_hint const*>(st.get_hint());
if (!h)
return expr_ref(m);
expr* e = h->get_hint(*this);
if (e)
return expr_ref(e, m);
return expr_ref(m);
}
std::ostream& solver::display_literals(std::ostream& out, unsigned n, literal const* lits) {
for (unsigned i = 0; i < n; ++i) {
expr* e = bool_var2expr(lits[i].var());
if (lits[i].sign())
m_clause_visitor.display_expr_def(std::cout << " (not ", e) << ")";
display_expr(out << " (not ", e) << ")";
else
m_clause_visitor.display_expr_def(std::cout << " ", e);
display_expr(out << " ", e);
}
std::cout << "))\n";
return out;
}
void solver::visit_expr(expr* e) {
void solver::visit_expr(std::ostream& out, expr* e) {
m_clause_visitor.collect(e);
m_clause_visitor.display_skolem_decls(std::cout);
m_clause_visitor.define_expr(std::cout, e);
if (m_display_all_decls)
m_clause_visitor.display_decls(out);
else
m_clause_visitor.display_skolem_decls(out);
m_clause_visitor.define_expr(out, e);
}
void solver::display_expr(expr* e) {
m_clause_visitor.display_expr_def(std::cout, e);
std::ostream& solver::display_expr(std::ostream& out, expr* e) {
return m_clause_visitor.display_expr_def(out, e);
}
}

View file

@ -0,0 +1,48 @@
/*++
Copyright (c) 2020 Microsoft Corporation
Module Name:
euf_proof_checker.cpp
Abstract:
Plugin manager for checking EUF proofs
Author:
Nikolaj Bjorner (nbjorner) 2020-08-25
--*/
#include "ast/ast_pp.h"
#include "sat/smt/euf_proof_checker.h"
#include "sat/smt/arith_proof_checker.h"
namespace euf {
proof_checker::proof_checker(ast_manager& m):
m(m) {
arith::proof_checker* apc = alloc(arith::proof_checker, m);
m_plugins.push_back(apc);
apc->register_plugins(*this);
}
proof_checker::~proof_checker() {}
void proof_checker::register_plugin(symbol const& rule, proof_checker_plugin* p) {
m_map.insert(rule, p);
}
bool proof_checker::check(expr_ref_vector const& clause, expr* e) {
if (!e || !is_app(e))
return false;
app* a = to_app(e);
proof_checker_plugin* p = nullptr;
if (m_map.find(a->get_decl()->get_name(), p))
return p->check(clause, a);
return false;
}
}

View file

@ -0,0 +1,46 @@
/*++
Copyright (c) 2022 Microsoft Corporation
Module Name:
euf_proof_checker.h
Abstract:
Plugin manager for checking EUF proofs
Author:
Nikolaj Bjorner (nbjorner) 2022-08-25
--*/
#pragma once
#include "util/map.h"
#include "util/scoped_ptr_vector.h"
#include "ast/ast.h"
namespace euf {
class proof_checker;
class proof_checker_plugin {
public:
virtual ~proof_checker_plugin() {}
virtual bool check(expr_ref_vector const& clause, app* jst) = 0;
virtual void register_plugins(proof_checker& pc) = 0;
};
class proof_checker {
ast_manager& m;
scoped_ptr_vector<proof_checker_plugin> m_plugins;
map<symbol, proof_checker_plugin*, symbol_hash_proc, symbol_eq_proc> m_map;
public:
proof_checker(ast_manager& m);
~proof_checker();
void register_plugin(symbol const& rule, proof_checker_plugin*);
bool check(expr_ref_vector const& clause, expr* e);
};
}

View file

@ -170,9 +170,6 @@ namespace euf {
TRACE("before_search", s().display(tout););
for (auto* s : m_solvers)
s->init_search();
if (get_config().m_lemmas2console)
get_drat().set_print_clause(*this);
}
bool solver::is_external(bool_var v) {

View file

@ -60,11 +60,10 @@ namespace euf {
std::ostream& display(std::ostream& out) const;
};
class solver : public sat::extension, public th_internalizer, public th_decompile, public sat::print_clause {
class solver : public sat::extension, public th_internalizer, public th_decompile, public sat::clause_eh {
typedef top_sort<euf::enode> deps_t;
friend class ackerman;
class user_sort;
// friend class sat::ba_solver;
struct stats {
unsigned m_ackerman;
unsigned m_final_checks;
@ -175,26 +174,22 @@ namespace euf {
void log_antecedents(std::ostream& out, literal l, literal_vector const& r);
void log_antecedents(literal l, literal_vector const& r);
void log_justification(literal l, th_explain const& jst);
void drat_log_decl(func_decl* f);
void drat_log_params(func_decl* f);
void drat_log_expr1(expr* n);
ptr_vector<expr> m_drat_todo;
obj_hashtable<ast> m_drat_asts;
bool m_drat_initialized{ false };
void init_drat();
ast_pp_util m_clause_visitor;
void on_clause(unsigned n, literal const* lits, sat::status st) override;
void def_add_arg(unsigned arg);
void def_end();
void def_begin(char id, unsigned n, std::string const& name);
bool m_proof_initialized = false;
void init_proof();
ast_pp_util m_clause_visitor;
bool m_display_all_decls = false;
void on_clause(unsigned n, literal const* lits, sat::status st) override;
void on_lemma(unsigned n, literal const* lits, sat::status st);
void on_proof(unsigned n, literal const* lits, sat::status st);
std::ostream& display_literals(std::ostream& out, unsigned n, sat::literal const* lits);
void display_assume(std::ostream& out, unsigned n, literal const* lits);
void display_redundant(std::ostream& out, unsigned n, literal const* lits, expr_ref& proof_hint);
void display_deleted(std::ostream& out, unsigned n, literal const* lits);
std::ostream& display_hint(std::ostream& out, expr* proof_hint);
expr_ref status2proof_hint(sat::status st);
// relevancy
//bool_vector m_relevant_expr_ids;
//bool_vector m_relevant_visited;
//ptr_vector<expr> m_relevant_todo;
//void init_relevant_expr_ids();
//void push_relevant(sat::bool_var v);
bool is_propagated(sat::literal lit);
// invariant
void check_eqc_bool_assignment() const;
@ -347,16 +342,16 @@ namespace euf {
// proof
bool use_drat() { return s().get_config().m_drat && (init_drat(), true); }
bool use_drat() { return s().get_config().m_drat && (init_proof(), true); }
sat::drat& get_drat() { return s().get_drat(); }
void drat_bool_def(sat::bool_var v, expr* n);
void drat_eq_def(sat::literal lit, expr* eq);
void drat_log_expr(expr* n);
void bool_def(bool_var v, unsigned n);
bool visit_clause(unsigned n, literal const* lits);
void display_clause(unsigned n, literal const* lits);
void visit_expr(expr* e);
void display_expr(expr* e);
void set_tmp_bool_var(sat::bool_var b, expr* e);
bool visit_clause(std::ostream& out, unsigned n, literal const* lits);
void display_assert(std::ostream& out, unsigned n, literal const* lits);
void visit_expr(std::ostream& out, expr* e);
std::ostream& display_expr(std::ostream& out, expr* e);
void on_instantiation(unsigned n, sat::literal const* lits, unsigned k, euf::enode* const* bindings);
scoped_ptr<std::ostream> m_proof_out;
// decompile
bool extract_pb(std::function<void(unsigned sz, literal const* c, unsigned k)>& card,

View file

@ -360,19 +360,7 @@ namespace q {
void solver::log_instantiation(unsigned n, sat::literal const* lits, justification* j) {
TRACE("q", for (unsigned i = 0; i < n; ++i) tout << literal2expr(lits[i]) << "\n";);
if (get_config().m_instantiations2console) {
ctx.visit_clause(n, lits);
if (j) {
for (unsigned i = 0; i < j->m_clause.num_decls(); ++i)
ctx.visit_expr(j->m_binding[i]->get_expr());
std::cout << "; (instantiation";
for (unsigned i = 0; i < j->m_clause.num_decls(); ++i) {
std::cout << " ";
ctx.display_expr(j->m_binding[i]->get_expr());
}
std::cout << ")\n";
}
ctx.display_clause(n, lits);
ctx.on_instantiation(n, lits, j ? j->m_clause.num_decls() : 0, j ? j->m_binding : nullptr);
}
}
}

View file

@ -125,7 +125,7 @@ namespace euf {
pop_core(n);
}
sat::status th_euf_solver::mk_status(sat::proof_hint const* ps) {
sat::status th_euf_solver::mk_status(th_proof_hint const* ps) {
return sat::status::th(m_is_redundant, get_id(), ps);
}
@ -149,7 +149,7 @@ namespace euf {
return add_clause(2, lits);
}
bool th_euf_solver::add_clause(sat::literal a, sat::literal b, sat::proof_hint const* ps) {
bool th_euf_solver::add_clause(sat::literal a, sat::literal b, th_proof_hint const* ps) {
sat::literal lits[2] = { a, b };
return add_clause(2, lits, ps);
}
@ -164,7 +164,7 @@ namespace euf {
return add_clause(4, lits);
}
bool th_euf_solver::add_clause(unsigned n, sat::literal* lits, sat::proof_hint const* ps) {
bool th_euf_solver::add_clause(unsigned n, sat::literal* lits, th_proof_hint const* ps) {
bool was_true = false;
for (unsigned i = 0; i < n; ++i)
was_true |= is_true(lits[i]);
@ -226,13 +226,14 @@ namespace euf {
return ctx.s().rand()();
}
size_t th_explain::get_obj_size(unsigned num_lits, unsigned num_eqs, sat::proof_hint const* pma) {
return sat::constraint_base::obj_size(sizeof(th_explain) + sizeof(sat::literal) * num_lits + sizeof(enode_pair) * num_eqs + (pma?pma->to_string().length()+1:1));
size_t th_explain::get_obj_size(unsigned num_lits, unsigned num_eqs) {
return sat::constraint_base::obj_size(sizeof(th_explain) + sizeof(sat::literal) * num_lits + sizeof(enode_pair) * num_eqs);
}
th_explain::th_explain(unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode_pair const& p, sat::proof_hint const* pma) {
th_explain::th_explain(unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode_pair const& p, th_proof_hint const* pma) {
m_consequent = c;
m_eq = p;
m_proof_hint = pma;
m_num_literals = n_lits;
m_num_eqs = n_eqs;
char * base_ptr = reinterpret_cast<char*>(this) + sizeof(th_explain);
@ -244,33 +245,24 @@ namespace euf {
m_eqs = reinterpret_cast<enode_pair*>(base_ptr);
for (i = 0; i < n_eqs; ++i)
m_eqs[i] = eqs[i];
base_ptr += sizeof(enode_pair) * n_eqs;
m_pragma = reinterpret_cast<char*>(base_ptr);
i = 0;
if (pma) {
std::string s = pma->to_string();
for (i = 0; s[i]; ++i)
m_pragma[i] = s[i];
}
m_pragma[i] = 0;
}
th_explain* th_explain::mk(th_euf_solver& th, unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode* x, enode* y, sat::proof_hint const* pma) {
th_explain* th_explain::mk(th_euf_solver& th, unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode* x, enode* y, th_proof_hint const* pma) {
region& r = th.ctx.get_region();
void* mem = r.allocate(get_obj_size(n_lits, n_eqs, pma));
void* mem = r.allocate(get_obj_size(n_lits, n_eqs));
sat::constraint_base::initialize(mem, &th);
return new (sat::constraint_base::ptr2mem(mem)) th_explain(n_lits, lits, n_eqs, eqs, c, enode_pair(x, y));
return new (sat::constraint_base::ptr2mem(mem)) th_explain(n_lits, lits, n_eqs, eqs, c, enode_pair(x, y), pma);
}
th_explain* th_explain::propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, sat::literal consequent, sat::proof_hint const* pma) {
th_explain* th_explain::propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, sat::literal consequent, th_proof_hint const* pma) {
return mk(th, lits.size(), lits.data(), eqs.size(), eqs.data(), consequent, nullptr, nullptr, pma);
}
th_explain* th_explain::propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, euf::enode* x, euf::enode* y, sat::proof_hint const* pma) {
th_explain* th_explain::propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, euf::enode* x, euf::enode* y, th_proof_hint const* pma) {
return mk(th, lits.size(), lits.data(), eqs.size(), eqs.data(), sat::null_literal, x, y, pma);
}
th_explain* th_explain::propagate(th_euf_solver& th, enode_pair_vector const& eqs, euf::enode* x, euf::enode* y, sat::proof_hint const* pma) {
th_explain* th_explain::propagate(th_euf_solver& th, enode_pair_vector const& eqs, euf::enode* x, euf::enode* y, th_proof_hint const* pma) {
return mk(th, 0, nullptr, eqs.size(), eqs.data(), sat::null_literal, x, y, pma);
}
@ -313,8 +305,8 @@ namespace euf {
out << "--> " << m_consequent;
if (m_eq.first != nullptr)
out << "--> " << m_eq.first->get_expr_id() << " == " << m_eq.second->get_expr_id();
if (m_pragma != nullptr)
out << " p " << m_pragma;
if (m_proof_hint != nullptr)
out << " p ";
return out;
}

View file

@ -58,6 +58,7 @@ namespace euf {
};
class th_decompile {
public:
virtual ~th_decompile() = default;
@ -138,6 +139,11 @@ namespace euf {
};
class th_proof_hint : public sat::proof_hint {
public:
virtual expr* get_hint(euf::solver& s) const = 0;
};
class th_euf_solver : public th_solver {
protected:
solver& ctx;
@ -150,16 +156,16 @@ namespace euf {
region& get_region();
sat::status mk_status(sat::proof_hint const* ps = nullptr);
sat::status mk_status(th_proof_hint const* ps = nullptr);
bool add_unit(sat::literal lit);
bool add_units(sat::literal_vector const& lits);
bool add_clause(sat::literal lit) { return add_unit(lit); }
bool add_clause(sat::literal a, sat::literal b);
bool add_clause(sat::literal a, sat::literal b, sat::proof_hint const* ps);
bool add_clause(sat::literal a, sat::literal b, th_proof_hint const* ps);
bool add_clause(sat::literal a, sat::literal b, sat::literal c);
bool add_clause(sat::literal a, sat::literal b, sat::literal c, sat::literal d);
bool add_clause(sat::literal_vector const& lits, sat::proof_hint const* ps = nullptr) { return add_clause(lits.size(), lits.data(), ps); }
bool add_clause(unsigned n, sat::literal* lits, sat::proof_hint const* ps = nullptr);
bool add_clause(sat::literal_vector const& lits, th_proof_hint const* ps = nullptr) { return add_clause(lits.size(), lits.data(), ps); }
bool add_clause(unsigned n, sat::literal* lits, th_proof_hint const* ps = nullptr);
void add_equiv(sat::literal a, sat::literal b);
void add_equiv_and(sat::literal a, sat::literal_vector const& bs);
@ -220,16 +226,16 @@ namespace euf {
* that retrieve literals on demand.
*/
class th_explain {
sat::literal m_consequent = sat::null_literal; // literal consequent for propagations
enode_pair m_eq = enode_pair(); // equality consequent for propagations
sat::literal m_consequent = sat::null_literal; // literal consequent for propagations
enode_pair m_eq = enode_pair(); // equality consequent for propagations
th_proof_hint const* m_proof_hint;
unsigned m_num_literals;
unsigned m_num_eqs;
sat::literal* m_literals;
enode_pair* m_eqs;
char* m_pragma = nullptr;
static size_t get_obj_size(unsigned num_lits, unsigned num_eqs, sat::proof_hint const* pma);
th_explain(unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode_pair const& eq, sat::proof_hint const* pma = nullptr);
static th_explain* mk(th_euf_solver& th, unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode* x, enode* y, sat::proof_hint const* pma = nullptr);
static size_t get_obj_size(unsigned num_lits, unsigned num_eqs);
th_explain(unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode_pair const& eq, th_proof_hint const* pma = nullptr);
static th_explain* mk(th_euf_solver& th, unsigned n_lits, sat::literal const* lits, unsigned n_eqs, enode_pair const* eqs, sat::literal c, enode* x, enode* y, th_proof_hint const* pma = nullptr);
public:
static th_explain* conflict(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs);
@ -240,9 +246,9 @@ namespace euf {
static th_explain* conflict(th_euf_solver& th, sat::literal lit, euf::enode* x, euf::enode* y);
static th_explain* conflict(th_euf_solver& th, euf::enode* x, euf::enode* y);
static th_explain* propagate(th_euf_solver& th, sat::literal lit, euf::enode* x, euf::enode* y);
static th_explain* propagate(th_euf_solver& th, enode_pair_vector const& eqs, euf::enode* x, euf::enode* y, sat::proof_hint const* pma = nullptr);
static th_explain* propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, sat::literal consequent, sat::proof_hint const* pma = nullptr);
static th_explain* propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, euf::enode* x, euf::enode* y, sat::proof_hint const* pma = nullptr);
static th_explain* propagate(th_euf_solver& th, enode_pair_vector const& eqs, euf::enode* x, euf::enode* y, th_proof_hint const* pma = nullptr);
static th_explain* propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, sat::literal consequent, th_proof_hint const* pma = nullptr);
static th_explain* propagate(th_euf_solver& th, sat::literal_vector const& lits, enode_pair_vector const& eqs, euf::enode* x, euf::enode* y, th_proof_hint const* pma = nullptr);
sat::ext_constraint_idx to_index() const {
return sat::constraint_base::mem2base(this);
@ -277,7 +283,7 @@ namespace euf {
enode_pair eq_consequent() const { return m_eq; }
sat::proof_hint const* get_pragma() const { return nullptr; } //*m_pragma ? m_pragma : nullptr; }
th_proof_hint const* get_pragma() const { return m_proof_hint; }
};