22.01.2014 Views

Dependently-Typed Programming with Scott Encoding

Dependently-Typed Programming with Scott Encoding

Dependently-Typed Programming with Scott Encoding

SHOW MORE
SHOW LESS

Create successful ePaper yourself

Turn your PDF publications into a flip-book with our unique Google optimized e-Paper software.

<strong>Dependently</strong>-<strong>Typed</strong> <strong>Programming</strong> <strong>with</strong> <strong>Scott</strong> <strong>Encoding</strong><br />

Peng Fu, Aaron Stump<br />

Computer Science, The University of Iowa<br />

Abstract<br />

We introduce Selfstar, a Curry-style dependent type system featuring<br />

self type ιx.t, together <strong>with</strong> mutually recursive definitions<br />

and ∗ : ∗. We show how to obtain <strong>Scott</strong>-encoded datatypes and the<br />

corresponding elimination schemes <strong>with</strong> Selfstar. Examples such<br />

as numerals, vector are given to demonstrate the power of Selfstar<br />

as a dependently-typed programming language. Standard metatheorems<br />

such as type preservation are proved.<br />

1. Introduction<br />

Self type is originated from our previous work on S [5], since then<br />

self type has been studied in combination <strong>with</strong> different typing principles.<br />

In this paper, we study a system called Selfstar, which combines<br />

self type together <strong>with</strong> ∗ : ∗ and mutually recursive definitions.<br />

In Selfstar, every type is inhabited, so Selfstar is inconsistent<br />

as a logic. The only logical feature in Selfstar is the Leibniz<br />

convertability, i.e. we define t 1 = A t 2 to be ΠC : A → ∗.Ct 1 →<br />

Ct 2. Note that we use “convertability” instead of “equality” to indicate<br />

one can not interpret t 1 = A t 2 as a formula. If we know<br />

the inhabitant of t 1 = A t 2 is normalized at the term λC.λx.x,<br />

then we can use t 1 = A t 2 to cast the type P t 1 to P t 2 by applying<br />

the term (λC.λx.x)P to the inhabitant of P t 1. Note that<br />

(λC.λx.x)P → β λx.x, so the casting will not affect the inhabitant<br />

of P t 1.<br />

<strong>Scott</strong> encoding (reported in [4]) does not suffer from the ineffeciency<br />

problem arised in Church encoding. For functional programming<br />

langauge, <strong>Scott</strong> encoding seems to be a better fit than<br />

Church encoding [9]. From the typing perspective, each <strong>Scott</strong>encoded<br />

data contains its subdata, one would need recursive definition<br />

in order to define a type for <strong>Scott</strong> encoded data. Elimination<br />

schemes for the <strong>Scott</strong>-encoded data are derivable in Selfstar, this<br />

means programmer can write down programs that have types like<br />

Πx : Nat.add x 0 = Nat x, which increases the flexibility of typelevel<br />

casting.<br />

The main contributions of this paper are:<br />

• We present Selfstar, which allows us to type <strong>Scott</strong>-encoded data<br />

and derive elimination schemes for <strong>Scott</strong>-encoded data. Selfstar<br />

simplifies the design of the functional programming language,<br />

since the primitive notion of inductive data and pattern matching<br />

is not needed in Selfstar.<br />

Permission to make digital or hard copies of all or part of this work for personal or<br />

classroom use is granted <strong>with</strong>out fee provided that copies are not made or distributed<br />

for profit or commercial advantage and that copies bear this notice and the full citation<br />

on the first page. Copyrights for components of this work owned by others than ACM<br />

must be honored. Abstracting <strong>with</strong> credit is permitted. To copy otherwise, or republish,<br />

to post on servers or to redistribute to lists, requires prior specific permission and/or a<br />

fee. Request permissions from Permissions@acm.org.<br />

Copyright c○ ACM [to be supplied]. . . $15.00<br />

• We prove type preservation and progress for Selfstar by applying<br />

the method we developed in the study of System S.<br />

1.1 Motivation<br />

In Curry style system T [6] equipped <strong>with</strong> polymorphic and dependent<br />

type, one has a primitive notion of recursor, namely, rec :<br />

Πx : Nat.∀U.(Nat → U → U) → U → U and two reductions<br />

rules: rec 0 f v → v and rec (Sn) f v → f n (rec n f v).<br />

The recursor can be emulated <strong>with</strong> lambda terms. For example,<br />

rec := λn.λf.λv.n f v, <strong>with</strong> the notion of numeral ¯0 := λs.λz.z<br />

and ¯n := λs.λz.s n − 1 (n − 1 s z). One can verify that the<br />

definition of rec in lambda calculus behaves the same as the one<br />

in system T. With recursive definition, we can define Nat :=<br />

∀U.(Nat → U → U) → U → U. Note that the type of ¯n is<br />

the same as the type of rec ¯n.<br />

So far the type of the recursor is elementary, i.e., not involving<br />

dependency. To make real use of dependent type, we ask if<br />

it is possible to obtain a type likes Πx : Nat.∀U : Nat →<br />

∗.(Πy : Nat.U y → U (Sy)) → U ¯0 → U x. Note that<br />

S := λn.λs.λz.s n (n s z). We want to emphasis the underlying<br />

computational behavior of this type should be the same as rec, thus<br />

we want a typing relation rec : ∀U : Nat → ∗.Πx : Nat.(Πy :<br />

Nat.U y → U (Sy)) → U ¯0 → U x. We also want the type of<br />

rec ¯n, namely (Πy : Nat.U y → U (Sy)) → U ¯0 → U ¯n to be<br />

the same as the type of ¯n. So we want the following typing relation:<br />

¯n : (Πy : Nat.U y → U (Sy)) → U ¯0 → U ¯n for any ¯n. We<br />

know the following self type mechanism:<br />

Γ ⊢ t : [t/x]T<br />

Γ ⊢ t : ιx.T selfGen Γ ⊢ t : ιx.T<br />

Γ ⊢ t : [t/x]T selfInst<br />

So it is not surprising we define Nat := ιx.(Πy : Nat.U y →<br />

U (Sy)) → U ¯0 → U x. With selfGen, selfInst and mutually<br />

recursive definition, one can verify that indeed, the type of ¯n is the<br />

same as the type of rec ¯n and the type of rec is indeed ∀U : Nat →<br />

∗.Πx : Nat.(Πy : Nat.U y → U (Sy)) → U ¯0 → U x. It is<br />

tempted to claim that rec represents the induction principle, but it<br />

is not for the following two reasons: 1. With mutually recursive<br />

definitions, the types can not be interpreted as formulas. 2. The<br />

dependent product Π is not exactly the first order quantifier ∀.<br />

System T is closed to the functional programming language, but<br />

still a little far from a functional programmer’s usual experience.<br />

In mordern functional programming language, one would want to<br />

write a plus 2 function in the following style:<br />

data Nat = Zero<br />

| Succ Nat<br />

plusTwo :: Nat -> Nat<br />

plusTwo n = case n of<br />

Succ p -> Succ ( plusTwo p)<br />

| Zero -> Succ ( Succ Zero )


With <strong>Scott</strong> numerals, we can achieve the effects above. Assume<br />

<strong>Scott</strong> numerals and mutually recursive definition. We define:<br />

Zero = lam s . lam z . z<br />

Succ n = lam s . lam z . s n<br />

plusTwo n = case n<br />

lam p. Succ ( plusTwo p)<br />

Succ ( Succ Zero )<br />

Of course, case := λn.λf.λa.n f a and lam denotes the usual<br />

λ. One can see the differences between the two programs above<br />

are mostly superficial. Now let us first give a elementary version of<br />

type for case, which is (Nat → U) → U → U. The dependency<br />

version is Πx : Nat.(Πy : Nat.U (Sy)) → U ¯0 → U x. And we<br />

define Nat := ιx.(Πy : Nat.U (Sy)) → U ¯0 → U x, so one can<br />

again check that case ¯n and ¯n have the same type for <strong>Scott</strong> numeral<br />

¯n using the selfInst rule.<br />

Observe that the term for rec and case (even the iterator in<br />

system F) is the term λn.λf.λa.n f a. We call the type of this<br />

term elimination scheme. By elemination scheme is derivable in<br />

Selfstar, we mean the elimination scheme is inhabited by the term<br />

λn.λf.λa.n f a in Selfstar (modulo type annotations).<br />

1.2 Overview<br />

In section 2, we present Selfstar. We show how to type <strong>Scott</strong>encoded<br />

data and its derivative. The corresponding elimination<br />

schemes for <strong>Scott</strong>-encoded data are derived. We also provide several<br />

examples (numerals and vectors) to demonstrate the power of<br />

Selfstar. In section 3, metatheorems such as type preservation are<br />

proved for Selfstar by applying the method we develped in S [5].<br />

2. <strong>Dependently</strong>-typed programming <strong>with</strong> Selfstar<br />

Selfstar use the self type mechanism to obtain inductive data, resulting<br />

a design that is simpler than most dependently-typed core<br />

languages. Intuitively, it is hard imagine how to emulate inductive<br />

datatype and pattern matching <strong>with</strong>out any build-in mechanisms.<br />

But as we observe in section 1.1, <strong>Scott</strong> encoding together <strong>with</strong> mutually<br />

recursive definitions are enough to perform pattern matching<br />

on inductive data. The real difficulty lies in the typing. We want<br />

to make sure that both <strong>Scott</strong>-encoded data and definable operations<br />

on these data are typable in Selfstar. The self type allows<br />

us to type <strong>Scott</strong> data and to derive the corresponding elimination<br />

schemes. Thus operations on <strong>Scott</strong> data is typable using the elimination<br />

scheme.<br />

2.1 System Selfstar<br />

We give the full specification of Selfstar in this section. We use<br />

gray boxes to hightlight certain important terms and rules.<br />

Definition 1 (Syntax).<br />

Terms t ::= ∗ | x | λx.t | tt ′ | µt | Πx : t 1.t 2 | ιx.t<br />

Closure µ ::= {x i ↦→ t i} i∈N<br />

Value v ::= ∗ | λx.t | Πx : t 1.t 2 | ιx.t | ⃗µ(Πx : t 1.t 2) | ⃗µ(ιx.t)<br />

Context Γ ::= · | Γ, x : t | Γ, ˜µ<br />

Remarks :<br />

• If µ is {x i ↦→ t i} i∈N , then ˜µ is {(x i : a i) ↦→ t i} i∈N for some<br />

term a i.<br />

• For {x i ↦→ t i} i∈N , we require for any 1 ≤ i ≤ n, the free<br />

variable set FV(t i) ⊆ dom(µ) = {x 1, ..., x n}. We also do not<br />

allow any reductions and substitutions inside the closure. We<br />

call this the locality restriction. Without locality requirement, it<br />

is hard to establish confluence for reductions(see [1]).<br />

• FV(µt) = FV(t) − dom(µ).<br />

· ⊢ wf<br />

Γ ⊢ wf Γ ⊢ t : ∗<br />

Γ, x : t ⊢ wf<br />

Γ ⊢ wf<br />

{Γ, ˜µ ⊢ t j : a j} (tj :a j )∈˜µ<br />

Γ, ˜µ ⊢ wf<br />

Figure 1. Well-formed Context Γ ⊢ wf<br />

Γ ⊢ ∗ : ∗ Star<br />

Γ, x : ιx.t ⊢ t : ∗<br />

Γ ⊢ ιx.t : ∗<br />

Self<br />

(x : t) ∈ Γ<br />

Γ ⊢ x : t<br />

Var<br />

Γ ⊢ t : ιx.t ′<br />

Γ ⊢ t : [t/x]t ′<br />

Γ, x : t 1 ⊢ t 2 : ∗ Γ ⊢ t 1 : ∗<br />

Γ ⊢ Πx : t 1.t 2 : ∗<br />

Γ ⊢ t : [t/x]t ′<br />

Γ ⊢ t : ιx.t ′<br />

Γ ⊢ ιx.t ′ : ∗<br />

Γ ⊢ t : t 1 Γ ⊢ t 1<br />

∼ = t2 Γ ⊢ t 2 : ∗<br />

Γ ⊢ t : t 2<br />

Γ, x : t 1 ⊢ t : t 2 Γ ⊢ t 1 : ∗<br />

Γ ⊢ λx.t : Πx : t 1.t 2<br />

Γ ⊢ t : Πx : t 1.t 2 Γ ⊢ t ′ : t 1<br />

Γ ⊢ tt ′ : [t ′ /x]t 2<br />

Pi<br />

SelfGen<br />

Lam<br />

App<br />

