arxiv: v1 [cs.cc] 18 Mar 2013

Similar documents
Space-aware data flow analysis

Dataflow Analysis - 2. Monotone Dataflow Frameworks

Static Program Analysis using Abstract Interpretation

Partial model checking via abstract interpretation

Discrete Fixpoint Approximation Methods in Program Static Analysis

The complexity of acyclic subhypergraph problems

Theory of Computation Chapter 1: Introduction

Disjointness conditions in free products of. distributive lattices: An application of Ramsay's theorem. Harry Lakser< 1)

Lecture 4. 1 Circuit Complexity. Notes on Complexity Theory: Fall 2005 Last updated: September, Jonathan Katz

arxiv:cond-mat/ v1 [cond-mat.stat-mech] 20 Jan 1997

Notes on Computer Theory Last updated: November, Circuits

Mean-Payoff Games and the Max-Atom Problem

GRAPH ALGORITHMS Week 7 (13 Nov - 18 Nov 2017)

ON THE NUMBER OF COMPONENTS OF A GRAPH

Min/Max-Poly Weighting Schemes and the NL vs UL Problem

3515ICT: Theory of Computation. Regular languages

BASIC CONCEPTS OF ABSTRACT INTERPRETATION

Polynomials in graph theory

Flow grammars a flow analysis methodology

Complexity Theory VU , SS The Polynomial Hierarchy. Reinhard Pichler

1 Non-deterministic Turing Machine

Outline. Complexity Theory EXACT TSP. The Class DP. Definition. Problem EXACT TSP. Complexity of EXACT TSP. Proposition VU 181.

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

Finite Automata and Regular Languages (part III)

Lecture #14: NP-Completeness (Chapter 34 Old Edition Chapter 36) Discussion here is from the old edition.

A Lower Bound for Boolean Satisfiability on Turing Machines

CPSC 421: Tutorial #1

Chapter 3: Proving NP-completeness Results

arxiv:cs/ v1 [cs.cc] 7 Sep 2006

Closure under the Regular Operations

Join Algorithms for the Theory of Uninterpreted Functions

NP-COMPLETE PROBLEMS. 1. Characterizing NP. Proof

DFA of non-distributive properties

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

Complexity. Complexity Theory Lecture 3. Decidability and Complexity. Complexity Classes

A Lower Bound of 2 n Conditional Jumps for Boolean Satisfiability on A Random Access Machine

Kernelization by matroids: Odd Cycle Transversal

CMSC 631 Program Analysis and Understanding. Spring Data Flow Analysis

arxiv: v1 [math.co] 20 Dec 2016

Equational Logic. Chapter Syntax Terms and Term Algebras

Lecture 8: Complete Problems for Other Complexity Classes

CSC D70: Compiler Optimization Dataflow-2 and Loops

A Characterisation of NL/poly via Nondeterministic Finite Automata.

Precise Interprocedural Analysis using Random Interpretation

Proving Inter-Program Properties

Peter Wood. Department of Computer Science and Information Systems Birkbeck, University of London Automata and Formal Languages

Reductions to Graph Isomorphism

The Maximum Flow Problem with Disjunctive Constraints

On the oriented chromatic index of oriented graphs

The Meet-Over-All-Paths Solution to Data-Flow Problems

Boolean complexity classes vs. their arithmetic analogs

Languages, regular languages, finite automata

Unmixed Graphs that are Domains

Precise Interprocedural Analysis through Linear Algebra

Model Checking & Program Analysis

Nested Interpolants. Matthias Heizmann Jochen Hoenicke Andreas Podelski. Abstract. 1. Introduction. University of Freiburg, Germany

Data flow analysis. DataFlow analysis

MIT Martin Rinard Laboratory for Computer Science Massachusetts Institute of Technology

Complexity of Regular Functions

An algebraic characterization of unary two-way transducers

Introduction. Foundations of Computing Science. Pallab Dasgupta Professor, Dept. of Computer Sc & Engg INDIAN INSTITUTE OF TECHNOLOGY KHARAGPUR

Enumerating the edge-colourings and total colourings of a regular graph

Equivalence of Regular Expressions and FSMs

On the Exponent of the All Pairs Shortest Path Problem

CpSc 421 Final Exam December 6, 2008

A note on monotone real circuits

Binary Decision Diagrams. Graphs. Boolean Functions

XMA2C011, Annual Examination 2012: Worked Solutions

