Standardization of a call-by-value lambda-calculus

Size: px
Start display at page:

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

Transcription

1 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 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,ronci}@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 as been proved tat tis extended calculus provides elegant caracterizations of many semantic properties, as for example solvability. We prove a standardization teorem for tis calculus by generalizing Takaasi s approac of parallel reductions. Te standardization property allows us to prove tat our calculus is conservative wit respect to te Plotkin s one. In particular, we sow tat te notion of solvability for tis calculus coincides wit tat for Plotkin s call-by-value lambda-calculus ACM Subject Classification D.3.1 Formal Definitions and Teory, F.3.2 Semantics of Programming Language, F.4.1 Matematical Logic. Keywords and prases standardization, sequentialization, lambda-calculus, sigma-reduction, parallel reduction, call-by-value, ead reduction, ernal reduction, solvability, potential valuability, observational equivalence. 1 Introduction Te λ v -calculus (λ v for sort) as been roduced by Plotkin in [15], in order to give a formal account of te call-by-value evaluation, wic is te most commonly used parameter passing policy for programming languages. λ v sares te syntax wit te classical, call-by-name, λ-calculus (λ for sort), but its reduction rule, β v, is a restriction of β, firing only in case te argument is a value (i.e., a variable or an abstraction). Wile β v is enoug for evaluation, it turned out to be too weak to study operational properties of terms. For example, in λ, te β-reduction is sufficient to caracterize solvability and (using extensionality) separability, but, in order to caracterize similar properties for λ v, it as been necessary to roduce different notions of reduction unsuitable for a correct call-by-value evaluation (see [13, 14]): tis is disappoing and requires complex reasoning. In tis paper we study λ σ v, te extension of λ v proposed in [3]. It keeps te λ v (and λ) syntax and it adds to te β v -reduction two commutation rules, called σ 1 and σ 3, wic unblock β v -redexes tat are idden by te yper-sequential structure of terms. It is well-known (see [14, 1]) tat in λ v tere are normal forms tat are unsolvable, e.g. (λyx.xx)(zz)(λx.xx). Te more evident benefit of λ σ v is tat te commutation rules make all normal forms solvable (indeed (λyx.xx)(zz)(λx.xx) is not a λ σ v normal form). More generally, te so obtained language, allows us to caracterize Tis work was partially supported by LINTEL TO_Call1_2012_0085, i.e. a Researc Project funded by te Compagnia di San Paolo. Giulio Guerrieri, Luca Paolini and Simona Ronci Della Rocca; licensed under Creative Commons License CC-BY Leibniz International Proceedings in Informatics Scloss Dagstul Leibniz-Zentrum für Informatik, Dagstul Publising, Germany

2 2 Standardization of a call-by-value lambda-calculus operational properties, like solvability and potential valuability, in an ernal and elegant way (see [3]). In tis paper we prove a standardization property in λ σ v, and some of its consequences, namely its soundness wit respect to te semantics of λ v. Let us recall te notion of standardization, wic as been first studied in te ordinary λ-calculus (see, for example [5, 8, 2]). A reduction sequence is standard if its redexes are ordered in a given way, and te corresponding standardization teorem establises tat every reduction sequence can constructively be transformed o a standard one. Standardization is a key tool to grasp te way in wic reductions work, tat seds some ligt on redexes relationsips and teir dependencies. It is useful for caracterization of semantic properties troug reduction strategies (te proof of operational semantics adequacy is a typical use). In te λ v setting standardization teorems ave been proved by Plotkin [15], Paolini and Ronci Della Rocca [14, 12] and Crary [4]. Te definition of standard sequence of reductions considered by Plotkin and Crary coincides, and it imposes a partial order on redexes, wile Paolini and Ronci Della Rocca define a total order on tem. All tese proofs are developed by using te notion of parallel reduction roduced by Tait and Martin-Löf (see Takaasi [17] for details and eresting tecnical improvements). We empasize tat tis metod does not involve te notion of residual of a redex, on wic many classical proofs for te λ-calculus are based (see for example [8, 2]). As in [15, 17, 14, 4], we use a suitable notion of parallel reduction for developing our standardization teorem for λ σ v. In particular we consider two groups of redexes, ead β v -redexes and ead σ-redexes (putting togeter σ 1 and σ 3 ), and we induce a total order on ead redexes of te two groups, witout imposing any order on ead σ-redexes temselves. More precisely, wen σ-redexes are missing, tis notion of standardization coincides wit tat presented in [14]. Moreover, we sow tat it is not possible to strengten 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 after te ead ones, for a non-standard definition of ead reduction. Sequentialization as eresting consequences: it allows us to prove tat fundamental operational properties in λ σ v, like observational equivalence, potential valuability and solvability, are conservative wit respect to te corresponding notions of λ v. Tis fully justifies te project in [3] were λ σ v as been roduced as a tool for studying te operational beaviour of λ v. Oter variants of λ v ave been roduced in te literature for modeling te call-by-value computation. We would like to cite ere at least te contributions of Moggi [10], Felleisen and Sabry [16], Maraist et al. [9], Herbelin and Zimmerman [7], Accattoli and Paolini [1]. All tese proposals are based on te roduction of new constructs to te syntax of λ v, so te comparison between tem is not easy wit respect to syntactical properties (some detailed comparison is given in [1]). We po out tat te calculi roduced in [10, 16, 9, 7] present some variants of our σ 1 and/or σ 3 rules, often in a setting wit explicit substitutions. Outline. In Section 2 we roduce te language λ σ v and its operational beaviour; in Section 3 te sequentialization property is proved; Section 4 contains te main result, i.e., standardization; in Section 5 some conservativity results wit respect to Plotkin s λ v -calculus are proved. Section 6 concludes te paper, wit some s for future work. 2 Te call-by-value lambda calculus wit sigma-rules In tis section we present λ σ v, a call-by-value λ-calculus roduced in [3] tat adds two σ-reduction rules to pure (i.e. witout constants) call-by-value λ-calculus defined by Plotkin in [15]. Te syntax of terms of λ σ v [3] is te same as te one of ordinary λ-calculus and Plotkin s

3 G. Guerrieri, L. Paolini and S. Ronci Della Rocca 3 call-by-value λ-calculus λ v [15] (witout constants). Given a countable set V of variables (denoted by x, y, z,... ), te 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. Te 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 te term obtained by te capture-avoiding simultaneous substitution of V i for eac free occurrence of x i in te term M (for all 1 i n). Note tat, 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 (wit exactly one ole ), denoted by C, are defined as usual via te grammar: C ::= λx.c CM MC. We use C M for te term obtained by te capture-allowing substitution of te term M for te ole in te context C. Notation. From now on, we set I = λx.x and = λx.xx. Te reduction rules of λ σ v consist of Plotkin s β v -reduction rule, roduced in [15], and two simple commutation rules called σ 1 and σ 3, studied in [3]. Definition 1 (Reduction rules). We define te 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 wit x / fv(l) V ((λx.l)n) σ3 (λx.v L)N wit x / fv(v ). For any r {β v, σ 1, σ 3 }, if M r M ten M is a r-redex and M is its r-contractum. In tis sense, a term of te sape (λx.m)n (for any M, N Λ) is a β-redex. We set σ = σ1 σ3 and v = βv σ. Te side conditions on σ1 and σ3 in Definition 1 can be always fulfilled by α-renaming. Obviously, any β v -redex is a β-redex but te converse does not old: (λx.z)(yi) is a β-redex but not a β v -redex. Redexes of different kind may overlap: for example, te term I is a σ 1 -redex and it contains te β v -redex I; te term (I )(xi) is a σ 1 -redex and it contains te σ 3 -redex (I ), wic contains in turn te β v -redex I. According to te 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]), te 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 = ) te reflexive-transitive (resp. transitive; reflexive) closure of R. Definition 2 (Reductions). Let r {β v, σ 1, σ 3, σ, v}. Te r-reduction r is te contextual closure of r, i.e. M r M iff tere is a context C and N, N Λ suc tat M = C N, M = C N and N r N. Te r-equivalence = r is te reflexive-transitive and symmetric closure of r. Let M be a term: M is r-normal if tere is no term N suc tat M r N; M is r-normalizable if tere is a r-normal term N suc tat M r N; M is strongly r-normalizing if tere is no sequence (N i ) i N suc tat 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.