Γ, ˜µ ⊢ t : t ′ {Γ, ˜µ ⊢ t j : a j} (tj :a j )∈˜µ<br />

Γ ⊢ ⃗µ∗ ∗<br />

(x i ↦→ t i) ∈ µ ∈ ⃗µ<br />

Γ ⊢ ⃗µx i ⃗µt i<br />

Γ ⊢ µt : µt ′<br />

Figure 2. Typing Γ ⊢ t : t ′<br />

Γ ⊢ ⃗µ(tt ′ ) (⃗µt)(⃗µt ′ )<br />

Γ ⊢ (λx.t)v [v/x]t<br />

Γ ⊢ t ′ t ′′<br />

Γ ⊢ (λx.t)t ′ (λx.t)t ′′<br />

(x i ↦→ t i) ∈ Γ<br />

Γ ⊢ x i t i<br />

x /∈ dom(⃗µ)<br />

Γ ⊢ ⃗µx x<br />

SelfInst<br />

Conv<br />

Mu<br />

Γ ⊢ ⃗µ(λx.t) λx.(⃗µt)<br />

Γ ⊢ t t ′′<br />

Γ ⊢ tt ′ t ′′ t ′<br />

Figure 3. Executions


Γ ⊢ t 1 ∗ t 2<br />

Γ ⊢ t 1 = t 2<br />

Γ ⊢ (λx.t)t ′ = [t ′ /x]t<br />

FV(t)#dom(µ)<br />

Γ ⊢ µt = t Γ ⊢ µ(ιx.t) = ιx.(µt ′ )<br />

Γ ⊢ µ(Πx : t 1.t 2) = Πx : µt 1.µt 2<br />

FV(t)#dom(µ)<br />

Γ ⊢ µt = t<br />

Γ ⊢ t = t ′′ Γ ⊢ t ′ = t ′′′<br />

Γ ⊢ tt ′ = t ′′ t ′′′ Γ, ˜µ ⊢ t = t ′<br />

Γ ⊢ µt = µt ′<br />

Γ ⊢ t = t ′<br />

Γ ⊢ t = t ′<br />

Γ ⊢ ιx.t = ιx.t ′ Γ ⊢ λx.t = λx.t ′<br />

Γ ⊢ t 2 = t 3 Γ ⊢ t 1 = t 2<br />

Γ ⊢ t 1 = t 3<br />

Γ ⊢ t 2 = t 1<br />

Γ ⊢ t 1 = t 2<br />

Remarks:<br />

µ ∈ Γ<br />

Γ ⊢ µt → o t<br />

Figure 4. Equality<br />

Γ ⊢ t → o t ′<br />

Γ ⊢ λx.t → o λx.t ′<br />

Γ ⊢ t → o t ′′<br />

Γ ⊢ t ′ → o t ′′<br />

Γ ⊢ tt ′ → o t ′′ t ′ Γ ⊢ tt ′ → o tt ′′<br />

Γ ⊢ t 1 → o t ′ 1<br />

Γ ⊢ Πx : t 1.t 2 → o Πx : t ′ 1.t ′ 2<br />

Γ ⊢ t 2 → o t ′ 2<br />

Γ ⊢ Πx : t 1.t 2 → o Πx : t ′ 1.t ′ 2<br />

Γ, ˜µ ⊢ t → o t ′<br />

Γ ⊢ µt → o µt ′<br />

Figure 5. Closure Reductions<br />

Γ ⊢ t → o t ′<br />

Γ ⊢ ιx.t → o ιx.t ′<br />

• (t i : a i) ∈ ˜µ means (x i : a i) ↦→ t i ∈ ˜µ. ⃗µt denotes µ 1...µ nt.<br />

• Typing does not depend on well-formness of the context, so the<br />

self type formation rule self is not circular in this sense. We will<br />

show: if Γ ⊢ wf and Γ ⊢ t : t ′ , then Γ ⊢ t ′ : ∗ (Appendix A).<br />

• We use call-by-value strategy for the execution.<br />

• ∼ = denotes =o ∪ =, where = o denotes the reflexive transitive<br />

and symmetry closure of → o.<br />

• The equality rules incorpates executions to automatize a portion<br />

of equality reasoning.<br />

• At type level, we want to have the ability to open the closure<br />

when it appear in the context. closure reduction allow us to do<br />

this, <strong>with</strong>out this type level reduction, we can not prove type<br />

preservation.<br />

2.2 <strong>Scott</strong> <strong>Encoding</strong>s in Selfstar<br />

Now let us see some concrete examples of <strong>Scott</strong> encodings in<br />

Selfstar. For convenient, we write a → b for Πx : a.b <strong>with</strong><br />

x /∈ FV(b).<br />

Definition 2 (<strong>Scott</strong>’s Derivative). Let ˜µ d be the following recursive<br />

defintions:<br />

(Nat : ∗) ↦→<br />

ιx .ΠC : Nat → ∗.(Πn : Nat.(C n) → (C (S n))) → (C 0) → C x<br />

(S : Nat → Nat) ↦→ λn.λC.λs.λz.s n (n C s z)<br />

(0 : Nat) ↦→ λC.λs.λz.z<br />

<strong>with</strong> s : Πn : Nat.(C n) → (C (S n)), z : C0, n : Nat, we have<br />

˜µ d ⊢ wf (using selfGen and selfInst rules).<br />

Definition 3 (Elimination Scheme for <strong>Scott</strong>’s Derivative).<br />

˜µ d ⊢ Rec : ΠC : Nat → ∗.Πn : Nat.((C n) → C (S n)) →<br />

C 0 → Πn : Nat.C n<br />

Rec := λC.λs.λz.λn.n C s z<br />

<strong>with</strong> s : Πn : Nat.(C n) → (C (S n)), z : C0, n : Nat.<br />

Typing: Let Γ = ˜µ d , C : Nat → ∗, s : Πn : Nat.(C n) →<br />

(C (S n)), z : C0, n : Nat. Since n : Nat, by selfInst, n : ΠC :<br />

(Nat → ∗).(Πy : Nat.(C y) → (C (S y))) → (C 0) → (C n).<br />

Thus n C s z : C n.<br />

Definition 4 (<strong>Scott</strong> Numerals). Let ˜µ s be the following recursive<br />

defintions:<br />

(Nat : ∗) ↦→<br />

ιx .ΠC : Nat → ∗. (Πn : Nat.C (S n)) → (C 0) → (C x)<br />

(S : Nat → Nat) ↦→ λn.λC.λs.λz.s n<br />

(0 : Nat) ↦→ λC.λs.λz.z<br />

<strong>with</strong> s : Πn : Nat.C (S n), z : C0, n : Nat, we have ˜µ s ⊢<br />

wf(using SelfGen and SelfInst rules).<br />

Definition 5 (Elimination Scheme for <strong>Scott</strong> Numerals).<br />

˜µ s ⊢ Case : ΠC : Nat → ∗.Πn : Nat.(C (S n)) → C 0 → Πn :<br />

Nat.C n<br />

Case := λC.λs.λz.λn.n C s z<br />

<strong>with</strong> s : Πn : Nat.(C (S n)), z : C0, n : Nat(using SelfInst rule).<br />

Typing: Let Γ = ˜µ s, C : Nat → ∗, s : Πn : Nat.C (S n), z :<br />

