Standardization of a call-by-value calculus

Size: px
Start display at page:

Download "Standardization of a call-by-value calculus"

Transcription

1 Standardization of a call-by-value calculus Giulio Guerrieri 1, Luca Paolini 2, and Simona Ronchi Della Rocca 2 1 Laboratoire PPS, UMR 7126, Université Paris Diderot, Sorbonne Paris Cité F Paris, France giulio.guerrieri@pps.univ-paris-diderot.fr 2 Dipartimento di Informatica, Università degli Studi di Torino C.so Svizzera 185, Torino, Italia {paolini,ronchi}@di.unito.it Abstract We study an extension of Plotkin s call-by-value lambda-calculus by means of two commutation rules (sigma-reductions). Recently, it has been proved that this calculus can provide elegant characterizations of many semantic properties. We prove a standardization theorem for this calculus by generalizing Takahashi s approach of parallel reductions. The standardization property allows us to prove that our calculus enjoys some meaningful conservation properties with respect to Plotkin s one. In particular, we show that the notion of solvability for this calculus coincides with that for Plotkin s call-by-value lambda-calculus ACM Subject Classification D.3.1 Formal Definitions and Theory, F.3.2 Semantics of Programming Language, F.4.1 Mathematical Logic. Keywords and phrases standardization, sequentialization, lambda-calculus, sigma-reduction, parallel reduction, call-by-value, head reduction, ernal reduction, solvability, potential valuability, observational equivalence. Digital Object Identifier /LIPIcs.xxx.yyy.p 1 Introduction The λ v -calculus has been roduced by Plotkin in [16], in order to give a formal account of the call-by-value evaluation, which is the most commonly used parameter passing policy for programming languages. λ v shares the syntax with the classical, call-by-name, λ-calculus, but its reduction rule, β v, is a restriction of the β one, firing only in case the argument is a value (i.e., a variable or an abstraction). While β v is enough for evaluation, it is turned out to be too weak in order to study operational properties of terms. For example, in λ, the β-reduction is sufficient to characterize solvability and (in addition with η) separability, but, in order to characterize similar properties for λ v, it has been necessary to roduce different notions of reduction unsuitable for a correct call-by-value evaluation (see [14, 15]), which is disappoing and requires complex reasoning. In this paper we study λ σ v, the extension of λ v proposed in [3]. It keeps the λ v syntax and it adds to the β v -reduction two commutation rules, called σ 1 and σ 3, which unblock β v -redexes which are hidden by the hyper-sequential structure of terms. It is well-known (see [15, 1]) that in λ v there are normal forms that are unsolvable, e.g. (λyx.xx)(zz)(λx.xx). The more evident benefit of λ σ v is that the commutation rules make all normal forms solvable (indeed (λyx.xx)(zz)(λx.xx) This work was partially supported by LINTEL TO_Call1_2012_0085, i.e. a Research Project funded by the Compagnia di San Paolo. Giulio Guerrieri, Luca Paolini and Simona Ronchi Della Rocca; licensed under Creative Commons License CC-BY Conference title on which this volume is based on. Editors: Billy Editor and Bill Editors; pp Leibniz International Proceedings in Informatics Schloss Dagstuhl Leibniz-Zentrum für Informatik, Dagstuhl Publishing, Germany

2 2 Standardization of a call-by-value calculus is not a λ σ v normal form). More generally, the so obtained language, allows to characterize operational properties, like solvability and potential valuability, in an ernal and elegant way (see [3]). In this paper we prove a standardization property in λ σ v, and some of its consequences, namely its soundness with respect to the semantics of λ v. Let us recall the notion of standardization, which has been first studied in the ordinary λ-calculus (see, for example [5, 8, 2]). A reduction sequence is standard if its redexes are ordered in a given way, and the corresponding standardization theorem establishes that every reduction sequence can constructively be transformed o a standard one. Standardization is a key tool to grasp the way in which reductions works that sheds some light on redexes relationships and their dependencies. It is useful for characterize semantic properties through reduction strategies (the proof of operational semantics adequacy is a typical use). In the λ v setting standardization theorems have been proved by Plotkin [16], Paolini and Ronchi Della Rocca [15, 13] and Crary [4]. The definition of standard sequence of reductions considered by Plotkin and Crary coincides, and it imposes a partial order between redexes, while Paolini and Ronchi Della Rocca define a total order between them. All these proofs are developed by using the notion of parallel reduction roduced by Tait and Martin-Löf (see Takahashi [19] for details and eresting technical improvements). We emphasize that this method does not involve the notion of residual of a redex, on which many classical proofs for the λ-calculus are based (see for example [2, 8]). As in [16, 19, 15, 4], we use a suitable notion of parallel reduction for developing our standardization theorem for λ σ v. In particular we consider two groups of redexes, head β v -redexes and head σ-redexes (putting together σ 1 and σ 3 ), and we induce a total order between head redexes of the two groups, without imposing any order between head σ-redexes themselves. More precisely, when σ-redexes are missing, this notion of standardization coincides with that presented in [15]. Moreover, we show that it is not possible to strength our standardization by (locally) ordering σ 1 -reduction to σ 3 -reduction (or viceversa). As usual, our standardization proof is based on a sequentialization result: inner reductions can always be postponed to the head ones, for a non-standard definition of head reduction. Sequentialization has eresting consequences: it allows us to prove that fundamental operational properties in λ σ v, like observational equivalence, potential valuability and solvability, are conservative with respect to the corresponding notions of λ v. This fully justify the project in [3] where λ σ v has been roduced as a tool for studying the operational behaviour of λ v. Other variants of λ v have been roduced in the literature for modeling the call-by-value computation. We would like here to cite at least the contributions of Moggi [10, 11], Felleisen and Sabry [17, 18], Maraist et al. [9], Herbelin and Zimmerman [7], Accattoli and Paolini [1]. All these proposals are based on the roduction of new constructs to the syntax of λ v, so the comparison between them is not easy with respect to syntactical properties (some detailed comparison is given in [1]). We po out that the calculi roduced in [10, 11, 17, 18, 9, 7] present some variants of our σ 1 and/or σ 3 rules, often in a setting with explicit substitutions. Outline. In Section 2 we roduce the language λ σ v and its operational behaviour; in Section 3 the sequentialization property is proved; Section 4 contains the main result, i.e., standardization; in Section 5 some conservativity results with respect to Plotkin s λ v -calculus are proved. Section 6 concludes the paper, with some h for a future work. 2 The call-by-value lambda calculus with sigma-rules In this section we present λ σ v, a call-by-value λ-calculus roduced in [3] that adds two σ-reduction rules to pure (i.e. without constants) call-by-value λ-calculus defined by Plotkin in [16].

3 G. Guerrieri, L. Paolini and S. Ronchi Della Rocca 3 The syntax of terms of λ σ v [3] is the same as the one of ordinary λ-calculus and Plotkin s call-by-value λ-calculus λ v [16] (without constants). Given a countable set V of variables (denoted by x, y, z,... ), the sets Λ of terms and Λ v of values are defined by mutual induction: (Λ v ) V, U ::= x λx.m values (Λ) M, N, L ::= V MN terms Clearly, Λ v Λ. All terms are considered up to α-conversion. The set of free variables of a term M is denoted by fv(m). Given V 1,..., V n Λ v and pairwise distinct variables x 1,..., x n, M{V 1 /x 1,..., V n /x n } denotes the term obtained by the capture-avoiding simultaneous substitution of V i for each free occurrence of x i in the term M (for all 1 i n). Note that, for all V, V 1,..., V n Λ v and pairwise distinct variables x 1,..., x n, V {V 1 /x 1,..., V n /x n } Λ v. Contexts (with exactly one hole ), denoted by C, are defined as usual via the grammar: C ::= λx.c CM MC. We use C M for the term obtained by the capture-allowing substitution of the term M for in the context C. Notation. For now on, we set I = λx.x and = λx.xx. The reduction rules of λ σ v consist of Plotkin s β v -reduction rule, roduced in [16], and two simple commutation rules called σ 1 and σ 3, studied in [3]. Definition 1 (Reduction rules). We define the following binary relations on Λ (for any M, N, L Λ and any V Λ v ): (λx.m)v βv M{V/x} (λx.m)nl σ1 (λx.ml)n with x / fv(l) V ((λx.l)n) σ3 (λx.v L)N with x / fv(v ). For any r {β v, σ 1, σ 3 }, if M r M then M is a r-redex and M is its r-contractum. In this sense, a term of the shape (λx.m)n (for any M, N Λ) is a β-redex. We set σ = σ1 σ3 and v = βv σ. The side conditions on σ1 and σ3 in Definition 1 can be always fulfilled by α-renaming. Obviously, any β v -redex is a β-redex but the converse does not hold: (λx.z)(yi) is a β-redex but not a β v -redex. Redexes of different type may overlap: for example, the term I is a σ 1 -redex and it contains the β v -redex I; the term (I )(xi) is a σ 1 -redex and it contains the σ 3 -redex (I ) which contains in turn the β v -redex I. According to the Girard s call-by-value boring translation ( ) v of terms o Intuitionistic Multiplicative Exponential Linear Logic proof-nets, defined by (A B) v =!A v!b v (see [6]), the images under ( ) v of a σ 1 -redex (resp. σ 3 -redex) and its contractum are equal modulo some bureaucratic steps of cut-elimination. Notation. Let R be a binary relation on Λ. We denote by R (resp. R + ; R = ) the reflexivetransitive (resp. transitive; reflexive) closure of R. Definition 2 (Reductions). Let r {β v, σ 1, σ 3, σ, v}. The r-reduction r is the contextual closure of r, i.e. M r M iff there is a context C and N, N Λ such that M = C N, M = C N and N r N. The r-equivalence = r is the reflexive-transitive and symmetric closure of r.