4 4 Standardization of a call-by-value lambda-calculus 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 ) ten M Λ v ; more precisely, V = λx.n and M = λx.n for some N, N Λ wit N r N (resp. N r N). Proposition 4 (See [3]). Te σ-reduction is confluent and strongly normalizing. Te v-reduction is confluent. Te λ σ v -calculus, λ σ v for sort, is te set Λ of terms endowed wit te v-reduction v. Te set Λ endowed wit te β v -reduction βv is te λ v -calculus (λ v for sort), i.e. te Plotkin s call-by-value λ-calculus [15] (witout constants), wic is tus 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 te only possible v-reduction pats from M and N respectively: M and N are not v-normalizable, and M = v N. Meanwile, M and N are β v -normal and different, ence M βv N (by confluence of βv, see [15]). Informally, σ-rules unblock β v -redexes wic are idden by te yper-sequential structure of terms. Tis approac is alternative to te one in [1] were idden β v -redexes are reduced using rules acting at a distance (troug explicit substitutions). It can be sown tat te call-by-value λ-calculus wit explicit substitution roduced in [1] can be embedded in λ σ v. 3 Sequentialization In tis section we aim to prove a sequentialization teorem (Teorem 22) for te λ σ v -calculus by adapting Takaasi s metod [17, 4] based on parallel reductions. Notation. From now on, we always assume tat V, V Λ v. Note tat te generic form of a term is V M 1... M m for some m N (in particular, values are obtained wen m = 0). Te sequentialization result is based on a partitioning of v-reduction between ead and ernal reduction. Definition 6 (Head β v -reduction). We define inductively te ead β v -reduction βv te following rules (m N in bot rules): by β v N βv N rigt (λx.m)v M 1... M m βv M{V/x}M 1... M m V NM 1... M m βv V N M 1... M m Te ead β v -reduction βv reduces exactly te same redexes (see also [13]) as te left reduction defined in [15, p. 136] for λ v and called evaluation in [16, 4]. If N βv N ten N is obtained from N by reducing te leftmost-outermost β v -redex, not in te scope of a λ: tus, te ead β v -reduction is deterministic (i.e., it is a partial function from Λ to Λ) and does not reduce values. Definition 7 (Head σ-reduction). We define inductively te ead σ-reduction σ by te following rules (m N in all te rules, x / fv(l) in te rule σ 1, x / fv(v ) in te rule σ 3 ): N σ N rigt (λx.m)nlm 1... M m σ (λx.ml)nm 1... M m V NM 1... M m σ V N M 1... M m V ((λx.l)n)m 1... M m σ 1 σ (λx.v L)NM 1... M m σ 3

5 G. Guerrieri, L. Paolini and S. Ronci Della Rocca 5 Te ead (v-)reduction is v = βv σ. Te ernal (v-)reduction is v = v v. Notice tat βv βv βv and σ σ σ and v v v. Values are normal forms for te ead reduction, but te converse does not old: xi / Λ v is ead-normal. Informally, if N σ N ten N is obtained from N by reducing one of te leftmost σ 1 - or σ 3 -redexes, not in te scope of a λ: in general, a term may contain several ead σ 1 - and σ 3 -redexes. Indeed, differently from βv, te ead σ-reduction σ is not deterministic, for example te leftmost-outermost σ 1 - and σ 3 -redexes may overlap: if M = (λy.y )( (xi))i ten M σ (λy.y I)( (xi)) = N 1 by applying te rule σ 1 and M σ (λz.(λy.y )(zz))(xi)i =N 2 by applying te rule σ 3. Note tat N 1 contains only a ead σ 3 -redex and N 1 σ (λz.(λy.y I)(zz))(xI) = N wic is normal for v ; meanwile N 2 contains only a ead σ 1 -redex and N 2 σ (λz.(λy.y )(zz)i)(xi) = N wic is normal for v : N N, ence te ead reduction v is not confluent and a term may ave several ead-normal forms (tis example does not contradict te confluence of σ-reduction because N σ N but by performing an ernal reduction step). Later, in Corollary 26.2 we sow tat if a term M as a ead normal form N Λ v ten N is te unique ead normal form of M. Definition 8 (Parallel reduction). We define inductively te parallel reduction by te following rules (x / fv(l) in te rule σ 1, x / fv(v ) in te rule σ 3 ): V V M i M i (m N, 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 (m N, 0 i m) σ1 (λx.m 0)NLM 1... M m (λx.m 0L )N M 1... M m V V N N L L M i M i (m N, 1 i m) σ3 V ((λx.l)n)m 1... M m (λx.v L )N M 1... M m M i M i (m N, 0 i m) λ (λx.m 0)M 1... M m (λx.m 0)M 1... M m M i M i (m N, 1 i m) var x M 1... M m x M 1... M m In Definition 8 te rule var as no premises wen m = 0: tis is te base case of te inductive definition of. Te rules σ 1 and σ 3 ave exactly tree premises wen m = 0. Intuitively, M M means tat M is obtained from M by reducing a number of β v -, σ 1 - and σ 3 -redexes (existing in M) simultaneously. Definition 9 (Internal and strong parallel reduction). We define inductively te ernal parallel reduction by te following rules: N N λx.n λx.n λ x x var V V N N M i M i (m N, 1 i m) rigt V NM 1... M m V N M 1... M m Te strong parallel reduction is defined by: M N iff M N and tere exist M, M Λ suc tat M β v M σ M N. Notice tat te rule rigt for as exactly two premises wen m = 0. Remark 10. Te relations, and are reflexive. Te reflexivity of follows immediately from te reflexivity of and. Te proofs of reflexivity of and are bot by structural induction on a term: in te case of, recall tat every term is of te form (λx.n)m 1... M m or x M 1... M m for some m N and ten apply te rule λ or var respectively, togeter wit te inductive ypotesis; in te case of, recall tat every term is of te form λx.m or x or V NM 1... M m for some m N and ten apply te rule λ (togeter wit te reflexivity of ) or var or rigt (togeter wit te reflexivity of and te inductive ypotesis) respectively.

6 6 Standardization of a call-by-value lambda-calculus One as (first, prove tat by induction on te derivation of M M, te oter inclusions follow from te definition of ) and, since is reflexive (Remark 10), βv and σ. Observe tat R for any R { βv, βv,,, }, even if for different reasons: for example, by reflexivity of (Remark 10), wereas βv by reducing te (leftmost-outermost) β v -redex. Next two furter remarks collect many minor properties tat can be easily proved. Remark Te ead β 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 as V βv M. 2. Te ead σ-reduction σ does neiter reduce a value nor reduce to a value, i.e., for any M Λ and any V Λ v, one as V σ M and M σ V. 3. Variables and abstractions are preserved under ( -expansion), i.e., if M x (resp. M λx.n ) ten M = x (resp. M = λx.n for some N Λ suc tat N N ). 4. If M M ten λx.m R λx.m for any R {,, }. Indeed, for R {, } apply te rule λ to conclude, ten λx.m λx.m according to te definition of. 5. For any V, V Λ v, V V iff V V. Te left-to-rigt direction olds because ; conversely, assume V V : if V is a variable ten necessarily V = V and ence V V by applying te rule var for ; oterwise V = λx.n for some N Λ, and ten necessarily V = λx.n wit N N, so V V by applying te rule λ for. Remark If M M and N N ten MN M N. For te proof, it is sufficient to consider te last rule of te derivation of M M. 2. If R { βv, σ } and M R M, ten MN R M N for any N Λ. For te proof, it is sufficient to consider te last rule of te derivation of M R M, for any R { βv, σ }. 3. If M M and N N were M / Λ v, ten MN M N : indeed, te last rule in te derivation of M M can be neiter λ nor var because M / Λ v. Te ypotesis M / Λ v is crucial: for example, x x and I but I and tus x(i ) x. 4. v v. As a consequence, = v and (by Proposition 4) is confluent. 5. v v, so = v. Tus, by Remark 11.3, variables and abstractions are preserved under v -expansion, i.e., if M v x (resp. M v λx.n ) ten M = x (resp. M = λx.n wit N v N ). 6. For any R { βv, σ }, if M R M ten M{V/x} R M {V/x} for any V Λ v. Te proof is by straigtforward induction on te derivation of M R M for any R { βv, σ }. As expected, a basic property of parallel reduction is te following: Lemma 13 (Substitution lemma for ). If M M and V V ten M{V/x} M {V /x}. Proof. By straigtforward induction on te derivation of M M. Te following lemma will play a crucial role in te proof of Lemmas and sows tat te ead σ-reduction σ can be postponed after te ead β v -reduction βv. Lemma 14 (Commutation of ead reductions). 1. If M σ L βv N ten tere exists L Λ suc tat M βv L = σ N. 2. If M σ L β v N ten tere exists L Λ suc tat M β v L σ N. 3. If M v M ten tere exists N Λ suc tat M β v N σ M. Proof. 1. By induction on te derivation of M σ L. Let us consider its last rule r. If r = σ 1 ten 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 were m N and x / fv(l 0 ). Since L βv N, tere are only two cases:

7 G. Guerrieri, L. Paolini and S. Ronci Della Rocca 7 eiter N 0 βv N 0 and N = (λx.m 0 L 0 )N 0M 1... M m (according to te rule rigt for βv ), ten 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 te rule β v, as x / fv(l 0 )), so M βv N. If r = σ 3 ten M = V ((λx.l 0 )N 0 )M 1... M m and L = (λx.v L 0 )N 0 M 1... M m wit m N and x / fv(v ). Since L βv N, tere are only two cases: eiter N 0 βv N 0 and N = (λx.v L 0 )N 0M 1... M m (according to te rule rigt for βv ), ten 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 te rule β v, as x / fv(v )), so M βv N. Finally, if r = rigt ten M = V N 0 M 1... M m and L = V N 0M 1... M m wit m N and N 0 σ N 0. By Remark 11.2, N 0 / Λ v and tus, since L βv N, te only possibility is tat N 0 βv N 0 and N = V N 0 M 1... M m (according to te rule rigt for βv ). By induction ypotesis, tere exists N 0 Λ suc tat N 0 βv N 0 = σ N 0. Terefore, M βv V N 0 M 1... M m = σ N. 2. Immediate consequence of Lemma 14.1, using standard tecniques of rewriting teory. 3. Immediate consequence of Lemma 14.2, using standard tecniques of rewriting teory. We are now able to travel over again Takaasi s metod [17, 4] in our setting wit β v - and σ-reduction. Te next four lemmas govern te strong parallel reduction and will be used to prove Lemma 19. Lemma 15. If M M and N N and M / Λ v, ten MN M N. Proof. One as MN M N by Remark 12.1 and since M M. By ypotesis, tere exist m, n N and M 0,..., M m, N 0,..., N n suc tat 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 as N n N M N by Remark Terefore, MN M N. Lemma 16. If M M and N N ten MN M N. Proof. If M / Λ v ten 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 ypotesis, tere are m, m, n, n N and M 0,..., M m, M 0,..., M m, N 0,..., N n, N 0,..., N n suc tat: M = M 0, M m = M 0, M m M, M i βv M i+1 for any 0 i < m, and M i σ M i +1 for any 0 i < m, N = N 0, N n = N 0, N n N, N j βv N j+1 for any 0 j < n and N j σ N j +1 for any 0 j < n. By Remark 11.3, M m Λ v since M Λ v, terefore m = 0 by Remark 11.2, and tus M m = M 0 M (and M m M since ) and M m Λ v. Using te rules rigt for βv and σ, one as 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 te rule rigt for, one as M m N n M N. Terefore, 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 ence MN M N. Lemma 17. If M M and V V, ten M{V/x} M {V /x}. Proof. By Lemma 13, one as M{V/x} M {V /x} since M M and V V. We proceed by induction on M Λ. Let us consider te last rule r of te derivation of M M. If r = var ten tere are two cases: eiter M = x and ten M{V/x}=V V =M {V /x}; or M = y x and ten 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 lambda-calculus If r = λ ten M = λy.n and M = λy.n wit N N ; we can suppose witout loss of generality tat y / fv(v ) {x}. One as N{V/x} N {V /x} according to Lemma 13. By applying te rule λ for, one as M{V/x} = λy.n{v/x} λy.n {V /x} = M {V /x} and tus M{V/x} M {V /x}. Finally, if r = rigt ten M = UNM 1... M m and M = U N M 1... M m for some m N wit U, U Λ v, U U, N N and M i M i for any 1 i m. By induction ypotesis, U{V/x} U {V /x} (indeed U U according to Remark 11.5) and N{V/x} N {V /x}. By Lemma 13, M i {V/x} M i {V /x} for any 1 i m. By Lemma 16, U{V/x}N{V/x} U {V /x}n {V/x} and ence, by applying Lemma 15 m times since U {V /x}n {V/x} / Λ v, one as 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 te proof of te two next lemmas, as well as in te proof of Corollary 21 and Teorem 22, our Lemma 14 plays a crucial role: indeed, since te ead σ-reduction well eract wit te ead β v -reduction, Takaasi s metod [17, 4] is still working wen adding te reduction rules σ 1 and σ 3 to Plotkin s β v -reduction. Lemma 18 (Substitution lemma for ). If M M and V V ten M{V/x} M {V /x}. Proof. By Lemma 13, one as M{V/x} M {V /x} since M M and V V. By ypotesis, tere exist m, n N and M 0,..., M m, N 0,..., N n suc tat 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 as N n {V/x} M {V /x}, tus tere exist L, N Λ suc tat M{V/x} β v N 0 {V/x} σ N n {V/x} β v N σ L M {V /x}. By Lemma 14.2, tere exists N Λ suc tat M{V/x} β v N 0 {V/x} β v N σ N σ L M {V /x}, terefore M{V/x} M {V /x}. Now we are ready to prove a key lemma, wic states tat parallel reduction coincides wit strong parallel reduction (te inclusion is trivial). Lemma 19 (Key Lemma). If M M ten M M. Proof. By induction on te derivation of M M. Let us consider its last rule r. If r = var ten M = x M 1... M m and M = x M 1... M m were m N and M i M i for any 1 i m. By reflexivity of (Remark 10), x x. By induction ypotesis, M i M i for any 1 i m. Terefore, M M by applying Lemma 16 m times. If r = λ ten M = (λx.m 0 )M 1... M m and M = (λx.m 0)M 1... M m were m N and M i M i for any 0 i m. By induction ypotesis, 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 ten M = (λx.m 0 )V M 1... M m and M = M 0{V /x}m 1... M m were m N, V V and M i M i for any 0 i m. By induction ypotesis, V V and M i M i for any 0 i m. By applying te rule β v for βv, one as 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, tus tere are L, N Λ suc tat M βv M 0 {V/x}M 1... M m β v L σ N M. So M M. If r = σ 1 ten M = (λx.m 0 )N 0 L 0 M 1... M m and M = (λx.m 0L 0)N 0M 1... M m were m N, L 0 L 0, N 0 N 0 and M i M i for any 0 i m. By induction ypotesis, N 0 N 0 and M i M i for any 1 i m. By applying te rule σ 1 for σ, one as M σ (λx.m 0 L 0 )N 0 M 1... M m. By Remark 12.1, M 0 L 0 M 0L 0 and tus λ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

