Discrete Fixpoint Approximation Methods in Program Static Analysis

Similar documents
Static Program Analysis using Abstract Interpretation

BASIC CONCEPTS OF ABSTRACT INTERPRETATION

Introduction to Abstract Interpretation. ECE 584 Sayan Mitra Lecture 18

Partial model checking via abstract interpretation

Cours M.2-6 «Interprétation abstraite: applications à la vérification et à l analyse statique» Examen partiel. Patrick Cousot.

«ATutorialon Abstract Interpretation»

Notes on Abstract Interpretation

Abstraction in Program Analysis & Model Checking. Abstraction in Model Checking. Motivations & Results

CS 6110 Lecture 21 The Fixed-Point Theorem 8 March 2013 Lecturer: Andrew Myers. 1 Complete partial orders (CPOs) 2 Least fixed points of functions

«Mathematical foundations: (6) Abstraction Part II» Exact fixpoint abstraction. Property transformer abstraction

CMSC 631 Program Analysis and Understanding Fall Abstract Interpretation

Logical Abstract Domains and Interpretations

Introduction to Abstract Interpretation

«Basic Concepts of Abstract Interpretation»

The non-logical symbols determine a specific F OL language and consists of the following sets. Σ = {Σ n } n<ω

Basic Concepts of Abstract Interpretation

Foundations of Abstract Interpretation

The Trace Partitioning Abstract Domain

An Abstract Interpretation Framework for Semantics and Diagnosis of Lazy Functional-Logic Languages

Order Theory, Galois Connections and Abstract Interpretation

Galois Connections. Roland Backhouse 3rd December, 2002

A Certified Denotational Abstract Interpreter (Proof Pearl)

AAA616: Program Analysis. Lecture 3 Denotational Semantics

MIT Martin Rinard Laboratory for Computer Science Massachusetts Institute of Technology

Abstract Interpretation: Fixpoints, widening, and narrowing

Model Checking & Program Analysis

Lecture Notes: Selected Topics in Discrete Structures. Ulf Nilsson

Making Abstract Interpretations Complete

The Knaster-Tarski Fixed Point Theorem for Complete Partial Orders

Generalized Strong Preservation by Abstract Interpretation

Abstract Interpretation: Fixpoints, widening, and narrowing

Bi-inductive Structural Semantics

Abstract Interpretation II

Operational Semantics

Space-aware data flow analysis

Denotational semantics

Software Verification

Mechanics of Static Analysis

Proving Fixed Points

Introduction to Axiomatic Semantics

Lecture 2. Least Fixed Points

Constructive design of a hierarchy of semantics of a transition system by abstract interpretation

Examples of Abstract Interpretations

Goal. Partially-ordered set. Game plan 2/2/2013. Solving fixpoint equations

Principles of Program Analysis: Abstract Interpretation

Notes on Ordered Sets

Verification of Real-Time Systems Numerical Abstractions

Data flow analysis. DataFlow analysis

Abstract Interpretation from a Topological Perspective

Tree sets. Reinhard Diestel

Reading: Chapter 9.3. Carnegie Mellon

A hierarchy of Constraint Systems for Data-Flow Analysis of Constraint Logic-Based Languages

Constructive Design of a Hierarchy of Semantics of a Transition System by Abstract Interpretation

Lecture 1: Lattice(I)

«Specification and Abstraction of Semantics» 1. Souvenir, Souvenir. Neil D. Jones. Contents. A Tribute Workshop and Festival to Honor Neil D.

Semantics and Verification of Software

Formal Methods in Software Engineering

arxiv:cs/ v1 [cs.lo] 22 Dec 2006

Static Program Analysis

Towards a quantum calculus

Introduction to Program Analysis and Abstract Interpretation (Part I)

A New Numerical Abstract Domain Based on Difference-Bound Matrices

Weak Relative Pseudo-Complements of Closure Operators

Accelerated Data-Flow Analysis

Postulate 2 [Order Axioms] in WRW the usual rules for inequalities

Synthetic Computability

CS611 Lecture 18 Fixed Points and CPOs

Least Common Subsumers and Most Specific Concepts in a Description Logic with Existential Restrictions and Terminological Cycles*

Groupe de travail. Analysis of Mobile Systems by Abstract Interpretation

Probabilistic Model Checking and [Program] Analysis (CO469)

Higher-Order Chaotic Iteration Sequences

