mirror of
https://github.com/Z3Prover/z3
synced 2025-08-08 04:01:22 +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:
parent
9922c766b9
commit
e2f4fc2307
37 changed files with 809 additions and 1078 deletions
|
@ -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,
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue