Talking Bananas Structural Recursion for Session Types

Sam Lindley J. Garrett Morris The University of Edinburgh, UK {Sam.Lindley,Garrett.Morris}@ed.ac.uk

Abstract commands, but a client that sends the recipient’s address before the Session types provide static guarantees that concurrent programs sender’s address is in violation of the protocol despite having sent respect communication protocols. We give a novel account of re- well-formed SMTP commands. cursive session types in the context of GV, a small concurrent ex- Session types, originally proposed by Honda [22], are an ap- tension of the linear λ-calculus. We extend GV with recursive types proach to statically verifying communicating concurrent programs. and , following the semantics of re- A session type specifies the expected communication along a chan- cursion, and show that doing so naturally gives rise to recursive nel. For example, consider a simplification of the client’s view of session types. We show that this principled approach to recursion the SMTP protocol. After being authenticated, a client has the op- resolves long-standing problems in the treatment of duality for re- tion of sending one or more messages, each consisting of a sender’s cursive session types. address, recipient’s address, and message, in that order. We could We characterize the expressiveness of GV concurrency by giv- express this with the following session type: ing a CPS translation to (non-concurrent) λ-calculus and proving Client def= !FromAddress.!ToAddress.!Message.Client ⊕! that reduction in GV is simulated by full reduction in λ-calculus. !Quit.end This shows that GV remains terminating in the presence of positive recursive types, and that such arguments extend to other extensions where type constants FromAddress, ToAddress, Message, and Quit of GV, such as polymorphism or non-linear types, by appeal to nor- denote the corresponding SMTP commands. This definition makes malization results for sequential λ-calculi. We also show that GV use of several session type constructors. The type !T.S denotes remains deadlock free and deterministic in the presence of recur- sending a value of type T before continuing with behavior S, S⊕! S0 sive types. denotes communicating a choice between behaviors S and S0, end Finally, we extend CP, a session-typed process calculus based denotes the end of a session, and finally we make use of recursive on linear logic, with recursive types, and show that doing so pre- definition to specify repetition in the protocol. A key aspect of serves the connection between reduction in GV and cut elimination session typing is duality. The session type of an SMTP server is in CP. dual to that of the client: Categories and Subject Descriptors D.3.2 [Language Classifi- Server def= ?FromAddress.?ToAddress.?Message.Server ⊕? cations]: Concurrent, distributed, and parallel languages; D.3.3 ?Quit.end [Language Constructs and Features]: Recursion This type definition uses dual features to those in the client’s type: Keywords Session types, recursion ?T.S denotes receiving a value of type T before continuing as S and S ⊕? S0 denotes receiving a choice between S and S. 1. Introduction We present a novel account of recursive and corecursive session types. Following initial algebra semantics, we characterize recur- Concurrency and communication have become central problems sive computation by catamorphisms (folds) rather than by an ar- in modern software design and engineering, from hand-held ap- bitrary fixed-point operator. Similarly, we characterize corecursive plications relying on remote services to provide key functionality, computation by anamorphisms (unfolds). This formulation differs through traditional applications now running on multi-core hard- from traditional presentations of recursive session types in three ware, to distributed applications running across data centers. As- ways. First, we identify dual notions of recursion, corresponding suring correct behavior for concurrent programs requires reasoning to producers and consumers, rather than having a single self-dual not just about the type of data communicated, but about the order in notion of recursion in session types. Second, as they are based on which communication takes place. For instance, the messages be- well-founded recursive data types, our recursive session types guar- tween an SMTP server and client are all strings representing SMTP antee termination and freedom from deadlock and livelock. Third, following algebraic ideas of recursion and duality leads to a sound syntactic characterization of duality for recursive session types. Many previous syntactic formulations of session type duality ei- ther incorrectly identify non-dual processes as dual [6,7] or rely on corecursive expansion of session types [8]. We present our formulation of recursive and corecursive ses- sion types as an extension, called µGV [26], to a core concur- rent λ-calculus called GV. We extend GV with (co)recursive types, (co)recursive session types, and (un)folds over both (co)recursive types and (co)recursive session types, and show that these are suf- ficient to write non-trivial programs using (co)recursive session lish a strong connection between µGV and linear logic, we present types. Previous work on GV has minimized the core calculus by an extension of CP, called µCP, which includes least and great- encoding session-typed features in terms of functional features and est fixed points and corresponding recursive and corecursive proof simple input and output primitives. We continue this thread, show- terms, and show semantics-preserving translations from µGV to ing that recursive session types can be encoded in terms of recursive µCP and vice versa (§6). We conclude by discussing related (§7) data types. This result simplifies the concurrent semantics of µGV; and future (§8) work. for example, it allows us to apply previous results on deadlock free- dom and determinism of GV to µGV unchanged. By identifying 2. Functional µGV least and greatest fixed points in the resulting calculus, we obtain a system that admits non-terminating communication, but still guar- We now describe Functional µGV, the functional fragment of antees deadlock freedom and productivity. µGV. We begin with Functional GV (§2.1), a linear λ-calculus We also seek to characterize the expressiveness of µGV’s con- without any form of recursion or iteration. We then discuss two currency. To do so, we give a CPS translation from GV into lin- extensions of GV. First, we extend Functional GV with recursive ear λ-calculus (without concurrency), and show that full reduction types and catamorphisms (§2.2). Second, we add corecursive types in the latter simulates reduction in the former. This also allows us and anamorphisms (§2.3), and discusses the relationship with non- to extend standard results on termination in sequential λ-calculi terminating recursion. to results for termination of extensions of GV. Most immediately, we can conclude that µGV is terminating. The approach applies 2.1 Functional GV equally well to other extensions of GV, such as with polymorphism Functional GV is based on the multiplicative-additive fragment or non-linear types. of intuitionistic linear logic. The syntax of Functional GV’s terms Recent work of Caires and Pfenning [13] and Wadler [34] and types is given at the top of Figure1. Types include linear develops a correspondence between reduction in process calculi implication (T ( U), binary and nullary multiplicative products and cut elimination in linear logic. GV is closely connected to (T ⊗ U and 1), additive sums (T ⊕ U and 0), and additive products Wadler’s logic-based process calculus CP: in prior work [26] we (T & U and >). We write M; N for let () = M in N. The syntax give translations between GV and CP such that reduction in each of terms includes (fully applied) constants KM, used to introduce simulates reduction in the other. We extend this observation to concurrency; we give the constants and their typing rules in §3. (co)recursive session types. We define an extension of CP to in- The typing rules for Functional GV are given in the center of clude (co)recursive types, following Baelde’s formulation of fixed Figure1; most are standard for linear λ-calculus. The variable points in classical linear logic [3], and then extend the semantics- rule insists on a singleton environment. Rules for the multiplicative preserving translation between CP and GV to include recursive combinators (T ( U and T ⊗U) split their hypotheses, while rules types. for the additive combinators (T & U and T ⊕ U) duplicate them. Recent work by Toninho et al [32] explores corecursive ses- In line with its interpretation as falsity, there is no introduction rule sion types from a propositions-as-types perspective. Despite having for 0, nor elimination rule for its dual >. similar aims, our approach differs from theirs in three significant A small-step operational semantics for Functional GV is given ways. First, we identify parallels between concurrent and sequen- at the bottom of Figure1. To maintain a close connection with tial abstractions, in this case between recursive and corecursive data cut-elimination, we define term reduction using weak explicit sub- types and recursive session types. Toninho et al., in contrast, de- stitutions [25]. In this approach, we capture substitutions at λ- velop corecursive session types directly from the corresponding abstractions instead of applying them directly to the body of the proof theory. This simplifies our concurrent semantics, as we do abstraction. Our values thus include closures λσx.M, which pair a not have to account for recursive communication directly. Second, function abstraction λx.M with a captured environment σ. We ex- we identify two forms of recursive session types—corresponding to tend the typing judgment to include closures by encodings based on recursive and corecursive data types—and that Γ, x : T ` Mσ : U dom(σ) = fv(M) \{x} their composition can provide unbounded computation. Toninho et σ al. identify one of these forms, but not the other. Third, as illustrated Γ ` λ x.M : T ( U by the equivalence with µCP, our session types are fundamentally where we write fv(M) to denote the free variables of M. The classical, while Toninho et al. build on intuitionistic proof theory. free variables of a closure are the free variables of the range of Thus, for example, our results on the duality of recursive session σ; capture-avoiding substitution Mσ is defined on the free vari- types do not arise from their approach. We see the coincidence of ables of M. We implicitly treat plain abstractions λx.M as closures our typing rules with theirs, despite the significant differences in σ 0 λ x.(Mσ ) where σ maps each free variable xi of M to a fresh vari- methodology and foundations, as reinforcing the relevance of both able x0, and σ0 is the inverse of σ. lines of inquiry. i The paper proceeds as follows. We begin with our session- 2.2 Recursion typed functional language, µGV. We present µGV’s functional fragment (§2), extending a core linear λ-calculus with both recur- Next, we extend Functional GV with recursive types and their sive and corecursive types, and discuss the connection between terms. Our treatment of recursive types is based on the initial al- µGV and languages with non-termination. We present the con- gebra semantics of recursion [21]. We extend our language with current fragment of the language (§3), including both recursive positive type F and their least fixed points µF. We intro- and corecursive session types, and compare our approach to tra- duce two new term forms: the operator in captures that µF is itself the carrier of an F-algebra, while the fold operator − captures ditional characterizations of duality for session types. We give en- L M codings for the concurrent features of µGV in terms of simpler con- that µF is initial. This presentation of recursive types has a long structs and give a concurrent small-step operational semantics for history in the community, dating at least µGV (§4). We characterize the expressivity of GV concurrency by from the treatment of lists in Squiggol [9], and generalized by Mei- a CPS translation to a non-concurrent λ-calculus (§5). In doing so, jer et al. [29]. we show that µGV is terminating, and thus (in combination with The syntax and typing for recursive types are given at the top existing work on GV) free from livelock and deadlock. To estab- of Figure2. We extend the language of types with type variables X and least fixed points µF, and terms as discussed before. We omit Syntax the (entirely standard) definition of positivity and well-formedness conditions for types and type operators. As Functional GV is a lin- Types T, U ::= T ( U | T ⊗ U | 1 | T ⊕ U | 0 ear calculus, the typing rule for M mandates an empty environ- | T & U | > ment: the evaluation of M mayL requireM arbitrarily many copies Terms L, M, N ::= x | KM | λx.M | MN of the expression M, andL duplicatingM M would require duplicat- | (M, N) | let (x, y) = M in N ing its assumptions. Functional GV contains only linear assump- | () | let () = M in N tions; adding the exponential modality would introduce non-linear | inl M | inr M assumptions, which could be used in the bodies of catamorphisms. | case L {inl x 7→ M; inr x 7→ N} The semantics of folds is given at the bottom of Figure2. The | absurd M extension of values and contexts is unsurprising. Each reduction of | hM, Ni | fst M | snd M | hi a fold M amounts to one unrolling of the body M, and depends Typing on theL actionM of F on terms. With the exception of identity, each type constructor gives rise to both covariant (F) and contravariant K : T U Γ ` M : T − ( functors (F ). We have the expected typings that if M : T ( U − x : T ` x : T Γ ` KM : U then F(M): F(T) ( F(U) and F (M): F(U) ( F(T). We use point-free notation for building functors over binary type 0 Γ, x : T ` M : U Γ ` M : T ( U Γ ` N : T constructors: for binary type constructor ∗, we write F ∗ G as 0 shorthand for X.F(X) ∗ G(X). Γ ` λx.M : T ( U Γ, Γ ` MN : U Example. We turn to natural numbers as a characteristic example 0 Γ ` M : T ⊗ T of recursive types. The definition of the type of naturals parallels Γ ` M : T Γ0 ` N : U Γ0, x : T, y : T0 ` N : U the standard (intuitionistic) definition: 0 0 Γ, Γ ` (M, N): T ⊗ U Γ, Γ ` let (x, y) = M in N : U N(X) = 1 ⊕ X Nat = µN

0 and we can give familiar definitions of the constructors: Γ ` M : 1 Γ ` N : T ` () : 1 Γ, Γ0 ` let () = M in N : T zero = in (inl ()) succ = λz.in (inr z) Now consider a standard recursive definition of addition: Γ ` M : T Γ ` M : U 0 + y = y (S x) + y = S (x + y) Γ ` inl M : T ⊕ U Γ ` inr M : T ⊕ U We can rewrite this definition as a curried function of one argument Γ ` M : T ⊕ T0 Γ0, x : T ` N : U Γ0, x : T0 ` N0 : U plus 0 = id plus (S x) = S ◦ plus x 0 0 Γ, Γ ` case M {inl x 7→ N; inr x 7→ N } : U where id is the identity function and ◦ is function composition. Finally, this version can be expressed using a fold: Γ ` M : 0 Γ ` M : T Γ ` N : U 0 plus = λx.case x {inl () 7→ id; inr f 7→ succ ◦ f } Γ, Γ ` absurd M : T Γ ` hM, Ni : T & U L M The body of the fold has type N(Nat ( Nat) ( (Nat ( Nat), Γ ` M : T & U Γ ` M : T & U so plus has type Nat ( Nat ( Nat. The product of x and y, unlike their sum, cannot be computed without duplicating (in some Γ ` fst M : T Γ ` snd M : U Γ ` hi : > way) either x or y. In an intuitionistic setting, we might accomplish Values and Contexts this by capturing x in the body of the fold, and using it in each V, W ::= x | λσx.M | (V, W) | () | inl V | inr V | hV, Wi | hi iteration. We cannot do the same in the linear setting. Instead, we will begin by demonstrating terms that duplicate and discard σ ::= {V1/x1,..., Vn/xn} naturals. That is, we show that contraction and weakening are where the xi are pairwise distinct E ::= [ ] | KE | EM | VE derivable for proposition Nat. This is an instance of a general result, | (E, M) | (V, E) | let (x, y) = E in M due to Filinski [18], that contraction and weakening are derivable | let () = E in M for the positive combinators in intuitionistic linear logic. In our | inl E | inr E setting this result is extended to include least fixed points µF. | case E {inl x 7→ N; inr x 7→ N0} dup = λx.case x {inl () 7→ (zero, zero) | absurd E L inr (y, z) 7→ (succ y, succ z)} | hE, Mi | hV, Ei | fst E | snd E drop = λx.case x {inl () 7→ () M L Reduction inr () 7→ ()} M σ We have that dup : Nat Nat × Nat, where the output naturals (λ x.M) V −→V M{V/x ] σ} ( let (x, y) = (V, W) in M −→ M{V/x, W/y} are equal to the input natural, and drop : Nat ( 1, where we V can then trivially eliminate the unit value. We can now implement let () = () in M −→ M V multiplication:  inl x 7→ N; case (inl V) 0 −→V N{V/x} body = λx.λy.case y {inl () 7→ (x, 0) inr y 7→ N 0 fst hV, Wi −→ V inr (x, y) 7→ let (x, x ) = dup x in} V (x0, plus x y); snd hV, Wi −→ W V times = λx.λy.let (x0, z) = body x y in drop x0; z 0 0 L M E[M] −→V E[M ] if M −→V M The body of the fold has type N(Nat × Nat) ( Nat × Nat; in the inductive case, the input pair will be (x, x(y − 1)), and the result is Figure 1: Functional GV Syntax and Semantics. then (x, xy). We duplicate x at each step: one copy is added to the Syntax Syntax Types T, U ::= · · · | X | µF Types T ::= · · · | νF Operators F, G ::= X.T Terms M ::= · · · | out M | M Terms L, M, N ::= · · · | in M | M Typing L M Typing Γ ` M : νF ` M : A ( F(A) Γ ` M : F(µF) ` M : F(T) ( T Γ ` out M : F(νF) ` M : A ( νF Γ ` in M : µF ` M : µF ( T Values and Contexts L M Values and Contexts V, W ::= M | M V | ... V, W ::= · · · | M E ::= out E | ... E ::= · · · | Lin EM Covariant Covariant Functors (X.νF)(M) = λx.(X.F(νF{T/X}))(M)(out x) (X.T)(M) = λx.x X not free in T if M : T ( U (X.X)(M) = M Contravariant Functor (F ⊗ G)(M) = λx.let (y, z) = x in (F(M) y, G(M) z) (X.νF)−(M) = λx.(X.F(νF{T/X}))−(M)(out x) (F G)(M) = λf .G(M) ◦ f ◦ F−(M) ( if M : T ( U (F ⊕ G)(M) = λx.case x {inl x 7→ F(M) x; inr x 7→ G(M) x} (F & G)(M) = λx.hF(M)(fst x), G(M)(snd x)i Reduction (X.µF)(M) = λx.in ((X.F(µF{U/X}))(M) x) out ( M V) −→V F( M )(MV) if M : T ( F(T) for some T L M if M : T ( U Contravariant Functors Figure 3: Extending GV with (X.T)−(M) = λx.x X not free in T (F × G)−(M) = λx.let (y, z) = x in (F−(M) y, G−(M) z) − − Example. As a canonical example of corecursive types, we con- (F ( G) (M) = λf .G (M) ◦ f ◦ F(M) sider streams of naturals. (F ⊕ G)−(M) = λx.case x {inl x 7→ F−(M) x; inr x 7→ G−(M) x} − − − S(X) = 1 &(Nat ⊗ X) Stream = νS (F & G) (M) = λx.hF (M)(fst x), G (M)(snd x)i (X.µF)−(M) = λx.in ((X.F(µF{U/X}))−(M) x) The definition here differs from the typical intuitionistic definition: L M if M : T ( U as our streams are linear, we need to include some provision for terminating the stream. Nevertheless, the choice of the length of the Reduction sum lies with the consumer of the stream, precisely dual to the case M (in V) −→V M (F( M ) V) if M : F(T) ( T for some T with recursive types, in which the choice lies with the producer. We L M L M can demonstrate this by showing a term of type νS, enumerating all Figure 2: Extending Functional GV with Recursion. the natural numbers: upFrom = λx.hdrop x, let (y, z) = dup x in (y, succ z)i

Here the body of nats has type Nat ( S(Nat), and so upFrom has product, while the other copy appears in the result. The result of the type Nat ( Stream. A consumer of the stream can choose how body x y is the pair (x, xy); we call drop to discard the last copy many elements it reads; for example, here is a term that reads the ofL x, andM return xy. first two naturals from a given stream: firstTwo s = let (x, s) = snd (out s) in 2.3 Corecursion and Nontermination let (y, s) = snd (out s) in fst (out s); (x, y) This section describes a further extension of Functional GV with corecursive data types (νF), the greatest fixed point of the functor Nontermination. Freyd [19] observed that: a) the greatest and F. Our treatment is dual to that of recursive types: we introduce least fixed points of functors coincide in many denotational models a type operator for νF for the greatest fixed point of positive of functional languages, and b) recognizing this coincidence gives functor F, and terms out, witnessing that it is an F-coalgebra, an interpretation to many non-terminating recursive programs. One and the anamorphism M , witnessing its finality. The full syntax can apply this observation to Functional µGV by identifying the and semantics of corecursive types are given in Figure3. The types µF and νF. Doing so has several consequences for the term typing rules and interpretation of corecursive types are dual to language. Observe that out and in now compose (in either order), those for recursive types: where recursive types provide an iterated giving the identity. This gives rise to two new reduction rules to fold operation and a finite number of folding steps, corecursive account for these compositions: types provide an iterated unfold operation and a finite number of unfolding steps. We restrict the type environment in unfolding to out (in V) −→V V in (out V) −→V V avoid duplicating linear resources. We extend the syntax of values It is now possible to compose folds and unfolds to define recursive with unfolds M both unapplied (of type A ( νF) and applied computations. Such compositions are sometimes called hylomor- (of type νF). The treatment of greatest fixed points as functors phisms [29]. Hylomorphisms are accounted for by the following is unsurprising; the reduction rule for anamorphisms unrolls the reduction: term M relying on the action of F as expected. Functional µGV is Functional GV extended with recursive and corecursive data types. N ( M V) −→V N (F( N )(F( M )(MV))). L M L M Intuitively, each evaluation of a hylomorphism corresponds to one Syntax folding step and one unfolding step. Such an extension would break Session types S ::= !T.S | ?T.S | end! | end? the logical interpretation of Functional GV, as all types are inhab- ! 0 ! ? 0 ? ited by the trivial hylomorphism. However, it does show a direct | S ⊕ S | 0 | S ⊕ S | 0 ! ? ! ? and intuitive connection between the (admittedly austere) setting | X | X | µ F | µ F | ν F | ν F of Functional GV and more practical programming languages. Types T, U::= · · · | S Session functors F ::= X .S Terms L, M, N::= · · · | inl! M | inr! M 3. Concurrent µGV | case? L {inl x 7→ M; inl y 7→ N} 3.1 Communication and Concurrency | M ? | in!M | out?M | M ! We now consider the concurrent fragment of µGV. The additional Constants K ::= sendL M | receive | fork | wait | link syntax of the concurrent fragment is listed at the top of Figure4. Duality Our primitive session types include input (?T.S), output (!T.S), and ! ? closed channels (end?, end!). Unlike many session type systems, !T.S = ?T.S end! = end? µ F = µ F ? ! but in keeping with logically-founded approaches, we have dual ?T.S = !T.S end? = end! µ F = µ F 0 types for closed channels rather than a single, self-dual type end. S ⊕! S0 = S ⊕? S 0! = 0? ν!F = ν?F 0 (In our prior work [26] we discuss the semantic and logical con- S ⊕? S0 = S ⊕! S 0? = 0! ν?F = ν!F sequences of providing a self-dual closed channel.) The remaining X.S = X.S X .S = X .S{X /X} features of concurrent µGV can be encoded in terms of the primi- S = S tive concurrent features, and the features of Functional µGV. These Positivity include selection (S ⊕! S0), branching (S ⊕? S0) and recursive ses- ξ ∈ {X, X} sion types (µ!F, µ?F). In traditional session typing notation, ⊕! p(ξ.?T.S) = p(ξ.T) ∧ p(ξ.S), where p ∈ {neg, pos} is written as ⊕, ⊕? as &, inl! as select inl, and case? as offer. pos(ξ.!T.S) = neg(ξ.T) ∧ pos(ξ.S) To avoid conflicting with the base features of Functional µGV and neg(ξ.!T.S) = pos(ξ.T) ∧ neg(ξ.S) to emphasize the uniformity of our extensions, we adopt notation which makes explicit the direction of communication. For example, Typing the ! denotes that inl! sends a left injection along a channel. Note Γ ` M : S ⊕! S0 that fork is the only term that introduces new session-typed chan- send : T ⊗ !T.S ( S nels; the remaining session-typed constructs all consume channels. Γ ` inl! M : S receive :?T.S ( T ⊗ S Thus, for example, the typing of inl! is inverted from the typing of fork :(S ( end!) ( S Γ ` M : S ⊕! S0 inl, eliminating the session type S ⊕! S0. wait : end? ( 1 ! 0 Our treatment of recursive session types is also guided by initial Γ ` inr M : S link : S ⊗ S ( end! algebra semantics. We introduce session type variables X and ses- ? 0 0 sion functors F. The argument and result of a session functor are Γ ` L : S ⊕ S Γ, x : S ` M : T Γ, x : S ` N : T both session types. We extend the standard notion of positivity to Γ ` case? L {inl x 7→ M; inr x 7→ N} : T session functors and ordinary functors of session type (Figure4). ? ! Just as we distinguished between consuming ( − ) and producing Γ ` L : 0 Γ ` M : µ F ` M : F(S) ( S L M ? ! ! ? ? (in) values of recursive types, we distinguish between consuming Γ ` case L {} : T Γ ` in M : F(µ F) ` M : µ F ( S and producing recursive communication. Thus, we have two dual ? L M ? Γ ` M : ν F ` L : S ( F(S) ( end! constructors for recursive session types, µ F for consuming recur- ? ? ! ! sive communication and µ!F for producing it. The terms inhabiting Γ ` out M : F(ν F) ` L : S ( ν F ( end! recursive session types are similar to those for recursive data types: (Shaded terms and types, and their typing and duality, can be ? in!M unfolds one iteration of a recursive session type, while M encoded in terms of the remaining terms and types.) consumes a recursive session type; the production and consump-L M Figure 4: Concurrent µGV Terms and Typing tion roles are indicated by direction of communication. The typing of in!M reflects its role as an eliminator of session types, parallel to the typing of inl! and inr! . As for recursive session types, we have corecursive session dualized session functor F. In the definition of F, note that we not types ν!F and ν?F. The typing rule for out?M is a direct reflection only dualize the body of F, but also the variable; this accounts for of the rule for out M. As the communication primitives all consume the duality between the two forms of recursion. We contrast this ! approach with standard approaches to duality for recursive session terms of session type, − consumes a channel of type ν!F, and types in §3.2. returns the remaining (empty) expectations of the channel. The differences between recursive and corecursive session types are Promises. While aesthetically appealing, our formation of recur- apparent both in the term formation rules and in the direction ? sive session types seems somewhat awkward to use. The typing of of communication: folds M consume recursive communication, M ? requires that M transform F(S) (a session) into S (itself a ses- ! L M while unfolds M produce corecursive communication. Though Lsion);M that is, it flattens nested sessions into single sessions. In con- their typing rules appear similar, the directions of communication trast, most uses of recursive session types transform the data values in in!M and out?M are opposite. carried by the session into a result value, only incidentally relying The notion of duality is central to session types: if the process on the nesting of sessions. We can relate these views by observing holding one end of a channel expects to send a value of some type that µGV provides a natural notion of promise, and that promises along that channel, the process holding the other end should expect allow us to treat arbitrary types as session types. Promises [28] in- to receive a value of the same type. The dual S of session type troduce asynchrony between the computation of a value and its use; S is defined in Figure4. The dual of recursive ( µ!F, µ?F) and a promise of type T denotes a value of type T which may not yet corecursive (ν!F, ν?F) session types is defined in terms of the have been computed. A promise of type T arises naturally in our setting as a channel of type ?T.end?. We write ?T to abbreviate received along a channel. ?T.end , and define mappings between promises and values: ? sum : Nats Nat ? ( un :?T ( T sum = λc.case? c { ? un M = let (z, c) = receive M in L inl c 7→ wait c; en? zero wait c; z inr c 7→ let (x, c) = receive c in ? en : T ( ?T let y = un? c in ? en M = fork (λx.send (M, x)) en? (plus x y)} ? M The operation un?M retrieves a value from promise M (blocking We wrap the running sum in a promise to lift it to session type; ? until it is available), while en M constructs a new promise, already the body of the fold has type NC(?Nat) ( ?Nat, so sum has type ? ? ? ? ? containing the value of M. Note that un ◦ en and en ◦ un are µ NC (?Nat. We could compose this with one of the producers both observationally equivalent to the identity function (the second above to compute a value, such as: is a trivial example of channel forwarding). The dual of the type ?T ? is the type !T.end!, which we will abbreviate !T. We can also define un (sum (fork (downFrom 4))) an operation to introduce and eliminate channels of type !T (that is, where we write n to indicate the representation of natural n. This to provide results to unfulfilled promises): term will evaluate to 10. ! un : ((S ( T) ⊗ !T) ( S We can also define channel transformers. For example, we could un!(L, M) = fork (λx.send (L x, M)) compute the running total of the stream, inserting the total (to that ! en : S ( !S point) after each element. ! en (M) = fork (λz.let (c, z) = receive z in running : Nats Nats end wait z; link (M, c)) ( ( ! running c = un? ( λc.case? c {inl c 7→ done c; The appearance of the continuation type S may be surprising; this is L inr c 7→ more c} ? c) 0 a consequence of the different treatment of the continuation in send done c = en? (λz.λd.drop z; link (c, inl! (in!dM))) ? and receive. As µGV lacks polymorphism, we will write un M more c = let (y, c) = receive c in ? to denote the substitution of M into the definition of un , rather let k = un? c in than to denote an application in µGV, and similarly for the other let (y, y0) = dup y in definitions in this section. en? (λz.λd.let (w, w0) = dup (plus y0 z) in Example: Recursive Sessions. We now present several examples let d = send (y, inr! (in!d)) in of channels of naturals, building on our earlier representation of let d = send (w, inr! (in!d)) in naturals numbers. We introduce type abbreviations for such chan- k w0 d) nels: ? ? Note that the body of the fold has type NC(?(Nat ( Nats ( NC(X) = end? ⊕ ?Nat.X Nats = µ NC end!)) ( ?(Nat ( Nats ( end!); the Nat argument stores the We begin with a term that sends two naturals along a given channel: running sum, and so is initialized to zero by running.

twoNats : Nats ( end! Example: Corecursive Sessions. The treatment of corecursive ! ! twoNats = λc.let c = send (zero, inr (in c)) in session types is similar to that of recursive session types. Never- let c = send (succ zero, inr! (in! c)) in theless, we give a short example of their use. The previous example inl! (in! c) showed a term downFrom that sent all the naturals below a start- ing value along a channel. With corecursive session types, we can Now we present a slightly more interesting example. Given some give a term that (potentially) sends all the naturals above a starting starting natural n, we send the sequence n, n − 1,..., 0 along a value along a channel. We begin with the types; unlike the previous channel. We rely on Nat itself being defined recursively. version, it is now the side sending naturals that offers to choice to downFrom : Nats ( end! continue: 0 0 downFrom n = let (n , k) = body n in drop n ; k 0 ? 0 ! 0 L M NC (X) = end! ⊕ !Nat.X Nats = ν NC body : N(Nat ⊗ (Nats ( end!)) ( (Nat ⊗ (Nats ( end!)) body z = case z {inl () 7→ zc; We can now define a process that sends an increasing stream of inr (y, k) 7→ sc y k} naturals: zc : Nat ⊗ (Nats ( end!) body = λm.λc.case? c {inl c 7→ drop m; c; ! ! zc = (zero, λc.let c = send (zero, inr (in c)) in inr c 7→ let (m, m0) = dup m in inl! (in!c)) let c = send (m, c) in 0 sc : Nat ( (Nats ( end!) ( (Nat ⊗ (Nats ( end!)) send (succ m , c) 0 sc y k = let (y, y ) = dup y in upFrom = body ! (succ y, λc.k (send (y0, inr! (in!c)))) As before, we rely on duplication and discard being defined for We have a similar challenge in defining body as we did in defining 0 natural numbers. We have that body : Nat ( NC! (Nat) ( end!, times: at each step of the recursion, we must both send a value along 0 and so upFrom : Nat ( Nats ( end!. A simple consumer of the channel and produce the same value for the next step. Observe such a channel might compute the sum of the first two values on that body has type N(Nat ⊗ (Nats ( end!)) ( Nat ⊗ (Nats ( the channel: end!), computing both the next natural in the sequence and the function that sends it along a channel. firstTwo c = let (x, c) = receive (inr! (out?c)) in We can also write functions that consume channels of naturals. let (y, c) = receive (inr! (out?c)) in For a simple example, we could compute the sum of the naturals wait (inl! (out?c)); x + y 3.2 Recursion, Duality, and Session Types Session functors

We relate our statement of duality to previous accounts of recursive F? = X.Q F(?X) F! = X.Q F(!X) session types. Most existing approaches use equirecursive, self- J K J K dual recursive session types, and do not include dualized type vari- Session types ables. Our system lacks self-dual constructs, and so cannot encode Q S ⊕! S0 = !(Q S ⊕ Q S0 ) Q µ?F = ?µF self-dual recursive types. However, we can imagine extending our ? J ? 0K J K J 0K QJ µ!FK = !µF S S S Q S ⊕ S = ?(Q S ⊕ Q S ) ? language with a construct µ F such that µ F = µ F. ? ! J Q 0 K = ?0 J K J K QJν FK = !νF! Honda et al. [23] originally proposed recursive session types J !K J ? K first order Q 0 = !0 Q ν F = ?νF ! for a system of session types in which messages did not J K J K include channel names. Duality was given by µSX.T = µSX.T, Terms where X = X. To distinguish this notion from ours, we write Q `! M = un!(λx.` x, Q M ) ? naive(T) instead of T. In contrast, our approach gives µ X.T = inl x 7→JM;K J K inl x 7→ Q M  ? Qscase? L {= case(un?Q L ) µ X.T{X/X}. It is not hard to see that logical duality coincides inr x 7→ N inr x 7→ QJNK S J K with naive duality for first-order session types. Intuitively, if µ X.T ! ! J K S Q in M = un (λx.in x, Q M ) is first-order, then if we compute µ X.T{X/X} each instance of X ? Q MJ NK = λy.en? (Q MJ(QK F (un?) y)) in T will first be dualised by T and then again by the substitution JL M K L ? J K J K M {X/X}. (un Q N ) ? ? J K Independently, Bono and Padovani [10] and Bernardi and Hen- Q out M = out (un Q M ) ! nessy [6] observe that naive duality is not enough for higher- Q LJ MNK = λy.Q F (Jen!K) order session types, that is, session types with support for del- J K (unJ ! (Kλc.fork (λd.Q L d c), egation. Consider S = µSX.?X.X. The logical dual of S is y)) J K µSX.!(µSX.?X.X).X, whereas the naive dual of S is µSX.!X.X, Q M (en! Q N ) which is (equirecursively) equivalent to µSX.!(µSX.!X.X).X. It is J K J K not difficult to show that the logical dual yields the correct be- Figure 5: Translation of µGV concurrency features into core µGV haviour, whereas the naive dual does not. They (each) proposed a new definition of duality for recursive session types, using a selec- tive form of substitution which applies only inside carried types. that introduce purely administrative reduction. The session functor Later, Bernardi and Hennessy [7] observed that even this approach F is translated to the ordinary function F?, using promises to lift an fails on examples such as µSX.µSY.?Y.X. They propose converting ordinary type variable to session type. This approach naturally ac- each recursive session type into a so-called m-closed recursive ses- counts for the use of dualized session type variables; for example, if sion type before applying native duality. A recursive session type F(X ) = !X .X , then we have that F?(X) = !(?X).?X = !(!X).?X. µSX.T is m-closed if X does not occur free inside a carried type in Recursive and corecursive session types are interpreted as promises T. It is straightforward to show that every recursive session type is of recursive and corecursive data types, and the interpretation of (equirecursively) equivalent to an m-closed one, and that, as they their terms is directed by the interpretation of their types. The defi- ? ! they are essentially first-order, naive duality and logical duality nitions of M and M may seem surprisingly complicated. In coincide on m-closed recursive session types. fact, we canL M present a different form of session-typed catamor- Duality for recursive session types clearly needs to be treated phisms and anamorphisms, directly encoded in terms of recursive carefully. We are encouraged that our definition coincides with types: the state of the art for equirecursive self-dual session types. We M S N = M (un?M) believe that this also shows the value of our deconstruction of L S M L M recursive session types into well-understood primitives: we are L MN = send ( λx.fork (λc.L x c) M, N) guided immediately to a correct, compact, and general definition with the following typing rule, which exchanges the restriction to of duality. session functors for a direct use of their encoding:

Remark. Dualized session type variables are redundant in equire- ` M : F?(T) ( T ` M : S ( F!(S) ( end! cursive session types, as every session type µSX.T is equivalent to S ? S ? ` M : µ F ( T ` M : S ( ν F ( end! µSX.T{µSY.T{X/Y}/X}, where Y is a fresh type variable. How- L M ? S ever, dualized session type variables do yield a cleaner composi- We can see that the encoding of − is an instance of − , that tional definition of duality. the encoding of − ! is an instanceL ofM − S, and that ourL examplesM can be written directly using the alternative forms (removing calls ? ? 4. Communication with Concurrency to en as necessary). Nevertheless, we have preferred − and − ! as they do not rely on details of our encoding andL areM closer 4.1 Encoding Concurrent Features to the algebraic intuition. Our prior work on GV [26] focuses on keeping the core language 4.2 Concurrent Semantics as simple as possible. For example, rather than include branching and choice in the concurrent semantics directly, a choice can be We give a concurrent semantics of µGV, building on the small- encoded as the promise of a (data type) sum. Kobayashi et al [24] step operational semantics for Functional µGV given in the last and Dardha et al. [17] make similar use of linear promises to relate section. As recursive session types and their terms can be encoded data types and session types in π-calculi. We extend this view to in terms of the core concurrency features, our semantics is mostly include recursive session types. There are two challenges in doing unchanged from that of our prior work [26]. Figures6 and7 give the so: a) we must encode session functors and their use of session type syntax and typing of configurations and configuration contexts; we variables, and b) we must encode their fixed points. will write Γ ` C : T to denote that there is some φ such that Γ `φ Our translation is given by the homomorphic extension of the C : T. Figure8 gives reductions and configuration equivalence. rules in Figure5. We underline those portions of the translation Because of the importance of promises in our interpretation of We have encoded recursive session types using features of Func- Configurations C ::= φM | (new x)C tional µGV, and so they do not appear in the concurrent semantics 0 | z = x ↔ y | C k C directly. We would like to confirm that their encoding matches the Flags φ ::= ◦ | • intuition of the original, unencoded forms. That is, we hope that a Configuration contexts D ::= [ ] | (new x)D | C k D configuration H[ M ? x] k H0[in!x] reduces to H[M (F( M ?)x)] k Thread evaluation contexts H ::= φE H0[x]. This reductionL M is blocked by the administrative stepsL M intro- duced in the encoding of − ?. However, we can show that it Figure 6: Configurations and Contexts. holds if we can suitably ignoreL administrativeM reductions. To do so, we adapt a notion of weak bisimulation to our setting. Unlike stan- dard presentations of concurrency, all µGV reductions are internal.

] φ Therefore, ignoring all internal reductions would trivially identify Γ ` M : T Γ ` M : end! Γ, x : S ` C : T all processes that compute the same results. We intend a finer char- • ◦ φ Γ ` • M : T Γ ` ◦ M : end! Γ ` (new x)C : T acterization, in which we ignore only administrative reductions. We have already identified (by underlining) the relevant sources

◦ of administrative reductions. We say that a reduction is adminis- x : S, y : S, z : end? ` z = x ↔ y : end! trative (−→) if all the reduced subexpressions are underlined. For example, the reduction of H[send (V, x)] k H0[receive x] to H[x] k φ 0 ◦ 0 0 Γ, x : S ` C : T Γ , x : S ` C : end! H [(V, x)] is administrative, but the reduction of H[send (M, x)] to Γ, Γ0, x : S] `φ C k C0 : T H[send (M0, x)] is not (unless M is itself identified as administra- tive). We write −→? for the reflexive, transitive closure of −→, Figure 7: Configuration Typing. and write C =⇒ C0 to denote C −→?−→−→? C0. Finally, we can adapt the standard notion of weak bisimulation to our setting. Definition 4. A relation R on configurations is an administrative µGV’s concurrent features, we give special cases of the functor weak bisimulation if, for each C RC , whenever C =⇒ C0 , then map for the promise functor. These have the same behavior as 1 2 1 1 there is a C0 such that C =⇒ C0 and C0 RC0 , and similarly for that given in the general case, but expose potentially administrative 2 2 2 1 2 reduction from C . We define administrative weak bisimiliarity ≈ reductions sooner. Our treatment of link repairs a defect in our 2 to be the union of all administrative weak bisimulations. prior account of GV and restores the diamond property for GV’s concurrent semantics. We can now relate the encoding of recursive session types to their The concurrent semantics of µGV is defined by a reduction re- expected semantics: lation on collections of parallel threads, called configurations. The Theorem 5. syntax of µGV configurations is given in Figure6, and includes threads, name restriction and composition of configurations. Be- 1. If ` M : F(S) ( S, then cause µGV is a functional language, we distinguish the “main” ? 0 ! + thread • M, which we expect to compute the result of the computa- (new x)(H[ M x] k H [in x]) −→ ≈ L M ? 0 tion, from the child threads ◦ M. We give a typing judgment for con- (new x)(H[M (F( M ) x)] k H [x]) L M figurations, based on the type system for the linear π-calculus [24] 2. If ` M : S F(S), then but with two significant differences. First, we ensure that there is ( at most one main thread. This constraint is enforced by the flags (◦ (new x)(H[ M ! V x] k H0[out? x]) −→+≈ • and •) on the derivations: a derivation Γ ` C : T indicates that (new x)(H[F( M ! (MV)) x] k H0[x]) configuration C contains the main thread (returning a value of type ◦ ? ? T), while Γ ` C : end! indicates that C does not. Second, we The key observations to establishing this result are that un and en require that exactly one channel is shared at each composition of introduce only incidental additional concurrency. processes. This latter constraint is sufficient to guarantee deadlock Lemma 6. freedom and progress, as we show in §5. 1. (new x)(H[un? x] k H0[un! (λx.M, x)]) ≈ Theorem 1 (Diamond property). If Γ ` C : T,C ≡−→≡ C1, and (new x)(H[M] k H0[x]) C ≡−→≡ C , then either C ≡ C or there exists C such that 2 1 2 3 2.E [F(un?) (F(λx.en? (M x)) N)] ≈ E[F(M) N] C1 ≡−→≡ C3, and C2 ≡−→≡ C3. 3.E [F(λx.un! (λx.M, x)) (F(en!) N)] ≈ E[F(M) N] This proof extends to any deterministic extension of the core func- tional calculus, such as the addition of the exponential modality or The first is entirely straightforward, the second and third can be polymorphism. The reader may be concerned that the WAIT rule shown by induction on the structure of F. The theorem follows does not apply in the case that x is returned from the main thread, directly from the lemmas and the definition of reduction. and similarly for z in the LINK1 rule. However, these cases can never occur in a closed, well-typed configuration. 5. Communication without Concurrency The other metatheoretic properties established by in our prior work hold here as well. In particular, reduction in µGV preserves We now show, via a CPS translation, that reduction in µGV can be typing. simulated by reduction in Functional µGV. We begin with a stan- dard left-to-right call-by-value CPS translation from the core calcu- Theorem 2. If Γ `φ C : T and C −→ C0 then Γ `φ C0 : T. lus into itself (Figures9 and 10), where R is a fixed return type. In While typing is not preserved by configuration equivalence, reduc- the rest of this subsection, we extend the CPS translation to session tion never produces or relies on ill-typed states. types and show that the CPS translation preserves reduction. As a corollary, we obtain that µGV is strongly normalising. 0 Theorem 3. If Γ ` C1 : T,C1 ≡ C2, and C2 −→ C2, then there Following Danvy and Nielson [15], we can mechanically trans- 0 0 0 0 0 is some C1 such that C2 ≡ C1,C1 −→ C1 and Γ ` C1 : T. form the naive CPS translation N − of Figure 10 into a composi- J K Covariant Functors Contravariant Functors (!F.G)(M) = λc.fork (λd.let (z, d) = receive d in (!F.G)−(M) = λc.fork (λd.let (z, d) = receive d in let c = send (F−(M) z, c) in let c = send (F(M) z, c) in − link (G(M) d, G(M) c)) link (G (M) d, G−(M) c)) (?F.G)(M) = λc.fork (λd.let (z, c) = receive c in (?F.G)−(M) = λc.fork (λd.let (z, c) = receive c in let d = send (F(M) z, d) in let d = send (F−(M) z, d) in − link (G(M) d, G(M) c)) link (G (M) d, G−(M) c)) ? ? (?F.end?)(M) = λc.en (F(M)(un c)) − ? − ? (?F.end?) (M) = λc.en (F (M)(un c)) Configuration Equivalence

H[link (x, y)] ≡ H[link (y, x)] C k C0 ≡ C0 k C C k (new x)C0 ≡ (new x)(C k C0), if x 6∈ fv(C) 0 0 z = x ↔ y ≡ z = y ↔ x C1 k (C2 k C3) ≡ (C1 k C2) k C3 D[C] ≡ D[C ], if C ≡ C Configuration Reduction LIFTV 0 0 0 SEND (new x)(H[receive x] k H [send (V, x)]) −→ (new x)(H[(V, x)] k H [x]) M −→V M σ 0 FORK H[fork (λ y.M)] −→ (new x)(H[x] k ◦ M({x/y}] σ)), x fresh φM −→ φM WAIT (new x)(H[wait x] k ◦ x) −→ H[()] INK LIFT L 0 H[link (x, y)] −→ (new z)(z = x ↔ y k H[z]), z fresh 0 LINK1 (new z x)((z = x ↔ y k ◦ z) k φM) −→ φM{y/x} C −→ C D[C] −→ D[C0]

Figure 8: Concurrent Semantics of µGV: Functors, Equivalences, and Reductions. tional first-order one-pass CPS transformation K − . By carefully of an appropriate configuration. The second is a continuation. The distinguishing between values and non-values, theJ one-passK trans- CPS translation respects decomposition of contexts. lation ensures that (most) administrative redexes are contracted by Lemma 9. the translation itself. Contracting these redexes enables a tight sim- ulation result (Theorem 10). Due to lack of space, we omit the (en- 1. If E 6= [ ] and Γ ` E[I]: T, then K E[I] k = I I (λx.E E k). tirely standard) details of the one-pass variant of the translation. 2. If Γ ` D[C] and D[C] is well-oriented,J K thenJ KK D[CJ] kK = Figure 11 gives the CPS translation of concurrent µGV. The D D K C k. J K translations of send, fork, and link depend on the polarities (input J K J K or output) of their arguments and results. We use subscripts to To simulate all reduction paths in µGV by reduction in Func- distinguish output and input session types. To give a compositional tional µGV, we must allow reduction under lambda abstractions. translation of configurations, we restrict attention to a canonical Otherwise, we would be limited to a single schedule in which order class of configurations. We write C k C to denote a parallel of communication is determined by the outermost input communi- 1 x 2 cation. For an intuition of this schedule, consider the translation of composition in which channel x has input session type in C1 and 0 0 C kx C . Reduction in C cannot proceed until C is ready to receive the dual output session type in C2. We say that a configuration C 0 is well-oriented (WO(C)) if all of the parallel compositions in C a result along x, even if C could perform some internal commu- are of this form, and in any link configuration z = x ↔ y in C, nication. Allowing reduction under lambda abstractions allows all x has input session type. Without loss of generality, we need only valid schedules to be simulated. We define M N by: consider reduction on well-oriented configurations. M −→V N M N σ σ Lemma 7. If Γ ` C : T, then there exists well-oriented C0 ≡ C. M N λ x.M λ x.N The translation of the main thread is the only place the continuation The following theorem states that the CPS translation simulates is actually used. The translation of a child thread supplies the iden- µGV reduction. tity continuation, which is well-typed as child threads always have Theorem 10 (Simulation). type end!. Name restrictions themselves are ignored, but names are + used in the translation of well-oriented parallel composition. It is 1. If Γ ` M : T and M −→ N, then K M k K N k. straightforward to verify that the CPS translation preserves typing. 2. If Γ ` C : T and C −→ C0, thenJ thereK existJ well-orientedK C00, C000 with C00 ≡ C and C000 ≡ C0 such that K C00 k + Theorem 8 (Type soundness). K C000 k. J K 1. If Γ ` M : T, then K Γ ` K M :(K T R) R. J K ( ( Thus we can simulate concurrent communication using only Func- 2. If Γ ` C : T and CJ isK well-oriented,J K J thenK K Γ ` K C : tional µGV. As a corollary, we obtain that µGV is strongly normal- (K T ( R) ( R. J K J K J K ising. To reason by induction over the reduction rules, which are Theorem 11 (Strong normalization). If Γ ` C : T, then there are defined in terms of evaluation contexts and configuration contexts, no infinite ≡−→≡ sequences starting from C. we extend the CPS translation to contexts (Figure 12). Evaluation contexts are interpreted as functions. The CPS translation of a The proof follows immediately from Theorem 10 and the quite configuration context takes two arguments. The first argument is a standard result that Functional µGV (linear λ-calculus with pos- meta-level function, which we instantiate with the CPS translation itive (co)recursive data types) is strongly normalising. (To show Types K T ( U = K T ( (K U ( R) ( R J K J K J K K end! = R K !T.S = K T ( K S ( R K T ⊗ U = K T ⊗ K U K 1 = 1 J K J K J K J K J K J K J K J K K end? = R ( R K ?T.S = (K T ( K S ( R) ( R K T ⊕ U = K T ⊕ K U K 0 = 0 J K J K J K J K KJT & UK = KJTK & KJUK K J>K = > Constants J K J K J K J K K µX.T = µX.K T K send! p k = let (x, c) = p in (c x) k K X = X J K KJνX.TK = νX.KJT K K send? p k = let (x, c) = p in k (c x) J K J K J K J K K receive c k = c (λx c.k (x, c)) Figure 9: CPS Translation for Core Types J K K fork! f k = k (λx.f x id) K Jfork?Kf k = (λx.f x id) k J K K wait c k = c (k ()) J K N x k = k x K link! p k = let (c, d) = p in k (c d) N KMJ Kk = N M (λx.N K x k) K Jlink?Kp k = let (c, d) = p in k (d c) {~V/~zJ} K J K J K J K N λ x.M k = k (λx k.N V~ (λ~z.N M k)) Shallow Polarization J N MNKk = N M (λx.JNK N (λyJ.x yK k)) N (MJ , N)Kk = N JMK(λx.N JNK(λy.k (x, y))) S! := !T.S | end! S? ::= ?T.S | end? J K J K J K N let (x, y) = M in N k = N M (λz.let (x, y) = z in N N k) send : T ⊗ !T.S S send : T ⊗ !T.S S J N ()Kk = k ()J K J K ! ! ( ! ? ? ( ? fork :(S end ) S fork :(S end ) S N let () = M inJNKk = N M (λz.let () = z in N N k) ! ! ( ! ( ! ? ? ( ! ( ? J N inl MKk = N JMK(λx.k (inl x)) J K link! : S! ⊗ S! ( end! link? : S? ⊗ S? ( end! N Jinr MKk = N JMK(λx.k (inr x)) Configurations J K J K   ucase M { } case z { K • M k = K M k K (new x)C k = K C k J K J K J K J K N inl x 7→N; k = N M λz. inl x 7→N N k;  K ◦ M k = K M id K z = x ↔ y k = z (x y) 0 0 0 0 v inr y 7→N }~ J K inr y 7→N JNK k} K C kJx C Kk = (KJ CK k){(λx.KJ C k)/x}K J K J K J K J K N absurd M k = N M (λz.absurd z) Figure 11: CPS Translation for Concurrent µGV and Contexts. JN hM, NiKk = N JMK(λx.N N (λy.k hx, yi)) NJ fst MKk = N JMK(λx.k (JfstKx)) N Jsnd MKk = N JMK(λx.k (snd x)) J N hiKk = k hiJ K Evaluation Contexts N inJMKk = N M (λx.k (in x)) E E k = λx.K E[e] k NJ M Kk = N JMK(λx.k ( x )) J K J K N outJL MMKk = N JMK(λx.k (outL M x)) Configuration Contexts NJ M Kk = N JMK(λx.k ( x )) D [] f k = f k J K J K D (new xJ)DKf k = K D f k Figure 10: Naive CPS Translation for Core Terms JD C kx DKf k = (KJ CK k){λx.D D f k/x} DJD kx CKf k = (DJDKf k){λx.KJ CK k/x} J K J K J K Figure 12: CPS Translation of µGV Contexts. the latter, map functional µGV into System F, forgetting linearity, and encoding the positive (co)recursive data types using polymor- phism.) An immediate consequence is that our calculus is free from 6. The µCP Language livelock; that is, that there are no oscillating sequences of configu- rations that diverge. In this section, we present the syntax (§6.1) and semantics (§6.2) The strong normalization result straightforwardly extends to of µCP, an extension of Wadler’s CP calculus with recursive and the extension of µGV with the exponential modality [26], and to corecursive types following Baelde’s approach to recursion and the setting where we allow reduction under lambdas in the source corecursion in classical linear logic [3]. We then argue that µCP and calculus. µGV are equally expressive via semantics-preserving translations As well as providing a means to prove termination, the CPS from µCP into µGV and vice versa. transformation is interesting in its own right as it provides insights 6.1 Syntax into the restricted nature of the concurrency provided by µGV. Furthermore, by composing the CPS translation with the translation Figure 13 gives the terms and typing of µCP, an extension of from µCP to µGV (§6.3), we obtain a translation from µCP into a Wadler’s process calculus CP with recursive and corecursive types. typed lambda calculus. This shows that the concurrency of µCP is The syntax of types is that of the propositions of linear logic, ex- equivalent to that provided by full reduction in the λ-calculus. tended with least (µF) and greatest (νF) fixed points. As in µGV, We can make the translation more uniform by factoring it we have omitted polymorphism and the exponential modality; their through a polarization phase. In particular, polarization allows us reintroduction is entirely orthogonal to our development. The def- to give a single translation for each of the send, fork, and link cases. inition of duality includes the duality of least and greatest fixed Polarization provides a way of encoding output session types as in- points; the dual of an operator is defined by F⊥(X) = (F(X⊥))⊥, put session types and vice-versa. It also leads to a clean way of as for µGV. The terms of µCP are restricted compared to π- handling polymorphic session types by uniformly choosing either calculus in several ways. Most significantly, composition and name a positive (output) or a negative (input) representation for session restriction are combined in a single syntactic form, and the com- type variables. posed processes are limited to share only the newly introduced Syntax presentations of introduction and elimination rules for induction and coinduction, as follows: Types A, B ::= A ⊗ B | A B | A ⊕ B | A & B | 1 | ⊥ | > |`0 | X | X⊥ | µF | νF Ψ ` F(µF) F(A) ` A Ψ, νF ` A A ` F(A) Operators F, G ::= X.A Ψ ` µF µF ` A Ψ, F(νF) ` A A ` νF Labels ` ∈ inl, inr Note that the hypotheses of the right rule for ν and left rule for µ are Processes P, Q, R ::= x[y].(P | Q) | x(y).P | x[].0 | x().P restricted to account for linearity. Baelde observes that, when using | x[`].P | case x {P; Q} | case x {} duality to convert these two-sided sequents to one-sided sequents, | x ↔ y | new x (P | Q) the left rule for µ and the right rule for ν collapse, and similarly the | rec x.P | corec x[y](P | Q) right rule for µ and the left rule for ν. This leaves us with only two Typing rules, with term assignments as follows: ⊥ P ` Ψ, x : AQ ` Ψ0, x : A⊥ P ` y : A, x : F(A) P ` Ψ, x : F(µF) ⊥ x ↔ y ` x : A, y : A⊥ new x (P | Q) ` Ψ, Ψ0 corec x(y).P ` y : A , x : νF rec x.P ` Ψ: x : µF However, there is a problem with this formulation. Suppose that we P ` Ψ, y : AQ ` Ψ0, x : B have some term Q ` A. We then have the composition new y (Q | x[y].(P|Q) ` Ψ, ∆0, x : A ⊗ B x[].0 ` x : 1 corec x(y).P) ` x : νF. However, we have no hope of reducing this cut, as we have no rule which can prove νF in isolation. We can P ` Ψ, x : B, y : A P ` Ψ address this problem by suspending the cut in question, moving it into the ν rule and giving the rule in Figure 13. In our prior x(y).P ` Ψ, x : A B x().P ` Ψ, x : ⊥ work [26] we observe a similar pattern in comparing the ⊗ rule ` to the typical process calculus rule for output. As in that case, P ` Ψ, x : A P ` Ψ, x : AQ ` Ψ, x : B the version without the suspended cut may expose reductions not x[inl].P ` Ψ, x : A ⊕ B case x {P; Q} ` Ψ, x : A & B present in the suspended version. Nevertheless, we can still define the simpler term as syntactic sugar: P ` Ψ, x : F(µF) corec x(y).P = corec x[y](y ↔ z | P) case x {} ` Ψ, x : > rec x.P ` Ψ, x : µF Examples. We return to the example of natural numbers to give ⊥ some flavor of the use of recursion and corecursion in µCP. We can P ` Ψ, y : AQ ` y : A , x : F(A) define the type of natural numbers much as before corec x[y](P | Q) ` Ψ, x : νF N(X) = 1 ⊕ X Nat = µN Duality and we can give very similar definitions of the constructors (A ⊗ B)⊥ = A⊥ B⊥ 1⊥ = ⊥ ⊥⊥ = 1 zerox = rec x.x[inl].x[].0 ⊥ ⊥ ⊥ ⊥ ⊥ (A B) = A `⊗ B 0 = > > = 0 succxy = rec x.x[inr].x ↔ y (A `⊕ B)⊥ = A⊥ & B⊥ (µF)⊥ = ν(F⊥) (X⊥)⊥ = X with the expected typings zero ` x : Nat and succ ` x : Nat, y : (A & B)⊥ = A⊥ ⊕ B⊥ (νF)⊥ = µ(F⊥) F⊥(X) = (F(X⊥))⊥ x xy Nat⊥. We can define the addition operation as follows:

Figure 13: µCP Typing Rules plusxyz = corec z[w].(whxi.w ↔ y; w(x).case z {z().w ↔ x; rec x.x[inr].zhxi.z ↔ w}) name. The forwarding construct x ↔ y corresponds to the axiom where the recursive body of the corec has type w : Nat Nat⊥, z : rule in linear logic; it is necessary for the treatment of recursion and ⊥ N(Nat ⊗ Nat) and so the term has typing plusxyz ` x`: Nat, y : for the extension of µCP to include polymorphism. ⊥ ⊥ Nat , z : Nat . Writing nz to denote the encoding of the natural A Simpler Send. The µCP rule for output is appealing because number n along channel z, we have that it corresponds exactly to the linear logic proof rule for ⊗. Its new z (2 | new y (2 | plus )) correspondence to π-calculus is less direct: the term for output z y xyz includes name restriction (introducing new name y), output of y will reduce to 4x along x, and finally a restricted composition (of P and Q). This complicates the reduction relation for µCP (which must account 6.2 Semantics for all three behaviors) and the correspondence to µGV (where the The semantics of µCP are given by the cut reduction rules in classi- rule for send is simpler). cal linear logic, extended to account for recursion and corecursion, An alternative presentation (following Boreale [12]) avoids the as shown in Figure 14. We write fv(P) for the free names of process name restriction and composition, as follows: P. Terms are identified up to congruence ≡. Many of the principle cut reductions (−→ ) correspond to process calculus reductions. P ` Ψ, x; B, y : A C The reduction of input against output is complicated by the im- ⊥ xhyi.P ` Ψ, x : A ⊗ B, y : A plicit name restriction and composition inherent in the term struc- While no longer identical to the ⊗ rule, this is closer to the formu- ture for output. The new rule for µCP is for rec against corec, and lation of output in π-calculus. As in our prior work [26], we write amounts to one unfolding of the corec term. In defining the unfold- ing, we rely on functoriality for the operators; if P ` x : A⊥, y : B, xhyi.P as syntactic sugar for x[z].(y ↔ z | P). F ⊥ ⊥ then mapx,y(P) ` x : F (A ), y : F(B). (We show functoriality Recursion and Corecursion. In µCP, recursion and corecursion for the positive combinators; the remaining cases can be obtained follow Baelde’s extension of classical linear logic to include induc- by switching the channels in the given cases.) We write −→ for ? ? tion and coinduction [3]. We begin by considering sequent calculus −→C−→CC. The following theorem is due to Baelde [3]: Structural Congruence x ↔ y ≡ y ↔ x new y (P | new x (Q | R)) ≡ new x (new y (P | Q) | R) if y 6∈ fv(R) new x (P | Q) ≡ new x (Q | P) new x (P1 | Q) ≡ new x (P2 | Q) if P1 ≡ P2 Functoriality (positive cases) X.A F⊕G F G mapx,y (P) = x ↔ y, X ∈/ FTV(A) mapx,y (P) = case x {y[inl].mapx,y(P); y[inr].mapx,y(P)} mapX.X(P) = P X.µF F⊥ X.F(µF) x,y mapx,y (P) = corec x(y).rec y.mapx,y (P) F⊗G 0 0 F 0 0 G mapx,y (P) = x(x ).y[y ].(mapx0,y0 (P{x /x, y /y}) | mapx,y(P)) Primary Cut Reductions

new x (x ↔ y | P) −→C P{y/x} new x (x[y].(P|Q) | x(y).R) −→C new x (Q | new y (P | R)) new x (p[inl].P | case x {Q; R}) −→C new x (P | Q) F F F new x (corec x[y](P | Q) | rec x.R) −→C new y (P | new z (Q{z/x} | new x (mapx,z(corec x[y](z ↔ y | Q)) | R))) Commuting Conversions

new z (x[y].(P | Q) | R) −→CC x[y].(new z (P | R) | Q) new z (case x {P; Q} | R) −→CC case x {new z (P | R); new z (Q | R)} new z (x[y].(P | Q) | R) −→CC x[y].(P | new z (Q | R)) new z (case x {} | P) −→CC case x {} new z (x(y).P | Q) −→CC x(y).new z (P | Q) new z (rec x.P | Q) −→CC rec x.new z (P | Q) new z (x().P | Q) −→CC x().new z (P | Q) new z (corec x[y](P | Q) | R) −→CC corec x[y](new z (P | R) | Q) new z (x[inl].P | Q) −→CC x[inl].new z (P | Q)

Figure 14: µCP Reduction Rules

Theorem 12 (Cut elimination). If P ` Ψ, then there is some P0 7. Related Work such that P −→ P0 and P0 is not of the form new x (Q | Q0) for any 0 Session Types and Linear Logic. Session types were originally x, Q, Q . introduced by Honda [22] as a typing discipline for a CCS-like This result corresponds to the termination and deadlock freedom process calculus. Takeuchi et al [30] and Honda et al. [23] ex- results for µGV: any well-typed process reduces to one that is tended the original approach to include delegation and recursion. blocked on external communication. The commuting conversions Honda’s system relied on a substructural type system, and bor- (−→CC) do not correspond to computational steps (and thus, do not rowed some syntax from linear logic, but did not draw a direct con- correspond to reductions in process calculi), but play a crucial role nection between the systems nor suggest the connection between in cut elimination by moving remaining internal communication the input and output session types and the ⊗ and connectives. behind any external communication. Abramsky [1] and Bellin and Scott [5] give interpretations` of lin- ear logic proofs as π-calculus processes, and of cut elimination as 6.3 Translations between µCP and µGV π-calculus reduction. Their interpretations of ⊗ and are very dif- We conclude our discussion of µCP by discussing its relationship ferent from the interpretations of input and output in` session types. with µGV. Our previous work [26] considers a functional calculus Caires and Pfenning [13] give the first formal correspondence be- GV and a process calculus CP, similar to µGV and µCP but lacking tween session types and linear logic, interpreting the propositions recursion and corecursion. In that setting we give translations C − of intuitionistic linear logic as session types, and showing that and G − from GV configurations to CP terms and vice versaJ andK π-calculus reduction corresponds with cut reduction. As a conse- show thatJ K these translations preserve both typing and semantics. quence of the latter correspondence, they show that cut elimination We have devised similar translations between µGV and µCP. in linear logic proves deadlock freedom for session-typed π calcu- We omit the details of the translations here. This is not only for lus terms. Vasconcelos et al. [33] and Gay and Vasconcelos [20] reasons of space; the details of the translations are identical but consider functional languages extended with session-typed con- for the treatment of recursion, and the recursive forms in µGV and currency. The functional fragments of their calculi are generally µCP are already quite similar. As in the case of the non-recursive less fully featured than ours (for example, they omit sums) while calculi, these translations preserve both typing and semantics. their concurrent fragments include non-determinism and deadlock. Wadler [34] presents a process calculus, called CP, similar to that of ◦ Theorem 13. If P ` Ψ, then G Ψ ` G P : end!. Caires and Pfenning, but based on classical rather than intuitionis- J K J K tic linear logic. He also gives a functional calculus, called GV, and Theorem 14. If P ` Ψ and P −→ Q then G P −→+≈ C C similar to that of Gay and Vasconcelos. He shows a type-preserving GC Q . J K J K translation from GV to CP; however, his GV is less expressive than Theorem 15. If Γ ` C : T, then C C z ` C Γ , z : C T ⊥. CP. Our prior work [26] introduces a more expressive variant of GV, + J K J K J K based on Wadler’s, and gives both a direct semantics and semantics- We write =⇒ for (≡−→≡) . preserving translations to and from Wadler’s CP. Theorem 16. If Γ ` C : T and C −→ C0, then C C z =⇒ C C0 z. J K J K Recursive and Corecursive Definition. The interpretation of re- The simulation of µGV evaluation by µCP evaluation is straight- cursive data types, and their connection to recursive functions, has forward. The simulation of µCP evaluation by µGV evaluation re- been studied extensively; we highlight the direct precursors of our lies on the same administrative weak bisimulation results needed approach. Goguen et al [21] proposed initial algebras and their cor- to show that the core µGV calculus simulates its extension with responding folds as a means for understanding recursive data types recursive session types, but introduces no other difficulties. and their use. Meijer et al. [29] characterized the use of both folds and unfolds, among other patterns, in the definition of recursive compass additional features. We have given a natural, semantically functional programs. The coincidence of least and greatest fixed justified approach to extending our system to non-terminating (but points for data type constructors in many models was first observed still productive) computation. Finally, we have related our calculus by Freyd [19]; he argues that this observation justifies the use of to a process calculus based on classical linear logic with induction such fixed points for recursive data types. Baelde and Miller [4] and coinduction, giving strong logical foundations to our work. first described an extension of linear logic with induction and coin- The model of concurrency in µGV (like much of the work on duction, encoded using the exponential modality and second-order logically founded session types) is somewhat limited. By ruling quantification. Baelde [3] treats induction and coinduction without out deadlock, livelock, and in particular data races, it also rules encoding; in particular, he gives a cut reduction rule for recursive out interesting forms of concurrency. For instance, there is no way and corecursive terms, and shows cut elimination directly. of modeling a non-deterministic stateful service such as an online book store in which it should be possible for one customer to ob- Recursive Session Types. There have been several recent devel- serve that a different customer bought the last copy of a book. This opments of recursive session types and their relationship with linear raises the question: can the logically founded approach be extended logic. We highlight three closely related to our development. to encompass more realistic forms of concurrency? Recently, Atkey Toninho et al [32] present a system with recursive session types et al. [2] demonstrated that conflating dual propositions in a process based on intuitionistic linear logic extended with corecursion. They calculus based on classical linear logic captures non-determinism, arrive at a similar (albeit intuitionistic) typing discipline for core- shared state, and more expressive concurrency patterns. On the one cursive session types to ours (§2.3), and give a direct proof of termi- hand we would like to transfer these results to GV. On the other nation for the resulting system (without encoding). However, their hand we would we would like to transfer ideas for the current pa- approach differs from ours in several significant ways. First, they per to extend the work of Atkey et al. to incorporate conflated fixed treat recursive processes as primitive, and so do not expose the con- points. nection with recursive data types. In contrast, we believe that the While the development of µGV in this paper is largely theo- parallels with data types (and thus, our ability to present a simple retical, we believe it can also inform practical implementations of core calculus) is one of the principal benefits of our approach. One ! ? session types, including our implementation of session types for consequence is that they have only corecursive processes (ν , ν in the Links web programming language [31]. Indeed, the Links im- ? ! our notation), but not recursive processes (µ , µ ) nor the possibil- plementation of session types is based on a core calculus FST (Sys- ity of identifying greatest and least fixed points. Finally, our session tem F with Session Types), a polymorphic variant of GV [27]. We types are classical, while theirs are intuitionistic. One consequence would like to investigate the relationship between µGV with FST, of our approach is that we are explicit about the role of duality, and while also taking into account practical considerations. thus identify a new notion of duality for recursive session types, In this paper we have focused on binary session types. Carbone while their notion of duality is implicit in the type system. We see et al. [14] give a logical account of multiparty session types based the similarities, despite theoretical and methodological differences, on CP. It would be interesting to adapt this work to the µGV setting as indicative of the strength of both approaches. by following our general approach to relating variants of CP and Dardha [16] gives an encoding of recursive session-typed π- GV. calculus into recursive (non-linear) π-calculus, and shows that this encoding preserves both typing and semantics. Her encoding is based on self-referential replicated processes, and thus supports ar- Acknowledgments bitrary non-termination, while not attempting to guarantee dead- Thanks to Giovanni Bernardi, Ornela Dardha, Simon Fowler, Philip lock or livelock freedom. She adopts a coinductive definition of Wadler, and the anonymous referees for helpful feedback. This duality from Bernardi et al. [8], which relies on partially unfolding work was funded by EPSRC grant number EP/K034413/1. recursive types at each computation of their duals. Bono and Padovani [10] and Bernardi and Hennessy [6] inde- pendently observed that the standard definition of duality for re- References cursive session types fails when recursion occurs in a carried type. [1] S. Abramsky. Proofs as processes. Theor. Comput. Sci., 135(1):5–9, Bono et al [11] present a session-typed functional language with 1994. self-dual recursion. They do not attempt to enforce termination or [2] R. Atkey, S. Lindley, and J. G. Morris. Conflation confers concurrency. distinguish recursion and corecursion, but do present a definition In S. Lindley, C. McBride, P. W. Trinder, and D. Sannella, editors, A of duality similar to the one that we propose for self-dual recursive List of Successes That Can Change the World - Essays Dedicated to session types (§3.2). Bernardi et al [8] systematically study several Philip Wadler on the Occasion of His 60th Birthday, volume 9600 of duality relations, and propose a notion of session typing indepen- Lecture Notes in Computer Science, pages 32–55. Springer, 2016. dent of the particular duality relation. They also give a coinductive [3] D. Baelde. Least and greatest fixed points in linear logic. ACM Trans. characterization of duality, and suggest a syntactic instance of their Comput. Logic, 13(1):2:1–2:44, Jan. 2012. characterization. A particular concern of their work, absent from [4] D. Baelde and D. Miller. Least and greatest fixed points in linear logic. ours, is subtyping: a process may offer more choices than those In N. Dershowitz and A. Voronkov, editors, Logic for Programming, from which its partner selects. However, their definitions are more Artificial Intelligence, and Reasoning, 14th International Conference, complex than ours even without considering subtyping; in particu- LPAR 2007, Yerevan, Armenia, October 15-19, 2007, Proceedings, lar, they rely on partially unfolding recursive types in each compu- volume 4790 of Lecture Notes in Computer Science, pages 92–106. tation of their duals. Springer, 2007. [5] G. Bellin and P. J. Scott. On the π-Calculus and linear logic. Theoret- ical Computer Science, 135(1):11–65, 1994. 8. Future Work [6] G. Bernardi and M. Hennessy. Using higher-order contracts to model We have presented a core concurrent linear λ-calculus with recur- session types (extended abstract). In P. Baldan and D. Gorla, editors, sive and corecursive data types, and shown how to encode recursive CONCUR 2014, volume 8704 of Lecture Notes in Computer Science, and corecursive session types and processes in this calculus. We pages 387–401. Springer, 2014. have shown that our type system guarantees termination and lock [7] G. Bernardi and M. Hennessy. Using higher-order contracts to model freedom, giving modular proofs which can easily be extended to en- session types. CoRR, abs/1310.6176v4, 2015. [8] G. Bernardi, O. Dardha, S. J. Gay, and D. Kouzapas. On duality 2010. relations for session types. In Trustworthy Global Computing - 9th [21] J. A. Goguen, J. W. Thatcher, E. G. Wagner, and J. B. Wright. Initial International Symposium, TGC 2014, Rome, Italy, September 5-6, algebra semantics and continuous algebras. J. ACM, 24(1):68–95, 2014. Revised Selected Papers, pages 51–66, 2014. 1977. [9] R. S. Bird and O. de Moor. Algebra of programming. Prentice Hall [22] K. Honda. Types for dyadic interaction. In E. Best, editor, CON- International series in computer science. Prentice Hall, 1997. CUR ’93, 4th International Conference on Concurrency Theory, [10] V. Bono and L. Padovani. Typing copyless message passing. Logical Hildesheim, Germany, August 23-26, 1993, Proceedings, volume 715 Methods in Computer Science, 8(1), 2012. of Lecture Notes in Computer Science, pages 509–523. Springer, 1993. [11] V. Bono, L. Padovani, and A. Tosatto. Polymorphic types for leak [23] K. Honda, V. T. Vasconcelos, and M. Kubo. Language primitives and detection in a session-oriented functional language. In D. Beyer and type discipline for structured communication-based programming. In M. Boreale, editors, Formal Techniques for Distributed Systems - Joint C. Hankin, editor, ESOP, volume 1381 of Lecture Notes in Computer IFIP WG 6.1 International Conference, FMOODS/FORTE 2013, Held Science, pages 122–138. Springer, 1998. as Part of the 8th International Federated Conference on Distributed [24] N. Kobayashi, B. C. Pierce, and D. N. Turner. Linearity and the pi- Computing Techniques, DisCoTec 2013, Florence, Italy, June 3-5, calculus. ACM Trans. Program. Lang. Syst., 21(5):914–947, 1999. 2013. Proceedings, volume 7892 of Lecture Notes in Computer Sci- ence, pages 83–98. Springer, 2013. [25]J.L evy´ and L. Maranget. Explicit substitutions and programming languages. In Foundations of Software Technology and Theoretical [12] M. Boreale. On the expressiveness of internal mobility in name- Computer Science, 1999, volume 1738 of LNCS. Springer, 1999. passing calculi. In U. Montanari and V. Sassone, editors, CONCUR ’96, Concurrency Theory, 7th International Conference, Pisa, Italy, [26] S. Lindley and J. G. Morris. A semantics for propositions as ses- August 26-29, 1996, Proceedings, volume 1119 of Lecture Notes in sions. In Programming Languages and Systems - 24th European Sym- Computer Science, pages 163–178. Springer, 1996. posium on Programming, ESOP 2015, Held as Part of the European Joint Conferences on Theory and Practice of Software, ETAPS 2015, [13] L. Caires and F. Pfenning. Session types as intuitionistic linear propo- London, UK, April 11-18, 2015. Proceedings, pages 560–584, 2015. sitions. In P. Gastin and F. Laroussinie, editors, CONCUR 2010 - Concurrency Theory, 21th International Conference, CONCUR 2010, [27] S. Lindley and J. G. Morris. Lightweight functional session types, Paris, France, August 31-September 3, 2010. Proceedings, volume 2015. Draft http://homepages.inf.ed.ac.uk/slindley/ 6269 of Lecture Notes in Computer Science, pages 222–236. Springer, papers/fst-draft-february2015.pdf. 2010. [28] B. Liskov and L. Shrira. Promises: Linguistic support for efficient [14] M. Carbone, S. Lindley, F. Montesi, C. Shurmann,¨ and P. Wadler. asynchronous procedure calls in distributed systems. In Proceedings Coherence generalises duality: a logical explanation of multiparty of the ACM SIGPLAN 1988 Conference on Programming Language session types. In CONCUR. LIPICS, 2016. To appear. Design and Implementation, PLDI ’88, pages 260–267, New York, NY, USA, 1988. ACM. [15] O. Danvy and L. R. Nielsen. A first-order one-pass CPS transfor- mation. In M. Nielsen and U. Engberg, editors, Foundations of Soft- [29] E. Meijer, M. M. Fokkinga, and R. Paterson. Functional programming ware Science and Computation Structures, 5th International Confer- with bananas, lenses, envelopes and barbed wire. In J. Hughes, editor, ence, FOSSACS 2002. Held as Part of the Joint European Conferences Functional Programming Languages and Computer Architecture, 5th on Theory and Practice of Software, ETAPS 2002 Grenoble, France, ACM Conference, Cambridge, MA, USA, August 26-30, 1991, Pro- April 8-12, 2002, Proceedings, volume 2303 of Lecture Notes in Com- ceedings, volume 523 of Lecture Notes in Computer Science, pages puter Science, pages 98–113. Springer, 2002. 124–144. Springer, 1991. [16] O. Dardha. Recursive session types revisited. In Proceedings Third [30] K. Takeuchi, K. Honda, and M. Kubo. An interaction-based language Workshop on Behavioural Types, BEAT 2014, Rome, Italy, 1st Septem- and its typing system. In C. Halatsis, D. G. Maritsas, G. Philokyprou, ber 2014., pages 27–34, 2014. and S. Theodoridis, editors, PARLE ’94: Parallel Architectures and Languages Europe, 6th International PARLE Conference, Athens, [17] O. Dardha, E. Giachino, and D. Sangiorgi. Session types revisited. Greece, July 4-8, 1994, Proceedings, volume 817 of Lecture Notes In D. D. Schreye, G. Janssens, and A. King, editors, Principles and in Computer Science, pages 398–413. Springer, 1994. Practice of Declarative Programming, PPDP’12, Leuven, Belgium - September 19 - 21, 2012, pages 139–150. ACM, 2012. [31] The Links Team. Links, 2016. http://groups.inf.ed.ac.uk/ [18] A. Filinski. Linear continuations. In R. Sethi, editor, Conference links. Record of the Nineteenth Annual ACM SIGPLAN-SIGACT Symposium [32] B. Toninho, L. Caires, and F. Pfenning. Corecursion and non- on Principles of Programming Languages, Albuquerque, New Mexico, divergence in session-typed processes. In Trustworthy Global Com- USA, January 19-22, 1992, pages 27–38. ACM Press, 1992. puting - 9th International Symposium, TGC 2014, Rome, Italy, [19] P. Freyd. Algebraically complete categories. In G. R. Aurelio Carboni, September 5-6, 2014. Revised Selected Papers, pages 159–175, 2014. Maria Cristina Pedicchio, editor, - Proceedings of [33] V. T. Vasconcelos, S. J. Gay, and A. Ravara. Type checking a multi- the International Conference held in Como, Italy, July 22–28, 1990. threaded functional language with session types. Theor. Comput. Sci., Springer, 1990. 368(1-2):64–87, 2006. [20] S. J. Gay and V. T. Vasconcelos. Linear type theory for asynchronous [34] P. Wadler. Propositions as sessions. J. Funct. Program., 24(2-3):384– session types. Journal of Functional Programming, 20(01):19–50, 418, 2014.