Analyse statique de programmes numériques avec calculs flottants



Similar documents
A Static Analyzer for Large Safety-Critical Software. Considered Programs and Semantics. Automatic Program Verification by Abstract Interpretation

Relational Abstract Domains for the Detection of Floating-Point Run-Time Errors

Numerology - A Case Study in Network Marketing Fractions

Static analysis of numerical programs

InvGen: An Efficient Invariant Generator

A Static Analyzer for Large Safety-Critical Software

Astrée: Building a Static Analyzer for Real-Life Programs

Fixpoint Computation in the Polyhedra Abstract Domain using Convex and Numerical Analysis Tools

RN-coding of Numbers: New Insights and Some Applications

Measures of Error: for exact x and approximation x Absolute error e = x x. Relative error r = (x x )/x.

Numerical Matrix Analysis

µz An Efficient Engine for Fixed points with Constraints

A Static Analyzer for Large Safety-Critical Software

CS321. Introduction to Numerical Methods

CHAPTER 5 Round-off errors

This Unit: Floating Point Arithmetic. CIS 371 Computer Organization and Design. Readings. Floating Point (FP) Numbers

CSE373: Data Structures and Algorithms Lecture 3: Math Review; Algorithm Analysis. Linda Shapiro Winter 2015

Proximal mapping via network optimization

Binary Division. Decimal Division. Hardware for Binary Division. Simple 16-bit Divider Circuit

What is Linear Programming?

RN-Codings: New Insights and Some Applications

Reducing Clocks in Timed Automata while Preserving Bisimulation

9 More on differentiation

IKOS: A Framework for Static Analysis based on Abstract Interpretation (Tool Paper)

Automatic Verification by Abstract Interpretation

ATheoryofComplexity,Conditionand Roundoff

6.2 Permutations continued

Continued Fractions and the Euclidean Algorithm

ECE 0142 Computer Organization. Lecture 3 Floating Point Representations

Chair of Software Engineering. Software Verification. Assertion Inference. Carlo A. Furia

Vector and Matrix Norms

Contrôle dynamique de méthodes d approximation

Geometrical Characterization of RN-operators between Locally Convex Vector Spaces

Pacific Journal of Mathematics

On strong fairness in UNITY

Modern Optimization Methods for Big Data Problems MATH11146 The University of Edinburgh

Static Program Transformations for Efficient Software Model Checking

YOU CAN COUNT ON NUMBER LINES

How To Develop A Static Analysis System For Large Programs

(Refer Slide Time: 01:11-01:27)

GenOpt (R) Generic Optimization Program User Manual Version 3.0.0β1

Numerical Analysis I

Computing Cubic Fields in Quasi-Linear Time

(Basic definitions and properties; Separation theorems; Characterizations) 1.1 Definition, examples, inner description, algebraic properties

Georgia Department of Education Kathy Cox, State Superintendent of Schools 7/19/2005 All Rights Reserved 1

Indiana State Core Curriculum Standards updated 2009 Algebra I

Math Review. for the Quantitative Reasoning Measure of the GRE revised General Test

THE NUMBER OF REPRESENTATIONS OF n OF THE FORM n = x 2 2 y, x > 0, y 0

Integer Factorization using the Quadratic Sieve

Method To Solve Linear, Polynomial, or Absolute Value Inequalities:

Høgskolen i Narvik Sivilingeniørutdanningen STE6237 ELEMENTMETODER. Oppgaver

Separation Properties for Locally Convex Cones

1 Introduction. Linear Programming. Questions. A general optimization problem is of the form: choose x to. max f(x) subject to x S. where.

Static Analysis of Dynamic Properties - Automatic Program Verification to Prove the Absence of Dynamic Runtime Errors

Tim Kerins. Leaving Certificate Honours Maths - Algebra. Tim Kerins. the date

arxiv: v2 [math.nt] 8 Aug 2009

FMA Implementations of the Compensated Horner Scheme

DNA Data and Program Representation. Alexandre David

a 1 x + a 0 =0. (3) ax 2 + bx + c =0. (4)

Mathematics (MAT) MAT 061 Basic Euclidean Geometry 3 Hours. MAT 051 Pre-Algebra 4 Hours

CHAPTER II THE LIMIT OF A SEQUENCE OF NUMBERS DEFINITION OF THE NUMBER e.

Lecture 7: Finding Lyapunov Functions 1

The Heat Equation. Lectures INF2320 p. 1/88

A new binary floating-point division algorithm and its software implementation on the ST231 processor

Static Taint-Analysis on Binary Executables

Correctly Rounded Floating-point Binary-to-Decimal and Decimal-to-Binary Conversion Routines in Standard ML. By Prashanth Tilleti

Notes on Factoring. MA 206 Kurt Bryan

Notes from Week 1: Algorithms for sequential prediction

NP-Hardness Results Related to PPAD

A Survey of Static Program Analysis Techniques