9 G. Guerrieri, L. Paolini and S. Ronci Della Rocca 9 times, ence tere are L, N Λ suc tat M σ (λx.m 0 L 0 )N 0 M 1... M m β v L σ N M. By Lemma 14.2, tere is L Λ suc tat M β v L σ L σ N M and tus M M. Finally, if r = σ 3 ten M = V ((λx.l 0 )N 0 )M 1... M m and M = (λx.v L 0)N 0M 1... M m wit m N, V V, L 0 L 0, N 0 N 0 and M i M i for any 1 i m. By induction ypotesis, N 0 N 0 and M i M i for any 1 i m. By te rule σ 3 for σ, one as M σ (λx.v L 0 )N 0 M 1... M m. By Remark 12.1, V L 0 V L 0 and tus λ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, ence tere are L, N Λ suc tat M σ (λx.v L 0 )N 0 M 1... M m β v L σ N M. By Lemma 14.2, tere is L Λ suc tat M β v L σ L σ N M, so M M. Next Lemma 20 and Corollary 21 sow tat ernal parallel reduction can be sifted after ead reductions. Lemma 20 (Postponement). If M L and L βv N (resp. L σ N) ten tere exists L Λ suc tat M βv L (resp. M σ L ) and L N. Proof. By induction on te derivation of M L. Let us consider its last rule r. If r = var, ten M = x = L wic contradicts L βv N and L σ N by Remarks If r = λ ten L = λx.l for some L Λ, wic contradicts L βv N and L σ N by Remarks Finally, if r = rigt ten M = V M 0 M 1... M m and L = V L 0 L 1... L m were m N, V V (so V V by Remark 11.5), M 0 L 0 (tus M 0 L 0 since ) and M i L i for any 1 i m. If L βv N ten tere are two cases, depending on te last rule r of te derivation of L βv N. If r = β v ten V = λx.n 0, L 0 Λ v and N = N 0{L 0 /x}l 1... L m, tus M 0 Λ v and V = λx.n 0 wit N 0 N 0 by Remark By Lemma 13, one as 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 te rule β v for βv ) and L N by applying Remark 12.1 m times. If r = rigt ten N = V N 0 L 1... L m wit L 0 βv N 0. By induction ypotesis, tere exists L 0 Λ suc tat M 0 βv L 0 N 0. Let L = V L 0M 1... M m : so M βv L (apply te rule rigt for βv ) and L N by applying Remark 12.1 m + 1 times. If L σ N ten tere are tree cases, depending on te last rule r of te derivation of L σ N. If r = σ 1 ten m > 0, V = λx.n 0 and N = (λx.n 0L 1 )L 0 L 2... L m, tus V = λx.n 0 wit N 0 N 0 by Remark Using Remarks 12.1 and 11.4, one as λ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 te rule σ 1 for σ ) and L N by applying Remark 12.1 m times. If r = σ 3 ten 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 wit M 01 L 01 and M 02 L 02 (so M 02 L 02 ). Using Remarks 12.1 and 11.4, one as λx.v M 01 λx.v L 01. Let L = (λx.v M 01 )M 02 M 1... M m : terefore M = V ((λx.m 01 )M 02 )M 1... M m σ L (apply te rule σ 3 for σ ) and L N by applying Remark 12.1 m + 1 times. If r = rigt ten N = V N 0 L 1... L m wit L 0 σ N 0. By induction ypotesis, tere exists L 0 Λ suc tat M 0 σ L 0 N 0. Let L = V L 0M 1... M m : so M σ L (apply te rule rigt for σ ) and L N by applying Remark 12.1 m + 1 times. Corollary 21. If M L and L βv N (resp. L σ N), ten tere exist L, L Λ suc tat M + β v L σ L N (resp. M β v L σ L N).