C0, n : Nat. Since n : Nat, by selfInst, n : ΠC : (Nat →<br />

∗).(Πy : Nat.(C (S y)) → (C 0) → (C n). Thus n C s z : C n.<br />

Definition 6 (Addition). We define µ +:<br />

(add : Nat → Nat → Nat) ↦→<br />

λn.λm.Case (λn.Nat) (λp.(S (add p m))) m n<br />

One can check that ˜µ s, µ + ⊢ wf.<br />

Definition 7 (Leibniz Convertability).<br />

Eq := λA.λx.λy.ΠC : A → ∗.C x → C y.<br />

Definition 8. ˜µ s ⊢ addZ : Πx : Nat.Eq Nat (add x 0) x.<br />

Typing: We are trying to show<br />

Πx : Nat.ΠC : Nat → ∗.C (add x 0) → C x<br />

is inhabited. We know that the type of case (λz.(Eq Nat (add z 0) z))<br />

is Πn : Nat.(Eq Nat (add (Sn) 0) Sn) → (Eq Nat (add 0 0) 0) →<br />

Πn : Nat.Eq Nat (add n 0) n.<br />

So case (λz.(Eq Nat (add z 0) z)) p 1 p 2 : Πn : Nat.Eq Nat (add n 0) n,<br />

<strong>with</strong> p 1 : Πn : Nat.(Eq Nat (add (Sn) 0) Sn) and p 2 :<br />

Eq Nat (add 0 0) 0.<br />

It is easy to see that p 2 := λC[: Nat → ∗].λx[: C(add 0 0)].x.<br />

We know addZ n (λq[: Nat].C(Sq)) : C(S(add n 0)) → C(Sn).<br />

Thus p 1 :=<br />

λn[: Nat].λC[: Nat → ∗].λz[: C(add (Sn) 0)].<br />

(addZ n (λq[: Nat].C(Sq))) z.<br />

So we arive at the following definition:<br />

addZ := case (λz.(Eq Nat (add z 0) z))<br />

(λn.λC.λz.(addZ n (λq.C(Sq))) z)<br />

(λC.λx.x) = β<br />

λy.y (λz.(Eq Nat (add z 0) z))<br />

(λn.λC.λz.(addZ n (λq.C(Sq))) z) (λC.λx.x)<br />

Observe that addZ is a recursive function that is equivalent to<br />

λC.λz.z for all input of <strong>Scott</strong> numerals. So it is safe to use addZ ¯n<br />

to convert add ¯n 0 to ¯n.<br />

Definition 9 (Vector).<br />

Let ˜µ v be the following recursive defintions:


(vec(U, n) : ∗) ↦→<br />

ιx .ΠC : Πp : Nat.vec(U, p) → ∗.<br />

(Πm : Nat.Πu : U.Πy : vec(U, m).C (S m) (cons m u y)) →<br />

(C 0 nil) → (C n x)<br />

(cons : Πn : Nat.U → vec(U, n) → vec(U, Sn)) ↦→<br />

λn.λv.λl.λC.λy.λx.y n v l<br />

(nil : vec(U, 0)) ↦→ λC.λy.λx.x<br />

where n : Nat, v : U, l : vec(U, n), C : Πp : Nat.vec(U, p) →<br />

∗, y : Πm : Nat.Πu : U.Πy : vec(U, m).(C (Sm) (cons m u y)), x :<br />

C 0 nil.<br />

Typing: It is easy to see that nil is typable to vec(U, 0). Now<br />

we show how cons is typable to Πn : Nat.U → vec(U, n) →<br />

vec(U, Sn). The type of y n v l is C (Sn) (cons n v l).<br />

So λC.λy.λx.y n v l :<br />

ΠC : (Πp : Nat.vec(U, p) → ∗).<br />

(Πm : Nat.Πu : U.Πy : vec(U, m).(C (Sm) (cons m u y))) →<br />

C 0 nil → C (Sn) (λC.λy.λx.y n v l) .<br />

So by selfGen, we have λC.λy.λx.y n v l : vec(U, Sn). Thus<br />

cons : Πn : Nat.U → vec(U, n) → vec(U, Sn).<br />

Definition 10 (Elimination Scheme for Vector).<br />

˜µ v ⊢ Case(U, n) :<br />

ΠC : (Πp : Nat.vec(U, p) → ∗).<br />

(Πm : Nat.Πu : U.Πy : vec(U, m).(C (Sm) (cons m u y)))<br />

→ C 0 nil → Πx : vec(U, n).(C n x)<br />

where Case(U, n) := λC.λs.λz.λx.x C s z<br />

C : (Πp : Nat.vec(U, p) → ∗), s : Πm : Nat.Πu : U.Πy :<br />

vec(U, m).(C (Sm) (cons m u y)), z : C 0 nil, x : vec(U, n).<br />

Definition 11 (Append).<br />

˜µ v ⊢ app :<br />

Πn 1 : Nat.Πn 2 : Nat.vec(U, n 1) → vec(U, n 2) → vec(U, n 1 +<br />

n 2)<br />

where app := λn 1.λn 2.λl 1.λl 2.<br />

Case(U, n 1)(λz.λq.vec(U, z + n 2))<br />

(λm.λh.λt.cons (m + n 2) h (app m n 2 t l 2))<br />

l 2 l 1<br />

(x ↦→ t) ∈ Γ<br />

Γ ⊢ x → β t<br />

Γ ⊢ (λx.t)t ′ → β [t ′ /x]t<br />

(x i ↦→ t i) ∈ µ<br />

Γ ⊢ t → β t ′<br />

Γ ⊢ µx i → β µt i Γ ⊢ λx.t → β λx.t ′<br />

Γ ⊢ t → β t ′′<br />

Γ ⊢ t ′ → β t ′′<br />

Γ ⊢ tt ′ → β t ′′ t ′ Γ ⊢ tt ′ → β tt ′′<br />

Γ, ˜µ ⊢ t → β t ′<br />

Γ ⊢ µt → β µt ′ Γ ⊢ t → β t ′<br />

Γ ⊢ ιx.t → β ιx.t ′<br />

Γ ⊢ t 2 → β t ′ 2<br />

Γ ⊢ Πx : t 1.t 2 → β Πx : t ′ 1.t ′ 2<br />

dom(µ)#FV(t)<br />

Γ ⊢ µt → µ t<br />

Γ ⊢ µ(t 1t 2) → µ (µt 1)(µt 2)<br />

Figure 6. Beta Reductions<br />

Γ ⊢ t 1 → β t ′ 1<br />

Γ ⊢ Πx : t 1.t 2 → β Πx : t ′ 1.t ′ 2<br />

Γ ⊢ µ(λx.t) → µ λx.µt<br />

Γ ⊢ µ(ιx.t) → µ ιx.µt<br />

Γ ⊢ µ(Πx : t 1.t 2) → µ Πx : µt 1.µt 2<br />

Γ ⊢ t → µ t ′<br />

Γ ⊢ λx.t → µ λx.t ′<br />

Γ ⊢ t ′ → µ t ′′<br />

Γ ⊢ t → µ t ′′<br />

Γ ⊢ tt ′ → µ tt ′′ Γ ⊢ tt ′ → µ t ′′ t ′<br />

Γ ⊢ t 1 → µ t ′ 1<br />

Γ ⊢ Πx : t 1t 2 → µ Πx : t 1.t ′ 2<br />

Γ ⊢ t 2 → µ t ′ 2<br />

Γ ⊢ Πx : t 1t 2 → µ Πx : t 1.t ′ 2<br />

Γ ⊢ t → µ t ′<br />

Γ ⊢ ιx.t → µ ιx.t ′<br />

Γ, ˜µ ⊢ t → µ t ′<br />

Γ ⊢ µt → µ µt ′<br />

Typing: We want to show app : Πn 1 : Nat.Πn 2 : Nat.vec(U, n 1) →<br />

vec(U, n 2) → vec(U, n 1+n 2). We instantiate C := λz.(λq.vec(U, z+<br />

n 2)) , where q free over vec(U, y + n 2), in Case(U, n 1). By<br />

beta reductions, we get Case(U, n 1) (λz.λq.vec(U, z + n 2)) :<br />

Πm : Nat.Πu : U.Πy : vec(U, m).(vec(U, (Sm) + n 2)) →<br />

vec(U, 0 + n 2) → Πx : vec(U, n 1).vec(U, n 1 + n 2).<br />

Also, λm.λh.λt.cons (m + n 2) h (app m n 2 t l 2) : Πm :<br />

Nat.Πh : U.Πt : vec(U, m).(vec(U, (S (m + n 2)))<br />

With l 1 : vec(U, n 1), l 2 : vec(U, n 2), we can see that it is the case.<br />

Definition 12 (Associativity).<br />

˜µ v ⊢ assoc : Π(n 1.n 2.n 3 : Nat).<br />

Π(v 1 : vec(U, n 1).v 2 : vec(U, n 2).v 3 : vec(U, n 3)).<br />

Eq vec(U, n 1+n 2+n 3) (app n 1 (n 2+n 3) v 1 (app n 2 n 3 v 2 v 3))<br />

(app (n 1 + n 2) n 3 (app n 1 n 2 v 1 v 2) v 3)<br />

3. Metatheory<br />

The proof of type preservation for Selfstar follows the same<br />

method for proving type preservation of S. Selfstar is a simpler<br />

system compare to S in the following senses: 1. With ∗ : ∗, we<br />

no longer have separate syntactic categories for types and kinds.<br />

2. Polymorphism is anotated in Selfstar, which leads to an easier<br />

proof of type preservation, namely, we do not need morph analysis<br />

[5] for Selfstar. These simplifications lead to a simpler proof of<br />

type preservation.<br />

In order to prove type preservation for Selfstar, we need confluence<br />

analysis for the type level transformation. We need to show<br />

Figure 7. Mu Reductions<br />

type level transformation is confluent. Thus the transformation<br />

from Πx : t 1.t 2 to Πx : t ′ 1.t ′ 2 implies t 1 can be transformed to t ′ 1<br />

and t 2 can be transformed to t ′ 2. Thus we establish the compatibility<br />

property for Selfstar, which is the major result in order to prove<br />

type preservation. The proofs for this section are in Appendix C, D<br />

and E. Once we prove type preservation, progress theorem is easy<br />

to prove, see Appendix B.<br />

3.1 The Analytical System<br />

It is combersome to directly prove the equality in Selfstar is<br />

Church-Rosser. We develop an analytical system and we prove that<br />

the analytical system is equivalent (theorem 16) to the equality system<br />

in Selfstar. Then we prove the analytical system is confluent,<br />

which implies the Church-Rosser of the equality in Selfstar.<br />

The beta-reductions (figure 6) include definition substitutions<br />

and the ordinary beta-reduction in lambda calculus. The mureductions<br />

(figure 7) are for moving the closure inside the term<br />

structure.<br />

Let → denote → β ∪ → µ. Let ↔ ∗ denote (→ ∪ → −1 ) ∗ . The<br />

following lemmas show the relation between → and =.<br />

Lemma 13. If Γ ⊢ t 1 t 2, then Γ ⊢ t 1 → t 2.<br />

Proof. By induction on derivation of Γ ⊢ t 1 t 2.


Lemma 14. If Γ ⊢ t 1 = t 2, then Γ ⊢ t 1 ↔ ∗ t 2.<br />

Proof. By induction on the derivation of Γ ⊢ t 1 = t 2.<br />

Lemma 15. If Γ ⊢ t 1 → t 2, then Γ ⊢ t 1 = t 2.<br />

Proof. By induction on the derivation of Γ ⊢ t 1 → t 2.<br />

The following theorem shows that the analytic system is equivalent<br />

to the equality system.<br />

Theorem 16. Γ ⊢ t 1 = t 2 iff Γ ⊢ t 1 ↔ ∗ t 2.<br />

Proof. By lemma 14, lemma 15.<br />

Suppose → is confluent. By theorem 16, we know that Γ ⊢<br />

Πx : t 1.t 2 = Πx : t ′ 1.t ′ 2 implies Γ ⊢ Πx : t 1.t 2 ↔ ∗ Πx : t ′ 1.t ′ 2.<br />

The confluence of → implies Church-Rosser of ↔ ∗ , namely, there<br />

exists a t such that Γ ⊢ Πx : t 1.t 2 → ∗ t and Γ ⊢ Πx : t ′ 1.t ′ 2 → ∗ t.<br />

By definition of →, we know t must be of the form Πx : t 3.t 4,<br />

<strong>with</strong> Γ ⊢ t 1 → t 3, Γ ⊢ t ′ 1 → t 3, Γ ⊢ t 2 → t 4 and Γ ⊢ t ′ 2 → t 4.<br />

So by lemma 15, we have Γ ⊢ t 1 = t ′ 1 and Γ ⊢ t 2 = t ′ 2.<br />

Now let us focus on the proof of the confluence of →. The<br />

confluence argument is similar to the one described in [3]. We are<br />

going to use the following lemma to conclude the confluence of<br />

→ β ∪ → µ.<br />

Lemma 17 (Hardin’s interpretation lemma[7]). Let → be → 1<br />

∪ → 2, → 1 being confluent and strongly normalizing. We denote by<br />

ν(a) the → 1-normal form of a. Suppose that there is some relation<br />

→ i on the → 1-normal forms satisfying:<br />

→ i⊆↠, 1 and a → 2 b implies ν(a)↠ iν(b) (†)<br />

Then the confluence of → i implies the confluence of →.<br />

Proof. Suppose → i is confluent. Assume a↠a ′ and a↠a ′′ . So<br />

by (†), ν(a)↠ iν(a ′ ) and ν(a)↠ iν(a ′′ ). Note that t→ ∗ 1t ′ implies<br />

ν(t) = ν(t ′ )(By the confluence and strong normalization of → 1).<br />

By the confluence of → i, there exists a b such that ν(a ′ )↠ ib<br />

and ν(a ′′ )↠ ib. Since → i, → 1⊆↠, we get a ′ ↠ν(a ′ )↠b and<br />

a ′′ ↠ν(a ′′ )↠b. Hence → is confluent.<br />

The idea behinds interpretation method is that it allows us to<br />

modulo the → 1-reduction, we only need to focus on proving the<br />

confluence of → i. This is essential since in our case, → β,µ can<br />

not be directly parallelized, namely, one can not use Tait-Martin<br />

Löf’s method (reported in [2]) directly to prove the confluence<br />

of → β,µ , the paralleled version does not enjoy diamond property.<br />

With the interpretation method, after modulo the → µ-reduction, we<br />

introduce a new reduction → βµ (corresponds to → i), we can then<br />

use the parallel reduction method to prove confluence of → βµ .<br />

Lemma 18. → µ is confluent and terminating.<br />

So → µ correspond to → 1 in the interpretation lemma. Since<br />

→ µ is strongly normalizing and confluent, we can define a normalization<br />

function which effectively computes the mu-normal form.<br />

Definition 19 (µ-Normal Forms).<br />

n ::= ∗ | x | µx i | λx.n | nn ′ | Πx : n.n ′ | ιx.n<br />

Note: for the µx i in definition 19, we assume x i ∈ dom(µ).<br />

1 ↠ is the reflexive symmetric transitive closure of →.<br />

Definition 20 (µ normalization function).<br />

m(∗) := ∗<br />

m(x) := x<br />

m(λy.t) := λy.m(t)<br />

m(t 1t 2) := m(t 1)m(t 2)<br />

m(ιx.t) := ιx.m(t)<br />

m(Πx : t.t ′ ) := Πx : m(t).m(t ′ ).<br />

m(⃗µy) := y if y /∈ dom(⃗µ).<br />

m(⃗µy) := µ iy if y ∈ dom(µ i).<br />

m(⃗µ(tt ′ )) := m(⃗µt)m(⃗µt ′ )<br />

m(⃗µ(λx.t)) := λx.m(⃗µt).<br />

m(⃗µ(ιx.t)) := ιx.m(⃗µt).<br />

m(⃗µ(Πx : t.t ′ )) := Πx : m(⃗µt).m(⃗µt ′ ).<br />

We shall devise a new notion of reduction on mu-normal form,<br />

then show that this reduction is confluent (corresponds to → i in<br />

the interpretation lemma and satisfying the † property), thus by the<br />

interpretation lemma, we can show → β,µ is confluent 2 . A natural<br />

way to define reduction on mu-normal form is that right after a betareduction,<br />

one immediately mu-normalizes the contractum, which<br />

can form a notion of reduction on mu-normal form.<br />

Definition 21 (β Reduction on µ-normal Forms).<br />

Γ ⊢ n → β t<br />

Γ ⊢ n → βµ m(t)<br />

The following lemma shows that → βµ corresponds to the → i<br />

in the interpretation lemma.<br />

Lemma 22. If Γ ⊢ t → β t ′ , then Γ ⊢ m(t) → βµ m(t ′ ).<br />

Lemma 23. → βµ is confluent.<br />

Theorem 24. → β,µ is confluent.<br />

Proof. We know → βµ is confluent. Since → µ is strongly normalizing<br />

and confluent, and by lemma 22 and Hardin’s interpretation<br />

lemma(lemma 17), we conclude → β,µ is confluent.<br />

3.2 Confluence Analysis<br />

Definition 25.<br />

Γ ⊢ t 1 → ι t 2 if t 1 ≡ ιx.t ′ and t 2 ≡ [t/x]t ′ for some fix term t.<br />

Note that → ι models the selfInst rule, → −1<br />

ι models the selfGen<br />

rule. The notion of ι-reduction does not build in structure congruence,<br />

namely, we do not allow reduction rules like: if T → ι T ′ ,<br />

then λx.T → ι λx.T ′ . The purpose of ι-reduction is to emulate the<br />

typing rule selfInst and selfGen. This rewriting point of view on<br />

typing is inspired by Kuan et.al. [10] and Stump et.al. [12].<br />

Lemma 26 (Confluence). → ι is confluent.<br />

Proof. This is obvious since → ι is deterministic.<br />

The goal of this section is to show → o,ι,β,µ is confluent. We<br />

make extensive use of the notion of commutativity, which provides<br />

a simple way to prove the confluence of a reduction system that has<br />

several confluent subreductions.<br />

Definition 27 (Commutativity). Let → 1, → 2 be two notions of<br />

reduction. → 1 (strongly) commute <strong>with</strong> → 2 if a → 1 b 1 and a → 2<br />

b 2, then there exists a c such that b 1 → 2 c and b 2 → 1 c.<br />

Proposition 28 (Hindley-Rosen [8] [11]). Let → 1, → 2 be two<br />

notions of reduction. Suppose both → 1 and → 2 are confluent, and<br />

→ ∗ 1 commutes <strong>with</strong> → ∗ 2. Then → 1 ∪ → 2 is confluent.<br />

2 → β,µ denotes → β ∪ → µ, we will use this convention throughout the<br />

paper.


Proposition 29 (Weak Commutativity [2]). Let ↩→ denote the<br />

reflexive closure of →. Let → 1, → 2 be two notions of reduction.<br />

→ 1 weak commutes <strong>with</strong> → 2 if a → 1 b 1 and a → 2 b 2, then there<br />

exists a c such that b 1 ↩→ 2 c and b 2 ↠ 1 c.<br />

If → 1 weak commutes <strong>with</strong> → 2, then → ∗ 1 and → ∗ 2 commute.<br />

Lemma 30. → β,µ commutes <strong>with</strong> → ι. Thus → β,µ,ι is confluent.<br />

Lemma 31. → o has diamond property, thus is confluent.<br />

Lemma 32. → o commutes <strong>with</strong> → ι, weak commutes <strong>with</strong> → β ,<br />

→ µ.<br />

Theorem 33. → o,ι,β,µ is confluent.<br />

Let = β,µ,ι,o denotes the reflexive transitive symmetry closure of<br />

→ o ∪ → ι ∪ → β ∪ → µ. The goal of confluence analysis is to<br />

establish the following theorem.<br />

Theorem 34 (ι-elimination, Compatibility).<br />

If Γ ⊢ Πx : t 1.t 2 = β,µ,ι,o Πx : t ′ 1.t ′ 2, then Γ ⊢ t 1 = β,µ,o t ′ 1<br />

and Γ ⊢ t 2 = β,µ,o t ′ 2.<br />

Proof. If Γ ⊢ Πx : t 1.t 2 = β,µ,ι,o Πx : t ′ 1.t ′ 2, then by the<br />

confluence of → β,µ,ι,o , there exists a t such that Γ ⊢ Πx :<br />

t 1.t 2(→ o,ι,β,µ ) ∗ t and Γ ⊢ Πx : t ′ 1.t ′ 2(→ o,ι,β,µ ) ∗ t. Since all the<br />

reductions on Πx : t 1.t 2 preserve the structure of the dependent<br />

type, one will never have a chance to use → ι-reduction, thus Γ ⊢<br />

Πx : t 1.t 2(→ o,β,µ ) ∗ t and Γ ⊢ Πx : t ′ 1.t ′ 2(→ o,β,µ ) ∗ t. So t must<br />

be of the form Πx : t 3.t 4. And Γ ⊢ t 1(→ o,β,µ ) ∗ t 3, Γ ⊢ t ′ 1(→ o,β,µ<br />

) ∗ t 3, Γ ⊢ t 2(→ o,β,µ ) ∗ t 4 and Γ ⊢ t ′ 2(→ o,β,µ ) ∗ t 4. Finally, we have<br />

Γ ⊢ t 1 = β,µ,o t ′ 1 and Γ ⊢ t 2 = β,µ,o t ′ 2.<br />

3.3 Type Preservation<br />

The proof of type preservation proceeds as usual. The inversion<br />

lemma and substitution lemma are standard. Note that in the final<br />

preservation proof, we use the compatibility theorem.<br />

Lemma 35 (Inversion).<br />

• If Γ ⊢ λx.t : t ′ , then Γ, x : t 1 ⊢ t : t 2 and Γ ⊢ Πx :<br />

t 1.t 2= β,µ,ι,o t ′ for some t 1, t 2.<br />

• If Γ ⊢ t 1t 2 : t ′ , then Γ ⊢ t 1 : Πx : t ′ 1.t ′ 2 and Γ ⊢ t 2 : t ′ 1,<br />

Γ ⊢ [t 2/x]t ′ 2= β,µ,ι,o t ′ for some t ′ 1, t ′ 2.<br />

• If Γ ⊢ x : t ′ , then x : t ∈ Γ and Γ ⊢ t= β,µ,ι,o t ′ for some t.<br />

Lemma 36 (Substitution). If Γ 1, x : t 1, Γ 2 ⊢ t : t 2 and Γ ⊢ t ′ : t 1,<br />

then Γ 1, [t ′ /x]Γ 2 ⊢ [t ′ /x]t : [t ′ /x]t 2.<br />

Theorem 37 (Type Preservation). If Γ ⊢ wf and Γ ⊢ t t ′ and<br />

Γ ⊢ t : t ′′ , then Γ ⊢ t ′ : t ′′ .<br />

Proof. We list one interesting case here.<br />

Γ ⊢ t ′ 1 : Πx : t ′′<br />

1 .t ′′<br />

2 Γ ⊢ t ′ 2 : t ′′<br />

1<br />

Γ ⊢ t ′ 1t ′ 2 : [t ′ 2/x]t ′′<br />

2<br />

Suppose Γ ⊢ (λx.t 1)v [v/x]t 1. Then we know Γ ⊢ (λx.t 1)v :<br />

[t/x]t ′′<br />

2 and Γ ⊢ λx.t 1 : Πx : t ′′<br />

1 .t ′′<br />

2 and Γ ⊢ v : t ′′<br />

1 . By inversion<br />

on Γ ⊢ λx.t 1 : Πx : t ′′<br />

1 .t ′′<br />

2 , we have Γ, x : a ⊢ t 1 : b and Γ ⊢ Πx :<br />

a.b= β,µ,ι,o Πx : t ′′<br />

1 .t ′′<br />

2 . By theorem 34, we have Γ ⊢ a = β,µ,o t ′′<br />

1<br />

and Γ ⊢ b = β,µ,o t ′′<br />

2 . So we have Γ, x : a ⊢ t 1 : t ′′<br />

2 and Γ ⊢ v : a.<br />

So by lemma 79, we have Γ ⊢ [v/x]t 1 : [v/x]t ′′<br />

2 , as required.<br />

4. Conclusion<br />

We introduce Selfstar, which incorporates the self type construct<br />

together <strong>with</strong> ∗ : ∗ and mutually recursion. <strong>Scott</strong>-encoded<br />

datatypes and the corresponding elimination schemes are derivable<br />

<strong>with</strong>in Selfstar. We also demonstrate the process of proving the<br />

type preservation theorem.<br />

References<br />

[1] Zena M. Ariola and Jan Willem Klop. Lambda calculus <strong>with</strong> explicit<br />

recursion. Information and Computation, 139(2):154 – 233, 1997.<br />

[2] H. P. Barendregt. The Lambda Calculus, Its Syntax and Semantics<br />

(Studies in Logic and the Foundations of Mathematics, Volume 103).<br />

Revised Edition. North Holland, revised edition, November 1985.<br />

[3] Pierre-Louis Curien, Thérèse Hardin, and Jean-Jacques Lévy. Confluence<br />

properties of weak and strong calculi of explicit substitutions. J.<br />

ACM, 43(2):362–397, 1996.<br />

[4] H. B. Curry, J. R. Hindley, and J. P. Seldin. Combinatory Logic,<br />

Volume II. North-Holland, 1972.<br />

[5] Peng Fu and Aaron Stump. Church encoding <strong>with</strong> dependent types.<br />

2013. submitted, available from first Author’s web page.<br />

[6] Jean-Yves Girard, Paul Taylor, and Yves Lafont. Proofs and types.<br />

Cambridge University Press, New York, NY, USA, 1989.<br />

[7] Thérèse Hardin. Confluence results for the pure strong categorical<br />

logic ccl. λ-calculi as subsystems of ccl. Theor. Comput. Sci.,<br />

65(3):291–342, July 1989.<br />

[8] James Roger Hindley. The Church-Rosser property and a result in<br />

combinatory logic. PhD thesis, University of Newcastle upon Tyne,<br />

1964.<br />

[9] J.M. Jansen, R. Plasmeijer, and P. Koopman. Functional pearl: Comprehensive<br />

encoding of data types and algorithms in the lambdacalculus.<br />

Internal report, NLDA, 2011.<br />

[10] George Kuan, David MacQueen, and Robert Bruce Findler. A rewriting<br />

semantics for type inference. In Proceedings of the 16th European<br />

conference on <strong>Programming</strong>, pages 426–440. Springer-Verlag, 2007.<br />

[11] Barry K. Rosen. Tree-manipulating systems and church-rosser theorems.<br />

J. ACM, 20(1):160–187, January 1973.<br />

[12] Aaron Stump, Garrin Kimmell, and Roba El Haj Omar. Type preservation<br />

as a confluence problem. In RTA, pages 345–360, 2011.<br />

A. Well-Form Type<br />

Lemma 38. If Γ ⊢ wf and Γ ⊢ t : t ′ , then Γ ⊢ t ′ : ∗.<br />

Proof. By induction on derivation of Γ ⊢ t : t ′ . We list a few<br />

nontrivial cases.<br />

Case:<br />

Γ ⊢ t : ιx.t ′<br />

Γ ⊢ t : [t/x]t ′<br />

SelfInst<br />

By IH, we have Γ ⊢ ιx.t ′ : ∗. So by inversion, we have Γ, x :<br />

ιx.t ′ ⊢ t ′ : ∗. So by lemma 79, we know Γ ⊢ [t/x]t ′ : ∗.<br />

Case:<br />

Γ, x : t 1 ⊢ t : t 2 Γ ⊢ t 1 : ∗<br />

Γ ⊢ λx.t : Πx : t 1.t 2<br />

Lam<br />

By IH, we know Γ, x : t 1 ⊢ t 2 : ∗. Since Γ ⊢ t 1 : ∗, by Pi rule, we<br />

have Γ ⊢ Πx : t 1.t 2 : ∗.<br />

Case:


Γ ⊢ t : Πx : t 1.t 2 Γ ⊢ t ′ : t 1<br />

Γ ⊢ tt ′ : [t ′ /x]t 2<br />

App<br />

By IH, we have Γ ⊢ Πx : t 1.t 2 : ∗. By inversion on Γ ⊢ Πx :<br />

t 1.t 2 : ∗, we have Γ, x : t 1 ⊢ t 2 : ∗. So by lemma 79, we have<br />

Γ ⊢ [t ′ /x]t 2 : ∗.<br />

Case:<br />

Γ, ˜µ ⊢ t : t ′ {Γ, ˜µ ⊢ t j : a j} (tj :a j )∈˜µ<br />

Γ ⊢ µt : µt ′<br />

Mu<br />

By IH, we have Γ, ˜µ ⊢ t ′ : ∗. So Γ ⊢ µt ′ : µ∗, thus Γ ⊢ µt ′ : ∗.<br />

B. Progress<br />

Lemma 39. If · ⊢ v : Πx : t 1.t 2, then v ≡ λx.t.<br />

Proof. Case analysis on v. Suppose v ≡ ∗. By inversion, · ⊢ ∗ : ∗<br />

and · ⊢ ∗ = β,µ,ι,o Πx : t 1.t 2, which contradicts Church-Rosser<br />

of = β,µ,ι,o . Suppose v ≡ ⃗µ(Πx : t 3.t 4). By inversion, we have<br />

˜⃗µ ⊢ Πx : t 3.t 4 : t a and · ⊢ ⃗µt a<br />

⃗µ(Πx:t 3 .t 4 )<br />

= β,µ,ι,o Πx : t 1.t 2. By<br />

inversion on ˜⃗µ ⊢ Πx : t 3.t 4 : t a, we have ˜⃗µ ⊢ ∗ Πx:t =<br />

3.t 4<br />

β,µ,ι,o t a.<br />

So we have · ⊢ ⃗µ∗ ⃗µ(Πx:t 3.t 4 )<br />

⃗µ(Πx:t 3 .t 4 )<br />

= β,µ,ι,o ⃗µt a = β,µ,ι,o Πx :<br />

t 1.t 2. Again, this contradicts Church-Rosser of = β,µ,ι,o . For other<br />

cases like: v ≡ Πx : t.t ′ , ιx.t, ⃗µ(ιx.t), we argue similarly.<br />

Theorem 40 (Progress). If · ⊢ t : t ′′ , then either · ⊢ t t ′ or t is<br />

a value.<br />

Proof. By induction on the derivation of · ⊢ t : t ′′ , we list a few<br />

cases.<br />

Case:<br />

˜µ ⊢ t : t ′ {˜µ ⊢ t j : a j} (tj :a j )∈˜µ<br />

· ⊢ µt : µt ′ Mu<br />

Identify t as ˙⃗µt ′′ , where t ′′ does not contains any closure at head<br />

position. Case analysis on t ′′ , if it is ∗, x, λx.t a, t at b , then there<br />

exist a t ′ such that · ⊢ t t ′ . If t ′′ ≡ Πx : t a.t b , ιx.t a, then it is<br />

already a value.<br />

Case:<br />

· ⊢ t : Πx : t 1.t 2 · ⊢ t ′ : t 1<br />

· ⊢ tt ′ : [t ′ /x]t 2<br />

App<br />

Since · ⊢ t : Πx : t 1.t 2 and · ⊢ t ′ : t 1, by IH, t either steps or is<br />

a value, likewise for t ′ . If t can take a step, then tt ′ can also take a<br />

step. If t is a value, by lemma 39, t must be of the form λx.t a. So if<br />

t ′ can take a step, then tt ′ can also take a step. If both t ′ is a value,<br />

then tt ′ can take a step.<br />

C. Proofs of Section 3.1<br />

Let ˙⃗µ denote 0 or more closures.<br />

Lemma 41. Let Φ denote the set of µ normal form. For any term<br />

t, m(t) ∈ Φ.<br />

Proof. One way to prove this is first identify t as ˙−→ µ1t ′ , here ˙−→ µ 1<br />

means there are zero or more closures and t ′ does not contains any<br />

closure at head position. Then we can proceed by induction on the<br />

structure of t ′ :<br />

Base Cases: t ′ = x, t ′ = ∗, obvious.<br />

Step Cases: If t ′ = λx.t ′′ , then m( µ ˙−→ 1(λx.t ′′ )) ≡ λx.m( µ ˙−→ 1t ′′ ).<br />

Now we can again identify t ′′ as µ ˙−→ 2t ′′′ , where t ′′′ does not have<br />

any closure at head position. Since t ′′′ is structurally smaller<br />

than λx.t ′′ , by IH, m( µ ˙−→ 1µ2t ˙−→ ′′′ ) ∈ Φ, thus m( µ ˙−→ 1(λx.t ′′ )) ≡<br />

λx.m( µ ˙−→ 1t ′′ ) ∈ Φ.<br />

For t ′ = t at b , t ′ = ιx.t ′′ , t ′ = Πx : t a.t b , we can argue<br />

similarly as above.<br />

In order to prove lemma 22, we prove the following more<br />

general lemma instead.<br />

Lemma 42. If Γ, ˙⃗µ ⊢ a → β b, then Γ ⊢ m( ˙⃗µa) → βµ m( ˙⃗µb).<br />

Proof. By induction on derivation of Γ, ˙⃗µ ⊢ a → β b.<br />

Base Case:<br />

(x ↦→ t) ∈ Γ, ˙⃗µ<br />

Γ, ˙⃗µ ⊢ x → β t<br />

If x ↦→ t ∈ ˙⃗µ, then Γ ⊢ m( ˙⃗µx) ≡ µx → βµ m(µt) ≡ m( ˙⃗µt).<br />

Techincally, the last equality need to be justified, informally we<br />

can justify that by locality of µ. If x ↦→ t ∈ Γ, then Γ ⊢ m( ˙⃗µx) ≡<br />

x → βµ m(t) ≡ m( ˙⃗µt).<br />

Base Case:<br />

(x i ↦→ t i) ∈ µ<br />

Γ, ˙⃗µ ⊢ µx i → β µt i<br />

We have Γ ⊢ m( ˙⃗µµx i) ≡ µx i → βµ m(µt i) ≡ m( ˙⃗µµt i).<br />

Base Case:<br />

Γ, ˙⃗µ ⊢ (λx.t)t ′ → β [t ′ /x]t<br />

We have Γ ⊢ m( ˙⃗µ((λx.t)t ′ )) ≡ (λx.m( ˙⃗µt))m( ˙⃗µt ′ ) → βµ<br />

m([m( ˙⃗µt)/x]m( ˙⃗µt ′ )) ≡ m([ ˙⃗µt/x] ˙⃗µt ′ ) ≡ m( ˙⃗µ([t ′ /x]t)). The<br />

last two equalities are by lemma 44, lemma 43.<br />

Step Case:<br />

Γ, ˙⃗µ ⊢ t → β t ′<br />

Γ, ˙⃗µ ⊢ λx.t → β λx.t ′<br />

Γ ⊢ m( ˙⃗µ(λx.t)) ≡ λx.m( ˙⃗µt)<br />

Step Case:<br />

Γ, ˙⃗µ, ˜µ ⊢ t → β t ′<br />

Γ, ˙⃗µ ⊢ µt → β µt ′<br />

IH<br />

→ βµ λx.m( ˙⃗µt ′ ) ≡ m( ˙⃗µ(λx.t ′ )).<br />

We want to show Γ ⊢ m( ˙⃗µµt) → βµ m( ˙⃗µµt ′ ). This is directly by<br />

IH.<br />

All the other cases are similar.<br />

Lemma 43.<br />

m(⃗µ⃗µt) ≡ m(⃗µt) and m(⃗µ([t 2/x]t 1)) ≡ m([⃗µt 2/x]⃗µt 1)


Proof. We can prove this using the same method as lemma 41,<br />

namely, identify t and then proceed by inducton.<br />

Lemma 44. m(m(t)) ≡ m(t) and m([m(t 1)/y]m(t 2)) ≡<br />

m([t 1/y]t 2).<br />

Proof. The first equality is by lemma 45 and lemma 41. For the<br />

second equality, we prove it using similar method as lemma 41: We<br />

identify t 2 as ˙−→ µ 1t ′ 2, where t ′ 2 does not contains any closure at head<br />

position. We proceed by induction on the structure of t ′ 2:<br />

Base Cases: t ′ 2 = ∗, obvious. For t ′ 2 = x, we use m(m(t)) ≡<br />

m(t).<br />

Step Cases: If t ′ 2 = λx.t ′′<br />

2 , then m( µ ˙−→ 1(λx.[t 1/y]t ′′<br />

2 )) ≡<br />

λx.m( µ ˙−→ 1([t 1/y]t ′′<br />

2 )) ≡ λx.m( µ ˙−→ 1µ2([t ˙−→ 1/y]t ′′′<br />

2 )), where t ′′<br />

2 is identified<br />

as µ ˙−→ 2t ′′′<br />

2 , and t ′′′<br />

2 does not have any closure at head position.<br />

Since t ′′′<br />

2 is structurally smaller than λx.t ′′<br />

2 , by IH,<br />

m( µ ˙−→ 1µ2([t ˙−→ 1/y]t ′′′<br />

2 )) ≡ m([t 1/y]( µ ˙−→ 1µ2t ˙−→ ′′′<br />

2 )) ≡<br />

m([m(t 1)/y]m( µ ˙−→ 1µ2t ˙−→ ′′′<br />

2 )).<br />

Thus λx.m( µ ˙−→ 1µ2([t ˙−→ 1/y]t ′′′<br />

2 )) ≡ λx.m([m(t 1)/y]m( µ ˙−→ 1µ2t ˙−→ ′′′<br />

2 )),<br />

implying m([t 1/y] µ ˙−→ 1(λx.t ′′<br />

2 )) ≡ m([m(t 1)/y]m(λx. µ ˙−→ 1t ′′<br />

2 )).<br />

Since m([m(t 1)/y]m(λx. µ ˙−→ 1t ′′<br />

2 )) ≡ m([m(t 1)/y]m( µ ˙−→ 1(λx.t ′′<br />

2 ))),<br />

we conclude m([m(t 1)/y]m( µ ˙−→ 1(λx.t ′′<br />

2 ))) ≡ m([t 1/y] µ ˙−→ 1(λx.t ′′<br />

2 )).<br />

For t ′ 2 = t at b , t ′ 2 = ιx.t ′′<br />

2 , t ′ 2 = Πx : t a.t b , we can argue<br />

similarly as above.<br />

Lemma 45. If n ∈ Φ, then m(n) ≡ n.<br />

Proof. By induction on the structure of n.<br />

Definition 46 (β Reduction on µ-normal Forms).<br />

Γ ⊢ n → β t<br />

Γ ⊢ n → βµ m(t)<br />

Note: From this definition we can conclude:<br />

Γ ⊢ n → βµ n ′<br />

Γ ⊢ λx.n → βµ λx.n ′ Γ ⊢ n ′ → βµ n ′′<br />

Γ ⊢ nn ′ → βµ nn ′′<br />

Γ ⊢ n ′ → βµ n ′′<br />

Γ ⊢ n → βµ n ′′<br />

Γ ⊢ Πx : n.n ′ → βµ Πx : n.n ′′ Γ ⊢ nn ′ → βµ n ′′ n ′<br />

Γ ⊢ n → βµ n ′′<br />

Γ ⊢ n → βµ n ′<br />

Γ ⊢ Πx : n.n ′ → βµ Πx : n ′′ .n ′ Γ ⊢ ιx.n → βµ ιx.n ′<br />

The first rule follows because: Assume Γ ⊢ n → βµ n ′ , say<br />

m(t) ≡ n ′ and Γ ⊢ n → β t. Then Γ ⊢ λx.n → β λx.t and<br />

m(λx.t) ≡ λx.m(t) ≡ λx.n ′ . The others follow similarly.<br />

Lemma 47. If Γ ⊢ n 1 → βµ n ′ 1, then Γ ⊢ m([n 2/x]n 1) → βµ<br />

m([n 2/x]n ′ 1).<br />

Proof. By induction on derivation of Γ ⊢ n 1 → β t 1, where<br />

m(t 1) ≡ n ′ 1. We will list a few nontrivial cases. Note that the we<br />

use lemma 44 implicitly.<br />

Base Case:<br />

(y ↦→ t 1) ∈ Γ<br />

Γ ⊢ y → β t 1<br />

In this case n 1 = y. By locality, we have Γ ⊢ m([n 2/x]y) ≡<br />

y → βµ m(t 1) ≡ m([n 2/x]t 1).<br />

Base Case:<br />

Γ ⊢ (λy.n)n ′ → β [n ′ /y]n<br />

n 1 = (λy.n)n ′ . So Γ ⊢ m([n 2/x]((λy.n)n ′ ))<br />

≡ m((λy.[n 2/x]n)[n 2/x]n ′ ) ≡ (λy.m([n 2/x]n))m([n 2/x]n ′ ) → βµ<br />

m([m([n 2/x]n ′ )/y]m([n 2/x]n)) ≡ m([[n 2/x]n ′ /y]([n 2/x]n)) ≡<br />

m([n 2/x]([n ′ /y]n)).<br />

Base Case:<br />

(x i ↦→ t i) ∈ µ<br />

Γ ⊢ µx i → β µt i<br />

n 1 = µx i. By locality, Γ ⊢ m([n 2/x]µx i) ≡ µx i → βµ<br />

m(µt i) ≡ m([n 2/x](µt i)).<br />

Step Case:<br />

Γ ⊢ n → β t ′<br />

Γ ⊢ λy.n → β λy.t ′<br />

n 1 = λy.n. By IH, we have Γ ⊢ m([n 2/x]n) → βµ m([n 2/x]t ′ ).<br />

So Γ ⊢ m(λy.[n 2/x]n) → βµ m(λy.[n 2/x]t ′ ).<br />

Step Case:<br />

Γ, ˜µ ⊢ t → β t ′<br />

Γ ⊢ µt → β µt ′<br />

This case will not arise since n 1 is already in µ normal form.<br />

The other cases are similar.<br />

Lemma 48. If Γ ⊢ n 2 → βµ n ′ 2, then Γ ⊢ m([n 2/x]n 1)<br />

m([n ′ 2/x]n 1).<br />

Proof. By induction on n 1.<br />

Definition 49 (Parallel Reductions).<br />

(x ↦→ t) ∈ Γ<br />

Γ ⊢ n ⇒ βµ n<br />

(x i ↦→ t i ) ∈ µ<br />

Γ ⊢ µx i ⇒ βµ m(µt i )<br />

Γ ⊢ x ⇒ βµ m(t)<br />

Γ ⊢ n 1 ⇒ βµ n ′ 1 Γ ⊢ n 2 ⇒ βµ n ′ 2<br />

Γ ⊢ (λx.n 1 )n 2 ⇒ βµ m([n ′ 2 /x]n′ 1 )<br />

Γ ⊢ n ⇒ βµ n ′ Γ ⊢ n ⇒ βµ n ′′ Γ ⊢ n ′ ⇒ βµ n ′′′<br />

Γ ⊢ λx.n ⇒ βµ λx.n ′ Γ ⊢ nn ′ ⇒ βµ n ′′ n ′′′<br />

Γ ⊢ n ⇒ βµ n ′ Γ ⊢ n ′ ⇒ βµ n ′′′ Γ ⊢ n ⇒ βµ n ′′<br />

Γ ⊢ ιx.n ⇒ βµ ιx.n ′ Γ ⊢ Πx : n.n ′ ⇒ βµ Πx : n ′′ .n ′′′<br />

Lemma 50. → βµ ⊆⇒ βµ ⊆→ ∗ βµ.<br />

∗<br />

→ βµ<br />

Proof. For → βµ ⊆⇒ βµ , by induction on the derivation of Γ ⊢<br />

n → β t, where Γ ⊢ n → βµ m(t).<br />

For ⇒ βµ ⊆→ ∗ βµ, by induction on the derivation of Γ ⊢ n ⇒ βµ<br />

n ′ . We show the case where(the other cases are obvious):<br />

Γ ⊢ n 1 ⇒ βµ n ′ 1 Γ ⊢ n 2 ⇒ βµ n ′ 2<br />

Γ ⊢ (λx.n 1)n 2 ⇒ βµ m([n ′ 2/x]n ′ 1)<br />

By lemma 52, we know that Γ ⊢ m([n 2/x]n 1) ⇒ βµ m([n 2/x]n 1),<br />

given Γ ⊢ n 1 ⇒ βµ n ′ 1, Γ ⊢ n 2 ⇒ βµ n ′ 2. Since → βµ ⊆⇒ βµ ,


we have: if Γ ⊢ n 1 → βµ n ′ 1, Γ ⊢ n 2 → βµ n ′ 2 , then Γ ⊢<br />

m([n 2/x]n 1) → βµ m([n ′ 2/x]n ′ 1)(†). By IH, we have Γ ⊢<br />

∗<br />

n 1 → βµ n ′ ∗<br />

1, Γ ⊢ n 2 → βµ n ′ 2. By lemma 47, lemma 48 and (†),<br />

∗<br />

we have Γ ⊢ (λx.n 1)n 2 → βµ m([n 2/x]n 1) → βµ m([n ′ 2/x]n ′ 1).<br />

Lemma 51. If Γ ⊢ n 2 ⇒ βµ n ′ 2, then Γ ⊢ m([n 2/x]n 1) ⇒ βµ<br />

m([n ′ 2/x]n 1).<br />

Proof. By induction on the structure of n 1.<br />

Base Cases: n 1 = x, n 1 = µx i, n 1 = ∗. Obvious.<br />

Step Case: n 1 = λy.n. We have Γ ⊢ m(λy.[n 2/x]n) ≡<br />

λy.m([n 2/x]n) ⇒ IH<br />

βµ λy.m([n ′ 2/x]n) ≡ m(λy.[n ′ 2/x]n).<br />

Step Case: n 1 = nn ′ . We have Γ ⊢ m([n 2/x]n[n 2/x]n ′ ) ≡<br />

m([n 2/x]n)m([n 2/x]n ′ IH<br />

) ⇒ βµ m([n ′ 2/x]n)m([n ′ 2/x]n ′ ) ≡<br />

m([n ′ 2/x]n[n ′ 2/x]n).<br />

Step Case: n 1 = ιx.n, Πx : n.n ′ . Similar as above.<br />

Lemma 52. If Γ ⊢ n 1 ⇒ βµ n ′ 1 and Γ ⊢ n 2 ⇒ βµ n ′ 2, then<br />

Γ ⊢ m([n 2/y]n 1) ⇒ βµ m([n ′ 2/y]n ′ 1).<br />

Proof. We prove this by induction on the derivation of Γ<br />

n 1 ⇒ βµ n ′ 1.<br />

Base Case:<br />

Γ ⊢ n ⇒ βµ n<br />

By lemma 51.<br />

Base Case:<br />

x i ↦→ t i ∈ µ<br />

Γ ⊢ µx i ⇒ βµ m(µt i)<br />

Since y /∈ FV(µx i) and µ is local, m([n 2/y]µx i) ≡ m(µx i),<br />

then m(µx 1) ≡ µx i ⇒ βµ m(µt i) ≡ m(m(µt i))(lemma 44).<br />

Base Case:<br />

(x ↦→ t) ∈ Γ<br />

Γ ⊢ x ⇒ βµ m(t)<br />

In this case, we assume x ≢ y, then we have m([n 2/y]x) ≡<br />

m(x) ≡ x ⇒ βµ m(t) ≡ m(m(t)) ≡ m([n 2/y]m(t)).<br />

Step Case:<br />

Γ ⊢ n a ⇒ βµ n ′ a Γ ⊢ n b ⇒ βµ n ′ b<br />

Γ ⊢ (λx.n a)n b ⇒ βµ m([n ′ a/x]n ′ b)<br />

We have Γ ⊢ m((λx.[n 2/y]n a)[n 2/y]n b ) ≡<br />

(λx.m([n 2/y]n a))m([n 2/y]n b ) ⇒ IH<br />

βµ<br />

m([m([n ′ 2/y]n ′ b)/x]m([n ′ 2/y]n ′ a)) ≡ m([n ′ 2/y]([n ′ b/x]n ′ a)).<br />

The last equality is by lemma 44. Here we first apply induction<br />

hypothesis to reduce, then apply ⇒ βµ .<br />

Step Case:<br />

⊢<br />

Γ ⊢ n ⇒ βµ n ′<br />

Γ ⊢ λx.n ⇒ βµ λx.n ′<br />

We have Γ ⊢ m(λx.[n 2/y]n) ≡ λx.m([n 2/y]n)<br />

λx.m([n ′ 2/y]n ′ ) ≡ m(λx.[n ′ 2/y]n ′ )<br />

Step Case:<br />

Γ ⊢ n a ⇒ βµ n ′ a<br />

Γ ⊢ n b ⇒ βµ n ′ b<br />

Γ ⊢ n an b ⇒ βµ n ′ an ′ b<br />

IH<br />

⇒ βµ<br />

We have Γ ⊢ m([n 2/y]n a[n 2/y]n b ) ≡ m([n 2/y]n a)m([n 2/y]n b )<br />

IH<br />

⇒ βµ m([n ′ 2/y]n ′ a)m([n ′ 2/y]n ′ b) ≡ m([n ′ 2/y](n ′ an ′ b)).<br />

The other cases are similar as above.<br />

Lemma 53 (Diamond Property). If Γ ⊢ n ⇒ βµ n ′ and Γ ⊢<br />

n ⇒ βµ n ′′ , then there exist n ′′′ such that Γ ⊢ n ′′ ⇒ βµ n ′′′ and<br />

Γ ⊢ n ′ ⇒ βµ n ′′′ .<br />

Proof. By induction on the derivation of Γ ⊢ n ⇒ βµ n ′ .<br />

Base Case:<br />

Γ ⊢ n ⇒ βµ n<br />

Obvious.<br />

Base Case:<br />

(x ↦→ t) ∈ Γ<br />

Γ ⊢ x ⇒ βµ m(t)<br />

Obvious.<br />

Base Case:<br />

Γ ⊢ µx i ⇒ βµ m(µt i)<br />

Obvious.<br />

Step Case:<br />

Γ ⊢ n 1 ⇒ βµ n ′ 1 Γ ⊢ n 2 ⇒ βµ n ′ 2<br />

Γ ⊢ (λx.n 1)n 2 ⇒ βµ m([n ′ 2/x]n ′ 1)<br />

Suppose Γ ⊢ (λx.n 1)n 2 ⇒ βµ (λx.n ′′<br />

1 )n ′′<br />

2 , where Γ ⊢ n 1 ⇒ βµ<br />

n ′′<br />

1 and Γ ⊢ n 2 ⇒ βµ n ′′<br />

2 . By IH, there exist n ′′′<br />

1 , n ′′′<br />

2 such that<br />

Γ ⊢ n ′′<br />

1 ⇒ βµ n ′′′<br />

1 and Γ ⊢ n ′ 1 ⇒ βµ n ′′′<br />

1 and Γ ⊢ n ′ 2 ⇒ βµ n ′′′<br />

2<br />

and Γ ⊢ n ′ 2 ⇒ βµ n ′′′<br />

2 . By lemma 52, Γ ⊢ m([n ′ 1/x]n ′ 2) ⇒ βµ<br />

m([n ′′′<br />

1 /x]n ′′′<br />

2 ), also Γ ⊢ (λx.n ′′<br />

1 )n ′′<br />

2 ⇒ βµ m([n ′′′<br />

1 /x]n ′′′<br />

2 ).<br />

Suppose Γ ⊢ (λx.n 1)n 2 ⇒ βµ m([n ′′<br />

2 /x]n ′′<br />

1 ), where Γ ⊢ n 1 ⇒ βµ<br />

n ′′<br />

1 and Γ ⊢ n 2 ⇒ βµ n ′′<br />

2 . By IH, there exist n ′′′<br />

1 , n ′′′<br />

2 such that<br />

Γ ⊢ n ′′<br />

1 ⇒ βµ n ′′′<br />

1 and Γ ⊢ n ′ 1 ⇒ βµ n ′′′<br />

1 and Γ ⊢ n ′ 2 ⇒ βµ n ′′′<br />

2<br />

and Γ ⊢ n ′ 2 ⇒ βµ n ′′′<br />

2 . By lemma 52, Γ ⊢ m([n ′ 1/x]n ′ 2) ⇒ βµ<br />

m([n ′′′<br />

1 /x]n ′′′<br />

2 ) and Γ ⊢ m([n ′′<br />

1 /x]n ′′<br />

2 ) ⇒ βµ m([n ′′′<br />

1 /x]n ′′′<br />

2 ).<br />

The other cases are either similar to the one above or easy.<br />

Theorem 54. → β ∪ → µ is confluent.<br />

Proof. We know by diamond property of ⇒ βµ , → βµ is confluent.<br />

Since → µ is strongly normalizing and confluent, and by lemma


22 and Hardin’s interpretation lemma(lemma 17), we conclude<br />

→ β ∪ → µ is confluent.<br />

D. Proofs of Section 3.2<br />

Lemma 55. Let → denote → β ∪ → µ, if Γ ⊢ t → t ′ , then<br />

Γ ⊢ [t 1/x]t → [t 1/x]t ′ for any t 1.<br />

Proof. Obvious.<br />

Lemma 56. Let → denote → β ∪ → µ, then → commutes <strong>with</strong> → ι.<br />

i.e. if Γ ⊢ t 1 → t 2 and Γ ⊢ t 1 → ι t 3, then there exist t 4 such that<br />

Γ ⊢ t 2 → ι t 4 and Γ ⊢ t 3 → t 4.<br />

Proof. Since Γ ⊢ t 1 → ι t 3, we know that t 1 ≡ ιx.t ′ and<br />

t 3 ≡ [t/x]t ′ . We also have Γ ⊢ t 1 ≡ ιx.t ′ → t 2. By inversion,<br />

we know that t 2 ≡ ιx.t ′′ <strong>with</strong> Γ ⊢ t ′ → t ′′ . By lemma 55,<br />

we know that Γ ⊢ [t/x]t ′ → [t/x]t ′′ . Thus t 4 ≡ [t/x]t ′′ and<br />

Γ ⊢ ιx.t ′′ → ι [t/x]t ′′ .<br />

Theorem 57. → ∪ → ι is confluent.<br />

Lemma 58. If Γ ⊢ t 1 → o t 2, then Γ ⊢ [t/x]t 1 → o [t/x]t 2.<br />

Proof. By induction on derivaton.<br />

Lemma 59. If Γ ⊢ t 1 → o t 2, then Γ ⊢ [t 1/x]t ↩→ o [t 2/x]t.<br />

Proof. By induction on the structure of t.<br />

Lemma 60. → o has diamond property, thus is confluent.<br />

Proof. Straightforward induction.<br />

Lemma 61. → o commutes <strong>with</strong> → ι.<br />

Proof. Suppose Γ ⊢ ιx.t ′ → ι [t/x]t ′ and Γ ⊢ ιx.t ′ → o ιx.t ′′<br />

<strong>with</strong> Γ ⊢ t ′ → o t ′′ . Then by lemma 58, we have Γ ⊢ [t/x]t ′ → o<br />

[t/x]t ′′ . We also have Γ ⊢ ιx.t ′′ → ι [t/x]t ′′ .<br />

Lemma 62. → o weak commutes <strong>with</strong> → β .<br />

Proof. By induction on → o.<br />

Case: Γ ⊢ µt → o t, where µ ∈ Γ.<br />

If Γ ⊢ µx i → β µt i, where x i ↦→ t i ∈ µ, then Γ ⊢ µx i → o x i. So<br />

we have Γ ⊢ µt i → o t i and Γ ⊢ x i → β t i since µ ∈ Γ.<br />

If Γ ⊢ µt → β µt ′ , <strong>with</strong> Γ ⊢ t → β t ′ . So we have Γ ⊢ t → β t ′ and<br />

Γ ⊢ µt ′ → o t ′ .<br />

Case: Γ ⊢ (λx.t 1)t 2 → o (λx.t ′ 1)t 2, where Γ ⊢ t 1 → o t ′ 1.<br />

Suppose Γ ⊢ (λx.t 1)t 2 → β [t 2/x]t 1. By lemma 58, we know that<br />

Γ ⊢ [t 2/x]t 1 → o [t 2/x]t ′ 1. And we also have Γ ⊢ (λx.t ′ 1)t 2 → β<br />

[t 2/x]t ′ 1.<br />

Case: Γ ⊢ (λx.t 1)t 2 → o (λx.t 1)t ′ 2, where Γ ⊢ t 2 → o t ′ 2.<br />

Suppose Γ ⊢ (λx.t 1)t 2 → β [t 2/x]t 1. By lemma 59, we know that<br />

Γ ⊢ [t 2/x]t 1 ↩→ o [t ′ 2/x]t 1. And we also have Γ ⊢ (λx.t 1)t ′ 2 → β<br />

[t ′ 2/x]t 1.<br />

The other cases are by induction.<br />

Lemma 63. → o weak commutes <strong>with</strong> → µ. i.e. if Γ ⊢ t → o t ′ and<br />

Γ ⊢ t → µ t ′′ , then there exist a t 1 such that Γ ⊢ t ′′ → ∗ o t 1 and<br />

Γ ⊢ t ′ ↩→ µ t 1.<br />

Proof. By induction on Γ ⊢ t → o t ′ .<br />

Case: Γ ⊢ µt → o t, where µ ∈ Γ.<br />

Suppose Γ ⊢ µt → µ t <strong>with</strong> dom(µ)#FV(t). This case is obvious.<br />

Suppose t ≡ λx.t 2 and Γ ⊢ µ(λx.t 2) → µ λx.µt 2. Then Γ ⊢<br />

λx.t 2 ↩→ µ λx.t 2 and Γ ⊢ λx.µt 2 → o λx.t 2.<br />

Suppose t ≡ t 2t 3 and Γ ⊢ µ(t 2t 3) → µ (µt 2)(µt 3). Then<br />

Γ ⊢ t 2t 3 ↩→ µ t 2t 3 and Γ ⊢ (µt 2)(µt 3) → ∗ o t 2t 3.<br />

For t ≡ ιx.t 2, Πx : t 2.t 3, we can argue similarly.<br />

The other cases are by induction.<br />

Theorem 64. → o ∪ → ι ∪ → β ∪ → µ is confluent.<br />

E. Proofs of Section 3.3<br />

Note: In this section we use t = β,µ,ι,o to mean the same thing as<br />

= β,µ,ι,o , but <strong>with</strong> an emphasis on the subject t.<br />

Lemma 65. If Γ ⊢ t 1 t = β,µ,ι,o t 2 and Γ ⊢ t : t 1 and Γ ⊢ t 2 : ∗,<br />

then Γ ⊢ t : t 2.<br />

Proof. By induction on length of Γ ⊢ t 1 t = β,µ,ι,o t 2.<br />

Lemma 66. If Γ ⊢ t 1 = t β,µ,ι,o t 2 and Γ ⊢ t = t ′ , then<br />

t<br />

Γ ⊢ t ′<br />

1 = β,µ,ι,o t 2.<br />

Proof. By induction on length of Γ ⊢ t 1 t = β,µ,ι,o t 2.<br />

Lemma 67. m(µ 1µ 2t) ≡ m(µ 2µ 1t), thus Γ ⊢ µ 1µ 2t = µ 2µ 1t.<br />

Proof. Identify t as ˙⃗µt ′ , where t ′ does not have any closure at<br />

head position. By induction on the structure of such t ′ . Also Γ ⊢<br />

µ 1µ 2t = m(µ 1µ 2t) = m(µ 2µ 1t) = µ 2µ 1t.<br />

Lemma 68. Γ ⊢ µ([t/x]t ′ ) = [µt/x]µt ′<br />

Proof. Γ ⊢ µ([t/x]t ′ ) = m(µ([t/x]t ′ )) = m([µt/x]µt ′ ) =<br />

[µt/x]µt ′ .<br />

Lemma 69. If Γ, ˜µ ⊢ t ′ t = β,µ,ι,o t ′′ , then Γ ⊢ µt ′ µt<br />

= β,µ,ι,o µt ′′<br />

Proof. By induction on length of Γ, ˜µ ⊢ t ′ = t β,µ,ι,o t ′′ . We list a<br />

few cases.<br />

Case: Γ, ˜µ ⊢ t ′ =t ′′ .<br />

We have Γ ⊢ µt ′ = µt ′′ .<br />

Case: Γ, ˜µ ⊢ ιx.t ′ → ι[t/x]t ′ .<br />

We know Γ ⊢ µιx.t ′ =ιx.µt ′ → ι[µt/x]µt ′ =µ[t/x]t ′ (the last<br />

equality is by lemma 68).<br />

Lemma 70 (Inversion I). If Γ ⊢ λx.t : t ′ , then Γ, x : t 1 ⊢ t : t 2<br />

λx.t<br />

and Γ ⊢ Πx : t 1.t 2 = β,µ,ι,o t ′ .<br />

Proof. By induction on the derivation of Γ ⊢ λx.t : t ′ .<br />

Lemma 71 (Inversion II). If Γ ⊢ t 1t 2 : t ′ , then Γ ⊢ t 1 : Πx : t ′ 1.t ′ 2<br />

and Γ ⊢ t 2 : t ′ 1, Γ ⊢ [t 2/x]t ′ t 1 t 2<br />

2 = β,µ,ι,o t ′ .


Lemma 72 (Inversion III). If Γ ⊢ ∗ : t, then Γ ⊢ ∗ ∗ = β,µ,ι,o t.<br />

Lemma 73 (Inversion IV). If Γ ⊢ x : t ′ , then x : t ∈ Γ and<br />

Γ ⊢ t x = β,µ,ι,o t ′ .<br />

Lemma 74 (Inversion V). If Γ, ˜µ ⊢ x j : t ′ and x j ∈ dom(µ),<br />

x j<br />

then x j : a j ∈ µ and Γ, ˜µ ⊢ a j = β,µ,ι,o t ′ .<br />

Lemma 75 (Inversion VI). If Γ ⊢ ⃗µt : t ′ and t does not have a<br />

closure at head position, then Γ, ˜⃗µ ⊢ t : t ′′ ′′ ⃗µt<br />

and Γ ⊢ ⃗µt = β,µ,ι,o<br />

t ′ .<br />

Lemma 76 (Inversion VII). If Γ ⊢ ιx.t : t ′ , then Γ, x : ιx.t ⊢ t : ∗<br />

and Γ ⊢ ∗ ιx.t<br />

= β,µ,ι,o t ′ .<br />

Lemma 77 (Inversion VIII). If Γ ⊢ Πx : t 1.t 2 : t ′ , then Γ, x :<br />

t 1 ⊢ t 2 : ∗ and Γ ⊢ t 1 : ∗ and Γ ⊢ ∗ Πx:t =<br />

1.t 2<br />

β,µ,ι,o t ′ .<br />

Lemma 78. If Γ, ˜µ, y : b ⊢ t : a , then Γ, y : µb, ˜µ ⊢ t : a.<br />

Proof. By induction on the derivation of Γ, ˜µ, y : µb ⊢ t : t ′′ .<br />

Lemma 79 (Substitution). If Γ 1, x : t 1, Γ 2 ⊢ t : t 2 and Γ ⊢ t ′ : t 1,<br />

then Γ 1, [t ′ /x]Γ 2 ⊢ [t ′ /x]t : [t ′ /x]t 2.<br />

Proof. By induction on the derivation of Γ 1, x : t 1, Γ 2 ⊢ t : t 2 .<br />

We will show a few nontrivial cases.<br />

Case:<br />

Γ, y : ιy.t ⊢ t : ∗<br />

Γ ⊢ ιy.t : ∗<br />

Let Γ = Γ 1, x : t 1, Γ 2. We want to show Γ 1, [t ′ /x]Γ 2 ⊢<br />

ιy.[t ′ /x]t : ∗. By IH, we have Γ 1, [t ′ /x]Γ 2, y : ιy.[t ′ /x]t ⊢<br />

[t ′ /x]t : ∗. So it is the case.<br />

Case:<br />

Γ ⊢ t : [t/y]t ′′<br />

Γ ⊢ t : ιy.t ′′<br />

Γ ⊢ ιy.t ′′ : ∗<br />

Let Γ = Γ 1, x : t 1, Γ 2. We want to show Γ 1, [t ′ /x]Γ 2 ⊢<br />

[t ′ /x]t : ιy.[t ′ /x]t ′′ . By IH, we have Γ 1, [t ′ /x]Γ 2 ⊢ [t ′ /x]t :<br />

[[t ′ /x]t/y]([t ′ /x]t ′′ ). So it is the case.<br />

Case:<br />

Γ ⊢ t : ιy.t ′′<br />

Γ ⊢ t : [t/y]t ′′<br />

Let Γ = Γ 1, x : t 1, Γ 2. We want to show Γ 1, [t ′ /x]Γ 2 ⊢ [t ′ /x]t :<br />

[[t ′ /x]t/y]([t ′ /x]t ′′ ). By IH, we have Γ 1, [t ′ /x]Γ 2 ⊢ [t ′ /x]t :<br />

ιy.[t ′ /x]t ′′ . So it is the case.<br />

Case:<br />

Γ, ˜µ ⊢ t : t ′′ {Γ, ˜µ ⊢ t j : a j} (tj :a j )∈˜µ<br />

Γ ⊢ µt : µt ′′<br />

Let Γ = Γ 1, x : t 1, Γ 2. We want to show Γ 1, [t ′ /x]Γ 2 ⊢ µ[t ′ /x]t :<br />

µ[t ′ /x]t ′′ . By IH, we have Γ 1, [t ′ /x]Γ 2, [t ′ /x]˜µ ⊢ [t ′ /x]t :<br />

[t ′ /x]t ′′ and {Γ 1, [t ′ /x]Γ 2, [t ′ /x]˜µ ⊢ t j : [t ′ /x]a j} (tj :[t ′ /x]a j )∈[t ′ /x]˜µ.<br />

Theorem 80 (Type Preservation). If Γ ⊢ wf and Γ ⊢ t t ′ and<br />

Γ ⊢ t : a, then Γ ⊢ t ′ : a.<br />

Mu<br />

Proof. By induction on the derivation of Γ ⊢ t : a, We list a few<br />

nontrivial cases.<br />

Case:<br />

Γ ⊢ ∗ : ∗<br />

This case will not arise.<br />

Case:<br />

x : a ∈ Γ<br />

Γ ⊢ x : a<br />

If Γ ⊢ x t ′ , this means (x : a) ↦→ t ′ ∈ Γ and Γ ⊢ t ′ : a since<br />

Γ ⊢ wf.<br />

Case:<br />

Γ ⊢ t : t 1 Γ ⊢ t 1<br />

∼ = t2 Γ ⊢ t 2 : ∗<br />

Γ ⊢ t : t 2<br />

In this case Γ ⊢ t t ′ . By IH, Γ ⊢ t ′ : t 1. Since Γ ⊢ t 1<br />

∼ = t2, we<br />

have Γ ⊢ t ′ : t 2.<br />

Case:<br />

Γ ⊢ t : ιx.t ′′<br />

Γ ⊢ t : [t/x]t ′′<br />

In this case Γ ⊢ t t ′ . By IH, Γ ⊢ t ′ : ιx.t ′′ . Thus we have<br />

Γ ⊢ t ′ : [t ′ /x]t ′′ . Since Γ ⊢ t ′ = t, we have Γ ⊢ t ′ : [t/x]t ′′ by<br />

Conv rule.<br />

Case:<br />

Γ ⊢ t : [t/x]t ′′<br />

Γ ⊢ t : ιx.t ′′<br />

Γ ⊢ ιx.t ′′ : ∗<br />

In this case Γ ⊢ t t ′ . By IH, Γ ⊢ t ′ : [t/x]t ′′ . Since<br />

Γ ⊢ [t/x]t ′′ = [t ′ /x]t ′′ , we have Γ ⊢ t ′ : [t ′ /x]t ′′ .Thus we have<br />

Γ ⊢ t ′ : ιx.t ′′ .<br />

Case:<br />

Γ ⊢ t ′ 1 : Πx : t ′′<br />

1 .t ′′<br />

2 Γ ⊢ t ′ 2 : t ′′<br />

1<br />

Γ ⊢ t ′ 1t ′ 2 : [t ′ 2/x]t ′′<br />

2<br />

Suppose Γ ⊢ (λx.t 1)v [v/x]t 1. Then we know Γ ⊢ (λx.t 1)v :<br />

[v/x]t ′′<br />

2 and Γ ⊢ λx.t 1 : Πx : t ′′<br />

1 .t ′′<br />

2 and Γ ⊢ v : t ′′<br />

1 . By<br />

inversion on Γ ⊢ λx.t 1 : Πx : t ′′<br />

1 .t ′′<br />

2 , we have Γ, x : a ⊢ t 1 : b<br />

and Γ ⊢ Πx : a.b λx.t 1<br />

= β,µ,ι,o Πx : t ′′<br />

1 .t ′′<br />

2 . By theorem 34,<br />

we have Γ ⊢ a = β,µ,o t ′′<br />

1 and Γ ⊢ b = β,µ,o t ′′<br />

2 . So we have<br />

Γ, x : a ⊢ t 1 : t ′′<br />

2 and Γ ⊢ v : a. So by lemma 79, we have<br />

Γ ⊢ [v/x]t 1 : [v/x]t ′′<br />

2 , as required.<br />

Suppose Γ ⊢ t 1t 2 t ′ 1t 2, where Γ ⊢ t 1 t ′ 1. We know<br />

Γ ⊢ t 1t 2 : [t 2/x]t ′′<br />

2 and Γ ⊢ t 1 : Πx : t ′′<br />

1 .t ′′<br />

2 and Γ ⊢ t 2 : t ′′<br />

1 .<br />

By IH, we know Γ ⊢ t ′ 1 : Πx : t ′′<br />

1 .t ′′<br />

2 . So Γ ⊢ t ′ 1t 2 : [t 2/x]t ′′<br />

2 .<br />

Suppose Γ ⊢ (λx.t 1)t 2 (λx.t 1)t ′ 2, where Γ ⊢ t 2 t ′ 2. We<br />

know Γ ⊢ (λx.t 1)t 2 : [t 2/x]t ′′<br />

2 and Γ ⊢ λx.t 1 : Πx : t ′′<br />

1 .t ′′<br />

2 and<br />

Γ ⊢ t 2 : t ′′<br />

1 . By IH, we know Γ ⊢ t ′ 2 : t ′′<br />

1 . So Γ ⊢ (λx.t 1)t ′ 2 :<br />

[t ′ 2/x]t ′′<br />

2 . And we know Γ ⊢ [t 2/x]t ′′<br />

2 = [t ′ 2/x]t ′′<br />

2 .


Case:<br />

Γ, ˜µ ⊢ t : t ′ {Γ, ˜µ ⊢ t j : a j} (tj :a j )∈˜µ<br />

Γ ⊢ µt : µt ′<br />

Suppose Γ ⊢ µx j µt j, where x j ↦→ t j ∈ µ. We have<br />

Γ, ˜µ ⊢ x j : t ′ . By inversion, Γ, ˜µ ⊢ x j : a j and Γ, ˜µ ⊢<br />

x j<br />

= β,µ,ι,o t ′ . Since Γ, ˜µ ⊢ x j = t j and by lemma 66, we<br />

a j<br />

get Γ, ˜µ ⊢ a j<br />

t j<br />

= β,µ,ι,o t ′ . Since Γ, ˜µ ⊢ t j : a j, by lemma 65,<br />

Γ, ˜µ ⊢ t j : t ′ . Thus we have Γ ⊢ µt j : µt ′ .<br />

µ([t ′ 2/x]t ′′<br />

2 ) µ(t′ 1 t′ 2<br />

= )<br />

β,µ,ι,o µt ′′ (lemma 69). By lemma 66, we have<br />

(µt ′ 1 )(µt′ 2 )<br />

Γ ⊢ [µt ′ 2/x]µt ′′<br />

2 = β,µ,ι,o µt ′′ . So Γ ⊢ (µt ′ 1)(µt ′ 2) :<br />

[µt ′ 2/x]µt ′′<br />

2 and then Γ ⊢ (µt ′ 1)(µt ′ 2) : µt ′′ (lemma 65).<br />

Suppose Γ ⊢ µ⃗µ(t ′ 1t ′ 2) (µ⃗µt ′ 1)(µ⃗µt ′ 2), we argue similar as the<br />

case for Γ ⊢ µ⃗µλx.t λx.µ⃗µt.<br />

Suppose Γ ⊢ µ⃗µx j µ⃗µt j, where x j ↦→ t j ∈ µ j. By inversion<br />

on Γ, ˜µ ⊢ ⃗µx j : t ′ , we have Γ, ˜µ, ˜⃗µ ⊢ x j : t a and Γ, ˜µ ⊢<br />

⃗µx j<br />

⃗µt a = β,µ,ι,o t ′ . By inversion on Γ, ˜µ, ˜⃗µ ⊢ x j : t a, we have<br />

Γ, ˜µ, ˜⃗µ ⊢ x j : b, where (x j : b) ∈ ˜µ ∪ ˜⃗µ and Γ, ˜µ, ˜⃗µ ⊢<br />

b x j<br />

= β,µ,ι,o t a. So Γ ⊢ µ⃗µb µ⃗µx j<br />

µ⃗µx j<br />

= β,µ,ι,o µ⃗µt a = β,µ,ι,o µt ′ .<br />

Since Γ ⊢ µ⃗µt j : µ⃗µb and Γ ⊢ µ⃗µx j = µ⃗µt j, so Γ ⊢ µ⃗µt j : µt ′ .<br />

Suppose Γ ⊢ µ∗ ∗. We have Γ, ˜µ ⊢ ∗ : t ′′ . We have<br />

Γ, ˜µ ⊢ ∗ = ∗ β,µ,ι,o t ′′ (by inversion). Thus we have Γ ⊢ µ∗ µ∗<br />

= β,µ,ι,o<br />

µt ′′ (lemma 69). We also know that Γ ⊢ ∗ : ∗ and Γ ⊢ µ∗ = ∗.<br />

So we have Γ ⊢ ∗ µ∗<br />

= β,µ,ι,o µt ′′ . Thus Γ ⊢ ∗ = ∗ β,µ,ι,o µt ′′ . So<br />

Γ ⊢ ∗ : µt ′′ (lemma 65).<br />

Suppose Γ ⊢ µ⃗µ∗ ∗. We argue similarly.<br />

Suppose Γ ⊢ µx x, where x /∈ dom(µ). We have Γ, ˜µ ⊢ x : t ′′ .<br />

We have Γ, ˜µ ⊢ a = x β,µ,ι,o t ′′ , where x : a ∈ Γ(by inversion).<br />

Thus we have Γ ⊢ µa µx<br />

= β,µ,ι,o µt ′′ (lemma 69). We also know<br />

that Γ ⊢ x : a and Γ ⊢ µx = x and Γ ⊢ µa = a. Thus<br />

Γ ⊢ a = x β,µ,ι,o µt ′′ . So Γ ⊢ x : µt ′′ (lemma 65).<br />

Suppose Γ ⊢ µ⃗µx x, where x /∈ dom(µ) ∪ dom(⃗µ). By<br />

inversion on Γ, ˜µ ⊢ ⃗µx : t ′ , we have Γ, ˜µ, ˜⃗µ ⊢ x : t a, where<br />

⃗µx<br />

Γ, ˜µ ⊢ ⃗µt a =β,µ,ι,o t ′ . By inversion on Γ, ˜µ, ˜⃗µ ⊢ x : t a, we have<br />

x : b ∈ Γ and Γ, ˜µ, ˜⃗µ ⊢ b = x β,µ,ι,o t a. So Γ, µ ⊢ ⃗µb ⃗µx<br />

= β,µ,ι,o<br />

⃗µx<br />

⃗µt a =β,µ,ι,o t ′ . So Γ ⊢ b = x β,µ,ι,o µt ′ . Thus Γ ⊢ x : µt ′ .<br />

Suppose Γ ⊢ µλx.t λx.µt. We have Γ, ˜µ ⊢ λx.t : t ′′ and<br />

Γ, ˜µ, x : t ′′<br />

1 ⊢ t : t ′′<br />

2 and Γ, ˜µ ⊢ Πx : t ′′<br />

1 .t ′′ λx.t<br />

2 = β,µ,ι,o t ′′ (by<br />

inversion). Thus we have Γ, x : µt ′′<br />

1 ⊢ µt : µt ′′<br />

2 (lemma 78) and<br />

Γ ⊢ µ(Πx : t ′′<br />

1 .t ′′<br />

2 ) µλx.t<br />

= β,µ,ι,o µt ′′ (lemma 69). By lemma 66,<br />

Γ ⊢ (Πx : µt ′′<br />

1 .µt ′′<br />

2 ) λx.µt<br />

= β,µ,ι,o µt ′′ . Also, Γ ⊢ λx.µt : Πx :<br />

(µt ′′<br />

1 ).(µt ′′<br />

2 ). So by lemma 65, Γ ⊢ λx.µt : µt ′′ .<br />

Suppose Γ ⊢ µ⃗µλx.t λx.µ⃗µt. By inversion on Γ, ˜µ ⊢<br />

⃗µ(λx.t) : t ′ , we have Γ, ˜µ, ˜⃗µ ⊢ λx.t : t a, where Γ, ˜µ ⊢<br />

⃗µ(λx.t)<br />

⃗µt a = β,µ,ι,o t ′ . By inversion on Γ, ˜µ, ˜⃗µ ⊢ λx.t : t a, then we<br />

have Γ, ˜µ, ˜⃗µ, x : t ′′<br />

1 ⊢ t : t ′′<br />

2 and Γ, ˜µ, ˜⃗µ ⊢ Πx : t ′′<br />

1 .t ′′<br />

2<br />

t a. So Γ, ˜µ ⊢ Πx : ⃗µt ′′<br />

1 .⃗µt ′′ ⃗µ(λx.t)<br />

⃗µ(λx.t)<br />

2 = β,µ,ι,o ⃗µt a<br />

µ⃗µ(λx.t)<br />

λx.t<br />

= β,µ,ι,o<br />

= β,µ,ι,o t ′ .<br />

Thus Γ ⊢ Πx : µ⃗µt ′′<br />

1 .µ⃗µt ′′<br />

2 = β,µ,ι,o µt ′ . Since Γ ⊢<br />

λx.µ⃗µt : Πx : µ⃗µt ′′<br />

1 .µ⃗µt ′′<br />

2 , we have Γ ⊢ λx.µ⃗µt : µt ′ .<br />

Suppose Γ ⊢ µ(t ′ 1t ′ 2) (µt ′ 1)(µt ′ 2). We have Γ, ˜µ ⊢ t ′ 1t ′ 2 :<br />

t ′′ . We have Γ, ˜µ ⊢ t ′ 1 : Πx : t ′′<br />

1 .t ′′<br />

2 and Γ, ˜µ ⊢ t ′ 2 : t ′′<br />

1<br />

t ′ 1 t′ 2<br />

and Γ, ˜µ ⊢ [t ′ 2/x]t ′′<br />

2 = β,µ,ι,o t ′′ (by inversion). Thus we have<br />

Γ ⊢ µt ′ 1 : µ(Πx : t ′′<br />

1 .t ′′<br />

2 ) and Γ ⊢ µt ′ 2 : µt ′′<br />

1 and Γ ⊢

Hooray! Your file is uploaded and ready to be published.

Saved successfully!

Ooh no, something went wrong!