4 4 Standardization of a call-by-value calculus Let M be a term: M is r-normal if there is no 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 (N i ) i N such that M = N 0 and N i r N i+1 for any i N. Finally, r is strongly normalizing if every N Λ is strongly r-normalizing. Patently, σ v and βv v. Remark 3. For any r {β v, σ 1, σ 3, σ, v} (resp. r {σ 1, σ 3, σ}), values are closed under r- reduction (resp. r-expansion): for any V Λ v, if V r M (resp. M r V ) then M Λ v ; more precisely, V = λx.n and M = λx.n for some N, N Λ with N r N (resp. N r N). Proposition 4 (see [3]). The σ-reduction is confluent and strongly normalizing. The v-reduction is confluent. The λ σ v -calculus, λ σ v for short, is the set Λ of terms endowed with the v-reduction v. The set Λ endowed with the β v -reduction βv is the λ v -calculus (λ v for short), i.e. the Plotkin s call-by-value λ-calculus [16] (without constants), which is thus a sub-calculus of λ σ v. Example 5. M = (λy. )(xi) σ1 (λy. )(xi) βv (λy. )(xi) βv... and N = ((λy. )(xi)) σ3 (λy. )(xi) βv (λy. )(xi) βv... are the only possible v-reduction paths from M and N respectively: M and N are not v-normalizable, and M = v N. Meanwhile, M and N are β v -normal and different, hence M βv N (by confluence of βv, see [16]). Informally, σ-rules unblock β v -redexes which are hidden by the hyper-sequential structure of terms. This approach is alternative to the one in [1] where hidden β v -redexes are reduced thanks to a rule acting at a distance (through explicit substitutions). It can be shown that the call-by-value λ-calculus with explicit substitution roduced in [1] can be embedded in λ σ v. 3 Sequentialization In this section we aim to prove a sequentialization theorem (Theorem 22) for the λ σ v -calculus by adapting the Takahashi s method [19, 4] based on parallel reductions. Notation. From now on, we always assume that V, V Λ v. Note that the generic form of a term is V M 1... M m for some m N (in particular, values are obtained when m = 0). The sequentialization result is based on a partitioning of v-reduction between head and ernal reduction. Definition 6 (Head β v -reduction). We define inductively the head β v -reduction βv the following rules (m N in both rules): by β v (λx.m)v M 1... M m βv M{V/x}M 1... M m N βv N right V NM 1... M m βv V N M 1... M m The head β v -reduction βv reduces exactly the same redexes (see [14]) as the left reduction defined in [16, p. 136], called evaluation in [17, 18, 4]. The head β v -reduction is deterministic and does not reduce values. Definition 7 (Head σ-reduction). We define inductively the head σ-reduction σ by the following rules (m N in all the rules, x / fv(l) in the rule σ 1, x / fv(v ) in the rule σ 3 ): σ 1 (λx.m)nlm 1... M m σ (λx.ml)nm 1... M m N σ N right V NM 1... M m σ V N M 1... M m V ((λx.l)n)m 1... M m σ (λx.v L)NM 1... M m σ 3

5 G. Guerrieri, L. Paolini and S. Ronchi Della Rocca 5 The head (β v σ-)reduction is βvσ = βv σ. The ernal reduction is v = v βvσ. Notice that βv βv βv and σ σ σ and v βvσ v. Values are normal forms for the head reduction, but the converse does not hold: xi / Λ v is head-normal. Informally, if N βv N (resp. N σ N ) then N is gotten from N by reducing the leftmost-outermost β v -redex (resp. σ 1 - or σ 3 -redex), not in the scope of a λ. Differently from βv, the head σ-reduction σ is not deterministic, because the leftmostoutermost σ 1 - and σ 3 -redexes may overlap: for example, if M = (λy.y )( (xi))i then M σ (λy.y I)( (xi))=n 1 by applying the rule σ 1 and M σ (λz.(λy.y )(zz))(xi)i =N 2 by applying the rule σ 3. Note that N 1 contains only a head σ 3 -redex and N 1 σ (λz.(λy.y I)(zz))(xI)= N which is normal for βvσ; meanwhile N 2 contains only a head σ 1 -redex and N 2 σ (λz.(λy.y )(zz)i)(xi) which is normal for βvσ: the head reduction βvσ is not confluent and a term may have several head-normal forms (this example does not contradict the confluence of σ-reduction because N 2 σ N but by performing an ernal reduction step). Anyway, next Corollary 26.2 implies that if a term M has a head normal form N Λ v then N is the unique head normal form of M. Definition 8 (Parallel reduction). We define inductively the parallel reduction by the following rules (m N in all the rules, x / fv(l) in the rule σ 1, x / fv(v ) in the rule σ 3 ): V V M i M i (0 i m) β v (λx.m 0)VM 1... M m M 0{V /x}m 1... M m N N L L M i M i (0 i m) (λx.m 0)NLM 1... M m (λx.m 0L )N M 1... M m σ 1 V V N N L L M i M i (1 i m) σ3 V ((λx.l)n)m 1... M m (λx.v L )N M 1... M m M i M i (0 i m) λ (λx.m 0)M 1... M m (λx.m 0)M 1... M m M i M i (1 i m) x M 1... M m x M 1... M m In Definition 8 the rule var has no premises when m = 0: this is the base case of the inductive definition of. The rules σ 1 and σ 3 have exactly three premises when m = 0. Intuitively, M M means that M is obtained from M by reducing a number of β v -, σ 1 - and σ 3 -redexes (existing in M) simultaneously. Definition 9 (Strong and ernal parallel reduction). We define inductively the ernal parallel reduction by the following rules (m N in the rule right): M M λx.m λx.m λ x x var var V V N N M i M i (1 i m) right V NM 1... M m V N M 1... M m The strong parallel reduction is defined by: M N iff M N and there exist M, M Λ such that M β v M σ M N. Notice that the rule right for has exactly two premises when m = 0. Remark 10. The relations, and are reflexive. The reflexivity of follows immediately from the reflexivity of and. The proofs of reflexivity of and are both by structural induction on a term: in the case of, remind that every term is of the form (λx.n)m 1... M m or x M 1... M m for some m N and then apply the rule λ or var respectively, together to the inductive hypothesis; in the case of, remind that every term is of the form λx.m or x or V NM 1... M m for some m N and then apply the rule λ (together to the reflexivity of ) or var or right (together to the reflexivity of and the inductive hypothesis) respectively.

6 6 Standardization of a call-by-value calculus One has (first, prove that by induction on the derivation of M M, the other inclusions follow from the definition of ) and, since is reflexive (Remark 10), βv and σ. Observe that R for any R { βv, βv,,, }, even if for different reasons: for example, by reflexivity of (Remark 10), whereas βv by reducing the (leftmost) β v -redex. Remark The head β v -reduction βv does not reduce a value (in particular, does not reduce under λ s), i.e., for any M Λ and any V Λ v, one has V βv M. 2. The head σ-reduction σ does neither reduce a value nor reduce to a value, i.e., for any M Λ and any V Λ v, one has V σ M and M σ V. 3. Variables and abstractions are preserved under ( -expansion), i.e., if M x (resp. M λx.n ) then M = x (resp. M = λx.n for some N Λ such that N N ). 4. If M M then λx.m R λx.m for any R {,, }. Indeed, for R {, } apply the rule λ to conclude, then λx.m λx.m according to the definition of. 5. For any V, V Λ v, V V iff V V. The left-to-right direction holds because ; conversely, assume V V : if V is a variable then necessarily V = V and hence V V by applying the rule var for ; otherwise V = λx.n for some N Λ, and then necessarily V = λx.n with N N, so V V by applying the rule λ for. See Appendix, p. 16 See Appendix, p. 16 Remark If M M and N N then MN M N. For the proof, it is sufficient to consider the last rule of the derivation of M M. 2. If R { βv, σ } and M R M, then MN R M N for any N Λ. For the proof, it is sufficient to consider the last rule of the derivation of M R M, for any R { βv, σ }. 3. If M M and N N where M / Λ v, then MN M N : indeed, the last rule in the derivation of M M can be neither λ nor var because M / Λ v. The hypothesis M / Λ v is crucial: for example, x x and I but I and thus x(i ) x. 4. v v. As a consequence, = v and (by Proposition 4) is confluent. 5. v v, so = v. Thus, by Remark 11.3, variables and abstractions are preserved under v -expansion, i.e., if M v x (resp. M v λx.n ) then M = x (resp. M = λx.n with N v N ). 6. For any R { βv, σ }, if M R M then M{V/x} R M {V/x} for any V Λ v. The proof is by straightforward induction on the derivation of M R M for any R { βv, σ }. As expected, a basic property of parallel reduction is the following: Lemma 13 (Substitution lemma for ). If M M and V V then M{V/x} M {V /x}. See Appendix, p. 16 Proof. By straightforward induction on the derivation of M M. The following lemma will play a crucial role in the proof of Lemmas and shows that the head σ-reduction σ can be postponed to the head β v -reduction βv. Lemma 14 (Commutation of head reductions). 1. If M σ L βv N then there exists L Λ such that M βv L = σ N. 2. If M σ L β v N then there exists L Λ such that M β v L σ N. 3. If M β vσ M then there exists N Λ such that M β v N σ M. Proof. 1. By induction on the derivation of M σ L. Let us consider its last rule r. If r = σ 1 then M = (λx.m 0 )N 0 L 0 M 1... M m and L = (λx.m 0 L 0 )N 0 M 1... M m where m N and x / fv(l 0 ). Since L βv N, there are only two cases:

7 G. Guerrieri, L. Paolini and S. Ronchi Della Rocca 7 either N 0 βv N 0 and N = (λx.m 0 L 0 )N 0M 1... M m (according to the rule right for βv ), then M βv (λx.m 0 )N 0L 0 M 1... M m σ N; or N 0 Λ v and N =M 0 {N 0 /x}l 0 M 1... M m (by the rule β v as x / fv(l 0 )), so M βv N. If r = σ 3 then M = V ((λx.l 0 )N 0 )M 1... M m and L = (λx.v L 0 )N 0 M 1... M m with m N and x / fv(v ). Since L βv N, there are only two cases: either N 0 βv N 0 and N = (λx.v L 0 )N 0M 1... M m (according to the rule right for βv ), then M βv V ((λx.l 0 )N 0)M 1... M m σ N; or N 0 Λ v and N =VL 0 {N 0 /x}m 1... M m (by the rule β v, as x / fv(v )), so M βv N. Finally, if r = right then M = V N 0 M 1... M m and L = V N 0M 1... M m with m N and N 0 σ N 0. By Remark 11.2, N 0 / Λ v and thus, since L βv N, the only possibility is that N 0 βv N 0 and N = V N 0 M 1... M m (according to the rule right for βv ). By induction hypothesis, there exists N 0 Λ such that N 0 βv N 0 = σ N 0. Therefore, M βv V N 0 M 1... M m = σ N. 2. Immediate consequence of Lemma 14.1, using standard techniques of rewriting theory. See Appendix, p Immediate consequence of Lemma 14.2, using standard techniques of rewriting theory. See Appendix, p. 17 We are now able to travel over again the Takahashi method [19, 4] in our setting with β v - and σ-reduction. The next four lemmas govern the strong parallel reduction. Lemma 15. If M M and N N and M / Λ v, then MN M N. Proof. One has MN M N by Remark 12.1 and since M M. By hypothesis, there exist m, n N and M 0,... M m, N 0,..., N n such that M = M 0, M m = N 0, N n M, M i βv M i+1 for any 0 i < m and N j σ N j+1 for any 0 j < n; by Remark 12.2, M i N βv M i+1 N for any 0 i < m and N j N σ N j+1 N for any 0 j < n. As M / Λ v, one has N n N M N by Remark Therefore, MN M N. Lemma 16. If M M and N N then MN M N. Proof. If M / Λ v then MN M N by Lemma 15 and since N N. Assume M Λ v : MN M N by Remark 12.1, as M M and N N. By hypothesis, there are m, m, n, n N and M 0,..., M m, M 0,..., M m, N 0,..., N n, N 0,..., N n such that: M = M 0, M m = M 0, M m for any 0 i < m, N = N 0, N n = N 0, N n any 0 j < n. M, M i βv M i+1 for any 0 i < m, and M i σ M i +1 N, N j βv N j+1 for any 0 j < n and N j σ N j +1 for By Remark 11.3, M m Λ v since M Λ v, therefore m = 0 by Remark 11.2, and thus M m = M 0 M (in particular, M m M ) and M m Λ v. Thanks to the rules right for βv and σ, one has M m N j βv M m N j+1 for any 0 j < n, and M m N j σ M m N j +1 for any 0 j < n. By Remark 12.2, M i N 0 βv M i+1 N 0 for any 0 i < m. By applying the rule right for, one has M m N n M N. Therefore, MN = M 0 N 0 β v M m N 0 β v M m N n = M m N 0 σ M m N n M N and hence MN M N. Lemma 17. If M M and V V, then M{V/x} M {V /x}. Proof. By Lemma 13, one has M{V/x} M {V /x} since M M and V V. We proceed by induction on M Λ. Let us consider the last rule r of the derivation of M M. If r = var then there are two cases: either M = x and then M{V/x}=V V =M {V /x}; or M = y x and then M{V/x} = y = M {V /x}, so M{V/x} M {V /x} by Remark 10.