10 10 Standardization of a call-by-value lambda-calculus Proof. Immediate by Lemma 20 and Lemma 19, applying Lemma 14.2 if L σ N. Now we obtain our first main result (Teorem 22): any v-reduction sequence can be sequentialized o a ead β v -reduction sequence followed by a ead σ-reduction sequence, followed by an ernal reduction sequence. In ordinary λ-calculus, te well-known result corresponding to our Teorem 22 says tat a β-reduction sequence can be factorized in a ead reduction sequence followed by an ernal reduction sequence (see for example [17, Corollary 2.6]). Teorem 22 (Sequentialization). If M v M ten tere exist L, N Λ suc tat M β v L σ N v M. Proof. By Remark 12.4, M M and tus tere are m N and M 0,..., M m Λ suc tat M = M 0, M m = M and M i M i+1 for any 0 i < m. We prove by induction on m N tat tere are L, N Λ suc tat M β v L σ N M, so N v M by Remark If m = 0 ten M = M 0 = M and ence we conclude by taking L = M = N. Finally, suppose m > 0. By induction ypotesis applied to M 1 M, tere exist L, N Λ suc tat M 1 β v L σ N M. By applying Lemma 19 to M, tere exist L 0, N 0 Λ suc tat M β v L 0 σ N 0 M 1. By applying Corollary 21 repeatedly, tere exists N Λ suc tat N 0 Lemma 14.3, tere exists L Λ suc tat M β v v N N and ence M v N L σ N M. M. According to It is wort noticing tat in Definition 7 tere is no distinction between ead σ 1 - and ead σ 3 -reduction steps, and, according to it, te sequentialization of Teorem 22 imposes no order between ead σ-reductions. We denote by σ1 and σ3 respectively te reduction relations σ1 σ and σ3 σ. So, a natural question arises: is it possible to sequentialize tem? Te answer is negative, as proved by te next two counterexamples. Let M = x((λy.z )(zi)) and N = (λy.xz )(zi): M σ3 (λy.xz )(zi) σ1 N, but tere exists no L suc tat M σ 1 L σ 3 N. In fact M contains only a ead σ 3 -redex and (λy.xz )(zi) contains only a ead σ 1 -redex Let M = x((λy.z )(zi) ) and N = (λy.x(z ))(zi): M σ1 x((λy.z )(zi)) σ3 N but tere is no L suc tat M σ 3 L σ 1 N. In fact M contains only a ead σ 1 -redex and x((λy.z )(zi)) contains only a ead σ 3 -redex. So, te impossibility of sequentializing a ead σ-reduction sequence is due to te fact tat a ead σ 1 -reduction step can create a ead σ 3 -redex, and viceversa. Tis is not a problem, since ead σ-reduction is strongly normalizing (by Proposition 4 and since σ σ ). Our approac does not force a strict order of ead σ-reductions. 4 Standardization Now we are able to prove te main result of tis paper, i.e., a standardization teorem for λ σ v (Teorem 25). In particular we provide a notion of standard reduction sequence tat avoids any auxiliary notion of residual redexes, by closely following te definition given in [15]. Notation. For any k, m N wit k m, we denote by M 0,, M k, M m ead a sequence of terms suc tat M i βv M i+1 wen 0 i < k, and M i σ M i+1 wen k i < m. It is easy to ceck tat M ead for any M Λ. Te notion of standard sequence of terms is defined by using te previous notion of ead-sequence. Our notion of standard reduction sequence is mutually defined togeter wit te notion of inner-sequence of terms (Definition 23). Tis definition allows us to avoid non-deterministic cases remarked in [7]