Zero-Sum Flows in Regular Graphs

Reading: Chapter 9.3. Carnegie Mellon

P is the class of problems for which there are algorithms that solve the problem in time O(n k ) for some constant k.

Single Source Shortest Paths

c 1998 Society for Industrial and Applied Mathematics Vol. 27, No. 4, pp , August

Design of Distributed Systems Melinda Tóth, Zoltán Horváth

Dataflow Analysis. Chapter 9, Section 9.2, 9.3, 9.4

Domain Extender for Collision Resistant Hash Functions: Improving Upon Merkle-Damgård Iteration

Ταουσάκος Θανάσης Αλγόριθμοι και Πολυπλοκότητα II 7 Φεβρουαρίου 2013

Computability and Complexity Theory: An Introduction

About the relationship between formal logic and complexity classes

Kleene Algebras and Algebraic Path Problems

Interprocedurally Analyzing Polynomial Identities

Finite Presentations of Pregroups and the Identity Problem

Binary Decision Diagrams

Semantics and Verification of Software

MAL TSEV CONSTRAINTS MADE SIMPLE

Automata Theory for Presburger Arithmetic Logic

Four-coloring P 6 -free graphs. I. Extending an excellent precoloring

Boolean Algebra and Propositional Logic

Advanced topic: Space complexity

Out-colourings of Digraphs

Chapter 6. Properties of Regular Languages

MINIMALLY NON-PFAFFIAN GRAPHS

Discrete Math, Spring Solutions to Problems V

Automata Theory, Computability and Complexity

Boolean Inner-Product Spaces and Boolean Matrices

U.C. Berkeley CS278: Computational Complexity Professor Luca Trevisan 9/6/2004. Notes for Lecture 3

CSE200: Computability and complexity Space Complexity

Precise Interprocedural Analysis using Random Interpretation (Revised version )

UNIT II REGULAR LANGUAGES

X-MA2C01-1: Partial Worked Solutions

Transcription:

On the computational complexity of Data Flow Analysis Gaurav Sood gauravsood0289@gmail.com K. Murali Krishnan kmurali@nitc.ac.in Department of Computer Science and Engineering, National Institute of Technology Calicut, Calicut - 67360, Kerala, India arxiv:303.435v [cs.cc] 8 Mar 203 ABSTRACT We consider the problem of Data Flow Analysis over monotone data flow frameworks with a finite lattice. The problem of computing the Maximum Fixed Point (MFP) solution is shown to be P-complete even when the lattice has just four elements. This shows that the problem is unlikely to be efficiently parallelizable. It is also shown that the problem of computing the Meet Over all Paths (MOP) solution is N L-complete (and hence efficiently parallelizable) when the lattice is finite even for non-monotone data flow frameworks. These results appear in contrast with the fact that when the lattice is not finite, solving the MOP problem is undecidable and hence significantly harder than the MFP problem which is polynomial time computable for lattices of finite height. Categories and Subject Descriptors F..3 [Computation by Abstract Devices]: Complexity Measures and Classes Reducibility and completeness; D.3.4 [Programming Languages]: Processors Optimization; F.3.2 [Logics and Meanings of Programs]: Semantics of Programming Languages Program analysis General Terms Theory Keywords Maximum fixed point solution, Meet over all paths solution, P-complete, N L-complete. INTRODUCTION The problem of data flow analysis over a monotone data framework with a bounded meet semilattice has been well studied in the context of static program analysis and machine independent compiler optimizations [, Sec. 9.2, 9.3]. Although the meet over all paths (MOP) solution is desirable, its computation is undecidable in general [6]. Iterative fixed point methods as in [7] are commonly used to find the maximum fixed point (MFP) solution as a conservative approximation to the MOP solution [6]. Several important problems like reaching definitions analysis, live variable analysis and available expressions analysis [, Sec. 9.2.4 9.2.6] are essentially data flow analysis problems over monotone data flow frameworks with a bounded meet semilattice. In this paper, the computational complexity of MFP and MOP data flow analysis problems over monotone data flow frameworks with a finite bounded meet semilattice is investigated. Since a finite bounded meet semilattice is essentially a finite lattice, we define the problem over finite lattices. We show that computing the MFP solution to data flow analysis problem over a monotone data flow framework with a finite lattice is P-complete. We further show that the problem of finding MOP solution is non-deterministic log space complete (N L-complete). In fact the proof in Section 6 shows that MOP is N L-complete even if the functions associated with the vertices of the control flow graph are nonmonotone. These results indicate that the MFP problem is unlikely to be in the complexity class N C (and hence fast parallel algorithms are unlikely to exist for the problem [5, Chap. 5]). The N L complexity of MOP problem yields O(log 2 n) depth, polynomial sized parallel circuit for the problem. This further leads to the observation that although MOP computation is harder than MFP computation in general, computing MOP solution appears significantly easier than computing MFP over finite lattices. 2. BACKGROUND Let (L, ) be a partially ordered set. Let and respectively denote the join and meet operations in L. A partially ordered set (L, ) is a meet semilattice, denoted by (L, ), if x y exists for all x,y L. A meet semilattice (L, ) is a lattice, denoted by (L,, ), if x y exists for all x,y L. A meet semilattice (L, ) is a bounded meet semilattice, denoted by (L,,), if there exists an element L suchthatl = l for all l L. Alattice (L,, )isa bounded lattice, denoted by (L,,,0,), if there exist elements 0, L such that l 0 = l and l = l for all l L. A finite bounded meet semilattice (L,,) is essentially a finite lattice where join operation is defined as follows: For all l,l L, l l = {l L l l and l l }. A lattice is complete lattice if S and S exist for all S L. It is easy to see that a finite lattice is complete. [3] Let L = {l,l 2,...l m} be a finite lattice. Let L n = { (l i,...l in ) j n,l ij L }. The tuple (l i,l i2,...l in ) L n will be denoted by l i n or simply by l i when there is no ambiguity about the index set. Let l j denote the j th element of l i. For all l i, l i L n, l i l i = l i l i and l i l i = l i l i. [3] Afunction f: A B is monotone if for all x,y A, x y implies f(x) f(y). Let G = (V,E) be a directed graph. Let deg (v) and deg + (v) respectively denote the indegree and outdegree of vertex v. The function pred: V 2 V is defined as follows pred(v) = {u (u,v) in E}

