Solver

class cvc5. Solver

A cvc5 solver.

Wrapper class for cvc5::Solver .

addSygusAssume ( self , Term t )

Add a formula to the set of Sygus assumptions.

SyGuS v2:

( assume <term> )
Parameters :

term – The formuula to add as an assumption.

addSygusConstraint ( self , Term t )

Add a formula to the set of SyGuS constraints.

SyGuS v2:

( constraint <term> )
Parameters :

term – The formula to add as a constraint.

addSygusInvConstraint ( self , Term inv_f , Term pre_f , Term trans_f , Term post_f )

Add a set of SyGuS constraints to the current state that correspond to an invariant synthesis problem.

SyGuS v2:

( inv-constraint <inv> <pre> <trans> <post> )
Parameters :
  • inv – The function-to-synthesize.

  • pre – The pre-condition.

  • trans – The transition relation.

  • post – The post-condition.

assertFormula ( self , Term term )

Assert a formula

SMT-LIB:

( assert <term> )
Parameters :

term – The formula to assert.

blockModel ( self , mode )

Block the current model. Can be called only if immediately preceded by a SAT or INVALID query.

SMT-LIB:

(block-model)

Requires enabling option produce-models to a mode other than none .

Warning

This method is experimental and may change in future versions.

Parameters :

mode – The mode to use for blocking

blockModelValues ( self , terms )

Block the current model values of (at least) the values in terms. Can be called only if immediately preceded by a SAT or NOT_ENTAILED query.

SMT-LIB:

(block-model-values ( <terms>+ ))

Requires enabling option produce-models .

Warning

This method is experimental and may change in future versions.

checkSat ( self )

Check satisfiability.

SMT-LIB:

( check-sat )
Returns :

The result of the satisfiability check.

checkSatAssuming ( self , * assumptions )

Check satisfiability assuming the given formula.

SMT-LIB:

( check-sat-assuming ( <prop_literal> ) )
Parameters :

assumptions – The formulas to assume.

Returns :

The result of the satisfiability check.

checkSynth ( self )

Try to find a solution for the synthesis conjecture corresponding to the current list of functions-to-synthesize, universal variables and constraints.

SyGuS v2:

( check-synth )
Returns :

The result of the check, which is “solution” if the check found a solution in which case solutions are available via getSynthSolutions, “no solution” if it was determined there is no solution, or “unknown” otherwise.

checkSynthNext ( self )

Try to find a next solution for the synthesis conjecture corresponding to the current list of functions-to-synthesize, universal variables and constraints. Must be called immediately after a successful call to check-synth or check-synth-next. Requires incremental mode.

SyGuS v2:

( check-synth )
Returns :

The result of the check, which is “solution” if the check found a solution in which case solutions are available via getSynthSolutions, “no solution” if it was determined there is no solution, or “unknown” otherwise.

declareDatatype ( self , unicode symbol , *ctors )

Create datatype sort.

SMT-LIB:

( declare-datatype <symbol> <datatype_decl> )
Parameters :
  • symbol – The name of the datatype sort.

  • ctors – The constructor declarations of the datatype sort.

Returns :

The datatype sort.

declareFun ( self , unicode symbol , list sorts , Sort sort )

Declare n-ary function symbol.

SMT-LIB:

( declare-fun <symbol> ( <sort>* ) <sort> )
Parameters :
  • symbol – The name of the function.

  • sorts – The sorts of the parameters to this function.

  • sort – The sort of the return value of this function.

Returns :

The function.

declarePool ( self , unicode symbol , Sort sort , initValue )

Declare a symbolic pool of terms with the given initial value.

SMT-LIB:

( declare-pool <symbol> <sort> ( <term>* ) )

Warning

This method is experimental and may change in future versions.

Parameters :
  • symbol – The name of the pool.

  • sort – The sort of the elements of the pool.

  • initValue – The initial value of the pool.

declareSepHeap ( self , Sort locType , Sort dataType )

When using separation logic, this sets the location sort and the datatype sort to the given ones. This method should be invoked exactly once, before any separation logic constraints are provided.

Warning

This method is experimental and may change in future versions.

Parameters :
  • locSort – The location sort of the heap.

  • dataSort – The data sort of the heap.

declareSort ( self , unicode symbol , int arity )

