<<

1 TLA in Pictures Leslie Lamport

Abstract— Predicate-action diagrams, which are similar to aproof. standard state-transition diagrams, are precisely defined as formulas of TLA (the Temporal Logic of Actions). We ex- II. TLA plain how these diagrams can be used to describe aspects of a specification—and those descriptions then proved correct— We now describe the syntax and semantics of TLA. The even when the complete specification cannot be written as description is illustrated with the formulas defined in Fig- a diagram. We also use the diagrams to illustrate proofs. ∆ ure 1. (The symbol =meansequals by definition.) Keywords— Concurrency, specification, state-transition di- agrams, temporal logic. We assume an infinite set of variables (such as x and y) and a class of semantic values. Our variables are the flexible variables of temporal logic, which are analogous to I. Introduction variables in a programming language. TLA also includes Pictures aid understanding. A simple flowchart is easier the rigid variables of predicate logic, which are analogous to understand than the equivalent programming-language to constant parameters of a program, but we ignore them text. However, complex pictures are confusing. A large, here. The class of values includes numbers, strings, sets, spaghetti-like flowchart is harder to understand than a and functions. properlystructured program text. A state is an assignment of values to variables. A behav- Pictures are inadequate for specifying complex systems, ior is an infinite sequence of states. Semantically, a TLA but theycan help us understand particular aspects of a formula is true or false of a behavior. Syntactically, TLA system. For a picture to provide more than an informal formulas are built up from state functions using Boolean comment, there must be a formal connection between the operators (¬, ∧, ∨, ⇒ [implication], and ≡ [equivalence]) complete specification and the picture. The assertion that and the operators and ✷, as described below. TLA also the picture is a correct description of (some aspect of) the has a hiding operator ∃, which we do not use here. system must be a precise mathematical statement. A state function is a nonBoolean expression built from We use TLA (the Temporal Logic of Actions) to specify variables, constants, and constant operators. Semantically, systems. In TLA, a specification is a logical formula de- it assigns a value to each state—for example x + 1 assigns scribing all possible correct behaviors of the system. As an to state s one plus the value that s assigns to the vari- aid to understanding TLA formulas, we introduce here a able x.Astate predicate (often called just a predicate)is type of picture called a predicate-action diagram. These di- a Boolean expression built from variables, constants, and agrams are similar to the various kinds of state-transition constant operators such as +. Semantically, it is true or diagrams that have been used for years to describe sys- false for a state—for example the predicate Init Φ is true of tems, starting with Mealyand Moore machines [1], [2]. We state s iff s assigns the value zero to both x and y. relate these pictures to TLA specifications byinterpreting An action is a Boolean expression containing primed and a predicate-action diagram as a TLA formula. A diagram unprimed variables. Semantically, an action is true or false denoting formula D is a correct description of a system of a pair of states, with primed variables referring to the with specification S iff (if and onlyif) S implies D.We second state—for example, action M1 is true for s, t iff therefore provide a precise statement of what it means for the value that state t assigns to x equals one plus the value a diagram to describe a specification. that state s assigns to x, and the values assigned to y by We use predicate-action diagrams in three ways that we states s and t are equal. A pair of states satisfying an believe are new for a preciselydefined formal notation: action A is called an A step.Thus,anM1 step is one that • To describe aspects of a specification even when it is increments x byone and leaves y unchanged. not feasible to write the complete specification as a If f is a state function or state predicate, we write f diagram. for the expression obtained bypriming all the variables of • To draw different diagrams that provide complemen- f. For example (x +1) equals x +1,andInit Φ equals taryviews of the same system. (x =0)∧(y =0).ForanactionA and a state function v, • To illustrate formal correctness proofs. Section II is a brief review of TLA; a more leisurelyin- ∆ Init Φ =(x =0) ∧ (y =0) troduction to TLA appears in [3]. Section III describes ∆   M1 =(x = x +1) ∧ (y = y) predicate-action diagrams, using an n-input Muller C- ∆   M2 =(y = y +1) ∧ (x = x) element as an example. It shows how diagrams are used to ∆ describe aspects of a complete specification, and to provide M = M1 ∨M2 ∆ complementaryviews of a system. Section IV gives another Φ = Init Φ ∧ ✷[M]x, y ∧ WFx, y(M1) ∧ WFx, y(M2) example of how predicate-action diagrams are used to de- Fig. 1. The TLA formula Φ describing a simple program that re- scribe a system, and shows how they are used to illustrate peatedly increments x or y. 2