Definition. Let L be a lattice and let f: L L. An element l L is called a fixed point of f if f(l) = l. An element l L is called the maximum fixed point of f if it is a fixed point of f and for every l L whenever f(l ) = l then l l. Let MFP(f) denote the maximum fixed point of f whenever it exists. 2. Data Flow Analysis (DFA) Definition 2. A control flow graph is a finite directed graph G = (V,E,v s,v t) where V = {v,v 2,...v n}, v s V, called entry, is a unique vertex satisfying deg (v s) = 0 and v t V, called exit, is a unique vertex satisfying deg + (v t) = 0. Every vertex v i V is reachable from v s. Definition 3. A monotone data flow framework [6] is a triple D = (L,,F) where (L,,,0,) is a finite lattice where L = {l,l 2,...l m} with l = 0 and l m = ; is the confluence operator; and F is a collection of monotone functions from L to L. Definition 4. A Data Flow Analysis (DFA) system is a 5-tuple α = (G,D,M,v θ,l φ ) where G is a control flow graph; D is a monotone data flow framework; M: V F assigns a function f i F to the vertex v i of G; and v θ V, l φ L. 2.2 Maximum Fixed Point (MFP) problem Let (G,D,M,v θ,l φ ) be a DFA system where G = (V,E,v s,v t) and D = (L,,F). Then f: L n L n is defined as follows li f( ) ( = f i v j pred(v i ) l j ) () Here we assume that S = when S =. With this convention, it is easy to see that f is well defined and monotone in L n. Theorem (Knaster Tarski theorem [2]). Let (L,,,0,) be a complete lattice and let f: L L be a monotone function. Then the MFP of f exists and is unique. Since every finite lattice is complete, it is clear from Theorem that the MFP of f exists and is unique. Suppose l i is the MFP of f, then we use the notation MFP(v j) for the element l j. Definition 5. Maximum Fixed Point (MFP) problem: Given a DFA system (G,D,M,v θ,l φ ), decide whether MFP(v θ ) = l φ. 2.3 Meet Over all Paths (MOP) problem A path p from vertex v i to vertex v ik in a graph G, called a v i -v ik path, is a non-empty alternating sequence v i e i v i2 e i2...e ik v ik of vertices and edges such that e ij = (v ij,v ij+ ) for all j k. A v i -v ik path is written simply as v i v i2...v ik when the edges in question are clear. It may be noted that vertices and edges on a path may not be distinct. The length of path p is denoted by len(p). Let f p = f ik f i be called the path function associated with path p. Let P ij be the set of all paths from vertex v i to vertex v j in G. Definition 6. Given a DFA system α = (G,D,M,v θ,l φ ), the meet over all paths solution, denoted MOP, is defined as follows MOP(v i) = p P si f p() Since L n is finite and hence complete, though there could be infinitely many v s-v i paths, MOP(v i) is well defined by taking the infimum of all path functions. Definition 7. Meet Over all Paths (MOP) problem: Given a DFA system (G,D,M,v θ,l φ ), decide whether MOP(v θ ) = l φ. 2.4 Monotone Circuit Value (MCV) problem This problem is used for reduction in Section 4 to prove that MFP is P-complete. Definition 8. A monotone Boolean circuit [5, p. 27, 22] is a 4-tuple C = (G,I,v θ,τ) where G = (V,E) is afinite directed acyclic graph where V = {v,v 2,...v n}, and for all v i V,deg (v i) {0,2}; I = { v i V deg (v i) = 0 } is the set of input vertices; v θ V, called output, is the unique vertex in G satisfying deg + (v θ ) = 0; and τ: V {,+} assigns either the Boolean AND function (denoted by ) or the Boolean OR function (denoted by +) to each vertex of G. Let u j be the j th input vertex of a Boolean circuit C and let x i I {0,} I bethe inputtothe circuit. The input value assignment is a function ν: I {0,} defined as follows ν(u j) = x j u j I,x j {0,} The function ν: I {0,} can be extended to the function ν: V {0,} called value of a node defined as follows ν(v i) ν(v j) if τ(v k ) = and pred(v ν(v k ) = k ) = {v i,v j}, ν(v i)+ν(v j) if τ(v k ) = + and pred(v k ) = {v i,v j} It is easy to see that ν is well defined when G is a directed acyclic graph. Definition 9. An instance of Monotone Circuit Value (MCV) problem is a pair (C,ν) with C = (G,I,v θ,τ) where C is a monotone Boolean circuit; and 2