The algorithmic analysis of hybrid system

Polynomial Precise Interval Analysis Revisited

Abstract Interpretation (Non-Standard Semantics) a.k.a. Picking The Right Abstraction

From Many Places to Few: Automatic Abstraction Refinement for Petri Nets

Introduction. Semantics and applications to verification. Xavier Rival. École Normale Supérieure. February 11, 2016

Abstract Interpretation and Static Analysis

Introduction. Semantics and applications to verification. Xavier Rival. École Normale Supérieure. February 11, 2015

COSE312: Compilers. Lecture 14 Semantic Analysis (4)

Universal Algebra for Logics

arxiv: v1 [cs.cc] 18 Mar 2013

Static Analysis by Policy Iteration on Relational Domains

Fuzzy Limits of Functions

Uniform Closures: Order-Theoretically Reconstructing Logic Program Semantics and Abstract Domain Refinements

A Semantic Basis for Specialising Domain Constraints

Restricted role-value-maps in a description logic with existential restrictions and terminological cycles

Computing (the smallest) fixed points of monotone maps arising in static analysis by AI

Dataflow Analysis - 2. Monotone Dataflow Frameworks

Interprocedurally Analyzing Polynomial Identities

Understanding Finiteness Analysis Using Abstract

Topology. Xiaolong Han. Department of Mathematics, California State University, Northridge, CA 91330, USA address:

Compositionality in SLD-derivations and their abstractions Marco Comini, Giorgio Levi and Maria Chiara Meo Dipartimento di Informatica, Universita di

COMPLETE LATTICES. James T. Smith San Francisco State University

An Abstract Domain to Infer Ordinal-Valued Ranking Functions

Set, functions and Euclidean space. Seungjin Han

A NEW CHARACTERIZATION OF COMPLETE HEYTING AND CO-HEYTING ALGEBRAS

Static Program Analysis

Formal Techniques for Software Engineering: Denotational Semantics

Computational Methods in Systems and Synthetic Biology

CS422 - Programming Language Design

Transcription:

Discrete Fixpoint Approximation Methods in Program Static Analysis P. Cousot Département de Mathématiques et Informatique École Normale Supérieure Paris <cousot@dmi.ens.fr> <http://www.dmi.ens.fr/~cousot> Example 0 1 n j { n:ω 1 ; i:ω; j:ω } read int(n); { n:!! 2 [0,+ 3 ]; i:ω; j:ω } i := n; { n:[0,+ ]; i:[0,+ ]; j:[1,+ ]? 4 } while (i <> 0) do { n:[0,+ ]; i:[1,+ ]; j:[1,+ ]? } j := 0; { n:[0,+ ]; i:[1,+ ]; j:[0,+ ] } while (j <> i) do { n:[0,+ ]; i:[1,+ ]; j:[0,1073741822]!! } j := (j + 1) { n:[0,+ ]; i:[1,+ ]; j:[1,+ ] } od; { n:[0,+ ]; i:[1,+ ]; j:[1,+ ] } i := (i - 1); { n:[0,+ ]; i:[0,1073741822]; j:[1,+ ] } od; { n:[0,+ ]; i:[0,0]; j:[1,+ ]? } i 1 Ω denotes uninitialization. 2!! denotes inevitable error when the invariant is violated. 3 + =1073741823, =-1073741824. 4 This questionmark indicates possible uninitialization. NACSA 98 1 P. Cousot NACSA 98 3 P. Cousot Static Program analysis Automatic determination of runtime properties of infinite state programs Applications: - compilation (dataflow analysis, type inference), - program transformation (partial evaluation, parallelization/vectorization,... ) - program verification (test generation, abstract debugging,... ) Problems: - text inspection only (excluding executions or simulations) - undecidable - necessarily approximate Abstract interpretation Abstract interpretation [1, 2]: design method for static analysis algorithms; effective approximation of the semantics of programs; often, the semantics maps the program text to a model of computation obtained as the least fixpoint of an operator on a partially ordered semantic domain; effective approximation of fixpoints of posets; References [1] P. Cousot and R. Cousot. Abstract interpretation: a unified lattice model for static analysis of programs by construction or approximation of fixpoints. In Conference Record of the Fourth Annual ACM SIGPLAN-SIGACT Symposium on Principles of Programming Languages,pages238 252,Los Angeles, California, 1977. ACM Press, New York, New York, usa. [2] P. Cousot and R. Cousot. Systematic design of program analysis frameworks. In Conference Record of the Sixth Annual ACM SIGPLAN-SIGACT Symposium on Principles of Programming Languages, pages269 282,SanAntonio, Texas, 1979. ACM Press, New York, New York, usa. NACSA 98 2 P. Cousot NACSA 98 4 P. Cousot