A A∨ A we define [ ]v to equal (v = v),soa[ ]v step is either E in[1] an A step or a step that leaves the value of v unchanged. n ✲ C M v - Thus, a [ 1]x,y step is one that increments x byone and i  . E leaves y unchanged, or else leaves the ordered pair x, y ✲ r . l ✏ unchanged. Since a tuple is unchanged iff each component o e n m is unchanged, a [M1]x,y step is one that increments x by m e in[n] e one and leaves y unchanged, or else leaves both x and y n ✲ n t unchanged. We define A v to equal A∧(v = v), so an t M M 1 x,y step is an 1 step that changes x or y.Sincean ✒ ✑ M1 step leaves y unchanged, an M1 x,y step is a step out that increments x by1, changes the value of x,andleaves y unchanged. Fig. 2. A Muller C-element. We saythat an action A is enabled in state s iff there exists a state t such that s, t is an A step. For example, often. Putting this all together, we see that Φ is true M1 is enabled iff it is possible to take a step that incre- of a behavior iff (i) x and y are initiallyzero, (ii) every ments x byone, changes x,andleavesy unchanged. Since step increments either x or y byone and leaves the other x +1= x for anynatural number x,actionM1 x,y is unchanged or else leaves both x and y unchanged, and enabledinanystateinwhichx is a natural number. If (iii) both x and y are incremented infinitelymanytimes. ∞ +1equals∞,thenM1 x,y is not enabled in a state The formula Init ∧ ✷[N ]v is a safetyproperty[4]. It in which x equals ∞. describes what steps are allowed, but it does not require A TLA formula is true or false of a behavior. A predicate anything to happen. (The formula is satisfied by a be- is true of a behavior iff it is true of the first state. An action havior satisfying the initial condition in which no variables is true of a behavior iff it is true of the first pair of states. ever change.) Fairness conditions are used to specifythat As usual in temporal logic, if F is a formula then ✷F is something must happen. the formula meaning that F is always true. Thus, ✷Init φ is true of a behavior iff x and y equal zero for everystate in III. Predicate-Action Diagrams the behavior. The formula ✷[M]x,y is true of a behavior A. An Example iff each step (pair of successive states) of the behavior is a We take as an example a Muller C-element [5]. This is a [M]x,y step. circuit with n binaryinputs in[1],...,in[n] and one binary ✷ Using and “enabled” predicates, we can define fairness output out, as shown in Figure 2. As the figure indicates, A operators WF and SF. The weak fairness formula WFv( ) we are considering the closed system consisting of the C- A asserts of a behavior that there are infinitelymany v element together with its environment. Initially, all the A steps, or there are infinitelymanystates in which v is inputs and the output are equal. The output becomes 0 A A not enabled. In other words, WFv( )assertsthatif v when all the inputs are 0, and it becomes 1 when all the A becomes enabled forever, then infinitelymany v steps inputs are 1. After an input changes, it must remain stable A occur. The strong fairness formula SFv( )assertsthatei- until the output changes. A ther there are infinitelymany v steps, or there are only The behavior of a 2-input C-element and its environment A finitelymanystates in which v is enabled. In other is described bythe predicate-action diagram of Figure 3(a), A A words, SFv( ) asserts that if v is enabled infinitelyof- where C is defined by ten, then infinitelymany A v steps occur. ∆ The standard form of a TLA specification is Init ∧ C(i, j, k) =(in[1] = i) ∧ (in[2] = j) ∧ (out = k) ✷[N ]v ∧ L,whereInit is a predicate, N is an action, v is a state function, and L is a conjunction of fairness con- The short arrows, with no originating node, identifythe ditions. This formula asserts of a behavior that (i) Init is nodes labeled C(0, 0, 0) and C(1, 1, 1) as initial nodes. true for the initial state, (ii) everystep of the behavior is an Theyindicate that the C-element starts in a state satis- N step or leaves v unchanged, and (iii) L holds. Formula fying C(0, 0, 0) or C(1, 1, 1). The arrows connecting nodes Φ of Figure 1 is in this form, asserting that (i) initially x indicate possible state transitions. For example, from a and y both equal zero, (ii) everystep either increments x state satisfying C(1, 1, 1), it is possible for the system byone and leaves y unchanged, increments y byone and to go to a state satisfying either C(0, 1, 1) or C(1, 0, 1). leaves x unchanged, or leaves both x and y unchanged, and More precisely, these arrows indicate all steps in which (iii) the fairness condition WFx, y(M1) ∧ WFx, y(M2) the triple in[1], in[2], out changes—that is, transitions in holds. Formula WFx, y(M1) asserts that there are in- which at least one of in[1], in[2], and out changes. Steps finitelymany M1 x,y steps or M1 x,y is infinitely that change other variables—for example, variables repre- often not enabled. Since (i) and (ii) implythat x is senting circuit elements inside the environment—but leave always a natural number, M1 x,y is always enabled. in[1], in[2], out unchanged are also possible. Hence, WFx, y(M1) implies that there are infinitelymany The predicate-action diagram of Figure 3(a) looks like a M1 x,y steps, so x is incremented infinitelyoften. Simi- standard state-transition diagram. However, we interpret larly, WFx, y(M2) implies that y is incremented infinitely it formallynot as a conventional state machine, but as the 3

