a semantical and operational account of call-by-value ...giuliog/fossacs.pdf · a semantical and...

27
A semantical and operational account of call-by-value solvability Alberto Carraro 1,2 and Giulio Guerrieri 2 1 DAIS, Universit` a Ca’ Foscari Venezia, [email protected] 2 PPS, Univ Paris Diderot, Sorbonne Paris Cit´ e [email protected] Abstract. In Plotkin’s call-by-value lambda-calculus, solvable terms are characterized syntactically by means of call-by-name reductions and there is no neat semantical characterization of such terms. Preserving confluence, we extend Plotkin’s original reduction without adding extra syntactical constructors, and we get a call-by-value operational characterization of solvable terms. Moreover, we give a semantical characterization of solvable terms in a relational model, based on Linear Logic, satisfying the Taylor expansion formula. As a technical tool, we also use a resource-sensitive calculus (with tests) in which the elements of the model are definable. Keywords: (resource) call-by-value lambda calculus, tests, potential valuability, solvability, relational semantics, weak and stratified reductions 1 Introduction In the theory of ordinary (i.e. untyped call-by-name) λ-calculus, the notion of solvability plays a crucial role. A λ-term M is solvable if there is a head context H such that HLM M β λx.x = I (the identity); M is unsolvable if it is not solvable. Solvability (see [1]) underlies the fundamental notions of approximants, B¨ohm- trees and separability; moreover, it is possible to encode partial recursive functions in λ-calculus in such a way that undefinedness is represented by unsolvable λ- terms ([1, Ch. 8]). Enforcing the idea of unsolvable-as-meaningless, it is consistent to equate all unsolvable λ-terms (but not all λ-terms having no β-normal form, [1, Ch. 16]). A fundamental theorem for ordinary λ-calculus (see [2,3]) states that for every λ-term M the following are equivalent: (1) M is solvable; (2) the head reduction of M terminates; (3) the semantics of M in the Scott’s model D is not the least element. Equivalence (1)(2) (resp. (1)(3)) gives a semantical (resp. syntactical or operational ) characterization of solvability in ordinary λ-calculus. The most common parameter passing policy for programming languages is call-by-value (CBV). Plotkin [4] introduced the λ v -calculus in order to grasp the CBV paradigm in a pure λ-calculus setting. The λ v -calculus (without constants) has the same syntax as ordinary λ-calculus but its β v -reduction rule allows the contraction of a β-redex only if the argument is a λ-value, i.e. a variable or an abstraction. As argued in [5], a good CBV λ-calculus should enjoy an internal op- erational characterization (i.e. by using CBV reduction rules) of CBV-solvability.

Upload: others

Post on 30-Apr-2020

3 views

Category:

Documents


0 download

TRANSCRIPT

Page 1: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account ofcall-by-value solvability

Alberto Carraro1,2 and Giulio Guerrieri2

1 DAIS, Universita Ca’ Foscari Venezia, [email protected] PPS, Univ Paris Diderot, Sorbonne Paris [email protected]

Abstract. In Plotkin’s call-by-value lambda-calculus, solvable terms arecharacterized syntactically by means of call-by-name reductions and thereis no neat semantical characterization of such terms. Preserving confluence,we extend Plotkin’s original reduction without adding extra syntacticalconstructors, and we get a call-by-value operational characterization ofsolvable terms. Moreover, we give a semantical characterization of solvableterms in a relational model, based on Linear Logic, satisfying the Taylorexpansion formula. As a technical tool, we also use a resource-sensitivecalculus (with tests) in which the elements of the model are definable.

Keywords: (resource) call-by-value lambda calculus, tests, potentialvaluability, solvability, relational semantics, weak and stratified reductions

1 Introduction

In the theory of ordinary (i.e. untyped call-by-name) λ-calculus, the notion ofsolvability plays a crucial role. A λ-term M is solvable if there is a head context H

such that HLMMβ λx.x = I (the identity); M is unsolvable if it is not solvable.Solvability (see [1]) underlies the fundamental notions of approximants, Bohm-trees and separability; moreover, it is possible to encode partial recursive functionsin λ-calculus in such a way that undefinedness is represented by unsolvable λ-terms ([1, Ch. 8]). Enforcing the idea of unsolvable-as-meaningless, it is consistentto equate all unsolvable λ-terms (but not all λ-terms having no β-normal form,[1, Ch. 16]). A fundamental theorem for ordinary λ-calculus (see [2,3]) states thatfor every λ-term M the following are equivalent: (1) M is solvable; (2) the headreduction of M terminates; (3) the semantics of M in the Scott’s model D∞ is notthe least element. Equivalence (1)⇔(2) (resp. (1)⇔(3)) gives a semantical (resp.syntactical or operational) characterization of solvability in ordinary λ-calculus.

The most common parameter passing policy for programming languages iscall-by-value (CBV). Plotkin [4] introduced the λv-calculus in order to grasp theCBV paradigm in a pure λ-calculus setting. The λv-calculus (without constants)has the same syntax as ordinary λ-calculus but its βv-reduction rule allows thecontraction of a β-redex only if the argument is a λ-value, i.e. a variable or anabstraction. As argued in [5], a good CBV λ-calculus should enjoy an internal op-erational characterization (i.e. by using CBV reduction rules) of CBV-solvability.

Page 2: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

2 A. Carraro, G. Guerrieri

This is not the case for Plotkin’s λv-calculus and the weakness of βv-reductionis widely recognized and accepted. Following [6,7], a λ-term M is λv-solvableif there is a head context H such that HLMM βv

I. Let ∆ = λx.xx: there isno head context sending (via βv-reduction) N = (λy.∆)(xI)∆ to I, thus N isλv-unsolvable and hence it should be divergent, whereas it is βv-normal. An oper-ational characterization of λv-solvability has been provided in [6,7] but through acall-by-name reduction; this result is improved in [8] where the characterizationis built upon strong normalization of the (call-by-name) lazy β-reduction.

There are many proposals of alternative CBV λ-calculi (see [9,10,11,12,5])extending Plotkin’s one by using explicit substitutions (constructors of theform let...in). In particular, Accattoli and Paolini [5] introduced recentlythe λvsub-calculus where the reduction rule acts at a distance by extending thenotion of βv-redex (with explicit substitutions). In this setting they give aninternal operational characterization of solvability and this characterization liftsto Herbelin and Zimmermann’s λCBV-calculus, another CBV λ-calculus withexplicit substitutions introduced in [9] (without rules acting at a distance butwith commutation rules for explicit substitutions).

Paolini and Ronchi Della Rocca [6,7] made major contributions to the studyof CBV-solvability through denotational semantics. In [6] they showed an inter-section type system that characterizes λv-potentially valuable3 (Thm. 6.4) andλv-solvable λ-terms (Thm. 6.5). We quote from [6, p. 28]: “The type assignmentsystem presented here is strongly related to the system presented in [13] forreasoning on the denotational semantics of the [Plotkin’s] λv-calculus. [. . . ] Thetwo systems have the same typability power”. It is not shown whether this typesystem is “legal” (see [7, Def. 10.1.5]), which is substantially a sufficient conditionto turn the type system into a filter model (i.e. a true domain model). In [7,Ch. 12] the same authors exhibit two models, V (§ 12.1) and VV (§ 12.2), bothbuilt from intersection type systems. The model V comes from a legal type systemand it is shown to be isomorphic to the one of [13]. All and only λv-potentiallyvaluable λ-terms have non trivial interpretation in V, but V gives only a partialsemantical characterization of λv-solvable λ-terms (Thm. 12.1.19). The modelVV characterizes observational equivalence (Thm. 12.2.14) but it is not a filtermodel. Recently, Ehrhard [14] used a relational model of the λv-calculus, basedon Linear Logic, to show that if the semantics of a λ-term M is not empty, thenM is strongly normalizing for the lazy βv-reduction (which does not reduce underabstractions); the converse is false (the aforesaid λ-term N is a counterexample).

The starting points of our work are [6,5,14]. We introduce the λσv -calculus, aCBV λ-calculus having the same syntax as ordinary (and hence Plotkin’s CBV)λ-calculus (there are no explicit substitutions) and extending the βv-reduction byadding two reduction rules, σ1 and σ3. For the λσv -calculus we give a semanticaland an internal operational characterization of solvability and potential valuability.We use the relational model of [14], which can also be seen as a model of ordinary

3 Following [6,7], a λ-term is λv-potentially valuable if there is a substitution sendingit (via βv ) into a λ-value. This notion is important for a CBV λ-calculus because ifwe want to manipulate some subterms, we need first to transform them into λ-values.

Page 3: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account of call-by-value solvability 3

λ-calculus (unlike the model V of [7]) and satisfies a version of the Taylor formula(see [14]). We also introduce a resource-sensitive calculus with tests in which theelements of the relational model are definable: this is a promising tool to facethe CBV full abstraction problem, along the lines of [15].

Our λσv -calculus springs from Girard’s call-by-value “boring” translation(·)v of λ-calculus into Intuitionistic Multiplicative Exponential Linear Logic(IMELL) proof-nets, identified by (A⇒ B)v = !Av ( !Bv (see [16]). The imagesof a σ1- or σ3-redex and its contractum under (·)v are equal modulo somespecified “immediate” steps of cut-elimination. Our σ-rules are related to (butpartly different from) Regnier’s σ-reduction defined in [17,18] for the ordinaryλ-calculus. Moreover, σ1 and σ3 correspond respectively to the commutationrules letapp and (a generalization of) let let in λCBV-calculus (see [9,5]). In somesense, they can be seen as a finer (and local) decomposition of the reductionrules acting at a distance in λvsub-calculus (it is possible to simulate λvsub- andλCBV-calculus in our λσv -calculus), but the absence of explicit substitutions inλσv -calculus prevents from lifting the internal operational characterization ofCBV-solvability from λvsub- or λCBV-calculus to our λσv -calculus.

Outline. In §2 we introduce our λσv -calculus. Then, §3, §4 and §5 are devoted tothe technical notions which are necessary in order to state our main results: in§3 we present two sub-reductions in the λσv -calculus, called w- and s-reduction;in §4 and §5 we present a resource-sensitive version of the λσv -calculus and therelational model of the (resource) λσv -calculus. In §6 we state and prove our maintheorems: the semantical (via the relational model) and syntactical (via w- ands-reductions) characterization of potential valuability and solvability; they sayalso that weak and strong normalizations coincide for both w- and s-reductions.

2 A CBV lambda-calculus with sigma-like-reductions

In this section we introduce λσv , our version of CBV λ-calculus. The syntax of λσvis the same as the one of ordinary λ-calculus. Given a countable set of variables(denoted by x, y, z, . . . ), the language of λσv is defined by the following grammar:

(Λv) V,U ::= x | λx.M λ-values(Λ) M,N,L ::= V | MN λ-terms

All λ-terms are considered up to α-conversion. The set of free variables ofa λ-term M is denoted by fv(M). Given pairwise distinct variables x1, . . . , xn,we denote by MV1/x1, . . . , Vn/xn the λ-term obtained by the capture-avoidingsimultaneous substitution of each free occurrence of xi in the λ-term M by theλ-value Vi (for 1 ≤ i ≤ n). Notice that, for all λ-values V, V1, . . . , Vn and pairwisedistinct variables x1, . . . , xn, V V1/x1, . . . , Vn/xn is a λ-value.

Contexts (with exactly one hole) are defined as usual via the grammar:

C ::= L·M | λx.C | CM | MC .

We use CLMM for the λ-term obtained by the capture-allowing substitution of theλ-term M for L·M in the context C.

Page 4: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

4 A. Carraro, G. Guerrieri

Definition 1. We define the following binary relations from Λ to Λ:

(λx.M)V 7→βvMV/x with V ∈ Λv

(λx.M)NL 7→σ1 (λx.ML)N with x 6∈ fv(L)V ((λx.L)N) 7→σ3 (λx.V L)N with x 6∈ fv(V ) and V ∈ Λv

For R ∈ βv, σ1, σ3, if M 7→R M′ then M is called R-redex.

We set 7→σ = 7→σ1∪ 7→σ3

and 7→v = 7→βv∪ 7→σ.

The side conditions on 7→σ in Def. 1 can be always fulfilled by α-renaming.

Notation. Let 7→R⊆ Λ× Λ. We use →R (called R-reduction) for the closure of7→R under all contexts; we denote by R (resp. →+

R) the reflexive-transitive(resp. transitive) closure of →R. Let M be a λ-term: M is R-normal if there isno λ-term N such that M →R N ; M is R-normalizable if there is a R-normal λ-term N such that M R N ; M is strongly R-normalizing if there is no sequence(Ni)i∈N such that M = N0 and Ni →R Ni+1 for every i ∈ N.

Notice that, for any λ-value V , if V →v M , then M is a λ-value.The λσv -calculus is the set Λ of λ-terms endowed with the v-reduction →v.

The set Λ endowed with→βvis Plotkin’s CBV λ-calculus ([4]) without constants.

Informally, σ-rules unblock βv-redexes which are hidden by the “hyper-sequential structure” of λ-terms. This approach is alternative to the one in[5] where hidden βv-redexes are reduced thanks to a rule acting at a distance.

Example. N = (λy.∆)(xI)∆→σ1(λy.∆∆)(xI)→βv

(λy.∆∆)(xI)→βv. . . is the

only possible v-reduction path from N : N is not v-normalizable but βv-normal.

2.1 Confluence of our CBV lambda-calculus

Our goal here is to prove that the v-reduction is confluent.

Proposition 2. The reduction →σ is strongly normalizing.Proof at p. 16

Proof. First, we define two sizes s(M) and #M by induction on the λ-term M :

s(x) = 2; #x = 1;

s(λx.M) = s(M) + 1; #λx.M = #M + s(M);

s(MN) = s(M) + s(N). #MN = #M + #N + 2s(M)s(N)− 1.

It is sufficient to show that if N →σ N′ then s(N) = s(N ′) and #N > #N ′. ut

Proposition 3. The reduction →σ is (not strongly) confluent.Proof at p. 16

Proof. By Newman’s Lemma and Prop. 2, it is sufficient to show that →σ islocally confluent. The proof of local confluence is by induction on M . The λ-termSee Remarks 41 and 42

at p. 18 Ξ = (λx.x′)((λy.y′I)(zI)

)(z′I) is an objection to strong confluence of →σ. ut

Lemma 4 (Hindley–Rosen, [1, p. 64]). Let →1,→2⊆ X2 (for any set X).If they are both confluent and they commute, i.e. if t 1 u1 and t 2 u2 thenthere exists s such that u1 2 s and u2 1 s, then →1 ∪ →2 is confluent.

Page 5: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account of call-by-value solvability 5

Lemma 5. Let M,M ′ ∈ Λ, V, V ′, V1, . . . , Vm ∈ Λv and R ∈ βv, σ, v.Proof at p. 17

(i) If V →R V′ then MV/xR MV ′/x.

(ii) If M →R M′ then MV1/x1, . . . , Vm/xm →R M

′V1/x1, . . . , Vm/xm.Lemma 6. The reductions →βv

and →σ commute. Proof at p. 18

Proof. It suffices to prove that if M →σ N1 and M →βv N2 then there is L s.t.N2 σ L and N1 →βv L. The proof of this statement is by induction on M . ut

By Lemmas 4 and 6, Prop. 3 and confluence of →βv (see [4]), we conclude:

Theorem 7. The reduction →v is (not strongly) confluent.

The λ-term Ξ (see proof of Prop. 3) is an objection to strong confluence of →v. See Remarks 41 and 42at p. 18If in the definition of 7→σ3

(Def. 1) we replace the λ-value V with any λ-term MSee Remark 43 at p. 18then →σ and →v are not (locally) confluent: consider (λx.x′)(zI)

((λy.y′)(z′I)

).

3 Weak and stratified CBV reductions

In this section we introduce two sub-reductions of →v: weak (or w-)reduction andstratified (or s-)reduction. We will show in §6 that they give an operational charac-terization of potential valuability and solvability: they are the “CBV counterpart”of head reduction for ordinary λ-calculus. Whereas head reduction is strictly de-terministic (any λ-term has at most one head redex), a λ-term might have several(overlapping) w- or s-redexes. Anyway, both w- and s-reductions are confluent(Prop. 10) and for them weak and strong normalization coincide (Thm. 24 and25). We have gathered our definition of w- and s-reductions from [5].

Definition 8. Weak and stratified contexts (denoted respectively by W and S)are contexts defined via the grammar:

W ::= L·M | WM | MW | (λx.W)M S ::= W | λx.S | SM

Notation. Let 7→R⊆ Λ × Λ: we use →w[R] (resp. →s[R]) for the closure underweak (resp. stratified) contexts of 7→R. We set w = w[v] and s = s[v]; for instance,→w =→w[v] (called w-reduction) and →s =→s[v] (called s-reduction).

Note that→w(→s(→v. In weak contexts, if the hole is under an abstractionthen this abstraction is the left-hand side of an application. Stratified contextsnever contain the hole under an abstraction which is in the right-hand side ofsome application, unless the abstraction is the left-hand side of an application.

Example. Let Ω = ∆∆: one has Ω →w Ω →w . . . , λy.Ω →s λy.Ω →s . . . , andx(λy.Ω)→v x(λy.Ω)→v . . . , whereas λy.Ω (resp. x(λy.Ω)) is w-(resp. s-)normal.

We will now prove that the w- and s-reductions are confluent.

Lemma 9. (i) The reductions →w[βv] and →s[βv ] are strongly confluent. Proof at p. 19(ii) The reductions →w[σ] and →s[σ] are confluent.(iii) The reductions →w[βv ] and →w[σ] (resp. →s[βv] and →s[σ]) commute.

By Lemmas 4 and 9 we can conclude:

Proposition 10. The reductions →w and →s are (not strongly) confluent.

The λ-term Ξ (see p. 4) is an objection to strong confluence of →w and →s.

Page 6: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

6 A. Carraro, G. Guerrieri

3.1 Characterization of w- and s-normal forms

Our goal here is to characterize w- and s-normal forms. Having no explicitsubstitutions, our characterization appears more concise than the one in [5].

Definition 11. We define the subsets anf, snf and wnf of Λ as follows:

(anf) Anf ::= xV | xAnf | AnfWnf

(wnf) Wnf ::= V | (λx.Wnf)Anf | Anf

(snf) Snf ::= x | λx.Snf | (λx.Snf)Anf | Anf

A β-redex is a λ-term of shape (λx.M)L. Notice that anf ( snf ( wnf and ifN ∈ anf then N has a free “head variable” and it is neither a value nor a β-redex.

Proposition 12. Let M be a λ-term.Proof at p. 20

(i) M is w-normal iff M ∈ wnf.(ii) M is s-normal iff M ∈ snf.(iii) M is w-(resp. s-)normal and is neither a value nor a β-redex iff M ∈ anf.

4 A resource CBV lambda-calculus

We now introduce the resource λσv -calculus, a valuable tool to prove some partsof our main results. It is an extension of the resource CBV λ-calculus introducedin [14, §5.2]. Its syntax is defined by the following grammar (the same as in [14]):

(rΛv) u, v ::= x | λx.t resource values(rΛt) s, t ::= st | [v1, . . . , vk] (k ≥ 0) resource terms(rΛ) e, f ::= v | t expressions

A resource term like [v1, . . . , vk] is a multiset of resource values (called bag).The resource-version of the βv-rule makes use of linear substitution, which

requires to enrich the syntax of the calculus with finite sets of resource terms.

Notation. Since the set Pf(A) of all finite subsets of a set A is the free module2〈A〉 generated by A over the boolean semiring 0, 1 with 1 + 1 = 1, we willuse algebraic notations for operations on its elements (+ for set unions, 0 for theempty set), as done in [15,14].

We denote by degx(e) the number of free occurrences of the variable x in theexpression e. Given e ∈ rΛ, v1, . . . , vk ∈ rΛv and an enumeration of the free occur-rences of variable x in e, if degx(e) = k then by

∑f∈Sk

evf(1)/x1, . . . , vf(k)/xkwe mean the sum of all expressions obtained by substituting vf(i) for the i-th freeoccurrence of x in e, as f varies over all elements of the set Sk of permutationsof 1, . . . , k. Finally, the linear substitution of [v1, . . . , vk] for x in e is

e〈[v1, . . . , vk]/x〉 =

∑f∈Sk

evf(1)/x1, . . . , vf(k)/xk if degx(e) = k

0 otherwise

Page 7: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account of call-by-value solvability 7

Notice that, for n ∈ v, t, if e ∈ rΛn then e〈[v1, . . . , vk]/x〉 ∈ 2〈rΛn〉.Resource contexts (with exactly one hole) are defined via the grammar:

R ::= L·M | Rt | tR | [λx.R, v1, . . . , vk] (k ≥ 0)

Let R be a resource context. We use RLtM for the resource term obtained bythe capture-allowing substitution of the resource term t for the hole L·M in R. IfT =

∑ni=1 ti (with t1, . . . , tn ∈ rΛt), then RLTM =

∑ni=1 RLtiM ∈ 2〈rΛt〉 (see also

[14, §5.2] and [15, §2.1]). For example, RL0M = 0 and [λx.[x]L[y][z] + [z][y]M, y] =[λx.[x]([y][z]), y] + [λx.[x]([z][y]), y].

Definition 13. We define the following binary relations from rΛt to 2〈rΛt〉:

[λx.t][v1, . . . , vk] 7→βv t〈[v1, . . . , vk]/x〉 [λx.t]ss′ 7→σ1 [λx.ts′]s if x /∈ fv(s′)

[v1, . . . , vn]t 7→0 0 if n 6= 1 [v]([λx.t]s) 7→σ3 [λx.[v]t]s if x /∈ fv(v)

We set 7→v = 7→βv ∪ 7→σ1 ∪ 7→σ3 ∪ 7→0.

According to the convention of §2,→v⊆ rΛt×2〈rΛt〉 is the reduction obtainedby resource-contextual closure of 7→v.

The resource λσv -calculus consists of the language rΛt and the reduction →v:it is the resource CBV λ-calculus of [14] plus the σ1- and σ3-rules.

As a technical simplification, we extend →v to a binary relation on 2〈rΛt〉 bylinearity, i.e. (

∑ni=1 ti) + S→v (

∑ni=1 Ti) + S iff ti →v Ti for every i = 1, . . . , n

(n ≥ 1). With this extension we can concisely state the following theorem:

Theorem 14. Reduction →v on 2〈rΛt〉 is strongly normalizing and confluent.

We omit the proof of Thm. 14. Strong normalization is evident (see [14] fora proof for the resource-contextual closure of 7→βv

∪ 7→0). The proof of localconfluence for the resource λσv -calculus is analogous to the one for v-reduction onλ-terms (see §2). Finally, confluence is obtained by Newman’s Lemma.

5 A relational model of (resource) CBV lambda-calculus

In this section we present a relational model for both the λσv -calculus and theresource λσv -calculus. This model is to be found in the category Rel of setsand relations (i.e. Rel(X,Y ) = P(X × Y )). In Rel identities are diagonalrelations and composition of morphisms is the standard composition of relations.This category has a symmetric monoidal structure given by 1 = 1 (arbitrarysingleton set) and X ⊗ Y = X × Y . This symmetric monoidal category is closed,with X ( Y = X ×Y , and ∗-autonomous with dualizing object ⊥ = 1. CategoryRel is cartesian, with X & Y = (1 ×X) ∪ (2 × Y ), and has an exponentialfunctor ! defined by !X = Mf(X) (the set of finite multisets on X) and !f =([α1, . . . , αn], [β1, . . . , βn]) : n ≥ 0, (αi, βi) ∈ f ∀ 1≤ i≤n for f ∈ Rel(X,Y ).

All this structure makes Rel a new-Seely category and hence a categoricalmodel of Linear Logic (LL). For more details we refer the reader to [19,14].

The model. We build inductively a family of sets (Un)n∈N given by U0 = ∅and Un+1 = Mf(Un) ×Mf(Un). Finally, we set U =

⋃n∈N Un. Notice that

Un ( Un+1 for all n ∈ N, and U =Mf(U)×Mf(U) = !U ( !U .

Page 8: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

8 A. Carraro, G. Guerrieri

5.1 Interpreting the CBV lambda-calculus

Using the fact that Rel has the structure of a LL model, we can give a concreteinterpretation of λ-terms as morphisms from Mf(U)n to Mf(U) in Rel (whereMf(U)n is the n-fold set-theoretic power of Mf(U)). This semantics can also bedescribed by type judgements (see [14]). With a ] b we indicate the union of the

multisets a and b (accounting for repetitions); if ~a and ~b are two finite sequences

(of the same length) of multisets, ~a ]~b is their component-wise union.

Definition 15. For every λ-term M and repetition-free list ~x ⊇ fv(M), wedefine, by induction on M , its interpretation JMK~x ⊆Mf(U)n ×Mf(U) (wheren is the length of ~x), as follows:

JxiK~x = (~a, ai) : ai ∈Mf(U), aj = [ ] for all 1 ≤ j ≤ n with j 6= i

Jλy.NK~x = (⊎ki=1 ~ai,

⊎ki=1[(bi, ci)]) : k≥0, ∀i = 1, . . . , k. ((~ai, bi), ci)∈JNK~x,y

JMNK~x = (~a0 ] ~a1, c) : ∃ b ∈Mf(U). (~a0, [(b, c)]) ∈ JMK~x, (~a1, b) ∈ JNK~x .

Notation. Hereafter, whenever we write JMK~x we suppose that ~x is a repetition-free list of variables containing fv(M). Moreover, we will sometimes silently usethe fact that JMK~x,y = ((~a, [ ]), b) : (~a, b) ∈ JMK~x whenever y 6∈ ~x.

Theorem 16 (soundness). Let M,N ∈ Λ. If M →v N , then JMK~x = JNK~x.Proof at p. 21

5.2 Interpreting the resource CBV lambda-calculus

In addition to the structure mentioned above, Rel is additive, and more preciselyits hom-sets are enriched over the category of complete lattices, with set-theoreticunion as join operation. The category Rel is a weak differential LL model (see[14]). Using this structure we can give the concrete interpretation of expressionsas morphisms from Mf(U)n to Mf(U) in Rel.

Definition 17. For every expression e and repetition-free list ~x ⊇ fv(e), wedefine, by induction on e, its interpretation JeK~x ⊆Mf(U)n ×Mf(U) (where nis the length of ~x), as follows:

JxiK~x = (~a, [α]) : α ∈ U, ai = [α], aj = [ ] for all 1≤j≤n with j 6= iJλz.tK~x = (~a, [(b, c)]) : ((~a, b), c) ∈ JtK~y,z

JstK~x = (~a0 ] ~a1, c) : ∃b ∈Mf(U). (~a0, [(b, c)]) ∈ JsK~x, (~a1, b) ∈ JtK~x

J[v1, . . . , vk]K~x = (⊎ki=1 ~ai,

⊎ki=1 bi) : k ≥ 0, ∀i = 1, . . . , k. (~ai, bi) ∈ JviK~x .

Finally, sums of expressions are interpreted by setting J∑ni=1 eiK~x =

⋃ni=1JeiK~x.

Notation. As for λ-terms, whenever we write JeK~x we suppose that ~x is arepetition-free list of variables containing fv(e), and similarly for the sums. Notethat J[ ]K~x = ([ ]n, [ ]) ⊆ Mf(U)n ×Mf(U), where [ ]n = ([ ], . . . , [ ]︸ ︷︷ ︸

n times

).

Page 9: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account of call-by-value solvability 9

Theorem 18 (soundness). Let S,T ∈ 2〈rΛt〉. If S→v T, then JSK~x = JTK~x.Proof at p. 21

The following notion of CBV Taylor expansion has been introduced in [14].

Definition 19 ([14], Taylor expansion). Given a λ-term M , we inductivelydefine a set T (M) of resource terms, called the Taylor expansion of M , as follows:

T (x) = [xn] : n ≥ 0 where [xn] = [

n times︷ ︸︸ ︷x, . . . , x]

T (λx.M) = [λx.t1, . . . , λx.tn] : n ≥ 0, ∀i. ti ∈ T (M)T (MN) = st : s ∈ T (M), t ∈ T (N) .

Theorem 20 ([14]). Let M be a λ-term. Then JMK~x =⋃t∈T (M)JtK~x .

Thm. 20 shows the semantical connection between λ-terms and their Taylorexpansion. In the next section (§6) it will be applied in Thm. 39.1, which is inturn a fundamental part of one of our main results Thm. 24.

Definition 21. For every expression e we define by induction the set strat(e) ofmultisets of resource values that occur in e in stratified position, as follows:

strat(x) = ∅; strat([v1, . . . , vn]) = [v1, . . . , vn] ∪⋃ni=1strat(vi) (n≥0);

strat(st) = strat(s); strat(λx.t) = strat(t) .

We set Strat = t ∈ rΛt : [ ] /∈ strat(t), whose elements are called stratifiedresource terms.

A stratified resource term t does not contain any [ ] in stratified position, i.e.every occurrence of [ ] in t is a subterm of some subterm of t in argument position.For instance: [x][ ], [x]([λz.[ ]][ ]) ∈ Strat but [ ], [ ][z], [λz.[ ][x, y]] /∈ Strat.

Stratified resource terms are not closed under v-reduction. For example, thestratified resource term [λx.[x]][λy.[ ]] v-reduces to the non-stratified [λy.[ ]].

Definition 22 (stratified Taylor expansion). Given a λ-term M , we defineits stratified Taylor expansion Ts(M) = t ∈ T (M) : if tv T, then T ⊆ Strat.

Example. The λ-term M = (λxy.x)Ω is neither w- nor s-normalizable and everyresource term in T (M) v-reduces to 0. Instead the non-s-normalizable (butw-normal) λ-term N = (λxy.Ω)(zz′) has infinitely many resource terms in T (N)that do not v-reduce to 0, like t = [λx.[ ]]([z][z′]) for example. However t 6∈ Ts(N)and Ts(N) contains only resource terms that v-reduce to 0, because all resourceterms in T (N) not v-reducing to 0 contain at least one [ ] in stratified position.

The semantical connection between λ-terms and their stratified Taylor expan-sion is illustrated in one of our main results, Thm. 25. In particular, Thm. 39.2 isthe step in which it is proved that the interpretation of Ts(M) actually witnessesthe strong s-normalization of M . Intuitively, if t ∈ Ts(M) then the v-normal formof t is a sum

∑ni=1 ti (n ≥ 0) of stratified resource terms, each of which does

not contain [ ] in stratified position: a subterm [ ] inside a ti does not “hide” anon-s-normalizable λ-term N such that M = SLNM. So, by Lemma 38.ii one canprove that if t 6= 0 then M is strongly s-normalizing.

Page 10: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

10 A. Carraro, G. Guerrieri

6 The main theorems

In this section we will present our main results: the semantical and internaloperational characterization of potential valuability (Thm. 24) and solvability(Thm. 25) for the λσv -calculus. See §1 for a overview of these notions.

Definition 23 (Potential valuability, solvability). Let M be a λ-term:

– M is potentially valuable if there exist variables x1, . . . , xm and λ-valuesV, V1, . . . , Vm (with m ≥ 0) such that MV1/x1, . . . , Vm/xmv V ;

– M is solvable if there exist variables x1, . . . , xm and λ-terms N1, . . . , Nn (forsome n,m ≥ 0) such that (λx1 . . . xm.M)N1 · · ·Nn v I.

We state now the two main theorems. In particular, Thm. 24 says thatw-normalizability (i.e. potential valuability) plays a role analogous to that ofhead-normalizability for many call-by-name models, like Scott’s D∞.

Theorem 24. Let M be a λ-term with ~x ⊇ fv(M). The following are equivalent:

(i) M is w-normalizable;(ii) M is potentially valuable;

(iii) JMK~x 6= ∅;(iv) M is strongly w-normalizing.

Theorem 25. Let M be a λ-term with ~x ⊇ fv(M). The following are equivalent:

(i) M is s-normalizable;

(ii) M is solvable;

(iii)⋃t∈Ts(M)JtK~x 6= ∅;

(iv) M is strongly s-normalizing.

An immediate corollary of Thm. 24 and 25 is that every solvable (i.e. s-normalizable) λ-term is also potentially valuable (i.e. w-normalizable).

The proofs of Thm. 24 and 25 are divided into parts, which are detailedseparately in the next subsections, due to the different techniques used for each oneof them. The splitting of the two proofs follows the same pattern. The implications(i)⇒ (ii) of both theorems are proved in §6.1 by purely syntactical means. Theimplication (ii)⇒ (iii) of Thm. 24 is shown in §6.2 using the resource λσv -calculusof §4; for this implication of Thm. 25 we use an extension of the resource λσv -calculus presented in §6.3. The implication (iii)⇒ (iv) of both theorems is provedin §6.4 using simulations of w- and s-reductions in λσv -calculus by the v-reductionof the resource λσv -calculus. Finally, (iv)⇒ (i) are trivial in both cases.

6.1 From weak and stratified normalization to solvability andpotential valuability

Our goal here is to prove the implication (i) ⇒ (ii) of Thm. 24 and 25. Ourapproach is largely inspired by [6,7,5].

For every n ∈ N, we set on = λxn . . . x0.x0. Notice that o0 = I and on is aclosed value for any n ∈ N. Moreover, onV 7→βv

on−1 for any n > 0 and V ∈ Λv.

Lemma 26. Let M ∈ wnf with fv(M) ⊆ x1, . . . , xm and let j ∈ N. Then thereProof at p. 21

exists h > 0 such that for all n1, . . . , nm ≥ j + h there exists a λ-term N suchthat Mon1/x1, . . . , o

nm/xmv λx.N and λx.N is closed.

Page 11: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account of call-by-value solvability 11

Lemma 27. Let M ∈ snf with fv(M) ⊆ x1, . . . , xm and let j ∈ N. Then thereProof at p. 22

exist h, k ∈ N such that for all n1, . . . , nm+k ≥ j + h there exists n ≥ j such thatMon1/x1, . . . , o

nm/xmonm+1 . . . onm+k v on.

Theorem 28. Let M be a λ-term.

1. [(i)⇒(ii) of Thm. 24] If M is w-normalizable then M is potentially valuable.2. [(i)⇒(ii) of Thm. 25] If M is s-normalizable then M is solvable.

Proof. For point 1 (resp. 2), hypothesis means that there is a w-(resp. s-)normalform M ′ such that M w M

′ (resp. M s M′), moreover M ′ ∈ wnf (resp. M ′ ∈

snf) by Prop. 12. Let fv(M) = x1, . . . , xm and thus fv(M ′) ⊆ x1, . . . , xm.

1. By Lemma 26 (taking j = 0) there exists h > 0 such that:M ′oh/x1, . . . , oh/xmv λx.N , for some closed λ-value λx.N . One hasMoh/x1, . . . , oh/xmv M

′oh/x1, . . . , oh/xm by Lemma 5.ii, so that Mis potentially valuable because λx.N is a closed λ-value.

2. By Lemma 27 (taking j = 0), there exist h, k, n ∈ N such that:(M ′oh/x1, . . . , oh/xm)oh . . . oh v o

n (oh is applied k times). We concludethat M is solvable because if we set H = (λx1 . . . xm.L·M) oh . . . oh︸ ︷︷ ︸

m+k times

I . . . I︸ ︷︷ ︸n times

, then

HLMMv HLM ′Mv (M ′oh/x1, . . . , oh/xm)oh . . . ohI . . . Iv o

nI . . . Iv I. ut

6.2 From potential valuability to non-emptyness

The following theorem proves the implication (ii) ⇒ (iii) of Thm. 24.

Theorem 29. Let M be a λ-term with ~x ⊇ fv(M). If M is potentially valuable,then JMK~x 6= ∅.

Proof. If M is potentially valuable (see Def. 23) there exist variables x1, . . . , xmand λ-values V, V1, . . . , Vm (for some m ≥ 0) s.t. MV1/x1, . . . , Vm/xmv V .Since variables are λ-values, we can suppose without loss of generality that~x = (x1, . . . , xm) ⊇ fv(M). Let ~y = fv(V ) ∪

⋃mi=1 fv(Vi). One can prove by

induction on M that See Lemma 48 at p. 24

JMV1/x1, . . . , Vm/xmK~y =

(⊎mi=1 ~ai, c) : ∃ b1, . . . , bm ∈Mf(U) :

((b1, . . . , bm), c) ∈ JMK~x, (~ai, bi) ∈ JViK~y for all 1 ≤ i ≤ m.

Since JV K~y 6= ∅ (this can be proved by simple inspection), by Thm. 16 we See Lemma 47 at p. 23

obtain that JMV1/x1, . . . , Vm/xmK~y 6= ∅ also holds, so that JMK~x 6= ∅. ut

6.3 From solvability to non-emptyness of stratified Taylor expansion

The implication (ii) ⇒ (iii) of Thm. 25 seems much more difficult to prove. Toaccomplish this task we introduce the resource λσv -calculus with tests, a CBV

Page 12: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

12 A. Carraro, G. Guerrieri

version of the resource calculus with tests defined in [15]. In this syntax allelements of the relational model are definable (see Def. 34).

The language extends that of resource λσv -calculus (see §4, p. 6) as follows:

(rΛv) u, v ::= x | λx.t resource values(rΛt) s, t ::= t ∗ p | st | [v1, . . . , vk] (k ≥ 0) resource terms(rΛτ ) p, q ::= τ [t1, . . . , tk] (k ≥ 0) tests

Note the overloaded use of rΛv and rΛt, which now (and until Lemma 36)indicate larger sets than those introduced in §4. We will use this extension toprove Lemma 36 (whose statement concerns only resource terms without tests).

Tests are – formally – multisets of resource terms, the “τ” being a tag fordistinguishing them from bags of values. Intuitively, they are constructions whichcan produce either success, represented by τ [ ], or failure, represented by 0.

Notation. We set ε = τ [ ] and τ [t1, . . . , tk] ‖ τ [tk+1, . . . , tn] = τ [t1, . . . , tn] (k≤n).

The test p ‖ q represents the (must-)parallel composition of p and q (i.e.,p ‖ q succeeds iff both p and q succeed). The composition is parallel in the sensethat the order of evaluation is inessential (remember that they are multisets).The binary operator ∗ allows to build a resource term out of a resource term anda test: intuitively, the resource term t ∗ p may be thought of as something thatoutputs the result of t only if p succeeds. Dually, the “cork construction” τ [t]may be thought of as a check that tests whether or not t v-reduces to [ ].

Resource, test-resource and test-test contexts (with exactly one hole), denotedresp. by R, Q and P, are defined by mutual induction via the grammar (k ≥ 0):

R ::= L·M | Rt | tR | t ∗ Q | [λx.R, v1, . . . , vk] (resource contexts);

Q ::= τ [R, t1, . . . , tk] (test-resource c.); P ::= L·M‖τ [t1, . . . , tk] (test-test c.).

Let t, t1, . . . , tn ∈ rΛt (resp. p, p1, . . . , pn ∈ rΛτ ). We use QLtM (resp. PLpM) forthe test obtained by the capture-allowing substitution of t (resp. p) for the holeL·M in Q (resp. P); similarly for RLtM (see p. 7). As usual, RL

∑i tiM =

∑i RLtiM,

QL∑i tiM =

∑i QLtiM and PL

∑i piM =

∑i PLpiM. E.g., t ∗ 0 = t ∗ QL0M = RL0M = 0.

Definition 30. The operational semantics of the resource λσv -calculus with testsextends the set of rules listed in Def. 13 with the following ones:

t(s ∗ p) 7→τ1 ts ∗ p τ [t ∗ p] 7→τ4 τ [t] ‖ p(t ∗ p)s 7→τ2 ts ∗ p τ [[v1, . . . , vn]] 7→τ5

ε if n = 0

0 otherwise(t ∗ p) ∗ q 7→τ3 t ∗ (p ‖ q)

We set 7→vτ = 7→v ∪ (⋃5i=1 7→τi) ⊆ (rΛt × 2〈rΛt〉) ∪ (rΛτ × 2〈rΛτ 〉). Then,

according to the convention of §2, →vτ ⊆ rΛτ × 2〈rΛτ 〉 is the reduction obtainedby test-contextual closure4 of 7→vτ . The resource λσv -calculus with tests consistsof the language rΛτ and the reduction →vτ .

4 This means that, for every p ∈ rΛτ and p′ ∈ 2〈rΛτ 〉, if p →vτ p′ then either there

exist a test-test context P, q ∈ rΛτ and q′ ∈ 2〈rΛτ 〉 such that p = PLqM, p′ = PLq′Mand q 7→τi q

′ with i ∈ 4, 5; or there exist a test-resource context Q, t ∈ rΛt andt′ ∈ 2〈rΛt〉 such that p = QLtM, p′ = QLt′M and t 7→vτ ′ t

′ with 7→vτ ′ = 7→v∪ (⋃3i=1 7→τi).

Page 13: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account of call-by-value solvability 13

As a technical simplification, we extend→vτ to a binary relation on 2〈rΛτ 〉 bylinearity, i.e., (

∑ni=1 qi)+P→vτ (

∑ni=1 Qi)+P iff qi →vτ Qi for every i = 1, . . . , n

(n ≥ 1). With this extension we can concisely state the following theorem:

Theorem 31. Reduction →vτ on 2〈rΛτ 〉 is strongly normalizing and confluent.

Definition 32. For every test p and repetition-free list ~x ⊇ fv(p), we define theinterpretation JpK~x ⊆ Mf(U)n × 1 of p, where n is the length of ~x, by mutualinduction with Def. 17 as follows:

JεK~x = ([ ]n, 1) Jp‖qK~x = (~a ]~b, 1) : (~a, 1)∈JpK~x, (~b, 1)∈JqK~x

Jτ [t]K~x = (~a, 1) : (~a, [ ])∈JtK~x Jt ∗ pK~x = (~a ]~b, c) : (~a, c)∈JtK~x, (~b, 1)∈JpK~x.

Finally, sums of tests are interpreted by setting J∑ni=1 piK~x =

⋃ni=1JpiK~x.

Theorem 33 (soundness). Let P,Q ∈ 2〈rΛτ 〉. If P→vτ Q, then JPK~x = JQK~x. Proof at p. 25

A key tool to connect the semantics with the vτ -reduction is the followingtransformation of elements of Mf(U) into resource terms and test contexts. Therole of this transformation is made clear in Lemma 35, used to prove Lemma 36.

Definition 34. Let c = [(a1, b1), . . . , (an, bn)] ∈Mf(U) (n ≥ 0). We define:

– the closed resource term c− = [λy1.b1−∗a1+L[ym1

1 ]M, . . . , λyn.bn−∗an+L[ymnn ]M],

where mi is the cardinality of the multiset ai (for i = 1, . . . , n);– the test-resource context c+ = τ [[λx.[ ]∗ ‖ni=1 τ [[λy.[ ] ∗ bi+L[yki ]M]([x]ai

−)]]L·M],where ki is the cardinality if the multiset bi (for i = 1, . . . , n).

Notation. For any a ∈Mf(U), #a indicates its cardinality. For ~a = (a1, . . . , an) ∈Mf(U)n and t ∈ rΛt, we write t〈~a−/~x〉 as a shorthand for t〈a1−/x1〉 · · · 〈an−/xn〉.

Lemma 35. Let (~a, b) ∈Mf(U)n ×Mf(U), k = #b and t ∈ rΛt with ~x ⊇ fv(t). Proof at p. 26

Then (~a, b) ∈ JtK~x iff τ [[λy.[ ] ∗ b+L[yk]M](t〈~a−/~x〉)]vτ ε.

Lemma 36. Let s and t be v-normal resource terms without tests (i.e., generatedby the grammar on §4, p. 6). If s ∈ Strat and t 6∈ Strat, then JsK~x ∩ JtK~x = ∅.

Proof. Let (~a, b) ∈Mf(U)n ×Mf(U) and QL·M = τ [[λy.[ ] ∗ b+L[yk]M](L·M〈~a−/~x〉)],with k = #b. One can prove by induction on the v-normal resource terms (withouttests) that: either QLtMvτ ε and QLsMvτ 0; or QLsMvτ ε and QLtMvτ 0; orQLsMvτ 0 and QLtMvτ 0. Hence, by Lemma 35, (~a, b) 6∈ JsK~x ∩ JtK~x. ut

Hereafter, when we will mention resource terms, we will refer to the oneswithout test (i.e., generated by the grammar on §4, p. 6).

The following theorem proves the implication (ii) ⇒ (iii) of Thm. 25.

Theorem 37. Let M be a λ-term and let ~x ⊇ fv(M). If M is solvable, then⋃t∈Ts(M)JtK~x 6= ∅.

Page 14: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

14 A. Carraro, G. Guerrieri

Proof. If M is solvable then there exists a context C = (λx1 . . . xm.L·M)N1 · · ·Nn(for some n,m ≥ 0) such that CLMMv I. By Thm. 16 and 20,

⋃t∈T (CLMM)JtK~x =

JCLMMK~x = JIK~x =⋃t∈T (I)JtK~x. Using Lemma 36 we infer that

⋃t∈Ts(CLMM)JtK~x =⋃

t∈Ts(I)JtK~x. Therefore⋃t∈Ts(CLMM)JtK~x 6= ∅ because it is easy to check that⋃

t∈Ts(I)JtK~x 6= ∅. By Thm. 18 and 14,⋃t∈Ts(CLMM)JtK~x 6= ∅ implies that there is a

resource term in Ts(CLMM) that v-reduces to a non-zero v-normal form. Now allresource terms in Ts(CLMM) are of the shape RLsM for some s ∈ Ts(M) (becausethe hole of C is in stratified position), so that if all resource terms in Ts(M)v-reduced to 0, then all resource terms in Ts(CLMM) would v-reduce to 0. Thus,there is t ∈ Ts(M) that v-reduces to a v-normal form T 6= 0. It is easy to provethat Jt′K~x 6= ∅ for every v-normal form t′, hence JtK~x = JTK~x 6= ∅ by Thm. 18. ut

6.4 From non-emptyness to strong normalization

Our goal here is to prove the implication (iii) ⇒ (iv) of Thm. 24 and 25.

Lemma 38. Let M,M ′ be λ-terms.Proof at p. 27

(i) If M →w M′ and t ∈ T (M), then there exists T ⊆ T (M ′) such that t→v T.

(ii) If M →s M′ and s ∈ Ts(M), then there exists S ⊆ Ts(M ′) such that s→+

v S.

Lemma 38.i is false if we replace the hypothesis M →w M′ with M →s M

′.For instance, take M = λx.Ω: then [ ] ∈ T (M) and M →s M , but [ ] is v-normal.

Theorem 39. Let M be a λ-term and let ~x ⊇ fv(M).

1. [(iii)⇒(iv) of Thm. 24] If JMK~x 6= ∅ then M is strongly w-normalizing.2. [(iii)⇒(iv) of Thm. 25] If

⋃t∈Ts(M)JtK~x 6= ∅, then M is strongly s-normalizing.

Proof. Let (~a, b) ∈ JMK~x (resp. (~a, b) ∈⋃t∈Ts(M)JtK~x). By Thm. 20 (resp. Then)

there exists t ∈ T (M) (resp. t ∈ Ts(M)) such that (~a, b) ∈ JtK~x. If M →w M′ (resp.

M →s M′), then by Lemma 38.i (resp. Lemma 38.ii) there exists T ⊆ T (M ′)

(resp. T ⊆ Ts(M ′)) such that t→+v T. According to Thm. 18, (~a, b) ∈ JTK~x, hence

T 6= ∅ and so there exists t′ ∈ T such that (~a, b) ∈ Jt′K~x. Therefore, if there wasan infinite reduction M →w M1 →w M2 →w . . . (resp. M →s M1 →s M2 →s . . . )then there would also be a an infinite reduction t→+

v T1 →+v T2 →+

v . . . , whichis impossible by Thm. 14. ut

Conclusions and future work

Our approach, that exploits the validity of the Taylor formula for a resource CBVλ-calculus, makes use of purely combinatorial proofs, rather than more standardapproaches based on reducibility or some specific machines. The interestingfeature of this approach is that it can be used for many different calculi alwaysusing a similar relational model and a suitable resource calculus.

We think that using the ordinary syntax of λ-calculus with our reduction willallow to develop a reasonable theory of CBV Bohm trees, never defined before(Paolini’s separability result in [20] for λv-calculus does not use Bohm trees),together with connections between equivalence of Bohm trees and observational

Page 15: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account of call-by-value solvability 15

equivalence. A future challenge is that of finding other fully abstract denotationalmodels, in view of Paolini and Ronchi Della Rocca’s proof of absence of fullyabstract filter models (see [7, Thm. 12.1.25]) built from legal type systems.

Another direction is relating two equivalence relations on λ-terms, the one gen-erated by our σ-rules and the one induced by Girard’s CBV “boring” translation(·)v of λ-calculus into IMELL proof-nets (along the lines of [17,18,21]).

References

1. Barendregt, H.P.: The Lambda Calculus: Its Syntax and Semantics. North-Holland,Amsterdam (1984)

2. Barendregt, H.: Solvability in lambda-calculi. J. Symb. Logic 75(3) (1975) 191–2313. Wadsworth, C.: The Relation Between Computational and Denotational Properties

for Scott’s D∞-Models of the λ-Calculus. SIAM J. Comput. 5(3) (1976) 488–5214. Plotkin, G.: Call-by-Name, Call-by-Value and the λ-Calculus. Theor. Comput. Sci.

1(2) (1975) 125–1595. Accattoli, B., Paolini, L.: Call-by-Value Solvability, Revisited. In: FLOPS. (2012)

4–166. Paolini, L., Ronchi Della Rocca, S.: Call-by-value Solvability. ITA 33(6) (1999)

507–5347. Paolini, L., Ronchi della Rocca, S.: The Parametric λ-calculus: a Metamodel for

Computation. Texts in Theoretical Computer Science. Springer, Berlin (2004)8. Paolini, L., Pimentel, E., Ronchi Della Rocca, S.: Strong Normalization from an

unusual point of view. Theor. Comput. Sci. 412(20) (2011) 1903–19159. Herbelin, H., Zimmermann, S.: An operational account of Call-by-Value Minimal

and Classical λ-calculus in ”Natural Deduction” form. In: TLCA. (2009) 142–15610. Hofmann, M.: Sound and Complete Axiomatisations of Call-by-Value Control

Operators. Mathematical Structures in Computer Science 5(4) (1995) 461–48211. Moggi, E.: Computational Lambda-Calculus and Monads. In: LICS. (1989) 14–2312. Dyckhoff, R., Lengrand, S.: Call-by-Value lambda-calculus and LJQ. J. Log.

Comput. 17(6) (2007) 1109–113413. Egidi, L., Honsell, F., Ronchi Della Rocca, S.: Operational, denotational and logical

descriptions: a case study. Fundamenta Informaticæ 16(2) (1992) 149–16914. Ehrhard, T.: Collapsing non-idempotent intersection types. In: CSL. (2012) 259–27315. Bucciarelli, A., Carraro, A., Ehrhard, T., Manzonetto, G.: Full Abstraction for

the Resource Lambda Calculus with Tests, through Taylor Expansion. LogicalMethods in Computer Science 8(4) (2012)

16. Girard, J.: Linear Logic. Theor. Comput. Sci. 50(1) (1987) 1–10217. Regnier, L.: Lambda calcul et reseaux. PhD thesis, Universite Paris 7 (1992)18. Regnier, L.: Une equivalence sur les lambda-termes. Theor. Comput. Sci. 126(2)

(1994) 281–29219. Mellies, P.: Categorical semantics of Linear Logic. In: Interactive models of

computation and program behaviour. Volume 27 of Panoramas et Syntheses. SocieteMathematique de France (2009) 1–196

20. Paolini, L.: Call-by-Value Separability and Computability. In: ICTCS. (2001) 74–8921. Accattoli, B., Kesner, D.: The Permutative λ-Calculus. In: LPAR. (2012) 23–36

Page 16: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

16 A. Carraro, G. Guerrieri

A Technical appendix

A.1 Proofs and remarks of Section 2

Proposition 2. The reduction →σ is strongly normalizing.Stated at p. 4

Proof. First we define two sizes s(M) and #M by induction on the λ-term M :

– s(x) = 2;– s(λx.M) = s(M) + 1;– s(MN) = s(M) + s(N).

– #x = 1;– #λx.M = #M + s(M);– #MN = #M + #N + 2s(M)s(N)− 1.

Notice that s(N) ≥ 2 and #N ≥ 1 for any λ-term N .In order to prove that σ-reduction is strongly normalizing, it suffices to show

that if M →σ M′ then s(M) = s(M ′) and #M > #M ′. We proceed by induction

on the definition of M →σ M′.

If M 7→σ1 M′ then M = (λx.M0)NL and M ′ = (λx.M0L)N , hence

#M = #M0 + #L+ #N + s(M0) + 2 s(N) + 2 s(L)

+ 2 s(M0)s(N) + 2 s(M0)s(L) + 2 s(L)s(N)− 2

#M ′ = #M0 + #L+ #N + s(M0) + 2 s(N) + s(L)

+ 2 s(M0)s(N) + 2 s(M0)s(L) + 2 s(L)s(N)− 2

so #M = #M ′+s(L) > #M ′. Moreover s(M) = s(M0)+s(L)+s(N)+1 = s(M ′).If M 7→σ3

M ′ then M = V ((λxL)N) and M ′ = (λx.V L)N , hence

#M = #V + #L+ #N + 2 s(V ) + 2 s(N) + s(L)

+ 2 s(V )s(N) + 2 s(V )s(L) + 2 s(L)s(N)− 2

#M ′ = #V + #L+ #N + s(V ) + 2 s(N) + s(L)

+ 2 s(V )s(N) + 2 s(V )s(L) + 2 s(L)s(N)− 2

so #M = #M ′+ s(V ) > #M ′. Moreover s(M) = s(V ) + s(L) + s(N) + 1 = s(M ′).If M = λx.N and M ′ = λx.N ′ with N →σ N

′, then #M = #N + s(N) and#M ′ = #N ′ + s(N ′); by induction hypothesis, #N > #N ′ and s(N) = s(N ′),hence #M > #N ′ + s(N) = #M ′ and s(M) = s(N) + 1 = s(M ′).

If M = LN and M ′ = L′N (resp. M = NL and M ′ = NL′) with L→σ L′,

then #M = #L+ #N + 2 s(L)s(N)− 1 and #M ′ = #L′+ #N + 2 s(L′)s(N)− 1;by induction hypothesis, #L > #L′ and s(L) = s(L′), hence #M > #L+ #N +2 s(L′)s(N)− 1 = #M ′ and s(M) = s(L) + s(N) = s(M ′). ut

Proposition 3. The reduction →σ is confluent.Stated at p. 4

Proof. By Newman’s lemma and Prop. 2, it suffices to show that →σ is locallyconfluent: if M →σ N1 and M →σ N2 then there is M ′ s.t. N1 σ M

′ andN2 σ M

′. We proceed by induction on M , the only interesting cases are:

– if M = (λx.M0)((λy.L)N

)L′ with M →σ1

(λx.M0L′)((λy.L)N

)= N1 and

M →σ3 (λy.(λx.M0)L)NL′ = N2, then N2 →σ1 (λy.(λx.M0)LL′)N →σ1

(λy.(λx.M0L′)L)N = M ′ and N1 →σ3 M

′;

Page 17: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account of call-by-value solvability 17

– if M = V((λx.L)((λx′.L′)N)

)with M →σ3 V

((λx′.(λx.L)L′)N

)= N1 and

M →σ3(λx.V L)

((λx′.L′)N

)= N2, then N1 →σ3

(λx′.V ((λx.L)L′)

)N →σ3

(λx′.(λx.V L)L′)N = M ′ and N2 →σ3M ′. ut

Lemma 5. Let M,M ′ ∈ Λ, V, V ′, V1, . . . , Vm ∈ Λv and R ∈ βv, σ, v. Stated at p. 5

(i) If V →R V′ then MV/xR MV ′/x.

(ii) If M →R M′ then MV1/x1, . . . , Vm/xm →R M

′V1/x1, . . . , Vm/xm.

Proof. For R = βv, the proofs of (i) and (ii) are in [4]. For R = v, the proof of(i) (resp. (ii)) is a consequence of the property (i) (resp. (ii)) for both →βv

and→σ, since →v =→βv

∪ →σ. Let us prove (i) and (ii) for R = σ.

(i) The proof is by induction on the λ-term M . If M = x, then MV/x = Vand MV ′/x = V ′, so MV/xσ MV ′/x by hypothesis.If M = y 6= x, then MV/x = y = MV ′/x, then MV/xσ MV ′/xby reflexivity of σ.If M = λy.N for some λ-term N , then we can suppose without loss ofgenerality that y /∈ fv(V ) ∪ x, so MV/x = λy.NV/x and MV ′/x =λy.NV ′/x; by induction hypothesis, NV/xσ NV ′/x and thereforeMV/xσ MV ′/x.If M = LN for some λ-terms L and N , then MV/x = LV/xNV/xand MV ′/x = LV ′/xNV ′/x; LV/x σ LV ′/x and NV/x σ

NV ′/x by induction hypothesis, thus MV/xσ LV ′/xNV/xσ

MV ′/x.(ii) The proof is by induction on the definition of M →σ M

′. For any λ-term Nwe set N∗ = NV1/x1, . . . , Vm/xm.If M 7→σ1 M

′ then M = (λy.M0)NL and M ′ = (λy.M0L)N with y /∈ fv(L);we can suppose without loss of generality that y /∈

⋃mi=1(fv(Vi) ∪ xi), so

M∗ = (λy.M∗0 )N∗L∗ and M ′∗ = (λy.M∗0L∗)N∗, hence M∗ →σ1

M ′∗ sincey /∈ (fv(L) \ x1, . . . , xm) ∪

⋃mi=1 fv(Vi) = fv(L∗).

If M 7→σ3 M′ then M = V ((λy.L)N) and M ′ = (λy.V L)N with y /∈ fv(V );

we can suppose without loss of generality that y /∈⋃mi=1(fv(Vi) ∪ xi), thus

M∗ = V ∗((λy.L∗)N∗) and M ′∗ = (λy.V ∗L∗)N∗, so M∗ →σ3M ′∗ since

y /∈ (fv(V ) \ x1, . . . , xm) ∪⋃mi=1 fv(Vi) = fv(V ∗).

If M = λy.N and M ′ = λy.N ′ with N →σ N′ then we can suppose without

loss of generality that y /∈⋃mi=1(fv(Vi) ∪ xi), hence M∗ = λy.N∗ and

M ′∗ = λy.N ′∗; by induction hypothesis, N∗ →σ N′∗ and so M∗ →σ M

′∗.If M = LN and M ′ = L′N (resp. M = NL and M ′ = NL′) with L→σ L

′,then M∗ = L∗N∗ and M ′∗ = L′∗N∗ (resp. M∗ = N∗L∗ and M ′∗ = N∗L′∗);by induction hypothesis, L∗ →σ L′∗, so M∗ →σ M

′∗. ut

Lemma 40. Let →1,→2⊆ X2 (for any set X) be such that if t →1 u1 andt→2 u2 then there is v ∈ X such that u2 1 v and u1 →2 v. Then they commute(i.e. if t1 u1 and t2 u2 then there is s ∈ X such that u1 2 s and u2 1 s).

Proof. For every t, u ∈ X, →⊆ X2 and n ∈ N, we write t →n u if there existv0, . . . , vn ∈ X such that t = v0, u = vn and vi → vi+1 for all 0 ≤ i ≤ n. We prove

Page 18: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

18 A. Carraro, G. Guerrieri

the following stronger statement, in order to apply the right induction hypothesis:if t1 u1 and t→m

2 u2 then there exists t′ ∈ X such that u2 1 t′ and u1 →m

2 t′.Let t→n

1 u1: the proof is by induction on (m,n) with the lexicographical orderon N2. If m = 0 or n = 0, we conclude easily.

Let m,n > 0: there exist u′1, u′2 ∈ X such that t→1 u

′1, t→2 u

′2, u′1 →n−1

1 u1and u′2 →m−1

2 u2. By hypothesis, there is v ∈ X such that u′1 →2 v and u′2 1 v.By induction hypothesis applied to u′2, there is v′ ∈ X such that u2 1 v

′ andv m−1

2 v′; thus u′1 →m2 v′, so there exists s ∈ X such that v′ 1 s and u1 →m

2 sby applying the induction hypothesis to u′1, therefore u2 1 s. ut

Lemma 6. The reductions →βvand →σ commute.Stated at p. 5

Proof. By lemma 40, it suffices to prove that if M →σ N1 and M →βvN2 then

there is L such that N2 σ L and N1 →βv L. The proof of this property is byinduction on M . The only interesting cases are:

– if M = (λx.N)V L′ with M →σ1 (λx.NL′)V = N1 and M →βv NV/xL′ =N2, then N1 →βv N2 since x /∈ fv(L′).

– ifM = U((λx.N)V ) withM→σ3(λx.UN)V = N1 andM→βv

U(NV/x) =N2, then N1 →βv

N2 since x /∈ fv(U).– if M = (λy.N ′)

((λx.N)V

)L′ with M →σ1

(λy.N ′L′)(λx.N)V = N1 andM →βv

(λy.N ′)NV/xL′ = N2, then N1→βv(λy.N ′L′)NV/x = L and

N2 →σ1L.

– if M = (λx.N)V with M →σ N1 = (λx.N)V ′ (resp. N1 = (λx.N ′)V ),M →βv

NV/x = N2 and V →σ V′ (resp. N →σ N

′), then we conclude byLemma 5.i (resp. 5.ii). ut

Remark 41. The reductions →σ and →v are not strongly confluent. For in-stance take N = zI, L = z′I and M = (λx.x′)

((λy.y′I)N

)L: one has M →σ1

(λx.x′L)((λy.y′I)N

)= N3 and M →σ3

(λy.(λx.x′)(y′I)

)NL = N1 (with N1 6=

N3) but for i ∈ 1, 3 the only v-redex in Ni is the σi-redex Ni itself; thereforeN1 →σ1

(λy.(λx.x′)(y′I)L

)N 6=

(λy.(λx.x′L)(y′I)

)N σ3

← N3.

Remark 42. If we define a Tait–Martin-Lof parallel reduction →ρ of →v inthe obvious way, then →ρ is not strongly confluent. For instance, take M =(λx.L)

((λy.N)((λz.N ′)N ′′)

)L′, M1 = (λx.LL′)

((λy.N)((λz.N ′)N ′′)

)and M2 =(

(λy.(λx.L)N)((λz.N ′)N ′′))L′: then M →ρ M1 and M →ρ M2 but there is no

λ-term M ′ such that M1 →ρ M′ρ←M2. Informally, →ρ is not able to reduce in

one step several “subsequent” σ1-redexes created by one σ3-step. Therefore, wecannot adapt the Tait–Martin-Lof technique in a natural way in order to provethat →v is confluent.

Remark 43. If in definition of 7→σ3(Def. 1) we replace the λ-value V with

any λ-term M then →σ and →v are not (locally) confluent. For instance, takeNi = (zi)I for i ∈ 1, 2 and M = (λx1.y1)N1

((λx2.y2)N2

): one would have

M →σ3(λx2.(λx1.y1)N1y2)N2 = L1 and M →σ1

(λx1.y1((λx2.y2)N2)

)N1 = L3

(with L1 6= L3) but the only v-redex in L1 (resp. L3) is the σ1-redex (λx1.y1)N1y2(resp. σ3-redex y1((λx2.y2)N2)); so L1 →σ1

(λx2.(λx1.y1y2)N1

)N2 = N12 and

L3 →σ3

(λx1.(λx2.y1y2)N2

)N1 = N21, with N12 6= N21 and N12, N21 v-normal.

Page 19: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account of call-by-value solvability 19

A.2 Proofs and remarks of Section 3

Lemma 44. Let M,M ′ ∈ Λ, V1, . . . , Vm ∈ Λv and R ∈ w[βv], w[σ], s[βv], s[σ].If M →R M

′ then MV1/x1, . . . , Vm/xm →R M′V1/x1, . . . , Vm/xm.

Proof. All the proofs are by induction on the definition of M →R M′.

For R ∈ w[σ], s[σ], the proof is analogous to that one for Lemma 5.ii.For R = w[βv] (for any λ-term N we set N∗ = NV1/x1, . . . , Vm/xm):

– If M 7→βv M ′ then M = (λy.N)V and M ′ = NV/y, moreover we cansuppose without loss of generality that y /∈

⋃mi=1(fv(Vi) ∪ xi); hence,

M∗ = (λy.N∗)V ∗ →βvN∗V ∗/y = M ′∗ (since V ∗ is a value).

– If M = LN and M ′ = L′N (resp. M = NL and M ′ = NL′) with L→w[βv ] L′,

then M∗ = L∗N∗ and M ′∗ = L′∗N∗ (resp. M∗ = N∗L∗ and M ′∗ = N∗L′∗);by induction hypothesis, L∗ →w[βv] L

′∗, so M∗ →w[βv] M′∗.

– If M = (λy.N)L and M ′ = (λy.N ′)L with N →w[βv] N′, moreover we can

suppose without loss of generality that y /∈⋃mi=1(fv(Vi)∪xi); hence, M∗ =

(λy.N∗)L∗ and M ′∗ = (λy.N ′∗)L∗; by induction hypothesis, N∗ →w[βv ] N′∗

and thus M∗ →w[βv] M′∗.

For R = s[βv] (for any λ-term N we set N∗ = NV1/x1, . . . , Vm/xm):

– If M →w[βv ] M′ then we have just proved that M∗ →w[βv] M

′∗ and soM∗ →s[βv ] M

′∗ (since →w[βv ]⊆→s[βv ]).– If M = LN and M ′ = L′N (resp. M = λy.L and M ′ = λy.L′) withL →s[βv] L

′, then M∗ = L∗N∗ and M ′∗ = L′∗N∗ (resp. M∗ = λy.L∗ andM ′∗ = λy.L′∗, by supposing without loss of generality that y /∈

⋃mi=1(fv(Vi)∪

xi)); by induction hypothesis, L∗ →s[βv ] L′∗, so M∗ →s[βv ] M

′∗. ut

Remark 45. There are no λ-values V and V ′ such that V →w V′.

For R ∈ s[βv], s[σ], if V →R V′ then V = λz.N , V ′ = λz.N ′ and N →R N

′.

Remark 46. For R ∈ s[βv], s[σ], the analogous of Lemma 5.i does not hold:V →R V ′ does not imply MV/x R MV ′/x. For instance, take M = yx:by Remark 45, for any λ-value V , MV/x = yV is R-normal.

Lemma 9. (i) The reductions →w[βv] and →s[βv ] are strongly confluent. Stated at p. 5

(ii) The reductions →w[σ] and →s[σ] are confluent.(iii) The reductions →w[βv ] and →w[σ] (resp. →s[βv] and →s[σ]) commute.

Proof.

(i) Let R ∈ w[βv], s[βv]. We prove by induction on M ∈ Λ that if M →R N1

and M →R N2 then there is M ′ ∈ Λ such that N1 →R M′ and N2 →R M

′.The only interesting case is when M = (λx.N)V with M →R NV/x = N1,M →R (λx.N ′)V = N2 and N →R N

′: N2 →R N′V/x and, by Lemma 44,

N1 →R N′V/x.

Page 20: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

20 A. Carraro, G. Guerrieri

(ii) By Newman’s lemma and Prop. 2 (since →w[σ],→s[σ]⊆→σ), it suffices toshow that →w[σ] and →s[σ] are locally confluent. The proof that →w[σ] and→s[σ] are locally confluent are analogous to that one for the local confluenceof →σ seen in Prop. 3.

(iii) The proof is analogous to that one for Lemma 6 (in particular, we useLemma 40). The only notable difference is that here there is not a case where,for R ∈ w, s, M = (λx.N)V , M →R[βv] NV/x = N2 and M →R[σ]

(λxN)V ′ = N1 with V →R[σ] V′ (see Remarks 45 and 46). ut

Notice that every s-normal forms is also a w-normal form, since →w⊆→s.Obviously, every βv-redex is also a β-redex (a λ-term of the form (λx.M)N).

Proposition 12. Let M be a λ-term.Stated at p. 6

(i) M is w-normal iff M ∈ wnf.

(ii) M is s-normal iff M ∈ snf.

(iii) M is w-(resp. s-)normal and is neither a value nor a β-redex iff M ∈ anf.

Proof.

⇒: We prove simultaneously the left-to-right part of the three statements, byinduction on the λ-term M .

If M is a λ-value then M ∈ wnf. Furthermore, if M is a variable then M ∈ snf;if M = λx.N is s-normal (for some λ-term N) then N is s-normal, henceN ∈ snf by induction hypothesis, and so M ∈ snf.

If M is not a λ-value then M = M1M2 for some λ-terms M1 and M2. Bysimple inspection of the definition of →w (resp. →s), the fact that M isw-(resp. s-)normal implies that M1 is w-(resp. s-)normal and M2 is w-normal,moreover M1 is not a β-redex (otherwise M would be a σ1-redex) and M isneither a βv- nor a σ3-redex. There are only three cases:

1. M1 is not a value: by induction hypothesis M1 ∈ anf and M2 ∈ wnf,therefore M ∈ anfwnf ⊆ anf.

2. M1 = λx.N : then M2 is neither a β-redex (otherwise M would be aσ3-redex) nor a value (otherwise M would be a βv-redex), so M2 ∈ anf byinduction hypothesis. Moreover, the fact that M1 is w-(resp. s-)normalentails that N is w-(resp. s-)normal and thus N ∈ wnf (resp. N ∈ snf),by induction hypothesis. Hence M ∈ anf.

3. M1 is a variable: if M2 is a value then M ∈ anf; if M2 is not a valuethen M2 is not a β-redex (otherwise M would be a σ3-redex) and thusM2 ∈ anf by induction hypothesis, therefore M ∈ anf.

⇐: The proof right-to-left part of the statement (i) (resp. (ii)) is by induction onM ∈ wnf (resp. M ∈ snf). The right-to-left part of the statement (iii) is animmediate consequence of (i) and (ii), since anf ⊆ snf ⊆ wnf and if M ∈ anfthen M is neither a value nor a β-redex. ut

Page 21: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account of call-by-value solvability 21

A.3 Proofs of Section 5

Theorem 16 (soundness). Let M,N ∈ Λ. If M →v N , then JMK~x = JNK~x.Stated at p. 8

Proof. For soundness w.r.t. the βv-rule we refer to [14] (see also Lemma 48).Regarding the σ-rules we have:

J(λy.M)NLK~x = (~a0 ] ~a1 ] ~a2, c) : ∃b ∈Mf(U). ∃d ∈Mf(U).

((~a1, d), [(b, c)]) ∈ JMK~x,y, (~a0, b) ∈ JLK~x, (~a2, d) ∈ JNK~x= J(λy.ML)NK~x

which validates the rule σ1, and

JV ((λy.L)N)K~x = (~a0 ] ~a1 ] ~a2, c) : ∃b ∈Mf(U). ∃d ∈Mf(U).

(~a0, [(b, c)]) ∈ JV K~x, ((~a1, d), b) ∈ JLK~x,y, (~a2, d) ∈ JNK~x= J(λy.V L)NK~x

which validates the rule σ3. Finally it is easy to check that the interpretation iscontextual. ut

Theorem 18 (soundness). Let S,T ∈ 2〈rΛt〉. If S→v T, then JSK~x = JTK~x. Stated at p. 9

Proof. For soundness w.r.t. the βv-rule and 0-rule, we refer to [14].Regarding the σ-rules we have:

J[λy.t]ss′K~x = (~a0 ] ~a1 ] ~a2, c) : ∃b ∈Mf(U). ∃d ∈Mf(U).

((~a1, d), [(b, c)]) ∈ JtK~x,y, (~a0, b) ∈ Js′K~x, (~a2, d) ∈ JsK~x= J[λy.ts′]sK~x

which validates the rule σ1, and

J[v]([λy.t]s)K~x = (~a0 ] ~a1 ] ~a2, c) : ∃b ∈Mf(U). ∃d ∈Mf(U).

(~a0, [(b, c)]) ∈ JvK~x, ((~a1, d), b) ∈ JtK~x,y, (~a2, d) ∈ JsK~x= J[λy.[v]t]sK~x

which validates the rule σ3. Finally it is easy to check that the interpretation iscontextual. ut

A.4 Proofs of Subsection 6.1

Lemma 26. Let M be a λ-term with fv(M) ⊆ x1, . . . , xm and let j ∈ N. Stated at p. 10

– If M ∈ anf then there exists h ∈ N∗ such that for all n1, . . . , nm ≥ j + h onehas Mon1/x1, . . . , o

nm/xmv ok for some k ≥ j.

– If M ∈ wnf then there is h ∈ N∗ such that for all n1, . . . , nm ≥ j+h one hasMon1/x1, . . . , o

nm/xmv λx.N for some λ-term N s.t. λx.N is closed.

Page 22: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

22 A. Carraro, G. Guerrieri

Proof. By mutual induction on M ∈ anf and M ∈ wnf. Notice that if M ∈ anfthen fv(M) 6= ∅ and thus m > 0.

If M = xV for some variable x and λ-value V , then x = xi for some1 ≤ i ≤ m. Let n1, . . . , nm ≥ j + 1. One has Mon1/x1, . . . , o

nm/xm =oniV on1/x1, . . . , o

nm/xm →v oni−1 (since V on1/x1, . . . , onm/xm is a λ-

value) where ni − 1 ≥ j. Hence we conclude by taking h = 1 and k = ni − 1.

If M = xN for some variable x and N ∈ anf, then x = xi for some 1 ≤ i ≤ m.By induction hypothesis there exists h ∈ N∗ such that for all n1, . . . , nm ≥j + h one has Non1/x1, . . . , o

nm/xm v ok for some k ≥ j. Hence for alln1, . . . , nm ≥ j + h one has Mon1/x1, . . . , o

nm/xmv oniok →v o

ni−1 whereni − 1 ≥ j + h− 1 ≥ j.

If M = N ′N ′′ for some N ′ ∈ anf and N ′′ ∈ wnf, then by induction hypothesisthere are h′, h′′ ∈ N∗ s.t. for all n′1, . . . , n

′m ≥ j + 1 +h′ and n′′1 , . . . , n

′′m ≥ j +h′′

one has N ′on′1/x1, . . . , o

n′m/xmv o

k and N ′′on′′1 /x1, . . . , o

n′′m/xmv λx.L

for some k ≥ j+1 and λ-term L such that λx.L is closed. Let h = maxh′+1, h′′:for all n1, . . . , nm ≥ j + h one has Non1/x1, . . . , o

nm/xmv okλx.L→v o

k−1

where k − 1 ≥ j.If M is a variable then M = xi for some 1 ≤ i ≤ m, hence for all n1, . . . , nm ≥

j + 1 one has Mon1/x1, . . . , onm/xm = oni which is a closed abstraction.

If M = λx.N for some λ-term N , then we can suppose without loss ofgenerality that x 6= xi for any 1 ≤ i ≤ m, so Mon1/x1, . . . , o

nm/xm =λx.Non1/x1, . . . , o

nm/xm which is closed because the oni ’s are closed andfv(N) ⊆ x, x1, . . . , xn.

If M = (λx.N ′)N ′′ for some λ-terms N ′ ∈ wnf and N ′′ ∈ anf then we cansuppose without loss of generality that x 6= xi for any 1 ≤ i ≤ m, moreoverfv(N ′) ⊆ x, x1, . . . , xm. By induction hypothesis, there exist h′, h′′ ∈ N∗ suchthat for all n′, n′1, . . . , n

′m ≥ j + h′ and n′′1 , . . . , n

′′m ≥ j + h′ + h′′ one has

N ′on′/x, on

′1/x1 . . . , o

n′m/xmv λxL and N ′′on

′′1 /x1, . . . , o

n′′m/xmv o

k

for some k ≥ j + h′ and λ-term L s.t. λx.L is closed. If h = h′ + h′′ then, for alln1, . . . , nm ≥ j+h, Mon1/x1, . . . , o

nm/xmv (λx.N ′on1/x1, . . . , onm/xm)ok

→v N′ok/x, on1/x1 . . . , o

nm/xmv λx.L. ut

Lemma 27. Let M ∈ snf with fv(M) ⊆ x1, . . . , xm and let j ∈ N. Then thereStated at p. 11

exist h, k ∈ N such that for all n1, . . . , nm+k ≥ j + h there exists n ≥ j such that

Mon1/x1, . . . , onm/xmonm+1 . . . onm+k v o

n.

Proof. By induction on M ∈ snf.

If M is a variable then M = xi with 1 ≤ i ≤ m, so Mon1/x1, . . . , onm/xm=

oni where ni ≥ j, hence we conclude by taking h = 0 = k.

If M = λx.N for some N ∈ snf then we can suppose without loss of generalitythat x 6= xi for any 1 ≤ i ≤ m, moreover fv(N) ⊆ x, x1, . . . , xm. By inductionhypothesis, there exist h, k′ ∈ N such that for all n′, n1, . . . , nm+k ≥ j+h one has

Page 23: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account of call-by-value solvability 23

(Non′/x, on1/x1 . . . , o

nm/xm)onm+1 . . . onm+k′ v on for some n ≥ j. Hence

(Mon1/x1, . . . , onm/xm)on

′onm+1 . . . onm+k′ =

(λxNon1/x1, . . . , onm/xm)on

′onm+1 . . . onm+k′ →v

(Non′/x, on1/x1 . . . , o

nm/xm)onm+1 . . . onm+k′ v on

where n ≥ j, thus we conclude by taking k = k′ + 1.

If M = (λx.N ′)N ′′ for some λ-terms N ′ ∈ snf and N ′′ ∈ anf then we cansuppose without loss of generality that x 6= xi for any 1 ≤ i ≤ m, moreoverfv(N ′) ⊆ x, x1, . . . , xm. By induction hypothesis, there exist h′, k ∈ N suchthat for all n′, n′1, . . . , n

′m+k ≥ j + h′ one has

(N ′on′/x, on

′1/x1 . . . , o

n′m/xm)on

′m+1 . . . on

′m+k v o

n

for some n ≥ j. By lemma 26 there exists h′′ ∈ N∗ such that for all n′′1 , . . . , n′′m ≥

j + h′ + h′′ one has N ′′on′′1 /x1, . . . , o

n′′m/xmv o

n′′for some n′′ ≥ j + h′. Let

h = h′ + h′′: for all n1, . . . , nm+k′ ≥ j + h one has (where n ≥ j)

(Mon1/x1, . . . , onm/xm)onm+1 . . . onm+k v

(λxN ′on1/x1, . . . , onm/xm)on

′′onm+1 . . . onm+k →v

(N ′on′′/x, on1/x1 . . . , o

nm/xm)onm+1 . . . onm+k v on

If M ∈ anf then there exists h ∈ N∗ such that for all n1, . . . , nm ≥ j + hone has Mon1/x1, . . . , o

nm/xm v on for some n ≥ j by lemma 26, thus we

conclude by taking k = 0. ut

A.5 Proofs of Subsection 6.2

The two following lemmas are used in the proof of Thm. 29 at p. 11.

Lemma 47. Let V be a λ-value and ~x = (x1, . . . , xn) ⊇ fv(V ) (with n ∈ N).

(i) For every ((a1, . . . , an), [ ]) ∈ JV K~x one has ai = [ ] for any 1 ≤ i ≤ n;

(ii) For any m ∈ N, if ((a11, . . . , a1n), a10), . . . , ((am1 , . . . , a

mn ), am0 ) ∈ JV K~x and

aj =⊎mi=1 a

ij for any 0 ≤ j ≤ n, then ((a1, . . . , an), a0) ∈ JV K~x (in particular,

([ ]n, [ ]) ∈ JV K~x).

Proof. We prove simultaneously points (i) and (ii) by simple inspection.

– If V is a variable then V = xk for some 1 ≤ k ≤ n, thus ((a1, . . . , an), [ ]) ∈JV K~x entails aj = [ ] for any 1 ≤ j ≤ n. If ((a11, . . . , a

1n), a10), . . . , ((am1 , . . . , a

mn ), am0 ) ∈

JxkK~x then, for every 1 ≤ i ≤ m, one has ai0 = aik (and so a0 =⊎mi=1 a

i0 =⊎m

i=1 aik = ak) and aij = [ ] (and so aj =

⊎mi=1 a

ij = [ ]) for any 1 ≤ j ≤ n with

j 6= k, therefore ((a1, . . . , an), a0) = (([ ], . . . , ak, . . . , [ ]), ak) ∈ JxkK~x = JV K~x .

Page 24: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

24 A. Carraro, G. Guerrieri

– If V = λy.N for some λ-termN then JV K~x = (⊎pk=1

~dk, [(b1, c1), . . . , (bp, cp)]) :

p ≥ 0, ∀k = 1, . . . , p. ((~dk, bk), ck) ∈ JNK~x,y. If ((a1, . . . , an), [ ]) ∈ JV K~x thenp = 0 and so aj = [ ] for any 1 ≤ j ≤ n. If ((a11, . . . , a

1n), a10), . . . , ((am1 , . . . , a

mn ), am0 ) ∈

JV K~x then, for any 1 ≤ i ≤ m, there exist pi ∈ N, bi1, . . . , bipi , c

i1, . . . , c

ipi ∈

Mf(U) and (ai1,1, . . . , ain,1), . . . , (ai1,pi , . . . , a

in,pi) ∈ Mf(U)n such that ai0 =

[(bi1, ci1), . . . , (bipi , c

ipi)], a

ij =

⊎pik=1 a

ij,k for any 1 ≤ j ≤ n and ((ai1,k, . . . , a

in,k, b

ik), cik) ∈

JNK~x,y for any 1 ≤ i ≤ m and 1 ≤ k ≤ pi; since aj =⊎mi=1 a

ij =⊎m

i=1

⊎pik=1 a

ij,k for any 1 ≤ j ≤ n and a0 =

⊎mi=1 a

i0 =

⊎mi=1

⊎pik=1[(bik, c

ik)],

one has ((a1, . . . , an), a0) ∈ Jλy.NK~x = JV K~x . ut

Lemma 48. Let M be a λ-term, let ~x = (x1, . . . , xm) and ~y = (y1, . . . , yn) be twofinite sequences of pairwise distinct variables such that fv(M) ⊆ x1, . . . , xm, y1, . . . , yn.If V1, . . . , Vm are λ-values such that

⋃mi=1 fv(Vi) ⊆ y1, . . . , yn, then

JMV1/x1, . . . , Vm/xmK~y =

(⊎mi=0 ~ai, c) : ∃ b1, . . . , bm ∈Mf(U).

((~a0, b1, . . . , bm), c) ∈ JMK~y,~x and (~ai, bi) ∈ JViK~y for all 1 ≤ i ≤ m

(1)

Proof. By induction on M . Let us denote by S the set in the right-hand side ofrelation (1), and by N∗ the λ-term NV1/x1, . . . , Vm/xm, for any λ-term N .

If M = xi for some 1 ≤ i ≤ m, then JM∗K~y = JViK~y = S by Lemma 47 andsince JxiK~y,~x = (([ ]n, [ ], . . . , c, . . . , [ ]), c) : c ∈Mf(U).

If M is a variable different from all xi’s, then M = yj for some 1 ≤ j ≤ n,hence JM∗K~y = JyjK~y = (([ ], . . . , c, . . . , [ ]), c) : c ∈ Mf(U) = S by Lemma 47and since JyjK~y,~x = (([ ], . . . , c, . . . , [ ], [ ]m), c) : c ∈Mf(U).

If M = λzN for some λ-term N then fv(N) ⊆ x1, . . . , xm, y1, . . . , yn, z andwe can suppose without loss of generality that z /∈ x1, . . . , xm, y1, . . . , yn, so

JM∗K~y = Jλz.N∗K~y =

(⊎kj=1 ~a

j , [(d1, c1), . . . , (dk, ck)]) : k ≥ 0,

((~aj , dj), cj) ∈ JN∗K~y,z for all 1 ≤ j ≤ k. (2)

By induction hypothesis, for every 1 ≤ j ≤ k, ((~aj , dj), cj) ∈ JN∗K~y,z if and

only if there exist bj1, . . . , bjm ∈Mf(U) and ~aj0, . . . ,~a

jm ∈Mf(U)n such that ~aj =⊎m

i=0 ~aji , ((~aj0, d

j , bj1, . . . , bjm), cj) ∈ JNK~y,z,~x and ((~aji , [ ]), bji ) ∈ JViK~y,z (which

is equivalent to (~aji , bji ) ∈ JViK~y because z /∈ fv(Vi)) for any 1 ≤ i ≤ m; let

c = [(d1, c1), . . . , (dk, ck)], ~ai =⊎kj=1 ~a

ji and bi =

⊎kj=1 b

ji : one has

⊎mi=0 ~ai =⊎m

i=0

⊎kj=1 ~a

ji =

⊎kj=1 ~a

j , ((~a0, b1, . . . , bm), c) ∈ Jλz.NK~y,~x and, by Lemma 47,(~ai, bi) ∈ JViK~y for all 1 ≤ i ≤ m. Therefore, according to relation (2), JM∗K~y = S.

If M = NL then fv(N), fv(L) ⊆ y1, . . . , yn and

JM∗K~y = JN∗L∗K~y = (~a′ ] ~a′′, c) | ∃ b ∈Mf(U) : (~a′, [(b, c)]) ∈ JN∗K~yand ( ~a′′, b) ∈ JL∗K~y. (3)

By induction hypothesis, (~a′, [(b, c)]) ∈ JN∗K~y iff there exist b′1, . . . , b′m ∈Mf(U)

and ~a′0, . . . ,~a′m ∈ Mf(U)n such that a′ =

⊎mi=0

~a′i, (( ~a′0, b′1, . . . , b

′m), [(b, c)]) ∈

Page 25: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account of call-by-value solvability 25

JNK~y,~x and (~a′i, b′i) ∈ JViK~y for all 1 ≤ i ≤ m; and ( ~a′′, b) ∈ JL∗K~y iff there

exist b′′1 , . . . , b′′m ∈ Mf(U) and ~a′′0 , . . . ,

~a′′m ∈ Mf(U)n such that a′′ =⊎mi=0

~a′′i ,

(( ~a′′0 , b′′1 , . . . , b

′′m), b) ∈ JLK~y,~x and ( ~a′′i , b

′′i ) ∈ JViK~y for all 1 ≤ i ≤ m; let ~ai = ~a′i] ~a′′i

and bi = b′i ] b′′i : one has⊎mi=0 ~ai =

⊎mi=0

~a′i ] ~a′′i = ~a′ ] ~a′′, ((~a0, b1, . . . , bm), c) ∈JNLK~y and, by Lemma 47, (~ai, bi) ∈ JViK~y for all 1 ≤ i ≤ m. Therefore, accordingto relation (3), JM∗K~y = S. ut

A.6 Proofs of Subsection 6.3

Theorem 33 (soundness). Let P,Q ∈ 2〈rΛτ 〉. If P→vτ Q, then JPK~x = JQK~x. Stated at p. 13

Proof. By Thm. 18 it suffices to prove that the τi-rules are sound. For example

Jt(s ∗ p)K~x = (~a0 ] ~a1 ] ~a2, c) : ∃b ∈Mf(U).

(~a0, [(b, c)]) ∈ JtK~x, (~a1, b) ∈ JsK~x, (~a2, 1) ∈ JpK~x= Jts ∗ pK~x

If n > 0, then

Jτ [[v1, . . . , vn]]K~x = (~a, 1) : (~a, [ ]) ∈ J[v1, . . . , vn]K~x = ∅

because the interpretation of a non-empty bag of values which does not reduceto 0 never contains an element like (~a, [ ]). Instead

Jτ [[ ]]K~x = (~a, 1) : (~a, [ ]) ∈ ([ ]n, [ ]) = JεK~x

This shows that the interpretation is invariant w.r.t. the rules 7→τ1 and 7→τ5 . Theother τi-rules are proved similarly. ut

Lemma 49. Let p be a closed test. Then:

(i) either pvτ ε or pvτ 0;(ii) JpK 6= ∅ iff pvτ ε.

Proof. (i) It suffices to show that for every closed resource term t, eitherτ [t] vτ ε or τ [t] vτ 0. As the rτλσv -calculus is strongly normalizing,

we have that t vτ

∑ki=1 si, where each si is a closed normal form. If

k = 0 then τ [t] vτ 0 since τ [0] = 0. Otherwise for each si there are twopossibilities:– si = [v1, . . . , vm] with vj not a variable. Then τ [si] reduces either to ε

or to 0, depending on the value of m.– si = τ [[v1, . . . , vm] ∗ ([v′1, . . . , v

′m′ ] ‖ q)] with vj , v

′l not a variable. Then

si →vτ τ [[v1, . . . , vm]] ‖ ([v′1, . . . , v′m′ ] ‖ q) that, using the induction

hypothesis, can only reduce to 0 or ε.We conclude since τ [t]vτ

∑ki=1 τ [si], and this latter expression reduces

to a finite (possibly empty) sum of ε’s, which is thus equal to 0 or ε.(ii) By the soundness of the model (Thm. 33) and item (i). ut

Page 26: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

26 A. Carraro, G. Guerrieri

The set U admits a well-founded ordering via the notion of rank of an elementα ∈ U : the rank of α, notation rk(α), is the smallest n ∈ N such that α ∈ Un.The rank of a multiset b ∈Mf(U), denoted by rk(b), is the greatest among theranks of its elements, if it is non-empty; the empty multiset has rank 0.

Lemma 50. Let a ∈Mf(U). Then:

(i) Ja−K = (1, a),(ii) Ja+L[xn]MKx = (a, 1), where n = #a.

Proof. The points (i) and (ii) are proved simultaneously by induction on rk(a).We write IH(i) and IH(ii) for the induction hypotheses concerning (i) and (ii),respectively.

If rk(a) = 0 then a = [ ], hence Ja−K = J[ ]K = [ ] and Ja+L[x0]MKx =Jτ [[λy.[ ] ∗ ε][ ]]Kx = JεKx = ([ ], 1). This proves the base cases for (i) and (ii).

For the inductive step, suppose rk(a) > 0 and #a = r, so that a = [(b1, c1), . . . , (br, cr)].

We prove (i). By definition Ja−K = J[λy1.c1− ∗ b1+L[ym11 ]M, . . . , λyr.cr− ∗ br+L[ymr

r ]M]K.So we have (1, a′) ∈ Ja−K iff a′ = [(b′1, c

′1), . . . , (b′r, c

′r)] and for all 1 ≤ j ≤ r

(b′j , c′j) ∈ Jcj− ∗ bj+L[ymj

j ]MK, i.e., (1, c′j) ∈ Jcj−K and (b′j , 1) ∈ Jbj+L[ymj

j ]MKyj . ByIH(i) and IH(ii) we have c′j = cj and b′j = bj . Therefore a′ = a.

We prove (ii). By definition a+L[xr]M = τ [[λz.[ ]∗ ‖ri=1 τ [[λy.[ ] ∗ ci+L[yki ]M]([z]bi−)]][xr]],where ki = #ci for all 1 ≤ i ≤ r. Hence, a+L[xr]Mvτ‖ri=1 τ [[λy.[ ] ∗ ci+L[yki ]M]([x]bi

−)].Using IH(i) and IH(ii) we have that Jbi−K = (1, bi) and Jci+L[yki ]MKy = (ci, 1)and therefore Jτ [[λy.[ ] ∗ ci+L[yki ]M]([x]bi

−)]Kx = ([(bi, ci)], 1), for each i =1, . . . , r. In conclusion Ja+L[xr]MKx = ([(b1, c1), . . . , (br, cr)], 1). ut

Lemma 51. Let b ∈Mf(U) and let t be a resource term with fv(t) ⊆ ~x. ThenJb+LtMK~x = (~a, 1) : (~a, b) ∈ JtK~x.

Proof. By induction on the structure of t. ut

Lemma 52. Let ~a ∈Mf(U)n and let p be a test with fv(p) ⊆ ~x. ThenJp〈~a−/~x〉K 6= ∅ iff (~a, 1) ∈ JpK~x.

Proof. By induction on the structure of p. ut

Lemma 53. Let b ∈ Mf(U) and let #b = n. If s is a resource term withfv(s) ⊆ ~x, then Jb+LsMK~x 6= ∅ iff Jτ [[λy.[ ] ∗ b+L[yn]M]s]K~x 6= ∅.

Proof. By induction on the rank of b. ut

Lemma 35. Let (~a, b) ∈Mf(U)n×Mf(U), #b = r and let t be a resource term.Stated at p. 13

Then (~a, b) ∈ JtK~x iff τ [[λy.[ ] ∗ b+L[yr]M](t〈~a−/~x〉)]vτ ε.

Proof. We have the following chain of equivalences:

(~a, b) ∈ JtK~x ⇔ (~a, 1) ∈ Jb+LtMK~x, by Lemma 51⇔ Jb+LtM〈~a−/~x〉K 6= ∅, by Lemma 52⇔ Jb+Lt〈~a−/~x〉MK 6= ∅, since fv(b+L·M) = ∅⇔ Jτ [[λy.[ ] ∗ b+L[yr]M](t〈~a−/~x〉)]K 6= ∅, by Lemma 53⇔ τ [[λy.[ ] ∗ b+L[yr]M](t〈~a−/~x〉)]vτ ε, by Lemma 49 ut

Page 27: A semantical and operational account of call-by-value ...giuliog/fossacs.pdf · A semantical and operational account of call-by-value solvability Alberto Carraro1; 2and Giulio Guerrieri

A semantical and operational account of call-by-value solvability 27

A.7 Proofs of Subsection 6.4

Lemma 54. Let M ∈ Λ, V ∈ Λv, t ∈ T (M) and v1, . . . , vm ∈ T (V ). For everyf ∈ Sm, if degx(t) = m then tvf(1)/x1, . . . , vf(m)/x

m ∈ T (NV/x).Proof. See Lemma 16 in [14]. ut

The following lemma is a generalization of Lemma 18 in [14].

Lemma 38. Let M,M ′ be λ-terms. Stated at p. 14

(i) If M →w M′ and t ∈ T (M), then there exists T ⊆ T (M ′) such that t→v T.

(ii) If M →s M′ and s ∈ Ts(M), then there exists S ⊆ Ts(M ′) such that s→+

v S.

Proof (of Lemma 38.i). By induction on the definition of M →w M′.

If M = (λx.N)V →w NV/x = M ′ then t = [λx.t1, . . . , λx.tn][v1, . . . , vm]for some n,m ∈ N, with ti ∈ T (N) and vj ∈ T (V ) for all 1≤ i≤n and 1≤j≤m.If n 6= 1, or n = 1 but degx(t1) 6= m, then t→v 0 ⊆ Ts(M ′). Otherwise, n = 1 anddegx(t1) = m, so t = [λx.t1][v1, . . . , vm]→v

∑f∈Sm

t1vf(1)/x1, . . . , vf(m)/xm =

T. By Lemma 54, T ⊆ T (M ′).If M = (λx.M0)NL →w (λx.M0L)N = M ′, then t = [λx.t1, . . . , λx.tn]s1s2

for some s1 ∈ T (N), s2 ∈ T (L) and n ∈ N, with ti ∈ T (M0) for all 1 ≤ i ≤ n. Ifn 6= 1 then t→v 0 ⊆ T (M ′), otherwise t = [λx.t1]s1s2 →v [λx.t1s2]s1 ⊆ T (M ′).

IfM = V ((λx.L)N)→w (λx.V L)N = M ′, then t = [v1, . . . , vn]([λx.t1, . . . , λx.tn]s)for some s ∈ T (N) and n,m ∈ N, with vi ∈ T (V ) and tj ∈ T (L) for all1 ≤ i ≤ n and 1 ≤ j ≤ m. If n 6= 1 or m 6= 1 then t →v 0 ⊆ T (M ′), otherwiset = [v1]([λx.t1]s′)→v [λx.v1t1]s′ ⊆ T (M ′).

If M = M0M1 →w M′0M1 = M ′ with M0 →w M

′0, then t = s0s1 where

si ∈ T (Mi) for i ∈ 1, 2. By induction hypothesis, there exists T0 ⊆ T (M ′0)such that s0 →v T0. Therefore t→v T0s1 ⊆ T (M ′).

The case where M = M0M1 →w M0M′1 = M ′ with M1 →w M

′1 is perfectly

similar to the previous one.IfM = (λx.L)N →w (λx.L′)N = M ′ with L→w L

′, then t = [λx.t1, . . . , λx.tn]sfor some s ∈ T (N) and n ∈ N, with ti ∈ T (L) for all 1 ≤ i ≤ n. If n 6= 1 thent →v 0 ⊆ T (M ′). Otherwise t = [λx.t1]s and, by induction hypothesis, thereexists L ⊆ T (L′) such that t1 →v L. Hence t→v (λx.L)s ⊆ T (M ′). utProof (of Lemma 38.ii). By induction on the definition of M →s M

′.If M →w M

′ then there exists T ⊆ T (M ′) such that t→v T, by Lemma 38.i.Since t ∈ Ts(M), for every T′ if T→∗v T′ then T′ ⊆ Strat. Therefore T ⊆ Ts(M ′).

If M = M0M1 →s M′0M1 = M ′ with M0 →s M

′0, then t = s0s1 where

s0 ∈ Ts(M0) and s1 ∈ T (M1). By induction hypothesis, there exists S0 ⊆ Ts(M ′0)such that s0 →+

v S0, thus t→+v S0s1 ⊆ T (M ′). Since t ∈ Ts(M), for every T′ if

S0s1 →∗v T′ then T′ ⊆ Strat. Therefore S0s1 ⊆ Ts(M ′).If M = λx.L →s λx.L

′ = M ′ with L →s L′, then t = [λx.t1, . . . , λx.tn]

for some n ≥ 1 with ti ∈ Ts(L) for all 1 ≤ i ≤ n. By induction hypothesis,for all 1 ≤ i ≤ n there exists Li ⊆ Ts(L′) such that ti →+

v Li. So, t →+v

[λx.L1, λx.t2, . . . , λx.tn] →+v . . . →+

v [λx.L1, . . . , λx.Ln] ⊆ T (M ′). Since t ∈Ts(M), for every T′ if [λx.L1, . . . , λx.Ln] →∗v T′ then T′ ⊆ Strat. Therefore[λx.L1, . . . , λx.Ln] ⊆ Ts(M ′). ut