Declare uninterpreted sort.

SMT-LIB:

( declare-sort <symbol> <numeral> )

Note

This corresponds to Solver.mkUninterpretedSort() if arity = 0, and to Solver.mkUninterpretedSortConstructorSort() if arity > 0.

Parameters :
  • symbol – The name of the sort.

  • arity – The arity of the sort.

Returns :

The sort.

declareSygusVar ( self , unicode symbol , Sort sort )

Append symbol to the current list of universal variables.

SyGuS v2:

( declare-var <symbol> <sort> )
Parameters :
  • sort – The sort of the universal variable.

  • symbol – The name of the universal variable.

Returns :

The universal variable.

defineFun ( self , unicode symbol , list bound_vars , Sort sort , Term term , glbl=False )

Define n-ary function.

SMT-LIB:

( define-fun <function_def> )
Parameters :
  • symbol – The name of the function.

  • bound_vars – The parameters to this function.

  • sort – The sort of the return value of this function.

  • term – The function body.

  • glbl – Determines whether this definition is global (i.e. persists when popping the context).

Returns :

The function.

defineFunRec ( self , sym_or_fun , bound_vars , sort_or_term , t = None , glbl = False )

Define recursive functions.

Supports the following arguments:

  • Term defineFunRec(str symbol, List[Term] bound_vars, Sort sort, Term term, bool glbl)

  • Term defineFunRec(Term fun, List[Term] bound_vars, Term term, bool glbl)

SMT-LIB:

( define-funs-rec ( <function_decl>^n ) ( <term>^n ) )

Create elements of parameter funs with mkConst() .

Parameters :
  • funs – The sorted functions.

  • bound_vars – The list of parameters to the functions.

  • terms – The list of function bodies of the functions.

  • global – Determines whether this definition is global (i.e. persists when popping the context).

Returns :

The function.

defineFunsRec ( self , funs , bound_vars , terms , glb = False )

Define recursive functions.

SMT-LIB:

( define-funs-rec ( <function_decl>^n ) ( <term>^n ) )

Create elements of parameter funs with mkConst() .

Parameters :
  • funs – The sorted functions.

  • bound_vars – The list of parameters to the functions.

  • terms – The list of function bodies of the functions.

  • glb – Determines whether this definition is global (i.e. persists when popping the context).

getAbduct ( self , Term conj , Grammar grammar=None )

Get an abduct.

SMT-LIB:

( get-abduct <conj> )
( get-abduct <conj> <grammar> )

Requires to enable option produce-abducts .

Warning

This method is experimental and may change in future versions.

Parameters :
  • conj – The conjecture term.

  • grammar – A grammar for the abduct.

Returns :

The abduct. See cvc5::Solver::getAbduct() for details.

getAbductNext ( self )

Get the next abduct.

Can only be called immediately after a succesful call to Solver.getAbduct() or Solver.getAbductNext() . Is guaranteed to produce a syntactically different abduct wrt the last returned abduct if successful.

SMT-LIB:

( get-abduct-next )

Requires to enable incremental mode, and option produce-abducts .

Warning

This method is experimental and may change in future versions.

Parameters :

output – The term where the result will be stored.

Returns :

True iff an abduct was found.

getAssertions ( self )

Get the list of asserted formulas.

SMT-LIB:

( get-assertions )
Returns :

The list of asserted formulas.

getBooleanSort ( self )
Returns :

Sort Boolean.

getDifficulty ( self )

Get a difficulty estimate for an asserted formula. This method is intended to be called immediately after any response to a Solver.checkSat() call.

Warning

This method is experimental and may change in future versions.

Returns :

A map from (a subset of) the input assertions to a real value that is an estimate of how difficult each assertion was to solver. Unmentioned assertions can be assumed to have zero difficulty.

getInfo ( self , unicode flag )

Get info from the solver.

SMT-LIB:

( get-info <info_flag> )
Parameters :

flag – The info flag.

Returns :

The info.

getInstantiations ( self )

Return a string that contains information about all instantiations made by the quantifiers module.

Warning

This method is experimental and may change in future versions.

getIntegerSort ( self )
Returns :

Sort Integer.

getInterpolant ( self , Term conj , Grammar grammar=None )

Get an interpolant.