(a) A predicate-action diagram. ✎ ✎ C(1, 0, 0) C(0, 1, 1) ✍ ✌ ✍ ✌ ✁✕ ❆ ✁✕ ❆ ❅❘✎ ✁ ❆✎ ❅❘✎ ✁ ❆✎ ✞✲ C(0, 0, 0) C(1, 1, 0) ✲ C(1, 1, 1) C(0, 0, 1) ¤ ✍ ✌ ✍ ✌ ✍ ✌ ✍ ✌ ❆ ✁✕ ❆ ✁✕ ❆✎ ✁ ❆✎ ✁ C(0, 1, 0) C(1, 0, 1) ✍ ✌ ✍ ✌ ✝ ✆

(b) The corresponding TLA formula. ∧ C(0, 0, 0) ∨ C(1, 1, 1)   ∧ ✷[C(0, 0, 0) ⇒ C(1, 0, 0) ∨ C(0, 1, 0) ]in[1],in[2],out  ∧ ✷[C(1, 0, 0) ⇒ C(1, 1, 0) ]in[1],in[2],out ...  ∧ ✷[C(0, 0, 1) ⇒ C(0, 0, 0) ]in[1],in[2],out Fig. 3. Predicate-Action diagram of in[1], in[2], out for a 2-input C-element, and the corresponding TLA formula.