ν: I {0,} is an input value assignment. Definition 0. Monotone Circuit Value (MCV) problem: Given an instance (C,ν) of MCV, decide whether ν(v θ ) = [5, p. 22]. 2.5 Graph Meet Reachability (GMR) problem This problem will be used as an intermediate problem in Section 6 for showing that MOP is N L-complete. Definition. Let A = {a,a 2,...a n} be a finite set and (L,,,0,) be a finite lattice where L = {l,l 2,...l m}. A directed graph G = (V,E) is said to be a product graph of A and L if V = {v ij a i A,l j L} is the set of vertices; and E V V is the set of directed edges. Definition 2. An instance of Graph Meet Reachability (GMR) problem is a 6-tuple (G,A,L,v θφ,a θ,l φ ) where G = (V,E) is a product graph of A and L; v θφ V where a θ A and l φ L; a θ A; and l φ L. Let R i = {l j v ij is reachable from v θφ } Definition 3. Graph Meet Reachability (GMR) problem: Given an instance (G,A,L,v θφ,a θ,l φ ) of GMR, decide whether l i = l φ l i R θ 2.6 Graph Reachability (GR) problem Graph Reachability problem is a well known N L- complete problem which will be used for reduction in this paper. Definition 4. An instance of Graph Reachability (GR) problem is a triple (G,v s,v t) where G = (V,E) is a directed graph; and v s,v t V Definition 5. Graph Reachability (GR) problem: Given an instance (G,v s,v t) of GR, decide whether v t is reachable from v s. Fact. GR is N L-complete [8, Theorem 6.2 on p. 398]. 3. RELATED WORK It is shown in [9] that the problem of finding meet over all valid paths (MVP) solution to the interprocedural data flow analysis over a distributive data flow framework with possibly infinite (resp. finite subset) semilattice is P-hard (resp. P-complete). It is shown in [9, 0] that the problem of finding MFP and MOP solution to data flow analysis over a distributive data flow framework with a distributive sublattice of the power set lattice of a finite set is reducible to graph reachability problem. Hence the problem is non-deterministic log space computable i.e., belongs to the complexity class N L (see [8, p. 42] for definition). Since N L N C [8, Theorem 6. on p. 395], [2] and N C admits fast parallel solutions, these results show that the above problem admits fast parallel algorithms. The following is an outline for rest of the paper. In Section 4, we show that MFP is P-complete by reduction from MCV. In Section 5, we give an N L algorithm for computing GMR. In Section 6, we prove that MOP is log space reducible to GMR thereby showing that MOP is in N L. Completeness of MOP w.r.t. the class N L follows easily by a log space reduction from GR to MOP. 4. MFP IS P-COMPLETE In this section, we give a log space reduction from MCV to MFP. Since MFP is in P [7] and MCV is P-complete [4], the reduction implies that MFP is also P-complete. 4. A reduction from MCV to MFP Given an instance α = (C,ν) of MCV with C = (G,I,v θ,τ). Construct an instance of MFP α = (G,D,M,v θ,(,)) as follows G = (V,E,v0,v θ) where V = { } { v0 v 0 i v i V(G)\I } { v i v i V(G) } E is defined as follows For each input vertex v i I add the edge (v 0,v i) to E for each vertex v i V add the edge (v 0 i,v i) to E for each vertex v k V \I with predecessors v i and v j with i < j add the edges (v i,v 0 k), (v j,v k) to E. Note that each v 0 k V \I has a unique predecessor in G. D = (L,,F) is defined as follows L = {(0,0),(0,),(,0),(,)} where = bitwise + operation in {0,} {0,} = bitwise operation in {0,} {0,} is the confluence operator F = {g I,g 0,g,g sw,g,g +} where g I: L L is the identity function g 0: L L is defined as follows g 0((a,a 2)) = (,0) g : L L is defined as follows g ((a,a 2)) = (,) (a,a 2) L (a,a 2) L The swap function g sw: L L is defined as follows g sw((a,a 2)) = (a 2,a ) g : L L is defined as follows g ((a,a 2)) = (,a a 2) (a,a 2) L (a,a 2) L 3