Powers of Two in Generalized Fibonacci Sequences

Regression Verification: Status Report

1. Prove that the empty set is a subset of every set.

Some representability and duality results for convex mixed-integer programs.

arxiv: v1 [math.pr] 5 Dec 2011

Copy in your notebook: Add an example of each term with the symbols used in algebra 2 if there are any.

FACTORING POLYNOMIALS IN THE RING OF FORMAL POWER SERIES OVER Z

s-convexity, model sets and their relation

24. The Branch and Bound Method

Abstract Interpretation-based Static Analysis Tools:

Introduction to Support Vector Machines. Colin Campbell, Bristol University

Abstract Simulation: a Static Analysis of Simulink Models

Termination Checking: Comparing Structural Recursion and Sized Types by Examples

Probability Generating Functions

Undergraduate Notes in Mathematics. Arkansas Tech University Department of Mathematics

CONTINUED FRACTIONS AND PELL S EQUATION. Contents 1. Continued Fractions 1 2. Solution to Pell s Equation 9 References 12

Numerical methods for American options

JUST THE MATHS UNIT NUMBER 1.8. ALGEBRA 8 (Polynomials) A.J.Hobson

CORRELATED TO THE SOUTH CAROLINA COLLEGE AND CAREER-READY FOUNDATIONS IN ALGEBRA

Support Vector Machines Explained

No: Bilkent University. Monotonic Extension. Farhad Husseinov. Discussion Papers. Department of Economics

a 11 x 1 + a 12 x a 1n x n = b 1 a 21 x 1 + a 22 x a 2n x n = b 2.

Introduction to Scheduling Theory

Binary Multiplication

An Abstract Monte-Carlo Method for the Analysis of Probabilistic Programs

What Every Computer Scientist Should Know About Floating-Point Arithmetic

FACTORING LARGE NUMBERS, A GREAT WAY TO SPEND A BIRTHDAY

6 EXTENDING ALGEBRA. 6.0 Introduction. 6.1 The cubic equation. Objectives

Cartesian Factoring of Polyhedra in Linear Relation Analysis

Transcription:

Analyse statique de programmes numériques avec calculs flottants 4ème Rencontres Arithmétique de l Informatique Mathématique Antoine Miné École normale supérieure 9 février 2011 Perpignan 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 1 / 63

Outline Introduction 1 Introduction Main goals Theoretical background 2 Rational Abstractions Interval domain Polyhedra domain 3 Floating-Point Abstractions Floating-point semantics Floating-point interval domain Expression linearization Floating-point polyhedra 4 Conclusion 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 2 / 63

Introduction Introduction 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 3 / 63

Static analysis Introduction Main Goals Goal: static analysis [CousotCousot-ISP76] Static (automatic) discovery of dynamic (semantic) properties of programs. Applications: compilation and optimisation, e.g.: array bound check elimination alias analysis verification, e.g.: infer invariants prove the absence of run-time errors (division by zero, overflow, invalid array access) prove functional properties We focus here on numerical properties of numerical variables. 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 4 / 63

Introduction Main Goals Example: discovering numerical invariants Insertion Sort for i=1 to 99 do p := T[i]; j := i+1; while j <= 100 and T[j] < p do T[j-1] := T[j]; j := j+1; end; T[j-1] := p; end; 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 5 / 63

Introduction Main Goals Example: discovering numerical invariants Interval analysis: Insertion Sort for i=1 to 99 do i [1, 99] p := T[i]; j := i+1; i [1, 99], j [2, 100] while j <= 100 and T[j] < p do i [1, 99], j [2, 100] T[j-1] := T[j]; j := j+1; i [1, 99], j [3, 101] end; i [1, 99], j [2, 101] T[j-1] := p; end; = there is no out of bound array access 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 5 / 63

Introduction Main Goals Example: discovering numerical invariants Linear inequality analysis: Insertion Sort for i=1 to 99 do i [1, 99] p := T[i]; j := i+1; i [1, 99], j = i + 1 while j <= 100 and T[j] < p do i [1, 99], i + 1 j 100 T[j-1] := T[j]; j := j+1; i [1, 99], i + 2 j 101 end; i [1, 99], i + 1 j 101 T[j-1] := p; end; 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 5 / 63

Theoretical background Introduction Theoretical Background Abstract interpretation: unifying theory of program semantics [CousotCousot-POPL77] Provide theoretical tools to design and compare static analyses that: always terminate are approximate (solve undecidability and efficiency issues) are sound by construction (no behavior is omitted) Analysis design roadmap: 1 concrete semantics 2 abstract domains 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 6 / 63