1 ∆ TLA formula of Figure 3(b). This formula has the form Init C = ∧ out ∈{0, 1} ∧ i ∈{ ,..., } → ∧ o in =[ 1 n out] Init o F ,whereInit is a state predicate and there is ∆ Input(i) = ∧ in[i]=out one conjunct Fo for each node o. The predicate Init is  ∨ ∧ in =[in except ![i]=1− in[i]] C(0, 0, 0) C(1, 1, 1). Each Fo describes the possible state ∧ out  = out changes starting from a state described bynode o.For ∆ Output = ∧∀i ∈{1,...,n} : in[i] = out example, the formula Fo for the node labeled C(1, 1, 0) is ∧ out  =1− out  ∧ in = in ✷[C(1, 1, 0) ⇒ C(1, 1, 1) ]in[1],in[2],out ∆ Next = Output ∨∃i ∈{1,...,n} : Input(i) A predicate-action diagram represents a safetyproperty;it ∆ ΠC = Init C ∧ ✷[Next]in,out ∧ WFin,out(Output) does not include anyfairness conditions. Figure 3(a) is a reasonable wayto describe a 2-input Fig. 4. A TLA specification of an n-input C-element. C-element. However, the corresponding diagram for a 3- input C-element would be quite complicated; and there is Next An action that is the disjunction of Output and all no wayto draw such a diagram for an n-input circuit. The the Input(i)actions,fori ∈{1,...,n}.Thus,aNext general specification is written directlyas a TLA formula step is either an Output step or an Input(i)stepfor in Figure 4. The arrayof inputs is represented formally some input line i. bya variable in whose value is a function with domain ΠC A temporal formula that is the specification of the {1,...,n}, where square brackets denote function applica- C-element (together with its environment). It asserts tion. (Formally, n is a rigid variable—one whose value is that (i) Init C holds initially, (ii) every step is either constant throughout a behavior.) We introduce two pieces a Next step or else leaves in, out unchanged, and of notation for representing functions: (iii) Output cannot be enabled forever without an Out- • [i ∈ S → e(i)] denotes the function f with domain S such that f[i]equalse(i) for every i in S. put step occurring. The fairness condition (iii) requires the output to change if all the inputs have; inputs • [f except ![i]=e] denotes the function g that is the same as f except that g[i]equalse. are not required to change. (Since predicate-action The formulas defined in Figure 4 have the following inter- diagrams describe onlysafetyproperties, the fairness pretation. condition is irrelevant to our explanation of the dia- Init C A state predicate asserting that out is either 0 or grams.) 1, and that in is the function with domain {1,...,n} The specification ΠC is short and precise. However, it is such that in[i]equalsout for all i in its domain. not as reader-friendlyas a predicate-action diagram. We Input(i) An action that is enabled iff in[i]equalsout. therefore use diagrams to help explain the specification, It complements in[i], leaves in[j] unchanged for j = i, beginning with the predicate-action diagram of Figure 5. and leaves out unchanged. (The symbol i is a param- It is a diagram of the state function in[i], out , meaning eter.) that it describes transitions that change in[i], out .Itis Output An action that is enabled iff all the in[i]aredif- a diagram for the formula ΠC , meaning that it represents ferent from out. It complements out and leaves in a formula that is implied byΠ C . The diagram shows the unchanged. synchronization between the C-element’s ith input and its 1 output. A list of formulas bulleted by ∧ or ∨ denotes their conjunction or disjunction; ∧ and ∨ are also used as ordinary infix operators. We can draw manydifferent predicate-action diagrams 4 ✛ ✘  ,  ∧ in[i]=1 (a) A predicate-action diagram of in out . ∧ out =0 ∃ j = i : Input(j) ∃ j = i : Input(j) ✚ ✙ ✡✣❏ ✛ ✘ ✛ ✘ Input(i) ✡ ❏ ✗ ✔ ❥✗ ✔ ✲ ∧ in[i]=0 ∧ in[i]=1 ✛ ✲ in[i]=out in[i] = out ∧ out =0 ∧ out =1 ✚ ✙ ✚ ✙ ✖ ✕ ✖ ✕ ❏ ✡ ❏✛ ✡✢✘ Output ∧ in[i]=0 (b) The corresponding TLA formula. ∧ out =1 ∧ i ✚ ✙ in[ ]=out  (in[i]=out) ⇒   i ,    Fig. 5. A predicate-action diagram of in[ ] out for the specification ∧ ✷ ∨ Input(i) ∧ (in [i] = out ) ΠC of an n-input C-element, where 1 ≤ i ≤ n. ∨ ∃ j  i j ∧  i   ( = : Input( )) (in [ ]=out ) in,out (in[i] = out) ⇒  ✗ ✔ ✗ ✔ ∧ ✷ ∨ Output ∧ (in[i]=out ) ❥ ∨ ∃ j  i j ∧  i   ( = : Input( )) (in [ ] = out ) in,out ✲ in[i]=out in[i] = out ✖ ✕ ✖ ✕ Fig. 8. A predicate-action diagram of in, out for ΠC ,andthe corresponding TLA formula, where 1 ≤ i ≤ n.

Fig. 6. Another predicate-action diagram of in[i], out for ΠC ,where 1 ≤ i ≤ n.