SMT-LIB:

( get-interpolant <conj> )
( get-interpolant <conj> <grammar> )

Requires option produce-interpolants to be set to a mode different from none .

Warning

This method is experimental and may change in future versions.

Parameters :
  • conj – The conjecture term.

  • grammar – A grammar for the inteprolant.

Returns :

The interpolant. See cvc5::Solver::getInterpolant() for details.

getInterpolantNext ( self )

Get the next interpolant.

Can only be called immediately after a successful call to Solver.getInterpolant() or Solver.getInterpolantNext() . Is guaranteed to produce a syntactically different interpolant wrt the last returned interpolant if successful.

SMT-LIB:

( get-interpolant-next )

Requires to enable incremental mode, and option produce-interpolants to be set to a mode different from none .

Warning

This method is experimental and may change in future versions.

Parameters :

output – The term where the result will be stored.

Returns :

True iff an interpolant was found.

getLearnedLiterals ( self , type = LearnedLitType.LEARNED_LIT_INPUT )

Get a list of literals that are entailed by the current set of assertions

SMT-LIB:

( get-learned-literals )

Warning

This method is experimental and may change in future versions.

Parameters :

type – The type of learned literals to return

Returns :

The list of literals.

getModel ( self , sorts , consts )

Get the model

SMT-LIB:

(get-model)

Requires to enable option produce-models .

Warning

This method is experimental and may change in future versions.

Parameters :
  • sorts – The list of uninterpreted sorts that should be printed in the model.

  • vars – The list of free constants that should be printed in the model. A subset of these may be printed based on Solver.isModelCoreSymbol() .

Returns :

A string representing the model.

getModelDomainElements ( self , Sort s )

Get the domain elements of uninterpreted sort s in the current model. The current model interprets s as the finite sort whose domain elements are given in the return value of this method.

Parameters :

s – The uninterpreted sort in question.

Returns :

The domain elements of s in the current model.

getNullSort ( self )
Returns :

A null sort object.

getOption ( self , unicode option )

Get the value of a given option.

SMT-LIB:

( get-option <keyword> )
Parameters :

option – The option for which the value is queried.

Returns :

A string representation of the option value.

getOptionInfo ( self , unicode option )

Get some information about the given option. Returns the information provided by the C++ OptionInfo as a dictionary.

Returns :

Information about the given option.

getOptionNames ( self )

Get all option names that can be used with Solver.setOption() , Solver.getOption() and Solver.getOptionInfo() .

Returns :

All option names.

getProof ( self , c = ProofComponent.PROOF_COMPONENT_FULL )

Get a proof associated with the most recent call to checkSat.

SMT-LIB:

(get-proof :c)

Requires to enable option produce-proofs .

Warning

This method is experimental and may change in future versions.

Parameters :

c – The component of the proof to return

Returns :

A string representing the proof. This takes into account

proof-format-mode when c is PROOF_COMPONENT_FULL.

getQuantifierElimination ( self , Term term )

Do quantifier elimination.

SMT-LIB:

( get-qe <q> )

Requires a logic that supports quantifier elimination. Currently, the only logics supported by quantifier elimination are LRA and LIA.

Warning

This method is experimental and may change in future versions.

Parameters :

q – A quantified formula of the form \(Q\bar{x_1}\dots Q\bar{x}_n. P( x_1 \dots x_i, y_1 \dots y_j)\) where \(Q\bar{x}\) is a set of quantified variables of the form \(Q x_1...x_k\) and \(P( x_1...x_i, y_1...y_j )\) is a quantifier-free formula

Returns :

A formula \(\phi\) such that, given the current set of formulas \(A\) asserted to this solver:

  • \((A \wedge q)\) \((A \wedge \phi)\) are equivalent

  • \(\phi\) is quantifier-free formula containing only free variables in \(y_1...y_n\) .

getQuantifierEliminationDisjunct ( self , Term term )

Do partial quantifier elimination, which can be used for incrementally computing the result of a quantifier elimination.

SMT-LIB:

( get-qe-disjunct <q> )

Requires a logic that supports quantifier elimination. Currently, the only logics supported by quantifier elimination are LRA and LIA.

Warning

This method is experimental and may change in future versions.

Parameters :