8 8 Standardization of a call-by-value calculus If r = λ then M = λy.n and M = λy.n with N N ; we can suppose without loss of generality that y / fv(v ) {x}. One has N{V/x} N {V /x} according to Lemma 13. By applying the rule λ for, one has M{V/x} = λy.n{v/x} λy.n {V /x} = M {V /x} and thus M{V/x} M {V /x}. Finally, if r = right then M = UNM 1... M m and M = U N M 1... M m for some m N with U, U Λ v, U U, N N and M i M i for any 1 i m. By induction hypothesis, U{V/x} U {V /x} (indeed U U according to Remark 11.5) and N{V/x} N {V /x}. By Lemma 16, U{V/x}N{V/x} U {V /x}n {V/x} and hence, by Lemma 15 since U {V /x}n {V/x} / Λ v, one has M{V/x} = U{V/x}N{V/x}M 1 {V/x}... M m {V/x} U {V /x}n {V /x}m 1{V /x}... M m{v /x} = M {V /x}. In the proof of the two next lemmas, as well as in the proof of Corollary 21 and Theorem 22, our Lemma 14 plays a crucial role: indeed, since the head σ-reduction well eract with the head β v -reduction, the Takahashi method [19, 4] is still working when adding the rules σ 1 and σ 3 to Plotkin s β v -reduction. Lemma 18 (Substitution lemma for ). If M M and V V then M{V/x} M {V /x}. Proof. By Lemma 13, one has M{V/x} M {V /x} since M M and V V. By hypothesis, there exist m, n N and M 0,... M m, N 0,..., N n such that M = M 0, M m = N 0, N n M, M i βv M i+1 for any 0 i < m and N j σ N j+1 for any 0 j < n; by Remark 12.6, M i {V/x} βv M i+1 {V/x} for any 0 i < m, and N j {V/x} σ N j+1 {V/x} for any 0 j < n. By Lemma 17, one has N n {V/x} M {V /x}, thus there exist L, N Λ such that M{V/x} β v N 0 {V/x} σ N n {V/x} β v N σ L M {V /x}. By Lemma 14.2, there exists N Λ such that M{V/x} β v N 0 {V/x} β v N σ N σ L M {V /x}, therefore M{V/x} M {V /x}. Now we are ready to prove the main lemma, which means that parallel reduction coincides with strong parallel reduction (the inclusion is trivial). Lemma 19 (Main Lemma). If M M then M M. Proof. By induction on the derivation of M M. Let us consider the last rule r. If r = var then M = x M 1... M m and M = x M 1... M m where m N and M i M i for any 1 i m. By reflexivity of (Remark 10), x x. By induction hypothesis, M i M i for any 1 i m. Therefore, M M by applying Lemma 16 m times. If r = λ then M = (λx.m 0 )M 1... M m and M = (λx.m 0)M 1... M m where m N and M i M i for any 0 i m. By induction hypothesis, M i M i for any 1 i m. According to Remark 11.4, λx.m 0 λx.m 0. So M M by applying Lemma 16 m times. If r = β v then M = (λx.m 0 )V M 1... M m and M = M 0{V /x}m 1... M m where m N, V V and M i M i for any 0 i m. By induction hypothesis, V V and M i M i for any 0 i m. By applying the rule β v for βv, one has M βv M 0 {V/x}M 1... M m ; moreover M 0 {V/x}M 1... M m M by Lemma 18 and by applying Lemma 16 m times, thus there are L, N Λ such that M βv M 0 {V/x}M 1... M m β v L σ N M. So M M. If r = σ 1 then M = (λx.m 0 )N 0 L 0 M 1... M m and M = (λx.m 0L 0)N 0M 1... M m where m N, L 0 L 0, N 0 N 0 and M i M i for any 0 i m. By induction hypothesis, N N and M i M i for any 1 i m. By applying the rule σ 1 for σ, one has M σ (λx.m 0 L 0 )N 0 M 1... M m. By Remark 12.1, M 0 L 0 M 0L 0 and thus λx.m 0 L 0 λx.m 0L 0 according to Remark So (λx.m 0 L 0 )N 0 M 1... M m M by applying Lemma 16 m + 1 times, hence there are L, N Λ such that M σ (λx.m 0 L 0 )N 0 M 1... M m β v L σ N M. By Lemma 14.2, there is L Λ such that M β v L σ L σ N M and thus M M.

9 G. Guerrieri, L. Paolini and S. Ronchi Della Rocca 9 Finally, if r = σ 3 then M = V ((λx.l 0 )N 0 )M 1... M m and M = (λx.v L 0)N 0M 1... M m with m N, V V, L 0 L 0, N 0 N 0 and M i M i for any 1 i m. By induction hypothesis, N 0 N 0 and M i M i for any 1 i m. By the rule σ 3 for σ, one has M σ (λx.v L 0 )N 0 M 1... M m. By Remark 12.1, V L 0 V L 0 and thus λx.v L 0 λx.v L 0 according to Remark So (λx.v L 0 )N 0 M 1... M m M by applying Lemma 16 m + 1 times, hence there are L, N Λ such that M σ (λx.(λy.m 0 )L 0 )N 0 M 1... M m β v L σ N M. By Lemma 14.2, there is L Λ such that M β v L σ L σ N M, so M M. Next Lemma 20 and Corollary 21 show that ernal parallel reduction can be shifted after head reductions. Lemma 20 (Postponement). If M L and L βv N (resp. L σ N) then there exist L Λ such that M βv L (resp. M σ L ) and L N. Proof. By induction on the derivation of M L. Let us consider its last rule r. If r = var, then M = x = L which contradicts L βv N and L σ N by Remarks If r = λ then L = λx.l for some L Λ, which contradicts L βv N and L σ N by Remarks Finally, if r = right then M = V M 0 M 1... M m and L = V L 0 L 1... L m where m N, V V (so V V by Remark 11.5), M 0 L 0 (so M 0 L 0 ) and M i L i for any 1 i m. If L βv N then there are two cases, depending on the last rule r of the derivation of L βv N. If r = β v then V = λx.n 0, L 0 Λ v and N = N 0{L 0 /x}l 1... L m, thus M 0 Λ v and V = λx.n 0 with N 0 N 0 by Remark By Lemma 13, one has N 0 {M 0 /x} N 0{L 0 /x}. Let L = N 0 {M 0 /x}m 1... M m : so M = (λx.n 0 )M 0 M 1... M m βv L (apply the rule β v for βv ) and L N by applying Remark 12.1 m times. If r = right then N = V N 0 L 1... L m with L 0 βv N 0. By induction hypothesis, there exists L 0 Λ such that M 0 βv L 0 N 0. Let L = V L 0M 1... M m : so M βv L (apply the rule right for βv ) and L N by applying Remark 12.1 m + 1 times. If L σ N then there are three cases, depending on the last rule r of the derivation of L σ N. If r = σ 1 then m > 0, V = λx.n 0 and N = (λx.n 0L 1 )L 0 L 2... L m, thus V = λx.n 0 with N 0 N 0 by Remark Thanks to Remarks 12.1 and 11.4, one has λx.n 0 M 1 λx.n 0L 1. Let L = (λx.n 0 M 1 )M 0 M 2... M m : so M = (λx.n 0 )M 0 M 1... M m σ L (apply the rule σ 1 for σ ) and L N by applying Remark 12.1 m times. If r = σ 3 then L 0 = (λx.l 01 )L 02 and N = (λx.v L 01 )L 02 L 1... L m. Since M 0 (λx.l 01 )L 02, necessarily M 0 = (λx.m 01 )M 02 with M 01 L 01 and M 02 L 02 (so M 02 L 02 ). Thanks to Remarks 12.1 and 11.4, one has λx.v M 01 λx.v L 01. Let L = (λx.v M 01 )M 02 M 1... M m : therefore M = V ((λx.m 01 )M 02 )M 1... M m σ L (apply the rule σ 3 for σ ) and L N by applying Remark 12.1 m + 1 times. If r = right then N = V N 0 L 1... L m with L 0 σ N 0. By induction hypothesis, there exists L 0 Λ such that M 0 σ L 0 N 0. Let L = V L 0M 1... M m : so M σ L (apply the rule right for σ ) and L N by applying Remark 12.1 m + 1 times. Corollary 21. If M L and either L βv such that M β v L σ L N. N or L σ N, then there exist L, L Λ Proof. Immediate by Lemma 20 and Lemma 19, applying Lemma 14.2 if L σ N.

10 10 Standardization of a call-by-value calculus Now we obtain our first main result (Theorem 22): any v-reduction sequence can be sequentialized o an head β v -reduction sequence followed by an head σ-reduction sequence, followed by an ernal reduction sequence. In ordinary λ-calculus, the well-known result corresponding to our Theorem 22 says that a β-reduction sequence can be factorized in a head reduction sequence followed by an ernal reduction sequence (see for example [19, Corollary 2.6]). Theorem 22 (Sequentialization). If M v M then there exist L, N Λ such that M β v L σ N v M. Proof. By Remark 12.4, M M and thus there are m N and M 0,..., M m Λ such that M = M 0, M m = M and M i M i+1 for any 0 i < m. We prove by induction on m N that there are L, N Λ such that M β v L σ N M, so N v M by Remark If m = 0 then M = M 0 = M and hence we conclude by taking L = M = N. Finally, suppose m > 0. By induction hypothesis applied to M 1 M, there exist L, N Λ such that M 1 β v L σ N M. By applying Lemma 19 to M, there exist L 0, N 0 Λ such that M β v L 0 σ N 0 M 1. By applying Corollary 21 repeatedly, there exists N Λ such that N 0 β N vσ N and hence M β N vσ M. According to Lemma 14.3, there exists L Λ such that M β v L σ N M. It is worth noticing that in Definition 7 there is no distinction between head σ 1 - and head σ 3 -reduction steps, and, according to it, the sequentialization of Theorem 22 imposes no order between head σ-reductions. We note σ1 and σ3 respectively the reductions relations σ1 σ and σ3 σ. So, a natural question arises: is it possible to sequentialize them? The answer is negative, as proved by the next two counterexamples. Let M = x((λy.z )(zi)) and N = (λy.xz )(zi): M σ3 (λy.xz )(zi) σ1 N, but there exists no L such that M σ 1 L σ 3 N. In fact M contains only a head σ 3 -redex and (λy.xz )(zi) contains only a head σ 1 -redex Let M = x((λy.z )(zi) ) and N = (λy.x(z ))(zi): M σ1 x((λy.z )(zi)) σ3 N but there is no L such that M σ 3 L σ 1 N. In fact M contains only a head σ 1 -redex and x((λy.z )(zi)) contains only a head σ 3 -redex. So, the impossibility of sequentializing a head σ-reduction sequence is due to the fact that a head σ 1 -reduction step can create a head σ 3 -redex, and viceversa. This is not a problem, since head σ-reduction is strongly normalizing (by Proposition 4 and since σ σ ). Our approach does not force a strict order of head σ-reductions. 4 Standardization Now we are able to prove the main result of this paper, i.e., a standardization theorem for λ σ v (Theorem 25). In particular we provide a notion of standard reduction sequence that avoid any auxiliary notion of residual redexes, by closely following the definition given in [16]. Notation. For any k, m N with k m, we note M 0,, M k, M m head a sequence of terms such that M i βv M i+1 when 0 i < k, and M i σ M i+1 when k i < m. It is easy to check that M head for any M Λ. The notion of standard sequence of terms is defined by using the previous notion of head-sequence. Our notion of standard reduction sequence is mutually defined together with the notion of inner-sequence of terms (Definition 23). This definition allows us to avoid non-deterministic cases remarked in [7] (we provide more details at the end of this section). We denote by M 0,, M m std (resp. M 0,, M m in ) a standard (resp. inner) sequence of terms.