11 G. Guerrieri, L. Paolini and S. Ronci Della Rocca 11 (we provide more details at te end of tis section). We denote by M 0,, M m std (resp. M 0,, M m in ) a standard (resp. inner) sequence of terms. 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 ead and M m,, M m+n in ten M 0,, M m+n std, were m, n N; 2. M in, for any M Λ; 3. if M 0,, M m std ten λz.m 0,, λz.m m in, were m N; 4. if V 0,, V std and N 0,, N n in ten V 0 N 0,, V 0 N n,, V N n in, were, n N; 5. if N 0,, N n in, M 0,, M m std and N 0 Λ v, ten N 0 M 0,, N n M 0,, N n M m in, were 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 ead ) ten N 0,, N n std. Indeed, N 0 ead (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. Teorem 25 (Standardization). 1. If M v M ten tere is a sequence M,, M std. 2. If M v M ten tere is a sequence M,, M in. Proof. Bot statements are proved simultaneously by induction on M Λ. 1. If M = z ten, by Teorem 22, M β v L σ N z for some L, N Λ. By Remarks 12.5 and 11.2, L = N = z; terefore M β v z and ence tere is a sequence M,, z ead. Tus, M,, z std by Remark 24. If M = λz.n ten, by Teorem 22, M β v L σ L v λz.n for some L, L Λ. By Remarks 12.5 and 11.2, L = L = λz.n wit N v N. So M β v λz.n and ence tere is a sequence M,, λz.n ead. By induction on (1), tere is a sequence N,, N std, tus λz.n,, λz.n in by Definition Terefore M,, λz.n,, λz.n std by Definition If M = N L ten, by Teorem 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. Tus tere is a sequence M,, M,, NL ead. By Remark 12.5, NL N L ; clearly, eac step of is an instance of te rule rigt of Definition 9. Tere are two sub-cases. If N Λ v ten N N and L L, so N v N and L v L by Remarks By induction respectively on (1) and (2), tere are sequences N,, N std and L,, L in, tus NL,, NL,, N L in by Definition Terefore M,, M,, NL,, NL,, N L std by Definition If N / Λ v (i.e., N =VM 1... M m wit m>0) ten N N and L L, so N v N and L v L by Remarks By induction respectively on (2) and (1), tere are sequences N,, N in and L,, L std. Hence NL,, N L,, N L in by Definition Tus M,, M,, NL,, N L,, N L std by Definition If M = z ten M = z by Remark 12.5, ence z in by Definition If M = λz.l ten M = λz.l and L v L by Remark Hence tere is a sequence L,, L std by induction on (1). By Definition 23.3, λz.l,, λz.l in. v

12 12 Standardization of a call-by-value lambda-calculus If M = N L ten M = NL for some N, L Λ by Remark 3, since v v and M / Λ v. By Remark 12.5, NL N L ; clearly, eac step of is an instance of te rule rigt of Definition 9. Tere are two sub-cases. If N Λ v ten N N and L L, so N v N and L v L by Remarks Tus tere are sequences N,, N std and L,, L in by induction respectively on (1) and (2). Terefore, by Definition 23.4, NL,, NL,, N L in. If N / Λ v (i.e. N = VM 1... M m wit m > 0) ten N N and L L, tus N v N and L v L by Remarks By induction respectively on (2) and (1), tere are sequences N,, N in and L,, L std. So NL,, N L,, N L in by Definition Due to non-sequentialization of ead σ 1 - and ead σ 3 -reductions, several standard sequences may ave te same starting term and ending term: for instance, if M = I( I)I and N = (λz.(λx.xi)(zz))i ten M v (λx.xi)( I) v N and M v (λz.i(zz))ii v (λz.i(zz)i)i v N are bot standard sequences from M to N. Finally, we can compare our notion of standardization wit tat given in [15]. To make te comparison possible we avoid σ-reductions and we recall tat βv is exactly te Plotkin s left-reduction [15, p. 136]. As remarked in [7, 1.5 p. 149], bot 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 [15]. On te oter and, only te second sequence is standard in our sense. It is easy to ceck tat collapsing te two notions of inner and standard sequence given in Definition 23, we get a notion of standard sequence tat accept bot te above sequences. 5 Some conservativity results Te sequentialization result (Teorem 22) as some eresting semantic consequences. It allows us to prove tat (Corollary 29) te λ σ v -calculus is sound wit respect to te call-byvalue observational equivalence roduced by Plotkin in [15] for λ v. Moreover we can prove tat some notions, like tat of potential valuability and solvability, roduced in [13] for λ v, coincide wit te respective notions for λ σ v (Teorem 31). Tis justifies te idea tat λ σ v is a useful tool for studying te properties of λ v. Our starting po is te following corollary. Corollary If M v V Λ v ten tere exists V Λ v suc tat M β v V v V. 2. For every V Λ v, M β v V if and only if M v V. Proof. Te first po is proved by observing tat, by Teorem 22, tere are N, L Λ suc tat M β v L σ N v V. By Remark 12.5, N Λ v and tus L = N according to Remark Concerning te second po, te rigt-to-left direction is a consequence of Lemma 14.3 and Remark 11.2; te left-to-rigt direction follows from βv v. Let us recall te notion of observational equivalence defined by Plotkin [15] for λ v. Definition 27 (Halting, observational equivalence). Let M Λ. We say tat ( te evaluation of) M alts if tere exists V Λ v suc tat M β v V. Te ( call-by-value) observational equivalence is an equivalence relation = on Λ defined by: M = N if, for every context C, one as tat C M alts iff C N alts. 1 1 Original Plotkin s definition of call-by-value observational equivalence (see [15]) also requires tat C M and C N are closed terms, according to te tradition identifying programs wit closed terms.

13 G. Guerrieri, L. Paolini and S. Ronci Della Rocca 13 Clearly, similar notions can be defined for λ σ v using v instead of βv. Head σ-reduction plays no role neiter in deciding te alting problem for evaluation (Corollary 26.1), nor in reacing a particular value (Corollary 26.2). So, we can conclude tat te notions of alting and observational equivalence in λ σ v coincide wit te ones in λ v, respectively. Now we compare te equational teory of λ σ v wit Plotkin s observational equivalence. Teorem 28 (Adequacy of v-reduction). If M v M ten: M alts iff M alts. Proof. If M β v V Λ v ten M v M v V since βv v. By Corollary 26.1, tere exists V Λ v suc tat M β v V. Tus M alts. Conversely, if M β v V Λ v ten M v V since βv v. By confluence of v (Proposition 4, since M v M ) and Remark 3 (since V Λ v ), tere is V Λ v suc tat V v V and M v V. By Corollary 26.1, tere is V Λ v suc tat M β v V. So M alts. Corollary 29 (Soundness). If M = v N ten M = N. Proof. Let C be a context. By confluence of v (Proposition 4), M = v N implies tat tere exists L Λ suc tat M v L and N v L, ence C M v C L and C N v C L. By Teorem 28, C M alts iff C L alts iff C N alts. Terefore, M = N. Plotkin [15, Teorem 5] as already proved tat M = βv N implies M = N, but our Corollary 29 is not obvious since our λ σ v -calculus equates more tan Plotkin s λ v -calculus (= βv = v since βv v, and Example 5 sows tat tis inclusion is strict). Te converse of Corollary 29 does not old since λx.x(λy.xy) = but λx.x(λy.xy) and are different v-normal forms and ence λx.x(λy.xy) v by confluence of v (Proposition 4). A furter remarkable consequence of Corollary 26.1 is tat te notions of potential valuability and solvability for λ σ v -calculus (studied in [3]) can be sown to coincide wit te ones for Plotkin s λ v -calculus (studied in [13, 14]), respectively. Let us recall teir definition. Definition 30 (Potential valuability, solvability). Let M be a term: M is v-potentially valuable (resp. β v -potentially valuable) if tere are m N, pairwise distinct variables x 1,..., x m and V, V 1,..., V m Λ v suc tat 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 tere are n,m N, variables x 1,..., x m and N 1,...,N n Λ suc tat (λx 1... x m.m)n 1 N n v I (resp. (λx 1... x m.m)n 1 N n β v I). Teorem 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 bot pos, te implication from rigt to left is trivial since βv v. Let us prove te oter direction. 1. Since M is v-potentially valuable, tere are variables x 1,..., x m and V, V 1,..., V m Λ v (wit m 0) suc tat M{V 1 /x 1,..., V m /x m } v V ; ten, tere exists V Λ v suc tat M{V 1 /x 1,..., V m /x m } β v V by Corollary 26.1 and because βv βv, terefore M is β v -potentially valuable. 2. Since M is v-solvable, tere exist variables x 1,..., x m and terms N 1,..., N n (for some n, m 0) suc tat (λx 1... x m.m)n 1 N n v I; ten, tere exists V Λ v suc tat (λx 1... x m.m)n 1 N n β v V v I by Corollary 26.1 and because βv βv. According to Remark 12.5, V = λx.n for some N Λ suc tat N v x. By Corollary 26.1, tere is V Λ v suc tat N β v V v x, ence V = x by Remark 12.5 again. Since βv βv, N β v x and tus V = λx.n β v I, terefore M is β v -solvable.

14 14 Standardization of a call-by-value lambda-calculus So, due to Teorem 31, te semantic (via a relational model) and operational (via two sub-reductions of v ) caracterization of v-potential valuability and v-solvability given in [3, Teorems 24-25] is also a semantic and operational caracterization of β v -potential valuability and β v -solvability. Te difference is tat in λ σ v tese notions can be studied operationally inside te calculus, wile it as been proved in [13, 14] tat te β v -reduction is too weak to caracterize tem: an operational caracterization of β v -potential valuability and β v -solvability cannot be given inside λ v. Hence, λ σ v is a useful, conservative and complete tool for studying semantic properties of λ v. 6 Conclusions In tis paper we ave proved a standardization teorem for te λ σ v -calculus roduced in [3]. Te used tecnique is a notion of parallel reduction. Let us recall tat parallel reduction in λ-calculus as been defined by Tait and Martin-Löf in order to prove confluence of te β-reduction, witout referring to te difficult notion of residuals. Takaasi in [17] as simplified tis tecnique and sowed tat it can be successfully applied to standardization. We would like to remark tat 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 : ten M M 1 and M M 2 but tere is no term M suc tat M 1 M and M 2 M. To sum up, does not enjoy te Diamond Property. Te standardization result allows us to formally verify te correctness of λ σ v wit respect to te semantics of λ v, so we can use λ σ v as a tool for studying properties of λ v. Tis is a remarkable result: in fact some properties, like potential valuability and solvability, cannot be caracterized in λ v by means of β v -reduction (as proved in [13, 14]), but tey ave a natural operational caracterization in λ σ v (via two sub-reductions of v ). We plan to continue to explore te call-by-value computation, using λ σ v. As a first step, we would like to revisit and improve te Separability Teorem given in [11] for λ v. Still te issue is more complex tan in te call-by-name, indeed in ordinary λ-calculus different βη-normal forms can be separated (by te Böm Teorem), wile in λ v tere are different normal forms tat cannot be separated, but wic are only semi-separable (e.g. I and λz.(λu.z)(zz)). We ope to completely caracterize separable and semi-separable normal forms in λ σ v. Tis sould be a first step aimed to define a semantically meaningful notion of approximants. Ten, we sould be able to provide a new insigt on te denotational analysis of te call-by-value, maybe overcoming limitations as tat of te absence of fully abstract filter models [14, Teorem ]. Last but not least, an unexplored but callenging researc direction is te use of commutation rules to improve te call-by-value evaluation. We do not ave concrete evidence supporting suc possibility, but since λ σ v is strongly related to te calculi presented in [7, 1], wic are endowed wit explicit substitutions, we are confident tat a sarp use of commutations can ave a relevant impact in te evaluation. References 1 Beniamino Accattoli and Luca Paolini. Call-by-Value Solvability, Revisited. In Tom Scrijvers and Peter Tiemann, editors, Functional and Logic Programming, volume 7294 of Lecture Notes in Computer Science, pages Springer-Verlag, Henk Barendregt. Te Lambda Calculus: Its Syntax and Semantics, volume 103 of Studies in logic and te foundation of matematics. Nort Holland, 1984.

15 G. Guerrieri, L. Paolini and S. Ronci Della Rocca 15 3 Alberto Carraro and Giulio Guerrieri. A Semantical and Operational Account of Call-by- Value Solvability. In Anca Muscoll, editor, Foundations of Software Science and Computation Structures, volume 8412 of Lecture Notes in Computer Science, pages Springer-Verlag, Karl Crary. A Simple Proof of Call-by-Value Standardization. Tecnical Report CMU-CS , Carnegie Mellon University, Haskell B. Curry and Robert Feys. Combinatory Logic, volume 1. Nort Holland, Jean-Yves Girard. Linear logic. Teoretical Computer Science, 50(1):1 102, Hugo Herbelin and Stépane Zimmermann. An Operational Account of Call-by-Value Minimal and Classical lambda-calculus in "Natural Deduction" Form. In Pierre-Louis Curien, editor, Typed Lambda Calculi and Applications, volume 5608 of Lecture Notes in Computer Science, pages Springer-Verlag, Roger Hindley. Standard and normal reductions. Transactions of te American Matematical Society, pages , Jon Maraist, Martin Odersky, David N. Turner, and Pilip Wadler. Call-by-name, call-byvalue, call-by-need and te linear lambda calculus. Teoretical Computer Science, 228(1 2): , Eugenio Moggi. Computational Lambda-Calculus and Monads. In Logic in Computer Science, pages IEEE Computer Society, Luca Paolini. Call-by-Value Separability and Computability. In Antonio Restivo, Simona Ronci Della Rocca, and Luca Roversi, editors, Italian Conference in Teoretical Computer Science, volume 2202 of Lecture Notes in Computer Science, pages Springer-Verlag, Luca Paolini and Simona Ronci Della Rocca. Parametric parameter passing lambdacalculus. Information and Computation, 189(1):87 106, Luca Paolini and Simona Ronci Della Rocca. Call-by-value Solvability. Teoretical Informatics and Applications, 33(6): , RAIRO Series, EDP-Sciences. 14 Luca Paolini and Simona Ronci Della Rocca. Te Parametric λ-calculus: a Metamodel for Computation. Texts in Teoretical Computer Science: An EATCS Series. Springer-Verlag, Gordon D. Plotkin. Call-by-name, call-by-value and te lambda-calculus. Teoretical Computer Science, 1(2): , Amr Sabry and Mattias Felleisen. Reasoning about programs in continuation-passing style. Lisp and symbolic computation, 6(3-4): , Masako Takaasi. Parallel reductions in lambda-calculus. Information and Computation, 118(1): , 1995.

Standardization of a call-by-value calculus

Standardization of a call-by-value calculus 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-75205 Paris,

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

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

Volume 29, Issue 3. Existence of competitive equilibrium in economies with multi-member households

Volume 29, Issue 3. Existence of competitive equilibrium in economies with multi-member households Volume 29, Issue 3 Existence of competitive equilibrium in economies wit multi-member ouseolds Noriisa Sato Graduate Scool of Economics, Waseda University Abstract Tis paper focuses on te existence of

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

Differentiation in higher dimensions

Differentiation in higher dimensions Capter 2 Differentiation in iger dimensions 2.1 Te Total Derivative Recall tat if f : R R is a 1-variable function, and a R, we say tat f is differentiable at x = a if and only if te ratio f(a+) f(a) tends

More information

1 The concept of limits (p.217 p.229, p.242 p.249, p.255 p.256) 1.1 Limits Consider the function determined by the formula 3. x since at this point

1 The concept of limits (p.217 p.229, p.242 p.249, p.255 p.256) 1.1 Limits Consider the function determined by the formula 3. x since at this point MA00 Capter 6 Calculus and Basic Linear Algebra I Limits, Continuity and Differentiability Te concept of its (p.7 p.9, p.4 p.49, p.55 p.56). Limits Consider te function determined by te formula f Note

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

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

A = h w (1) Error Analysis Physics 141

A = h w (1) Error Analysis Physics 141 Introduction In all brances of pysical science and engineering one deals constantly wit numbers wic results more or less directly from experimental observations. Experimental observations always ave inaccuracies.

More information

A Reconsideration of Matter Waves

A Reconsideration of Matter Waves A Reconsideration of Matter Waves by Roger Ellman Abstract Matter waves were discovered in te early 20t century from teir wavelengt, predicted by DeBroglie, Planck's constant divided by te particle's momentum,

More information

The derivative function

The derivative function Roberto s Notes on Differential Calculus Capter : Definition of derivative Section Te derivative function Wat you need to know already: f is at a point on its grap and ow to compute it. Wat te derivative

More information

Efficient algorithms for for clone items detection

Efficient algorithms for for clone items detection Efficient algoritms for for clone items detection Raoul Medina, Caroline Noyer, and Olivier Raynaud Raoul Medina, Caroline Noyer and Olivier Raynaud LIMOS - Université Blaise Pascal, Campus universitaire

More information

Math 31A Discussion Notes Week 4 October 20 and October 22, 2015

Math 31A Discussion Notes Week 4 October 20 and October 22, 2015 Mat 3A Discussion Notes Week 4 October 20 and October 22, 205 To prepare for te first midterm, we ll spend tis week working eamples resembling te various problems you ve seen so far tis term. In tese notes

More information

Combining functions: algebraic methods

Combining functions: algebraic methods Combining functions: algebraic metods Functions can be added, subtracted, multiplied, divided, and raised to a power, just like numbers or algebra expressions. If f(x) = x 2 and g(x) = x + 2, clearly f(x)

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

Pre-Calculus Review Preemptive Strike

Pre-Calculus Review Preemptive Strike Pre-Calculus Review Preemptive Strike Attaced are some notes and one assignment wit tree parts. Tese are due on te day tat we start te pre-calculus review. I strongly suggest reading troug te notes torougly

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

University Mathematics 2

University Mathematics 2 University Matematics 2 1 Differentiability In tis section, we discuss te differentiability of functions. Definition 1.1 Differentiable function). Let f) be a function. We say tat f is differentiable at