q – A quantified formula of the form \(Q\bar{x_1} ... Q\bar{x_n}. P( x_1...x_i, y_1...y_j)\) where \(Q\bar{x}\) is a set of quantified variables of the form \(Q x_1...x_k\) and \(P( x_1...x_i, y_1...y_j )\) is a quantifier-free formula.

Returns :

A formula \(\phi\) such that, given the current set of formulas \(A\) asserted to this solver:

  • \((A \wedge q \implies A \wedge \phi)\) if \(Q\) is \(\forall\) , and \((A \wedge \phi \implies A \wedge q)\) if \(Q\) is \(\exists\)

  • \(\phi\) is quantifier-free formula containing only free variables in \(y_1...y_n\)

  • If \(Q\) is \(\exists\) , let \((A \wedge Q_n)\) be the formula \((A \wedge \neg (\phi \wedge Q_1) \wedge ... \wedge \neg (\phi \wedge Q_n))\) where for each \(i = 1...n\) , formula \((\phi \wedge Q_i)\) is the result of calling getQuantifierEliminationDisjunct() for \(q\) with the set of assertions \((A \wedge Q_{i-1})\) . Similarly, if \(Q\) is \(\forall\) , then let \((A \wedge Q_n)\) be \((A \wedge (\phi \wedge Q_1) \wedge ... \wedge (\phi \wedge Q_n))\) where \((\phi \wedge Q_i)\) is the same as above. In either case, we have that \((\phi \wedge Q_j)\) will eventually be true or false, for some finite \(j\) .

getRealSort ( self )
Returns :

Sort Real.

getRegExpSort ( self )
Returns :

The sort of regular expressions.

getRoundingModeSort ( self )
Returns :

Sort RoundingMode.

getStatistics ( self )

Returns a snapshot of the current state of the statistic values of this solver. The returned object is completely decoupled from the solver and will not change when the solver is used again.

getStringSort ( self )
Returns :

Sort String.

getSynthSolution ( self , Term term )

Get the synthesis solution of the given term. This method should be called immediately after the solver answers unsat for sygus input.

Parameters :

term – The term for which the synthesis solution is queried.

Returns :

The synthesis solution of the given term.

getSynthSolutions ( self , list terms )

Get the synthesis solutions of the given terms. This method should be called immediately after the solver answers unsat for sygus input.

Parameters :

terms – The terms for which the synthesis solutions is queried.

Returns :

The synthesis solutions of the given terms.

getUnsatAssumptions ( self )

Get the set of unsat (“failed”) assumptions.

SMT-LIB:

( get-unsat-assumptions )

Requires to enable option produce-unsat-assumptions .

Returns :

The set of unsat assumptions.

getUnsatCore ( self )

Get the unsatisfiable core.

SMT-LIB:

(get-unsat-core)

Requires to enable option produce-unsat-cores .

Note

In contrast to SMT-LIB, the API does not distinguish between named and unnamed assertions when producing an unsatisfiable core. Additionally, the API allows this option to be called after a check with assumptions. A subset of those assumptions may be included in the unsatisfiable core returned by this method.

Returns :

A set of terms representing the unsatisfiable core.

getValue ( self , term_or_list )

Get the value of the given term or list of terms in the current model.

SMT-LIB:

( get-value ( <term>* ) )
Parameters :

term_or_list – The term or list of terms for which the value is queried.

Returns :

The value or list of values of the given term or list of terms.

getValueSepHeap ( self )

When using separation logic, obtain the term for the heap.

Warning

This method is experimental and may change in future versions.

Returns :

The term for the heap.

getValueSepNil ( self )

When using separation logic, obtain the term for nil.

Warning

This method is experimental and may change in future versions.

Returns :

The term for nil.

isModelCoreSymbol ( self , Term v )

This returns False if the model value of free constant v was not essential for showing the satisfiability of the last call to checkSat using the current model. This method will only return false (for any v) if the model-cores option has been set.

Warning

This method is experimental and may change in future versions.

Parameters :

v – The term in question.

Returns :

True if v is a model core symbol.

mkArraySort ( self , Sort indexSort , Sort elemSort )

Create an array sort.

Parameters :
  • indexSort – The array index sort.

  • elemSort – The array element sort.

Returns :

The array sort.

mkBagSort ( self , Sort elemSort )