Concrete semantics Introduction Theoretical Background Concrete semantics: most precise mathematical expression of the program behavior Example: from programs (CFGs) to equation systems entry X 1 = D (initial states) 1 X:=?(0,10) X 2 = C X :=?(0, 10) X 1 2 X 3 = C Y := 100 X 2 Y:=100 loop 3 X<0 C Y := Y + 10 X 5 invariant X>=0 6 X 4 = C X 0 X 3 4 X 5 = C X := X 1 X 4 X:=X 1 Y:=Y+10 X 6 = C X < 0 X 3 5 V = {X, Y} variables D = P(V Q) sets of environments X i D reachable environments at location i C c X models the effect of command c on X 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 7 / 63

Introduction Concrete semantics (cont.) Theoretical Background Semantics of (non-deterministic) expressions and commands. E e : (V Q) P(Q) E c ρ E?(c, c ) ρ E V ρ E e ρ E e 1 e 2 ρ (expression semantics) = { c } = { x c x c } = { ρ(v) } = { v v E e ρ } = { v 1 + v 2 v 1 E e 1 ρ, v 2 E e 2 ρ } {+,,, /} ( / v 2 0) C c : P(V Q) P(V Q) C V :=e X C e 0 X (command semantics) = { ρ[ V v ] ρ X, v E e ρ } = { ρ ρ X, v E e ρ, v 0 } (can be extended to actual programming languages!) 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 8 / 63

Introduction Concrete semantics (cont.) Theoretical Background equations have the form: X i = F i (X 1,..., X n ) (P(V Q),,, ) is a complete lattice all the F i are monotonic (A B = C c A C c B) Constructive version of Tarski s theorem by [Tarski-PJM55] and [CousotCousot-PJM79] the system has a least solution (least fixpoint of F i ) it is the limit of: X 0 i = X k+1 i = F i (X1 k,..., X n k ) for successor ordinals Xi o = δ<o X i δ for limit ordinals (many kinds of semantics can be expressed in fixpoint form [Cousot-ENTCS97]) 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 9 / 63

Undecidability Introduction Theoretical Background elements in P(V Q) are not computer-representable C and are not computable least solutions of equations are not computable (requiring transfinite iterations) = we use computable abstractions i.e.: computable sound over-approximations 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 10 / 63

Abstract domains Introduction Theoretical Background Abstract elements: D set of computer-representable elements γ : D D concretization approximation order: X Y = γ(x ) γ(y ) Abstract operators: C c : D D and : (D D ) D soundness: (C c γ)(x ) (γ C c )(X ) γ(x ) γ(y ) γ(x Y ) Fixpoint extrapolation : (D D ) D widening soundness: γ(x ) γ(y ) γ(x Y ) termination: sequence (Y i ) i N the sequence X 0 = Y 0, X i+1 = X i Y i+1 stabilizes in finite time: n < ω, X n+1 = X n Both semantics and algorithmic aspects. 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 11 / 63