Fixpoint semantics Program semantics can be defined as least fixpoints [3]: lfp F where F(lfp F) = lfp F F(x) =x = lfp F x of amonotonic operator F L m L on a complete partial order (CPO): L,,, where L, is a poset with infimum and the least upper bound (lub) of increasing chains exists. Reference [3] P. Cousot. Design of semantics by abstract interpretation, invited address. In Mathematical Foundations of Programming Semantics, Thirteenth Annual Conference (MFPS XIII),CarnegieMellonUniversity,Pittsburgh,Pennsylvania, USA, 23 26 March 1997. Kleenian fixpoint theorem 5 Amapϕ L c L on a cpo L,,, is upper-continuous iff it preserves lubs of increasing chains x i,i N: ϕ( x i )= ϕ(x i ); i N i N The least fixpoint of an upper-continuous map ϕ L c L on a cpo L,,, is: lfp ϕ = n 0 ϕ n ( ) where the iterates ϕ n (x) of ϕ from x are: - ϕ 0 (x) def = x; - ϕ n+1 (x) def = ϕ(ϕ n (x)) for all x L. 5 Can be generalized to monotonic non-continuous maps by considering transfinite iterates. NACSA 98 5 P. Cousot NACSA 98 7 P. Cousot Tarski s Fixpoint Theorem Amonotonicmapϕ L L on a complete lattice: has a least fixpoint: L,,,,, lfp ϕ = {x L ϕ(x) x} and, dually, a greatest fixpoint: gfp ϕ = {x L x ϕ(x)} Chaotic/asynchronous iterations Convergent iterates L = n 0 F n (P ) of a monotonic system of equations on a poset: X = F (X) X 1 = F 1 (X 1,...,X n )... X n = F n (X 1,...,X n ) starting from a prefixpoint (P F (P )) always converge to the same limit L whichever chaotic or asynchronous iteration strategy is used. NACSA 98 6 P. Cousot NACSA 98 8 P. Cousot

Example: reachability analysis Program: { X 1 } x := 1; { X 2 } while (x < 1000) do { X 3 } x := x + 1; { X 4 } od; { X 5 } System of equations: X 1 = {Ω} X 2 = {1} X 4 X 3 = {x X 2 x<1000} X 4 = {x +1 x X 3 } X 5 = {x X 2 x 1000} Reachable states: X 1 = {Ω} X 2 = {x 1 x 1000} X 3 = {x 1 x<1000} X 4 = {x +1 x X 3 } X 5 = {1000} Definition of Galois connections Given posets P, and Q,, agalois connection is a pair of maps such that: P Q Q P x P : y Q : (x) y x (y) in which case we write: P, Q, NACSA 98 9 P. Cousot NACSA 98 11 P. Cousot Effective fixpoint approximation Simplify the fixpoint system of semantic equations: Galois connections; Accelerate convergence of the iterates: widening/narrowing; Equivalent definition of Galois connections P, Q, Galois connection [ ] D m, Q, [ ] m Q, P, monotone monotone [ x P : x (x)] extensive [ y Q : (y) y] reductive NACSA 98 10 P. Cousot NACSA 98 12 P. Cousot

Duality principle We write 1 or for the inverse of the partial order. Observe that: P, Q, if and only if Q( ) P( ) duality principle: if a theorem is true for all posets, so is its dual obtained by substituting, >,,,,,, etc. respectively for, <,,,,,,, etc. If Example 2 of Galois connection ρ P Q (P) (Q) (X) =post[ρ]x post-image def = {y x X : x, y ρ} (Q) (P) (Y )= pre[ρ]y dual pre-image def = {x y : x, y ρ y Y } (P), (Q), NACSA 98 13 P. Cousot NACSA 98 15 P. Cousot Example 1 of Galois connection Example 3 of Galois connections If @ P Q (P) (Q) (X) def = {@(x) x X} (Q) (P) (Y ) def = {x @(x) Y } (P), (Q), direct image inverse image If S and T are sets (S T ), S (T ), where: (F ) def = λx {f(x) f F } (ϕ) def = {f S T x S : f(x) ϕ(x)} NACSA 98 14 P. Cousot NACSA 98 16 P. Cousot