in[i]=1− in[i] ✗ ✔ ❥✗ ✔ for the same specification. Figure 6 shows another diagram ✲ in[i]=out in[i] = out of in[i], out for ΠC . It is simpler than the one in Figure 5, ✖ ✕ ✖ ✕ but it contains less information. It does not indicate that out  =1− out the values of in[i]andout are always 0 or 1, and it does not show which variable is changed byeach transition. The Fig. 9. Yet another predicate-action diagram of in, out for ΠC . latter information is added in the diagram of Figure 7(a), where each transition is labeled with an action. The label Input(i) on the left-to-right arrow indicates that a transi- draw multiple diagrams to illustrate different aspects of a tion from a state satisfying in[i]=out to a state satisfying system. Actual specifications are written as TLA formulas, in[i] = out is an Input(i) step. This diagram represents which are much more expressive than pictures. the TLA formula of Figure 7(b). B. A Formal Treatment Even more information is conveyed by a predicate-action diagram of in, out , which also shows transitions that leave B.1 Definition in[i]andout unchanged but change in[j]forsomej = i. We first define preciselythe TLA formula represented by Such a diagram is drawn in Figure 8(a). Figure 8(b) gives a diagram. Formally, a predicate-action diagram consists the corresponding TLA formula. of a directed graph, with a subset of the nodes identified as There are innumerable predicate-action diagrams that initial nodes, where each node is labeled bya state pred- can be drawn for a specification. Figure 9 shows yet an- icate and each edge is labeled byan action. We assume other diagram for the C-element specification ΠC .Since a given diagram of a state function v and introduce the we are not relying on these diagrams as our specification, following notation. but simplyto help explain the specification, we can show as much or as little information in them as we wish. We can N The set of nodes. I The set of initial nodes. E(n) The set of edges originating at node n. (a) A predicate-action diagram of in[i], out. d(e) The destination node of edge e. Input(i) Pn The predicate labeling node n. ✗ ✔ ✗ ✔ ❥ Ee The action labeling edge e. ✲ in[i]=out in[i] = out ✖ ✕ ✖ ✕ The formula ∆ represented bythe diagram is defined as Output follows. ∆ ∃ ∈ (b) The corresponding TLA formula. Init ∆ = n I : Pn ∆ ∧ in[i]=out An = ∃ e ∈ E(n):Ee ∧ Pd(e) ∧ ✷ i ⇒ i ∧  i   ∆ [(in[ ]=out) Input( ) (in [ ] = out )]in[i],out ∆ ∧∀ ∈ ✷ n ⇒An v   ∆ = Init n N : [P ] ∧ ✷[(in[i] = out) ⇒ Output ∧ (in [i]=out )]in[i],out Fig. 7. A more informative predicate-action diagram of in[i], out When no explicit label is attached to an edge e,wetake for ΠC , and the corresponding TLA formula. Ee to be true. When no set of initial nodes is explicitly 5 indicated, we take I to be N. With the usual convention The first condition is an assertion about predicates; it is for quantification over an emptyset, An is defined to equal generallyeasyto prove. To prove the second condition, false if there are no edges originating at node n. one usuallyfinds an invariant Inv such that InitΠ ∧ ✷[M]u implies ✷Inv, so Π implies ✷[M∧Inv]u. The second con- B.2 Another Interpretation dition is then proved byshowing that [ M∧Inv]u implies Another possible interpretation of the predicate-action [Pn ⇒An]v, for each node n. Usually, u and v are tu- diagram is the formula ∆, defined by ples and everycomponent of v is a component of u,so u = u implies v = v. In this case, one need show only  ∆ ∆ = Init∆ ∧ ✷[∃ n ∈ N : Pn ∧An]v that M∧Inv implies [Pn ⇒An]v,foreachn. Bydefinition of An, this means proving This is perhaps a more obvious interpretation—especially if the diagram is viewed as a description of a next-state Pn ∧M∧Inv ⇒ (∃ m ∈ E(n):Em ∧ Pd(m)) ∨ (v = v) relation. We now show that ∆ always implies ∆, and that for each node n. This formula asserts that an M step the converse implication holds if the predicates labeling the E nodes are disjoint. that starts with Pn and Inv true and changes v is an m step that ends in a state satisfying Pd(m), for some edge m  (A) ∆ implies ∆. originating at node n. Proof : A simple invariance proof, using rule INV1 of [3, IV. Illustrating Proofs Figure 5, page 888], shows that ∆ implies ✷(∃ n ∈ N : Pn). We then have: In TLA, there is no distinction between a specification and a property; they are both formulas. Verification means ∆ ∆ = Init ∆ ∧∀n ∈ N : ✷[Pn ⇒An]v proving that one formula implies another. A practical, rel- ≡ Init ∆ ∧ ✷([∃ n ∈ N : Pn]v) ativelycomplete set of rules for proving such implications ∧∀n ∈ N : ✷[Pn ⇒An]v is described in [3]. We show here how predicate-action dia- [ because ∆ implies ✷(∃ n ∈ N : Pn)] grams can be used to illustrate these proofs. We take as our example the same one treated in [3], that the specification ≡ Init ∆ ∧ ✷[(∃ n ∈ N : Pn) Ψ defined in Section IV-A below implies the specification ∧∀n ∈ N :(Pn ⇒An)]v [ because ✷ distributes over conjunction and ∀,and Φ defined in Section II above. [X]v ∧∀n ∈ N :[Yn]v is equivalent to [X ∧∀n ∈ N : Yn]v ] A. Another Specification ⇒ Init ∆ ∧ ✷[∃ n ∈ N : Pn ∧An]v We define a TLA formula Ψ describing a program with [ by predicate logic, since B⇒Cimplies ✷[B]v ⇒ ✷[C]v ] two processes, each of which repeatedlyloops through the ∆ = ∆ sequence of operations P (sem); increment; V (sem), where one process increments x byone and the other increments (B) If ¬(Pm ∧Pn)holdsforallm, n in N with m = n,then y byone. Here, P (sem)andV (sem) denote the usual op- ∆ implies ∆. erations on a sem. To describe this program formally, we introduce a variable pc that indicates the con- Proof : Bypropositional logic, the hypothesis implies trol state. Each process has three control points, which we call “a”, “b”,and“g”. (Quotes indicate string values.) (∃ n ∈ N : Pn ∧An) ⇒ (∀ n ∈ N : Pn ⇒An) We motivate the definition of Ψ with the three predicate- The result then follows from simple temporal reasoning, action diagrams for Ψ in Figure 10. In these diagrams, the essentiallybythe reverse of the string of equivalences and predicate PC (p, q) asserts that control is at p in process 1 implication used to prove (A). and at q in process 2. Figure 10(a) shows how the control state changes when the P (sem), V (sem), and increment We usuallylabel the nodes of a predicate-action diagram actions are performed. Variables other than pc not men- with disjoint predicates, in which case (A) and (B) imply  tioned in an edge label are left unchanged bythe indicated that the interpretations ∆ and ∆ are equivalent. Diagrams steps—for example, steps described bythe edge labeled with nondisjoint node labels mayoccasionallybe useful; ∆ x = x +1 leave y and sem unchanged—but this is not is the more convenient interpretation of such diagrams. asserted bythe diagram. The next-state action N is writ- ten as the disjunction N1 ∨N2 of the next-state actions C. Proving a Predicate-Action Diagram of each process; and each Ni is written as the disjunction Saying that a diagram is a predicate-action diagram for αi ∨ βi ∨ γi. Figure 10(b) illustrates this decomposition. a specification Π asserts that Π implies the formula ∆ rep- Finally, the predicate-action diagram of Figure 10(c) de- resented bythe diagram. Formula Π will usuallyhave the scribes how the semaphore variable sem changes. form InitΠ ∧ ✷[M]u ∧ L,whereL is a fairness condition. To write the specification Ψ, we let pc be a function with Formula ∆ equals Init ∆ ∧∀n ∈ N : ✷[Pn ⇒An]v.To domain {1, 2},withpc[i] indicating where control resides prove Π ⇒ ∆, we prove: in process i. The formula PC (p, q) can then be defined by 1. Init Π ⇒ Init ∆ ∆ 2. Init Π ∧ ✷[M]u ⇒ ✷[Pn ⇒An]v, for each node n. PC (p, q) =(pc[1] = p) ∧ (pc[2] = q) 6

