SlideShare a Scribd company logo
1 of 23
Download to read offline
Satisfiability Modulo Theories
                      Lezione 1 - Overview
                   (slides revision: Thursday 20th October, 2011, 16:33)




                            Roberto Bruttomesso

                   Seminario di Logica Matematica
                     (Corso Prof. Silvio Ghilardi)

                                 20 Ottobre 2011




R. Bruttomesso (SMT)                    Overview                           20 Ottobre 2011   1 / 23
FAQ

Le slides seguono la dispensa (in fase di scrittura !)

                         “Satisfiability Modulo Theories”

e sono entrambi disponibili da
http://www.oprover.org/roberto/teaching/smt, dove trovate
anche i puntatori ai tool e agli esempi usati qui

Per chi cercasse un libro su questi argomenti, forse quello che si
avvicina di piu’ e’

Decision Procedures - An Algorithmic Point of View

(www.decision-procedures.org)

Per ricevimento la mia stanza e’ la S206, II piano, via Comelico, previa
richiesta via email roberto.bruttomesso@gmail.com

  R. Bruttomesso (SMT)               Overview              20 Ottobre 2011   2 / 23
Outline




1   A gentle introduction to SMT
     Introduction
     The Eager and the Lazy approaches


2   SMT-LIB and SMT-solvers
     SMT-LIB
     SMT-solvers




    R. Bruttomesso (SMT)      Overview   20 Ottobre 2011   3 / 23
Intro


Satisfiability Modulo Theories, SMT, studies practical methods to solve logical formulæ




These formulæ are defined/interpreted modulo a background theory, such as Linear Real
Arithmetic (LRA), Arrays (A), Bit-Vectors (BV), etc.




For instance, we want to determine the satisfiability modulo LRA of

                  (x + y ≤ 0) ∧ (x = 0) ∧ (¬a ∨ (x = 1) ∨ (y ≥ 0))                          (1)

where x, y are arithmetic variables, while a is a Boolean variable. Intuitively, (1) is
satisfiable iff we can find values for x and y in R and for a in B such that it evaluates to




   R. Bruttomesso (SMT)                   Overview                    20 Ottobre 2011       4 / 23
Definitions (syntax)

In SMT a theory T is defined over a signature Σ, which is a set of
function and predicate symbols such as {0, 1, . . . , +, −, . . . , ≤}. The
equality symbol = is assumed to be included in every signature.

Variables and function symbols in Σ can be used to build
theory-terms (T -term): a T -term is either a variable or, recursively,
an application of a function symbol in Σ to terms

Predicate symbols in Σ can be used to build theory-atoms (T -atom):
a T -atom is the application of a predicate symbol in Σ to T -terms

A theory-literal (T -literal) is either a T -atom or its negation

A formula is any Boolean combination of T -atoms and Boolean
variables

  R. Bruttomesso (SMT)              Overview               20 Ottobre 2011    5 / 23
Definitions (semantic)
In SMT the interpretation of the symbols in Σ is fixed, and it corresponds to the usual
semantic of the operators. For instance, in LIA (Linear Integer Arithmetic):
       numerals are mapped to the corresponding value in Z
       + is interpreted as the function
                                            (0, 0) → 0
                                            (0, 1) → 1
                                            ...

The only unspecified entities are variables, for which we have to build an assignment, a
mapping from variables to concrete values in Z
Now everything is specified and we can evaluate T -terms, T -atoms and formulæ. For
instance, the T -atom

                      (x + y ≤ 0) ∧ (x = 0) ∧ (¬a ∨ (x = 1) ∨ (y ≥ 0))

evaluates to  under the assignment {x → 0, y → 0, a → ⊥}, and it evaluates to ⊥ under
the assignment {x → 5, y → −10, a → ⊥}

We say that a formula ϕ is satisfiable modulo T , if there is an assignment M that evaluates
ϕ to   . In that case we say that M is a model

   R. Bruttomesso (SMT)                   Overview                  20 Ottobre 2011    6 / 23
Solving SMT formulæ by reduction to SAT
Approaches to solve SMT formulæ are based on the observation that SMT can be reduced
to SAT, i.e., the purely Boolean Satisfiability Problem
Consider for instance the LIA formula

              ϕ ≡ (x − y ≤ 0) ∧ (y − z ≤ 0) ∧ ((z − x ≤ −1) ∨ (z − x ≤ −2))

We may use a Boolean variable a to mean “x − y ≤ 0” evaluates to        in the model.
Similarly we could use b, c, d for the other T -atoms.
First of all, we notice that it does not hold in LIA that

                          x−y ≤0        y−z ≤0           z − x ≤ −1

evaluate to   at the same time, because this is not possible in LIA. This translates to
the Boolean relation
                                    ¬(a ∧ b ∧ c)
Similarly we may derive

                   ¬(a ∧ b ∧ d)    ¬(¬a ∧ ¬b ∧ ¬c)         ¬(¬a ∧ ¬b ∧ ¬d)


Moreover, because of the “structure” of ϕ, it holds that

                                       a ∧ b ∧ (c ∨ d)

   R. Bruttomesso (SMT)                   Overview                     20 Ottobre 2011   7 / 23
Solving SMT formulæ by reduction to SAT

               ϕ ≡ (x − y ≤ 0) ∧ (y − z ≤ 0) ∧ ((z − x ≤ −1) ∨ (z − x ≤ −2))

                                 a represents x − y ≤ 0
                                 b represents y − z ≤ 0
                                 c represents z − x ≤ −1
                                 d represents z − x ≤ −2


Putting all the conditions together we get the Boolean formula

ψ ≡ a ∧ b ∧ (c ∨ d) ∧ ¬(a ∧ b ∧ c) ∧ ¬(a ∧ b ∧ d) ∧ ¬(¬a ∧ ¬b ∧ ¬c) ∧ ¬(¬a ∧ ¬b ∧ ¬d)

Because of our translation, we have that ϕ is LIA-satisfiable if and only if ψ is satisfiable.
This is true because
  1 we have exhaustively encoded incompatible relations between T -atoms
  2 we have encoded the structure of ϕ
Therefore we may run any off-the-shelf SAT-solver to determine the satisfiability of ψ (and
therefore that of ϕ)


   R. Bruttomesso (SMT)                   Overview                    20 Ottobre 2011    8 / 23
Exercizes

  1   Check that ϕ is LIA-unsatisfiable, and that ψ is also unsatisfiable

  2   Check that ¬(¬a ∧ ¬b ∧ ¬c) and ¬(¬a ∧ ¬b ∧ ¬d) are actually
      redundant in ψ. Why it is so ?

  3   Substitute z − x ≤ −2 with z − x ≤ 2 into ϕ, recompute the correct
      ψ, and check that ϕ is LIA-satisfiable and that ψ is also satisfiable

  4   Prove that the encoding into SAT is correct and complete, i.e.,
      that if
        (i) we have exhaustively encoded incompatible relations between
            T -atoms
       (ii) we have encoded the structure of ϕ
      then ϕ is T -satisfiable if and only if ψ is satisfiable


  R. Bruttomesso (SMT)             Overview               20 Ottobre 2011   9 / 23
