ecfb521a01d176d944f9cb75c1eb6f85e3c1acce
[cvc5.git] / src / smt / solver_engine.h
1 /******************************************************************************
2 * Top contributors (to current version):
3
4 *
5 * This file is part of the cvc5 project.
6 *
7 * Copyright (c) 2009-2021 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.
11 * ****************************************************************************
12 *
13 * SolverEngine: the main public entry point of libcvc5.
14 */
15
16 #include "cvc5_public.h"
17
18 #ifndef CVC5__SMT__SOLVER_ENGINE_H
19 #define CVC5__SMT__SOLVER_ENGINE_H
20
21 #include <map>
22 #include <memory>
23 #include <string>
24 #include <vector>
25
26 #include "context/cdhashmap_forward.h"
27 #include "cvc5_export.h"
28 #include "options/options.h"
29 #include "smt/smt_mode.h"
30 #include "theory/logic_info.h"
31 #include "util/result.h"
32
33 namespace cvc5 {
34
35 template <bool ref_count>
36 class NodeTemplate;
37 typedef NodeTemplate<true> Node;
38 typedef NodeTemplate<false> TNode;
39 class TypeNode;
40
41 class Env;
42 class NodeManager;
43 class TheoryEngine;
44 class UnsatCore;
45 class StatisticsRegistry;
46 class Printer;
47 class ResourceManager;
48 struct InstantiationList;
49
50 /* -------------------------------------------------------------------------- */
51
52 namespace api {
53 class Solver;
54 } // namespace api
55
56 /* -------------------------------------------------------------------------- */
57
58 namespace context {
59 class Context;
60 class UserContext;
61 } // namespace context
62
63 /* -------------------------------------------------------------------------- */
64
65 namespace preprocessing {
66 class PreprocessingPassContext;
67 }
68
69 /* -------------------------------------------------------------------------- */
70
71 namespace prop {
72 class PropEngine;
73 } // namespace prop
74
75 /* -------------------------------------------------------------------------- */
76
77 namespace smt {
78 /** Utilities */
79 class SolverEngineState;
80 class AbstractValues;
81 class Assertions;
82 class ResourceOutListener;
83 class SmtNodeManagerListener;
84 class CheckModels;
85 /** Subsolvers */
86 class SmtSolver;
87 class SygusSolver;
88 class AbductionSolver;
89 class InterpolationSolver;
90 class QuantElimSolver;
91
92 struct SolverEngineStatistics;
93 class SolverEngineScope;
94 class PfManager;
95 class UnsatCoreManager;
96
97 } // namespace smt
98
99 /* -------------------------------------------------------------------------- */
100
101 namespace theory {
102 class TheoryModel;
103 class Rewriter;
104 class QuantifiersEngine;
105 } // namespace theory
106
107 /* -------------------------------------------------------------------------- */
108
109 class CVC5_EXPORT SolverEngine
110 {
111 friend class ::cvc5::api::Solver;
112 friend class ::cvc5::smt::SolverEngineState;
113 friend class ::cvc5::smt::SolverEngineScope;
114
115 /* ....................................................................... */
116 public:
117 /* ....................................................................... */
118
119 /**
120 * Construct an SolverEngine with the given expression manager.
121 * If provided, optr is a pointer to a set of options that should initialize
122 * the values of the options object owned by this class.
123 */
124 SolverEngine(NodeManager* nm, const Options* optr = nullptr);
125 /** Destruct the SMT engine. */
126 ~SolverEngine();
127
128 //--------------------------------------------- concerning the state
129
130 /**
131 * This is the main initialization procedure of the SolverEngine.
132 *
133 * Should be called whenever the final options and logic for the problem are
134 * set (at least, those options that are not permitted to change after
135 * assertions and queries are made).
136 *
137 * Internally, this creates the theory engine, prop engine, decision engine,
138 * and other utilities whose initialization depends on the final set of
139 * options being set.
140 *
141 * This post-construction initialization is automatically triggered by the
142 * use of the SolverEngine; e.g. when the first formula is asserted, a call
143 * to simplify() is issued, a scope is pushed, etc.
144 */
145 void finishInit();
146 /**
147 * Return true if this SolverEngine is fully initialized (post-construction)
148 * by the above call.
149 */
150 bool isFullyInited() const;
151 /**
152 * Return true if a checkEntailed() or checkSatisfiability() has been made.
153 */
154 bool isQueryMade() const;
155 /** Return the user context level. */
156 size_t getNumUserLevels() const;
157 /** Return the current mode of the solver. */
158 SmtMode getSmtMode() const;
159 /**
160 * Whether the SmtMode allows for get-value, get-model, get-assignment, etc.
161 * This is equivalent to:
162 * getSmtMode()==SmtMode::SAT || getSmtMode()==SmtMode::SAT_UNKNOWN
163 */
164 bool isSmtModeSat() const;
165 /**
166 * Returns the most recent result of checkSat/checkEntailed or
167 * (set-info :status).
168 */
169 Result getStatusOfLastCommand() const;
170 //--------------------------------------------- end concerning the state
171
172 /**
173 * Set the logic of the script.
174 * @throw ModalException, LogicException
175 */
176 void setLogic(const std::string& logic);
177
178 /**
179 * Set the logic of the script.
180 * @throw ModalException, LogicException
181 */
182 void setLogic(const char* logic);
183
184 /**
185 * Set the logic of the script.
186 * @throw ModalException
187 */
188 void setLogic(const LogicInfo& logic);
189
190 /** Get the logic information currently set. */
191 const LogicInfo& getLogicInfo() const;
192
193 /** Get the logic information set by the user. */
194 LogicInfo getUserLogicInfo() const;
195
196 /**
197 * Set information about the script executing.
198 */
199 void setInfo(const std::string& key, const std::string& value);
200
201 /** Return true if given keyword is a valid SMT-LIB v2 get-info flag. */
202 bool isValidGetInfoFlag(const std::string& key) const;
203
204 /** Query information about the SMT environment. */
205 std::string getInfo(const std::string& key) const;
206
207 /**
208 * Set an aspect of the current SMT execution environment.
209 * @throw OptionException, ModalException
210 */
211 void setOption(const std::string& key, const std::string& value);
212
213 /** Set is internal subsolver.
214 *
215 * This function is called on SolverEngine objects that are created
216 * internally. It is used to mark that this SolverEngine should not
217 * perform preprocessing passes that rephrase the input, such as
218 * --sygus-rr-synth-input or
219 * --sygus-abduct.
220 */
221 void setIsInternalSubsolver();
222 /** Is this an internal subsolver? */
223 bool isInternalSubsolver() const;
224
225 /**
226 * Block the current model. Can be called only if immediately preceded by
227 * a SAT or INVALID query. Only permitted if produce-models is on, and the
228 * block-models option is set to a mode other than "none".
229 *
230 * This adds an assertion to the assertion stack that blocks the current
231 * model based on the current options configured by cvc5.
232 *
233 * The return value has the same meaning as that of assertFormula.
234 */
235 Result blockModel();
236
237 /**
238 * Block the current model values of (at least) the values in exprs.
239 * Can be called only if immediately preceded by a SAT or NOT_ENTAILED query.
240 * Only permitted if produce-models is on, and the block-models option is set
241 * to a mode other than "none".
242 *
243 * This adds an assertion to the assertion stack of the form:
244 * (or (not (= exprs[0] M0)) ... (not (= exprs[n] Mn)))
245 * where M0 ... Mn are the current model values of exprs[0] ... exprs[n].
246 *
247 * The return value has the same meaning as that of assertFormula.
248 */
249 Result blockModelValues(const std::vector<Node>& exprs);
250
251 /**
252 * Declare heap. For smt2 inputs, this is called when the command
253 * (declare-heap (locT datat)) is invoked by the user. This sets locT as the
254 * location type and dataT is the data type for the heap. This command should
255 * be executed only once, and must be invoked before solving separation logic
256 * inputs.
257 */
258 void declareSepHeap(TypeNode locT, TypeNode dataT);
259
260 /**
261 * Get the separation heap types, which extracts which types were passed to
262 * the method above.
263 *
264 * @return true if the separation logic heap types have been declared.
265 */
266 bool getSepHeapTypes(TypeNode& locT, TypeNode& dataT);
267
268 /** When using separation logic, obtain the expression for the heap. */
269 Node getSepHeapExpr();
270
271 /** When using separation logic, obtain the expression for nil. */
272 Node getSepNilExpr();
273
274 /**
275 * Get an aspect of the current SMT execution environment.
276 * @throw OptionException
277 */
278 std::string getOption(const std::string& key) const;
279
280 /**
281 * Define function func in the current context to be:
282 * (lambda (formals) formula)
283 * This adds func to the list of defined functions, which indicates that
284 * all occurrences of func should be expanded during expandDefinitions.
285 *
286 * @param func a variable of function type that expects the arguments in
287 * formal
288 * @param formals a list of BOUND_VARIABLE expressions
289 * @param formula The body of the function, must not contain func
290 * @param global True if this definition is global (i.e. should persist when
291 * popping the user context)
292 */
293 void defineFunction(Node func,
294 const std::vector<Node>& formals,
295 Node formula,
296 bool global = false);
297
298 /**
299 * Define functions recursive
300 *
301 * For each i, this constrains funcs[i] in the current context to be:
302 * (lambda (formals[i]) formulas[i])
303 * where formulas[i] may contain variables from funcs. Unlike defineFunction
304 * above, we do not add funcs[i] to the set of defined functions. Instead,
305 * we consider funcs[i] to be a free uninterpreted function, and add:
306 * forall formals[i]. f(formals[i]) = formulas[i]
307 * to the set of assertions in the current context.
308 * This method expects input such that for each i:
309 * - func[i] : a variable of function type that expects the arguments in
310 * formals[i], and
311 * - formals[i] : a list of BOUND_VARIABLE expressions.
312 *
313 * @param global True if this definition is global (i.e. should persist when
314 * popping the user context)
315 */
316 void defineFunctionsRec(const std::vector<Node>& funcs,
317 const std::vector<std::vector<Node>>& formals,
318 const std::vector<Node>& formulas,
319 bool global = false);
320 /**
321 * Define function recursive
322 * Same as above, but for a single function.
323 */
324 void defineFunctionRec(Node func,
325 const std::vector<Node>& formals,
326 Node formula,
327 bool global = false);
328 /**
329 * Add a formula to the current context: preprocess, do per-theory
330 * setup, use processAssertionList(), asserting to T-solver for
331 * literals and conjunction of literals. Returns false if
332 * immediately determined to be inconsistent. Note this formula will
333 * be included in the unsat core when applicable.
334 *
335 * @throw TypeCheckingException, LogicException, UnsafeInterruptException
336 */
337 Result assertFormula(const Node& formula);
338
339 /**
340 * Reduce an unsatisfiable core to make it minimal.
341 */
342 std::vector<Node> reduceUnsatCore(const std::vector<Node>& core);
343
344 /**
345 * Check if a given (set of) expression(s) is entailed with respect to the
346 * current set of assertions. We check this by asserting the negation of
347 * the (big AND over the) given (set of) expression(s).
348 * Returns ENTAILED, NOT_ENTAILED, or ENTAILMENT_UNKNOWN result.
349 *
350 * @throw Exception
351 */
352 Result checkEntailed(const Node& assumption);
353 Result checkEntailed(const std::vector<Node>& assumptions);
354
355 /**
356 * Assert a formula (if provided) to the current context and call
357 * check(). Returns SAT, UNSAT, or SAT_UNKNOWN result.
358 *
359 * @throw Exception
360 */
361 Result checkSat();
362 Result checkSat(const Node& assumption);
363 Result checkSat(const std::vector<Node>& assumptions);
364
365 /**
366 * Returns a set of so-called "failed" assumptions.
367 *
368 * The returned set is a subset of the set of assumptions of a previous
369 * (unsatisfiable) call to checkSatisfiability. Calling checkSatisfiability
370 * with this set of failed assumptions still produces an unsat answer.
371 *
372 * Note that the returned set of failed assumptions is not necessarily
373 * minimal.
374 */
375 std::vector<Node> getUnsatAssumptions(void);
376
377 /*---------------------------- sygus commands ---------------------------*/
378
379 /**
380 * Add sygus variable declaration.
381 *
382 * Declared SyGuS variables may be used in SyGuS constraints, in which they
383 * are assumed to be universally quantified.
384 *
385 * In SyGuS semantics, declared functions are treated in the same manner as
386 * declared variables, i.e. as universally quantified (function) variables
387 * which can occur in the SyGuS constraints that compose the conjecture to
388 * which a function is being synthesized. Thus declared functions should use
389 * this method as well.
390 */
391 void declareSygusVar(Node var);
392
393 /**
394 * Add a function-to-synthesize declaration.
395 *
396 * The given sygusType may not correspond to the actual function type of func
397 * but to a datatype encoding the syntax restrictions for the
398 * function-to-synthesize. In this case this information is stored to be used
399 * during solving.
400 *
401 * vars contains the arguments of the function-to-synthesize. These variables
402 * are also stored to be used during solving.
403 *
404 * isInv determines whether the function-to-synthesize is actually an
405 * invariant. This information is necessary if we are dumping a command
406 * corresponding to this declaration, so that it can be properly printed.
407 */
408 void declareSynthFun(Node func,
409 TypeNode sygusType,
410 bool isInv,
411 const std::vector<Node>& vars);
412 /**
413 * Same as above, without a sygus type.
414 */
415 void declareSynthFun(Node func, bool isInv, const std::vector<Node>& vars);
416
417 /**
418 * Add a regular sygus constraint or assumption.
419 * @param n The formula
420 * @param isAssume True if n is an assumption.
421 */
422 void assertSygusConstraint(Node n, bool isAssume = false);
423
424 /**
425 * Add an invariant constraint.
426 *
427 * Invariant constraints are not explicitly declared: they are given in terms
428 * of the invariant-to-synthesize, the pre condition, transition relation and
429 * post condition. The actual constraint is built based on the inputs of these
430 * place holder predicates :
431 *
432 * PRE(x) -> INV(x)
433 * INV() ^ TRANS(x, x') -> INV(x')
434 * INV(x) -> POST(x)
435 *
436 * The regular and primed variables are retrieved from the declaration of the
437 * invariant-to-synthesize.
438 */
439 void assertSygusInvConstraint(Node inv, Node pre, Node trans, Node post);
440 /**
441 * Assert a synthesis conjecture to the current context and call
442 * check(). Returns sat, unsat, or unknown result.
443 *
444 * The actual synthesis conjecture is built based on the previously
445 * communicated information to this module (universal variables, defined
446 * functions, functions-to-synthesize, and which constraints compose it). The
447 * built conjecture is a higher-order formula of the form
448 *
449 * exists f1...fn . forall v1...vm . F
450 *
451 * in which f1...fn are the functions-to-synthesize, v1...vm are the declared
452 * universal variables and F is the set of declared constraints.
453 *
454 * @param isNext Whether we are asking for the next synthesis solution (if
455 * using incremental).
456 *
457 * @throw Exception
458 */
459 Result checkSynth(bool isNext = false);
460
461 /*------------------------- end of sygus commands ------------------------*/
462
463 /**
464 * Declare pool whose initial value is the terms in initValue. A pool is
465 * a variable of type (Set T) that is used in quantifier annotations and does
466 * not occur in constraints.
467 *
468 * @param p The pool to declare, which should be a variable of type (Set T)
469 * for some type T.
470 * @param initValue The initial value of p, which should be a vector of terms
471 * of type T.
472 */
473 void declarePool(const Node& p, const std::vector<Node>& initValue);
474 /**
475 * Simplify a formula without doing "much" work. Does not involve
476 * the SAT Engine in the simplification, but uses the current
477 * definitions, assertions, and the current partial model, if one
478 * has been constructed. It also involves theory normalization.
479 *
480 * @throw TypeCheckingException, LogicException, UnsafeInterruptException
481 *
482 * @todo (design) is this meant to give an equivalent or an
483 * equisatisfiable formula?
484 */
485 Node simplify(const Node& e);
486
487 /**
488 * Expand the definitions in a term or formula.
489 *
490 * @param n The node to expand
491 *
492 * @throw TypeCheckingException, LogicException, UnsafeInterruptException
493 */
494 Node expandDefinitions(const Node& n);
495
496 /**
497 * Get the assigned value of an expr (only if immediately preceded by a SAT
498 * or NOT_ENTAILED query). Only permitted if the SolverEngine is set to
499 * operate interactively and produce-models is on.
500 *
501 * @throw ModalException, TypeCheckingException, LogicException,
502 * UnsafeInterruptException
503 */
504 Node getValue(const Node& e) const;
505
506 /**
507 * Same as getValue but for a vector of expressions
508 */
509 std::vector<Node> getValues(const std::vector<Node>& exprs) const;
510
511 /**
512 * @return the domain elements for uninterpreted sort tn.
513 */
514 std::vector<Node> getModelDomainElements(TypeNode tn) const;
515
516 /**
517 * @return true if v is a model core symbol
518 */
519 bool isModelCoreSymbol(Node v);
520
521 /**
522 * Get a model (only if immediately preceded by an SAT or unknown query).
523 * Only permitted if the model option is on.
524 *
525 * @param declaredSorts The sorts to print in the model
526 * @param declaredFuns The free constants to print in the model. A subset
527 * of these may be printed based on isModelCoreSymbol.
528 * @return the string corresponding to the model. If the output language is
529 * smt2, then this corresponds to a response to the get-model command.
530 */
531 std::string getModel(const std::vector<TypeNode>& declaredSorts,
532 const std::vector<Node>& declaredFuns);
533
534 /** print instantiations
535 *
536 * Print all instantiations for all quantified formulas on out,
537 * returns true if at least one instantiation was printed. The type of output
538 * (list, num, etc.) is determined by printInstMode.
539 */
540 void printInstantiations(std::ostream& out);
541 /**
542 * Print the current proof. This method should be called after an UNSAT
543 * response. It gets the proof of false from the PropEngine and passes
544 * it to the ProofManager, which post-processes the proof and prints it
545 * in the proper format.
546 */
547 void printProof();
548
549 /**
550 * Get synth solution.
551 *
552 * This method returns true if we are in a state immediately preceded by
553 * a successful call to checkSynth.
554 *
555 * This method adds entries to solMap that map functions-to-synthesize with
556 * their solutions, for all active conjectures. This should be called
557 * immediately after the solver answers unsat for sygus input.
558 *
559 * Specifically, given a sygus conjecture of the form
560 * exists x1...xn. forall y1...yn. P( x1...xn, y1...yn )
561 * where x1...xn are second order bound variables, we map each xi to
562 * lambda term in solMap such that
563 * forall y1...yn. P( solMap[x1]...solMap[xn], y1...yn )
564 * is a valid formula.
565 */
566 bool getSynthSolutions(std::map<Node, Node>& solMap);
567 /**
568 * Same as above, but used for getting synthesis solutions from a "subsolver"
569 * that has been initialized to assert the synthesis conjecture as a
570 * normal assertion.
571 *
572 * This method returns true if we are in a state immediately preceded by
573 * a successful call to checkSat, where this SolverEngine has an asserted
574 * synthesis conjecture.
575 */
576 bool getSubsolverSynthSolutions(std::map<Node, Node>& solMap);
577
578 /**
579 * Do quantifier elimination.
580 *
581 * This function takes as input a quantified formula q
582 * of the form:
583 * Q x1...xn. P( x1...xn, y1...yn )
584 * where P( x1...xn, y1...yn ) is a quantifier-free
585 * formula in a logic that supports quantifier elimination.
586 * Currently, the only logics supported by quantifier
587 * elimination is LRA and LIA.
588 *
589 * This function returns a formula ret such that, given
590 * the current set of formulas A asserted to this SolverEngine :
591 *
592 * If doFull = true, then
593 * - ( A ^ q ) and ( A ^ ret ) are equivalent
594 * - ret is quantifier-free formula containing
595 * only free variables in y1...yn.
596 *
597 * If doFull = false, then
598 * - (A ^ q) => (A ^ ret) if Q is forall or
599 * (A ^ ret) => (A ^ q) if Q is exists,
600 * - ret is quantifier-free formula containing
601 * only free variables in y1...yn,
602 * - If Q is exists, let A^Q_n be the formula
603 * A ^ ~ret^Q_1 ^ ... ^ ~ret^Q_n
604 * where for each i=1,...n, formula ret^Q_i
605 * is the result of calling doQuantifierElimination
606 * for q with the set of assertions A^Q_{i-1}.
607 * Similarly, if Q is forall, then let A^Q_n be
608 * A ^ ret^Q_1 ^ ... ^ ret^Q_n
609 * where ret^Q_i is the same as above.
610 * In either case, we have that ret^Q_j will
611 * eventually be true or false, for some finite j.
612 *
613 * The former feature is quantifier elimination, and
614 * is run on invocations of the smt2 extended command get-qe.
615 * The latter feature is referred to as partial quantifier
616 * elimination, and is run on invocations of the smt2
617 * extended command get-qe-disjunct, which can be used
618 * for incrementally computing the result of a
619 * quantifier elimination.
620 *
621 * The argument strict is whether to output
622 * warnings, such as when an unexpected logic is used.
623 *
624 * throw@ Exception
625 */
626 Node getQuantifierElimination(Node q, bool doFull, bool strict = true);
627
628 /**
629 * This method asks this SMT engine to find an interpolant with respect to
630 * the current assertion stack (call it A) and the conjecture (call it B). If
631 * this method returns true, then interpolant is set to a formula I such that
632 * A ^ ~I and I ^ ~B are both unsatisfiable.
633 *
634 * The argument grammarType is a sygus datatype type that encodes the syntax
635 * restrictions on the shapes of possible solutions.
636 *
637 * This method invokes a separate copy of the SMT engine for solving the
638 * corresponding sygus problem for generating such a solution.
639 */
640 bool getInterpol(const Node& conj,
641 const TypeNode& grammarType,
642 Node& interpol);
643
644 /** Same as above, but without user-provided grammar restrictions */
645 bool getInterpol(const Node& conj, Node& interpol);
646
647 /**
648 * This method asks this SMT engine to find an abduct with respect to the
649 * current assertion stack (call it A) and the conjecture (call it B).
650 * If this method returns true, then abd is set to a formula C such that
651 * A ^ C is satisfiable, and A ^ ~B ^ C is unsatisfiable.
652 *
653 * The argument grammarType is a sygus datatype type that encodes the syntax
654 * restrictions on the shape of possible solutions.
655 *
656 * This method invokes a separate copy of the SMT engine for solving the
657 * corresponding sygus problem for generating such a solution.
658 */
659 bool getAbduct(const Node& conj, const TypeNode& grammarType, Node& abd);
660
661 /** Same as above, but without user-provided grammar restrictions */
662 bool getAbduct(const Node& conj, Node& abd);
663
664 /**
665 * Get list of quantified formulas that were instantiated on the last call
666 * to check-sat.
667 */
668 void getInstantiatedQuantifiedFormulas(std::vector<Node>& qs);
669
670 /**
671 * Get instantiation term vectors for quantified formula q.
672 *
673 * This method is similar to above, but in the example above, we return the
674 * (vectors of) terms t1, ..., tn instead.
675 *
676 * Notice that these are not guaranteed to come in the same order as the
677 * instantiation lemmas above.
678 */
679 void getInstantiationTermVectors(Node q,
680 std::vector<std::vector<Node>>& tvecs);
681 /**
682 * As above but only the instantiations that were relevant for the
683 * refutation.
684 */
685 void getRelevantInstantiationTermVectors(
686 std::map<Node, InstantiationList>& insts, bool getDebugInfo = false);
687 /**
688 * Get instantiation term vectors, which maps each instantiated quantified
689 * formula to the list of instantiations for that quantified formula. This
690 * list is minimized if proofs are enabled, and this call is immediately
691 * preceded by an UNSAT or ENTAILED query
692 */
693 void getInstantiationTermVectors(
694 std::map<Node, std::vector<std::vector<Node>>>& insts);
695
696 /**
697 * Get an unsatisfiable core (only if immediately preceded by an UNSAT or
698 * ENTAILED query). Only permitted if cvc5 was built with unsat-core support
699 * and produce-unsat-cores is on.
700 */
701 UnsatCore getUnsatCore();
702
703 /**
704 * Get a refutation proof (only if immediately preceded by an UNSAT or
705 * ENTAILED query). Only permitted if cvc5 was built with proof support and
706 * the proof option is on. */
707 std::string getProof();
708
709 /**
710 * Get the current set of assertions. Only permitted if the
711 * SolverEngine is set to operate interactively.
712 */
713 std::vector<Node> getAssertions();
714
715 /**
716 * Get difficulty map, which populates dmap, mapping input assertions
717 * to a value that estimates their difficulty for solving the current problem.
718 */
719 void getDifficultyMap(std::map<Node, Node>& dmap);
720
721 /**
722 * Push a user-level context.
723 * throw@ ModalException, LogicException, UnsafeInterruptException
724 */
725 void push();
726
727 /**
728 * Pop a user-level context. Throws an exception if nothing to pop.
729 * @throw ModalException
730 */
731 void pop();
732
733 /** Reset all assertions, global declarations, etc. */
734 void resetAssertions();
735
736 /**
737 * Interrupt a running query. This can be called from another thread
738 * or from a signal handler. Throws a ModalException if the SolverEngine
739 * isn't currently in a query.
740 *
741 * @throw ModalException
742 */
743 void interrupt();
744
745 /**
746 * Set a resource limit for SolverEngine operations. This is like a time
747 * limit, but it's deterministic so that reproducible results can be
748 * obtained. Currently, it's based on the number of conflicts.
749 * However, please note that the definition may change between different
750 * versions of cvc5 (as may the number of conflicts required, anyway),
751 * and it might even be different between instances of the same version
752 * of cvc5 on different platforms.
753 *
754 * A cumulative and non-cumulative (per-call) resource limit can be
755 * set at the same time. A call to setResourceLimit() with
756 * cumulative==true replaces any cumulative resource limit currently
757 * in effect; a call with cumulative==false replaces any per-call
758 * resource limit currently in effect. Time limits can be set in
759 * addition to resource limits; the SolverEngine obeys both. That means
760 * that up to four independent limits can control the SolverEngine
761 * at the same time.
762 *
763 * When an SolverEngine is first created, it has no time or resource
764 * limits.
765 *
766 * Currently, these limits only cause the SolverEngine to stop what its
767 * doing when the limit expires (or very shortly thereafter); no
768 * heuristics are altered by the limits or the threat of them expiring.
769 * We reserve the right to change this in the future.
770 *
771 * @param units the resource limit, or 0 for no limit
772 * @param cumulative whether this resource limit is to be a cumulative
773 * resource limit for all remaining calls into the SolverEngine (true), or
774 * whether it's a per-call resource limit (false); the default is false
775 */
776 void setResourceLimit(uint64_t units, bool cumulative = false);
777
778 /**
779 * Set a per-call time limit for SolverEngine operations.
780 *
781 * A per-call time limit can be set at the same time and replaces
782 * any per-call time limit currently in effect.
783 * Resource limits (either per-call or cumulative) can be set in
784 * addition to a time limit; the SolverEngine obeys all three of them.
785 *
786 * Note that the per-call timer only ticks away when one of the
787 * SolverEngine's workhorse functions (things like assertFormula(),
788 * checkEntailed(), checkSat(), and simplify()) are running.
789 * Between calls, the timer is still.
790 *
791 * When an SolverEngine is first created, it has no time or resource
792 * limits.
793 *
794 * Currently, these limits only cause the SolverEngine to stop what its
795 * doing when the limit expires (or very shortly thereafter); no
796 * heuristics are altered by the limits or the threat of them expiring.
797 * We reserve the right to change this in the future.
798 *
799 * @param millis the time limit in milliseconds, or 0 for no limit
800 */
801 void setTimeLimit(uint64_t millis);
802
803 /**
804 * Get the current resource usage count for this SolverEngine. This
805 * function can be used to ascertain reasonable values to pass as
806 * resource limits to setResourceLimit().
807 */
808 unsigned long getResourceUsage() const;
809
810 /** Get the current millisecond count for this SolverEngine. */
811 unsigned long getTimeUsage() const;
812
813 /**
814 * Get the remaining resources that can be consumed by this SolverEngine
815 * according to the currently-set cumulative resource limit. If there
816 * is not a cumulative resource limit set, this function throws a
817 * ModalException.
818 *
819 * @throw ModalException
820 */
821 unsigned long getResourceRemaining() const;
822
823 /** Permit access to the underlying NodeManager. */
824 NodeManager* getNodeManager() const;
825
826 /**
827 * Print statistics from the statistics registry in the env object owned by
828 * this SolverEngine. Safe to use in a signal handler.
829 */
830 void printStatisticsSafe(int fd) const;
831
832 /**
833 * Print the changes to the statistics from the statistics registry in the
834 * env object owned by this SolverEngine since this method was called the last
835 * time. Internally prints the diff and then stores a snapshot for the next
836 * call.
837 */
838 void printStatisticsDiff() const;
839
840 /** Get the options object (const and non-const versions) */
841 Options& getOptions();
842 const Options& getOptions() const;
843
844 /** Get a pointer to the UserContext owned by this SolverEngine. */
845 context::UserContext* getUserContext();
846
847 /** Get a pointer to the Context owned by this SolverEngine. */
848 context::Context* getContext();
849
850 /** Get a pointer to the TheoryEngine owned by this SolverEngine. */
851 TheoryEngine* getTheoryEngine();
852
853 /** Get a pointer to the PropEngine owned by this SolverEngine. */
854 prop::PropEngine* getPropEngine();
855
856 /** Get the resource manager of this SMT engine */
857 ResourceManager* getResourceManager() const;
858
859 /** Get the printer used by this SMT engine */
860 const Printer& getPrinter() const;
861
862 /** Get a pointer to the Rewriter owned by this SolverEngine. */
863 theory::Rewriter* getRewriter();
864 /**
865 * Get expanded assertions.
866 *
867 * Return the set of assertions, after expanding definitions.
868 */
869 std::vector<Node> getExpandedAssertions();
870
871 /**
872 * !!!!! temporary, until the environment is passsed to all classes that
873 * require it.
874 */
875 Env& getEnv();
876 /* ....................................................................... */
877 private:
878 /* ....................................................................... */
879
880 // disallow copy/assignment
881 SolverEngine(const SolverEngine&) = delete;
882 SolverEngine& operator=(const SolverEngine&) = delete;
883
884 /** Set solver instance that owns this SolverEngine. */
885 void setSolver(api::Solver* solver) { d_solver = solver; }
886
887 /** Get a pointer to the (new) PfManager owned by this SolverEngine. */
888 smt::PfManager* getPfManager() { return d_pfManager.get(); };
889
890 /** Get a pointer to the StatisticsRegistry owned by this SolverEngine. */
891 StatisticsRegistry& getStatisticsRegistry();
892
893 /**
894 * Internal method to get an unsatisfiable core (only if immediately preceded
895 * by an UNSAT or ENTAILED query). Only permitted if cvc5 was built with
896 * unsat-core support and produce-unsat-cores is on. Does not dump the
897 * command.
898 */
899 UnsatCore getUnsatCoreInternal();
900
901 /**
902 * Check that a generated proof checks. This method is the same as printProof,
903 * but does not print the proof. Like that method, it should be called
904 * after an UNSAT response. It ensures that a well-formed proof of false
905 * can be constructed by the combination of the PropEngine and ProofManager.
906 */
907 void checkProof();
908
909 /**
910 * Check that an unsatisfiable core is indeed unsatisfiable.
911 */
912 void checkUnsatCore();
913
914 /**
915 * Check that a generated Model (via getModel()) actually satisfies
916 * all user assertions.
917 */
918 void checkModel(bool hardFailure = true);
919
920 /**
921 * Check that a solution to an interpolation problem is indeed a solution.
922 *
923 * The check is made by determining that the assertions imply the solution of
924 * the interpolation problem (interpol), and the solution implies the goal
925 * (conj). If these criteria are not met, an internal error is thrown.
926 */
927 void checkInterpol(Node interpol,
928 const std::vector<Node>& easserts,
929 const Node& conj);
930
931 /**
932 * This is called by the destructor, just before destroying the
933 * PropEngine, TheoryEngine, and DecisionEngine (in that order). It
934 * is important because there are destruction ordering issues
935 * between PropEngine and Theory.
936 */
937 void shutdown();
938
939 /**
940 * Quick check of consistency in current context: calls
941 * processAssertionList() then look for inconsistency (based only on
942 * that).
943 */
944 Result quickCheck();
945
946 /**
947 * Get the (SMT-level) model pointer, if we are in SAT mode. Otherwise,
948 * return nullptr.
949 *
950 * This ensures that the underlying theory model of the SmtSolver maintained
951 * by this class is currently available, which means that cvc5 is producing
952 * models, and is in "SAT mode", otherwise a recoverable exception is thrown.
953 *
954 * @param c used for giving an error message to indicate the context
955 * this method was called.
956 */
957 theory::TheoryModel* getAvailableModel(const char* c) const;
958 /**
959 * Get available quantifiers engine, which throws a modal exception if it
960 * does not exist. This can happen if a quantifiers-specific call (e.g.
961 * getInstantiatedQuantifiedFormulas) is called in a non-quantified logic.
962 *
963 * @param c used for giving an error message to indicate the context
964 * this method was called.
965 */
966 theory::QuantifiersEngine* getAvailableQuantifiersEngine(const char* c) const;
967
968 // --------------------------------------- callbacks from the state
969 /**
970 * Notify push pre, which is called just before the user context of the state
971 * pushes. This processes all pending assertions.
972 */
973 void notifyPushPre();
974 /**
975 * Notify push post, which is called just after the user context of the state
976 * pushes. This performs a push on the underlying prop engine.
977 */
978 void notifyPushPost();
979 /**
980 * Notify pop pre, which is called just before the user context of the state
981 * pops. This performs a pop on the underlying prop engine.
982 */
983 void notifyPopPre();
984 /**
985 * Notify post solve pre, which is called once per check-sat query. It
986 * is triggered when the first d_state.doPendingPops() is issued after the
987 * check-sat. This method is called before the contexts pop in the method
988 * doPendingPops.
989 */
990 void notifyPostSolvePre();
991 /**
992 * Same as above, but after contexts are popped. This calls the postsolve
993 * method of the underlying TheoryEngine.
994 */
995 void notifyPostSolvePost();
996 // --------------------------------------- end callbacks from the state
997
998 /**
999 * Internally handle the setting of a logic. This function should always
1000 * be called when d_logic is updated.
1001 */
1002 void setLogicInternal();
1003
1004 /*
1005 * Check satisfiability (used to check satisfiability and entailment).
1006 */
1007 Result checkSatInternal(const std::vector<Node>& assumptions,
1008 bool isEntailmentCheck);
1009
1010 /**
1011 * Check that all Expr in formals are of BOUND_VARIABLE kind, where func is
1012 * the function that the formal argument list is for. This method is used
1013 * as a helper function when defining (recursive) functions.
1014 */
1015 void debugCheckFormals(const std::vector<Node>& formals, Node func);
1016
1017 /**
1018 * Checks whether formula is a valid function body for func whose formal
1019 * argument list is stored in formals. This method is
1020 * used as a helper function when defining (recursive) functions.
1021 */
1022 void debugCheckFunctionBody(Node formula,
1023 const std::vector<Node>& formals,
1024 Node func);
1025
1026 /**
1027 * Helper method to obtain both the heap and nil from the solver. Returns a
1028 * std::pair where the first element is the heap expression and the second
1029 * element is the nil expression.
1030 */
1031 std::pair<Node, Node> getSepHeapAndNilExpr();
1032 /**
1033 * Get assertions internal, which is only called after initialization. This
1034 * should be used internally to get the assertions instead of getAssertions
1035 * or getExpandedAssertions, which may trigger initialization and SMT state
1036 * changes.
1037 */
1038 std::vector<Node> getAssertionsInternal();
1039
1040 /**
1041 * Return a reference to options like for `EnvObj`.
1042 */
1043 const Options& options() const;
1044
1045 /* Members -------------------------------------------------------------- */
1046
1047 /** Solver instance that owns this SolverEngine instance. */
1048 api::Solver* d_solver = nullptr;
1049
1050 /**
1051 * The environment object, which contains all utilities that are globally
1052 * available to internal code.
1053 */
1054 std::unique_ptr<Env> d_env;
1055 /**
1056 * The state of this SolverEngine, which is responsible for maintaining which
1057 * SMT mode we are in, the contexts, the last result, etc.
1058 */
1059 std::unique_ptr<smt::SolverEngineState> d_state;
1060
1061 /** Abstract values */
1062 std::unique_ptr<smt::AbstractValues> d_absValues;
1063 /** Assertions manager */
1064 std::unique_ptr<smt::Assertions> d_asserts;
1065 /** Resource out listener */
1066 std::unique_ptr<smt::ResourceOutListener> d_routListener;
1067
1068 /** The SMT solver */
1069 std::unique_ptr<smt::SmtSolver> d_smtSolver;
1070
1071 /**
1072 * The utility used for checking models
1073 */
1074 std::unique_ptr<smt::CheckModels> d_checkModels;
1075
1076 /**
1077 * The proof manager, which manages all things related to checking,
1078 * processing, and printing proofs.
1079 */
1080 std::unique_ptr<smt::PfManager> d_pfManager;
1081
1082 /**
1083 * The unsat core manager, which produces unsat cores and related information
1084 * from refutations. */
1085 std::unique_ptr<smt::UnsatCoreManager> d_ucManager;
1086
1087 /** The solver for sygus queries */
1088 std::unique_ptr<smt::SygusSolver> d_sygusSolver;
1089
1090 /** The solver for abduction queries */
1091 std::unique_ptr<smt::AbductionSolver> d_abductSolver;
1092 /** The solver for interpolation queries */
1093 std::unique_ptr<smt::InterpolationSolver> d_interpolSolver;
1094 /** The solver for quantifier elimination queries */
1095 std::unique_ptr<smt::QuantElimSolver> d_quantElimSolver;
1096
1097 /**
1098 * The logic set by the user. The actual logic, which may extend the user's
1099 * logic, lives in the Env class.
1100 */
1101 LogicInfo d_userLogic;
1102
1103 /** Whether this is an internal subsolver. */
1104 bool d_isInternalSubsolver;
1105
1106 /** The statistics class */
1107 std::unique_ptr<smt::SolverEngineStatistics> d_stats;
1108
1109 /**
1110 * The global scope object. Upon creation of this SolverEngine, it becomes the
1111 * SolverEngine in scope. It says the SolverEngine in scope until it is
1112 * destructed, or another SolverEngine is created.
1113 */
1114 std::unique_ptr<smt::SolverEngineScope> d_scope;
1115 }; /* class SolverEngine */
1116
1117 /* -------------------------------------------------------------------------- */
1118
1119 } // namespace cvc5
1120
1121 #endif /* CVC5__SMT_ENGINE_H */