Create a bag sort.

Parameters :

elemSort – The sort of the bag elements.

Returns :

The bag sort.

mkBitVector ( self , int size , *args )

Create bit-vector value.

Supports the following arguments:

  • Term mkBitVector(int size, int val=0)

  • Term mkBitVector(int size, string val, int base)

Returns :

A Term representing a bit-vector value.

Parameters :
  • size – The bit-width.

  • val – An integer representating the value, in the first form. In the second form, a string representing the value.

  • base – The base of the string representation (second form only).

mkBitVectorSort ( self , uint32_t size )

Create a bit-vector sort.

Parameters :

size – The bit-width of the bit-vector sort

Returns :

The bit-vector sort

mkBoolean ( self , bool val )

Create a Boolean constant.

Returns :

The Boolean constant.

Parameters :

val – The value of the constant.

mkCardinalityConstraint ( self , Sort sort , int index )

Create cardinality constraint.

Warning

This method is experimental and may change in future versions.

Parameters :
  • sort – Sort of the constraint.

  • index – The upper bound for the cardinality of the sort.

mkConst ( self , Sort sort , symbol=None )

Create (first-order) constant (0-arity function symbol).

SMT-LIB:

( declare-const <symbol> <sort> )
( declare-fun <symbol> ( ) <sort> )
Parameters :
  • sort – The sort of the constant.

  • symbol – The name of the constant. If None, a default symbol is used.

Returns :

The first-order constant.

mkConstArray ( self , Sort sort , Term val )

Create a constant array with the provided constant value stored at every index

Parameters :
  • sort – The sort of the constant array (must be an array sort).

  • val – The constant value to store (must match the sort’s element sort).

Returns :

The constant array term.

mkDatatypeConstructorDecl ( self , unicode name )

Create datatype constructor declaration.

Parameters :

name – The name of the constructor.

Returns :

The datatype constructor declaration.

mkDatatypeDecl ( self , unicode name , sorts_or_bool=None , isCoDatatype=None )

Create a datatype declaration.

Parameters :
  • name – The name of the datatype.

  • isCoDatatype – True if a codatatype is to be constructed.

Returns :

The datatype declaration.

mkDatatypeSort ( self , DatatypeDecl dtypedecl )

Create a datatype sort.

Parameters :

dtypedecl – The datatype declaration from which the sort is created.

Returns :

The datatype sort.

mkDatatypeSorts ( self , list dtypedecls )

Create a vector of datatype sorts using unresolved sorts. The names of the datatype declarations in dtypedecls must be distinct.

When constructing datatypes, unresolved sorts are replaced by the datatype sort constructed for the datatype declaration it is associated with.

Parameters :

dtypedecls – The datatype declarations from which the sort is created.

Returns :

The datatype sorts.

mkEmptyBag ( self , Sort s )

Create a constant representing an empty bag of the given sort.

Parameters :

sort – The sort of the bag elements.

Returns :

The empty bag constant.

mkEmptySequence ( self , Sort sort )

Create an empty sequence of the given element sort.

Parameters :

sort – The element sort of the sequence.

Returns :

The empty sequence with given element sort.

mkEmptySet ( self , Sort s )

Create a constant representing an empty set of the given sort.

Parameters :

sort – The sort of the set elements.

Returns :

The empty set constant.

mkFalse ( self )

Create a Boolean false constant.

Returns :

The false constant.

mkFloatingPoint ( self , int exp , int sig , Term val )

Create a floating-point constant.

Parameters :
  • exp – Size of the exponent.

  • sig – Size of the significand.

  • val – Value of the floating-point constant as a bit-vector term.

mkFloatingPointNaN ( self , int exp , int sig )

Create a not-a-number (NaN) floating-point constant.

Parameters :
  • exp – Number of bits in the exponent.

  • sig – Number of bits in the significand.

Returns :

The floating-point constant.

mkFloatingPointNegInf ( self , int exp , int sig )

Create a negative infinity floating-point constant.

Parameters :
  • exp – Number of bits in the exponent.

  • sig – Number of bits in the significand.

Returns :

The floating-point constant.

mkFloatingPointNegZero ( self , int exp , int sig )

Create a negative zero (+0.0) floating-point constant.