Moore families A Moore family is a subset of a complete lattice L,,,,, containing and closed under arbitrary glbs ; If P, Q, and P,,,,, is a complete lattice (Q) is a Moore family. Aconsequenceisthatonecanreasonupon the abstract semantics using only P and the image of P by the upper closure operator (instead of Q). Intuition: - The upper-approximation of x P is any y (Q) such that x y; - The best approximation of x is (x). Preservation of lubs/glbs If P, Q,, preserves existing lubs: if X exists, ( X ) is the lub of {(x) x X}. By the duality principle: If P, Q, preserves existing glbs: if Y Q and Y exists, ( Y ) is the glb of {(y) y Y }. NACSA 98 17 P. Cousot NACSA 98 19 P. Cousot Unique adjoint In a Galois connection, one function uniquely determines the other: If P, 1 Q, and P, 2 1 2 Q,,( 1 = 2 ) if and only if ( 1 = 2 ). x P : (x) = {y x (y)} y Q : (y) = {x (x) y} Complete join preserving abstraction function and complete meet preserving concretization function Let P, and Q, be posets. If a 1. P( ) Q( ) 2. {x (x) y} exists for all y Q, P, Q, where y Q : (y) = {x (x) y} By duality, if a 1. Q( ) P( ) 2. {y x (y)} exists for all x P P, Q, where x P : (x) = {y x (y)} NACSA 98 18 P. Cousot NACSA 98 20 P. Cousot

Galois surjection & injection If P, Q,, : is onto iff is one-to-one iff is the identity By the duality principle, if P, Q,, : is one-to-one iff is onto iff is the identity Notation: P, Q, P, Q, P, Q, P, Q, Galois connection Galois surjection Galois injection Galois bijection with denoting into and denoting onto. The image of a complete lattice by a Galois surjection is a complete lattice If P, Q, and P,,,,, is a complete lattice, so is Q, with 0=( ) 1=( ) Y = ( (y)) y Y Y = ( (y)) y Y infimum supremum lub glb NACSA 98 21 P. Cousot NACSA 98 23 P. Cousot The image of a Cpo by a Galois surjection is a Cpo If P,,, is a cpo, Q, is a poset and P, Q, Q,, 0, is a cpo with: 0 def = ( ) X def = ( (x)) x X Pointwise extension of Galois connections If P, Q, : S P, S Q, where: (f) def = f (g) def = g NACSA 98 22 P. Cousot NACSA 98 24 P. Cousot