11 G. Guerrieri, L. Paolini and S. Ronchi Della Rocca 11 Definition 23 (Standard and inner sequences). Standard and inner sequences of terms are defined by mutual induction as follows: 1. if M 0,, M m head and M m,, M m+n in then M 0,, M m+n std, where m, n N; 2. M in, for any M Λ; 3. if M 0,, M m std then λz.m 0,, λz.m m in, where m N; 4. if V 0,, V h std and N 0,, N n in then V 0 N 0,, V 0 N n,, V h N n in, where h, n N; 5. if N 0,, N n in, M 0,, M m std and N 0 Λ v, then N 0 M 0,, N n M 0,, N n M m in, where m, n N. For instance, let M = (λy.ix)(z( I))(II): M v (λy.ix)(z( I))I v (λy.x)(z( I))I and M v (λy.ix(ii))(z( I)) v (λy.ix(ii))(z(ii)) are not standard sequences; M v (λy.ix)(z( I))I and M v (λy.ix)(z(ii))(ii) v (λy.ix)(zi)(ii) v (λy.ix(ii))(zi) v (λy.x(ii))(zi) v (λy.xi)(zi) are standard sequences. Remark 24. For any n N, if N 0,, N n in (resp. N 0,, N n head ) then N 0,, N n std. Indeed, N 0 head (resp. N n in by Definition 23.2), so N 0,, N n std by Definition In particular, N std for any N Λ: apply Definition 23.2 and Remark 24 for n = 0. Theorem 25 (Standardization). 1. If M v M then there is a sequence M,, M std. 2. If M v M then there is a sequence M,, M in. Proof. Both statements are proved simultaneously by induction on M Λ. 1. If M = z then, by Theorem 22, M β v L σ N v z for some L, N Λ. By Remarks 12.5 and 11.2, L = N = z; therefore M β v z and hence there is a sequence M,, z head. Thus, M,, z std by Remark 24. If M = λz.n then, by Theorem 22, M β v L σ L v λz.n for some L, L Λ. By Remarks 12.5 and 11.2, L = L = λz.n with N v N. So M β v λz.n and hence there is a sequence M,, λz.n head. By induction on (1), there is a sequence N,, N std, thus λz.n,, λz.n in by Definition Therefore M,, λz.n,, λz.n std by Definition If M = N L then, by Theorem 22, M β v M σ M 0 v N L for some M, M 0 Λ. By Remark 3, M 0 = NL for some N, L Λ, since v v and M / Λ v. Thus there is a sequence M,, M,, NL head. By Remark 12.5, NL N L ; clearly, each step of is an instance of the rule right of Definition 9. There are two sub-cases. If N Λ v then N N and L L, so N v N and L v L by Remarks By induction respectively on (1) and (2), there are sequences N,, N std and L,, L in, thus NL,, NL,, N L in by Definition Therefore M,, M,, NL,, NL,, N L std by Definition If N / Λ v (i.e., N =VM 1... M m with m>0) then N N and L L, so N v N and L v L by Remarks By induction respectively on (2) and (1), there are sequences N,, N in and L,, L std. Hence NL,, N L,, N L in by Definition Thus M,, M,, NL,, N L,, N L std by Definition If M = z then M = z by Remark 12.5, hence z in by Definition If M = λz.l then M = λz.l and L v L by Remark Hence there is a sequence L,, L std by induction on (1). By Definition 23.3, λz.l,, λz.l in. If M = N L then M = NL for some N, L Λ by Remark 3, since v v and M / Λ v. By Remark 12.5, NL N L ; clearly, each step of is an instance of the rule right of Definition 9. There are two sub-cases.

12 12 Standardization of a call-by-value calculus If N Λ v then N N and L L, so N v N and L v L by Remarks Thus there are sequences N,, N std and L,, L in by induction respectively on (1) and (2). Therefore, by Definition 23.4, NL,, NL,, N L in. If N / Λ v (i.e. N = VM 1... M m with m > 0) then N N and L L, thus N v N and L v L by Remarks By induction respectively on (2) and (1), there are sequences N,, N in and L,, L std. So NL,, N L,, N L in by Definition Due to non-sequentialization of head σ 1 - and head σ 3 -reductions, several standard sequences may have the same starting term and ending term: for instance, if M = I( I)I and N = (λz.(λx.xi)(zz))i then M v (λx.xi)( I) v N and M v (λz.i(zz))ii v (λz.i(zz)i)i v N are both standard sequences from M to N. Finally, we can compare our notion of standardization with that given in [16]. To make the comparison possible we avoid σ-reductions and we recall that βv is exactly the Plotkin s left-reduction [16, p. 146]. As remarked in [7, 1.5 p. 149], both sequences (λz.ii)(ii) v (λz.i)(ii) v (λz.i)i and (λz.ii)(ii) v (λz.ii)i v (λz.i)i are standard according to [16]. On the other hand, only the second sequence is standard in our sense. It is easy to check that collapsing the two notions of inner and standard sequence given in Definition 23, we get a notion of standard sequence that accept both the above sequences. 5 Some conservativity results The sequentialization result (Theorem 22) has some eresting semantic consequences. It allows us to prove that the λ σ v -calculus is sound with respect to the call-by-value observational equivalence roduced by Plotkin in [16]. Moreover we can prove that some notions, like that of potential valuability and solvability, roduced in [15] for λ v, coincide with the same notions in λ σ v. This justify the idea that λ σ v is an useful tool for studying the properties of λ v, which was inspiring its definition in [3]. Our starting po is the following corollary. Corollary If M v V Λ v then there exists V Λ v such that M β v V v V. 2. For any V Λ v, M β v V if and only if M β V. vσ Proof. The first po is proved by observing that, by Theorem 22, there are N, L Λ such that M β v L σ N v V. By Remarks 12.5, N Λ v and thus L = N according to Remark The second po is a consequence of Lemma 14.3 and Remark Let us recall the notion of observational equivalence defined by Plotkin [16]. Definition 27 (Halting, observational equivalence). Let M Λ. We say that ( the evaluation of) M halts if there exists V Λ v such that M β v V. The ( call-by-value) observational equivalence is an equivalence relation = on Λ defined by: M = N if, for every context C, one has that C M halts iff C N halts. 1 Clearly, similar notions can be defined for λ σ v using βvσ instead of βv. Corollary 26 proves that head σ-reduction plays no role neither in deciding the halting problem for evaluation, nor in reaching a particular value. So we can conclude that the observational equivalence in λ σ v is conservative with respect to that of λ v. 1 Original Plotkin s definition of call-by-value observational equivalence (see [16]) also requires that C M and C N are closed terms, according to the tradition identifying programs with closed terms.

13 G. Guerrieri, L. Paolini and S. Ronchi Della Rocca 13 Theorem 28 (Adequacy of v-reduction). If M v M then M halts iff M halts. Proof. If M β v V Λ v then M v M v V since βv v. By Corollary 26.1, there exists V Λ v such that M β v V. Thus M halts. Conversely, if M β v V Λ v then M v V since βv v. By confluence of v (Proposition 4, since M v M ) and Remark 3 (since V Λ v ), there is V Λ v such that V v V and M v V. By Corollary 26.1, there is V Λ v such that M β v V. So M halts. Corollary 29 (Soundness). If M = v N then M = N. Proof. Let C be a context. By confluence of v (Proposition 4), M = v N implies that there exists L Λ such that M v L and N v L, hence C M v C L and C N v C L. By Theorem 28, C M halts iff C L halts iff C N halts. Therefore, M = N. Plotkin [16] has already proved that M = βv N implies M = N, but our Corollary 29 is not obvious since our λ σ v -calculus equates more than Plotkin s λ v -calculus (see Example 5). The converse of Corollary 29 does not hold since λx.x(λy.xy) = but λx.x(λy.xy) and are different v-normal forms and hence λx.x(λy.xy) v. A further remarkable consequence of Corollary 26.1 is that the notions of potential valuability and solvability for λ σ v -calculus (studied in [3]) can be shown coincide with the ones for Plotkin s λ v -calculus (studied in [14, 15]), respectively. Let us recall their definition. Definition 30 (Potential valuability, solvability). Let M be a term: M is v-potentially valuable (resp. β v -potentially valuable) if there are m N, pairwise distinct variables x 1,..., x m and V, V 1,..., V m Λ v such that M{V 1 /x 1,..., V m /x m } v V (resp. M{V 1 /x 1,..., V m /x m } β v V ); M is v-solvable (resp.β v -solvable) if there are n,m N, variables x 1,..., x m and N 1,...,N n Λ such that (λx 1... x m.m)n 1 N n v I (resp. (λx 1... x m.m)n 1 N n β v I). Theorem 31. Let M be a term: 1. M is v-potentially valuable if and only if M is β v -potentially valuable; 2. M is v-solvable if and only if M is β v -solvable. Proof. In both pos, the implication from right to left is trivial since βv v. Let us prove the other direction. 1. Since M is v-potentially valuable, there are variables x 1,..., x m and V, V 1,..., V m Λ v (with m 0) such that M{V 1 /x 1,..., V m /x m } v V ; then, there exist V Λ v such that M{V 1 /x 1,..., V m /x m } β v V by Corollary 26.1 and because βv βv, therefore M is β v -potentially valuable. 2. Since M is v-solvable, there exist variables x 1,..., x m and terms N 1,..., N n (for some n, m 0) such that (λx 1... x m.m)n 1 N n v I; then, there exists V Λ v such that (λx 1... x m.m)n 1 N n β v V v I by Corollary 26.1 and because βv βv. According to Remarks 12.5, V = I and therefore M is β v -solvable. So, thanks to Theorem 31, the semantic (via a relational model) and operational (via two sub-reductions of v ) characterization of v-potential valuability and v-solvability given in [3, Theorems 24-25] is also a semantic and operational characterization of β v -potential valuability and β v -solvability. The difference is that in λ σ v these notions can be studied inside the calculus, while it has been proved in [14, 15] that the β v -reduction is too weak in order to characterize them. So λ σ v is a useful tool for studying semantic properties of λ v.