Parameters :
  • exp – Number of bits in the exponent.

  • sig – Number of bits in the significand.

Returns :

The floating-point constant.

mkFloatingPointPosInf ( self , int exp , int sig )

Create a positive infinity floating-point constant.

Parameters :
  • exp – Number of bits in the exponent.

  • sig – Number of bits in the significand.

Returns :

The floating-point constant.

mkFloatingPointPosZero ( self , int exp , int sig )

Create a positive zero (+0.0) floating-point constant.

Parameters :
  • exp – Number of bits in the exponent.

  • sig – Number of bits in the significand.

Returns :

The floating-point constant.

mkFloatingPointSort ( self , uint32_t exp , uint32_t sig )

Create a floating-point sort.

Parameters :
  • exp – The bit-width of the exponent of the floating-point sort.

  • sig – The bit-width of the significand of the floating-point sort.

mkFunctionSort ( self , sorts , Sort codomain )

Create function sort.

Parameters :
  • sorts – The sort of the function arguments.

  • codomain – The sort of the function return value.

Returns :

The function sort.

mkGrammar ( self , boundVars , ntSymbols )

Create a SyGuS grammar. The first non-terminal is treated as the starting non-terminal, so the order of non-terminals matters.

Parameters :
  • boundVars – The parameters to corresponding synth-fun/synth-inv.

  • ntSymbols – The pre-declaration of the non-terminal symbols.

Returns :

The grammar.

mkInteger ( self , val )

Create an integer constant.

Parameters :

val – Representation of the constant: either a string or integer.

Returns :

A constant of sort Integer.

mkOp ( self , k , * args )

Create operator.

Supports the following arguments:

  • Op mkOp(Kind kind)

  • Op mkOp(Kind kind, const string& arg)

  • Op mkOp(Kind kind, uint32_t arg0, ...)

mkParamSort ( self , unicode symbolname=None )

Create a sort parameter.

Warning

This method is experimental and may change in future versions.

Parameters :

symbol – The name of the sort.

Returns :

The sort parameter.

mkPi ( self )

Create a constant representing the number Pi.

Returns :

A constant representing PI .

mkPredicateSort ( self , * sorts )

Create a predicate sort.

Parameters :

sorts – The list of sorts of the predicate.

Returns :

The predicate sort.

mkReal ( self , numerator , denominator = None )

Create a real constant from a numerator and an optional denominator.

First converts the arguments to a temporary string, either "<numerator>" or "<numerator>/<denominator>" . This temporary string is forwarded to cvc5::Solver::mkReal() and should thus represent an integer, a decimal number or a fraction.

Parameters :
  • numerator – The numerator.

  • denominator – The denominator, or None .

Returns :

A real term with literal value.

mkRecordSort ( self , * fields )

Create a record sort

Warning

This method is experimental and may change in future versions.

Parameters :

fields – The list of fields of the record.

Returns :

The record sort.

mkRegexpAll ( self )

Create a regular expression all ( re.all ) term.

Returns :

The all term.

mkRegexpAllchar ( self )

Create a regular expression allchar ( re.allchar ) term.

Returns :

The allchar term.

mkRegexpNone ( self )

Create a regular expression none ( re.none ) term.

Returns :

The none term.

mkRoundingMode ( self , rm )

Create a roundingmode constant.

Parameters :

rm – The floating point rounding mode this constant represents.

mkSepEmp ( self )

Create a separation logic empty term.

Warning

This method is experimental and may change in future versions.

Returns :

The separation logic empty term.

mkSepNil ( self , Sort sort )

Create a separation logic nil term.

Warning

This method is experimental and may change in future versions.

Parameters :

sort – The sort of the nil term.

Returns :

The separation logic nil term.

mkSequenceSort ( self , Sort elemSort )

Create a sequence sort.

Parameters :

elemSort – The sort of the sequence elements

Returns :

The sequence sort.

mkSetSort ( self , Sort elemSort )

Create a set sort.

Parameters :

elemSort – The sort of the set elements.

Returns :

The set sort.

mkString ( self , unicode s , useEscSequences=None )

Create a String constant from a str which may contain SMT-LIB compatible escape sequences like \u1234 to encode unicode characters.

Parameters :
  • s – The string this constant represents.

  • useEscSequences – Determines whether escape sequences in s should be converted to the corresponding unicode character