Abstract analysis Introduction Theoretical Background Principles: compute entirely in D replace C c with C c in equations iterate, using widening at loop heads W X i, 0 X i, k+1 = { F i (X 1, k,..., X n, k ) = X i, k F i (X 1, k,..., X n, k ) if i / W if i W Theorem: the iterations stabilize in finite time δ < ω: X i, δ+1 = X i, δ the result is sound: X i γ(x i, δ ) 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 12 / 63

Introduction Theoretical Background Some existing numerical abstract domains Intervals Xi [a i, b i ] [CousotCousot-ISP76] Simple Congruences Xi a i [b i ] [Granger-JCM89] Linear Equalities i α ix i = β [Karr-AI76] Linear Congruences i α ix i β [γ] [Granger-TAPSOFT91] 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 13 / 63

Introduction Theoretical Background Some existing numerical abstract domains (cont.) Polyhedra i α ix i β [CousotHalbwachs-POPL78] Octagons ±Xi ± X j β [Miné-WCRE01] Ellipsoids αx 2 i + βx 2 j + γx i Y i δ [Feret-ESOP04] Varieties P( X) = 0, P R[V] [SankaranarayananAl-POPL04] 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 14 / 63

Introduction Precision vs. cost tradeoff Theoretical Background Example: three abstractions of the same set of points Worst-case time cost per operation wrt. number of variables: polyhedra: exponential octagons: cubic intervals: linear 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 15 / 63

Rational Abstractions Rational Numerical Abstract Domains 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 16 / 63

Rational Abstractions Interval Domain Interval Domain 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 17 / 63

Rational Abstractions Interval abstract elements Interval Domain Abstract elements: I = { (a, b) a Q { }, b Q {+ }, a b } (intervals as pairs of bounds) D Concretization: = (V I) { } γ( ) = γ(x ) = { ρ V, ρ(v) γ I (X (V)) } if X where γ I (a, b) = { x Q a x b } Order: X Y where (a, b) I (a, b ) { X = X, Y V, X (V) I Y (V) a a b b 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 18 / 63

Rational Abstractions Interval abstract operators Interval Domain Interval arithmetics in I I (a, b) (a, b) + I (c, d) (a, b) I (c, d) (a, b) I (c, d) (a, b) I (c, d). = ( b, a) = (a + c, b + d) = (a d, b c) = (min(ac, ad, bc, bd), max(ac, ad, bc, bd)) = (min(a, c), max(b, d)) Join in D X Y = X Y λv. X (V) I Y (V) if Y = if X = if X, Y (optimal, but not exact: γ(x ) γ(y ) γ(x Y )) 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 19 / 63

Interval assignment Rational Abstractions Interval Domain E e : D I (abstract expression evaluation, X ) E c X E?(c, c ) X E V X E e X E e 1 e 2 X = (c, c) = (c, c ) = X (V) = I E e X = E e 1 X I E e 2 X C c : D D C V :=e X = (abstract command semantics) { if X = X [ V E e X ] if X 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 20 / 63

Interval test example Rational Abstractions Interval Domain C X + Y Z 0 X, with X = {X (0, 10), Y (2, 10), Z (3, 5)} X (0, 10) + (, + ) - (, + ) Y (2, 10) Z (3, 5) X (0, 10) + (2, 20) - ( 3, 17) Y (2, 10) Z (3, 5) X (0, 10) + (2, 20) - ( 3, 0) Y (2, 10) Z (3, 5) X (0, 3) + (2, 5) - ( 3, 0) Y (2, 5) Z (3, 5) 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 21 / 63

Interval widening Rational Abstractions Interval Domain Classic widening in D X X Y where = (a, b) I (c, d) = Y λv. X (V) I Y (V) ({ a if a c otherwise, if Y = if X = if X, Y { b if b d + otherwise ) Unstable bounds are set to ± Widening with thresholds: Parametrized by a finite set T. If b < d, bump to the next value in T greater than d (or + ). 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 22 / 63

Rational Abstractions Interval analysis example Interval Domain Analysis example with W = {2} X<40 1 2 3 X:=0 X>=40 4 X:=X+1 l 1 2 = 0 = 0 0 0 0 3 = 0 = 0 [0, 39] [0, 39] 4 40 40 X 0 l More precisely, at the widening point: X 1 X 2 X 3 X 4 X 1 l X 2 l X 3 l X 4 l 2 = ([0, 0] ) = [0, 0] = [0, 0] 2 = [0, 0] ([0, 0] ) = [0, 0] [0, 0] = [0, 0] 2 = [0, 0] ([0, 0] [1, 1]) = [0, 0] [0, 1] = [0, + [ 2 = [0, + [ ([0, 0] [1, 40]) = [0, + [ [0, 40] = [0, + [ Note that the most precise interval abstraction would be X [0, 40] at 2, and X = 40 at 4. 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 23 / 63 X 5 l

Rational Abstractions Polyhedra Domain Polyhedra Domain 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 24 / 63

Polyhedra domain Rational Abstractions Polyhedra Domain Domain proposed by [CousotHalbwachs-POPL78] to infer conjunctions of linear inequalities j ( n i=1 α ijv i β j ). Abstract elements: LinCons = linear constraints over V with coefficients in Q D Concretization: = P finite (LinCons) γ(x ) = { ρ V Q c X, ρ = c } γ(x ) is a closed convex polyhedron of (V Q) Q V γ(x ) may be empty, bounded, or unbounded γ is not injective 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 25 / 63

Polyhedra algorithms Rational Abstractions Polyhedra Domain Fourier-Motzkin elimination: Fourier(X, V k ) eliminates V k from all the constraints in X : Fourier(X, V k ) = { ( i α iv i β) X α k = 0 } { ( α k )c+ + α + k c c + = ( i α+ i V i β + ) X, α + k > 0, c = ( i α i V i β ) X, α k < 0 } Semantics γ(fourier(x, V k )) = { ρ[v k v] v Q, ρ γ(x ) } i.e., forget the value of V k 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 26 / 63

Polyhedra algorithms Rational Abstractions Polyhedra Domain Linear programming: [Schrijver-86] simplex(x, α) = min { i α iρ(v i ) ρ γ(x ) } Application: remove redundant constraints: for each c = ( i α iv i β) X if β simplex(x \ {c}, α), then remove c from X (e.g., Fourier causes a quadratic growth in constraint number, most of which are redundant) Note: calling simplex many times can be costly use fast syntactic checks first check against the bounding-box first use simplex as a last resort 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 27 / 63

Rational Abstractions Polyhedra abstract operators Polyhedra Domain Order: X Y X = Y ( i α iv i β) Y, simplex(x, α) β γ(x ) γ(y ) X Y Y X Join: [BenoyKing-LOPSTR96] We introduce temporaries V X j, V Y j, σ X, σ Y : = X Y Fourier( { ( j α jv X j βσ X 0) ( j α jv j β) X } { ( j α jv Y j βσ Y 0) ( j α jv j β) Y } { V j = V X j + V Y j V j V } { σ X 0, σ Y 0, σ X + σ Y = 1 }, { V X j, V Y j V j V } { σ X, σ Y } ) γ(x Y ) is the topological closure of the convex hull of γ(x ) and γ(y ) (optimal). 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 28 / 63

Rational Abstractions Polyhedra Domain Polyhedra abstract operators (cont.) Precise abstract commands: C i α iv i + β 0 X (exact) = X {( i α iv i + β 0)} C V j := [, + ] X = Fourier(X, V j )) C V j := i α iv i + β X = subst(v V i, Fourier((X {V = i α iv i + β}), V j )) Fallback abstract commands: C e 0 X = X (coarse but sound) C V j := e X = Fourier(X, V j ) alternate solution: apply interval abstract commands to the bounding box 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 29 / 63

Polyhedra widening Rational Abstractions Polyhedra Domain Classic widening in D X Y = { c X Y {c} } { c Y c X, X = (X \ c ) {c} } suppress unstable constraints c X, Y {c} add back constraints c Y equivalent to those in X i.e., when c X, X = (X \ c ) {c}. (X and Y must have no redundant constraint) Example: 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 30 / 63

Analysis example Rational Abstractions Polyhedra Domain Rate limiter Y:=0; while true do X:=?(-128,128); D:=?(0,16); S:=Y; Y:=X; R:=X-S; if R+D<=0 then Y:=S-D fi; if R-D>=0 then Y:=S+D fi od X: input signal Y: output signal S: last output R: delta Y-S D: max. allowed for R Polyhedra analysis: result: at, Y [ 128, 128] to prove, e.g., Y 128, the analysis needs to: represent the properties R = X S and R D combine them to deduce S X D, and then Y = S D X 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 31 / 63

Analysis example Rational Abstractions Polyhedra Domain Rate limiter Y:=0; while true do X:=?(-128,128); D:=?(0,16); S:=Y; Y:=X; R:=X-S; if R+D<=0 then Y:=S-D fi; if R-D>=0 then Y:=S+D fi od X: input signal Y: output signal S: last output R: delta Y-S D: max. allowed for R Interval analysis: iterations without widening: X,0 X,1 X,2... X,n Y = 0 Y 144 Y 160... Y 128 + 16n i.e., not bound on Y can be found. 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 31 / 63

Floating-point abstractions Floating-point abstractions 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 32 / 63

Floating-point abstractions Floating-point uses Two independent problems: Implement the analyzer using floating-point goal: trade precision for efficiency exact rational arithmetics can be costly coefficients can grow large (polyhedra) Analyze floating-point programs goal: catch run-time errors caused by rounding (overflow, division by 0,... ) Also: a floating-point analyzer for floating-point programs. Challenge: how to stay sound? 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 33 / 63

Floating-point abstractions Floating-point semantics Floating-point semantics 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 34 / 63

Floating-point abstractions Floating-point numbers Floating-point semantics IEEE 754-1985 standard is the most widespread format (supported by most processors and programming languages) IEEE Binary representation: a number is a triple s, e, f a 1-bit sign s a x-bit exponent e, with a bias (e represents e bias) a p-bit fraction f =.b 1... b p, (f represents i 2 i b i ) IEEE format examples given by the choice of x, bias, p: x = 8 32-bit single precision float: bias = 127 p = 23 Other widespread formats: 64-bit double, 80-bit double extended, 128-bit quad 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 35 / 63

Floating-point abstractions Floating-point representation Floating-point semantics Semantics s, e, f represents either: a normalized number: ( 1) s 2 e bias 1.f (if 1 e 2 x 2); a denormalized number: ( 1) s 2 1 bias 0.f (if e = 0, f 0); +0 or 0 (if e = 0, f = 0); + or (if e = 2 x 1, f = 0); an error code NaN (if e = 2 x 1, f 0). Visual representation (positive part) +0 mf Mf + mf Mf denormalized normalized = 2 1 bias p smallest positive = (2 2 p ) 2 2x bias 2 largest non- 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 36 / 63

Floating-point abstractions Floating-point computations Floating-point semantics The set of floating-point numbers is not closed under +,,, /: every result is rounded to a representable float, an overflow or division by 0 generates + or (overflow); small numbers are truncated to +0 or 0 (underflow); some operations are invalid (0/0, (+ ) + ( ), etc.) and return NaN. Simplified semantics: overflows and NaNs halt the program with an error Ω, rounding and underflow are not errors, we do not distinguish between +0 and 0. = variable values live in a finite subset F of Q, expression values live in F {Ω}. 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 37 / 63

Floating-point abstractions Floating-point expressions Floating-point semantics Floating-point expressions exp f exp f ::= [c, c ] constant range c, c F, c c V variable V V exp f negation exp f r exp f operator {,,, } (we use circled operators to distinguish them from operators in Q) 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 38 / 63

Floating-point abstractions Concrete semantics of expressions Floating-point semantics Semantics of rounding: R r : Q F {Ω}. 4 rounding modes r: towards +,, 0, or to nearest n. Example inition: R + (x) = R (x) = { min { y F y x } if x Mf Ω if x > Mf { max { y F y x } if x Mf Ω if x < Mf Notes: x, r, R (x) R r (x) R + (x) r, R r is monotonic 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 39 / 63

Floating-point abstractions Floating-point semantics Concrete semantics of expressions (cont.) E e f : (V F) P(F {Ω}) (expression semantics) E V ρ = { ρ(v) } E [c, c ] ρ = { x F c x c } E e f ρ = { x x E e f ρ F } ({Ω} E e f ρ) E e f r e f ρ = { R r (x y) x E e f ρ F, y E e f ρ F } { Ω if Ω E e f ρ E e f ρ } { Ω if 0 E e f ρ and = } C c : P(V F) P((V F) {Ω}) (command semantics) C X := e f X = { ρ[ X v ] ρ X, v E e f ρ F } ({Ω} E e f X ) C e f 0 X = { ρ ρ X, v E e f ρ F, v 0 } ({Ω} E e f X ) 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 40 / 63

Floating-point abstractions Floating-point interval domain Floating-point interval domain 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 41 / 63

Floating-point abstractions Floating-point interval domain Floating-point interval abstract elements Goals: analyze floating-point programs abstracted using a floating-point implementation report infinities and NaN as errors Abstract elements: I = { (a, b) a, b F, a b } (floating-point bounds) D Concretization: = (V I) { } γ( ) = γ(x ) = { ρ V, ρ(v) γ I (X (V)) } if X where γ I (a, b) = { x F a x b } 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 42 / 63

Floating-point abstractions Floating-point interval domain Floating-point interval domain Interval arithmetics I (a, b) (a, b) I (a, b ) (a, b) I (a, b ) (a, b) I (a, b ) = ( b, a) = (a a, b + b ) = (a b, b + a ) = (min(a a, a b, b a, b b ), max(a + a, a + b, b + a, b + b )) (a, b) I (a, b ) = (min(a, a ), max(b, b )) I sets unstable bounds to ± Mf (we suppose r is unknown and assume a worst case rounding) Error management If some bound in E e f evaluates to ± or NaN, we report an alarm to the user and continue the evaluation with ( Mf, Mf ) 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 43 / 63

Floating-point abstractions Floating-point interval domain Floating-point interval analysis example Filter with reinitialisation Z:=0; while true do if?(0,1) then Z:=?(-12,12) fi; Z:=(0.3? Z)??(-10,10) od in Q: Z < 10/0.7 ( 14.29). in F: Z B, with B = R + (10/0.7). Interval analysis: using a widening with thresholds T: Z min { x T x B }. = proof of absence of overflow if T has a value larger than B. 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 44 / 63

Floating-point abstractions Expression linearization Expression linearization 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 45 / 63

Floating-point abstractions Expression linearization Floating-point issues in relational domains Relational domains assume many powerful properties on Q: associativity, distributivity,... that are not true on F! Example: Fourier-Motzkin elimination X Y c Y Z d = X Z c + d X n Y c Y n Z d X n Z c n d (X = 1, Y = 10 38, Z = 1, c = X n Y = 10 38, d = Y n Z = 10 38, c n d = 0, X n Z = 2 > 0) We cannot manipulate float expressions as easily as rational ones! Solution: keep representing and manipulating rational expressions abstract float expressions from programs into rational ones feed them to a rational abstract domain (optional) implement the rational domain using floats 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 46 / 63

Floating-point abstractions Affine interval forms Expression linearization We put expressions in affine interval form: [Miné-ESOP04] Semantics: exp l ::= [a 0, b 0 ] + k [a k, b k ] V k E e l ρ = { c 0 + k c k ρ(v k ) i, c i [a i, b i ] } (evaluated in Q) Advantages: affine expressions are easy to manipulate interval coefficients allow non-determinism in expressions, hence, the opportunity for abstraction intervals can easily model rounding errors easy to design algorithms for C X :=e l and C e l 0 in most domains 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 47 / 63

Floating-point abstractions Affine interval form algebra Expression linearization Operations on affine interval forms: adding and subtracting two forms multiplying and dividing a form by an interval Using interval arithmetics + I, I, I : (i 0 + k i k V k ) (i 0 + k i k V k) = (i 0 + I i 0 ) + k (i k + I i k ) V k i (i 0 + k i k V k ) = (i I i 0 ) + k (i I i k ) V k... Projection: π k : D exp l We suppose we are given an abstract interval projection operator π k such that: π k (X ) = [a, b] such that [a, b] { ρ(v k ) ρ γ(x ) }. 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 48 / 63

Floating-point abstractions Linearization of rational expressions Expression linearization Intervalization: ι : (exp l D ) exp l Intervalization flattens the expression into a single interval: ι(i 0 + k i k V k, X ) = i 0 + I I, k (i k I π k (X )). Linearization without rounding errors: Defined by induction on the syntax of expressions: l(v, X ) = [1, 1] V l([a, b], X ) = [a, b] l(e 1 +e 2, X ) = l(e 1, X ) l(e 2, X ) l(e 1 e 2, X ) = l(e 1, X ) l(e 2, X ) l(e 1 /e 2, X ) = l(e 1, X ) ι(l(e 2, X ), X ) l(e 1 e 2, X ) = can be l : (exp D ) exp l { either ι(l(e1, X ), X ) l(e 2, X ) or ι(l(e 2, X ), X ) l(e 1, X ) 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 49 / 63

Floating-point abstractions Expression linearization Linearization of floating-point expressions Rounding an affine interval form: if the result is normalized: we have a relative error ε with magnitude 2 23 : ε([a 0, b 0 ] + k [a k, b k ] V k ) = max( a 0, b 0 ) [ 2 23, 2 23 ] + k (max( a k, b k ) [ 2 23, 2 23 ] V k ) if the result is denormalized, we have an absolute error ω = [ 2 159, 2 159 ]. = we sum these two sources of rounding errors Linearization with rounding errors: l : (exp f D ) exp l l(e 1 r e 2, X ) = l(e 1, X ) l(e 2, X ) ε(l(e 1, X )) ε(l(e 2, X )) ω l(e 1 r e 2, X ) = ι(l(e 1, X ), X ) (l(e 2, X ) ε(l(e 2, X ))) ω etc. 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 50 / 63

Floating-point abstractions Expression linearization Applications of the floating-point linearization Soundness of the linearization e, X D, ρ γ(x ), if Ω / E e ρ, then E e ρ E l(e, X ) ρ Application: C V :=e X Note: check that Ω / E e ρ for ρ γ(x ) with interval arithmetic compute C V :=e X as C V :=l(e, X ) X (use C V :=[ Mf, Mf ] X if Ω E e ρ) I, I, I, I are sound abstractions of + I, I, I, / I = use I, etc. to get a sound approximation of l(e, X ) in floats 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 51 / 63

Floating-point abstractions Sound floating-point polyhedra Sound floating-point polyhedra 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 52 / 63

Floating-point abstractions Sound floating-point polyhedra Sound floating-point polyhedra Algorithms to adapt: [ChenAl-APLAS08] linear programming: simplex f (X, α) simplex(x, α) simplex(x, α) = min { k α kρ(v k ) ρ γ(x ) } Fourier-Motzkin elimination: Fourier f (X, V k ) = Fourier(X, V k ) Fourier(X, V k ) = { ( i α iv i β) X α k = 0 } { ( α k )c+ + α + k c c + = ( i α+ i V i β + ) X, α + k > 0, c = ( i α i V i β ) X, α k < 0 } 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 53 / 63

Floating-point abstractions Sound floating-point polyhedra Sound floating-point linear programming Guaranteed linear programming: [NeumaierShcherbina-MP04] Goal: under-approximate µ = min { c x M x b } knowing that x [ x l, x h ] (bounding-box for γ(x )). compute any approximation µ of the dual problem: µ µ = max { b y t M y = c, y 0 } and the corresponding vector y (e.g. using an off-the-shelf solver; µ may over-approximate or under-approximate µ) compute with intervals safe bounds [ r l, r h ] for A y c: [ r l, r h ] = ( t A I y) I c and then: ν = inf(( b I y) I ([ r l, r h ] I [ x l, x h ])) then: ν µ. 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 54 / 63

Floating-point abstractions Sound floating-point polyhedra Sound floating-point Fourier-Motzkin elimination Given: c + = ( i α + i V i β + ) with α + k > 0 c = ( i α i V i β ) with α k < 0 a bounding-box of γ(x ): [ x l, x h ] We wish to compute i k α iv i β in F implied by ( α k )c+ + α + k c in γ(x ). normalize c + and c using interval arithmetics: { Vk + i k (α+ i I α + k )V i β + I α + k V k + i k (α i I ( α k ))V i β I ( α k ) (interval affine forms) add them using interval arithmetics: i k [a i, b i ]V i [a 0, b 0 ] where [a i, b i ] = (α + i I α + k ) I (α i I α k ), [a 0, b 0 ] = (β + I α + k ) I (β I α k ). 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 55 / 63

Floating-point abstractions Sound floating-point polyhedra Sound floating-point Fourier-Motzkin elimination (cont.) linearize the interval linear form into i k α iv i β where { αi [a i, b i ] β = sup ([a 0, b 0 ] I I, i k ( α i I [a i, b i ] ) I [ x l, x h ] ) Soundness: For all choices of α i [a i, b i ], i k α iv k β holds in Fourier(X, V k ). (e.g. α i = (a i n b i ) 2) 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 56 / 63

Floating-point abstractions Consequences of rounding Sound floating-point polyhedra Precision loss: Projection: γ(fourier(x, V k )) { ρ[v k v] v Q, ρ γ(x ) } = C V k := [, + ] γ(x ) Order: X Y = γ(x ) γ(y ) Join: ( ) γ(x Y ) ConvexHull(γ(X ) γ(y )) ( ) Efficiency loss: cannot remove all redundant constraints 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 57 / 63

Floating-point abstractions Floating-point polyhedra widening Sound floating-point polyhedra Widening : X Y = { c X Y {c} } (drop { c Y c X, X = (X \ c ) {c}} as X and Y may have redundant constraints) Stability improvement: robust strategies to choose α i [a i, b i ] during Fourier-Motzkin: choose simple α i (e.g., integer nearest (a i + b i )/2) reuse the same (or a multiple of) α i used for other variables 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 58 / 63

Conclusion Conclusion 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 59 / 63

Conclusion Abstraction summary Floating-point polyhedra analyzer for floating-point programs expression abstraction environment abstraction float expression e f linearization P(V F) affine form e l in Q abstract domain float implementation polyhedra in Q affine form e l in F float implementation polyhedra in F widening polyhedra in F 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 60 / 63

Conclusion Academic implementation Interproc: on-line analyzer for a toy language, using Apron. http://pop-art.inrialpes.fr/interproc/interprocweb.cgi 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 61 / 63

Conclusion Industrial application Astrée static analyzer: [BertraneAl-AIAA10] developed at ENS, now industrialized by analyzes embedded critical control/command C code checks for run-time errors (arithmetics, arrays, pointers) applied to industrial Airbus code, up to 1 M lines zero alarm, 14h computation time Based on abstract interpretation: uses intervals and octagons (no polyhedra) and many more abstract domains (some domain-specific) uses linearization of float expressions More information: http://www.astree.ens.fr/ 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 62 / 63

Conclusion The End 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 63 / 63

Bibliography Bibliography 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 64 / 63

Bibliography Bibliography [Tarski-PJM55] A. Tarski. A lattice theoretical fixpoint theorem and its applications. In Pacific J. Math., 5 (1995), 285 310. [CousotCousot-ISP76] P. Cousot & R. Cousot. Static determination of dynamic properties of programs. In Proc. of the 2d Int. Symp. on Prog. Dunod, 1976. [CousotCousot-POPL77] P. Cousot & R. Cousot. Abstract interpretation: a unified lattice model for static analysis of programs by construction or approximation of fixpoints. In Proc. of POPL 77, 238 252, ACM Press, 1977. [CousotCousot-PJM79] P. Cousot & R. Cousot. Constructive versions of Tarski s fixed point theorems. In Pacific J. Math., 82:1 (1979), 43 57. [CousotHalbwachs-POPL78] P. Cousot & N. Halbwachs. Automatic discovery of linear restraints among variables of a program. In Proc. of POPL 78, 84 96. ACM Press, 1978. 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 65 / 63

Bibliography Bibliography (cont.) [Schrijver-86] A. Schrijver. Theory of linear and integer programming. John Wiley & Sons, Inc., 1986. [BenoyKing-LOPSTR96] F. Benoy & A. King. Inferring argument size relationships with CLP(R). In Proc. of LOPSTR 96, vol. 1207 of LNCS, pp. 204 223. Springer, 1996. [Cousot-ENTCS97] P. Cousot. Constructive design of a hierarchy of semantics of a transition system by abstract interpretation. In ENTCS 6, 25 p., 1997. [Miné-AST01] A. Miné. The octagon abstract domain. In AST 01, pp. 310 319. IEEE CS Press, 2001. [Miné-ESOP04] A. Miné. Relational abstract domains for the detection of floating-point run-time errors. In Proc. of ESOP 04, vol. 2986 of LNCS, pp. 3 17. Springer, 2004. 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 66 / 63

Bibliography Bibliography (cont.) [NeumaieShcherbina-MP04] A. Neumaier & O. Shcherbina. Safe bounds in linear and mixed-integer linear program ming. In Math. Program., 99(2):283 296 (2004). [Miné-ESOP04] A. Miné. Relational abstract domains for the detection of floating-point run-time errors. In Proc. of ESOP 04, vol. 2986 of LNCS, pp. 3 17. Springer, 2004. [ChenAl-APLAS08] L. Chen, A. Miné & P. Cousot. A sound floating-point polyhedra abstract domain. In Proc. of APLAS 08, vol. 5356 of LNCS, pp. 3 18. Springer, 2008. [BertraneAl-AIAA10] J. Bertrane, P. Cousot, R. Cousot, J. Feret, L. Mauborgne, A. Miné, X. Rival. Static Analysis and Verification of Aerospace Software by Abstract Interpretation. In AIAA Infotech@Aerospace (I@A 2010), 38 p. 9/02/2011 Analyse statique de programmes numériques Antoine Miné p. 67 / 63