14 14 Standardization of a call-by-value calculus 6 Conclusions In this paper we have proved a standardization theorem for the λ σ v -calculus, roduced in [3]. The used technique is a notion of parallel reduction. Let us recall that parallel reduction in λ-calculus has been defined by Tait and Martin-Löf in order to prove confluence of the β-reduction, without referring to the difficult notion of residuals. Takahashi in [19] has simplified this technique and showed that it can be successfully applied to standardization. We would like to remark that our parallel reduction cannot be used to prove confluence of v. Indeed, take M = (λx.l) ( (λy.n)((λz.n )N ) ) L, M 1 = (λx.ll ) ( (λy.n)((λz.n )N ) ) and M 2 = ( (λy.(λx.l)n)((λz.n )N ) ) L : then M M 1 and M M 2 but there is no term M such that M 1 M and M 2 M. To sum up, does not enjoy the Diamond Property. The standardization result allows us to formally verify the correctness of λ σ v with respect to the semantics of λ v, so we can use λ σ v as a tool for studying properties of λ v. This is a remarkable result: in fact some properties, like potential valuability and solvability, cannot be characterized in λ v by means of β v -reduction (as proved in [14, 15]), but they have a natural operational characterization in λ σ v (via two sub-reductions of v ). We plan to continue to explore the call-by-value computation, using λ σ v. As a first step, we would like to revisit and improve the Separability Theorem given in [12] for λ v. Still the issue is more complex than in the call-by-name, indeed in ordinary λ-calculus different βη-normal forms can be separated (by the Böhm Theorem), while in λ v there are different normal forms that cannot be separated, but which are only semi-separable (e.g. I and λz.(λu.z)(zz)). We hope to completely characterize separable and semi-separable normal forms in λ σ v. This should be a first step aimed to define a semantically meaningful notion of approximants. Then, we should be able to provide a new insight on the denotational analysis of the call-by-value, maybe overcoming limitations as that of the absence of fully abstract filter models [15, Theorem ]. Last but not least, an unexplored but challenging research direction is the use of commutation rules to improve the call-by-value evaluation. We do not have concrete evidence supporting such possibility, but since λ σ v is strongly related to the calculi presented in [7, 1], which are endowed with explicit substitutions, we are confident that a sharp use of commutations can have a relevant impact in the evaluation. References 1 Beniamino Accattoli and Luca Paolini. Call-by-Value Solvability, Revisited. In Tom Schrijvers and Peter Thiemann, editors, Functional and Logic Programming, volume 7294 of Lecture Notes in Computer Science, pages Springer Berlin Heidelberg, Henk Barendregt. The Lambda Calculus: Its Syntax and Semantics, volume 103 of Studies in logic and the foundation of mathematics. North Holland, Amsterdam, Alberto Carraro and Giulio Guerrieri. A Semantical and Operational Account of Call-by- Value Solvability. In Anca Muscholl, editor, Foundations of Software Science and Computation Structures, volume 8412 of Lecture Notes in Computer Science, pages Springer Berlin Heidelberg, April An extended version (with all the proofs) is available at 4 Karl Crary. A Simple Proof of Call-by-Value Standardization. Technical Report CMU-CS , Carneige Mellon University, Haskell B. Curry and Robert Feys. Combinatory Logic, volume 1. North Holland, Amsterdam, Jean-Yves Girard. Linear logic. Theoretical Computer Science, 50(1):1 102, Hugo Herbelin and Stéphane Zimmermann. An Operational Account of Call-by-Value Minimal and Classical lambda-calculus in "Natural Deduction" Form. In Pierre-Louis Curien,

15 G. Guerrieri, L. Paolini and S. Ronchi Della Rocca 15 editor, Typed Lambda Calculi and Applications, 9th International Conference, TLCA 2009, Lecture Notes in Computer Science, pages , Roger Hindley. Standard and normal reductions. Transactions of the American Mathematical Society, pages , John Maraist, Martin Odersky, David N. Turner, and Philip Wadler. Call-by-name, call-byvalue, call-by-need and the linear lambda calculus. Theoretical Computer Science, 228(1 2): , Eugenio Moggi. Computational lambda-calculus and monads. Technical report, Edinburgh University, Tech. Report ECS-LFCS Eugenio Moggi. Computational Lambda-Calculus and Monads. In Proceedings of the Fourth Annual Symposium on Logic in Computer Science (LICS 89), pages IEEE Computer Society, Luca Paolini. Call-by-Value Separability and Computability. In Antonio Restivo, Simona Ronchi Della Rocca, and Luca Roversi, editors, Proceedings of the 7th Italian Conference in Theoretical Computer Science - ICTCS 01 (Torino, Italy, October 4-6, 2001), volume 2202 of Lecture Notes in Computer Science, pages Springer-Verlag, Luca Paolini and Simona Ronchi Della Rocca. Parametric parameter passing lambdacalculus. Information and Computation, 189(1):87 106, February Luca Paolini and Simona Ronchi Della Rocca. Call-by-value Solvability. Theoretical Informatics and Applications, 33(6): , November RAIRO Series, EDP-Sciences. 15 Luca Paolini and Simona Ronchi Della Rocca. The Parametric λ-calculus: a Metamodel for Computation. Texts in Theoretical Computer Science: An EATCS Series. Springer-Verlag, Berlin, Gordon D. Plotkin. Call-by-name, call-by-value and the lambda-calculus. Theoretical Computer Science, 1(2): , Amr Sabry and Matthias Felleisen. Reasoning About Programs in Continuation-passing Style. SIGPLAN Lisp Poers, V(1): , January Amr Sabry and Matthias Felleisen. Reasoning about programs in continuation-passing style. Lisp and symbolic computation, 6(3-4): , Masako Takahashi. Parallel reductions in lambda-calculus. Information and Computation, 118(1): , April 1995.

16 16 Standardization of a call-by-value calculus A Technical appendix A.1 Proofs omitted in the article Stated at p. 6 Remark Certainly, v v. 5. Certainly, v v. Proof. 4. The proof that M v M implies M M is by induction on M Λ. Recall that M = V N 1... N n for some n N, V Λ v and N i Λ for any 1 i n. By reflexivity of (Remark 10), N i N i for any 1 i n. There are five cases: either n > 0, V = λx.n, N 1 Λ v and M = N{N 1 /x}n 2... N n, then N N by reflexivity of (Remark 10), so M M by applying the rule β v. or n > 1, V = λx.n and M = (λx.nn 2 )N 1 N 3... N n, then N N by reflexivity of (Remark 10), so M M by applying the rule σ 1. or n > 0, N 1 = (λx.l)n and M = (λx.v L)NN 2... N n, then V V, L L and N N (Remark 10), so M M by applying the rule σ 3. or M = V N 1... N n with V v V, then V = λx.n and V = λx.n with N v N by Remark 3, so N N by induction hypothesis; thus M M by the rule λ; or n > 0 and M = V N 1... N i... N n with N i v N i for some 1 i n, then N i N i by induction hypothesis, V V by reflexivity of (Remark 10); hence M M by applying the rule var or λ, according as V is a variable or an abstraction, respectively. The proof that M M implies M v M is by straightforward induction on the derivation of M M. 5. The proof that M v M implies M M is by induction on M Λ. Remember that M = V N 1... N n for some n N, V Λ v and N i Λ for any 1 i n. As M v M and M βvσ M, one has: either M = V N 1... N n with V v V, then V = λx.n and V = λx.n with N v N by Remark 3, so N N according to Remark 12.4, thus M M by the rule λ for ; or n > 0 and M = V N 1N 2... N n with N 1 v N 1, then N 1 N 1 by induction hypothesis, and V V and N i N i for any 2 i n by reflexivity of (Remark 10); hence M M by applying the rule right for ; or n > 1 and M = V N 1... N i... N n with N i v N i for some 2 i n, then N i N i by Remark 12.4, V V and N j N j for any 2 j n with j i and N 1 N 1 by reflexivity of and (Remark 10); hence M M by applying the rule right for. The proof that M M implies M derivation of M M. v M is by straightforward induction on the Stated at p. 6 Lemma 13 (Substitution lemma for ). If M M and V V then M{V/x} M {V /x}. Proof. By induction on the derivation of M M. Let us consider the last rule r. If r = var then M = y M 1... M m and M = y M 1... M m with m N and M i M i for any 1 i m. By induction hypothesis, M i {V/x} M i {V /x} for any 1 i m. If y x

Standardization of a call-by-value lambda-calculus

Standardization of a call-by-value lambda-calculus Standardization of a call-by-value lambda-calculus Giulio Guerrieri 1, Luca Paolini 2, and Simona Ronci Della Rocca 2 1 Laboratoire PPS, UMR 7126, Université Paris Diderot, Sorbonne Paris Cité F-75205

More information

Lazy Strong Normalization

Lazy Strong Normalization Lazy Strong Normalization Luca Paolini 1,2 Dipartimento di Informatica Università di Torino (ITALIA) Elaine Pimentel 1,2 Departamento de Matemática Universidade Federal de Minas Gerais (BRASIL) Dipartimento

More information

The Lambda-Calculus Reduction System

The Lambda-Calculus Reduction System 2 The Lambda-Calculus Reduction System 2.1 Reduction Systems In this section we present basic notions on reduction systems. For a more detailed study see [Klop, 1992, Dershowitz and Jouannaud, 1990]. Definition

More information

arxiv: v2 [cs.lo] 23 Nov 2016

arxiv: v2 [cs.lo] 23 Nov 2016 STANDARDIZATION AND CONSERVATIVITY OF A REFINED CALL-BY-VALUE LAMBDA-CALCULUS arxiv:1611.07255v2 [cs.lo] 23 Nov 2016 GIULIO GUERRIERI, LUCA PAOLINI, AND SIMONA RONCHI DELLA ROCCA Aix Marseille Université,

More information

Call-by-value non-determinism in a linear logic type discipline

Call-by-value non-determinism in a linear logic type discipline Call-by-value non-determinism in a linear logic type discipline Alejandro Díaz-Caro? Giulio Manzonetto Université Paris-Ouest & INRIA LIPN, Université Paris 13 Michele Pagani LIPN, Université Paris 13

More information

Variations on a theme: call-by-value and factorization

Variations on a theme: call-by-value and factorization Variations on a theme: call-by-value and factorization Beniamino Accattoli INRIA & LIX, Ecole Polytechnique Accattoli (INRIA Parsifal) Variations on a theme: call-by-value and factorization 1 / 31 Outline

More information

summer school Logic and Computation Goettingen, July 24-30, 2016