The Eager and Lazy Approaches

Recall that in our reduction to SAT we need to encode
  (i) incompatible relations between T -atoms exhaustively
 (ii) the structure of ϕ

Condition (ii) is easy to encode. The critical condition is (i). If we
have 3 T -atoms a, b, c, then we need to check whether
    a and b are incompatible
    a and ¬b are incompatible
    ...
    a and b and c are incompatible
    ...

Potentially, this leads to checking O(2n ) relations, if n T -atoms are in
the formula

  R. Bruttomesso (SMT)            Overview              20 Ottobre 2011   10 / 23
The Eager and Lazy Approaches
There are (at least) two ways to discover incompatibilities
    eagerly adding them before calling a SAT-solver (eager approach)
             + Easy to implement: SAT-solver used as black-box
             + Good for bit-vectors theories
             − Potentially generates too big encoding: needs
                heuristics to make it efficient
             − Bad for arithmetic theories
    lazily, by adding them during the SAT-solver’s search (lazy
    approach)
              + Generates smaller encodings
              + Good for arithmetic theories
              + Modular approach: allows easy theory combination
              − Trickier to implement: SAT-solver has to be “openen”
Most of this course will be devoted to the lazy approach, which is
nowadays the most successful technique available
  R. Bruttomesso (SMT)           Overview            20 Ottobre 2011   11 / 23
The Eager Approach


          SMT formula ϕ




                                    ψ
                        Encoder               SAT-solver




                                             sat / unsat



 R. Bruttomesso (SMT)             Overview                 20 Ottobre 2011   12 / 23
The Lazy Approach

     SMT formula ϕ




               Extract Struct.


                                 not good
                 SAT-solver                  T -solver

                                 Candidate
                                  model
                                                     good


                    unsat                      sat

 R. Bruttomesso (SMT)             Overview           20 Ottobre 2011   13 / 23
Plan of (the rest of) the course

    The eager approach: solving bit-vectors

    Modern SAT-solvers: conflict analysis, clause learning, and
    heuristics

    The Lazy approach: generalities

    A theory-solver for IDL

    A theory-solver for UF

    A theory-solver for LRA

    (see if there is time left)


  R. Bruttomesso (SMT)            Overview         20 Ottobre 2011   14 / 23
Outline




1   A gentle introduction to SMT
     Introduction
     The Eager and the Lazy approaches


2   SMT-LIB and SMT-solvers
     SMT-LIB
     SMT-solvers




    R. Bruttomesso (SMT)      Overview   20 Ottobre 2011   15 / 23
SMT-LIB (v2) http://www.smtlib.org

The SMT-LIB initiative
    defines a standard input language for SMT-solvers
    defines theories and logics in which formulæ can be written
    collects benchmarks

The SMT-LIB language allows to write formulæ in a lisp-like format.
E.g.:

     (< (+ x y) 0)
     (= (f x y) (g z))


stand for x + y < 0 and f (x, y) = g(z) respectively
An SMT-LIB file looks more similar to a set of commands for an
SMT-solver, rather then a logic formula

  R. Bruttomesso (SMT)           Overview              20 Ottobre 2011   16 / 23
SMT-LIB Theories
An SMT-LIB theory consists of some sorts, (e.g., Int) and of some functions (e.g., −, +).
Predicates are also considered functions, with codomain in Bool (e.g., <, ≤). For instance
     (theory Ints
                                                      (theory Core
         :sorts ((Int 0))
                                                          :sorts ((Bool 0))
         :funs ((NUMERAL Int)
                 (- Int Int)                              :funs ((true Bool)
                 (- Int Int Int :left-assoc)                     (false Bool)
                 (+ Int Int Int :left-assoc)                     (not Bool Bool)
                 (* Int Int Int :left-assoc)                     (=> Bool Bool Bool :right-assoc)
                 (div Int Int Int :left-assoc)                   (and Bool Bool Bool :left-assoc)
                 (mod Int Int Int)                               (or Bool Bool Bool :left-assoc)
                 (abs Int Int)                                   (xor Bool Bool Bool :left-assoc)
                 (<= Int Int Bool :chainable)                    (par (A) (= A A Bool :chainable))
                 (< Int Int Bool :chainable)                     (par (A) (distinct A A Bool :pairwise))
                 (>= Int Int Bool :chainable)                    (par (A) (ite Bool A A A))
                 (> Int Int Bool :chainable)                    )
               )
                                                          [...]
          [...]                                       )
     )

                       These definitions can be found at www.smtlib.org

The sorts and the function symbols declared in a theory are always interpreted. This
means that a to specify a model for a formula ϕ, we just need to specify the assignment of
the variables to the concrete values in the sorts.

   R. Bruttomesso (SMT)                          Overview                         20 Ottobre 2011          17 / 23
SMT-LIB Logics
    The difference between “logic” and “theory” might look very subtle. An SMT-LIB logic
    includes a theory definition, plus it describes some restrictions on how formulæ can be built.
(logic QF_LIA
                                                         (logic QF_IDL
    :theories (Ints)
                                                             :theories (Ints)
    :language
    "Closed quantifier-free formulas built
                                                             :language
    over an arbitrary expansion of the
                                                             "Closed quantifier-free formulas with
    Ints signature with free constant symbols,
                                                             atoms of the form:
    but whose terms of sort Int are all linear,
                                                             - q
    that is, have no occurrences of the function
                                                             - (op (- x y) n),
    symbols *, /, div, mod, and abs, except as
                                                             - (op (- x y) (- n)), or
    specified the :extensions attribute.
                                                             - (op x y)
    "
                                                             where
                                                             - q is a variable or free constant symbol of sort Bool,
    :extensions
                                                             - op is <, <=, >, >=, =, or distinct,
    "Terms with _concrete_ coefficients are also
                                                             - x, y are free constant symbols of sort Int,
    allowed, that is, terms of the form c, (* c x),
                                                             - n is a numeral.
    or (* x c) where x is a free constant and c
                                                             "
    is a term of the form n or (- n) for some numeral n.
                                                         )
    "
)


    In the following we will not be so strict, and we will not make any distinction between
    “theories” and “logics”, calling both “theories”. For instance when we will say that we
    reason modulo the theory LIA we mean that we are working with QF LIA formulæ

       R. Bruttomesso (SMT)                          Overview                          20 Ottobre 2011      18 / 23
Writing an SMT-LIB file
The logic can be specified with the command
    (set-logic QF_LIA)

Variables are declared with
    (declare-fun x ( ) Int)

A formula is specified with
    (assert (<= (+ x y) 0))