v v 2 v 3 v 0 g I x x 2 x 3 v 4 v g v2 x g x2 g x3 v3 x x 2 (,x ) (,x 2) (,x 3) + v 5 v 0 4 g s (x x 2)+x 3 (x,) Figure : An instance of MCV v 4 g g +: L L is defined as follows (,x x 2) g +((a,a 2)) = (,a +a 2) (a,a 2) L It is easy to see that all functions in F are monotone. g s v 0 5 (x 3,) M: V F is defined as follows M(v i) = f i = M(v 0) = f 0 = g I M(v 0 i) = f 0 i = g sw v 0 i V { g0 if v i I and ν(v i) = 0 g if v i I and ν(v i) = More compactly fi = g ν(vi ) if i I. { M(vi) = fi g if v = i V \I and τ(v i) = g + if v i V \I and τ(v i) = + Figure shows an instance of MCV where I = {v,v 2,v 3}. Figure 2 shows an instance of MFP constructed from the MCV instance of Figure. 4.2 Proof of correctness Let G = (V,E) be a directed acyclic graph. level(v): V N is defined as follows { level(v i) = + max v j pred(v i ) 0 if deg (v i) = 0, level(v j) if deg (v i) > 0 It is easy to see that level function is well defined. Lemma. Let α = (C,ν) be an instance of MCV with C = (G,I,v θ,τ) and G = (V,E). Let α = (G,D,M,v θ,(,)) be the instance of MFP as constructed in Section 4.. Then f has a unique fixed point. For all v i V, MFP(v i) = (,ν(v i)) and for all v i V \I, if v j is the predecessor of v 0 i in G, then MFP(v 0 i) = (ν(v j),). Proof. MFP(f) exists by Theorem. Therefore, f has at least one fixed point. Let l i be an arbitrary fixed point of f. Let l 0 j and l j denote the elements of l i corresponding to vertices v 0 j and v j respectively. We first prove that l k i is uniquely defined for all vertices v k i of G. g + v 5 (,(x x 2)+x 3 ) Figure 2: Data flow graph corresponding to MCV instance in Figure Since l i is a fixed point of f, so, f( l i ) = l i. By Equation (), fi k k ( j ) = l i. So, l k i = f k i ( v k j pred(vk i ) l v k j pred(vk i ) l k j ) v k i V (2) l 0 = g I( ) = g I() = is uniquely defined. Let v i be an arbitrary vertex of V. We prove the uniqueness of l k i by induction on level(v i). Base case: level(v i) = 0 i.e. v i I. So, f i = g ν(vi ) and g ν(vi )((a,a 2)) = (,ν(v i)) (a,a 2) L by definition. From Equation 2, l i = f i (l 0) = g ν(vi )(l 0) = (,ν(v i)) is uniquely defined. Inductive step: Let the theorem be true v i V such that level(v i) < m. Let level(v k ) = m. Let v i,v j,i < j be predecessors of v k in G. By definition of level function, level(v i) < level(l k ) = m and level(v j) < level(l k ) = m. By induction hypothesis, l i = (,ν(v i)) and l j = (,ν(v j)). From Equation 2, l 0 k = g sw(l i) = g sw((,ν(v i))) = (ν(v i),) is uniquely defined. Let τ(v k ) =. From Equation 2, l k = f k(l 0 k l j ) = g ((ν(v i),) (,ν(v j))) = g (ν(v i),ν(v j)) = (,ν(v i) ν(v j)) = (,ν(v k )) is uniquely defined. The case τ(v k ) = + is proved similarly. 4

Since l k i is uniquely defined for all vertices v k i of G, l i is unique and hence l i is the maximum fixed point of f. So, for all v i V, MFP(v i) = l i = (,ν(v i)) and for all v i V \I, if v j is the predecessor of v 0 i in G, then MFP(v 0 i) = l 0 i = (ν(v j),). Corollary. Let α = (C,ν) be an instance of MCV with C = (G,I,v θ,τ) and G = (V,E). Let α = (G,D,M,v θ,(,)) be the instance of MFP as constructed in Section 4.. ν(v θ ) = MFP(v θ) = (,). Theorem 2. MFP is P-complete. Proof. A polynomial time algorithm for MFP is given in [7]. MCV is shown P-complete in [4]. It is easy to see that the above reduction is computable in log space. Hence MFP is P-complete. 5. AN ALGORITHM FOR GMR Algorithm is an algorithm for deciding GMR. Algorithm Algorithm for GMR : procedure GMRA(G,A,L,v θφ,a θ,l φ ) 2: temp 3: for i,n do 4: 5: if v θ i is reachable from v θφ then temp temp l i 6: end if 7: end for 8: if temp = l φ then 9: return True 0: else : return False 2: end if 3: end procedure The observation below is a direct consequence of the above algorithm. Lemma 2. Let α = (G,A,L,v θφ,a θ,l φ ) be an instance of GMR. Let R i = {l j v ij is reachable from v θφ }. Then Algorithm returns true l i = l φ. l i R θ Lemma 3. GMR is computable in non-deterministic log space. Proof. Variables temp and i take Θ(log L ) space. Since graph reachability takes up only non-deterministic log space [8, Example 2.0 on p. 48]), Line 4 takes non-deterministic O(log V ) = O(log( A L )) space. So, GMR is computable in non-deterministic log space. 6. MOP IS NL-COMPLETE In this section, we give a log space reduction from MOP to GMR. Since GMR is non-deterministic log space computable, this implies that MOP can also be computed in non-deterministic log space. 6. A reduction from MOP to GMR Given an instance α = (G,D,M,v θ,l φ ) of MOP with G = (V,E,v s,v t), D = (L,,F) and L = {0 = l,l 2,...l m = }. Construct an instance of GMR α = (G,A,L,v 0 sm,v θ,l φ ) as follows This reduction is similar to the reduction in [0]. A = { v 0 i v i V(G) } { v i v i V(G) } G = (V,E ) where V = { v k ij v k i A,l j L } E is defined as follows For each vertex v i V,l j L, if f i(l j) = l k, add the edge (v 0 ij,v ik) to E. For each edge (v i,v j) E,l k L, add the edge (v ik,v 0 jk) to E. Example. Figure 3 shows a Data Flow Graph and a lattice. Let the set F of monotone functions be defined as follows: F = {f,f 2,f 3,f 4,f 5,f 6} where f = {(l,l ),(l 2,l 4),(l 3,l 4),(l 4,l 3),(l 5,l 5)} f 2 = {(l,l ),(l 2,l 3),(l 3,l 3),(l 4,l 5),(l 5,l 5)} f 3 = {(l,l 2),(l 2,l 3),(l 3,l 3),(l 4,l 2),(l 5,l 3)} f 4 = {(l,l ),(l 2,l 3),(l 3,l 5),(l 4,l 4),(l 5,l 5)} f 5 = {(l,l 2),(l 2,l 2),(l 3,l 5),(l 4,l 3),(l 5,l 5)} f 6 = {(l,l ),(l 2,l 4),(l 3,l 5),(l 4,l 4),(l 5,l 5)} Figure 4 shows the corresponding product graph. 6.2 Proof of correctness Lemma 4. Let α = (G,D,M,v θ,l φ ) be an instance of MOP with G = (V,E,v s,v t) and D = (L,,F). Let α = (G,A,L,v 0 sm,v θ,l φ ) be an instance of GMR as constructed in Section 6.. Then for all v i V, there exists a v s-v i path p in G such that f p() = l j v ij is reachable from v 0 sm in G. Proof. If part: Let v i be an arbitrary vertex in V. Let p be a v s-v i path in G and let f p() = l j for some l j L. We prove the if part by induction on len(p). Base case: Let len(p) = 0 i.e. v i = v s. Then l j = f p() = f s(). Then (v 0 sm,v sj) E. So, v sj is reachable from v 0 sm in G. Inductive step: Let len(p) = k and let the if part be true for all paths from v s in G with length less than k. Let p = v s,...,v i,v i for some v i V. Let p be the path p with v i excluded i.e. p = p v i where is the path concatenation operator. Therefore, p is a path from v s with length k. Let f p() = l j and f i(l j ) = l j for some l j L. By induction hypothesis, v i j is reachable from v0 sm in G. By construction of E, (v i j,v0 ij ),(v0 ij,v ij) E. So, v ij is reachable from v 0 sm in G. So, the if part is true for all v s-v i paths in G. Only if part: Let v ij is reachable from v 0 sm in G. Let p be a path in G from v 0 sm to v ij. It is easy to show that length of p is odd. So, we prove the only if part by induction on len(p) where len(p) is odd. Base case: Let len(p) = i.e. (v 0 sm,v ij) E. By construction of E, v i = v s and f s() = l j. So, there exists a trivial v s-v i path q in G, the path having only one node v s, such that f q() = l j. 5

v s = v f v 0 v3 0 v5 0 v v 5 v 2 f 2 = l 5 v 0 2 v 0 25 l 3 v 2 v 25 v 3 f 3 f 4 v 4 l 4 l 2 v 0 3 v 0 35 v 0 4 v 0 45 0 = l v 3 v 35 v 4 v 45 v 5 f 5 v 0 5 v 0 55 v 5 v 55 v t = v 6 f 6 Figure 3: A Data Flow Graph and a finite lattice v 0 6 v 0 65 Inductive step: Let len(p) = 2k + for some integer k and let the only if part be true for all odd length paths from v sm in G with length less than 2k+. Let p = v sm,...,v i j,v0 ij,v ij for some v i V,l j L. Let p be the path p with v 0 ij and v ij excluded i.e. p = p v 0 ij v ij where is the path concatenation operator. By induction hypothesis, there exists a v s-v i path q in G such that f q() = l j. By construction of E, (v i,v i) E and f i(l j ) = l j. So, there exists a v s-v i path q = q v i in G such that f q () = f i f q() = f i(f q()) = f i(l j ) = l j. So, the only if part is true. So, the theorem is true. Corollary 2. Let α = (G,D,M,v θ,l φ ) be an instance of MOP with G = (V,E,v s,v t) and D = (L,,F). Let α = (G,A,L,v 0 sm,v θ,l φ ) be an instance of GMR as constructed in Section 6.. Let R θ = { l i v θi is reachable from v 0 sm}. Then MOP(v θ ) = l φ l i R θ l i = l φ. Theorem 3. MOP is N L-complete. Proof. Algorithm is N L computable. It is easy to see that the above reduction is log space computable. So, MOP is in N L. Since Graph Reachability (GR) problem is an instance of MOP and GR is N L-complete, it follows that MOP is also N L-complete. It is easy to see that the above proofs do not use the monotonicity of the data flow framework. So, the N L- completeness result holds even if the data flow framework is not monotone. v6 v63 v65 Figure 4: The product graph corresponding to MOP instance in Figure 3 7. ACKNOWLEDGMENT We would like to thank Dr. Vineeth Paleri and Ms. Rekha R. Pai for introducing us to the problem and for helpful discussions. We would also like to thank Dr. Priya Chandran for reviewing our work. 8. REFERENCES [] A. V. Aho, M. S. Lam, R. Sethi, and J. D. Ullman. Compilers: Principles, Techniques, and Tools (2nd Edition). Prentice Hall, 2 edition, Sept. 2006. [2] A. Borodin. On relating time and space to size and depth. SIAM Journal on Computing, 6(4):733 744, 977. [3] B. A. Davey and H. A. Priestley. Introduction to Lattices and Order. Cambridge University Press, 2 edition, May 2002. [4] L. M. Goldschlager. The monotone and planar circuit value problems are log space complete for p. SIGACT News, 9(2):25 29, July 977. [5] R. Greenlaw, H. J. Hoover, and W. L. Ruzzo. Limits to Parallel Computation: P-Completeness Theory. Oxford University Press, USA, Apr. 995. [6] J. B. Kam and J. D. Ullman. Monotone data flow analysis frameworks. Acta Informatica, 7:305 37, 977. 0.007/BF00290339. [7] G. A. Kildall. A unified approach to global program optimization. In Proceedings of the st annual ACM 6

SIGACT-SIGPLAN symposium on Principles of programming languages, POPL 73, pages 94 206, New York, NY, USA, 973. ACM. [8] C. H. Papadimitriou. Computational Complexity. Addison Wesley, 993. [9] T. Reps. On the sequential nature of interprocedural program-analysis problems. Acta Informatica, 33:739 757, 996. 0.007/BF03036473. [0] T. Reps, S. Horwitz, and M. Sagiv. Precise interprocedural dataflow analysis via graph reachability. In Proceedings of the 22nd ACM SIGPLAN-SIGACT symposium on Principles of programming languages, POPL 95, pages 49 6, New York, NY, USA, 995. ACM. [] M. Sipser. Introduction to the Theory of Computation. Course Technology, 2 edition, Feb. 2005. [2] A. Tarski. A lattice-theoretical fixpoint theorem and its applications. Pacific journal of Mathematics, 5(2):285 309, 955. APPENDIX A. LATTICE REPRESENTATION Algorithm 2 Converts a lattice given as a covering relation to a lattice as a poset : procedure CovRel-to-Poset(L, ) 2: for i,n do 3: for j,n do 4: Set l i l j to value False 5: end for 6: end for 7: for i,n do 8: for j,n do 9: if l j is reachable from l i then 0: Set l i l j to value True : end if 2: end for 3: end for 4: end procedure A lattice can be represented as a poset (L, ) [3, p. 33], as a covering relation for a poset (L, ) [3, p. ] or as an algebraic structure (L,, ) [3, p. 39]. In this section, we give a non-deterministic log space algorithms to convert a lattice given as a poset or a covering relation to a lattice as an algebraic structure. This makes the completeness result of MOP w.r.t. the class N L independent of the particular representation of the lattice. A covering relation (L, ) of a poset can be viewed as a graph where L is the set of vertices and is the set of edges. So, reachability is defined as it is done for a graph. Algorithm 3 Converts a lattice given as a poset to a lattice as an algebraic structure : procedure Poset-to-AlgStr(L, ) 2: for i,n do 3: for j,n do 4: Set l i lj to value 5: Set l i lj to value 0 6: end for 7: end for 8: for i,n do 9: for j,n do 0: for k,n do : if l i l k and l j l k then 2: if l k l i lj then 3: Set l i lj to value l k 4: end if 5: else if l k l i and l k l j then 6: if l i lj l k then 7: Set l i lj to value l k 8: end if 9: end if 20: end for 2: end for 22: end for 23: end procedure Algorithm 2 (resp. Algorithm 3) converts a lattice given as a covering relation of a poset (resp. a poset) to the lattice as a poset (resp. an algebraic structure). The composition of the two algorithms converts a lattice given as a covering relation of a poset to the lattice as an algebraic structure. Lemma 5. Given a poset (L, ) or a covering relation (L, )) representation of a lattice, its algebraic structure representation (L,, ) can be computed in non-deterministic log space. Proof. Line 9 of Algorithm 2 takes non-deterministic log space since GR takes non-deterministic log space [8, Example 2.0 on p. 48]. All other lines of the two algorithms take at most log space. So, Algorithm 2 takes non-deterministic log space while Algorithm 3 takes log space. The composition of the two algorithms takes non-deterministic log space [, Theorem 8.23 on p. 324]. So, the conversions can be done in non-deterministic log space. It may be noted that the lattice in Section 4 is of constant size. So, it can be converted to a poset or a covering relation in constant time. So, the completeness result of MFP w.r.t. the class P is also independent of the particular representation of the lattice. 7