V (a) ✞ (sem) ¤  ✏  ✏ x = x +1 PC (“b”, “a”) ✲ PC (“g”, “a”) ✆ P (sem) ✟✟✯✒ ✑ ✒ ✑ ✟✟ ❘ ✟✏ ✲ PC (“a”, “a”) ❍ ✒✒ ✑❍ ❍  ✏  ✏ P (sem) ❍❍❥ PC (“a”, “b”) ✲ PC (“a”, “g”) ¤ ✒ ✑y = y +1 ✒ ✑ ✝ ✆ V (sem)

(b) ✞ γ1 ¤  ✏  ✏ , β1 ✲ , ✆ ✟✯ PC (“b” “a”) PC (“g” “a”) α✟1 ✟✒ ✑ ✒ ✑ ❘ ✟✏✟ ✲ PC (“a”, “a”) ✒✒ ❍✑❍ ❍  ✏  ✏ α2 ❍❍❥ PC (“a”, “b”) ✲ PC (“a”, “g”) ¤ ✒ ✑β2 ✒ ✑ ✝ ✆ γ2

(c) β1 ∨ β2 α1 ∨ α2 ✗ ✔ ❥✗ ✔ ✲ sem =1 sem =0 ✖ ✕ ✖ ✕

γ1 ∨ γ2

Fig. 10. Three predicate-action diagrams of x, y, pc, sem for Ψ.