Asks the tool to compute satisfiability of assertions
    (check-sat)

Asks the tool to return a model (in case of sat result)
    (set-option :produce-models true)
    ...
    (get-value (x y))

Disable annoying printouts
    (set-option :print-success false)
   R. Bruttomesso (SMT)                    Overview       20 Ottobre 2011   19 / 23
Example

     (set-logic QF_LIA)
     (declare-fun x ( ) Int)
     (declare-fun y ( ) Int)
     (declare-fun a ( ) Bool)
     (assert (<= (+ x y) 0))
     (assert (= x 0))
     (assert (or (not a) (= x 1) (>= y 0)))
     (assert (not (= (+ y 1) 0)))
     (check-sat)
     (exit)


which stands for the LIA formula

   (x + y ≤ 0) ∧ (x = 0) ∧ ((¬a ∨ (x = 1) ∨ (y ≥ 0)) ∧ ¬(y + 1 = 0)

  R. Bruttomesso (SMT)         Overview             20 Ottobre 2011   20 / 23
SMT-solvers
An SMT-solver is a tool that can parse and solve an SMT-LIB
benchmark.
There are many such tools available online. In this course we will use
Yices (developed at SRI, Stanford, closed source), Z3 (developed at
MSR, Redmond, closed source) and OpenSMT (developed here, open
source). Other available tools are MathSAT, CVC4, Boolector,
veriT, STP.

   roberto@moriarty:examples$ smtlib2yices < test1.smt2
   success
   success
   success
   success
   success
   success
   success
   success
   sat



  R. Bruttomesso (SMT)             Overview               20 Ottobre 2011   21 / 23
SMT-LIB script
The SMT-LIB language allows specification of scripts. A script is a
benchmark that may contain many check-sat commands. Also, it
may include push and pop commands which can be used to control the
assertion stack
   (set-option :print-success false)
   (set-logic QF_LIA)
   (declare-fun x ( ) Int)
   (declare-fun y ( ) Int)
   (assert (<= (+ x y) 0))
   (assert (= x 0))
   (assert (or (= x 1) (>= y 0)))
   (check-sat)
   (push 1)
   (assert (not (= y 0)))
   (check-sat)
   (pop 1)
   (check-sat)
   (exit)



  R. Bruttomesso (SMT)             Overview      20 Ottobre 2011   22 / 23
Exercizes

  1   Translate the following LIA formula SMT-LIB, and evaluate it
      with an SMT-solver

          (x − y ≤ 0) ∧ (y − z ≤ 0) ∧ ((z − x ≤ −1) ∨ (z − x ≤ −2))


  2   Translate the following LRA formula SMT-LIB, and evaluate it
      with an SMT-solver

                         (b ∨ (x + y ≤ 0)) ∧ (¬b ∨ (x + z ≤ 10))


  3   For the satisfiable formulæ above print out a model

  4   For the satisfiable formulæ above, add constraints such that they
      become unsatisfiable

  R. Bruttomesso (SMT)                Overview              20 Ottobre 2011   23 / 23

More Related Content

What's hot

Formal methods 1 - introduction
Formal methods   1 - introductionFormal methods   1 - introduction
Formal methods 1 - introductionVlad Patryshev
 
Proof-Theoretic Semantics: Point-free meaninig of first-order systems
Proof-Theoretic Semantics: Point-free meaninig of first-order systemsProof-Theoretic Semantics: Point-free meaninig of first-order systems
Proof-Theoretic Semantics: Point-free meaninig of first-order systemsMarco Benini
 
A brief introduction to Hartree-Fock and TDDFT
A brief introduction to Hartree-Fock and TDDFTA brief introduction to Hartree-Fock and TDDFT
A brief introduction to Hartree-Fock and TDDFTJiahao Chen
 
Introduction to Max-SAT and Max-SAT Evaluation
Introduction to Max-SAT and Max-SAT EvaluationIntroduction to Max-SAT and Max-SAT Evaluation
Introduction to Max-SAT and Max-SAT EvaluationMasahiro Sakai
 
Lesson 16: Inverse Trigonometric Functions
Lesson 16: Inverse Trigonometric FunctionsLesson 16: Inverse Trigonometric Functions
Lesson 16: Inverse Trigonometric FunctionsMatthew Leingang
 
Writing a SAT solver as a hobby project
Writing a SAT solver as a hobby projectWriting a SAT solver as a hobby project
Writing a SAT solver as a hobby projectMasahiro Sakai
 
Free Ebooks Download ! Edhole
Free Ebooks Download ! EdholeFree Ebooks Download ! Edhole
Free Ebooks Download ! EdholeEdhole.com
 
P805 bourgeois
P805 bourgeoisP805 bourgeois
P805 bourgeoiskklub
 
Continuity of a Function
Continuity of a Function Continuity of a Function
Continuity of a Function Vishvesh Jasani
 
Lesson 16: Inverse Trigonometric Functions (Section 041 slides)
Lesson 16: Inverse Trigonometric Functions (Section 041 slides)Lesson 16: Inverse Trigonometric Functions (Section 041 slides)
Lesson 16: Inverse Trigonometric Functions (Section 041 slides)Matthew Leingang
 
Computability - Tractable, Intractable and Non-computable Function
Computability - Tractable, Intractable and Non-computable FunctionComputability - Tractable, Intractable and Non-computable Function
Computability - Tractable, Intractable and Non-computable FunctionReggie Niccolo Santos
 
Noisy Optimization combining Bandits and Evolutionary Algorithms
Noisy Optimization combining Bandits and Evolutionary AlgorithmsNoisy Optimization combining Bandits and Evolutionary Algorithms
Noisy Optimization combining Bandits and Evolutionary AlgorithmsOlivier Teytaud
 
Asymptotic Notations
Asymptotic NotationsAsymptotic Notations
Asymptotic NotationsNagendraK18
 
A Numerical Analytic Continuation and Its Application to Fourier Transform
A Numerical Analytic Continuation and Its Application to Fourier TransformA Numerical Analytic Continuation and Its Application to Fourier Transform
A Numerical Analytic Continuation and Its Application to Fourier TransformHidenoriOgata
 
Doering Savov
Doering SavovDoering Savov
Doering Savovgh
 
NEIU Permutation Algebra
NEIU Permutation AlgebraNEIU Permutation Algebra
NEIU Permutation Algebrachrislt521
 
Scala: Pattern matching, Concepts and Implementations
Scala: Pattern matching, Concepts and ImplementationsScala: Pattern matching, Concepts and Implementations
Scala: Pattern matching, Concepts and ImplementationsMICHRAFY MUSTAFA
 
Algorithm chapter 10
Algorithm chapter 10Algorithm chapter 10
Algorithm chapter 10chidabdu
 
Inversion Theorem for Generalized Fractional Hilbert Transform
Inversion Theorem for Generalized Fractional Hilbert TransformInversion Theorem for Generalized Fractional Hilbert Transform
Inversion Theorem for Generalized Fractional Hilbert Transforminventionjournals
 

What's hot (20)

Formal methods 1 - introduction
Formal methods   1 - introductionFormal methods   1 - introduction
Formal methods 1 - introduction
 
Proof-Theoretic Semantics: Point-free meaninig of first-order systems
Proof-Theoretic Semantics: Point-free meaninig of first-order systemsProof-Theoretic Semantics: Point-free meaninig of first-order systems
Proof-Theoretic Semantics: Point-free meaninig of first-order systems
 
A brief introduction to Hartree-Fock and TDDFT
A brief introduction to Hartree-Fock and TDDFTA brief introduction to Hartree-Fock and TDDFT
A brief introduction to Hartree-Fock and TDDFT
 
Introduction to Max-SAT and Max-SAT Evaluation
Introduction to Max-SAT and Max-SAT EvaluationIntroduction to Max-SAT and Max-SAT Evaluation
Introduction to Max-SAT and Max-SAT Evaluation
 
Lesson 16: Inverse Trigonometric Functions
Lesson 16: Inverse Trigonometric FunctionsLesson 16: Inverse Trigonometric Functions
Lesson 16: Inverse Trigonometric Functions
 
Writing a SAT solver as a hobby project
Writing a SAT solver as a hobby projectWriting a SAT solver as a hobby project
Writing a SAT solver as a hobby project
 
Free Ebooks Download ! Edhole
Free Ebooks Download ! EdholeFree Ebooks Download ! Edhole
Free Ebooks Download ! Edhole
 
P805 bourgeois
P805 bourgeoisP805 bourgeois
P805 bourgeois
 
Continuity of a Function
Continuity of a Function Continuity of a Function
Continuity of a Function
 
Computability
Computability Computability
Computability
 
Lesson 16: Inverse Trigonometric Functions (Section 041 slides)
Lesson 16: Inverse Trigonometric Functions (Section 041 slides)Lesson 16: Inverse Trigonometric Functions (Section 041 slides)
Lesson 16: Inverse Trigonometric Functions (Section 041 slides)
 
Computability - Tractable, Intractable and Non-computable Function
Computability - Tractable, Intractable and Non-computable FunctionComputability - Tractable, Intractable and Non-computable Function
Computability - Tractable, Intractable and Non-computable Function
 
Noisy Optimization combining Bandits and Evolutionary Algorithms
Noisy Optimization combining Bandits and Evolutionary AlgorithmsNoisy Optimization combining Bandits and Evolutionary Algorithms
Noisy Optimization combining Bandits and Evolutionary Algorithms
 
Asymptotic Notations
Asymptotic NotationsAsymptotic Notations
Asymptotic Notations
 
A Numerical Analytic Continuation and Its Application to Fourier Transform
A Numerical Analytic Continuation and Its Application to Fourier TransformA Numerical Analytic Continuation and Its Application to Fourier Transform
A Numerical Analytic Continuation and Its Application to Fourier Transform
 
Doering Savov
Doering SavovDoering Savov
Doering Savov
 
NEIU Permutation Algebra
NEIU Permutation AlgebraNEIU Permutation Algebra
NEIU Permutation Algebra
 
Scala: Pattern matching, Concepts and Implementations
Scala: Pattern matching, Concepts and ImplementationsScala: Pattern matching, Concepts and Implementations
Scala: Pattern matching, Concepts and Implementations
 
Algorithm chapter 10
Algorithm chapter 10Algorithm chapter 10
Algorithm chapter 10
 
Inversion Theorem for Generalized Fractional Hilbert Transform
Inversion Theorem for Generalized Fractional Hilbert TransformInversion Theorem for Generalized Fractional Hilbert Transform
Inversion Theorem for Generalized Fractional Hilbert Transform
 

Similar to Overview of Satisfiability Modulo Theories (SMT

Winter 10 Undecidability.pptx
Winter 10 Undecidability.pptxWinter 10 Undecidability.pptx
Winter 10 Undecidability.pptxHarisPrince
 
from_data_to_differential_equations.ppt
from_data_to_differential_equations.pptfrom_data_to_differential_equations.ppt
from_data_to_differential_equations.pptashutoshvb1
 
2_GLMs_printable.pdf
2_GLMs_printable.pdf2_GLMs_printable.pdf
2_GLMs_printable.pdfElio Laureano
 
Herbrand-satisfiability of a Quantified Set-theoretical Fragment (Cantone, Lo...
Herbrand-satisfiability of a Quantified Set-theoretical Fragment (Cantone, Lo...Herbrand-satisfiability of a Quantified Set-theoretical Fragment (Cantone, Lo...
Herbrand-satisfiability of a Quantified Set-theoretical Fragment (Cantone, Lo...Cristiano Longo
 
Workshop presentations l_bworkshop_reis
Workshop presentations l_bworkshop_reisWorkshop presentations l_bworkshop_reis
Workshop presentations l_bworkshop_reisTim Reis
 
Accelerated approximate Bayesian computation with applications to protein fol...
Accelerated approximate Bayesian computation with applications to protein fol...Accelerated approximate Bayesian computation with applications to protein fol...
Accelerated approximate Bayesian computation with applications to protein fol...Umberto Picchini
 
Intro to Approximate Bayesian Computation (ABC)
Intro to Approximate Bayesian Computation (ABC)Intro to Approximate Bayesian Computation (ABC)
Intro to Approximate Bayesian Computation (ABC)Umberto Picchini
 

Similar to Overview of Satisfiability Modulo Theories (SMT (20)

20100822 opensmt bruttomesso
20100822 opensmt bruttomesso20100822 opensmt bruttomesso
20100822 opensmt bruttomesso
 
smtlecture.5
smtlecture.5smtlecture.5
smtlecture.5
 
smtlecture.10
smtlecture.10smtlecture.10
smtlecture.10
 
smtlecture.7
smtlecture.7smtlecture.7
smtlecture.7
 
smtlecture.6
smtlecture.6smtlecture.6
smtlecture.6
 
20100822 satandsmt bruttomesso
20100822 satandsmt bruttomesso20100822 satandsmt bruttomesso
20100822 satandsmt bruttomesso
 
Winter 10 Undecidability.pptx
Winter 10 Undecidability.pptxWinter 10 Undecidability.pptx
Winter 10 Undecidability.pptx
 
smtlecture.3
smtlecture.3smtlecture.3
smtlecture.3
 
Signals and Systems Assignment Help
Signals and Systems Assignment HelpSignals and Systems Assignment Help
Signals and Systems Assignment Help
 
from_data_to_differential_equations.ppt
from_data_to_differential_equations.pptfrom_data_to_differential_equations.ppt
from_data_to_differential_equations.ppt
 
Signal Processing Homework Help
Signal Processing Homework HelpSignal Processing Homework Help
Signal Processing Homework Help
 
2_GLMs_printable.pdf
2_GLMs_printable.pdf2_GLMs_printable.pdf
2_GLMs_printable.pdf
 
Herbrand-satisfiability of a Quantified Set-theoretical Fragment (Cantone, Lo...
Herbrand-satisfiability of a Quantified Set-theoretical Fragment (Cantone, Lo...Herbrand-satisfiability of a Quantified Set-theoretical Fragment (Cantone, Lo...
Herbrand-satisfiability of a Quantified Set-theoretical Fragment (Cantone, Lo...
 
satandsmt.stpetersburg
satandsmt.stpetersburgsatandsmt.stpetersburg
satandsmt.stpetersburg
 
Workshop presentations l_bworkshop_reis
Workshop presentations l_bworkshop_reisWorkshop presentations l_bworkshop_reis
Workshop presentations l_bworkshop_reis
 
APAL2032
APAL2032APAL2032
APAL2032
 
Big o
Big oBig o
Big o
 
Signals and Systems Assignment Help
Signals and Systems Assignment HelpSignals and Systems Assignment Help
Signals and Systems Assignment Help
 
Accelerated approximate Bayesian computation with applications to protein fol...
Accelerated approximate Bayesian computation with applications to protein fol...Accelerated approximate Bayesian computation with applications to protein fol...
Accelerated approximate Bayesian computation with applications to protein fol...
 
Intro to Approximate Bayesian Computation (ABC)
Intro to Approximate Bayesian Computation (ABC)Intro to Approximate Bayesian Computation (ABC)
Intro to Approximate Bayesian Computation (ABC)
 

Recently uploaded

“Oh GOSH! Reflecting on Hackteria's Collaborative Practices in a Global Do-It...
“Oh GOSH! Reflecting on Hackteria's Collaborative Practices in a Global Do-It...“Oh GOSH! Reflecting on Hackteria's Collaborative Practices in a Global Do-It...
“Oh GOSH! Reflecting on Hackteria's Collaborative Practices in a Global Do-It...Marc Dusseiller Dusjagr
 
Employee wellbeing at the workplace.pptx
Employee wellbeing at the workplace.pptxEmployee wellbeing at the workplace.pptx
Employee wellbeing at the workplace.pptxNirmalaLoungPoorunde1
 
SOCIAL AND HISTORICAL CONTEXT - LFTVD.pptx
SOCIAL AND HISTORICAL CONTEXT - LFTVD.pptxSOCIAL AND HISTORICAL CONTEXT - LFTVD.pptx
SOCIAL AND HISTORICAL CONTEXT - LFTVD.pptxiammrhaywood
 
Mastering the Unannounced Regulatory Inspection
Mastering the Unannounced Regulatory InspectionMastering the Unannounced Regulatory Inspection
Mastering the Unannounced Regulatory InspectionSafetyChain Software
 
Industrial Policy - 1948, 1956, 1973, 1977, 1980, 1991
Industrial Policy - 1948, 1956, 1973, 1977, 1980, 1991Industrial Policy - 1948, 1956, 1973, 1977, 1980, 1991
Industrial Policy - 1948, 1956, 1973, 1977, 1980, 1991RKavithamani
 
Q4-W6-Restating Informational Text Grade 3
Q4-W6-Restating Informational Text Grade 3Q4-W6-Restating Informational Text Grade 3
Q4-W6-Restating Informational Text Grade 3JemimahLaneBuaron
 
mini mental status format.docx
mini    mental       status     format.docxmini    mental       status     format.docx
mini mental status format.docxPoojaSen20
 
How to Make a Pirate ship Primary Education.pptx
How to Make a Pirate ship Primary Education.pptxHow to Make a Pirate ship Primary Education.pptx
How to Make a Pirate ship Primary Education.pptxmanuelaromero2013
 
Incoming and Outgoing Shipments in 1 STEP Using Odoo 17
Incoming and Outgoing Shipments in 1 STEP Using Odoo 17Incoming and Outgoing Shipments in 1 STEP Using Odoo 17
Incoming and Outgoing Shipments in 1 STEP Using Odoo 17Celine George
 
Crayon Activity Handout For the Crayon A
Crayon Activity Handout For the Crayon ACrayon Activity Handout For the Crayon A
Crayon Activity Handout For the Crayon AUnboundStockton
 
APM Welcome, APM North West Network Conference, Synergies Across Sectors
APM Welcome, APM North West Network Conference, Synergies Across SectorsAPM Welcome, APM North West Network Conference, Synergies Across Sectors
APM Welcome, APM North West Network Conference, Synergies Across SectorsAssociation for Project Management
 
Alper Gobel In Media Res Media Component
Alper Gobel In Media Res Media ComponentAlper Gobel In Media Res Media Component
Alper Gobel In Media Res Media ComponentInMediaRes1
 
Organic Name Reactions for the students and aspirants of Chemistry12th.pptx
Organic Name Reactions  for the students and aspirants of Chemistry12th.pptxOrganic Name Reactions  for the students and aspirants of Chemistry12th.pptx
Organic Name Reactions for the students and aspirants of Chemistry12th.pptxVS Mahajan Coaching Centre
 
POINT- BIOCHEMISTRY SEM 2 ENZYMES UNIT 5.pptx
POINT- BIOCHEMISTRY SEM 2 ENZYMES UNIT 5.pptxPOINT- BIOCHEMISTRY SEM 2 ENZYMES UNIT 5.pptx
POINT- BIOCHEMISTRY SEM 2 ENZYMES UNIT 5.pptxSayali Powar
 
The Most Excellent Way | 1 Corinthians 13
The Most Excellent Way | 1 Corinthians 13The Most Excellent Way | 1 Corinthians 13
The Most Excellent Way | 1 Corinthians 13Steve Thomason
 
Software Engineering Methodologies (overview)
Software Engineering Methodologies (overview)Software Engineering Methodologies (overview)
Software Engineering Methodologies (overview)eniolaolutunde
 
Introduction to AI in Higher Education_draft.pptx
Introduction to AI in Higher Education_draft.pptxIntroduction to AI in Higher Education_draft.pptx
Introduction to AI in Higher Education_draft.pptxpboyjonauth
 
CARE OF CHILD IN INCUBATOR..........pptx
CARE OF CHILD IN INCUBATOR..........pptxCARE OF CHILD IN INCUBATOR..........pptx
CARE OF CHILD IN INCUBATOR..........pptxGaneshChakor2
 
BASLIQ CURRENT LOOKBOOK LOOKBOOK(1) (1).pdf
BASLIQ CURRENT LOOKBOOK  LOOKBOOK(1) (1).pdfBASLIQ CURRENT LOOKBOOK  LOOKBOOK(1) (1).pdf
BASLIQ CURRENT LOOKBOOK LOOKBOOK(1) (1).pdfSoniaTolstoy
 

Recently uploaded (20)

“Oh GOSH! Reflecting on Hackteria's Collaborative Practices in a Global Do-It...
“Oh GOSH! Reflecting on Hackteria's Collaborative Practices in a Global Do-It...“Oh GOSH! Reflecting on Hackteria's Collaborative Practices in a Global Do-It...
“Oh GOSH! Reflecting on Hackteria's Collaborative Practices in a Global Do-It...
 
Employee wellbeing at the workplace.pptx
Employee wellbeing at the workplace.pptxEmployee wellbeing at the workplace.pptx
Employee wellbeing at the workplace.pptx
 
SOCIAL AND HISTORICAL CONTEXT - LFTVD.pptx
SOCIAL AND HISTORICAL CONTEXT - LFTVD.pptxSOCIAL AND HISTORICAL CONTEXT - LFTVD.pptx
SOCIAL AND HISTORICAL CONTEXT - LFTVD.pptx
 
Mastering the Unannounced Regulatory Inspection
Mastering the Unannounced Regulatory InspectionMastering the Unannounced Regulatory Inspection
Mastering the Unannounced Regulatory Inspection
 
Industrial Policy - 1948, 1956, 1973, 1977, 1980, 1991
Industrial Policy - 1948, 1956, 1973, 1977, 1980, 1991Industrial Policy - 1948, 1956, 1973, 1977, 1980, 1991
Industrial Policy - 1948, 1956, 1973, 1977, 1980, 1991
 
Q4-W6-Restating Informational Text Grade 3
Q4-W6-Restating Informational Text Grade 3Q4-W6-Restating Informational Text Grade 3
Q4-W6-Restating Informational Text Grade 3
 
mini mental status format.docx
mini    mental       status     format.docxmini    mental       status     format.docx
mini mental status format.docx
 
How to Make a Pirate ship Primary Education.pptx
How to Make a Pirate ship Primary Education.pptxHow to Make a Pirate ship Primary Education.pptx
How to Make a Pirate ship Primary Education.pptx
 
Incoming and Outgoing Shipments in 1 STEP Using Odoo 17
Incoming and Outgoing Shipments in 1 STEP Using Odoo 17Incoming and Outgoing Shipments in 1 STEP Using Odoo 17
Incoming and Outgoing Shipments in 1 STEP Using Odoo 17
 
Crayon Activity Handout For the Crayon A
Crayon Activity Handout For the Crayon ACrayon Activity Handout For the Crayon A
Crayon Activity Handout For the Crayon A
 
APM Welcome, APM North West Network Conference, Synergies Across Sectors
APM Welcome, APM North West Network Conference, Synergies Across SectorsAPM Welcome, APM North West Network Conference, Synergies Across Sectors
APM Welcome, APM North West Network Conference, Synergies Across Sectors
 
Alper Gobel In Media Res Media Component
Alper Gobel In Media Res Media ComponentAlper Gobel In Media Res Media Component
Alper Gobel In Media Res Media Component
 
Organic Name Reactions for the students and aspirants of Chemistry12th.pptx
Organic Name Reactions  for the students and aspirants of Chemistry12th.pptxOrganic Name Reactions  for the students and aspirants of Chemistry12th.pptx
Organic Name Reactions for the students and aspirants of Chemistry12th.pptx
 
POINT- BIOCHEMISTRY SEM 2 ENZYMES UNIT 5.pptx
POINT- BIOCHEMISTRY SEM 2 ENZYMES UNIT 5.pptxPOINT- BIOCHEMISTRY SEM 2 ENZYMES UNIT 5.pptx
POINT- BIOCHEMISTRY SEM 2 ENZYMES UNIT 5.pptx
 
The Most Excellent Way | 1 Corinthians 13
The Most Excellent Way | 1 Corinthians 13The Most Excellent Way | 1 Corinthians 13
The Most Excellent Way | 1 Corinthians 13
 
Software Engineering Methodologies (overview)
Software Engineering Methodologies (overview)Software Engineering Methodologies (overview)
Software Engineering Methodologies (overview)
 
Model Call Girl in Tilak Nagar Delhi reach out to us at 🔝9953056974🔝
Model Call Girl in Tilak Nagar Delhi reach out to us at 🔝9953056974🔝Model Call Girl in Tilak Nagar Delhi reach out to us at 🔝9953056974🔝
Model Call Girl in Tilak Nagar Delhi reach out to us at 🔝9953056974🔝
 
Introduction to AI in Higher Education_draft.pptx
Introduction to AI in Higher Education_draft.pptxIntroduction to AI in Higher Education_draft.pptx
Introduction to AI in Higher Education_draft.pptx
 
CARE OF CHILD IN INCUBATOR..........pptx
CARE OF CHILD IN INCUBATOR..........pptxCARE OF CHILD IN INCUBATOR..........pptx
CARE OF CHILD IN INCUBATOR..........pptx
 
BASLIQ CURRENT LOOKBOOK LOOKBOOK(1) (1).pdf
BASLIQ CURRENT LOOKBOOK  LOOKBOOK(1) (1).pdfBASLIQ CURRENT LOOKBOOK  LOOKBOOK(1) (1).pdf
BASLIQ CURRENT LOOKBOOK LOOKBOOK(1) (1).pdf
 

Overview of Satisfiability Modulo Theories (SMT

  • 1. Satisfiability Modulo Theories Lezione 1 - Overview (slides revision: Thursday 20th October, 2011, 16:33) Roberto Bruttomesso Seminario di Logica Matematica (Corso Prof. Silvio Ghilardi) 20 Ottobre 2011 R. Bruttomesso (SMT) Overview 20 Ottobre 2011 1 / 23
  • 2. FAQ Le slides seguono la dispensa (in fase di scrittura !) “Satisfiability Modulo Theories” e sono entrambi disponibili da http://www.oprover.org/roberto/teaching/smt, dove trovate anche i puntatori ai tool e agli esempi usati qui Per chi cercasse un libro su questi argomenti, forse quello che si avvicina di piu’ e’ Decision Procedures - An Algorithmic Point of View (www.decision-procedures.org) Per ricevimento la mia stanza e’ la S206, II piano, via Comelico, previa richiesta via email roberto.bruttomesso@gmail.com R. Bruttomesso (SMT) Overview 20 Ottobre 2011 2 / 23
  • 3. Outline 1 A gentle introduction to SMT Introduction The Eager and the Lazy approaches 2 SMT-LIB and SMT-solvers SMT-LIB SMT-solvers R. Bruttomesso (SMT) Overview 20 Ottobre 2011 3 / 23
  • 4. Intro Satisfiability Modulo Theories, SMT, studies practical methods to solve logical formulæ These formulæ are defined/interpreted modulo a background theory, such as Linear Real Arithmetic (LRA), Arrays (A), Bit-Vectors (BV), etc. For instance, we want to determine the satisfiability modulo LRA of (x + y ≤ 0) ∧ (x = 0) ∧ (¬a ∨ (x = 1) ∨ (y ≥ 0)) (1) where x, y are arithmetic variables, while a is a Boolean variable. Intuitively, (1) is satisfiable iff we can find values for x and y in R and for a in B such that it evaluates to R. Bruttomesso (SMT) Overview 20 Ottobre 2011 4 / 23
  • 5. Definitions (syntax) In SMT a theory T is defined over a signature Σ, which is a set of function and predicate symbols such as {0, 1, . . . , +, −, . . . , ≤}. The equality symbol = is assumed to be included in every signature. Variables and function symbols in Σ can be used to build theory-terms (T -term): a T -term is either a variable or, recursively, an application of a function symbol in Σ to terms Predicate symbols in Σ can be used to build theory-atoms (T -atom): a T -atom is the application of a predicate symbol in Σ to T -terms A theory-literal (T -literal) is either a T -atom or its negation A formula is any Boolean combination of T -atoms and Boolean variables R. Bruttomesso (SMT) Overview 20 Ottobre 2011 5 / 23
  • 6. Definitions (semantic) In SMT the interpretation of the symbols in Σ is fixed, and it corresponds to the usual semantic of the operators. For instance, in LIA (Linear Integer Arithmetic): numerals are mapped to the corresponding value in Z + is interpreted as the function (0, 0) → 0 (0, 1) → 1 ... The only unspecified entities are variables, for which we have to build an assignment, a mapping from variables to concrete values in Z Now everything is specified and we can evaluate T -terms, T -atoms and formulæ. For instance, the T -atom (x + y ≤ 0) ∧ (x = 0) ∧ (¬a ∨ (x = 1) ∨ (y ≥ 0)) evaluates to under the assignment {x → 0, y → 0, a → ⊥}, and it evaluates to ⊥ under the assignment {x → 5, y → −10, a → ⊥} We say that a formula ϕ is satisfiable modulo T , if there is an assignment M that evaluates ϕ to . In that case we say that M is a model R. Bruttomesso (SMT) Overview 20 Ottobre 2011 6 / 23
  • 7. Solving SMT formulæ by reduction to SAT Approaches to solve SMT formulæ are based on the observation that SMT can be reduced to SAT, i.e., the purely Boolean Satisfiability Problem Consider for instance the LIA formula ϕ ≡ (x − y ≤ 0) ∧ (y − z ≤ 0) ∧ ((z − x ≤ −1) ∨ (z − x ≤ −2)) We may use a Boolean variable a to mean “x − y ≤ 0” evaluates to in the model. Similarly we could use b, c, d for the other T -atoms. First of all, we notice that it does not hold in LIA that x−y ≤0 y−z ≤0 z − x ≤ −1 evaluate to at the same time, because this is not possible in LIA. This translates to the Boolean relation ¬(a ∧ b ∧ c) Similarly we may derive ¬(a ∧ b ∧ d) ¬(¬a ∧ ¬b ∧ ¬c) ¬(¬a ∧ ¬b ∧ ¬d) Moreover, because of the “structure” of ϕ, it holds that a ∧ b ∧ (c ∨ d) R. Bruttomesso (SMT) Overview 20 Ottobre 2011 7 / 23
  • 8. Solving SMT formulæ by reduction to SAT ϕ ≡ (x − y ≤ 0) ∧ (y − z ≤ 0) ∧ ((z − x ≤ −1) ∨ (z − x ≤ −2)) a represents x − y ≤ 0 b represents y − z ≤ 0 c represents z − x ≤ −1 d represents z − x ≤ −2 Putting all the conditions together we get the Boolean formula ψ ≡ a ∧ b ∧ (c ∨ d) ∧ ¬(a ∧ b ∧ c) ∧ ¬(a ∧ b ∧ d) ∧ ¬(¬a ∧ ¬b ∧ ¬c) ∧ ¬(¬a ∧ ¬b ∧ ¬d) Because of our translation, we have that ϕ is LIA-satisfiable if and only if ψ is satisfiable. This is true because 1 we have exhaustively encoded incompatible relations between T -atoms 2 we have encoded the structure of ϕ Therefore we may run any off-the-shelf SAT-solver to determine the satisfiability of ψ (and therefore that of ϕ) R. Bruttomesso (SMT) Overview 20 Ottobre 2011 8 / 23
  • 9. Exercizes 1 Check that ϕ is LIA-unsatisfiable, and that ψ is also unsatisfiable 2 Check that ¬(¬a ∧ ¬b ∧ ¬c) and ¬(¬a ∧ ¬b ∧ ¬d) are actually redundant in ψ. Why it is so ? 3 Substitute z − x ≤ −2 with z − x ≤ 2 into ϕ, recompute the correct ψ, and check that ϕ is LIA-satisfiable and that ψ is also satisfiable 4 Prove that the encoding into SAT is correct and complete, i.e., that if (i) we have exhaustively encoded incompatible relations between T -atoms (ii) we have encoded the structure of ϕ then ϕ is T -satisfiable if and only if ψ is satisfiable R. Bruttomesso (SMT) Overview 20 Ottobre 2011 9 / 23
  • 10. The Eager and Lazy Approaches Recall that in our reduction to SAT we need to encode (i) incompatible relations between T -atoms exhaustively (ii) the structure of ϕ Condition (ii) is easy to encode. The critical condition is (i). If we have 3 T -atoms a, b, c, then we need to check whether a and b are incompatible a and ¬b are incompatible ... a and b and c are incompatible ... Potentially, this leads to checking O(2n ) relations, if n T -atoms are in the formula R. Bruttomesso (SMT) Overview 20 Ottobre 2011 10 / 23
  • 11. The Eager and Lazy Approaches There are (at least) two ways to discover incompatibilities eagerly adding them before calling a SAT-solver (eager approach) + Easy to implement: SAT-solver used as black-box + Good for bit-vectors theories − Potentially generates too big encoding: needs heuristics to make it efficient − Bad for arithmetic theories lazily, by adding them during the SAT-solver’s search (lazy approach) + Generates smaller encodings + Good for arithmetic theories + Modular approach: allows easy theory combination − Trickier to implement: SAT-solver has to be “openen” Most of this course will be devoted to the lazy approach, which is nowadays the most successful technique available R. Bruttomesso (SMT) Overview 20 Ottobre 2011 11 / 23
  • 12. The Eager Approach SMT formula ϕ ψ Encoder SAT-solver sat / unsat R. Bruttomesso (SMT) Overview 20 Ottobre 2011 12 / 23
  • 13. The Lazy Approach SMT formula ϕ Extract Struct. not good SAT-solver T -solver Candidate model good unsat sat R. Bruttomesso (SMT) Overview 20 Ottobre 2011 13 / 23
  • 14. Plan of (the rest of) the course The eager approach: solving bit-vectors Modern SAT-solvers: conflict analysis, clause learning, and heuristics The Lazy approach: generalities A theory-solver for IDL A theory-solver for UF A theory-solver for LRA (see if there is time left) R. Bruttomesso (SMT) Overview 20 Ottobre 2011 14 / 23
  • 15. Outline 1 A gentle introduction to SMT Introduction The Eager and the Lazy approaches 2 SMT-LIB and SMT-solvers SMT-LIB SMT-solvers R. Bruttomesso (SMT) Overview 20 Ottobre 2011 15 / 23
  • 16. SMT-LIB (v2) http://www.smtlib.org The SMT-LIB initiative defines a standard input language for SMT-solvers defines theories and logics in which formulæ can be written collects benchmarks The SMT-LIB language allows to write formulæ in a lisp-like format. E.g.: (< (+ x y) 0) (= (f x y) (g z)) stand for x + y < 0 and f (x, y) = g(z) respectively An SMT-LIB file looks more similar to a set of commands for an SMT-solver, rather then a logic formula R. Bruttomesso (SMT) Overview 20 Ottobre 2011 16 / 23
  • 17. SMT-LIB Theories An SMT-LIB theory consists of some sorts, (e.g., Int) and of some functions (e.g., −, +). Predicates are also considered functions, with codomain in Bool (e.g., <, ≤). For instance (theory Ints (theory Core :sorts ((Int 0)) :sorts ((Bool 0)) :funs ((NUMERAL Int) (- Int Int) :funs ((true Bool) (- Int Int Int :left-assoc) (false Bool) (+ Int Int Int :left-assoc) (not Bool Bool) (* Int Int Int :left-assoc) (=> Bool Bool Bool :right-assoc) (div Int Int Int :left-assoc) (and Bool Bool Bool :left-assoc) (mod Int Int Int) (or Bool Bool Bool :left-assoc) (abs Int Int) (xor Bool Bool Bool :left-assoc) (<= Int Int Bool :chainable) (par (A) (= A A Bool :chainable)) (< Int Int Bool :chainable) (par (A) (distinct A A Bool :pairwise)) (>= Int Int Bool :chainable) (par (A) (ite Bool A A A)) (> Int Int Bool :chainable) ) ) [...] [...] ) ) These definitions can be found at www.smtlib.org The sorts and the function symbols declared in a theory are always interpreted. This means that a to specify a model for a formula ϕ, we just need to specify the assignment of the variables to the concrete values in the sorts. R. Bruttomesso (SMT) Overview 20 Ottobre 2011 17 / 23
  • 18. SMT-LIB Logics The difference between “logic” and “theory” might look very subtle. An SMT-LIB logic includes a theory definition, plus it describes some restrictions on how formulæ can be built. (logic QF_LIA (logic QF_IDL :theories (Ints) :theories (Ints) :language "Closed quantifier-free formulas built :language over an arbitrary expansion of the "Closed quantifier-free formulas with Ints signature with free constant symbols, atoms of the form: but whose terms of sort Int are all linear, - q that is, have no occurrences of the function - (op (- x y) n), symbols *, /, div, mod, and abs, except as - (op (- x y) (- n)), or specified the :extensions attribute. - (op x y) " where - q is a variable or free constant symbol of sort Bool, :extensions - op is <, <=, >, >=, =, or distinct, "Terms with _concrete_ coefficients are also - x, y are free constant symbols of sort Int, allowed, that is, terms of the form c, (* c x), - n is a numeral. or (* x c) where x is a free constant and c " is a term of the form n or (- n) for some numeral n. ) " ) In the following we will not be so strict, and we will not make any distinction between “theories” and “logics”, calling both “theories”. For instance when we will say that we reason modulo the theory LIA we mean that we are working with QF LIA formulæ R. Bruttomesso (SMT) Overview 20 Ottobre 2011 18 / 23
  • 19. Writing an SMT-LIB file The logic can be specified with the command (set-logic QF_LIA) Variables are declared with (declare-fun x ( ) Int) A formula is specified with (assert (<= (+ x y) 0)) Asks the tool to compute satisfiability of assertions (check-sat) Asks the tool to return a model (in case of sat result) (set-option :produce-models true) ... (get-value (x y)) Disable annoying printouts (set-option :print-success false) R. Bruttomesso (SMT) Overview 20 Ottobre 2011 19 / 23
  • 20. Example (set-logic QF_LIA) (declare-fun x ( ) Int) (declare-fun y ( ) Int) (declare-fun a ( ) Bool) (assert (<= (+ x y) 0)) (assert (= x 0)) (assert (or (not a) (= x 1) (>= y 0))) (assert (not (= (+ y 1) 0))) (check-sat) (exit) which stands for the LIA formula (x + y ≤ 0) ∧ (x = 0) ∧ ((¬a ∨ (x = 1) ∨ (y ≥ 0)) ∧ ¬(y + 1 = 0) R. Bruttomesso (SMT) Overview 20 Ottobre 2011 20 / 23
  • 21. SMT-solvers An SMT-solver is a tool that can parse and solve an SMT-LIB benchmark. There are many such tools available online. In this course we will use Yices (developed at SRI, Stanford, closed source), Z3 (developed at MSR, Redmond, closed source) and OpenSMT (developed here, open source). Other available tools are MathSAT, CVC4, Boolector, veriT, STP. roberto@moriarty:examples$ smtlib2yices < test1.smt2 success success success success success success success success sat R. Bruttomesso (SMT) Overview 20 Ottobre 2011 21 / 23
  • 22. SMT-LIB script The SMT-LIB language allows specification of scripts. A script is a benchmark that may contain many check-sat commands. Also, it may include push and pop commands which can be used to control the assertion stack (set-option :print-success false) (set-logic QF_LIA) (declare-fun x ( ) Int) (declare-fun y ( ) Int) (assert (<= (+ x y) 0)) (assert (= x 0)) (assert (or (= x 1) (>= y 0))) (check-sat) (push 1) (assert (not (= y 0))) (check-sat) (pop 1) (check-sat) (exit) R. Bruttomesso (SMT) Overview 20 Ottobre 2011 22 / 23
  • 23. Exercizes 1 Translate the following LIA formula SMT-LIB, and evaluate it with an SMT-solver (x − y ≤ 0) ∧ (y − z ≤ 0) ∧ ((z − x ≤ −1) ∨ (z − x ≤ −2)) 2 Translate the following LRA formula SMT-LIB, and evaluate it with an SMT-solver (b ∨ (x + y ≤ 0)) ∧ (¬b ∨ (x + z ≤ 10)) 3 For the satisfiable formulæ above print out a model 4 For the satisfiable formulæ above, add constraints such that they become unsatisfiable R. Bruttomesso (SMT) Overview 20 Ottobre 2011 23 / 23