If Lifting Galois connections at higher-order P 1, 1 1 Q 1, 1 1 P 2, 2 2 Q 2, 2 2 m P 1 P 2, 2 where Q 1 m Q 2, 2 ϕ ψ def = x : ϕ(x) ψ(x) (ϕ) def = 2 ϕ 1 (ψ) def = 2 ψ 1 Example: interval analysis Concrete/exact: D def = {x N min int x max int} def D Ω = D {Ω} values & uninitialization n 1 program points V variables S def =[1,n] (V D Ω ) states Abstract/approximate: I def = {[a, b] {x N a x b}intervals (Ω) def = {Ω} concretization ([a, b]) def = {x N a x b} ( Ω, [a, b] ) def = (Ω) ([a, b]) L def abstract =[1,n] (V A) domain A (D Ω ) concretization (P ) def = {ρ i [1,n]: v V : ρ(i)(v) (P (i)(v))} P Q def = (P ) (Q) ordering Galois connexion: (S), L, NACSA 98 25 P. Cousot NACSA 98 27 P. Cousot Composition of Galois connections The composition of Galois connections is a Galois connection: ( P, 1 P, 1 P, 2 Q, ) 2 P, 1 2 Q, 2 1 Kleenian fixpoint abstraction If D,,, is a cpo, Q, is a poset, F P m D, F Q m Q, and F = F D, D, (lfp F )=lfp F NACSA 98 26 P. Cousot NACSA 98 28 P. Cousot

Kleenian fixpoint approximation If D,,, is a cpo, Q, is a poset, F P m D, F A m A, and F F D, D, (lfp F ) lfp F Infinite strictly increasing chains Because of infinite (or very long) strictly increasing chains, thefixpointiteratesmay not converge (or very slowly); Because of infinite (or very long) strictly decreasing chains,thelocaldecreasingiterates may not converge (or not rapidly enough); The design strategy of using a more abstract domain satisfying the ACC often yields too imprecise results; It is often both more precise and faster to speed up convergence using widenings along increasing chains and narrowings along deceasing ones. NACSA 98 29 P. Cousot NACSA 98 31 P. Cousot Interval lattice Slow fixpoint iterations program: 0: x := 1; 1: while true do 2: x := (x + 1) 3: od {false} 4: forward abstract equations: X0 = (INIT 0) X1 = assign[ x, 1 ](X0) U X3 X2 = assert[ true ](X1) X3 = assign[ x, (x + 1) ](X2) X4 = assert[ false ](X1) iterations from: X0 = { x:_o_ } X1 = _ _ X2 = _ _ X3 = _ _ X4 = _ _ X0 = { x:_o_ } X1 = { x:[1,1] } X2 = { x:[1,1] } X3 = { x:[2,2] } X1 = { x:[1,2] } X2 = { x:[1,2] } X3 = { x:[2,3] } X1 = { x:[1,3] } X2 = { x:[1,3] } X3 = { x:[2,4] } X1 = { x:[1,4] } X2 = { x:[1,4] } X3 = { x:[2,5] }... NACSA 98 30 P. Cousot NACSA 98 32 P. Cousot

Widening definition: A widening P P P on a poset P, satisfies: x, y P : x (x y) y (x y) For all increasing chains x 0 x 1... the increasing chain y 0 def = x 0,...,yn+A def = y n x n+1,... is not strictly increasing. use: Approximate missing lubs. Convergence acceleration; Fixpoint upper approximation by widening Any iteration sequence with widening is increasing and stationary after finitely many iteration steps; Its limit L is a post-fixpoint of F,whence an upper-approximation of the least fixpoint lfp F 6 : lfp F L 6 if lfp F does exist e.g. if P,,, is a cpo. NACSA 98 33 P. Cousot NACSA 98 35 P. Cousot Iteration sequence with widening Let F be a monotonic operator on a poset P, ; Let P P P be a widening; The iteration sequence with widening for F from is X n,n N: - X 0 = - X n+1 = X n if F (X n ) (X n ) - X n+1 = X n F (X n ) if F (X n ) X n Example of widening for intervals [a, b] [a,b ] def = [(a >= a? a a >=1?1 a >=0?0 a >= 1? 1 min int), (b <= b? b b <= 1? 1 b <=0?0 b <=1?1 max int] y def = y x def = x Ω Ω def = Ω Ω [a, b] def = Ω, [a, b] Ω Ω, [a, b] def = Ω, [a, b] [a, b] Ω def = Ω, [a, b] Ω, [a, b] Ω def = Ω, [a, b] [a, b] Ω, [a,b ] def = Ω, [a, b] [a,b ] Ω, [a, b] [a,b ] def = Ω, [a, b] [a,b ] Ω, [a, b] Ω, [a,b ] def = Ω, [a, b] [a,b ] NACSA 98 34 P. Cousot NACSA 98 36 P. Cousot

Widening for systems of equations Averyroughidea: compute the dependence graph of the system of equations; widen at cut-points; iterate accordingto the weak topological ordering Interval program analysis example with widening labelled program: 0: x := 1; 1: y := 1000; 2: while (x < y) do 3: x := (x + 1) 4: od 5: iterations with widening from: X0 = { x:_o_; y:_o_ } X1 = _ _ X2 = _ _ X3 = _ _ X4 = _ _ X5 = _ _ X0 = { x:_o_; y:_o_ } X1 = { x:[1,1]; y:_o_ } widening at 2 by { x:[1,1]; y:[1000,1000] } X2 = { x:[1,1]; y:[1000,1000] } X3 = { x:[1,1]; y:[1000,1000] } X4 = { x:[2,2]; y:[1000,1000] } widening at 2 by { x:[1,2]; y:[1000,1000] } X2 = { x:[1,+oo]; y:[1000,1000] } X3 = { x:[1,999]; y:[1000,1000] } X4 = { x:[2,1000]; y:[1000,1000] } X2 = { x:[1,1000]; y:[1000,1000] } X3 = { x:[1,999]; y:[1000,1000] } X4 = { x:[2,1000]; y:[1000,1000] } X5 = { x:[1000,1000]; y:[1000,1000] } NACSA 98 37 P. Cousot NACSA 98 39 P. Cousot Example labelled program: 0: x := 1; 1: y := 1000; 2: while (x < y) do 3: x := (x + 1) 4: od 5: forward abstract equations: X0 = (INIT 0) X1 = assign[ x, 1 ](X0) X2 = assign[ y, 1000 ](X1) U X4 X3 = assert[ (x < y) ](X2) X4 = assign[ x, (x + 1) ](X3) X5 = assert[ ((y < x) (x = y)) ](X2) forward graph with 6 vertices: 0 : {1} 1 : {2} 2 : {3, 5} 3 : {4} 4 : {2} 5 : {} forward weak topological order: 0 1 ( 2 3 4 ) 5 forward cut & check points: {2} Narrowing Since we have got a postfixpoint L of F P P,itsiterates F n (L ) are all upper approximations of lfp F. To accelerate convergence of this decreasing chain, we use a narrowing P P P on the poset P, satisfying: - x, y P : y x = y x y x - For all decreasing chains x 0 x 1... the decreasing chain y 0 def = x 0,...,yn+A def = y n x n+1,... is not strictly decreasing. NACSA 98 38 P. Cousot NACSA 98 40 P. Cousot

Decreasing iteration sequence with narrowing Let F be a monotonic operator on a poset P, ; Let P P P be a narrowing; The iteration sequence with narrowing for F from the postfixpoint P 7 is Y n,n N: - Y 0 = P - Y n+1 = Y n if F (X n )=X n - Y n+1 = Y n F (X n ) if F (X n ) X n Example of narrowing for intervals if x x y y [x, y] [x,y ]= narrow xyx y let narrow x y x y = (if (x = min_int) x else x), (if (y = max_int) y else y) ;; Trivially extended to initialization & interval analysis. 7 F (P ) P. NACSA 98 41 P. Cousot NACSA 98 43 P. Cousot Fixpoint upper approximation by narrowing Any iteration sequence with narrowing starting from a postfixpoint P of F 8 is decreasing and stationary after finitely many iteration steps; if lfp F does exist 9 and lfp F P its limit L is a fixpoint of F,whencean upper-approximation of the least fixpoint lfp F : lfp F L P 8 F (P ) P 9 e.g. if P,,, is a cpo. NACSA 98 42 P. Cousot Program analysis example with narrowing labelled program: 0: x := 1; 1: y := 1000; 2: while (x < y) do 3: x := (x + 1) 4: od {((y < x) (x = y))} 5: iterations with narrowing from: X0 = { x:_o_; y:_o_ } X1 = { x:[1,1]; y:_o_ } X2 = { x:[1,1000]; y:[1000,1000] } X3 = { x:[1,999]; y:[1000,1000] } X4 = { x:[2,1000]; y:[1000,1000] } X5 = { x:[1000,1000]; y:[1000,1000] } X0 = { x:_o_; y:_o_ } X1 = { x:[1,1]; y:_o_ } narrowing at 2 by { x:[1,1000]; y:[1000,1000] } X2 = { x:[1,1000]; y:[1000,1000] } X3 = { x:[1,999]; y:[1000,1000] } X4 = { x:[2,1000]; y:[1000,1000] } X5 = { x:[1000,1000]; y:[1000,1000] } stable NACSA 98 44 P. Cousot

Widenings and narrowings are not dual The iteration with widening starts from below the least fixpoints and stabilizes above; The iteration with narrowing starts from above the least fixpoints and stabilizes above; In general, widenings and narrowing are not monotonic. Conclusion Averyelementaryintroduction to abstract interpretation; For more details, see e.g. http://www.dmi.ens.fr/~cousot NACSA 98 45 P. Cousot NACSA 98 46 P. Cousot Improving the precision of widenings/narrowings Threshold; Widening/narrowing (and stabilization checks) at cut points; Computation history-based extrapolation: Asimpleexample: - Do not widen/narrow if a component of the system of fixpoint equations was computed for the first time since the last widening/narrowing ; - Otherwise, do not widen/narrow the abstract values of variables which were not assigned to 10 since the last widening / narrowing. 10 more precisely which did not appear in abstract equations corresponding to an assignment to these variables. NACSA 98 46 P. Cousot