summer school Logic and Computation Goettingen, July 24-30, 2016 Università degli Studi di Torino summer school Logic and Computation Goettingen, July 24-30, 2016 A bit of history Alonzo Church (1936) The as formal account of computation. Proof of the undecidability

More information

Call-by-Value Non-determinism in a Linear Logic Type Discipline

Call-by-Value Non-determinism in a Linear Logic Type Discipline Call-by-Value Non-determinism in a Linear Logic Type Discipline Alejandro Díaz-Caro 1,, Giulio Manzonetto 1,2, and Michele Pagani 1,2 1 Université Paris 13, Sorbonne Paris Cité, LIPN, F-93430, Villetaneuse,

More information

A Terminating and Confluent Linear Lambda Calculus

A Terminating and Confluent Linear Lambda Calculus A Terminating and Confluent Linear Lambda Calculus Yo Ohta and Masahito Hasegawa Research Institute for Mathematical Sciences, Kyoto University Kyoto 606-8502, Japan Abstract. We present a rewriting system

More information

Parametric parameter passing -calculus

Parametric parameter passing -calculus Information and Computation 189 (2004) 87 106 www.elsevier.com/locate/ic Parametric parameter passing -calculus Luca Paolini a and Simona Ronchi Della Rocca b, *,1 a Università di Genova - DISI, Université

More information

Call-by-Value Non-determinism in a Linear Logic Type Discipline

Call-by-Value Non-determinism in a Linear Logic Type Discipline Call-by-Value Non-determinism in a Linear Logic Type Discipline Alejandro Díaz-Caro 1,, Giulio Manzonetto 1,2, and Michele Pagani 1,2 1 Université Paris 13, Sorbonne Paris Cité, LIPN, F-93430, Villetaneuse,

More information

Equivalence of Algebraic λ -calculi extended abstract

Equivalence of Algebraic λ -calculi extended abstract Equivalence of Algebraic λ -calculi extended abstract Alejandro Díaz-Caro LIG, Université de Grenoble, France Alejandro.Diaz-Caro@imag.fr Christine Tasson CEA-LIST, MeASI, France Christine.Tasson@cea.fr

More information

Parallel Reduction in Type Free Lambda-mu- Calculus

Parallel Reduction in Type Free Lambda-mu- Calculus 九州大学学術情報リポジトリ Kyushu University Institutional Repository Parallel Reduction in Type Free Lambda-mu- Calculus Baba, Kensuke Graduate School of Information Science and Electrical Engineering, Kyushu University

More information

The Minimal Graph Model of Lambda Calculus

The Minimal Graph Model of Lambda Calculus The Minimal Graph Model of Lambda Calculus Antonio Bucciarelli 1 and Antonino Salibra 2 1 Université Paris 7, Equipe PPS, 2 place Jussieu, 72251 Paris Cedex 05, France buccia@pps.jussieu.fr, 2 Università

More information

The Permutative λ-calculus

The Permutative λ-calculus The Permutative λ-calculus Beniamino Accattoli 1 Delia Kesner 2 INRIA and LIX (École Polytechnique) PPS (CNRS and Université Paris-Diderot) 1 / 34 Outline 1 λ-calculus 2 Permutative extension 3 Confluence

More information

Completeness and Partial Soundness Results for Intersection & Union Typing for λµ µ

Completeness and Partial Soundness Results for Intersection & Union Typing for λµ µ Completeness and Partial Soundness Results for Intersection & Union Typing for λµ µ Steffen van Bakel Department of Computing, Imperial College London, 180 Queen s Gate, London SW7 2BZ, UK Abstract This

More information

Call-by-value solvability, revisited

Call-by-value solvability, revisited Call-by-value solvability, revisited Beniamino Accattoli 1 and Luca Paolini 2 1 INRIA and LIX (École Polytechnique), France 2 Dipartimento di Informatica, Università degli Studi di Torino, Italy Abstract.

More information

SHARING IN THE WEAK LAMBDA-CALCULUS REVISITED

SHARING IN THE WEAK LAMBDA-CALCULUS REVISITED SHARING IN THE WEAK LAMBDA-CALCULUS REVISITED TOMASZ BLANC, JEAN-JACQUES LÉVY, AND LUC MARANGET INRIA e-mail address: tomasz.blanc@inria.fr INRIA, Microsoft Research-INRIA Joint Centre e-mail address:

More information

Parametric λ-theories

Parametric λ-theories Parametric λ-theories Luca Paolini 1 Dipartimento di Informatica Università degli Studi di Torino Corso Svizzera, 185 10149 Torino - Italy e-mail: paolini@di.unito.it Abstract The parametric lambda calculus

More information

On Upper Bounds on the Church-Rosser Theorem

On Upper Bounds on the Church-Rosser Theorem On Upper Bounds on the Church-Rosser Theorem Ken-etsu Fujita Department of Computer Science Gunma University Kiryu, Japan fujita@cs.gunma-u.ac.jp The Church-Rosser theorem in the type-free λ-calculus is

More information

REVISITING CALL-BY-VALUE BÖHM TREES IN LIGHT OF THEIR TAYLOR EXPANSION

REVISITING CALL-BY-VALUE BÖHM TREES IN LIGHT OF THEIR TAYLOR EXPANSION REVISITING CALL-BY-VALUE BÖHM TREES IN LIGHT OF THEIR TAYLOR EXPANSION EMMA KERINEC, GIULIO MANZONETTO, AND MICHELE PAGANI Université de Lyon, ENS de Lyon, Université Claude Bernard Lyon 1, LIP e-mail

More information

On the Standardization Theorem for λβη-calculus

On the Standardization Theorem for λβη-calculus On the Standardization Theorem for λβη-calculus Ryo Kashima Department of Mathematical and Computing Sciences Tokyo Institute of Technology Ookayama, Meguro, Tokyo 152-8552, Japan. e-mail: kashima@is.titech.ac.jp

More information

EasyChair Preprint. Normalization and Taylor expansion of lambda-terms

EasyChair Preprint. Normalization and Taylor expansion of lambda-terms EasyChair Preprint 165 Normalization and Taylor expansion of lambda-terms Federico Olimpieri EasyChair preprints are intended for rapid dissemination of research results and are integrated with the rest

More information

Diagrams for Meaning Preservation

Diagrams for Meaning Preservation Diagrams for Meaning Preservation 2003-06-13 Joe Wells Detlef Plump Fairouz Kamareddine Heriot-Watt University University of York www.cee.hw.ac.uk/ultra www.cs.york.ac.uk/ det Diagrams for Meaning Preservation

More information

Sequent Combinators: A Hilbert System for the Lambda Calculus

Sequent Combinators: A Hilbert System for the Lambda Calculus Sequent Combinators: A Hilbert System for the Lambda Calculus Healfdene Goguen Department of Computer Science, University of Edinburgh The King s Buildings, Edinburgh, EH9 3JZ, United Kingdom Fax: (+44)

More information

Lecture Notes on Combinatory Modal Logic

Lecture Notes on Combinatory Modal Logic Lecture Notes on Combinatory Modal Logic 15-816: Modal Logic Frank Pfenning Lecture 9 February 16, 2010 1 Introduction The connection between proofs and program so far has been through a proof term assignment

More information

Parallel Reduction in Type Free λµ-calculus

Parallel Reduction in Type Free λµ-calculus Electronic Notes in Theoretical Computer Science 42 (2001) URL: http://www.elsevier.nl/locate/entcs/volume42.html 15 pages Parallel Reduction in Type Free λµ-calculus Kensuke Baba 1 Graduate School of

More information

An Implicit Characterization of PSPACE

An Implicit Characterization of PSPACE An Implicit Characterization of PSPACE Marco Gaboardi Dipartimento di Scienze dell Informazione, Università degli Studi di Bologna - Mura Anteo Zamboni 7, 40127 Bologna, Italy, gaboardi@cs.unibo.it and

More information

Classical Combinatory Logic

Classical Combinatory Logic Computational Logic and Applications, CLA 05 DMTCS proc. AF, 2006, 87 96 Classical Combinatory Logic Karim Nour 1 1 LAMA - Equipe de logique, Université de Savoie, F-73376 Le Bourget du Lac, France Combinatory

More information

A Behavioural Model for Klop s Calculus

A Behavioural Model for Klop s Calculus Replace this file with prentcsmacro.sty for your meeting, or with entcsmacro.sty for your meeting. Both can be found at the ENTCS Macro Home Page. A Behavioural Model for Klop s Calculus Mariangiola Dezani-Ciancaglini

More information

Lecture Notes on Classical Linear Logic

Lecture Notes on Classical Linear Logic Lecture Notes on Classical Linear Logic 15-816: Linear Logic Frank Pfenning Lecture 25 April 23, 2012 Originally, linear logic was conceived by Girard [Gir87] as a classical system, with one-sided sequents,

More information

Solvability in Resource Lambda-Calculus

Solvability in Resource Lambda-Calculus Solvability in Resource Lambda-Calculus Michele Pagani and Simona Ronchi della Rocca Dipartimento di Informatica Università di Torino C.so Svizzera 185 10149 Torino (IT) {pagani,ronchi}@di.unito.it Abstract.

More information

A General Technique for Analyzing Termination in Symmetric Proof Calculi

A General Technique for Analyzing Termination in Symmetric Proof Calculi A General Technique for Analyzing Termination in Symmetric Proof Calculi Daniel J. Dougherty 1, Silvia Ghilezan 2 and Pierre Lescanne 3 1 Worcester Polytechnic Institute, USA, dd@cs.wpi.edu 2 Faculty of

More information

TERMS FOR NATURAL DEDUCTION, SEQUENT CALCULUS AND CUT ELIMINATION IN CLASSICAL LOGIC

TERMS FOR NATURAL DEDUCTION, SEQUENT CALCULUS AND CUT ELIMINATION IN CLASSICAL LOGIC TERMS FOR NATURAL DEDUCTION, SEQUENT CALCULUS AND CUT ELIMINATION IN CLASSICAL LOGIC SILVIA GHILEZAN Faculty of Engineering, University of Novi Sad, Novi Sad, Serbia e-mail address: gsilvia@uns.ns.ac.yu

More information

3.2 Equivalence, Evaluation and Reduction Strategies

3.2 Equivalence, Evaluation and Reduction Strategies 3.2 Equivalence, Evaluation and Reduction Strategies The λ-calculus can be seen as an equational theory. More precisely, we have rules i.e., α and reductions, for proving that two terms are intensionally

More information

Observability for Pair Pattern Calculi

Observability for Pair Pattern Calculi Observability for Pair Pattern Calculi Antonio Bucciarelli 1, Delia Kesner 2, and Simona Ronchi Della Rocca 3 1,2 Univ Paris Diderot, Sorbonne Paris Cité, PPS, UMR 7126, CNRS, Paris, France 3 Dipartimento

More information

arxiv: v1 [cs.lo] 29 May 2014

arxiv: v1 [cs.lo] 29 May 2014 An Introduction to the Clocked Lambda Calculus Jörg Endrullis, Dimitri Hendriks, Jan Willem Klop, and Andrew Polonsky VU University Amsterdam, The Netherlands Abstract We give a brief introduction to the