The semaphore actions P and V are defined by 3. Ψ ⇒ WFx,y(Mi), for i =1, 2 We illustrate the proofs of conditions 2 and 3 with the ∆ P (sem) = ∧ 0 < sem predicate-action diagram of x, y, sem, pc for Ψ in Fig- ∧ sem = sem − 1 ure 12, where Q is defined by ∆ V (sem) = sem = sem +1 ∆ Q i(p, q) = ∧ PC (p, q) Missing from Figure 10 are a specification of the initial val- ∧ sem = i ues of x and y, which we take to be zero, and a fairness con- ∧ (x ∈ Nat) ∧ (y ∈ Nat) dition. One could augment predicate-action diagrams with some notation for indicating fairness conditions. However, and Nat is the set of natural numbers. the conditions that are easyto represent with a diagram First, we must show that the diagram in Figure 12 are not expressive enough to describe the varietyof fairness is a predicate-action diagram for Ψ. This is easyusing requirements that arise in practice. The WF and SF formu- the definition in Section III-B.1; no invariant is needed. las, which are expressive enough, are not easyto represent For example, the condition to be proved for the node la- graphically. So, we have not attempted to represent fair- beled Q 0(“b”, “a”)isthatanN step that starts with ness in our diagrams. We take as the fairness condition for Q0(“b”, “a”)trueisanM1 step (one that increments x and our specification Ψ strong fairness on the next-state action leaves y unchanged) that makes Q 0(“g”, “a”)true.This Ni of each process. The complete definition of Ψ appears follows easilyfrom the definitions of Q and N ,sinceanN in Figure 11. step starting with PC (“b”, “a”)truemustbeaβ1 step. To prove condition 2, it suffices to prove that everystep B. An Illustrated Proof allowed bythe diagram of Figure 12 is a [ M]x,y step. The The proof of Ψ ⇒ Φ is broken into the proof of three steps not shown explicitlybythe diagram are ones that conditions: leave w unchanged. Such steps leave x, y unchanged, so 1. Init Ψ ⇒ Init Φ theyare [ M]x,y steps. The actions labeling all the edges 2. Init Ψ ∧ ✷[N ]w ⇒ ✷[M]x,y of the diagram imply[ M]x,y, so all the steps shown ex- 7

∆ ∆ Init Ψ = ∧ pc =[i ∈{1, 2} →“a”] β1 = ∧ pc[1] = “b” ∧ (x =0)∧ (y =0) ∧ pc  =[pc except ![1] = “g”] ∧ sem =1 ∧ x = x +1  ∆ ∧y, sem = y, sem αi = ∧ (pc [i]=“a”) ∧ (0 < sem)  ∆ ∧ pc =[pc except ![i]=“b”] β2 = ∧ pc[2] = “b” ∧ sem = sem − 1 ∧ pc  =[pc except ![2] = “g”] ∧x, y = x, y ∧ y = y +1  ∆ ∧x, sem = x, sem γi = ∧ pc[i]=“g”  ∆ ∧ pc =[pc except ![i]=“a”] Ni = αi ∨ βi ∨ γi  ∧ ∆ sem = sem +1 N N ∨N ∧x, y x, y = 1 2 = ∆ w = x, y, sem, pc ∆ Ψ = Init Ψ ∧ ✷[N ]w ∧ SFw(N1) ∧ SFw(N2)

Fig. 11. The specification Ψ.

 ✞ x, y = x, y ¤  ✏  ✏ M1 ✲ ✆ Q0(“b”, “a”) Q0(“g”, “a”) x, y = x, y ✟✟✯✒ ✑ ✒ ✑ ✟✟ ❘ ✟✏ ✲ Q (“a”, “a”) 1 ❍ ✒✒ ✑❍ ❍  ✏  ✏  ❍ x, y = x, y ❍❥ ✲ ¤ Q0(“a”, “b”) Q0(“a”, “g”) ✒ ✑M2 ✒ ✑ ✝ ✆ x, y = x, y

Fig. 12. Another predicate-action diagram of x, y, sem, pc  for Ψ.