More information

1 Calculus. 1.1 Gradients and the Derivative. Q f(x+h) f(x)

1 Calculus. 1.1 Gradients and the Derivative. Q f(x+h) f(x) Calculus. Gradients and te Derivative Q f(x+) δy P T δx R f(x) 0 x x+ Let P (x, f(x)) and Q(x+, f(x+)) denote two points on te curve of te function y = f(x) and let R denote te point of intersection of

More information

Research Article New Results on Multiple Solutions for Nth-Order Fuzzy Differential Equations under Generalized Differentiability

Research Article New Results on Multiple Solutions for Nth-Order Fuzzy Differential Equations under Generalized Differentiability Hindawi Publising Corporation Boundary Value Problems Volume 009, Article ID 395714, 13 pages doi:10.1155/009/395714 Researc Article New Results on Multiple Solutions for Nt-Order Fuzzy Differential Equations

More information

Continuity and Differentiability Worksheet

Continuity and Differentiability Worksheet Continuity and Differentiability Workseet (Be sure tat you can also do te grapical eercises from te tet- Tese were not included below! Typical problems are like problems -3, p. 6; -3, p. 7; 33-34, p. 7;

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

On convexity of polynomial paths and generalized majorizations

On convexity of polynomial paths and generalized majorizations On convexity of polynomial pats and generalized majorizations Marija Dodig Centro de Estruturas Lineares e Combinatórias, CELC, Universidade de Lisboa, Av. Prof. Gama Pinto 2, 1649-003 Lisboa, Portugal

More information

Analytic Functions. Differentiable Functions of a Complex Variable

Analytic Functions. Differentiable Functions of a Complex Variable Analytic Functions Differentiable Functions of a Complex Variable In tis capter, we sall generalize te ideas for polynomials power series of a complex variable we developed in te previous capter to general

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

Math 161 (33) - Final exam

Math 161 (33) - Final exam Name: Id #: Mat 161 (33) - Final exam Fall Quarter 2015 Wednesday December 9, 2015-10:30am to 12:30am Instructions: Prob. Points Score possible 1 25 2 25 3 25 4 25 TOTAL 75 (BEST 3) Read eac problem carefully.

More information

232 Calculus and Structures

232 Calculus and Structures 3 Calculus and Structures CHAPTER 17 JUSTIFICATION OF THE AREA AND SLOPE METHODS FOR EVALUATING BEAMS Calculus and Structures 33 Copyrigt Capter 17 JUSTIFICATION OF THE AREA AND SLOPE METHODS 17.1 THE

More information

REVIEW LAB ANSWER KEY

REVIEW LAB ANSWER KEY REVIEW LAB ANSWER KEY. Witout using SN, find te derivative of eac of te following (you do not need to simplify your answers): a. f x 3x 3 5x x 6 f x 3 3x 5 x 0 b. g x 4 x x x notice te trick ere! x x g

More information

MVT and Rolle s Theorem

MVT and Rolle s Theorem AP Calculus CHAPTER 4 WORKSHEET APPLICATIONS OF DIFFERENTIATION MVT and Rolle s Teorem Name Seat # Date UNLESS INDICATED, DO NOT USE YOUR CALCULATOR FOR ANY OF THESE QUESTIONS In problems 1 and, state

More information

Symmetry Labeling of Molecular Energies

Symmetry Labeling of Molecular Energies Capter 7. Symmetry Labeling of Molecular Energies Notes: Most of te material presented in tis capter is taken from Bunker and Jensen 1998, Cap. 6, and Bunker and Jensen 2005, Cap. 7. 7.1 Hamiltonian Symmetry

More information

Characterization of reducible hexagons and fast decomposition of elementary benzenoid graphs

Characterization of reducible hexagons and fast decomposition of elementary benzenoid graphs Discrete Applied Matematics 156 (2008) 1711 1724 www.elsevier.com/locate/dam Caracterization of reducible exagons and fast decomposition of elementary benzenoid graps Andrej Taranenko, Aleksander Vesel

More information

How to Find the Derivative of a Function: Calculus 1

How to Find the Derivative of a Function: Calculus 1 Introduction How to Find te Derivative of a Function: Calculus 1 Calculus is not an easy matematics course Te fact tat you ave enrolled in suc a difficult subject indicates tat you are interested in te

More information

3.4 Worksheet: Proof of the Chain Rule NAME

3.4 Worksheet: Proof of the Chain Rule NAME Mat 1170 3.4 Workseet: Proof of te Cain Rule NAME Te Cain Rule So far we are able to differentiate all types of functions. For example: polynomials, rational, root, and trigonometric functions. We are

More information

ch (for some fixed positive number c) reaching c

ch (for some fixed positive number c) reaching c GSTF Journal of Matematics Statistics and Operations Researc (JMSOR) Vol. No. September 05 DOI 0.60/s4086-05-000-z Nonlinear Piecewise-defined Difference Equations wit Reciprocal and Cubic Terms Ramadan

More information

Poisson Equation in Sobolev Spaces

Poisson Equation in Sobolev Spaces Poisson Equation in Sobolev Spaces OcMountain Dayligt Time. 6, 011 Today we discuss te Poisson equation in Sobolev spaces. It s existence, uniqueness, and regularity. Weak Solution. u = f in, u = g on

More information

Generic maximum nullity of a graph

Generic maximum nullity of a graph Generic maximum nullity of a grap Leslie Hogben Bryan Sader Marc 5, 2008 Abstract For a grap G of order n, te maximum nullity of G is defined to be te largest possible nullity over all real symmetric n

More information

Mathematics 5 Worksheet 11 Geometry, Tangency, and the Derivative

Mathematics 5 Worksheet 11 Geometry, Tangency, and the Derivative Matematics 5 Workseet 11 Geometry, Tangency, and te Derivative Problem 1. Find te equation of a line wit slope m tat intersects te point (3, 9). Solution. Te equation for a line passing troug a point (x

More information

lecture 26: Richardson extrapolation

lecture 26: Richardson extrapolation 43 lecture 26: Ricardson extrapolation 35 Ricardson extrapolation, Romberg integration Trougout numerical analysis, one encounters procedures tat apply some simple approximation (eg, linear interpolation)

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

Stationary Gaussian Markov Processes As Limits of Stationary Autoregressive Time Series

Stationary Gaussian Markov Processes As Limits of Stationary Autoregressive Time Series Stationary Gaussian Markov Processes As Limits of Stationary Autoregressive Time Series Lawrence D. Brown, Pilip A. Ernst, Larry Sepp, and Robert Wolpert August 27, 2015 Abstract We consider te class,

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

Derivation Of The Schwarzschild Radius Without General Relativity

Derivation Of The Schwarzschild Radius Without General Relativity Derivation Of Te Scwarzscild Radius Witout General Relativity In tis paper I present an alternative metod of deriving te Scwarzscild radius of a black ole. Te metod uses tree of te Planck units formulas:

More information

Exam 1 Review Solutions

Exam 1 Review Solutions Exam Review Solutions Please also review te old quizzes, and be sure tat you understand te omework problems. General notes: () Always give an algebraic reason for your answer (graps are not sufficient),

More information

Order of Accuracy. ũ h u Ch p, (1)

Order of Accuracy. ũ h u Ch p, (1) Order of Accuracy 1 Terminology We consider a numerical approximation of an exact value u. Te approximation depends on a small parameter, wic can be for instance te grid size or time step in a numerical

More information

SECTION 3.2: DERIVATIVE FUNCTIONS and DIFFERENTIABILITY

SECTION 3.2: DERIVATIVE FUNCTIONS and DIFFERENTIABILITY (Section 3.2: Derivative Functions and Differentiability) 3.2.1 SECTION 3.2: DERIVATIVE FUNCTIONS and DIFFERENTIABILITY LEARNING OBJECTIVES Know, understand, and apply te Limit Definition of te Derivative

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

HOW TO DEAL WITH FFT SAMPLING INFLUENCES ON ADEV CALCULATIONS

HOW TO DEAL WITH FFT SAMPLING INFLUENCES ON ADEV CALCULATIONS HOW TO DEAL WITH FFT SAMPLING INFLUENCES ON ADEV CALCULATIONS Po-Ceng Cang National Standard Time & Frequency Lab., TL, Taiwan 1, Lane 551, Min-Tsu Road, Sec. 5, Yang-Mei, Taoyuan, Taiwan 36 Tel: 886 3

More information

Justifications for Common Knowledge

Justifications for Common Knowledge - Justifications for Common Knowledge Samuel Buceli Roman Kuznets Tomas Studer Institut für Informatik und angewandte Matematik, Universität Bern Neubrückstrasse 10, CH-3012 Bern (Switzerland) {buceli,

More information

A SHORT INTRODUCTION TO BANACH LATTICES AND

A SHORT INTRODUCTION TO BANACH LATTICES AND CHAPTER A SHORT INTRODUCTION TO BANACH LATTICES AND POSITIVE OPERATORS In tis capter we give a brief introduction to Banac lattices and positive operators. Most results of tis capter can be found, e.g.,

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

Some Review Problems for First Midterm Mathematics 1300, Calculus 1

Some Review Problems for First Midterm Mathematics 1300, Calculus 1 Some Review Problems for First Midterm Matematics 00, Calculus. Consider te trigonometric function f(t) wose grap is sown below. Write down a possible formula for f(t). Tis function appears to be an odd,

More information

DIGRAPHS FROM POWERS MODULO p

DIGRAPHS FROM POWERS MODULO p DIGRAPHS FROM POWERS MODULO p Caroline Luceta Box 111 GCC, 100 Campus Drive, Grove City PA 1617 USA Eli Miller PO Box 410, Sumneytown, PA 18084 USA Clifford Reiter Department of Matematics, Lafayette College,

More information

Numerical Differentiation

Numerical Differentiation Numerical Differentiation Finite Difference Formulas for te first derivative (Using Taylor Expansion tecnique) (section 8.3.) Suppose tat f() = g() is a function of te variable, and tat as 0 te function

More information

Domination Problems in Nowhere-Dense Classes of Graphs

Domination Problems in Nowhere-Dense Classes of Graphs LIPIcs Leibniz International Proceedings in Informatics Domination Problems in Nowere-Dense Classes of Graps Anuj Dawar 1, Stepan Kreutzer 2 1 University of Cambridge Computer Lab, U.K. anuj.dawar@cl.cam.ac.uk

More information

7 Semiparametric Methods and Partially Linear Regression

7 Semiparametric Methods and Partially Linear Regression 7 Semiparametric Metods and Partially Linear Regression 7. Overview A model is called semiparametric if it is described by and were is nite-dimensional (e.g. parametric) and is in nite-dimensional (nonparametric).

More information

3.1 Extreme Values of a Function

3.1 Extreme Values of a Function .1 Etreme Values of a Function Section.1 Notes Page 1 One application of te derivative is finding minimum and maimum values off a grap. In precalculus we were only able to do tis wit quadratics by find

More information

THE IDEA OF DIFFERENTIABILITY FOR FUNCTIONS OF SEVERAL VARIABLES Math 225

THE IDEA OF DIFFERENTIABILITY FOR FUNCTIONS OF SEVERAL VARIABLES Math 225 THE IDEA OF DIFFERENTIABILITY FOR FUNCTIONS OF SEVERAL VARIABLES Mat 225 As we ave seen, te definition of derivative for a Mat 111 function g : R R and for acurveγ : R E n are te same, except for interpretation:

More information

MA455 Manifolds Solutions 1 May 2008

MA455 Manifolds Solutions 1 May 2008 MA455 Manifolds Solutions 1 May 2008 1. (i) Given real numbers a < b, find a diffeomorpism (a, b) R. Solution: For example first map (a, b) to (0, π/2) and ten map (0, π/2) diffeomorpically to R using

More information

7.1 Using Antiderivatives to find Area

7.1 Using Antiderivatives to find Area 7.1 Using Antiderivatives to find Area Introduction finding te area under te grap of a nonnegative, continuous function f In tis section a formula is obtained for finding te area of te region bounded between

More information

Lecture XVII. Abstract We introduce the concept of directional derivative of a scalar function and discuss its relation with the gradient operator.

Lecture XVII. Abstract We introduce the concept of directional derivative of a scalar function and discuss its relation with the gradient operator. Lecture XVII Abstract We introduce te concept of directional derivative of a scalar function and discuss its relation wit te gradient operator. Directional derivative and gradient Te directional derivative

More information

NUMERICAL DIFFERENTIATION. James T. Smith San Francisco State University. In calculus classes, you compute derivatives algebraically: for example,

NUMERICAL DIFFERENTIATION. James T. Smith San Francisco State University. In calculus classes, you compute derivatives algebraically: for example, NUMERICAL DIFFERENTIATION James T Smit San Francisco State University In calculus classes, you compute derivatives algebraically: for example, f( x) = x + x f ( x) = x x Tis tecnique requires your knowing

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

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

Problem Solving. Problem Solving Process

Problem Solving. Problem Solving Process Problem Solving One of te primary tasks for engineers is often solving problems. It is wat tey are, or sould be, good at. Solving engineering problems requires more tan just learning new terms, ideas and

More information

HOMEWORK HELP 2 FOR MATH 151

HOMEWORK HELP 2 FOR MATH 151 HOMEWORK HELP 2 FOR MATH 151 Here we go; te second round of omework elp. If tere are oters you would like to see, let me know! 2.4, 43 and 44 At wat points are te functions f(x) and g(x) = xf(x)continuous,

More information

EFFICIENCY OF MODEL-ASSISTED REGRESSION ESTIMATORS IN SAMPLE SURVEYS

EFFICIENCY OF MODEL-ASSISTED REGRESSION ESTIMATORS IN SAMPLE SURVEYS Statistica Sinica 24 2014, 395-414 doi:ttp://dx.doi.org/10.5705/ss.2012.064 EFFICIENCY OF MODEL-ASSISTED REGRESSION ESTIMATORS IN SAMPLE SURVEYS Jun Sao 1,2 and Seng Wang 3 1 East Cina Normal University,

More information

WYSE Academic Challenge 2004 Sectional Mathematics Solution Set

WYSE Academic Challenge 2004 Sectional Mathematics Solution Set WYSE Academic Callenge 00 Sectional Matematics Solution Set. Answer: B. Since te equation can be written in te form x + y, we ave a major 5 semi-axis of lengt 5 and minor semi-axis of lengt. Tis means

More information

Complexity of Decoding Positive-Rate Reed-Solomon Codes

Complexity of Decoding Positive-Rate Reed-Solomon Codes Complexity of Decoding Positive-Rate Reed-Solomon Codes Qi Ceng 1 and Daqing Wan 1 Scool of Computer Science Te University of Oklaoma Norman, OK73019 Email: qceng@cs.ou.edu Department of Matematics University

More information

Brazilian Journal of Physics, vol. 29, no. 1, March, Ensemble and their Parameter Dierentiation. A. K. Rajagopal. Naval Research Laboratory,

Brazilian Journal of Physics, vol. 29, no. 1, March, Ensemble and their Parameter Dierentiation. A. K. Rajagopal. Naval Research Laboratory, Brazilian Journal of Pysics, vol. 29, no. 1, Marc, 1999 61 Fractional Powers of Operators of sallis Ensemble and teir Parameter Dierentiation A. K. Rajagopal Naval Researc Laboratory, Wasington D. C. 2375-532,

More information

f a h f a h h lim lim

f a h f a h h lim lim Te Derivative Te derivative of a function f at a (denoted f a) is f a if tis it exists. An alternative way of defining f a is f a x a fa fa fx fa x a Note tat te tangent line to te grap of f at te point

More information

= 0 and states ''hence there is a stationary point'' All aspects of the proof dx must be correct (c)

= 0 and states ''hence there is a stationary point'' All aspects of the proof dx must be correct (c) Paper 1: Pure Matematics 1 Mark Sceme 1(a) (i) (ii) d d y 3 1x 4x x M1 A1 d y dx 1.1b 1.1b 36x 48x A1ft 1.1b Substitutes x = into teir dx (3) 3 1 4 Sows d y 0 and states ''ence tere is a stationary point''

More information

5 Ordinary Differential Equations: Finite Difference Methods for Boundary Problems

5 Ordinary Differential Equations: Finite Difference Methods for Boundary Problems 5 Ordinary Differential Equations: Finite Difference Metods for Boundary Problems Read sections 10.1, 10.2, 10.4 Review questions 10.1 10.4, 10.8 10.9, 10.13 5.1 Introduction In te previous capters we

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

Time (hours) Morphine sulfate (mg)

Time (hours) Morphine sulfate (mg) Mat Xa Fall 2002 Review Notes Limits and Definition of Derivative Important Information: 1 According to te most recent information from te Registrar, te Xa final exam will be eld from 9:15 am to 12:15

More information

1 + t5 dt with respect to x. du = 2. dg du = f(u). du dx. dg dx = dg. du du. dg du. dx = 4x3. - page 1 -

1 + t5 dt with respect to x. du = 2. dg du = f(u). du dx. dg dx = dg. du du. dg du. dx = 4x3. - page 1 - Eercise. Find te derivative of g( 3 + t5 dt wit respect to. Solution: Te integrand is f(t + t 5. By FTC, f( + 5. Eercise. Find te derivative of e t2 dt wit respect to. Solution: Te integrand is f(t e t2.

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

Consider a function f we ll specify which assumptions we need to make about it in a minute. Let us reformulate the integral. 1 f(x) dx.

Consider a function f we ll specify which assumptions we need to make about it in a minute. Let us reformulate the integral. 1 f(x) dx. Capter 2 Integrals as sums and derivatives as differences We now switc to te simplest metods for integrating or differentiating a function from its function samples. A careful study of Taylor expansions

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

Introduction to Derivatives

Introduction to Derivatives Introduction to Derivatives 5-Minute Review: Instantaneous Rates and Tangent Slope Recall te analogy tat we developed earlier First we saw tat te secant slope of te line troug te two points (a, f (a))

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

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

Section 2.7 Derivatives and Rates of Change Part II Section 2.8 The Derivative as a Function. at the point a, to be. = at time t = a is

Section 2.7 Derivatives and Rates of Change Part II Section 2.8 The Derivative as a Function. at the point a, to be. = at time t = a is Mat 180 www.timetodare.com Section.7 Derivatives and Rates of Cange Part II Section.8 Te Derivative as a Function Derivatives ( ) In te previous section we defined te slope of te tangent to a curve wit

More information

Differential Calculus (The basics) Prepared by Mr. C. Hull

Differential Calculus (The basics) Prepared by Mr. C. Hull Differential Calculus Te basics) A : Limits In tis work on limits, we will deal only wit functions i.e. tose relationsips in wic an input variable ) defines a unique output variable y). Wen we work wit

More information

Financial Econometrics Prof. Massimo Guidolin

Financial Econometrics Prof. Massimo Guidolin CLEFIN A.A. 2010/2011 Financial Econometrics Prof. Massimo Guidolin A Quick Review of Basic Estimation Metods 1. Were te OLS World Ends... Consider two time series 1: = { 1 2 } and 1: = { 1 2 }. At tis

More information

SAHLQVIST THEOREM FOR MODAL FIXED POINT LOGIC

SAHLQVIST THEOREM FOR MODAL FIXED POINT LOGIC SAHLQVIST THEOREM FOR MODAL FIXED POINT LOGIC NICK BEZHANISHVILI AND IAN HODKINSON Abstract. We define Salqvist fixed point formulas. By extending te tecnique of Sambin and Vaccaro we sow tat (1) for eac

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

1. Questions (a) through (e) refer to the graph of the function f given below. (A) 0 (B) 1 (C) 2 (D) 4 (E) does not exist

1. Questions (a) through (e) refer to the graph of the function f given below. (A) 0 (B) 1 (C) 2 (D) 4 (E) does not exist Mat 1120 Calculus Test 2. October 18, 2001 Your name Te multiple coice problems count 4 points eac. In te multiple coice section, circle te correct coice (or coices). You must sow your work on te oter

More information

Improved Algorithms for Largest Cardinality 2-Interval Pattern Problem

Improved Algorithms for Largest Cardinality 2-Interval Pattern Problem Journal of Combinatorial Optimization manuscript No. (will be inserted by te editor) Improved Algoritms for Largest Cardinality 2-Interval Pattern Problem Erdong Cen, Linji Yang, Hao Yuan Department of

More information

5.1 We will begin this section with the definition of a rational expression. We

5.1 We will begin this section with the definition of a rational expression. We Basic Properties and Reducing to Lowest Terms 5.1 We will begin tis section wit te definition of a rational epression. We will ten state te two basic properties associated wit rational epressions and go

More information

SECTION 1.10: DIFFERENCE QUOTIENTS LEARNING OBJECTIVES

SECTION 1.10: DIFFERENCE QUOTIENTS LEARNING OBJECTIVES (Section.0: Difference Quotients).0. SECTION.0: DIFFERENCE QUOTIENTS LEARNING OBJECTIVES Define average rate of cange (and average velocity) algebraically and grapically. Be able to identify, construct,

More information

Integral Calculus, dealing with areas and volumes, and approximate areas under and between curves.

Integral Calculus, dealing with areas and volumes, and approximate areas under and between curves. Calculus can be divided into two ke areas: Differential Calculus dealing wit its, rates of cange, tangents and normals to curves, curve sketcing, and applications to maima and minima problems Integral

More information

4. The slope of the line 2x 7y = 8 is (a) 2/7 (b) 7/2 (c) 2 (d) 2/7 (e) None of these.

4. The slope of the line 2x 7y = 8 is (a) 2/7 (b) 7/2 (c) 2 (d) 2/7 (e) None of these. Mat 11. Test Form N Fall 016 Name. Instructions. Te first eleven problems are wort points eac. Te last six problems are wort 5 points eac. For te last six problems, you must use relevant metods of algebra

More information

Influence of the Stepsize on Hyers Ulam Stability of First-Order Homogeneous Linear Difference Equations

Influence of the Stepsize on Hyers Ulam Stability of First-Order Homogeneous Linear Difference Equations International Journal of Difference Equations ISSN 0973-6069, Volume 12, Number 2, pp. 281 302 (2017) ttp://campus.mst.edu/ijde Influence of te Stepsize on Hyers Ulam Stability of First-Order Homogeneous

More information

1 Lecture 13: The derivative as a function.

1 Lecture 13: The derivative as a function. 1 Lecture 13: Te erivative as a function. 1.1 Outline Definition of te erivative as a function. efinitions of ifferentiability. Power rule, erivative te exponential function Derivative of a sum an a multiple

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

On the Concept of Returns to Scale: Revisited

On the Concept of Returns to Scale: Revisited 3 J. Asian Dev. Stud, Vol. 5, Issue, (Marc 206) ISSN 2304-375X On te Concept of Returns to Scale: Revisited Parvez Azim Abstract Tis paper sows w it is tat in Economics text books and literature we invariabl

More information

Chapter 2 Limits and Continuity

Chapter 2 Limits and Continuity 4 Section. Capter Limits and Continuity Section. Rates of Cange and Limits (pp. 6) Quick Review.. f () ( ) () 4 0. f () 4( ) 4. f () sin sin 0 4. f (). 4 4 4 6. c c c 7. 8. c d d c d d c d c 9. 8 ( )(

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

2.8 The Derivative as a Function

2.8 The Derivative as a Function .8 Te Derivative as a Function Typically, we can find te derivative of a function f at many points of its domain: Definition. Suppose tat f is a function wic is differentiable at every point of an open

More information