More information

Computational Soundness of a Call by Name Calculus of Recursively-scoped Records. UMM Working Papers Series, Volume 2, Num. 3.

Computational Soundness of a Call by Name Calculus of Recursively-scoped Records. UMM Working Papers Series, Volume 2, Num. 3. Computational Soundness of a Call by Name Calculus of Recursively-scoped Records. UMM Working Papers Series, Volume 2, Num. 3. Elena Machkasova Contents 1 Introduction and Related Work 1 1.1 Introduction..............................

More information

3.2 Reduction 29. Truth. The constructor just forms the unit element,. Since there is no destructor, there is no reduction rule.

3.2 Reduction 29. Truth. The constructor just forms the unit element,. Since there is no destructor, there is no reduction rule. 32 Reduction 29 32 Reduction In the preceding section, we have introduced the assignment of proof terms to natural deductions If proofs are programs then we need to explain how proofs are to be executed,

More information

Consequence Relations and Natural Deduction

Consequence Relations and Natural Deduction Consequence Relations and Natural Deduction Joshua D. Guttman Worcester Polytechnic Institute September 9, 2010 Contents 1 Consequence Relations 1 2 A Derivation System for Natural Deduction 3 3 Derivations

More information

Henk Barendregt and Freek Wiedijk assisted by Andrew Polonsky. Radboud University Nijmegen. March 5, 2012

Henk Barendregt and Freek Wiedijk assisted by Andrew Polonsky. Radboud University Nijmegen. March 5, 2012 1 λ Henk Barendregt and Freek Wiedijk assisted by Andrew Polonsky Radboud University Nijmegen March 5, 2012 2 reading Femke van Raamsdonk Logical Verification Course Notes Herman Geuvers Introduction to

More information

The Call-by-Need Lambda Calculus

The Call-by-Need Lambda Calculus The Call-by-Need Lambda Calculus John Maraist, Martin Odersky School of Computer and Information Science, University of South Australia Warrendi Road, The Levels, Adelaide, South Australia 5095 fmaraist,oderskyg@cis.unisa.edu.au

More information

Partially commutative linear logic: sequent calculus and phase semantics

Partially commutative linear logic: sequent calculus and phase semantics Partially commutative linear logic: sequent calculus and phase semantics Philippe de Groote Projet Calligramme INRIA-Lorraine & CRIN CNRS 615 rue du Jardin Botanique - B.P. 101 F 54602 Villers-lès-Nancy

More information

Extended Abstract: Reconsidering Intuitionistic Duality

Extended Abstract: Reconsidering Intuitionistic Duality Extended Abstract: Reconsidering Intuitionistic Duality Aaron Stump, Harley Eades III, Ryan McCleeary Computer Science The University of Iowa 1 Introduction This paper proposes a new syntax and proof system

More information

Reducibility proofs in λ-calculi with intersection types

Reducibility proofs in λ-calculi with intersection types Reducibility proofs in λ-calculi with intersection types Fairouz Kamareddine, Vincent Rahli, and J. B. Wells ULTRA group, Heriot-Watt University, http://www.macs.hw.ac.uk/ultra/ March 14, 2008 Abstract

More information

Sub-λ-calculi, Classified

Sub-λ-calculi, Classified Electronic Notes in Theoretical Computer Science 203 (2008) 123 133 www.elsevier.com/locate/entcs Sub-λ-calculi, Classified François-Régis Sinot Universidade do Porto (DCC & LIACC) Rua do Campo Alegre

More information

Constructive approach to relevant and affine term calculi

Constructive approach to relevant and affine term calculi Constructive approach to relevant and affine term calculi Jelena Ivetić, University of Novi Sad, Serbia Silvia Ghilezan,University of Novi Sad, Serbia Pierre Lescanne, University of Lyon, France Silvia

More information

Origin in Mathematical Logic

Origin in Mathematical Logic Lambda Calculus Origin in Mathematical Logic Foundation of mathematics was very much an issue in the early decades of 20th century. Cantor, Frege, Russel s Paradox, Principia Mathematica, NBG/ZF Origin

More information

Subtyping and Intersection Types Revisited

Subtyping and Intersection Types Revisited Subtyping and Intersection Types Revisited Frank Pfenning Carnegie Mellon University International Conference on Functional Programming (ICFP 07) Freiburg, Germany, October 1-3, 2007 Joint work with Rowan

More information

Characterizing Polynomial and Exponential Complexity Classes in Elementary Lambda-Calculus

Characterizing Polynomial and Exponential Complexity Classes in Elementary Lambda-Calculus Characterizing Polynomial and Exponential Complexity Classes in Elementary Lambda-Calculus Patric Baillot 1, Eria De Benedetti 1,2, and Simona Ronchi Della Rocca 2 1 CNRS, ENS de Lyon, INRIA, UCBL, Université

More information

Commutative Locative Quantifiers for Multiplicative Linear Logic

Commutative Locative Quantifiers for Multiplicative Linear Logic Commutative Locative Quantifiers for Multiplicative Linear Logic Stefano Guerrini 1 and Patrizia Marzuoli 2 1 Dipartimento di Informatica Università degli Studi Roma La Sapienza Via Salaria, 113 00198

More information

arxiv: v1 [cs.lo] 27 Dec 2018

arxiv: v1 [cs.lo] 27 Dec 2018 Towards a Semantic Measure of the Execution Time in Call-by-Value lambda-calculus (Long Version) Giulio Guerrieri Dipartimento di Informatica Scienza e Ingegneria (DISI), Università di Bologna, Bologna,

More information

The duality of computation

The duality of computation The duality of computation (notes for the 3rd International Workshop on Higher-Order Rewriting) Hugo Herbelin LIX - INRIA-Futurs - PCRI Abstract. The work presented here is an extension of a previous work

More information

Step-Indexed Biorthogonality: a Tutorial Example

Step-Indexed Biorthogonality: a Tutorial Example Step-Indexed Biorthogonality: a Tutorial Example Andrew Pitts University of Cambridge Computer Laboratory 1 Introduction The purpose of this note is to illustrate the use of step-indexing [2] combined

More information

Lambda-Calculus (I) 2nd Asian-Pacific Summer School on Formal Methods Tsinghua University, August 23, 2010

Lambda-Calculus (I) 2nd Asian-Pacific Summer School on Formal Methods Tsinghua University, August 23, 2010 Lambda-Calculus (I) jean-jacques.levy@inria.fr 2nd Asian-Pacific Summer School on Formal Methods Tsinghua University, August 23, 2010 Plan computation models lambda-notation bound variables conversion

More information

Intersection Types for

Intersection Types for Intersection Types for -Trees Steffen van Bakel Franco Barbanera Mariangiola Dezani-Ciancaglini Fer-Jan de Vries Department of Computing, Imperial College, 180 Queen s Gate, London SW7 2BZ, UK E-mail:

More information

FROM COHERENT TO FINITENESS SPACES

FROM COHERENT TO FINITENESS SPACES FROM COHERENT TO FINITENESS SPACES PIERRE HYVERNAT Laboratoire de Mathématiques, Université de Savoie, 73376 Le Bourget-du-Lac Cedex, France. e-mail address: Pierre.Hyvernat@univ-savoie.fr Abstract. This

More information

Intersection Synchronous Logic

Intersection Synchronous Logic UnB 2007 p. 1/2 Intersection Synchronous Logic Elaine Gouvêa Pimentel Simona Ronchi della Rocca Luca Roversi UFMG/UNITO, 2007 UnB 2007 p. 2/2 Outline Motivation UnB 2007 p. 2/2 Outline Motivation Intuitionistic

More information

Simply Typed λ-calculus

Simply Typed λ-calculus Simply Typed λ-calculus Lecture 1 Jeremy Dawson The Australian National University Semester 2, 2017 Jeremy Dawson (ANU) COMP4630,Lecture 1 Semester 2, 2017 1 / 23 A Brief History of Type Theory First developed

More information

Calculi for Intuitionistic Normal Modal Logic

Calculi for Intuitionistic Normal Modal Logic Calculi for Intuitionistic Normal Modal Logic Yoshihiko Kakutani Department of Information Science, University of Tokyo kakutani@is.s.u-tokyo.ac.jp Abstract This paper provides a call-by-name and a call-by-value

More information

COMP6463: λ-calculus

COMP6463: λ-calculus COMP6463: λ-calculus 1. Basics Michael Norrish Michael.Norrish@nicta.com.au Canberra Research Lab., NICTA Semester 2, 2015 Outline Introduction Lambda Calculus Terms Alpha Equivalence Substitution Dynamics

More information

Investigation of Prawitz s completeness conjecture in phase semantic framework

Investigation of Prawitz s completeness conjecture in phase semantic framework Investigation of Prawitz s completeness conjecture in phase semantic framework Ryo Takemura Nihon University, Japan. takemura.ryo@nihon-u.ac.jp Abstract In contrast to the usual Tarskian set-theoretic

More information

Mathematical Synthesis of Equational Deduction Systems. Marcelo Fiore. Computer Laboratory University of Cambridge

Mathematical Synthesis of Equational Deduction Systems. Marcelo Fiore. Computer Laboratory University of Cambridge Mathematical Synthesis of Equational Deduction Systems Marcelo Fiore Computer Laboratory University of Cambridge TLCA 2009 3.VII.2009 Context concrete theories meta-theories Context concrete theories meta-theories

More information

Intersection Types and Lambda Theories

Intersection Types and Lambda Theories Intersection Types and Lambda Theories M.Dezani-Ciancaglini S.Lusin Abstract We illustrate the use of intersection types as a semantic tool for showing properties of the lattice of λ-theories. Relying

More information

Elementary Affine Logic and the Call by Value Lambda Calculus

Elementary Affine Logic and the Call by Value Lambda Calculus Elementary Affine Logic and the Call by Value Lambda Calculus Paolo Coppola 1, Ugo Dal Lago 2, and Simona Ronchi Della Rocca 3 1 Università di Udine 2 Università di Bologna 3 Università di Torino Abstract.

More information

Origin in Mathematical Logic

Origin in Mathematical Logic Lambda Calculus Origin in Mathematical Logic Foundation of mathematics was very much an issue in the early decades of 20th century. Cantor, Frege, Russel s Paradox, Principia Mathematica, NBG/ZF The Combinatory

More information

Pairing Transitive Closure and Reduction to Efficiently Reason about Partially Ordered Events

Pairing Transitive Closure and Reduction to Efficiently Reason about Partially Ordered Events Pairing Transitive Closure and Reduction to Efficiently Reason about Partially Ordered Events Massimo Franceschet Angelo Montanari Dipartimento di Matematica e Informatica, Università di Udine Via delle

More information

Canonical Calculi: Invertibility, Axiom expansion and (Non)-determinism

Canonical Calculi: Invertibility, Axiom expansion and (Non)-determinism Canonical Calculi: Invertibility, Axiom expansion and (Non)-determinism A. Avron 1, A. Ciabattoni 2, and A. Zamansky 1 1 Tel-Aviv University 2 Vienna University of Technology Abstract. We apply the semantic

More information

Intersection and Singleton Type Assignment Characterizing Finite Böhm-Trees

