[Lecture Notes in Computer Science] Typed Lambda Calculi and Applications Volume 2044 || Parigot’s...

15
Parigot’s Second Order λμ-Calculus and Inductive Types Ralph Matthes ? Institut f¨ ur Informatik der Ludwig-Maximilians-Universit¨at M¨ unchen Oettingenstraße 67, D-80538 M¨ unchen, Germany [email protected] Abstract. A new proof of strong normalization of Parigot’s (second order) λμ-calculus is given by a reduction-preserving embedding into system F (second order polymorphic λ-calculus). The main idea is to use the least stable supertype for any type. These non-strictly positive inductive types and their associated iteration principle are available in system F, and allow to give a translation vaguely related to CPS trans- lations (corresponding to the Kolmogorov embedding of classical logic into intuitionistic logic). However, they simulate Parigot’s μ-reductions whereas CPS translations hide them. As a major advantage, this embedding does not use the idea of reducing stability (¬¬φ φ) to that for atomic formulae. Therefore, it even extends to non-interleaving positive fixed-point types. As a non-trivial application, strong normalization of λμ-calculus, extended by primitive recursion on monotone inductive types, is established. 1 Introduction λμ-calculus [12] essentially is the extension of natural deduction by “reductio ad absurdum” (RAA), i. e., by a term formation rule corresponding to stability (¬¬ρ ρ, also called “duplex negatio affirmat”) and by rewrite rules for the simplification of the application of elimination rules to RAA (in the case of - elimination this corresponds to the fact that the stability of ρ σ is derivable from that of σ). In [14]—besides a direct proof of strong normalization via saturated sets— we find a reduction of the proof of strong normalization of λμ-calculus to the well-known strong normalization of system F. The proof is quite intricate since the considered CPS translation maps the RAA redexes and their contracta to the same term, and hence needs additional arguments on RAA reductions alone to guarantee strong normalization of the whole calculus. Since these μ-reductions have the flavour of iteration (the application occurs in the contractum in a controlled fashion), it has been tempting to explain μ- reductions via inductive types. This is achieved by studying the “stabilization” ? I am grateful for an invitation to present a preliminary version of the present results at the “S´ eminaire Preuves, Programmes et Syst` emes” at Paris VII in October 2000. S. Abramsky (Ed.): TLCA 2001, LNCS 2044, pp. 329–343, 2001. c Springer-Verlag Berlin Heidelberg 2001

Transcript of [Lecture Notes in Computer Science] Typed Lambda Calculi and Applications Volume 2044 || Parigot’s...

Parigot’s Second Order λµ-Calculus andInductive Types

Ralph Matthes?

Institut fur Informatik der Ludwig-Maximilians-Universitat MunchenOettingenstraße 67, D-80538 Munchen, Germany

[email protected]

Abstract. A new proof of strong normalization of Parigot’s (secondorder) λµ-calculus is given by a reduction-preserving embedding intosystem F (second order polymorphic λ-calculus). The main idea is touse the least stable supertype for any type. These non-strictly positiveinductive types and their associated iteration principle are available insystem F, and allow to give a translation vaguely related to CPS trans-lations (corresponding to the Kolmogorov embedding of classical logicinto intuitionistic logic). However, they simulate Parigot’s µ-reductionswhereas CPS translations hide them.As a major advantage, this embedding does not use the idea of reducingstability (¬¬φ → φ) to that for atomic formulae. Therefore, it evenextends to non-interleaving positive fixed-point types. As a non-trivialapplication, strong normalization of λµ-calculus, extended by primitiverecursion on monotone inductive types, is established.

1 Introduction

λµ-calculus [12] essentially is the extension of natural deduction by “reductioad absurdum” (RAA), i. e., by a term formation rule corresponding to stability(¬¬ρ → ρ, also called “duplex negatio affirmat”) and by rewrite rules for thesimplification of the application of elimination rules to RAA (in the case of →-elimination this corresponds to the fact that the stability of ρ → σ is derivablefrom that of σ).

In [14]—besides a direct proof of strong normalization via saturated sets—we find a reduction of the proof of strong normalization of λµ-calculus to thewell-known strong normalization of system F. The proof is quite intricate sincethe considered CPS translation maps the RAA redexes and their contracta tothe same term, and hence needs additional arguments on RAA reductions aloneto guarantee strong normalization of the whole calculus.

Since these µ-reductions have the flavour of iteration (the application occursin the contractum in a controlled fashion), it has been tempting to explain µ-reductions via inductive types. This is achieved by studying the “stabilization”? I am grateful for an invitation to present a preliminary version of the present results

at the “Seminaire Preuves, Programmes et Systemes” at Paris VII in October 2000.

S. Abramsky (Ed.): TLCA 2001, LNCS 2044, pp. 329–343, 2001.c© Springer-Verlag Berlin Heidelberg 2001

330 R. Matthes

]ρ of any type ρ. It is the least stable type (i. e., there is a constant of type¬¬]ρ → ]ρ) such that ρ is included in ]ρ. From the minimality, we get aniteration principle which in fact simulates µ-reductions and therefore illustratesnicely what can be achieved with non-strictly positive inductive types.

There is an even easier embedding into system F [6] which exploits the factthat the elimination rules deconstruct the type to be eliminated. Our approachis not based on this observation, and it even turns out that also type conceptscan be treated where this observation cannot be made any longer, i. e., where theelimination does not deconstruct the type to be eliminated. This is exemplifiedwith the addition of non-interleaving positive fixed-point types. It is very un-likely that they can be embedded into system F [17]. Therefore, we cannot expectto get an embedding into system F but we do get an embedding into system Faugmented with those fixed-point types. µ-reductions for fixed-point types stillhave a connection between the type eliminated and the result type. This is nottrue of primitive recursion: The inductive type and the type of results of thefunctional defined by primitive recursion may be completely unrelated. Never-theless, we can also treat this extension. Firstly, the most general formulationof primitive recursion is introduced: primitive recursion on monotone inductivetypes. Secondly, an embedding into non-interleaving positive fixed-point typesis established, which moreover interacts nicely with µ-reductions. Therefore, wecan even prove that λµ-calculus with primitive recursion on monotone inductivetypes (and µ-reductions for every type construct) is strongly normalizing.

The next section recalls system F, in section 3 we review λµ-calculus anddiscuss variations on the definition. Section 4 presents the stabilization ]ρ andthe associated iteration principle and explains it via an impredicative encod-ing, whereas in section 5 this system becomes the target of an embedding ofλµ-calculus. Substitution lemmas are proven in great detail in order to give afeeling for the embedding. Fixed-point types are introduced in section 6, and theembedding of the previous section is extended to cover fixed-point types (in anyof the systems studied before). This extension is astonishingly straightforward.In the final section 7 monotone inductive types with primitive recursion aredescribed. A new embedding of monotone inductive types with primitive recur-sion into non-interleaving positive fixed-point types gives the final normalizationtheorem.

2 System F

We consider system F (see e. g. [5]) only with function types and universal types,i. e., we have infinitely many type variables (denoted by α, β, . . . ) and withtypes ρ and σ we also have the function type ρ → σ. Moreover, given a variableα and a type ρ we form the universal type ∀αρ. The ∀ binds α in ρ. The renamingconvention for bound variables is adopted. Let FV(ρ) be the set of type variablesoccurring free in ρ.

The terms of F are presented as in [5], i. e., without contexts and with fixedtypes (see [1, p. 159] for comments on this original typing a la Church). We

Parigot’s Second Order λµ-Calculus and Inductive Types 331

have infinitely many term variables with types (denoted e. g. by xρ), lambdaabstraction (λxρrσ)ρ→σ for terms binding xρ in rσ, term application (rρ→σsρ)σ,lambda abstraction (Λαrρ)∀αρ for types (under the usual proviso that α /∈ FV(σ)for any xσ free in rρ) and type application (r∀αρσ)ρ[α:=σ]. We freely use therenaming convention for bound term and type variables of terms, and moreoveromit type superscripts of the terms. We do this even in case of typed variableswhich are in fact pairs of variable names and types, but only if the type can bereconstructed from a lambda abstraction binding a variable with the same name.(The interested reader may consult the discussion of these issues in [7, sections2.1.2 and 2.2.6.].) Instead of rρ we often write that r has type ρ or even r : ρ.We let application associate to the left and write λxρ.r to avoid parenthesizingof r.

Beta reduction . for system F is as usual given by

(β→) (λxρr)s � r[xρ := s](β∀) (Λαr)σ � r[α := σ] .

Here, we used the result r[xρ := s] of the capture-free substitution of the typedvariable xρ by the term s of type ρ in r and the result r[α := σ] of the capture-free substitution of the type variable α by the type σ in r. The reduction relation→ is defined as the term closure of �. The main theorem on F (due to Girard)is that → is strongly normalizing, i. e., there are no infinite reduction sequencesr1 → r2 → r3 → · · · (In [7] all the definitions and lemmas are given quitecarefully so as to make sure that the main theorem can also be proved in thispresentation without contexts.)

3 Second Order λµ-Calculus

We present λµ-calculus [12] not with sequents but in the same style as system F.Although λµ-calculus extends system F by the classical law of reductio ad absur-dum, there is neither falsity nor negation in the type system. Hence, we keep thetype system of F. However, there is a second kind of variables (called µ-variablesin [12]), denoted by a, b, c, . . . They are also paired with types, e. g. aρ, whoseinterpretation is the assumption of the negation of ρ.

The term system is essentially that of F, extended by the rule of reductio adabsurdum: µaσ.[bρ]rρ is a term of type σ, binding aσ in [bρ]rρ. The latter is noterm but will be called a named term as in [12].

The named term [bρ]rρ is to be understood as the application of the µ-variablebρ—assuming the negation of ρ—to the term r proving ρ. Hence, bρ is free in[bρ]rρ. Moreover, a named term would prove falsity (if it were a legal term andfalsity were included into the type system), hence [bρ]rρ gives falsity under theassumption of the negation of σ, expressed by the bound variable aσ. Reductioad absurdum yields σ, hence justifying the type assignment for µaσ.[bρ]rρ.

As a word of caution, aρ is not a term and not legal for lambda abstraction.(Although we allow renaming of bound variables without explicit mention, we

332 R. Matthes

have to keep to the same sort: either normal variable—called λ-variable in [12]—or µ-variable.)

The beta reduction relation � of system F is extended by µ-reductions asfollows:

(µ→) (µaρ→σ.r)s � µbσ.r[xρ→σ.[aρ→σ]x := [b](xs)

]

(µ∀) (µa∀αρ.r)σ � µbρ[α:=σ].r[x∀αρ.[a∀αρ]x := [b](xσ)

].

Here, r is always a named term, and r[xρ.[aρ]x := [bσ]tσ

]denotes the result

of replacing inductively (bottom-up) in r each subexpression of the form [aρ]uρ

by [bσ](t[xρ := u]) for any term u of type ρ (“xρ.” binds xρ in this substitutionnotation, especially in t). In fact, this is only another notation for the samesubstitution concept in [12].

Remark 1. For our present purposes it seems that this unusual notion of sub-stitution cannot be avoided. If we had type ⊥, we could formulate reductio adabsurdum without reference to µ-variables: Write ¬ρ for ρ → ⊥ and extend Fonly by terms µx¬ρ.r⊥ of type ρ which bind x¬ρ in r. (We do not need the extraconcept of named terms since we simply have x¬ρrρ of type ⊥ instead of somenamed term [aρ]rρ.) Now, we can extend � of system F by

(µ→)′ (µx¬(ρ→σ).r)s � µy¬σ.r[x¬(ρ→σ) := λzρ→σ.y(zs)](µ∀)′ (µx¬∀αρ.r)σ � µy¬ρ[α:=σ].r[x¬∀αρ := λz∀αρ.y(zσ)] .

This formulation1 only needs standard substitution and is obviously slightlymore general than λµ (λµ imposes a stronger discipline on term formation).By the method of saturated sets, it is possible to show strong normalization ofthe term closure → of �. Of course, this requires an appropriate formulation ofthe notion of saturation. Moreover, we would reprove strong normalization of Finstead of using this fact via an embedding into system F. Unfortunately, ourembedding shown below does not extend to this reformulation which in somesense abuses the function types (via negation) for the explanation of reductio adabsurdum. Therefore, we abandon this reformulation.

Another approach to avoid the peculiar substitution is taken in [18, pp.17–20] and [2]: Instead of named terms [a]r, we have responses/commands [C]r(resp. 〈r | C〉) where C is a stack of terms with a µ-variable a at the bottom.This more general view of continuations/contexts allows to define the equalityrelation associated with µ-reduction by help of ordinary substitution. Whilethis is sufficient for the study of equality in [18], we are interested in strongnormalization. Hence, also the call-by-value and call-by-name formulations in[2] which are possible with ordinary substitution do not suffice.

Second order λµ-calculus is strongly normalizing [13].1 It amounts to a λ-calculus notation for classical natural deduction in the style of

Prawitz [16].

Parigot’s Second Order λµ-Calculus and Inductive Types 333

4 Extension of F by Iteration on Stabilization

If ¬¬ρ → ρ is provable, then ρ is called stable. We consider an extension F]

of system F by a least stable supertype ]ρ for any type ρ. This is expressedas follows: We add a type constant ⊥ for falsity2 (and set ¬ρ := ρ → ⊥) andfor every type ρ we assume the type ]ρ (] is a unary type former) called thestabilization of ρ.

The term system reflects that ]ρ is stable, that ρ embeds into ]ρ and that ]ρis the least type with these properties.3 We assume constants I]ρ of type ρ → ]ρand S]ρ of type ¬¬]ρ → ]ρ and add a term formation rule: If r has type ]ρ, τ isa type and s1 : ρ → τ and s2 : ¬¬τ → τ then rEτs1s2 is a term of type τ .

I]ρ : ρ → ]ρ S]ρ : ¬¬]ρ → ]ρ

r : ]ρ s1 : ρ → τ s2 : ¬¬τ → τ

rEτs1s2 : τ

Extend � of system F by

(]I) (I]ρr)Eτs1s2 � s1r

(]S) (S]ρr)Eτs1s2 � s2

(λy¬τ .r(λz]ρ.y(zEτs1s2))

).

In the second rule, we assume that y¬τ and z]ρ do not occur free in s1 or s2. Let→ be the term closure of �. Clearly, it enjoys subject reduction, i. e., if rρ → sσ

then ρ = σ.The presently defined system F] will be called F with iteration on stabiliza-

tion. Let F⊥ be F extended only by the type constant ⊥. Trivially, F⊥ inheritsstrong normalization from F since we consider neither a term formation nor arewrite rule for ⊥.

Definition 1. A type-respecting reduction-preserving embedding (embedding forshort) of a typed term rewrite system S into a typed term rewrite system S ′ isa function −′ (the − sign represents the indefinite argument of the function ′)which assigns to every type ρ of S a type ρ′ of S ′ and to every term r of type ρof S a term r′ of the (image) type ρ′ of S ′ such that the following implicationholds: If r → s in S, then r′ →+ s′ in S ′. (→+ denotes the transitive closure of→.)

Obviously, if there is an embedding of S into S ′, then strong normalization ofS ′ is inherited by S.2 ⊥ is just some constant: We do not assume an elimination rule expressing ex falsum

quodlibet.3 The author has recently been informed that already [15, p. 110] describes a similar

idea: An inductively defined predicate may be “made classical” by adding stabilityas another clause to the definition. This turns the definition into a non-strictlypositive one and enforces stability. Note, however, that non-strictly positive inductivedefinitions lead to inconsistencies in higher-order predicate logic, see the examplereported in [15, p. 108]. In the framework of system F, we are in the fortunatesituation that arbitrary types may be stabilized without harm to consistency.

334 R. Matthes

Lemma 1. There is an embedding of F] into F⊥.

Proof. By the very description, ]ρ is nothing but the non-strictly positive induc-tive type µα.ρ+¬¬α with α /∈ FV(ρ), formulated without the sum type (do notmix up the notation µαρ for inductive types with µar for reductio ad absurdum).Its canonical polymorphic encoding would be ∀α.((ρ + ¬¬α) → α) → α whichcan be simplified to ∀α.(ρ → α) → (¬¬α → α) → α.

By iteration on the type ρ of F] define the type ρ′ of F⊥. This shall be donehomomorphically in all cases except:

(]ρ)′ := ∀α.(ρ′ → α) → (¬¬α → α) → α

with α /∈ FV(ρ). Clearly, FV(ρ′) = FV(ρ) and (ρ[α := σ])′ = ρ′[α := σ′].By iteration on the term r : ρ of F] define the term r′ : ρ′ of F⊥. (Simulta-

neously, one has to show that the free variables of r′ are the xρ′with xρ free in

r.) We only consider the non-homomorphic cases.

(I]ρ)′ := λxρ′Λαλxρ′→α

1 λx¬¬α→α2 .x1x

(S]ρ)′ := λx¬¬(]ρ)′Λαλxρ′→α

1 λx¬¬α→α2 .x2

(λy¬α.x(λz(]ρ)′

.y(zαx1x2)))

(rEτs1s2)′ := r′τ ′s′1s

′2 .

Since (r[xρ := s])′ = r′[xρ′:= s′] and (r[α := ρ])′ = r′[α := ρ′], this translation

respects (β→) and (β∀). It is a trivial calculation to prove that

((I]ρr)Eτs1s2)′ →4 (s1r)′ and((S]ρr)Eτs1s2

)′→4

(s2

(λy¬τ .r(λz]ρ.y(zEτs1s2))

))′.

(We use →n to express n steps of →.) Therefore, −′ is indeed an embedding. ut

5 Embedding Second Order λµ-Calculus into F withIteration on Stabilization

In contrast to the proof of strong normalization of second order λµ-calculus in[14] by a CPS translation which maps the terms to the left and the right side of(µ→) and (µ∀) to the same term, respectively, and therefore needs an additionalargument for the strong normalization of (µ→) and (µ∀) alone (without (β→)and (β∀)), our translation given below also simulates those µ-reductions andhence is an embedding. Note, however, that [6] presents an embedding which iseven easier—the only non-homomorphic rule is (∀αρ)′ := ∀α.(¬¬α → α) → ρ′—but which heavily uses the fact that stability may be proved for those translatedtypes from stability of their free type variables. This will rule out the extensionto fixed-point types to be studied in the next section.

Define the type ρ∗ of F] by iteration on the type ρ of second order λµ-calculus(in the sequel denoted by λµ) as follows:

α∗ := α(ρ → σ)∗ := ]ρ∗ → ]σ∗

(∀αρ)∗ := ∀α ]ρ∗

Parigot’s Second Order λµ-Calculus and Inductive Types 335

By induction on ρ one verifies that (ρ[α := σ])∗ = ρ∗[α := σ∗] and FV(ρ∗) =FV(ρ).The type translation of our embedding is given by ρ′ := ]ρ∗. Therefore,

α′ = ]α(ρ → σ)′ = ](ρ′ → σ′)

(∀αρ)′ = ](∀αρ′)

and (ρ[α := σ])′ = ρ′[α := σ∗], and also FV(ρ′) = FV(ρ).Of course, ρ′ could have been defined directly by iteration on ρ without

reference to ρ∗. However, the substitution property would not have looked sonatural. Note that if we had used ¬¬ρ instead of ]ρ everywhere in this definition,we would have arrived at Kolmogorov’s negative translation used in [3,14] (thecorresponding term translation would have been a translation in continuation-passing style).

By iteration on the term r : ρ of λµ define the term r′ : ρ′ of F]. (Simul-taneously, one has to show that the free variables of r′ are the xρ′

with xρ afree normal variable in r and the a¬σ′

with aσ a free µ-variable in r. Hence, weassume that the names of the µ-variables are also names of variables of our F].)

(xρ)′ := xρ′

(λxρrσ)′ := I(ρ→σ)′(λxρ′r′)

(rρ→σsρ)′ := r′Eσ′(λzρ′→σ′.zs′)Sσ′

(Λαrρ)′ := I(∀αρ)′(Λαr′)(r∀αρσ)′ := r′E(ρ[α:=σ])′(λz∀αρ′

.zσ∗)S(ρ[α:=σ])′

(µaσ.[bρ]rρ)′ := Sσ′(λa¬σ′.b¬ρ′

r′)

In the third clause, we assume that zρ→σ is not free in s. Note that the fourthclause is legal since the proviso on the formation of Λαr′ is fulfilled by ourstatement which is proved simultaneously with the definition. The substitutionproperty (ρ[α := σ])′ = ρ′[α := σ∗] is heavily used in the fifth clause. Finallyobserve that no definition is given for a′ since a is no term.

It will now be useful to treat named terms as if they were terms. Therefore,the sixth rule decomposes into:

([bρ]rρ)′ := b¬ρ′r′ : ⊥

(µaσ.r)′ := Sσ′(λa¬σ′.r′)

Lemma 2. (r[xρ := s])′ = r′[xρ′:= s′] and (r[α := ρ])′ = r′[α := ρ∗].

Proof. Induction on r. ut

Corollary 1. ((λxρr)s)′ →3 (r[xρ := s])′ and ((Λαr)σ)′ →3 (r[α := σ])′.

Let →∗ denote the reflexive transitive closure of →.

Lemma 3. r′[a¬ρ′:= λxρ′

.b¬σ′t′] →∗ (

r[xρ.[aρ]x := [bσ]t])′ for t of type σ.

336 R. Matthes

Proof. By induction on named terms and terms r. (“Special” substitution canreadily be extended to terms r.) We show the only non-trivial case where r hasthe form [aρ]sρ, hence with the same µ-variable aρ. The left-hand side becomes

(λxρ′.b¬σ′

t′)s′[a¬ρ′:= λxρ′

.b¬σ′t′].

One beta reduction step yields

b¬σ′t′[xρ′

:= s′[a¬ρ′:= λxρ′

.b¬σ′t′]

].

By induction hypothesis,

s′[a¬ρ′:= λxρ′

.b¬σ′t′] →∗ (

s[xρ.[aρ]x := [bσ]t])′

.

Hence, →∗ leads to

b¬σ′t′[xρ′

:=(s[xρ.[aρ]x := [bσ]t]

)′].

By the previous lemma,

t′[xρ′

:=(s[xρ.[aρ]x := [bσ]t]

)′] =(t[xρ := s[xρ.[aρ]x := [bσ]t]

])′.

To sum up, the left-hand side is in relation →∗ to([bσ]t

[xρ := s[xρ.[aρ]x := [bσ]t

])′

which is the right-hand side by the definition of “special” substitution. ut

Theorem 1. −′ is an embedding of λµ into F].

Proof. We only check that µ-reduction steps give rise to at least one rewritestep of F with stabilization. Since we already convinced ourselves that there areno problems with types, we will neglect the types altogether. As an additionalbenefit, we may treat both µ-reductions uniformly (and will later profit fromthis uniformity in the extension by fixed-point types). Write R for a term or atype. Define the term or type R as follows: r := r′ and ρ := ρ∗. The µ-reductionrules now both become:

(µa.[c]r)R � µb.([c]r)[x.[a]x := [b](xR)].

Moreover, we uniformly have (rR)′ = r′E(λz.zR)S. Therefore,

((µa.[c]r)R)′ = (S(λa.cr′))E(λz.zR)S

and one application of (]S) leads to

S(λb.(λa.cr′)

(λx.b(xE(λz.zR)S)

))= S

(λb.(λa.([c]r)′)(λx.b(xR)′)

).

Parigot’s Second Order λµ-Calculus and Inductive Types 337

One beta reduction step yields

S(λb.([c]r)′[a := λx.b(xR)′]

)→∗ S

(λb.

(([c]r)[x.[a]x := [b](xR)]

)′)

by the previous lemma. This is(µb.([c]r)[x.[a]x := [b](xR)]

)′. ut

Hence, second order λµ-calculus has been proven to be strongly normalizing oncemore since also F] has been embedded into the strongly normalizing system F⊥.

6 Extension to Fixed-Point Types

We extend each system by non-interleaving positive fixed-point types: systemF, system F] (with iteration on stabilization) and second order λµ-calculus. Ouraim is to extend the embedding of λµ via F] into F⊥ to the variants λµf , F]f

and F⊥f with those fixed-point types.System F with non-interleaving positive fixed-point types (in the sequel called

Ff ) essentially has been studied in [4] under the name Fret, and its strong nor-malization shown by an embedding into Mendler’s system [10]. A direct proofof strong normalization by saturated sets has been given in [9] under the nameNPF. There is strong evidence that no embedding into system F exists [17].

We now extend system F by types fαρ which are supposed to describe ar-bitrary fixed-points of λαρ, i. e., of the operation σ 7→ ρ[α := σ]. We confineourselves to (non-strictly) positive dependencies which moreover have to be non-interleaved, i. e., fαρ may only be formed when every occurrence of α in ρ is“to the left of an even number of →” and not free in some subexpression fβσof fαρ. The last clause may be rephrased as follows: If fixed-point types fβσare formed with a free parameter α then the formation of a fixed-point typefαρ—hence w. r. t. that parameter α—is forbidden.4 More formally:

Definition 2. We inductively define the set Tnpf of non-interleaving positivefixed-point types and simultaneously for every ρ ∈ Tnpf the sets N+(ρ) and N−(ρ)of type variables which occur only positively or occur only negatively, respectively,and moreover do not occur in the scope of a fixed-point type formation (the setFV(ρ) of free type variables is defined as before with the additional FV(fαρ) :=FV(ρ)\{α}). Let always range p over the set {+,−} of polarities and set −+ :=− and −− := +. Let TV be the set of type variables.

α ∈ Tnpf . N+(α) := TV. N−(α) := TV \ {α}.If ρ, σ ∈ Tnpf then ρ → σ ∈ Tnpf and Np(ρ → σ) := N−p(ρ) ∩ Np(σ).If ρ ∈ Tnpf then ∀αρ ∈ Tnpf and Np(∀αρ) := Np(ρ) ∪ {α}.If ρ ∈ Tnpf and α ∈ N+(ρ) (the only place where the Np(ρ) enter the condi-tions) then fαρ ∈ Tnpf and Np(fαρ) := TV \ FV(fαρ).

4 Note that otherwise there would be a very high degree of freedom in the interpreta-tion of fαρ since fβσ is intended only to model an arbitrary fixed-point.

338 R. Matthes

Note the change of the polarity in the second rule which substantiates the sloganthat α’s occurrences may only be to the left of an even number of →. In the lastrule we achieve non-interleavedness by removing any free variable of fαρ.

It is somewhat awkward to prove that Tnpf is closed under substitution [9,p. 303]. Since it nevertheless holds, we may extend system F to Tnpf and moreoveradd the following term formation rules: If t : ρ[α := fαρ] then Cfαρt : fαρ. Ifr : fαρ then rEf : ρ[α := fαρ].

t : ρ[α := fαρ]Cfαρt : fαρ

r : fαρ

rEf : ρ[α := fαρ]

Beta reduction for fixed-point types will extend � of F by

(βf ) (Cfαρt)Ef � t .

This constitutes Ff . It has been mentioned above that the term closure of � isstrongly normalizing, i. e., Ff is strongly normalizing. Let F⊥f be its extensionby the type constant ⊥. As before, strong normalization is inherited.

Likewise extend system F] by non-interleaving positive fixed-point types toyield system F]f : We also write Tnpf for its set of types which additionally hasthe rules:

⊥ ∈ Tnpf . Np(⊥) := TV.If ρ ∈ Tnpf then ]ρ ∈ Tnpf and Np(]ρ) := Np(ρ).

We may now add the same term formation rules and (βf ) as above. It is clearthat the embedding of F] into F⊥ immediately extends to an embedding of F]f

into F⊥f : one only has to add homomorphic clauses for the new type former andthe new term formation rules. (Note that (ρ[α := fαρ])′ = ρ′[α := (fαρ)′] =ρ′[α := fαρ′] indicates that the new type former does not pose any problemwith the embedding.)

λµ may as well be extended to the types Tnpf (the original definition). Weagain add the two term formation rules and (βf ) but also a µ-rule pertaining tothe fixed-point types:

(µf ) (µafαρ.r)Ef � µbρ[α:=fαρ].r[xfαρ.[afαρ]x := [b](xEf )

]

with r a named term. It strictly follows the pattern given in the proof of The-orem 1 if we also consider Ef as a possible value of R. Hence, the µ-reductionrules are still uniformly described by (the untyped pattern)

(µa.[c]r)R � µb.([c]r)[x.[a]x := [b](xR)].

The resulting system shall be denoted by λµf . Let us extend the embedding ofsection 5. Define the type ρ∗ of F]f by iteration on the type ρ of λµf as follows(and simultaneously prove that Np(ρ∗) = Np(ρ) and FV(ρ∗) = FV(ρ)):

α∗ := α(ρ → σ)∗ := ]ρ∗ → ]σ∗

(∀αρ)∗ := ∀α ]ρ∗

(fαρ)∗ := fα ]ρ∗

Parigot’s Second Order λµ-Calculus and Inductive Types 339

The last clause is legal since α ∈ N+(ρ) = N+(ρ∗) = N+(]ρ∗) by the simultane-ously proved statement.

By induction on ρ one again verifies that (ρ[α := σ])∗ = ρ∗[α := σ∗].Again set ρ′ := ]ρ∗ which implies that (fαρ)′ = ](fαρ′). Also we still have

(ρ[α := σ])′ = ρ′[α := σ∗] and FV(ρ′) = FV(ρ).The previous translation −′ of the terms is extended by clauses for the new

term formation rules. If we set Ef := Ef , then the crucial rule follows the usual(untyped) pattern:

(rR)′ := r′E(λz.zR)S

The new rules are

(Cfαρtρ[α:=fαρ])′ := I(fαρ)′(Cfαρ′t′)(rfαρEf )′ := r′E(ρ[α:=fαρ])′(λzfαρ′

.zEf )S(ρ[α:=fαρ])′

Notice that the first term is well-typed since t′ : (ρ[α := fαρ])′ and

(ρ[α := fαρ])′ = ρ′[α := (fαρ)∗] = ρ′[α := fαρ′].

The crucial equation (ρ[α := fαρ])′ = ρ′[α := fαρ′] also justifies the seconddefinition.

Theorem 2. −′ is an embedding of λµf into F]f .

Proof. Lemma 2 and Lemma 3 clearly still hold, ((Cfαρt)Ef )′ →2 (Cfαρ′t′)Ef →t′, and the treatment of (µf ) is already captured by the uniform proof of Theo-rem 1. ut

Corollary 2. The system λµf of second order λµ-calculus with non-interleavingpositive fixed-point types is strongly normalizing.

Remark 2. There seems to be a widespread belief that in some sense µ-reductionsare nothing but an exploitation of the fact that stability of a type may be provedfrom the stability of its atoms. This immediately works for first-order λµ-calculus(only function types) since the stability of ρ → σ is derivable from stability ofσ. Since also the stability of ∀αρ is derivable from that of ρ, one may expectthat the universal quantifier also is well-behaved. But notice that the set ofatoms (type variables) varies with quantification. Nevertheless, it is possible togive an embedding on grounds of this view [6]. As remarked in the introductionto section 5, the crucial clause is (∀αρ)′ := ∀α.(¬¬α → α) → ρ′, hence arelativization which neatly solves the problem.

What is the problem with fixed-points? It is again possible to derive thestability of fαρ from that of ρ[α := fαρ]:

λu¬¬ρ[α:=fαρ]→ρ[α:=fαρ]λx¬¬fαρ.Cfαρ(u(λy¬ρ[α:=fαρ].x(λzfαρ.y(zEf )))).

But the latter type is usually more complex than the former! Therefore, I donot see a way to extend the idea of reducing the proof of stability to that of thefree type variables. How could one define stability proofs for ρ′ by iteration onρ with some easy translation −′? Our embedding shown above does not at allcare about such a definition since every ρ′ is stable by construction.

340 R. Matthes

Remark 3. One could ask for other type constructions where stability is not in-herited from the constituent types. It is well-known that sum types (disjunction)provide an example of this phenomenon. Unfortunately, the embedding into F]

does not extend to sum types with permutative conversions but only withoutthem. A solution could be to introduce permutative conversions for the stabi-lization types. At present, we could as well take the impredicative encoding ofsums and also get the respective µ-reduction for free (inside λµ). This idea willbe demonstrated in the next section for the more interesting case of monotoneinductive types with primitive recursion.

7 Second Order λµ-Calculus with Primitive Recursionon Monotone Inductive Types Is Strongly Normalizing

Inductive types are a syntactic representation of least pre-fixed-points of op-erations σ 7→ ρ[α := σ]. Typically, they are studied as long as α only oc-curs positively in ρ. (Often, even interleaving is ruled out.) Nevertheless, itturned out [15, sect. 6.3 in chap. 2] that the only needed ingredient for a use-ful notion of inductive type is a monotonicity witness, i. e., a term of type∀α∀β.(α → β) → ρ → ρ[α := β] (in [15] monotone specifications are consid-ered instead). If those terms are not given beforehand but incorporated into theterm system they even do not need to be closed in order to guarantee strongnormalization of the rewrite rules associated with them [7]. There is a choicewhether the monotonicity witnesses are attached to the introduction rule or tothe elimination rule which is primitive recursion. In both cases one has strongnormalization [7], in the second case this even has been shown by an embeddinginto system Ff [8]. However, for practical purposes the first variant seems moreadequate. Fortunately, it also embeds into system Ff which will be the key tothis section’s result.

Recall that the product type ρ×σ may be impredicatively encoded in systemF by ρ × σ := ∀α.(ρ → σ → α) → α for α /∈ FV(ρ) ∪ FV(σ). If r : ρ and s : σthen 〈r, s〉 := Λαλzρ→σ→α.zrs : ρ × σ.

Second order λµ-calculus with primitive recursion on monotone inductivetypes (in the sequel denoted by λµµ) is defined as an extension of second orderλµ-calculus by arbitrary types µαρ (µ binds α in ρ) and by the following termformation rules:If m : ∀α∀β.(α → β) → ρ → ρ[α := β] and t : ρ[α := µαρ] then Cµαρmt : µαρ.If r : µαρ and s : ρ[α := µαρ × σ] → σ then rEµσs : σ.

m : ∀α∀β.(α → β) → ρ → ρ[α := β] t : ρ[α := µαρ]Cµαρmt : µαρ

r : µαρ s : ρ[α := µαρ × σ] → σ

rEµσs : σ

The associated beta reduction rule of primitive recursion is

(βµ) (Cµαρmt)Eµσs � s(m(µαρ)(µαρ × σ)

(λxµαρ.〈x, (λxµαρ.xEµσs)x〉

)t)

Parigot’s Second Order λµ-Calculus and Inductive Types 341

The µ-reduction rule (µ in the sense of Parigot) follows the standard pattern ifwe now even allow R to be Eµσs:

(µµ) (µaµαρ.r)Eµσs � µbσ.r[xµαρ.[aµαρ]x := [b](xEµσs)

]

Again, r denotes a named term in this rule. Note that σ is not at all related toµαρ.

Theorem 3. There is an embedding of λµµ into λµf .

Proof. ρ′ ∈ Tnpf is defined by iteration on ρ. The only non-homomorphic clauseis that for µαρ:

(µαρ)′ := fβ∀γ.((

∀α.(β × γ → α) → ρ′)

→ γ)

→ γ .

We assume that β, γ /∈ {α} ∪ FV(ρ). In fact, the only occurrence of β is 6 timesto the left of → (do not forget that the coding of β × γ provides 2 of them). Itis easy to see that (ρ[α := σ])′ = ρ′[α := σ′] and FV(ρ′) = FV(ρ).

The translation of the terms is also defined homomorphically, with the ex-ception of the two clauses for monotone inductive types:

(Cµαρmt)′ := C(µαρ)′

(Λγλz(∀α.((µαρ)′×γ→α)→ρ′)→γ .z

(Λαλu(µαρ)′×γ→α.

m′(µαρ)′α(λx(µαρ)′

.u〈x, (λx(µαρ)′.xEfγz)x〉

)t′))

and

(rEµσs)′ := r′Efσ′(λz∀α.((µαρ)′×σ′→α)→ρ′

.s′(z((µαρ)′ × σ′)(λx(µαρ)′×σ′

x)))

Since (r[xρ := s])′ = r′[xρ′:= s′] and (r[α := ρ])′ = r′[α := ρ′], also this

translation respects (β→) and (β∀). It is an interesting exercise to prove thatalso (βµ) is simulated.

Since the other term formation rules including reductio ad absurdum aretranslated homomorphically, we may set ([bρ]r)′ := [bρ′

]r′ and hence have atranslation also for named terms. Consequently, (µaσ.r)′ = µaσ′

.r′ with r anamed term. In order to treat the µ-reduction rules we need (r[x.[a]x := [b]t])′ =r′[x.[a]x := [b]t′] for named terms r, but this is proved for named terms and forordinary terms r simultaneously by induction on the size of r.

By induction on natural numbers n, one easily gets (only with µ-reductions)that

(µa.r)R1 . . . Rn →n µb.r[x.[a]x := [b](xR1 . . . Rn)]

(where R1, . . . , Rn are terms or types or objects of the form Eµσs). Of course,this is compatible with the typing requirements.

Therefore, also (µµ) is simulated:((µa.r)Eµσs))′ = (µa.r′)Efσ′(λz....s′ . . .) →3

µb.r′[x.[a]x := [b](xEfσ′(λz....s′ . . .))] = µb.r′[x.[a]x := [b](xEµσs)′] =µb.(r[x.[a]x := [b](xEµσs)])′ = (µb.r[x.[a]x := [b](xEµσs)])′.The other µ-reductions are treated slightly easier. ut

342 R. Matthes

Remark 4. There is no hope for an embedding in the style of the previous sectionwith fixed-point types replaced by inductive types, i. e., for a direct embeddingof λµµ into F], extended by primitive recursion on monotone inductive types.Firstly, the function spaces are overly used for the purpose of defining the typingfor rEµσs. Secondly, if we replaced it by

r : µαρ s0 : σ

rEµσ(xρ[α:=µαρ×σ].s0) : σ

where xρ[α:=µαρ×σ] is bound in s0, we cannot use (ρ[α := µαρ × σ])′ since ×is encoded. Assume it were explicitly included into the system. Then we wouldget (ρ[α := µαρ × σ])′ = ρ′[α := (µαρ × σ)∗] = ρ′[α := (µαρ)′ × σ′]. But wewould need µαρ′ instead of (µαρ)′. And this would have to be lifted with amonotonicity witness for ρ′ w. r. t. α. But we do not even have a monotonicitywitness for µαρ at hand since those only come with the introduction rule forµαρ. Consequently, we first have to get rid of the inductive types (in favour offixed-point types) before we can attack reductio ad absurdum.

Corollary 3. Second order λµ-calculus with primitive recursion on monotoneinductive types is strongly normalizing.

8 Conclusions and Future Work

An alternative to the Kolmogorov translation of classical logic into minimal logichas been presented which simplifies proofs of normalization for a classical versionof λ-calculus (Parigot’s λµ). The translation using stabilization types properlysimulates Parigot’s µ-reductions and carries over to extensions of system F. Thelogical reading of the main theorem gives consistency for classical second-orderpropositional logic with “extended induction” on monotone inductive proposi-tions (where extended induction is given by reading the typing rule for Eµ as aninference rule of natural deduction).

On the computational side, the result gives an application of iteration onnon-strictly positive inductive types. However, the exact nature of the compu-tation involved in this translation should be further studied. Does it exemplifya programming style comparable to continuation-passing style? Moreover, doesthe method help in understanding other λ-calculi for classical logic such as sym-metric λ-calculus?

References

1. Henk P. Barendregt. The Lambda Calculus: Its Syntax and Semantics. North–Holland, Amsterdam, second revised edition, 1984.

2. Pierre-Louis Curien and Hugo Herbelin. The duality of computation. In Proceed-ings of the fifth ACM SIGPLAN International Conference on Functional Program-ming (ICFP ’00), Montreal, pages 233–243. ACM Press, 2000.

Parigot’s Second Order λµ-Calculus and Inductive Types 343

3. Philippe de Groote. A CPS-translation of the λµ-calculus. In Sophie Tison, editor,Trees in Algebra and Programming - CAAP’94, 19th International Colloquium,volume 787 of Lecture Notes in Computer Science, pages 85–99, Edinburgh, 1994.Springer Verlag.

4. Herman Geuvers. Inductive and coinductive types with iteration and recursion.In Bengt Nordstrom, Kent Pettersson, and Gordon Plotkin, editors, Proceedingsof the 1992 Workshop on Types for Proofs and Programs, Bastad, Sweden, June1992, pages 193–217, 1992. Only published electronically:ftp://ftp.cs.chalmers.se/pub/cs-reports/baastad.92/proc.dvi.Z

5. Jean-Yves Girard, Yves Lafont, and Paul Taylor. Proofs and Types, volume 7 ofCambridge Tracts in Theoretical Computer Science. Cambridge University Press,1989.

6. Thierry Joly. Un plongement de la logique classique du 2nd ordre dans AF2.Unpublished manuscript. In French, 5 pp., January 1996.

7. Ralph Matthes. Extensions of System F by Iteration and Primitive Recursion onMonotone Inductive Types. Doktorarbeit (PhD thesis), University of Munich, 1998.Available via http://www.tcs.informatik.uni-muenchen.de/ matthes/.

8. Ralph Matthes. Monotone (co)inductive types and positive fixed-point types. The-oretical Informatics and Applications, 33(4/5):309–328, 1999.

9. Ralph Matthes. Monotone fixed-point types and strong normalization. In GeorgGottlob, Etienne Grandjean, and Katrin Seyr, editors, Computer Science Logic,12th International Workshop, Brno, Czech Republic, August 24–28, 1998, Proceed-ings, volume 1584 of Lecture Notes in Computer Science, pages 298–312. SpringerVerlag, 1999.

10. Nax P. Mendler. Recursive types and type constraints in second-order lambda cal-culus. In Proceedings of the Second Annual IEEE Symposium on Logic in ComputerScience, Ithaca, N.Y., pages 30–36. IEEE Computer Society Press, 1987. Forms apart of [11].

11. Paul F. Mendler. Inductive definition in type theory. Technical Report 87-870,Cornell University, Ithaca, N.Y., September 1987. Ph.D. Thesis (Paul F. Mendler= Nax P. Mendler).

12. Michel Parigot. λµ-calculus: an algorithmic interpretation of classical natural de-duction. In Andrei Voronkov, editor, Logic Programming and Automated Rea-soning, International Conference LPAR’92, St. Petersburg, Russia, volume 624 ofLecture Notes in Computer Science, pages 190–201. Springer Verlag, 1992.

13. Michel Parigot. Strong normalization for second order classical natural deduction.In Proceedings, Eighth Annual IEEE Symposium on Logic in Computer Science,pages 39–46, Montreal, Canada, 1993. IEEE Computer Society Press.

14. Michel Parigot. Proofs of strong normalisation for second order classical naturaldeduction. The Journal of Symbolic Logic, 62(4):1461–1479, 1997.

15. Christine Paulin-Mohring. Definitions Inductives en Theorie des Types d’OrdreSuperieur. Habilitation a diriger les recherches, ENS Lyon, 1996.

16. Dag Prawitz. Natural Deduction. A Proof-Theoretical Study. Almquist and Wiksell,1965.

17. Zdzis law Sp lawski and Pawe l Urzyczyn. Type Fixpoints: Iteration vs. Recursion.SIGPLAN Notices, 34(9):102–113, 1999. Proceedings of the 1999 InternationalConference on Functional Programming (ICFP), Paris, France.

18. Thomas Streicher and Bernhard Reus. Classical logic, continuation semantics andabstract machines. Journal of Functional Programming, 8(6):543–572, 1998.