Returns :

The String constant.

mkTerm ( self , kind_or_op , * args )

Create a term.

Supports the following arguments:

  • Term mkTerm(Kind kind)

  • Term mkTerm(Op op)

  • Term mkTerm(Kind kind, *args)

  • Term mkTerm(Op op, *args)

where *args is a comma-separated list of terms.

mkTrue ( self )

Create a Boolean true constant.

Returns :

The true constant.

mkTuple ( self , sorts , terms )

Create a tuple term. Terms are automatically converted if sorts are compatible.

Parameters :
  • sorts – The sorts of the elements in the tuple.

  • terms – The elements in the tuple.

Returns :

The tuple Term.

mkTupleSort ( self , * sorts )

Create a tuple sort.

Parameters :

sorts – Of the elements of the tuple.

Returns :

The tuple sort.

mkUninterpretedSort ( self , unicode name=None )

Create an uninterpreted sort.

Parameters :

symbol – The name of the sort.

Returns :

The uninterpreted sort.

mkUninterpretedSortConstructorSort ( self , size_t arity , unicode symbol=None )

Create a sort constructor sort.

An uninterpreted sort constructor is an uninterpreted sort with arity > 0.

Parameters :
  • symbol – The symbol of the sort.

  • arity – The arity of the sort (must be > 0).

Returns :

The sort constructor sort.

mkUniverseSet ( self , Sort sort )

Create a universe set of the given sort.

Parameters :

sort – The sort of the set elements

Returns :

The universe set constant

mkUnresolvedDatatypeSort ( self , unicode name , size_t arity=0 )

Create an unresolved datatype sort.

This is for creating yet unresolved sort placeholders for mutually recursive datatypes.

Parameters :
  • symbol – The name of the sort.

  • arity – The number of sort parameters of the sort.

Returns :

The unresolved sort.

mkVar ( self , Sort sort , symbol=None )

Create a bound variable to be used in a binder (i.e. a quantifier, a lambda, or a witness binder).

Parameters :
  • sort – The sort of the variable.

  • symbol – The name of the variable.

Returns :

The variable.

pop ( self , nscopes = 1 )

Pop nscopes level(s) from the assertion stack.

SMT-LIB:

( pop <numeral> )
Parameters :

nscopes – The number of levels to pop.

push ( self , nscopes = 1 )

Push nscopes level(s) to the assertion stack.

SMT-LIB:

( push <numeral> )
Parameters :

nscopes – The number of levels to push.

resetAssertions ( self )

Remove all assertions.

SMT-LIB:

( reset-assertions )
setInfo ( self , unicode keyword , unicode value )

Set info.

SMT-LIB:

( set-info <attribute> )
Parameters :
  • keyword – The info flag.

  • value – The value of the info flag.

setLogic ( self , unicode logic )

Set logic.

SMT-LIB:

( set-logic <symbol> )
Parameters :

logic – The logic to set.

setOption ( self , unicode option , unicode value )

Set option.

SMT-LIB:

( set-option <option> )
Parameters :
  • option – The option name.

  • value – The option value.

simplify ( self , Term t )

Simplify a formula without doing “much” work. Does not involve the SAT Engine in the simplification, but uses the current definitions, assertions, and the current partial model, if one has been constructed. It also involves theory normalization.

Warning

This method is experimental and may change in future versions.

Parameters :

t – The formula to simplify.

Returns :

The simplified formula.

synthFun ( self , unicode symbol , bound_vars , Sort sort , Grammar grammar=None )

Synthesize n-ary function following specified syntactic constraints.

SyGuS v2:

( synth-fun <symbol> ( <boundVars>* ) <sort> <g> )
Parameters :
  • symbol – The name of the function.

  • boundVars – The parameters to this function.

  • sort – The sort of the return value of this function.

  • grammar – The syntactic constraints.

Returns :

The function.

synthInv ( self , symbol , bound_vars , Grammar grammar=None )

Synthesize invariant.

SyGuS v2:

( synth-inv <symbol> ( <boundVars>* ) <grammar> )
Parameters :
  • symbol – The name of the invariant.

  • boundVars – The parameters to this invariant.

  • grammar – The syntactic constraints.

Returns :

The invariant.