Intersection and Singleton Type Assignment Characterizing Finite Böhm-Trees Information and Computation 178, 1 11 (2002) doi:101006/inco20022907 Intersection and Singleton Type Assignment Characterizing Finite Böhm-Trees Toshihiko Kurata 1 Department of Mathematics, Tokyo Metropolitan

More information

Consequence Relations and Natural Deduction

Consequence Relations and Natural Deduction Consequence Relations and Natural Deduction Joshua D Guttman Worcester Polytechnic Institute September 16, 2010 Contents 1 Consequence Relations 1 2 A Derivation System for Natural Deduction 3 3 Derivations

More information

hal , version 1-21 Oct 2009

hal , version 1-21 Oct 2009 ON SKOLEMISING ZERMELO S SET THEORY ALEXANDRE MIQUEL Abstract. We give a Skolemised presentation of Zermelo s set theory (with notations for comprehension, powerset, etc.) and show that this presentation

More information

Delayed substitutions

Delayed substitutions Delayed substitutions José Espírito Santo Departamento de Matemática Universidade do Minho Portugal jes@math.uminho.pt Abstract. This paper investigates an approach to substitution alternative to the implicit

More information

Type Theory and Constructive Mathematics. Type Theory and Constructive Mathematics Thierry Coquand. University of Gothenburg

Type Theory and Constructive Mathematics. Type Theory and Constructive Mathematics Thierry Coquand. University of Gothenburg Type Theory and Constructive Mathematics Type Theory and Constructive Mathematics Thierry Coquand University of Gothenburg Content An introduction to Voevodsky s Univalent Foundations of Mathematics The

More information

Computer Science at Kent

Computer Science at Kent Computer Science at Kent New eta-reduction and Church-Rosser Yong Luo Technical Report No. 7-05 October 2005 Copyright c 2005 University of Kent Published by the Computing Laboratory, University of Kent,

More information

Semantics with Intersection Types

Semantics with Intersection Types Semantics with Intersection Types Steffen van Bakel Department of Computing, Imperial College of Science, Technology and Medicine, 180 Queen s Gate, London SW7 2BZ, U.K., E-mail: svb@doc.ic.ac.uk (Sections

More information

Silvia Ghilezan and Jelena Ivetić

Silvia Ghilezan and Jelena Ivetić PUBLICATIONS DE L INSTITUT MATHÉMATIQUE Nouvelle série, tome 82(96) (2007), 85 91 DOI 102298/PIM0796085G INTERSECTION TYPES FOR λ Gtz -CALCULUS Silvia Ghilezan and Jelena Ivetić Abstract. We introduce

More information

Scuola Estiva di Logica Gargnano, August

Scuola Estiva di Logica Gargnano, August Linear Logic, Linear Logic, Università degli Studi di Torino Scuola Estiva di Logica 2015 Gargnano, August 28 2015 Complexity theory Linear Logic, study of complexity classes and their relations define

More information

A characterization of the Taylor expansion of λ-terms

A characterization of the Taylor expansion of λ-terms A characterization of the Taylor expansion of λ-terms Pierre Boudes, Fanny He, and Michele Pagani LIPN - University Paris 13 Villetaneuse, France {boudes, he, pagani}@lipn.univ-paris13.fr Abstract The

More information

Alonzo Church ( ) Lambda Calculus. λ-calculus : syntax. Grammar for terms : Inductive denition for λ-terms

Alonzo Church ( ) Lambda Calculus. λ-calculus : syntax. Grammar for terms : Inductive denition for λ-terms Alonzo Church (1903-1995) Lambda Calculus 2 λ-calculus : syntax Grammar for terms : t, u ::= x (variable) t u (application) λx.t (abstraction) Notation : Application is left-associative so that t 1 t 2...

More information

Using models to model-check recursive schemes

Using models to model-check recursive schemes Using models to model-check recursive schemes S Salvati and I Walukiewicz Université de Bordeaux, INRIA, CNRS, LaBRI UMR5800 Abstract We propose a model-based approach to the model checking problem for

More information

Mathematical Logic IV

Mathematical Logic IV 1 Introduction Mathematical Logic IV The Lambda Calculus; by H.P. Barendregt(1984) Part One: Chapters 1-5 The λ-calculus (a theory denoted λ) is a type free theory about functions as rules, rather that

More information

The Bang Calculus: an untyped lambda-calculus generalizing Call-By-Name and Call-By-Value

The Bang Calculus: an untyped lambda-calculus generalizing Call-By-Name and Call-By-Value The Bang Calculus: an untyped lambda-calculus generalizing Call-By-Name and Call-By-Value Thomas Ehrhard CNRS, IRIF, UMR 8243, Univ Paris Diderot, Sorbonne Paris Cité, F-75205 Paris, France thomas.ehrhard@pps.univ-paris-diderot.fr

More information

arxiv: v2 [cs.lo] 19 Sep 2016

arxiv: v2 [cs.lo] 19 Sep 2016 Open Call-by-Value (Extended Version) Beniamino Accattoli 1 and Giulio Guerrieri 2 1 INRIA, UMR 7161, LIX, École Polytechnique, beniamino.accattoli@inria.fr, 2 Aix Marseille Université, CNRS, Centrale

More information

The Lambek-Grishin calculus for unary connectives

The Lambek-Grishin calculus for unary connectives The Lambek-Grishin calculus for unary connectives Anna Chernilovskaya Utrecht Institute of Linguistics OTS, Utrecht University, the Netherlands anna.chernilovskaya@let.uu.nl Introduction In traditional

More information

Deriving natural deduction rules from truth tables (Extended version)

Deriving natural deduction rules from truth tables (Extended version) Deriving natural deduction rules from truth tables (Extended version) Herman Geuvers 1 and Tonny Hurkens 2 1 Radboud University & Technical University Eindhoven, The Netherlands herman@cs.ru.nl Abstract

More information

Logic for Computational Effects: work in progress

Logic for Computational Effects: work in progress 1 Logic for Computational Effects: work in progress Gordon Plotkin and John Power School of Informatics University of Edinburgh King s Buildings Mayfield Road Edinburgh EH9 3JZ Scotland gdp@inf.ed.ac.uk,

More information

Full Abstraction for Resource Calculus with Tests

Full Abstraction for Resource Calculus with Tests Full Abstraction for Resource Calculus with Tests Antonio Bucciarelli 1, Alberto Carraro 1,3, Thomas Ehrhard 1, and Giulio Manzonetto 2 1 Laboratoire PPS, CNRS, Université Paris-Diderot, Paris, France

More information

The Sensible Graph Theories of Lambda Calculus

The Sensible Graph Theories of Lambda Calculus The Sensible Graph Theories of Lambda Calculus Antonio Bucciarelli Université Paris 7 Equipe PPS 2 place Jussieu, 72251 Paris Cedex 05, France buccia@ppsjussieufr Antonino Salibra Università Ca Foscari

More information

Rewriting, Explicit Substitutions and Normalisation

Rewriting, Explicit Substitutions and Normalisation Rewriting, Explicit Substitutions and Normalisation XXXVI Escola de Verão do MAT Universidade de Brasilia Part 1/3 Eduardo Bonelli LIFIA (Fac. de Informática, UNLP, Arg.) and CONICET eduardo@lifia.info.unlp.edu.ar

More information

Bruno DINIS and Gilda FERREIRA INSTANTIATION OVERFLOW

Bruno DINIS and Gilda FERREIRA INSTANTIATION OVERFLOW RPORTS ON MATHMATICAL LOGIC 51 (2016), 15 33 doi:104467/20842589rm160025279 Bruno DINIS and Gilda FRRIRA INSTANTIATION OVRFLOW A b s t r a c t The well-known embedding of full intuitionistic propositional

More information

Natural Deduction for Propositional Logic

Natural Deduction for Propositional Logic Natural Deduction for Propositional Logic Bow-Yaw Wang Institute of Information Science Academia Sinica, Taiwan September 10, 2018 Bow-Yaw Wang (Academia Sinica) Natural Deduction for Propositional Logic

More information

Pairing Transitive Closure and Reduction to Efficiently Reason about Partially Ordered Events

Pairing Transitive Closure and Reduction to Efficiently Reason about Partially Ordered Events Pairing Transitive Closure and Reduction to Efficiently Reason about Partially Ordered Events Massimo Franceschet Angelo Montanari Dipartimento di Matematica e Informatica, Università di Udine Via delle

More information

Communication Errors in the π-calculus are Undecidable

Communication Errors in the π-calculus are Undecidable Communication Errors in the π-calculus are Undecidable Vasco T. Vasconcelos Department of Informatics Faculty of Sciences, University of Lisbon António Ravara Department of Mathematics Lisbon Institute

More information

Modularity of Confluence: A Simplified Proof

Modularity of Confluence: A Simplified Proof 1 Modularity of Confluence: A Simplified Proof Jan Willem Klop 1,2,5,6 Aart Middeldorp 3,5 Yoshihito Toyama 4,7 Roel de Vrijer 2 1 Department of Software Technology CWI, Kruislaan 413, 1098 SJ Amsterdam

More information

Five Basic Concepts of. Axiomatic Rewriting Theory

Five Basic Concepts of. Axiomatic Rewriting Theory Five Basic Concepts of Axiomatic Rewriting Theory Paul-André Melliès Institut de Recherche en Informatique Fondamentale (IRIF) CNRS & Université Paris Denis Diderot 5th International Workshop on Confluence

More information

Beyond First-Order Logic

Beyond First-Order Logic Beyond First-Order Logic Software Formal Verification Maria João Frade Departmento de Informática Universidade do Minho 2008/2009 Maria João Frade (DI-UM) Beyond First-Order Logic MFES 2008/09 1 / 37 FOL

More information

A probabilistic lambda calculus - Some preliminary investigations

A probabilistic lambda calculus - Some preliminary investigations A probabilistic lambda calculus - Some preliminary investigations Ugo Dal Lago, Margherita Zorzi Università di Bologna, Università di Verona June, 9-11, 2010, Torino Introduction: Λ P We present some results

More information

Graph lambda theories

Graph lambda theories Under consideration for publication in Math. Struct. in Comp. Science Graph lambda theories A N T O N I O B U C C I A R E L L I 1 and A N T O N I N O S A L I B R A 2 1 Equipe PPS (case 7014), Université

More information

Completeness and Partial Soundness Results for Intersection & Union Typing for λµ µ

Completeness and Partial Soundness Results for Intersection & Union Typing for λµ µ Completeness and Partial Soundness Results for Intersection & Union Typing for λµ µ (Annals of Pure and Applied Logic 161, pp 1400-1430, 2010) Steffen van Bakel Department of Computing, Imperial College

More information

Introduction to Kleene Algebra Lecture 14 CS786 Spring 2004 March 15, 2004

Introduction to Kleene Algebra Lecture 14 CS786 Spring 2004 March 15, 2004 Introduction to Kleene Algebra Lecture 14 CS786 Spring 2004 March 15, 2004 KAT and Hoare Logic In this lecture and the next we show that KAT subsumes propositional Hoare logic (PHL). Thus the specialized

More information