plicitlybythe diagram are also [ M]x,y steps. This proves to visualize the proof, which can help us to understand it. condition 2. V. Conclusion We now sketch the proof of condition 3. To prove WFx,y(Mi), it suffices to show that infinitelymany We have described three uses of diagrams that we believe Mi x,y steps occur. We first observe that each of the are new for diagrams with a precise formal semantics: predicates labeling a node in the diagram implies that ei- • To describe particular aspects of a complex specifica- ther N1 w or N2 w is enabled. The fairness condition tion with a simple diagram. An n-input C-element of Ψ then implies that a behavior cannot remain forever cannot be specified with a simple picture. However, at anynode, but must keep moving through the diagram. we explained the specification with diagrams describ- Hence, the behavior must infinitelyoften pass through ing the synchronization between the output and each the Q 1(“a”, “a”) node. The predicate Q 1(“a”, “a”) implies individual input. that both N1 w and N2 w are enabled. Hence, the fair- • To provide complementaryviews of the same system. ness condition SFw(N1) ∧ SFw(N2) implies that infinitely Diagrams (b) and (c) of Figure 10 look quite different, many N1 w steps and infinitelymany N2 w steps must but theyare diagrams for the same specification. occur. Action N1 w is enabled onlyin the three nodes of • To illustrate proofs. The disjunction of the predicates the top loop. Taking infinitelymany N1 w steps is there- labeling the nodes in Figure 12 equals the invariant I fore possible onlybygoing around the top loop infinitely of the proof in Section 7.2 of [3]. The diagram provides manytimes, which implies that infinitelymany M1 steps a graphical representation of the invariance proof. occur, each starting in a state with Q 0(“b”, “a”) true. Since TLA differs from traditional specification methods in two Q 0(“b”, “a”) implies x ∈ Nat,anM1 step starting with important ways. First, all TLA specifications are inter- Q 0(“b”, “a”) true changes x,soitisanM1 x,y step. preted over the same set of states. Instead of assigning Hence, infinitelymany M1 x,y steps occur. Similarly, values just to the variables that appear in the specifica- taking infinitelymany N2 w steps implies that infinitely tion, a state assigns values to all of the infinite number many M2 x,y steps occur. This completes the proof of of variables that can appear in anyspecification. Second, condition 3. TLA specifications are invariant under stuttering. A for- Using the predicate-action diagram does not simplifythe mula can neither require nor rule out finite sequences of proof. If we were to make the argument given above rigor- steps that do not change anyvariables mentioned in the ous, we would go through preciselythe same steps as in the formula. (The state-function subscripts in TLA formulas proof described in [3]. However, the diagram does allow us are there to guarantee invariance under stuttering.) 8

These two differences lead to two major differences be- Leslie Lamport Leslie Lamport attended the tween traditional state-transition diagrams and predicate- Bronx High School of Science, where he took a course in mechanical drawing. He later re- action diagrams. In traditional diagrams, each node repre- ceived a Ph.D. in mathematics from Bran- sents a single state. Because states in TLA assign values to deis University, where he studied the propa- gation of singularities in the Cauchy problem an infinite number of variables, it is impossible to describe for analytic partial differential equations. Since a single state with a formula. Anyformula can specify 1985, he has been a member of Digital Equip- the values of onlya finite number of variables. To draw ment Corporation’s Systems Research Labora- tory, where he has written several biographical diagrams of TLA formulas, we let each node represent a sketches. predicate, which describes a set of states. In traditional di- agrams, everypossible state change is indicated byan edge. Because TLA formulas are invariant under stuttering, we draw diagrams of particular state functions—usuallytuples of variables.

TLA differs from most specification methods because it is a logic. It uses simple logical operations like implication and conjunction instead of more complicated automata- based notions of simulation and composition [6]. Every- thing we have done with predicate-action diagrams can be done with state-transition diagrams in anypurelystate- based formalism. However, conventional formalisms must use some notion of homomorphism between diagrams to describe what is expressed in TLA as logical implication.

Most formalisms employing state-transition diagrams are not purelystate-based, but use both states and events. Nodes represent states, and edges describe input and out- put events. The meaning of a diagram is the sequence of events it allows; the states are effectivelyhidden. In TLA, there are onlystates, not events. Systemsare described in terms of changes to interface variables rather than in terms of interface events. Variables describing the internal state are hidden with the existential quantifier ∃ describedin[3]. Changes to anyvariable, whether internal or interface, can be indicated bynode labels or edge labels. Hence, a purely state-based approach like TLA allows more flexibilityin how diagrams are drawn than a method based on states and events.

References

[1] G. H. Mealy, “A method for synthesizing sequential circuits,” Bell System Technical Journal, vol. 34, pp. 1045–1079, Sept. 1955. [2] E. F. Moore, “Gedanken-experiments on sequential machines,” in Automata Studies (C. E. Shannon and J. McCarthy, eds.), pp. 129–153, Princeton, New Jersey: Princeton University Press, 1956. [3] L. Lamport, “The temporal logic of actions,” ACM Trans. Pro- gramming Languages and Systems, vol. 16, pp. 872–923, May 1994. [4] B. Alpern and F. B. Schneider, “Defining liveness,” Information Processing Letters, vol. 21, pp. 181–185, Oct. 1985. [5] C. Mead and L. Conway, Introduction to VLSI Systems,ch.7. Reading, Massachusetts: Addison-Wesley, 1980. [6] M. Abadi and L. Lamport, “Conjoining specifications,” Research Report 118, Digital Equipment Corporation, Systems Research Center, 1993. To appear in ACM Transactions on Programming Languages and Systems.