3
0
Fork 0
mirror of https://github.com/Z3Prover/z3 synced 2025-08-23 11:37:54 +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

@ -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();