Fix for bug 734
[cvc5.git] / src / theory / bv / theory_bv.cpp
1 /********************* */
2 /*! \file theory_bv.cpp
3 ** \verbatim
4 ** Top contributors (to current version):
5 ** Liana Hadarean, Tim King, Dejan Jovanovic
6 ** This file is part of the CVC4 project.
7 ** Copyright (c) 2009-2016 by the authors listed in the file AUTHORS
8 ** in the top-level source directory) and their institutional affiliations.
9 ** All rights reserved. See the file COPYING in the top-level source
10 ** directory for licensing information.\endverbatim
11 **
12 ** [[ Add lengthier description here ]]
13 ** \todo document this file
14 **/
15
16 #include "theory/bv/theory_bv.h"
17
18 #include "options/bv_options.h"
19 #include "options/smt_options.h"
20 #include "smt/smt_statistics_registry.h"
21 #include "theory/bv/abstraction.h"
22 #include "theory/bv/bv_eager_solver.h"
23 #include "theory/bv/bv_subtheory_algebraic.h"
24 #include "theory/bv/bv_subtheory_bitblast.h"
25 #include "theory/bv/bv_subtheory_core.h"
26 #include "theory/bv/bv_subtheory_inequality.h"
27 #include "theory/bv/slicer.h"
28 #include "theory/bv/theory_bv_rewrite_rules_normalization.h"
29 #include "theory/bv/theory_bv_rewrite_rules_simplification.h"
30 #include "theory/bv/theory_bv_rewriter.h"
31 #include "theory/bv/theory_bv_utils.h"
32 #include "theory/theory_model.h"
33 #include "proof/theory_proof.h"
34 #include "proof/proof_manager.h"
35 #include "theory/valuation.h"
36
37 using namespace CVC4::context;
38 using namespace CVC4::theory::bv::utils;
39 using namespace std;
40
41 namespace CVC4 {
42 namespace theory {
43 namespace bv {
44
45 TheoryBV::TheoryBV(context::Context* c, context::UserContext* u,
46 OutputChannel& out, Valuation valuation,
47 const LogicInfo& logicInfo, std::string name)
48 : Theory(THEORY_BV, c, u, out, valuation, logicInfo, name),
49 d_context(c),
50 d_alreadyPropagatedSet(c),
51 d_sharedTermsSet(c),
52 d_subtheories(),
53 d_subtheoryMap(),
54 d_statistics(getFullInstanceName()),
55 d_staticLearnCache(),
56 d_BVDivByZero(),
57 d_BVRemByZero(),
58 d_funcToArgs(),
59 d_funcToSkolem(u),
60 d_lemmasAdded(c, false),
61 d_conflict(c, false),
62 d_invalidateModelCache(c, true),
63 d_literalsToPropagate(c),
64 d_literalsToPropagateIndex(c, 0),
65 d_propagatedBy(c),
66 d_eagerSolver(NULL),
67 d_abstractionModule(new AbstractionModule(getFullInstanceName())),
68 d_isCoreTheory(false),
69 d_calledPreregister(false),
70 d_needsLastCallCheck(false),
71 d_extf_range_infer(u),
72 d_extf_collapse_infer(u)
73 {
74 d_extt = new ExtTheory( this );
75 d_extt->addFunctionKind( kind::BITVECTOR_TO_NAT );
76 d_extt->addFunctionKind( kind::INT_TO_BITVECTOR );
77
78 if (options::bitblastMode() == theory::bv::BITBLAST_MODE_EAGER) {
79 d_eagerSolver = new EagerBitblastSolver(this);
80 return;
81 }
82
83 if (options::bitvectorEqualitySolver()) {
84 SubtheorySolver* core_solver = new CoreSolver(c, this);
85 d_subtheories.push_back(core_solver);
86 d_subtheoryMap[SUB_CORE] = core_solver;
87 }
88
89 if (options::bitvectorInequalitySolver()) {
90 SubtheorySolver* ineq_solver = new InequalitySolver(c, u, this);
91 d_subtheories.push_back(ineq_solver);
92 d_subtheoryMap[SUB_INEQUALITY] = ineq_solver;
93 }
94
95 if (options::bitvectorAlgebraicSolver()) {
96 SubtheorySolver* alg_solver = new AlgebraicSolver(c, this);
97 d_subtheories.push_back(alg_solver);
98 d_subtheoryMap[SUB_ALGEBRAIC] = alg_solver;
99 }
100
101 BitblastSolver* bb_solver = new BitblastSolver(c, this);
102 if (options::bvAbstraction()) {
103 bb_solver->setAbstraction(d_abstractionModule);
104 }
105 d_subtheories.push_back(bb_solver);
106 d_subtheoryMap[SUB_BITBLAST] = bb_solver;
107 }
108
109 TheoryBV::~TheoryBV() {
110 if (d_eagerSolver) {
111 delete d_eagerSolver;
112 }
113 for (unsigned i = 0; i < d_subtheories.size(); ++i) {
114 delete d_subtheories[i];
115 }
116 delete d_abstractionModule;
117 delete d_extt;
118 }
119
120 void TheoryBV::setMasterEqualityEngine(eq::EqualityEngine* eq) {
121 if (options::bitblastMode() == theory::bv::BITBLAST_MODE_EAGER) {
122 return;
123 }
124 if (options::bitvectorEqualitySolver()) {
125 dynamic_cast<CoreSolver*>(d_subtheoryMap[SUB_CORE])->setMasterEqualityEngine(eq);
126 }
127 }
128
129 void TheoryBV::spendResource(unsigned ammount) throw(UnsafeInterruptException) {
130 getOutputChannel().spendResource(ammount);
131 }
132
133 TheoryBV::Statistics::Statistics(const std::string &name):
134 d_avgConflictSize(name + "theory::bv::AvgBVConflictSize"),
135 d_solveSubstitutions(name + "theory::bv::NumberOfSolveSubstitutions", 0),
136 d_solveTimer(name + "theory::bv::solveTimer"),
137 d_numCallsToCheckFullEffort(name + "theory::bv::NumberOfFullCheckCalls", 0),
138 d_numCallsToCheckStandardEffort(name + "theory::bv::NumberOfStandardCheckCalls", 0),
139 d_weightComputationTimer(name + "theory::bv::weightComputationTimer"),
140 d_numMultSlice(name + "theory::bv::NumMultSliceApplied", 0)
141 {
142 smtStatisticsRegistry()->registerStat(&d_avgConflictSize);
143 smtStatisticsRegistry()->registerStat(&d_solveSubstitutions);
144 smtStatisticsRegistry()->registerStat(&d_solveTimer);
145 smtStatisticsRegistry()->registerStat(&d_numCallsToCheckFullEffort);
146 smtStatisticsRegistry()->registerStat(&d_numCallsToCheckStandardEffort);
147 smtStatisticsRegistry()->registerStat(&d_weightComputationTimer);
148 smtStatisticsRegistry()->registerStat(&d_numMultSlice);
149 }
150
151 TheoryBV::Statistics::~Statistics() {
152 smtStatisticsRegistry()->unregisterStat(&d_avgConflictSize);
153 smtStatisticsRegistry()->unregisterStat(&d_solveSubstitutions);
154 smtStatisticsRegistry()->unregisterStat(&d_solveTimer);
155 smtStatisticsRegistry()->unregisterStat(&d_numCallsToCheckFullEffort);
156 smtStatisticsRegistry()->unregisterStat(&d_numCallsToCheckStandardEffort);
157 smtStatisticsRegistry()->unregisterStat(&d_weightComputationTimer);
158 smtStatisticsRegistry()->unregisterStat(&d_numMultSlice);
159 }
160
161 Node TheoryBV::getBVDivByZero(Kind k, unsigned width) {
162 NodeManager* nm = NodeManager::currentNM();
163 if (k == kind::BITVECTOR_UDIV) {
164 if (d_BVDivByZero.find(width) == d_BVDivByZero.end()) {
165 // lazily create the function symbols
166 ostringstream os;
167 os << "BVUDivByZero_" << width;
168 Node divByZero = nm->mkSkolem(os.str(),
169 nm->mkFunctionType(nm->mkBitVectorType(width), nm->mkBitVectorType(width)),
170 "partial bvudiv", NodeManager::SKOLEM_EXACT_NAME);
171 d_BVDivByZero[width] = divByZero;
172 }
173 return d_BVDivByZero[width];
174 }
175 else if (k == kind::BITVECTOR_UREM) {
176 if (d_BVRemByZero.find(width) == d_BVRemByZero.end()) {
177 ostringstream os;
178 os << "BVURemByZero_" << width;
179 Node divByZero = nm->mkSkolem(os.str(),
180 nm->mkFunctionType(nm->mkBitVectorType(width), nm->mkBitVectorType(width)),
181 "partial bvurem", NodeManager::SKOLEM_EXACT_NAME);
182 d_BVRemByZero[width] = divByZero;
183 }
184 return d_BVRemByZero[width];
185 }
186
187 Unreachable();
188 }
189
190
191 void TheoryBV::collectFunctionSymbols(TNode term, TNodeSet& seen) {
192 if (seen.find(term) != seen.end())
193 return;
194 if (term.getKind() == kind::APPLY_UF) {
195 TNode func = term.getOperator();
196 storeFunction(func, term);
197 } else if (term.getKind() == kind::SELECT) {
198 TNode func = term[0];
199 storeFunction(func, term);
200 } else if (term.getKind() == kind::STORE) {
201 AlwaysAssert(false, "Cannot use eager bitblasting on QF_ABV formula with stores");
202 }
203 for (unsigned i = 0; i < term.getNumChildren(); ++i) {
204 collectFunctionSymbols(term[i], seen);
205 }
206 seen.insert(term);
207 }
208
209 void TheoryBV::storeFunction(TNode func, TNode term) {
210 if (d_funcToArgs.find(func) == d_funcToArgs.end()) {
211 d_funcToArgs.insert(make_pair(func, NodeSet()));
212 }
213 NodeSet& set = d_funcToArgs[func];
214 if (set.find(term) == set.end()) {
215 set.insert(term);
216 Node skolem = utils::mkVar(utils::getSize(term));
217 d_funcToSkolem.addSubstitution(term, skolem);
218 }
219 }
220
221 void TheoryBV::mkAckermanizationAsssertions(std::vector<Node>& assertions) {
222 Debug("bv-ackermanize") << "TheoryBV::mkAckermanizationAsssertions\n";
223
224 Assert(options::bitblastMode() == theory::bv::BITBLAST_MODE_EAGER);
225 AlwaysAssert(!options::incrementalSolving());
226 TNodeSet seen;
227 for (unsigned i = 0; i < assertions.size(); ++i) {
228 collectFunctionSymbols(assertions[i], seen);
229 }
230
231 FunctionToArgs::const_iterator it = d_funcToArgs.begin();
232 NodeManager* nm = NodeManager::currentNM();
233 for (; it!= d_funcToArgs.end(); ++it) {
234 TNode func = it->first;
235 const NodeSet& args = it->second;
236 NodeSet::const_iterator it1 = args.begin();
237 for ( ; it1 != args.end(); ++it1) {
238 for(NodeSet::const_iterator it2 = it1; it2 != args.end(); ++it2) {
239 TNode args1 = *it1;
240 TNode args2 = *it2;
241 Node args_eq;
242
243 if (args1.getKind() == kind::APPLY_UF) {
244 AlwaysAssert (args1.getKind() == kind::APPLY_UF &&
245 args1.getOperator() == func);
246 AlwaysAssert (args2.getKind() == kind::APPLY_UF &&
247 args2.getOperator() == func);
248 AlwaysAssert (args1.getNumChildren() == args2.getNumChildren());
249
250 std::vector<Node> eqs(args1.getNumChildren());
251
252 for (unsigned i = 0; i < args1.getNumChildren(); ++i) {
253 eqs[i] = nm->mkNode(kind::EQUAL, args1[i], args2[i]);
254 }
255 args_eq = eqs.size() == 1 ? eqs[0] : nm->mkNode(kind::AND, eqs);
256 } else {
257 AlwaysAssert (args1.getKind() == kind::SELECT &&
258 args1[0] == func);
259 AlwaysAssert (args2.getKind() == kind::SELECT &&
260 args2[0] == func);
261 AlwaysAssert (args1.getNumChildren() == 2);
262 AlwaysAssert (args2.getNumChildren() == 2);
263 args_eq = nm->mkNode(kind::EQUAL, args1[1], args2[1]);
264 }
265 Node func_eq = nm->mkNode(kind::EQUAL, args1, args2);
266 Node lemma = nm->mkNode(kind::IMPLIES, args_eq, func_eq);
267 assertions.push_back(lemma);
268 }
269 }
270 }
271
272 // replace applications of UF by skolems (FIXME for model building)
273 for(unsigned i = 0; i < assertions.size(); ++i) {
274 assertions[i] = d_funcToSkolem.apply(assertions[i]);
275 }
276 }
277
278 Node TheoryBV::expandDefinition(LogicRequest &logicRequest, Node node) {
279 Debug("bitvector-expandDefinition") << "TheoryBV::expandDefinition(" << node << ")" << std::endl;
280
281 switch (node.getKind()) {
282 case kind::BITVECTOR_SDIV:
283 case kind::BITVECTOR_SREM:
284 case kind::BITVECTOR_SMOD:
285 return TheoryBVRewriter::eliminateBVSDiv(node);
286 break;
287
288 case kind::BITVECTOR_UDIV:
289 case kind::BITVECTOR_UREM: {
290 NodeManager* nm = NodeManager::currentNM();
291 unsigned width = node.getType().getBitVectorSize();
292
293 if (options::bitvectorDivByZeroConst()) {
294 Kind kind = node.getKind() == kind::BITVECTOR_UDIV ? kind::BITVECTOR_UDIV_TOTAL : kind::BITVECTOR_UREM_TOTAL;
295 return nm->mkNode(kind, node[0], node[1]);
296 }
297
298 TNode num = node[0], den = node[1];
299 Node den_eq_0 = nm->mkNode(kind::EQUAL, den, nm->mkConst(BitVector(width, Integer(0))));
300 Node divTotalNumDen = nm->mkNode(node.getKind() == kind::BITVECTOR_UDIV ? kind::BITVECTOR_UDIV_TOTAL :
301 kind::BITVECTOR_UREM_TOTAL, num, den);
302
303 // if (options::bitblastMode() == theory::bv::BITBLAST_MODE_EAGER) {
304 // // Ackermanize UF if using eager bit-blasting
305 // Node ackerman_var = nm->mkNode(node.getKind() == kind::BITVECTOR_UDIV ? kind::BITVECTOR_ACKERMANIZE_UDIV : kind::BITVECTOR_ACKERMANIZE_UREM, num);
306 // node = nm->mkNode(kind::ITE, den_eq_0, ackerman_var, divTotalNumDen);
307 // return node;
308 // } else {
309 Node divByZero = getBVDivByZero(node.getKind(), width);
310 Node divByZeroNum = nm->mkNode(kind::APPLY_UF, divByZero, num);
311 node = nm->mkNode(kind::ITE, den_eq_0, divByZeroNum, divTotalNumDen);
312 logicRequest.widenLogic(THEORY_UF);
313 return node;
314 //}
315 }
316 break;
317
318 default:
319 return node;
320 break;
321 }
322
323 Unreachable();
324 }
325
326
327 void TheoryBV::preRegisterTerm(TNode node) {
328 d_calledPreregister = true;
329 Debug("bitvector-preregister") << "TheoryBV::preRegister(" << node << ")" << std::endl;
330
331 if (options::bitblastMode() == theory::bv::BITBLAST_MODE_EAGER) {
332 // the aig bit-blaster option is set heuristically
333 // if bv abstraction is not used
334 if (!d_eagerSolver->isInitialized()) {
335 d_eagerSolver->initialize();
336 }
337
338 if (node.getKind() == kind::BITVECTOR_EAGER_ATOM) {
339 Node formula = node[0];
340 d_eagerSolver->assertFormula(formula);
341 }
342 // nothing to do for the other terms
343 return;
344 }
345
346 for (unsigned i = 0; i < d_subtheories.size(); ++i) {
347 d_subtheories[i]->preRegister(node);
348 }
349
350 // AJR : equality solver currently registers all terms to ExtTheory, if we want a lazy reduction without the bv equality solver, need to call this
351 //getExtTheory()->registerTermRec( node );
352 }
353
354 void TheoryBV::sendConflict() {
355 Assert(d_conflict);
356 if (d_conflictNode.isNull()) {
357 return;
358 } else {
359 Debug("bitvector") << indent() << "TheoryBV::check(): conflict " << d_conflictNode << std::endl;
360 d_out->conflict(d_conflictNode);
361 d_statistics.d_avgConflictSize.addEntry(d_conflictNode.getNumChildren());
362 d_conflictNode = Node::null();
363 }
364 }
365
366 void TheoryBV::checkForLemma(TNode fact) {
367 if (fact.getKind() == kind::EQUAL) {
368 if (fact[0].getKind() == kind::BITVECTOR_UREM_TOTAL) {
369 TNode urem = fact[0];
370 TNode result = fact[1];
371 TNode divisor = urem[1];
372 Node result_ult_div = mkNode(kind::BITVECTOR_ULT, result, divisor);
373 Node divisor_eq_0 = mkNode(kind::EQUAL,
374 divisor,
375 mkConst(BitVector(getSize(divisor), 0u)));
376 Node split = utils::mkNode(kind::OR, divisor_eq_0, mkNode(kind::NOT, fact), result_ult_div);
377 lemma(split);
378 }
379 if (fact[1].getKind() == kind::BITVECTOR_UREM_TOTAL) {
380 TNode urem = fact[1];
381 TNode result = fact[0];
382 TNode divisor = urem[1];
383 Node result_ult_div = mkNode(kind::BITVECTOR_ULT, result, divisor);
384 Node divisor_eq_0 = mkNode(kind::EQUAL,
385 divisor,
386 mkConst(BitVector(getSize(divisor), 0u)));
387 Node split = utils::mkNode(kind::OR, divisor_eq_0, mkNode(kind::NOT, fact), result_ult_div);
388 lemma(split);
389 }
390 }
391 }
392
393 void TheoryBV::check(Effort e)
394 {
395 if (done() && e<Theory::EFFORT_FULL) {
396 return;
397 }
398
399 //last call : do reductions on extended bitvector functions
400 if( e==Theory::EFFORT_LAST_CALL ){
401 std::vector< Node > nred;
402 getExtTheory()->getActive( nred );
403 doExtfReductions( nred );
404 return;
405 }
406
407 TimerStat::CodeTimer checkTimer(d_checkTime);
408 Debug("bitvector") << "TheoryBV::check(" << e << ")" << std::endl;
409 TimerStat::CodeTimer codeTimer(d_statistics.d_solveTimer);
410 // we may be getting new assertions so the model cache may not be sound
411 d_invalidateModelCache.set(true);
412 // if we are using the eager solver
413 if (options::bitblastMode() == theory::bv::BITBLAST_MODE_EAGER) {
414 // this can only happen on an empty benchmark
415 if (!d_eagerSolver->isInitialized()) {
416 d_eagerSolver->initialize();
417 }
418 if (!Theory::fullEffort(e))
419 return;
420
421 std::vector<TNode> assertions;
422 while (!done()) {
423 TNode fact = get().assertion;
424 Assert (fact.getKind() == kind::BITVECTOR_EAGER_ATOM);
425 assertions.push_back(fact);
426 }
427 Assert (d_eagerSolver->hasAssertions(assertions));
428
429 bool ok = d_eagerSolver->checkSat();
430 if (!ok) {
431 if (assertions.size() == 1) {
432 d_out->conflict(assertions[0]);
433 return;
434 }
435 Node conflict = NodeManager::currentNM()->mkNode(kind::AND, assertions);
436 d_out->conflict(conflict);
437 return;
438 }
439 return;
440 }
441
442
443 if (Theory::fullEffort(e)) {
444 ++(d_statistics.d_numCallsToCheckFullEffort);
445 } else {
446 ++(d_statistics.d_numCallsToCheckStandardEffort);
447 }
448 // if we are already in conflict just return the conflict
449 if (inConflict()) {
450 sendConflict();
451 return;
452 }
453
454 while (!done()) {
455 TNode fact = get().assertion;
456
457 checkForLemma(fact);
458
459 for (unsigned i = 0; i < d_subtheories.size(); ++i) {
460 d_subtheories[i]->assertFact(fact);
461 }
462 }
463
464 bool ok = true;
465 bool complete = false;
466 for (unsigned i = 0; i < d_subtheories.size(); ++i) {
467 Assert (!inConflict());
468 ok = d_subtheories[i]->check(e);
469 complete = d_subtheories[i]->isComplete();
470
471 if (!ok) {
472 // if we are in a conflict no need to check with other theories
473 Assert (inConflict());
474 sendConflict();
475 return;
476 }
477 if (complete) {
478 // if the last subtheory was complete we stop
479 break;
480 }
481 }
482
483 //check extended functions
484 if (Theory::fullEffort(e)) {
485 //do inferences (adds external lemmas) TODO: this can be improved to add internal inferences
486 std::vector< Node > nred;
487 if( getExtTheory()->doInferences( 0, nred ) ){
488 return;
489 }
490 d_needsLastCallCheck = false;
491 if( !nred.empty() ){
492 //other inferences involving bv2nat, int2bv
493 if( options::bvAlgExtf() ){
494 if( doExtfInferences( nred ) ){
495 return;
496 }
497 }
498 if( !options::bvLazyReduceExtf() ){
499 if( doExtfReductions( nred ) ){
500 return;
501 }
502 }else{
503 d_needsLastCallCheck = true;
504 }
505 }
506 }
507 }
508
509 bool TheoryBV::doExtfInferences( std::vector< Node >& terms ) {
510 bool sentLemma = false;
511 eq::EqualityEngine * ee = getEqualityEngine();
512 std::map< Node, Node > op_map;
513 for( unsigned j=0; j<terms.size(); j++ ){
514 TNode n = terms[j];
515 Assert( n.getKind()==kind::BITVECTOR_TO_NAT || kind::INT_TO_BITVECTOR );
516 if( n.getKind()==kind::BITVECTOR_TO_NAT ){
517 //range lemmas
518 if( d_extf_range_infer.find( n )==d_extf_range_infer.end() ){
519 d_extf_range_infer.insert( n );
520 unsigned bvs = n[0].getType().getBitVectorSize();
521 Node min = NodeManager::currentNM()->mkConst( Rational( 0 ) );
522 Node max = NodeManager::currentNM()->mkConst( Rational( Integer(1).multiplyByPow2(bvs) ) );
523 Node lem = NodeManager::currentNM()->mkNode( kind::AND, NodeManager::currentNM()->mkNode( kind::GEQ, n, min ), NodeManager::currentNM()->mkNode( kind::LT, n, max ) );
524 Trace("bv-extf-lemma") << "BV extf lemma (range) : " << lem << std::endl;
525 d_out->lemma( lem );
526 sentLemma = true;
527 }
528 }
529 Node r = ( ee && ee->hasTerm( n[0] ) ) ? ee->getRepresentative( n[0] ) : n[0];
530 op_map[r] = n;
531 }
532 for( unsigned j=0; j<terms.size(); j++ ){
533 TNode n = terms[j];
534 Node r = ( ee && ee->hasTerm( n[0] ) ) ? ee->getRepresentative( n ) : n;
535 std::map< Node, Node >::iterator it = op_map.find( r );
536 if( it!=op_map.end() ){
537 Node parent = it->second;
538 //Node cterm = parent[0]==n ? parent : NodeManager::currentNM()->mkNode( parent.getOperator(), n );
539 Node cterm = parent[0].eqNode( n );
540 Trace("bv-extf-lemma-debug") << "BV extf collapse based on : " << cterm << std::endl;
541 if( d_extf_collapse_infer.find( cterm )==d_extf_collapse_infer.end() ){
542 d_extf_collapse_infer.insert( cterm );
543
544 Node t = n[0];
545 if( n.getKind()==kind::INT_TO_BITVECTOR ){
546 Assert( t.getType().isInteger() );
547 //congruent modulo 2^( bv width )
548 unsigned bvs = n.getType().getBitVectorSize();
549 Node coeff = NodeManager::currentNM()->mkConst( Rational( Integer(1).multiplyByPow2(bvs) ) );
550 Node k = NodeManager::currentNM()->mkSkolem( "int_bv_cong", t.getType(), "for int2bv/bv2nat congruence" );
551 t = NodeManager::currentNM()->mkNode( kind::PLUS, t, NodeManager::currentNM()->mkNode( kind::MULT, coeff, k ) );
552 }
553 Node lem = parent.eqNode( t );
554
555 if( parent[0]!=n ){
556 Assert( ee->areEqual( parent[0], n ) );
557 lem = NodeManager::currentNM()->mkNode( kind::IMPLIES, parent[0].eqNode( n ), lem );
558 }
559 Trace("bv-extf-lemma") << "BV extf lemma (collapse) : " << lem << std::endl;
560 d_out->lemma( lem );
561 sentLemma = true;
562 }
563 }
564 }
565 return sentLemma;
566 }
567
568 bool TheoryBV::doExtfReductions( std::vector< Node >& terms ) {
569 std::vector< Node > nredr;
570 if( getExtTheory()->doReductions( 0, terms, nredr ) ){
571 return true;
572 }
573 Assert( nredr.empty() );
574 return false;
575 }
576
577 bool TheoryBV::needsCheckLastEffort() {
578 return d_needsLastCallCheck;
579 }
580
581 void TheoryBV::collectModelInfo( TheoryModel* m, bool fullModel ){
582 Assert(!inConflict());
583 if (options::bitblastMode() == theory::bv::BITBLAST_MODE_EAGER) {
584 d_eagerSolver->collectModelInfo(m, fullModel);
585 }
586 for (unsigned i = 0; i < d_subtheories.size(); ++i) {
587 if (d_subtheories[i]->isComplete()) {
588 d_subtheories[i]->collectModelInfo(m, fullModel);
589 return;
590 }
591 }
592 }
593
594 Node TheoryBV::getModelValue(TNode var) {
595 Assert(!inConflict());
596 for (unsigned i = 0; i < d_subtheories.size(); ++i) {
597 if (d_subtheories[i]->isComplete()) {
598 return d_subtheories[i]->getModelValue(var);
599 }
600 }
601 Unreachable();
602 }
603
604 void TheoryBV::propagate(Effort e) {
605 Debug("bitvector") << indent() << "TheoryBV::propagate()" << std::endl;
606 if (options::bitblastMode() == theory::bv::BITBLAST_MODE_EAGER) {
607 return;
608 }
609
610 if (inConflict()) {
611 return;
612 }
613
614 // go through stored propagations
615 bool ok = true;
616 for (; d_literalsToPropagateIndex < d_literalsToPropagate.size() && ok; d_literalsToPropagateIndex = d_literalsToPropagateIndex + 1) {
617 TNode literal = d_literalsToPropagate[d_literalsToPropagateIndex];
618 // temporary fix for incremental bit-blasting
619 if (d_valuation.isSatLiteral(literal)) {
620 Debug("bitvector::propagate") << "TheoryBV:: propagating " << literal <<"\n";
621 ok = d_out->propagate(literal);
622 }
623 }
624
625 if (!ok) {
626 Debug("bitvector::propagate") << indent() << "TheoryBV::propagate(): conflict from theory engine" << std::endl;
627 setConflict();
628 }
629 }
630
631
632 eq::EqualityEngine * TheoryBV::getEqualityEngine() {
633 CoreSolver* core = (CoreSolver*)d_subtheoryMap[SUB_CORE];
634 if( core ){
635 return core->getEqualityEngine();
636 }else{
637 return NULL;
638 }
639 }
640
641 bool TheoryBV::getCurrentSubstitution( int effort, std::vector< Node >& vars, std::vector< Node >& subs, std::map< Node, std::vector< Node > >& exp ) {
642 eq::EqualityEngine * ee = getEqualityEngine();
643 if( ee ){
644 //get the constant equivalence classes
645 bool retVal = false;
646 for( unsigned i=0; i<vars.size(); i++ ){
647 Node n = vars[i];
648 if( ee->hasTerm( n ) ){
649 Node nr = ee->getRepresentative( n );
650 if( nr.isConst() ){
651 subs.push_back( nr );
652 exp[n].push_back( n.eqNode( nr ) );
653 retVal = true;
654 }else{
655 subs.push_back( n );
656 }
657 }else{
658 subs.push_back( n );
659 }
660 }
661 //return true if the substitution is non-trivial
662 return retVal;
663 }
664 return false;
665 }
666
667 int TheoryBV::getReduction( int effort, Node n, Node& nr ) {
668 Trace("bv-ext") << "TheoryBV::checkExt : non-reduced : " << n << std::endl;
669 if( n.getKind()==kind::BITVECTOR_TO_NAT ){
670 //taken from rewrite code
671 const unsigned size = utils::getSize(n[0]);
672 NodeManager* const nm = NodeManager::currentNM();
673 const Node z = nm->mkConst(Rational(0));
674 const Node bvone = nm->mkConst(BitVector(1u, 1u));
675 NodeBuilder<> result(kind::PLUS);
676 Integer i = 1;
677 for(unsigned bit = 0; bit < size; ++bit, i *= 2) {
678 Node cond = nm->mkNode(kind::EQUAL, nm->mkNode(nm->mkConst(BitVectorExtract(bit, bit)), n[0]), bvone);
679 result << nm->mkNode(kind::ITE, cond, nm->mkConst(Rational(i)), z);
680 }
681 nr = Node(result);
682 return -1;
683 }else if( n.getKind()==kind::INT_TO_BITVECTOR ){
684 //taken from rewrite code
685 const unsigned size = n.getOperator().getConst<IntToBitVector>().size;
686 NodeManager* const nm = NodeManager::currentNM();
687 const Node bvzero = nm->mkConst(BitVector(1u, 0u));
688 const Node bvone = nm->mkConst(BitVector(1u, 1u));
689 std::vector<Node> v;
690 Integer i = 2;
691 while(v.size() < size) {
692 Node cond = nm->mkNode(kind::GEQ, nm->mkNode(kind::INTS_MODULUS_TOTAL, n[0], nm->mkConst(Rational(i))), nm->mkConst(Rational(i, 2)));
693 cond = Rewriter::rewrite( cond );
694 v.push_back(nm->mkNode(kind::ITE, cond, bvone, bvzero));
695 i *= 2;
696 }
697 NodeBuilder<> result(kind::BITVECTOR_CONCAT);
698 result.append(v.rbegin(), v.rend());
699 nr = Node(result);
700 return -1;
701 }
702 return 0;
703 }
704
705 Theory::PPAssertStatus TheoryBV::ppAssert(TNode in, SubstitutionMap& outSubstitutions) {
706 switch(in.getKind()) {
707 case kind::EQUAL:
708 {
709 if (in[0].isVar() && !in[1].hasSubterm(in[0])) {
710 ++(d_statistics.d_solveSubstitutions);
711 outSubstitutions.addSubstitution(in[0], in[1]);
712 return PP_ASSERT_STATUS_SOLVED;
713 }
714 if (in[1].isVar() && !in[0].hasSubterm(in[1])) {
715 ++(d_statistics.d_solveSubstitutions);
716 outSubstitutions.addSubstitution(in[1], in[0]);
717 return PP_ASSERT_STATUS_SOLVED;
718 }
719 Node node = Rewriter::rewrite(in);
720 if ((node[0].getKind() == kind::BITVECTOR_EXTRACT && node[1].isConst()) ||
721 (node[1].getKind() == kind::BITVECTOR_EXTRACT && node[0].isConst())) {
722 Node extract = node[0].isConst() ? node[1] : node[0];
723 if (extract[0].getKind() == kind::VARIABLE) {
724 Node c = node[0].isConst() ? node[0] : node[1];
725
726 unsigned high = utils::getExtractHigh(extract);
727 unsigned low = utils::getExtractLow(extract);
728 unsigned var_bitwidth = utils::getSize(extract[0]);
729 std::vector<Node> children;
730
731 if (low == 0) {
732 Assert (high != var_bitwidth - 1);
733 unsigned skolem_size = var_bitwidth - high - 1;
734 Node skolem = utils::mkVar(skolem_size);
735 children.push_back(skolem);
736 children.push_back(c);
737 } else if (high == var_bitwidth - 1) {
738 unsigned skolem_size = low;
739 Node skolem = utils::mkVar(skolem_size);
740 children.push_back(c);
741 children.push_back(skolem);
742 } else {
743 unsigned skolem1_size = low;
744 unsigned skolem2_size = var_bitwidth - high - 1;
745 Node skolem1 = utils::mkVar(skolem1_size);
746 Node skolem2 = utils::mkVar(skolem2_size);
747 children.push_back(skolem2);
748 children.push_back(c);
749 children.push_back(skolem1);
750 }
751 Node concat = utils::mkNode(kind::BITVECTOR_CONCAT, children);
752 Assert (utils::getSize(concat) == utils::getSize(extract[0]));
753 outSubstitutions.addSubstitution(extract[0], concat);
754 return PP_ASSERT_STATUS_SOLVED;
755 }
756 }
757 }
758 break;
759 case kind::BITVECTOR_ULT:
760 case kind::BITVECTOR_SLT:
761 case kind::BITVECTOR_ULE:
762 case kind::BITVECTOR_SLE:
763
764 default:
765 // TODO other predicates
766 break;
767 }
768 return PP_ASSERT_STATUS_UNSOLVED;
769 }
770
771 Node TheoryBV::ppRewrite(TNode t)
772 {
773 Debug("bv-pp-rewrite") << "TheoryBV::ppRewrite " << t << "\n";
774 Node res = t;
775 if (RewriteRule<BitwiseEq>::applies(t)) {
776 Node result = RewriteRule<BitwiseEq>::run<false>(t);
777 res = Rewriter::rewrite(result);
778 } else if (d_isCoreTheory && t.getKind() == kind::EQUAL) {
779 std::vector<Node> equalities;
780 Slicer::splitEqualities(t, equalities);
781 res = utils::mkAnd(equalities);
782 } else if (RewriteRule<UltPlusOne>::applies(t)) {
783 Node result = RewriteRule<UltPlusOne>::run<false>(t);
784 res = Rewriter::rewrite(result);
785 } else if( res.getKind() == kind::EQUAL &&
786 ((res[0].getKind() == kind::BITVECTOR_PLUS &&
787 RewriteRule<ConcatToMult>::applies(res[1])) ||
788 (res[1].getKind() == kind::BITVECTOR_PLUS &&
789 RewriteRule<ConcatToMult>::applies(res[0])))) {
790 Node mult = RewriteRule<ConcatToMult>::applies(res[0])?
791 RewriteRule<ConcatToMult>::run<false>(res[0]) :
792 RewriteRule<ConcatToMult>::run<true>(res[1]);
793 Node factor = mult[0];
794 Node sum = RewriteRule<ConcatToMult>::applies(res[0])? res[1] : res[0];
795 Node new_eq =utils::mkNode(kind::EQUAL, sum, mult);
796 Node rewr_eq = RewriteRule<SolveEq>::run<true>(new_eq);
797 if (rewr_eq[0].isVar() || rewr_eq[1].isVar()){
798 res = Rewriter::rewrite(rewr_eq);
799 } else {
800 res = t;
801 }
802 }
803
804
805 // if(t.getKind() == kind::EQUAL &&
806 // ((t[0].getKind() == kind::BITVECTOR_MULT && t[1].getKind() == kind::BITVECTOR_PLUS) ||
807 // (t[1].getKind() == kind::BITVECTOR_MULT && t[0].getKind() == kind::BITVECTOR_PLUS))) {
808 // // if we have an equality between a multiplication and addition
809 // // try to express multiplication in terms of addition
810 // Node mult = t[0].getKind() == kind::BITVECTOR_MULT? t[0] : t[1];
811 // Node add = t[0].getKind() == kind::BITVECTOR_PLUS? t[0] : t[1];
812 // if (RewriteRule<MultSlice>::applies(mult)) {
813 // Node new_mult = RewriteRule<MultSlice>::run<false>(mult);
814 // Node new_eq = Rewriter::rewrite(utils::mkNode(kind::EQUAL, new_mult, add));
815
816 // // the simplification can cause the formula to blow up
817 // // only apply if formula reduced
818 // if (d_subtheoryMap.find(SUB_BITBLAST) != d_subtheoryMap.end()) {
819 // BitblastSolver* bv = (BitblastSolver*)d_subtheoryMap[SUB_BITBLAST];
820 // uint64_t old_size = bv->computeAtomWeight(t);
821 // Assert (old_size);
822 // uint64_t new_size = bv->computeAtomWeight(new_eq);
823 // double ratio = ((double)new_size)/old_size;
824 // if (ratio <= 0.4) {
825 // ++(d_statistics.d_numMultSlice);
826 // return new_eq;
827 // }
828 // }
829
830 // if (new_eq.getKind() == kind::CONST_BOOLEAN) {
831 // ++(d_statistics.d_numMultSlice);
832 // return new_eq;
833 // }
834 // }
835 // }
836
837 if (options::bvAbstraction() && t.getType().isBoolean()) {
838 d_abstractionModule->addInputAtom(res);
839 }
840 Debug("bv-pp-rewrite") << "to " << res << "\n";
841 return res;
842 }
843
844 void TheoryBV::presolve() {
845 Debug("bitvector") << "TheoryBV::presolve" << endl;
846 }
847
848 static int prop_count = 0;
849
850 bool TheoryBV::storePropagation(TNode literal, SubTheory subtheory)
851 {
852 Debug("bitvector::propagate") << indent() << getSatContext()->getLevel() << " " << "TheoryBV::storePropagation(" << literal << ", " << subtheory << ")" << std::endl;
853 prop_count++;
854
855 // If already in conflict, no more propagation
856 if (d_conflict) {
857 Debug("bitvector::propagate") << indent() << "TheoryBV::storePropagation(" << literal << ", " << subtheory << "): already in conflict" << std::endl;
858 return false;
859 }
860
861 // If propagated already, just skip
862 PropagatedMap::const_iterator find = d_propagatedBy.find(literal);
863 if (find != d_propagatedBy.end()) {
864 return true;
865 } else {
866 bool polarity = literal.getKind() != kind::NOT;
867 Node negatedLiteral = polarity ? literal.notNode() : (Node) literal[0];
868 find = d_propagatedBy.find(negatedLiteral);
869 if (find != d_propagatedBy.end() && (*find).second != subtheory) {
870 // Safe to ignore this one, subtheory should produce a conflict
871 return true;
872 }
873
874 d_propagatedBy[literal] = subtheory;
875 }
876
877 // Propagate differs depending on the subtheory
878 // * bitblaster needs to be left alone until it's done, otherwise it doesn't know how to explain
879 // * equality engine can propagate eagerly
880 bool ok = true;
881 if (subtheory == SUB_CORE) {
882 d_out->propagate(literal);
883 if (!ok) {
884 setConflict();
885 }
886 } else {
887 d_literalsToPropagate.push_back(literal);
888 }
889 return ok;
890
891 }/* TheoryBV::propagate(TNode) */
892
893
894 void TheoryBV::explain(TNode literal, std::vector<TNode>& assumptions) {
895 Assert (wasPropagatedBySubtheory(literal));
896 SubTheory sub = getPropagatingSubtheory(literal);
897 d_subtheoryMap[sub]->explain(literal, assumptions);
898 }
899
900
901 Node TheoryBV::explain(TNode node) {
902 Debug("bitvector::explain") << "TheoryBV::explain(" << node << ")" << std::endl;
903 std::vector<TNode> assumptions;
904
905 // Ask for the explanation
906 explain(node, assumptions);
907 // this means that it is something true at level 0
908 if (assumptions.size() == 0) {
909 return utils::mkTrue();
910 }
911 // return the explanation
912 Node explanation = utils::mkAnd(assumptions);
913 Debug("bitvector::explain") << "TheoryBV::explain(" << node << ") => " << explanation << std::endl;
914 Debug("bitvector::explain") << "TheoryBV::explain done. \n";
915 return explanation;
916 }
917
918
919 void TheoryBV::addSharedTerm(TNode t) {
920 Debug("bitvector::sharing") << indent() << "TheoryBV::addSharedTerm(" << t << ")" << std::endl;
921 d_sharedTermsSet.insert(t);
922 if (options::bitvectorEqualitySolver()) {
923 for (unsigned i = 0; i < d_subtheories.size(); ++i) {
924 d_subtheories[i]->addSharedTerm(t);
925 }
926 }
927 }
928
929
930 EqualityStatus TheoryBV::getEqualityStatus(TNode a, TNode b)
931 {
932 if (options::bitblastMode() == theory::bv::BITBLAST_MODE_EAGER)
933 return EQUALITY_UNKNOWN;
934 Assert (options::bitblastMode() == theory::bv::BITBLAST_MODE_LAZY);
935 for (unsigned i = 0; i < d_subtheories.size(); ++i) {
936 EqualityStatus status = d_subtheories[i]->getEqualityStatus(a, b);
937 if (status != EQUALITY_UNKNOWN) {
938 return status;
939 }
940 }
941 return EQUALITY_UNKNOWN; ;
942 }
943
944
945 void TheoryBV::enableCoreTheorySlicer() {
946 Assert (!d_calledPreregister);
947 d_isCoreTheory = true;
948 if (d_subtheoryMap.find(SUB_CORE) != d_subtheoryMap.end()) {
949 CoreSolver* core = (CoreSolver*)d_subtheoryMap[SUB_CORE];
950 core->enableSlicer();
951 }
952 }
953
954
955 void TheoryBV::ppStaticLearn(TNode in, NodeBuilder<>& learned) {
956 if(d_staticLearnCache.find(in) != d_staticLearnCache.end()){
957 return;
958 }
959 d_staticLearnCache.insert(in);
960
961 if (in.getKind() == kind::EQUAL) {
962 if((in[0].getKind() == kind::BITVECTOR_PLUS && in[1].getKind() == kind::BITVECTOR_SHL) ||
963 (in[1].getKind() == kind::BITVECTOR_PLUS && in[0].getKind() == kind::BITVECTOR_SHL)) {
964 TNode p = in[0].getKind() == kind::BITVECTOR_PLUS ? in[0] : in[1];
965 TNode s = in[0].getKind() == kind::BITVECTOR_PLUS ? in[1] : in[0];
966
967 if(p.getNumChildren() == 2
968 && p[0].getKind() == kind::BITVECTOR_SHL
969 && p[1].getKind() == kind::BITVECTOR_SHL ){
970 unsigned size = utils::getSize(s);
971 Node one = utils::mkConst(size, 1u);
972 if(s[0] == one && p[0][0] == one && p[1][0] == one){
973 Node zero = utils::mkConst(size, 0u);
974 TNode b = p[0];
975 TNode c = p[1];
976 // (s : 1 << S) = (b : 1 << B) + (c : 1 << C)
977 Node b_eq_0 = b.eqNode(zero);
978 Node c_eq_0 = c.eqNode(zero);
979 Node b_eq_c = b.eqNode(c);
980
981 Node dis = utils::mkNode(kind::OR, b_eq_0, c_eq_0, b_eq_c);
982 Node imp = in.impNode(dis);
983 learned << imp;
984 }
985 }
986 }
987 }else if(in.getKind() == kind::AND){
988 for(size_t i = 0, N = in.getNumChildren(); i < N; ++i){
989 ppStaticLearn(in[i], learned);
990 }
991 }
992 }
993
994 bool TheoryBV::applyAbstraction(const std::vector<Node>& assertions, std::vector<Node>& new_assertions) {
995 bool changed = d_abstractionModule->applyAbstraction(assertions, new_assertions);
996 if (changed &&
997 options::bitblastMode() == theory::bv::BITBLAST_MODE_EAGER &&
998 options::bitvectorAig()) {
999 // disable AIG mode
1000 AlwaysAssert (!d_eagerSolver->isInitialized());
1001 d_eagerSolver->turnOffAig();
1002 d_eagerSolver->initialize();
1003 }
1004 return changed;
1005 }
1006
1007 void TheoryBV::setProofLog( BitVectorProof * bvp ) {
1008 if( options::bitblastMode() == theory::bv::BITBLAST_MODE_EAGER ){
1009 d_eagerSolver->setProofLog( bvp );
1010 }else{
1011 for( unsigned i=0; i< d_subtheories.size(); i++ ){
1012 d_subtheories[i]->setProofLog( bvp );
1013 }
1014 }
1015 }
1016
1017 void TheoryBV::setConflict(Node conflict) {
1018 if (options::bvAbstraction()) {
1019 Node new_conflict = d_abstractionModule->simplifyConflict(conflict);
1020
1021 std::vector<Node> lemmas;
1022 lemmas.push_back(new_conflict);
1023 d_abstractionModule->generalizeConflict(new_conflict, lemmas);
1024 for (unsigned i = 0; i < lemmas.size(); ++i) {
1025 lemma(utils::mkNode(kind::NOT, lemmas[i]));
1026 }
1027 }
1028 d_conflict = true;
1029 d_conflictNode = conflict;
1030 }
1031
1032 } /* namespace CVC4::theory::bv */
1033 } /* namespace CVC4::theory */
1034 } /* namespace CVC4 */