3
0
Fork 0
mirror of https://github.com/Z3Prover/z3 synced 2025-04-24 01:25:31 +00:00
Signed-off-by: Nikolaj Bjorner <nbjorner@microsoft.com>
This commit is contained in:
Nikolaj Bjorner 2018-10-14 15:16:22 -07:00
commit e9d615e309
204 changed files with 4620 additions and 2435 deletions

View file

@ -1862,8 +1862,9 @@ namespace sat {
return p;
}
ba_solver::ba_solver(): m_solver(0), m_lookahead(0), m_unit_walk(0),
m_allocator("ba"), m_constraint_id(0), m_ba(*this), m_sort(m_ba) {
ba_solver::ba_solver()
: m_solver(nullptr), m_lookahead(nullptr), m_unit_walk(nullptr),
m_constraint_id(0), m_ba(*this), m_sort(m_ba) {
TRACE("ba", tout << this << "\n";);
m_num_propagations_since_pop = 0;
}
@ -3838,7 +3839,7 @@ namespace sat {
}
init_visited();
for (wliteral l : p1) {
SASSERT(m_weights.get(l.second.index(), 0) == 0);
SASSERT(m_weights.size() <= l.second.index() || m_weights[l.second.index()] == 0);
m_weights.setx(l.second.index(), l.first, 0);
mark_visited(l.second);
}
@ -4394,6 +4395,12 @@ namespace sat {
}
if (slack >= k) {
#if 0
return active2constraint();
active2pb(m_A);
std::cout << "not asserting\n";
display(std::cout, m_A, true);
#endif
return nullptr;
}

View file

@ -134,7 +134,7 @@ namespace sat {
virtual void set_k(unsigned k) { VERIFY(k < 4000000000); m_k = k; }
virtual unsigned get_coeff(unsigned i) const { UNREACHABLE(); return 0; }
unsigned k() const { return m_k; }
virtual bool well_formed() const;
bool well_formed() const override;
};
class card : public pb_base {
@ -146,13 +146,13 @@ namespace sat {
literal& operator[](unsigned i) { return m_lits[i]; }
literal const* begin() const { return m_lits; }
literal const* end() const { return static_cast<literal const*>(m_lits) + m_size; }
virtual void negate();
virtual void swap(unsigned i, unsigned j) { std::swap(m_lits[i], m_lits[j]); }
virtual literal_vector literals() const { return literal_vector(m_size, m_lits); }
virtual bool is_watching(literal l) const;
virtual literal get_lit(unsigned i) const { return m_lits[i]; }
virtual void set_lit(unsigned i, literal l) { m_lits[i] = l; }
virtual unsigned get_coeff(unsigned i) const { return 1; }
void negate() override;
void swap(unsigned i, unsigned j) override { std::swap(m_lits[i], m_lits[j]); }
literal_vector literals() const override { return literal_vector(m_size, m_lits); }
bool is_watching(literal l) const override;
literal get_lit(unsigned i) const override { return m_lits[i]; }
void set_lit(unsigned i, literal l) override { m_lits[i] = l; }
unsigned get_coeff(unsigned i) const override { return 1; }
};
@ -179,14 +179,14 @@ namespace sat {
void update_max_sum();
void set_num_watch(unsigned s) { m_num_watch = s; }
bool is_cardinality() const;
virtual void negate();
virtual void set_k(unsigned k) { m_k = k; VERIFY(k < 4000000000); update_max_sum(); }
virtual void swap(unsigned i, unsigned j) { std::swap(m_wlits[i], m_wlits[j]); }
virtual literal_vector literals() const { literal_vector lits; for (auto wl : *this) lits.push_back(wl.second); return lits; }
virtual bool is_watching(literal l) const;
virtual literal get_lit(unsigned i) const { return m_wlits[i].second; }
virtual void set_lit(unsigned i, literal l) { m_wlits[i].second = l; }
virtual unsigned get_coeff(unsigned i) const { return m_wlits[i].first; }
void negate() override;
void set_k(unsigned k) override { m_k = k; VERIFY(k < 4000000000); update_max_sum(); }
void swap(unsigned i, unsigned j) override { std::swap(m_wlits[i], m_wlits[j]); }
literal_vector literals() const override { literal_vector lits; for (auto wl : *this) lits.push_back(wl.second); return lits; }
bool is_watching(literal l) const override;
literal get_lit(unsigned i) const override { return m_wlits[i].second; }
void set_lit(unsigned i, literal l) override { m_wlits[i].second = l; }
unsigned get_coeff(unsigned i) const override { return m_wlits[i].first; }
};
class xr : public constraint {
@ -197,13 +197,13 @@ namespace sat {
literal operator[](unsigned i) const { return m_lits[i]; }
literal const* begin() const { return m_lits; }
literal const* end() const { return begin() + m_size; }
virtual void negate() { m_lits[0].neg(); }
virtual void swap(unsigned i, unsigned j) { std::swap(m_lits[i], m_lits[j]); }
virtual bool is_watching(literal l) const;
virtual literal_vector literals() const { return literal_vector(size(), begin()); }
virtual literal get_lit(unsigned i) const { return m_lits[i]; }
virtual void set_lit(unsigned i, literal l) { m_lits[i] = l; }
virtual bool well_formed() const;
void negate() override { m_lits[0].neg(); }
void swap(unsigned i, unsigned j) override { std::swap(m_lits[i], m_lits[j]); }
bool is_watching(literal l) const override;
literal_vector literals() const override { return literal_vector(size(), begin()); }
literal get_lit(unsigned i) const override { return m_lits[i]; }
void set_lit(unsigned i, literal l) override { m_lits[i] = l; }
bool well_formed() const override;
};
@ -524,44 +524,44 @@ namespace sat {
public:
ba_solver();
virtual ~ba_solver();
virtual void set_solver(solver* s) { m_solver = s; }
virtual void set_lookahead(lookahead* l) { m_lookahead = l; }
virtual void set_unit_walk(unit_walk* u) { m_unit_walk = u; }
~ba_solver() override;
void set_solver(solver* s) override { m_solver = s; }
void set_lookahead(lookahead* l) override { m_lookahead = l; }
void set_unit_walk(unit_walk* u) override { m_unit_walk = u; }
void add_at_least(bool_var v, literal_vector const& lits, unsigned k);
void add_pb_ge(bool_var v, svector<wliteral> const& wlits, unsigned k);
void add_xr(literal_vector const& lits);
virtual bool propagate(literal l, ext_constraint_idx idx);
virtual lbool resolve_conflict();
virtual void get_antecedents(literal l, ext_justification_idx idx, literal_vector & r);
virtual void asserted(literal l);
virtual check_result check();
virtual void push();
virtual void pop(unsigned n);
virtual void simplify();
virtual void clauses_modifed();
virtual lbool get_phase(bool_var v);
virtual bool set_root(literal l, literal r);
virtual void flush_roots();
virtual std::ostream& display(std::ostream& out) const;
virtual std::ostream& display_justification(std::ostream& out, ext_justification_idx idx) const;
virtual void collect_statistics(statistics& st) const;
virtual extension* copy(solver* s);
virtual extension* copy(lookahead* s, bool learned);
virtual void find_mutexes(literal_vector& lits, vector<literal_vector> & mutexes);
virtual void pop_reinit();
virtual void gc();
virtual double get_reward(literal l, ext_justification_idx idx, literal_occs_fun& occs) const;
virtual bool is_extended_binary(ext_justification_idx idx, literal_vector & r);
virtual void init_use_list(ext_use_list& ul);
virtual bool is_blocked(literal l, ext_constraint_idx idx);
virtual bool check_model(model const& m) const;
bool propagate(literal l, ext_constraint_idx idx) override;
lbool resolve_conflict() override;
void get_antecedents(literal l, ext_justification_idx idx, literal_vector & r) override;
void asserted(literal l) override;
check_result check() override;
void push() override;
void pop(unsigned n) override;
void simplify() override;
void clauses_modifed() override;
lbool get_phase(bool_var v) override;
bool set_root(literal l, literal r) override;
void flush_roots() override;
std::ostream& display(std::ostream& out) const override;
std::ostream& display_justification(std::ostream& out, ext_justification_idx idx) const override;
void collect_statistics(statistics& st) const override;
extension* copy(solver* s) override;
extension* copy(lookahead* s, bool learned) override;
void find_mutexes(literal_vector& lits, vector<literal_vector> & mutexes) override;
void pop_reinit() override;
void gc() override;
double get_reward(literal l, ext_justification_idx idx, literal_occs_fun& occs) const override;
bool is_extended_binary(ext_justification_idx idx, literal_vector & r) override;
void init_use_list(ext_use_list& ul) override;
bool is_blocked(literal l, ext_constraint_idx idx) override;
bool check_model(model const& m) const override;
ptr_vector<constraint> const & constraints() const { return m_constraints; }
void display(std::ostream& out, constraint const& c, bool values) const;
virtual bool validate();
bool validate() override;
};

View file

@ -152,6 +152,8 @@ namespace sat {
m_gc_burst = p.gc_burst();
m_gc_defrag = p.gc_defrag();
m_force_cleanup = p.force_cleanup();
m_minimize_lemmas = p.minimize_lemmas();
m_core_minimize = p.core_minimize();
m_core_minimize_partial = p.core_minimize_partial();

View file

@ -146,6 +146,8 @@ namespace sat {
bool m_gc_burst;
bool m_gc_defrag;
bool m_force_cleanup;
bool m_minimize_lemmas;
bool m_dyn_sub_res;

View file

@ -26,7 +26,7 @@ Notes:
namespace sat {
drat::drat(solver& s):
s(s),
m_out(0),
m_out(nullptr),
m_inconsistent(false),
m_check_unsat(false),
m_check_sat(false),

View file

@ -529,7 +529,7 @@ namespace sat {
}
lbool local_search::check() {
return check(0, 0);
return check(0, nullptr);
}
#define PROGRESS(tries, flips) \

View file

@ -277,7 +277,7 @@ namespace sat {
lbool check();
lbool check(unsigned sz, literal const* assumptions, parallel* p = 0);
lbool check(unsigned sz, literal const* assumptions, parallel* p = nullptr);
local_search_config& config() { return m_config; }

View file

@ -16,6 +16,8 @@ Author:
Notes:
--*/
#include <cmath>
@ -31,7 +33,7 @@ namespace sat {
}
lookahead::scoped_ext::~scoped_ext() {
if (p.m_s.m_ext) p.m_s.m_ext->set_lookahead(0);
if (p.m_s.m_ext) p.m_s.m_ext->set_lookahead(nullptr);
}
lookahead::scoped_assumptions::scoped_assumptions(lookahead& p, literal_vector const& lits): p(p), lits(lits) {
@ -1215,7 +1217,7 @@ namespace sat {
lookahead& lh;
public:
lookahead_literal_occs_fun(lookahead& lh): lh(lh) {}
double operator()(literal l) { return lh.literal_occs(l); }
double operator()(literal l) override { return lh.literal_occs(l); }
};
// Ternary clause managagement:
@ -2055,6 +2057,15 @@ namespace sat {
return h;
}
bool lookahead::should_cutoff(unsigned depth) {
return depth > 0 &&
((m_config.m_cube_cutoff == depth_cutoff && depth == m_config.m_cube_depth) ||
(m_config.m_cube_cutoff == freevars_cutoff && m_freevars.size() <= m_init_freevars * m_config.m_cube_freevars) ||
(m_config.m_cube_cutoff == psat_cutoff && psat_heur() >= m_config.m_cube_psat_trigger) ||
(m_config.m_cube_cutoff == adaptive_freevars_cutoff && m_freevars.size() < m_cube_state.m_freevars_threshold) ||
(m_config.m_cube_cutoff == adaptive_psat_cutoff && psat_heur() >= m_cube_state.m_psat_threshold));
}
lbool lookahead::cube(bool_var_vector& vars, literal_vector& lits, unsigned backtrack_level) {
scoped_ext _scoped_ext(*this);
lits.reset();
@ -2100,22 +2111,13 @@ namespace sat {
}
backtrack_level = UINT_MAX;
depth = m_cube_state.m_cube.size();
if ((m_config.m_cube_cutoff == depth_cutoff && depth == m_config.m_cube_depth) ||
(m_config.m_cube_cutoff == freevars_cutoff && m_freevars.size() <= m_init_freevars * m_config.m_cube_freevars) ||
(m_config.m_cube_cutoff == psat_cutoff && psat_heur() >= m_config.m_cube_psat_trigger) ||
(m_config.m_cube_cutoff == adaptive_freevars_cutoff && m_freevars.size() < m_cube_state.m_freevars_threshold) ||
(m_config.m_cube_cutoff == adaptive_psat_cutoff && psat_heur() >= m_cube_state.m_psat_threshold)) {
if (should_cutoff(depth)) {
double dec = (1.0 - pow(m_config.m_cube_fraction, depth));
m_cube_state.m_freevars_threshold *= dec;
m_cube_state.m_psat_threshold *= 2.0 - dec;
set_conflict();
m_cube_state.inc_cutoff();
#if 0
// return cube of all literals, not just the ones in the main cube
lits.append(m_trail.size() - init_trail, m_trail.c_ptr() + init_trail);
#else
lits.append(m_cube_state.m_cube);
#endif
vars.reset();
for (auto v : m_freevars) if (in_reduced_clause(v)) vars.push_back(v);
backtrack(m_cube_state.m_cube, m_cube_state.m_is_decision);
@ -2135,6 +2137,8 @@ namespace sat {
if (lit == null_literal) {
vars.reset();
for (auto v : m_freevars) if (in_reduced_clause(v)) vars.push_back(v);
m_model.reset();
init_model();
return l_true;
}
TRACE("sat", tout << "choose: " << lit << " cube: " << m_cube_state.m_cube << "\n";);

View file

@ -558,6 +558,8 @@ namespace sat {
double psat_heur();
bool should_cutoff(unsigned depth);
public:
lookahead(solver& s) :
m_s(s),

View file

@ -1,45 +0,0 @@
/*++
Copyright (c) 2017 Microsoft Corporation
Module Name:
sat_par.cpp
Abstract:
Utilities for parallel SAT solving.
Author:
Nikolaj Bjorner (nbjorner) 2017-1-29.
Revision History:
--*/
#include "sat/sat_par.h"
namespace sat {
par::par() {}
void par::exchange(literal_vector const& in, unsigned& limit, literal_vector& out) {
#pragma omp critical (par_solver)
{
if (limit < m_units.size()) {
// this might repeat some literals.
out.append(m_units.size() - limit, m_units.c_ptr() + limit);
}
for (unsigned i = 0; i < in.size(); ++i) {
literal lit = in[i];
if (!m_unit_set.contains(lit.index())) {
m_unit_set.insert(lit.index());
m_units.push_back(lit);
}
}
limit = m_units.size();
}
}
};

View file

@ -1,39 +0,0 @@
/*++
Copyright (c) 2017 Microsoft Corporation
Module Name:
sat_par.h
Abstract:
Utilities for parallel SAT solving.
Author:
Nikolaj Bjorner (nbjorner) 2017-1-29.
Revision History:
--*/
#ifndef SAT_PAR_H_
#define SAT_PAR_H_
#include "sat/sat_types.h"
#include "util/hashtable.h"
#include "util/map.h"
namespace sat {
class par {
typedef hashtable<unsigned, u_hash, u_eq> index_set;
literal_vector m_units;
index_set m_unit_set;
public:
par();
void exchange(literal_vector const& in, unsigned& limit, literal_vector& out);
};
};
#endif

View file

@ -15,7 +15,7 @@ def_module_params('sat',
('restart.margin', DOUBLE, 1.1, 'margin between fast and slow restart factors. For ema'),
('restart.emafastglue', DOUBLE, 3e-2, 'ema alpha factor for fast moving average'),
('restart.emaslowglue', DOUBLE, 1e-5, 'ema alpha factor for slow moving average'),
('variable_decay', UINT, 110, 'multiplier (divided by 100) for the VSIDS activity increement'),
('variable_decay', UINT, 110, 'multiplier (divided by 100) for the VSIDS activity increment'),
('inprocess.max', UINT, UINT_MAX, 'maximal number of inprocessing passes'),
('branching.heuristic', SYMBOL, 'vsids', 'branching heuristic vsids, lrb or chb'),
('branching.anti_exploration', BOOL, False, 'apply anti-exploration heuristic for branch selection'),
@ -31,6 +31,7 @@ def_module_params('sat',
('gc.burst', BOOL, False, 'perform eager garbage collection during initialization'),
('gc.defrag', BOOL, True, 'defragment clauses when garbage collecting'),
('simplify.delay', UINT, 0, 'set initial delay of simplification by a conflict count'),
('force_cleanup', BOOL, False, 'force cleanup to remove tautologies and simplify clauses'),
('minimize_lemmas', BOOL, True, 'minimize learned clauses'),
('dyn_sub_res', BOOL, True, 'dynamic subsumption resolution for minimizing learned clauses'),
('core.minimize', BOOL, False, 'minimize computed core'),
@ -52,17 +53,39 @@ def_module_params('sat',
('unit_walk', BOOL, False, 'use unit-walk search instead of CDCL'),
('unit_walk_threads', UINT, 0, 'number of unit-walk search threads to find satisfiable solution'),
('lookahead.cube.cutoff', SYMBOL, 'depth', 'cutoff type used to create lookahead cubes: depth, freevars, psat, adaptive_freevars, adaptive_psat'),
# - depth: the maximal cutoff is fixed to the value of lookahead.cube.depth.
# So if the value is 10, at most 1024 cubes will be generated of length 10.
# - freevars: cutoff based on a variable fraction of lookahead.cube.freevars.
# Cut if the number of current unassigned variables drops below a fraction of number of initial variables.
# - psat: Let psat_heur := (Sum_{clause C} (psat.clause_base ^ {-|C|+1})) / |freevars|^psat.var_exp
# Cut if the value of psat_heur exceeds psat.trigger
# - adaptive_freevars: Cut if the number of current unassigned variables drops below a fraction of free variables
# at the time of the last conflict. The fraction is increased every time the cutoff is created.
# - adative_psat: Cut based on psat_heur in an adaptive way.
('lookahead.cube.fraction', DOUBLE, 0.4, 'adaptive fraction to create lookahead cubes. Used when lookahead.cube.cutoff is adaptive_freevars or adaptive_psat'),
('lookahead.cube.depth', UINT, 1, 'cut-off depth to create cubes. Used when lookahead.cube.cutoff is depth.'),
('lookahead.cube.freevars', DOUBLE, 0.8, 'cube free fariable fraction. Used when lookahead.cube.cutoff is freevars'),
('lookahead.cube.freevars', DOUBLE, 0.8, 'cube free variable fraction. Used when lookahead.cube.cutoff is freevars'),
('lookahead.cube.psat.var_exp', DOUBLE, 1, 'free variable exponent for PSAT cutoff'),
('lookahead.cube.psat.clause_base', DOUBLE, 2, 'clause base for PSAT cutoff'),
('lookahead.cube.psat.trigger', DOUBLE, 5, 'trigger value to create lookahead cubes for PSAT cutoff. Used when lookahead.cube.cutoff is psat'),
('lookahead_search', BOOL, False, 'use lookahead solver'),
('lookahead.preselect', BOOL, False, 'use pre-selection of subset of variables for branching'),
('lookahead_simplify', BOOL, False, 'use lookahead solver during simplification'),
('lookahead.use_learned', BOOL, False, 'use learned clauses when selecting lookahead literal'),
('lookahead_simplify.bca', BOOL, True, 'add learned binary clauses as part of lookahead simplification'),
('lookahead.global_autarky', BOOL, False, 'prefer to branch on variables that occur in clauses that are reduced'),
('lookahead.reward', SYMBOL, 'march_cu', 'select lookahead heuristic: ternary, heule_schur (Heule Schur), heuleu (Heule Unit), unit, or march_cu')))
('lookahead.reward', SYMBOL, 'march_cu', 'select lookahead heuristic: ternary, heule_schur (Heule Schur), heuleu (Heule Unit), unit, or march_cu'))
# reward function used to determine which literal to cube on.
# - ternary: reward function useful for random 3-SAT instances. Used by Heule and Knuth in March.
# - heule_schur: reward function based on "Schur Number 5", Heule, AAAI 2018
# The score of a literal lit is:
# Sum_{C in Clauses | lit in C} 2 ^ (- |C|+1)
# * Sum_{lit' in C | lit' != lit} lit_occs(~lit')
# / | C |
# where lit_occs(lit) is the number of clauses containing lit.
# - heuleu: The score of a literal lit is: Sum_{C in Clauses | lit in C} 2 ^ (-|C| + 1)
# - unit: heule_schur + also counts number of unit clauses.
# - march_cu: default reward function used in a version of March
# Each reward function also comes with its own variant of "mix_diff", which
# is the function for combining reward metrics for the positive and negative variant of a literal.
)

View file

@ -66,18 +66,20 @@ namespace sat {
m_next_simplify = 0;
m_num_checkpoints = 0;
m_simplifications = 0;
m_ext = 0;
m_ext = nullptr;
m_cuber = nullptr;
m_mc.set_solver(this);
}
solver::~solver() {
m_ext = 0;
m_ext = nullptr;
SASSERT(check_invariant());
TRACE("sat", tout << "Delete clauses\n";);
del_clauses(m_clauses);
TRACE("sat", tout << "Delete learned\n";);
del_clauses(m_learned);
dealloc(m_cuber);
m_cuber = nullptr;
}
void solver::del_clauses(clause_vector& clauses) {
@ -1008,14 +1010,39 @@ namespace sat {
}
lbool solver::cube(bool_var_vector& vars, literal_vector& lits, unsigned backtrack_level) {
if (!m_cuber) {
bool is_first = !m_cuber;
if (is_first) {
m_cuber = alloc(lookahead, *this);
}
lbool result = m_cuber->cube(vars, lits, backtrack_level);
m_cuber->update_cube_statistics(m_aux_stats);
if (result == l_false) {
switch (result) {
case l_false:
dealloc(m_cuber);
m_cuber = nullptr;
if (is_first) {
pop_to_base_level();
set_conflict(justification());
}
break;
case l_true: {
lits.reset();
pop_to_base_level();
model const& mdl = m_cuber->get_model();
for (bool_var v = 0; v < mdl.size(); ++v) {
if (value(v) != l_undef) {
continue;
}
literal l(v, false);
if (mdl[v] != l_true) l.neg();
push();
assign_core(l, justification());
}
mk_model();
break;
}
default:
break;
}
return result;
}
@ -1053,7 +1080,7 @@ namespace sat {
init_assumptions(num_lits, lits);
propagate(false);
if (check_inconsistent()) return l_false;
cleanup();
cleanup(m_config.m_force_cleanup);
if (m_config.m_unit_walk) {
return do_unit_walk();
@ -1127,7 +1154,7 @@ namespace sat {
srch.config().set_config(m_config);
srch.import(*this, false);
scoped_rl.push_child(&srch.rlimit());
lbool r = srch.check(num_lits, lits, 0);
lbool r = srch.check(num_lits, lits, nullptr);
m_model = srch.get_model();
// srch.collect_statistics(m_aux_stats);
return r;
@ -1264,7 +1291,7 @@ namespace sat {
if (!canceled) {
rlimit().reset_cancel();
}
set_par(0, 0);
set_par(nullptr, 0);
ls.reset();
uw.reset();
if (finished_id == -1) {
@ -1425,7 +1452,7 @@ namespace sat {
if (should_restart())
return l_undef;
if (at_base_lvl()) {
cleanup(); // cleaner may propagate frozen clauses
cleanup(false); // cleaner may propagate frozen clauses
if (inconsistent()) {
TRACE("sat", tout << "conflict at level 0\n";);
return l_false;
@ -1621,7 +1648,7 @@ namespace sat {
SASSERT(at_base_lvl());
m_cleaner();
m_cleaner(m_config.m_force_cleanup);
CASSERT("sat_simplify_bug", check_invariant());
m_scc();
@ -3662,10 +3689,10 @@ namespace sat {
// Simplification
//
// -----------------------
void solver::cleanup() {
void solver::cleanup(bool force) {
if (!at_base_lvl() || inconsistent())
return;
if (m_cleaner() && m_ext)
if (m_cleaner(force) && m_ext)
m_ext->clauses_modifed();
}

View file

@ -37,7 +37,6 @@ Revision History:
#include "sat/sat_drat.h"
#include "sat/sat_parallel.h"
#include "sat/sat_local_search.h"
#include "sat/sat_par.h"
#include "util/params.h"
#include "util/statistics.h"
#include "util/stopwatch.h"
@ -550,7 +549,7 @@ namespace sat {
//
// -----------------------
public:
void cleanup();
void cleanup(bool force);
void simplify(bool learned = true);
void asymmetric_branching();
unsigned scc_bin();

View file

@ -8,4 +8,6 @@ z3_add_component(sat_solver
core_tactics
sat_tactic
solver
TACTIC_HEADERS
inc_sat_solver.h
)

View file

@ -259,7 +259,7 @@ public:
return m_num_scopes;
}
void assert_expr_core2(expr * t, expr * a) override {
void assert_expr_core2(expr * t, expr * a) override {
if (a) {
m_asmsf.push_back(a);
assert_expr_core(m.mk_implies(a, t));
@ -308,10 +308,19 @@ public:
return nullptr;
}
expr_ref_vector last_cube(bool is_sat) {
expr_ref_vector result(m);
result.push_back(is_sat ? m.mk_true() : m.mk_false());
return result;
}
expr_ref_vector cube(expr_ref_vector& vs, unsigned backtrack_level) override {
if (!is_internalized()) {
lbool r = internalize_formulas();
if (r != l_true) return expr_ref_vector(m);
if (r != l_true) {
IF_VERBOSE(0, verbose_stream() << "internalize produced " << r << "\n");
return expr_ref_vector(m);
}
}
convert_internalized();
obj_hashtable<expr> _vs;
@ -323,14 +332,18 @@ public:
}
sat::literal_vector lits;
lbool result = m_solver.cube(vars, lits, backtrack_level);
if (result == l_false || lits.empty()) {
expr_ref_vector result(m);
result.push_back(m.mk_false());
return result;
switch (result) {
case l_true:
return last_cube(true);
case l_false:
return last_cube(false);
default:
break;
}
if (result == l_true) {
if (lits.empty()) {
set_reason_unknown(m_solver.get_reason_unknown());
return expr_ref_vector(m);
}
}
expr_ref_vector fmls(m);
expr_ref_vector lit2expr(m);
lit2expr.resize(m_solver.num_vars() * 2);
@ -473,6 +486,7 @@ public:
}
void convert_internalized() {
m_solver.pop_to_base_level();
if (!is_internalized() && m_fmls_head > 0) {
internalize_formulas();
}

View file

@ -28,6 +28,9 @@ solver* mk_inc_sat_solver(ast_manager& m, params_ref const& p, bool incremental_
tactic* mk_psat_tactic(ast_manager& m, params_ref const& p);
/*
ADD_TACTIC('psat', '(try to) solve goal using a parallel SAT solver.', 'mk_psat_tactic(m, p)')
*/
void inc_sat_display(std::ostream& out, solver& s, unsigned sz, expr*const* soft, rational const* _weights);

View file

@ -72,7 +72,7 @@ struct goal2sat::imp {
imp(ast_manager & _m, params_ref const & p, sat::solver & s, atom2bool_var & map, dep2asm_map& dep2asm, bool default_external):
m(_m),
pb(m),
m_ext(0),
m_ext(nullptr),
m_solver(s),
m_map(map),
m_dep2asm(dep2asm),
@ -452,7 +452,15 @@ struct goal2sat::imp {
unsigned sz = m_result_stack.size();
if (root) {
m_result_stack.reset();
m_ext->add_pb_ge(sat::null_bool_var, wlits, k.get_unsigned());
unsigned k1 = k.get_unsigned();
if (sign) {
k1 = 1 - k1;
for (wliteral& wl : wlits) {
wl.second.neg();
k1 += wl.first;
}
}
m_ext->add_pb_ge(sat::null_bool_var, wlits, k1);
}
else {
sat::bool_var v = m_solver.mk_var(true);
@ -477,7 +485,15 @@ struct goal2sat::imp {
unsigned sz = m_result_stack.size();
if (root) {
m_result_stack.reset();
m_ext->add_pb_ge(sat::null_bool_var, wlits, k.get_unsigned());
unsigned k1 = k.get_unsigned();
if (sign) {
k1 = 1 - k1;
for (wliteral& wl : wlits) {
wl.second.neg();
k1 += wl.first;
}
}
m_ext->add_pb_ge(sat::null_bool_var, wlits, k1);
}
else {
sat::bool_var v = m_solver.mk_var(true);
@ -1048,7 +1064,7 @@ void sat2goal::mc::insert(sat::bool_var v, app * atom, bool aux) {
expr_ref sat2goal::mc::lit2expr(sat::literal l) {
if (!m_var2expr.get(l.var())) {
app* aux = m.mk_fresh_const(0, m.mk_bool_sort());
app* aux = m.mk_fresh_const(nullptr, m.mk_bool_sort());
m_var2expr.set(l.var(), aux);
if (!m_gmc) m_gmc = alloc(generic_model_converter, m, "sat2goal");
m_gmc->hide(aux->get_decl());
@ -1092,7 +1108,7 @@ struct sat2goal::imp {
SASSERT(m_lit2expr.get((~l).index()) == 0);
app* aux = mc ? mc->var2expr(l.var()) : nullptr;
if (!aux) {
aux = m.mk_fresh_const(0, m.mk_bool_sort());
aux = m.mk_fresh_const(nullptr, m.mk_bool_sort());
if (mc) {
mc->insert(l.var(), aux, true);
}

View file

@ -86,7 +86,7 @@ public:
public:
mc(ast_manager& m);
virtual ~mc() {}
~mc() override {}
// flush model converter from SAT solver to this structure.
void flush_smc(sat::solver& s, atom2bool_var const& map);
void operator()(model_ref& md) override;