• Aucun résultat trouvé

The previous sections provide a theoretical basis for compiling high-level languages to typed assembly language. In this section we discuss some issues that arise when putting this technology into practice.

Texpγ,∆,Γ[[letx=uτ ine]] def= hH,(movr,Tvalγ[[uτ]];I)i

wherehH, Ii = Tγ{x7→r},∆,Γ{r:T[[τ]]}

exp [[e]]

ris fresh

Texpγ,∆,Γ[[letx=πi(u1ϕ1,...,τnϕni)ine]] def= hH,(movr,Tvalγ[[u1ϕ1,...,τnϕni]];

ldr, r[i1];I)i

wherehH, Ii = Tγ{x7→r},∆,Γ{r:T[[τi]]}

exp [[e]]

ris fresh Texpγ,∆,Γ[[letx=v1p v2ine]] def= hH,(movr,Tvalγ[[v1]];

arithpr, r,Tvalγ[[v2]];I)i

wherehH, Ii = Tγ{x7→r},∆,Γ{r:int}

exp [[e]]

arith+ = add arith = sub arith× = mul ris fresh

Texpγ,∆,Γ[[let[α, x] =unpacku∃α.τine]] def= hH,(unpack[α, r],Tvalγ[[u∃α.τ]];I)i

wherehH, Ii = Texpγ{x7→r},∆{α},Γ{r:T[[τ]]}[[e]]

α, rare fresh

Texpγ,∆,Γ[[letx=malloc[τ1, . . . , τn]ine]] def= hH,(mallocr[T[[τ1]], . . . ,T[[τn]]];I)i where

hH, Ii = Tγ{x7→r},∆,Γ{r:hT[[τ1]]0,...,T[[τn]]0i}

exp [[e]]

ris fresh

Texpγ,∆,Γ[[letx=u1ϕ1,...,τnϕni[i]vine]] def= hH,(movr,Tvalγ[[u1ϕ1,...,τnϕni]];

movr0,Tvalγ[[v]];

str[i1], r0;I)i where

hH, Ii = Tγ{x7→r},∆,Γ0

exp [[e]]

Γ0 = Γ{r:T[[1ϕ1, . . . , τi−ϕi−11 , τi1, τi+1ϕi+1, . . . , τnϕni]]} r, r0are fresh

Texpγ,∆,Γ[[v(v1, . . . , vn)]] def= h∅,(movr00,Tvalγ[[v]];

movr10,Tvalγ[[v1]];. . .; movrn0,Tvalγ[[vn]];

mov r1, r01;. . . mov rn, rn0; jmpr00)i

wherer0iare fresh andr0i6∈ {r1, . . . ,rn}

Texpγ,∆,Γ[[if0(v, e1, e2)]] def= hH1H2{`7→h},(movr,Tvalγ[[v]];

bnzr, `[∆];I1)i wherehHi, Iii = Texpγ,∆,Γ[[ei]]

h = code[∆]Γ.I2

`, rare fresh Texpγ,∆,Γ[[halt[τ]v]] def= h∅,(mov r1,Tvalγ[[v]];

halt[T[[τ]]])i

Fig. 20. Translation of expressions fromλAto TAL.

(H,{}, I) where

H =

l fact:

code[ ]{r1:hi,r2:int,r3:τk}.

bnz r2,l nonzero

unpack [α,r3],r3 % zero branch: callk(inr3) with 1 ld r4,r3[0] % project kcode

ld r1,r3[1] % project kenvironment mov r2,1

jmp r4 % jump with{r1= env,r2= 1}

l nonzero:

code[ ]{r1:hi,r2:int,r3:τk}.

sub r4,r2,1 % n−1

malloc r5[int, τk] % create environment for cont inr5 st r5[0],r2 % storeninto environment

st r5[1],r3 % storekinto environment

malloc r3[∀[ ].{r1:hint1, τk1i,r2:int},hint1, τk1i] % create cont closure inr3 mov r2,l cont

st r3[0],r2 % store cont code

st r3[1],r5 % store environmenthn, ki

mov r2,r4 % arg :=n−1

mov r3,pack[hint1, τk1i,r3]asτk % abstract the type of the environment jmp l fact % callfact(env, n−1,l cont)

l cont:

code[ ]{r1:hint1, τk1i,r2:int}. % r2contains (n−1)!

ld r3,r1[0] % retrieven

ld r4,r1[1] % retrievek

mul r2,r3,r2 % (n1)!

unpack [α,r4],r4 % unpackk ld r3,r4[0] % project kcode ld r1,r4[1] % project kenvironment

jmp r3 % jump tokwith{r1= env,r2=n!} l halt:

code[ ]{r1:hi,r2:int}.

mov r1,r2

halt[int] % halt with result inr1 and I =

malloc r1[ ] % create an empty environment (hi) malloc r2[ ] % create another empty environment malloc r3[∀[ ].{r1:hi,r2:int},hi] % create halt closure inr3

mov r4,l halt

st r3[0],r4 % store halt code

st r3[1],r2 % store halt environmenthi

mov r2,6 % load argument (6)

mov r3,pack[hi,r3]asτk % abstract the type of the environment jmp l fact % callfact(hi,6,l halt)

and τk = ∃α.h∀[ ].{r1:α,r2:int}1, α1i

Fig. 21. Typed assembly code for factorial.

8.1 Implementation

In order to investigate the applicability of our approach to realistic modern pro-gramming languages, we have implemented a version of TAL for the Intel 32-bit Architecture (IA32) [Intel 1996], and have compilers for a number of different source languages including a safe C-like language [Morrisett et al. 1999] and a higher-order, dynamically typed language (a subset of Scheme). Compilers for Standard ML and a small object-oriented language are also in development.

TALx86, the target language for our compilers, is a strongly typed version of the IA32 assembly language. Our type checker verifies standard MASM assembly code in which type annotations and complex instructions such as malloc are assembly language macros. The MASM assembler processes this annotated code as it would any other assembly code, expanding the instruction macros as their definitions dictate and erasing types as it translates the assembly code into concrete machine instructions. We have also implemented our own assembler and are extending it to produced typed object files. Such typed object files include code/data segments and a type segment similar to Necula and Lee’s code and proof segments in their PCC binaries [Necula 1997]. We have implemented a tool that reads TALx86 files, type checks them, and assembles them into object files or invokes MASM.

The TALx86 type system is based on the type system described in this article but enriched with a variety of standard constructs including floats, sums, arrays, references, recursive types, and higher-order type constructors. In order to deal with floating-point values correctly in the presence of polymorphism, we use a kind structure that distinguishes types of objects that are 32 bits wide (such as pointers and integers) from types of objects possibly of other sizes. If a polymorphic type variableα has the 32-bit kind, then objects of type α can be passed in general-purpose registers, and tuple offsets may be computable for fields appearing after a field of typeα. If, on the other hand,αhas the more general kind “Type,” the type checker cannot tell how large objects of type αare, and these operations are disallowed.

To support separate compilation and type-safe linking, we have also augmented our typed assembly language with a module system [Glew and Morrisett 1999]. A TAL interface file specifies the types and terms that a TAL implementation file defines. The types may either be opaque to support information hiding and mod-ularity, or transparent to allow information sharing and admit some cross-module optimizations. Our system performs a series of checks to ensure that implementa-tions are well formed and that their interfaces are compatible and complete. Once interface compatibility and completeness have been verified, we assemble the code as described above and invoke a standard untyped linker.

To deal with the creation and examination of exception packets TALx86 includes a type-tagging mechanism [Glew 1999]. The basic idea is that freshly created heap pointers may be associated with a type, and that a tag for an unknown typeαcan be tested against a tag for a known typeτ. If the test succeeds, the unknown type is refined to the known type. Using these tags, we implement an exception packet as an existentially packaged pair containing a tag of the hidden type (serving as the exception name) and a value of that type.

TALx86 also contains some support for compiling objects. The type system

has a more general notion of subtyping than this article that includes the usual contravariant rule for code, right-extension and depth subtyping for tuples, and a variance mechanism for arrays and references. Furthermore, the type-tagging mech-anism can also be used to tag objects with their class. This mechmech-anism provides us with a way to implement down-casting. However, while TALx86 contains the necessary constructs to admit some simple object encodings, we are still developing the theoretical and practical tools we need to admit efficient object encodings.

Although this article describes a CPS-based compiler, all of the compilers we have built use a stack-based compilation model. Both standard continuations and exceptions are allocated on the stack, and the stack is also used to store spilled temporary values. The details of our stack typing discipline are discussed in Mor-risett et al. [1998]. The primary mechanisms are as follows. The size of the stack and the types of its contents are specified bystack types,and code blocks indicate stack types describing the state of the stack they expect. Since code is typically expected to work with stacks of varying size, functions may quantify over stack type variables, resulting in stack polymorphism. The combination of stack types and our register typing discipline allows us to model almost any standard calling conven-tion. Arguments, results, and continuations (or return addresses) may be placed in registers, on the stack, or both, and the implementer may specialize conventions for known functions for better register allocation.

Real machines also have a finite amount of heap space. It is straightforward to link TALx86 to a conservative garbage collector [Boehm and Weiser 1988] and reclaim dead heap values. It is worth noting that our use of conservative collection is sound. Conservative collectors make assumptions about the way pointers can be used in programs that untyped assembly language programs can violate. However, the TAL type system guarantees that these assumptions do hold because labels are a strong abstraction; labels cannot be synthesized out of integers, and operations like pointer arithmetic are disallowed. TAL guarantees that other GC constraints hold because values that disobey the constraints cannot be constructed. For example, TAL disallows pointers into the middle of objects and ensures alignment constraints are obeyed.

Support for an accurate collector would require introducing tags so that we may distinguish pointers from integers, or else require a type-passing interpretation [Tol-mach 1994; Morrisett and Harper 1997]. In the former case, we would have to ensure that all values are properly tagged and fully initialized at every potential garbage collection site. We believe that it is feasible to devise a type system to ensure these constraints, but we have not seriously investigated this option.

8.2 Future Work

There remain several directions in which TAL could be improved. One of the most important is to make array manipulation efficient. In order to ensure safe access to arrays, TALx86 uses complex instructions (which expand into three real instructions) that perform subscript and update operations after checking that the array offset is in bounds. These bounds checks cannot be eliminated in the cur-rent TAL framework. As a result, array-intensive applications will suffer the same performance penalties that they do in Java just-in-time compilers where there is no time to perform analyses to eliminate the checks. However, Xi [1999] and Xi

and Pfenning [1998; 1999] have shown how to eliminate array bounds checks ef-fectively using dependent types. TALx86 can be extended with similar constructs.

We have implemented a prototype version in which these checks can be eliminated, but we have not yet added compiler support for generating code with unchecked array subscripts.

Another important direction is to augment our compiler with data-layout opti-mizations such as those used in the TIL compiler [Tarditi et al. 1996]. As discussed in Section 5, such optimizations require programs to have the ability to analyze types at run-time, which is not directly compatible with the type-erasure interpre-tation adopted here. To make such optimizations permissible, we are augmenting the TALx86 language so that TAL programs can construct values that represent types and analyze those values when necessary, following the work of Crary et al. [1998; 1999].

Although we believe our translations are operationally correct, we are still search-ing for robust proofs of correctness. Similar CPS [Danvy and Filinski 1992] and closure conversion [Minamide et al. 1996] translations have already been proven cor-rect, but these results do not easily extend to languages that include recursive types or objects. The principal problem is that these arguments are based on inductively defined, type-indexed logical relations between source and target language terms.

Extending this framework so that it supports recursive types or objects is difficult because the relations can no longer be constructed in a simple inductive fashion. A syntactic proof of correctness seems possible (we have constructed such arguments for the CPS and closure conversion phases), but the proofs are overly specific to the details of the translation. Moreover, security-conscious applications might require translations that are not only operationally correct but also fully abstract. We hope further research on the proof theory of similar systems will eventually allow us to construct these arguments.

Other avenues of future research include extension of our type system to the same level of generality as PCC through the use of a dependent type theory, an investigation of the support required to compile Java classes and objects into TAL, and an exploration of type-theoretic mechanisms for performing explicit memory management.

9. SUMMARY

We have given a compiler from System F to a statically typed assembly language.

The type system for the assembly language ensures that source-level abstractions such as closures and polymorphic functions are enforced at the machine-code level.

Furthermore, although the type system may preclude some advanced optimiza-tions, many common compiler-introduced, low-level optimizaoptimiza-tions, such as register allocation, instruction selection, or instruction scheduling, are largely unaffected.

Furthermore, programmers concerned with efficiency can hand-code routines in as-sembly, as long as the resulting code passes the type checker. Consequently, TAL provides a foundation for high-performance computing in environments where un-trusted code must be checked for safety before being executed.

APPENDIX

Lemma (Context Strengthening). If0 then

(1) If`TALτ then0 `TALτ

(2) If`TALτ1≤τ2 then0`TALτ1≤τ2.

Proof. Part 1 is immediate by (type). Part 2 is by induction on derivations.

Lemma (Subtyping Regularity). If`TAL τ τ0 then`TAL τ and

`TALτ0.

Proof. By induction on derivations.

Lemma (Heap Extension). If `TALH : Ψ,∅ `TALτ,Ψ{`:τ} `TALh:τ hval, and `6∈H then

(1) `TALΨ{`:τ}

(2) `TALH{`7→h}: Ψ{`:τ}

(3) IfΨ`TALR: ΓthenΨ{`:τ} `TALR: Γ (4) IfΨ; ∆; Γ`TALI then Ψ{`:τ}; ∆; Γ`TALI (5) IfΨ`TALh:σ hvalthenΨ{`:τ} `TALh:σ hval (6) IfΨ; ∆`TALw:σϕ thenΨ{`:τ}; ∆`TALw:σϕ

(7) IfΨ; ∆`TALw:σwvalthenΨ{`:τ}; ∆`TALw:σwval (8) IfΨ; ∆; Γ`TALv:σ thenΨ{`:τ}; ∆; Γ`TALv:σ.

Proof. Part 1 is immediate by (htype). Part 2 follows from parts 1 and 5. Parts 3–8 are by induction on derivations.

Lemma (Heap Update). If `TAL H : Ψ,∅ `TAL τ Ψ(`), and Ψ{` :τ} `TAL

h:τ then

(1) `TALΨ{`:τ}

(2) `TALH{`7→h}: Ψ{`:τ}

(3) IfΨ`TALR: ΓthenΨ{`:τ} `TALR: Γ (4) IfΨ; ∆; Γ`TALI then Ψ{`:τ}; ∆; Γ`TALI (5) IfΨ`TALh:σ hvalthenΨ{`:τ} `TALh:σ hval (6) IfΨ; ∆`TALw:σϕ thenΨ{`:τ}; ∆`TALw:σϕ

(7) IfΨ; ∆`TALw:σwvalthenΨ{`:τ}; ∆`TALw:σwval (8) IfΨ; ∆; Γ`TALv:σ thenΨ{`:τ}; ∆; Γ`TALv:σ.

Proof. Part 1 is immediate by (htype) and Subtyping Regularity. Part 2 follows from parts 1 and 5. Parts 3–8 are by induction on derivations. The only interesting case is the case for the rule (label). The derivation must end with the following:

`TALσ0 ≤σ

Ψ; ∆`TAL`0 :σwval (Ψ(`0) =σ0)

If ` 6= `0 then clearly the inference also holds for Ψ{` : τ}. Suppose ` = `0. By hypothesis and Context Strengthening, we deduce ∆ `TAL τ σ0. Then the conclusion may be proven with the (trans) rule, as follows:

`TALτ ≤σ0`TALσ0 ≤σ

`TALτ≤σ

Ψ{`:τ}; ∆`TAL`:σwval (Ψ{`:τ}(`) =τ)

Lemma (Register File Update). If Ψ`TAL R : Γ and Ψ;∅ `TAL w :τ wval thenΨ`TALR{r7→w}: Γ{r:τ}.

Proof. SupposeRis{r17→w1, . . . , rn7→wn}and Γ is{r17→τ1, . . . , rm7→τm} wherermay or may not be in{r1, . . . , rn}. Since Ψ`TALR: Γ, by the rule (reg) it must be the case thatn≥mand Ψ;∅ `TALwi:τiwval (for all 1≤i≤nand some τm+1, . . . , τn). So certainly fori such thatri 6=r, we have Ψ;∅ `TAL wi :τi wval, and by hypothesis we have Ψ;∅ `TAL w :τ wval; so by rule (reg) Ψ`TAL R{r 7→

w}: Γ{r7→τ}.

Lemma (Canonical Heap Forms). IfΨ`TALh:τ hvalthen (1) Ifτ =[α~].Γ then

(a) h=code[.I (b) Ψ;α~; Γ`TALI.

(2) Ifτ =0ϕ0, . . . , τn−ϕn−11 ithen (a) h=hw0, . . . , wn−1i (b) Ψ;∅ `TALwi:τiϕi. Proof. By inspection.

Lemma (Canonical Word Forms). If `TALH : Ψand Ψ;∅ `TALw: τ wval then

(1) Ifτ =int then w=i. (2) Ifτ =[β1, . . . , βm].Γthen

(a) w=`[σ1, . . . , σn]

(b) H(`) =code[α1, . . . , αn, β1, . . . , βm0.I (c) Γ = Γ0[~σ/~α]

(d) Ψ;α1, . . . , αn, β1, . . . , βm; Γ0 `TALI. (3) Ifτ =0ϕ0, . . . , τn−ϕn−11 ithen

(a) w=`

(b) H(`) =hw0, . . . , wn−1i (c) Ψ;∅ `TALwi:τiϕi.

(4) Ifτ =∃α.τ thenw=pack[τ0, w0]as∃α.τ and Ψ;∅ `TALw0 :τ[τ0] wval.

Proof. (1) By inspection.

(2) By induction on the derivation of Ψ;∅ `TAL w : τ wval: The derivation must end with either the (label) or the (tapp-word) rule. Suppose the former. Then we havew=`, Ψ(`) =τ0, and∅ `TALτ0 ≤ ∀[~β].Γ. Inspection of the subtyping rules then reveals that τ0 = [~β].Γ. Since `TAL H : Ψ, we may deduce that Ψ`TALH(`) :[~β].Γ hval. The conclusion follows by Canonical Heap Forms.

Alternatively, suppose the derivation ends with (tapp-word). Thenw=w0[σ] and Ψ;∅ `TALw0:[α, ~β].Γ0 wval with Γ = Γ0[σ/α]. The conclusion follows by induction.

(3) The derivation Ψ;∅ `TALw:τ wval must be shown by use of the (label) rule.

Thus, w = `, Ψ(`) = τ0, and ∅ `TAL τ0 ≤ hτ0ϕ0, . . . , τn−ϕn−11 i. Let us say that ϕ≤ϕ and 10. Then inspection of the subtype rules reveals that τ0 must

be of the form0ϕ0, . . . , τn−n−11 i with ϕ0i ≤ϕi (for each 0≤i≤n−1). Since

`TAL H : Ψ, we may deduce that Ψ `TAL H(`) :0ϕ00, . . . , τn−ϕ0n−11 i hval. Thus H(`) =hw0, . . . , wn−1i and Ψ;∅ `TAL wi : τiϕ0i by Canonical Heap Forms. It remains to show that Ψ;∅ `TAL wi:τiϕi for all 0≤i≤n−1. Supposeϕ0i= 1 andϕi= 0 (otherwise the conclusion is immediate). Then Ψ;∅ `TALwi:τi1 is shown by the (init) rule, which also permits the deduction of Ψ;∅ `TALwi:τi0. (4) By inspection.

Lemma (Rˆ Typing). If Ψ `TAL R : Γ and Ψ;; Γ `TAL v : τ then Ψ;∅ `TAL

Rˆ(v) :τ wval.

Proof. The proof is by induction on the syntax ofv. Consider the cases forv: Case v=w. Immediate.

Case v=r. The only rule that can type v is (reg-val), and this rule requires τ = Γ(r). The only rule that can typeR is (reg), and this rule requires Ψ;∅ `TAL

R(r) :τ wval. The conclusion follows, since ˆR(r) =R(r).

Case v=v0[σ]. The only rule that can typev is (tapp-val), soτ=[β~].Γ0[σ/α] and Ψ;; Γ`TALv0:[α, ~β].Γ0. By induction we deduce Ψ;∅ `TAL Rˆ(v0) :[α, ~β].Γ0 wval, and then the rule (tapp-word) proves Ψ;∅ `TAL Rˆ(v0)[σ] :[β~].Γ0[σ/α] wval.

The result follows, since ˆR(v0[σ]) = ˆR(v0)[σ].

Case v=pack[σ, v0]as∃α.τ0. The only rule that can type v is (pack-val), so τ =∃α.τ0 and Ψ;; Γ`TAL v0 :τ0[σ/α]. By induction we deduce Ψ;∅ `TAL Rˆ(v0) : τ0[σ/α] wval, and then the rule (pack-word) proves Ψ;∅ `TAL pack[σ,Rˆ(v0)] as

∃α.τ0:∃α.τ0wval. The result follows, since ˆR(pack[σ, v0]as∃α.τ0) =pack[σ,Rˆ(v0)]

as∃α.τ0.

Lemma (Canonical Forms). If `TALH: Ψ hval,Ψ`TALR: Γ, and Ψ;; Γ`TALv:τ then

(1) Ifτ =int then Rˆ(v) =i. (2) Ifτ =[β1, . . . , βm].Γthen

(a) ˆR(v) =`[σ1, . . . , σn]

(b) H(`) =code[α1, . . . , αn, β1, . . . , βm0.I (c) Γ = Γ0[~σ/~α]

(d) Ψ;α1, . . . , αn, β1, . . . , βm; Γ0 `TALI. (3) Ifτ =0ϕ0, . . . , τn−ϕn−11 ithen

(a) ˆR(v) =`

(b) H(`) =hw0, . . . , wn−1i (c) Ψ;∅ `TALwi:τiϕi.

(4) Ifτ =∃α.τ thenRˆ(v) =pack[τ0, w]as∃α.τ and Ψ;∅ `TALw:τ[τ0] wval.

Proof. Immediate from ˆRTyping and Canonical Word Forms.

Lemma (Type Substitution). If β~`TALτi then (1) IfΨ;~α, ~β; Γ`TALI thenΨ;β~; Γ[~τ/~α]`TALI[~τ/~α]

(2) IfΨ;~α, ~β; Γ`TALv:τ thenΨ;β~; Γ[~τ/~α]`TALv[~τ/~α] :τ[~τ/~α]

(3) IfΨ;~α, ~β `TALw:τ wvalthenΨ;β~`TALw[~τ/~α] :τ[~τ/~α] wval (4) Ifα, ~~ β`TALΓ1Γ2 then ~β`TALΓ1[~τ/~α]Γ2[~τ/~α]

(5) Ifα, ~~ β`TALτ1≤τ2 thenβ~`TALτ1[~τ/~α]≤τ2[~τ/~α] (6) Ifα, ~~ β`TALτ thenβ~`TALτ[~τ/~α].

Proof. By induction on derivations. The only interesting case is the case for the rule (type):

FTV(τ)⊆ {~α, ~β}

~

α, ~β `TALτ

The hypothesis must also be proven with the rule (type), so FTV(τi) ⊆ {β}~ . Consequently

FTV(τ[~τ/~α]) FTV(τ)\ {~α} ∪([

iFTV(τi))

⊆ {~α, ~β} \ {~α} ∪ {β}~

= {β}.~

Hence we may proveβ~ `TALτ[~τ/~α] using the (type) rule.

Lemma (Register File Weakening). If`TAL Γ1Γ2 and Ψ; ∆`TAL R: Γ1 then Ψ; ∆`TALR: Γ2.

Proof. By inspection of the rules (weaken) and (reg).

Theorem (Subject Reduction). If `TALP and P 7−→P0 then`TALP0. Proof. P has the form (H, R, ι;I) or (H, R,jmpv). LetT D be the derivation of`TALP. Consider the following cases forjmporι:

Case jmp. T D has the form

`TALH: Ψ Ψ`TALR: Γ

Ψ;; Γ`TALv:[ ].Γ0 ∅ `TALΓΓ0 Ψ;; Γ`TALjmpv

`TAL P .

By the operational semantics, P0 = (H, R, I[~σ/~α]) where ˆR(v) =`[] andH(`) = code[00.I. Then

(1) `TALH : Ψ is inT D.

(2) From∅ `TAL Γ Γ0 and Ψ`TAL R : Γ it follows by Register File Weakening that Ψ`TALR: Γ0.

(3) By Canonical Forms it follows from Ψ;; Γ`TAL v : [ ].Γ0 that Γ0 = Γ00[~σ/~α] and Ψ;; Γ00`TALI. By Type Substitution we conclude Ψ;; Γ0`TALI[~σ/~α].

Case add,mul,sub. T D has the form

`TALH : Ψ Ψ`TALR: Γ

Ψ;; Γ`TALrs:int Ψ;; Γ`TALv:int Ψ;; Γ0`TALI

Ψ;; Γ`TALarithprd, rs, v;I

`TALP

where Γ0 = Γ{rd : int}. By the operational semantics, P0 = (H, R0, I) where R0=R{rd7→R(rs)pRˆ(v)}. Then

(1) `TALH : Ψ is inT D.

(2) By Canonical Forms it follows that R(rs) and ˆR(v) are integer literals, and therefore Ψ;∅ `TAL R(rs)pRˆ(v) : int wval. We conclude Ψ `TAL R0 : Γ0 by Register File Update.

(3) Ψ;; Γ0`TALI is inT D. Case bnz. T D has the form

`TALH : Ψ Ψ`TALR: Γ

Ψ;; Γ`TALr:int Ψ;; Γ`TALv:[ ].Γ0

∅ `TALΓΓ0 Ψ;; Γ`TALI Ψ;; Γ`TAL bnzr, v;I

`TALP .

IfR(r) = 0 thenP0= (H, R, I) and`TAL P0 follows, since Ψ;; Γ`TALI is inT D. Otherwise the reasoning is exactly as in the case forjmp.

Case ld. T Dhas the form

`TALH : Ψ Ψ`TALR: Γ

0≤i≤n−1 ϕi= 1 Ψ;; Γ`TALrs:0ϕ0, . . . , τn−ϕn−11 i

Ψ;; Γ0`TALI Ψ;; Γ`TALldrd, rs[i];I

`TAL P

where Γ0 = Γ{rd :τi}. By the operational semantics,P0 = (H, R0, I) where R0 = R{rd7→wi},R(rs) =`, H(`) =hw0, . . . , wm−1}, and 0≤i < m. Then

(1) `TALH : Ψ is inT D.

(2) By Canonical Forms it follows from Ψ;; Γ `TAL rs : 0ϕ0, . . . , τn−ϕn−11 i that m=nand Ψ;∅ `TALwj:τjϕj for 0≤j < n. Since ϕi = 1 it must be the case (by inspection of the (init) rule) that Ψ;∅ `TAL wi:τi wval. By Register File we conclude Ψ`TALR0: Γ0.

(3) Ψ;; Γ0`TALI is inT D. Case malloc. T Dhas the form

`TALH : Ψ Ψ`TALR: Γ

∅ `TALτi Ψ;; Γ0`TALI Ψ;; Γ`TALmallocrd[τ1, . . . , τn];I

`TALP

where σ = 10, . . . , τn0i, Ψ0 = Ψ{` : σ}, and Γ0 = Γ{rd : σ}. By the operational semantics, P0 = (H0, R0, I) where H0 =H{` 7→ h?τ1, . . . ,?τni}, R0 =R{rd 7→`}, and`6∈H. Then

(1) By the (tuple) and (uninit) rules we may deduce Ψ0`TALh?τ1, . . . ,?τni:σhval.

By Heap Extension it follows that `TALH0: Ψ0.

(2) By the (type), (reflex), and (label) rules we may deduce that Ψ0;∅ `TAL ` : σ wval. By Heap Extension we deduce that Ψ0 `TAL R : Γ, and it follows by Register File Update that Ψ0 `TALR0 : Γ0.

(3) By Heap Extension, Ψ0;; Γ0`TALI. Case mov. T D has the form

`TALH : Ψ Ψ`TALR: Γ

Ψ;; Γ`TALv:τ Ψ;; Γ0`TALI Ψ;; Γ`TALmovr, v;I

`TALP

where Γ0 = Γ{r : τ}. By the operational semantics, P0 = (H, R0, I) where R0 = R{r7→Rˆ(v)}. Then

(1) `TALH : Ψ is inT D.

(2) By ˆR Typing it follows from Ψ;; Γ `TAL v : τ that Ψ;∅ `TAL Rˆ(v) :τ wval.

Using Register File Update we conclude that Ψ`TALR0: Γ0. (3) Ψ;; Γ0`TALI is inT D.

Case st. T Dhas the form

`TALH: Ψ Ψ`TALR: Γ

0≤i≤n−1 Ψ;; Γ`TALrd:σ0

Ψ;; Γ`TALrs:τi Ψ;; Γ0`TALI Ψ;; Γ`TALstrd[i], rs;I

`TALP where

σ0 = 0ϕ0, . . . , τn−ϕn−11 i

σ1 = 0ϕ0, . . . , τi−ϕi−11 , τi1, τiϕ+1i+1, . . . , τn−ϕn−11 i Γ0 = Γ{rd:σ1}.

By the operational semantics,P0= (H0, R, I) where

H0=H{`7→ hw0, . . . , wi−1, R(rs), wi+1, . . . , wm−1i}

andR(rd) =`,H(`) =hw0, . . . , wmi, and 0≤i < m. Then

(1) Since Ψ;; Γ`TAL rd :σ0, it must be the case that Γ(rd) =σ0, and thus since Ψ`TALR : Γ andR(rd) =` we may deduce Ψ;∅ `TAL `:σ0wval. The latter judgment must be proven with the (label) rule; hence ∅ `TAL σ00 ≤σ0 where Ψ(`) =σ00. Note that it follows from Subtyping Regularity and the definition ofσ0 that∅ `TALτj for each 0≤j < n.

Let us say thatϕ≤ϕand 10. Inspection of the subtyping rules reveals that σ00 must be of the form0ϕ00, . . . , τϕ

0n−1

n−1 iwithϕ0j≤ϕj. Let σ01=0ϕ00, . . . , τϕ

0i−1

i−1 , τi1, τϕ

0i+1

i+1 , . . . , τϕ

0n−1

n−1 i.

Then ∅ `TALσ01 ≤σ00 and ∅ `TAL σ01 ≤σ1. Since `TALH : Ψ, we may deduce that m =n and Ψ;∅ `TAL wj : τϕ

0j

j for 0 j < n. Let Ψ0 = Ψ{`: σ01}. By Heap Update it follows that Ψ0;∅ `TAL wj :τϕ

0j

j .

Using ˆR Typing and Heap Update, we may deduce that Ψ0;∅ `TAL R(rs) : τi wval, and by applying the (init) and (tuple) rules we may conclude that

Ψ0`TALhw0, . . . , wi−1, R(rs), wi+1, . . . , wm−1i:σ01 hval. Hence `TALH0: Ψ0 by Heap Update.

(2) By Heap Update we may deduce that Ψ0 `TALR: Γ. Recall that∅ `TAL σ10 σ1. Thus, Ψ0;∅ `TAL`:σ1wval, and by Register File Update we may conclude that Ψ0 `TALR: Γ0 (since R=R{rd7→`}).

(3) By Heap Update, Ψ0;; Γ0`TALI.

Case unpack. T Dhas the form

`TALH: Ψ Ψ`TALR: Γ

Ψ;; Γ`TALv:∃α.τ0 Ψ;α; Γ{r:τ0} `TALI Ψ;; Γ`TALunpack[α, r], v;I

`TAL P .

By the operational semantics,P0 = (H, R0, I0) whereR0=R{r7→w},I0 =I[τ /α] and ˆR(v) =pack[τ, w]as∃α.τ0. Then

(1) `TALH : Ψ is inT D.

(2) By Canonical Forms it follows from Ψ;; Γ `TAL v : ∃α.τ0 that Ψ;∅ `TAL w : τ0[τ /α] wval. Let Γ0= Γ{r:τ0[τ /α]}. By Register File Update if follows that Ψ`TALR0: Γ0.

(3) By Type Substitution it follows from Ψ;α; Γ{r:τ0} `TALI that Ψ;; Γ0`TALI0.

Theorem (Progress). If `TAL P then either there exists P0 such that P 7−→

P0, or P is of the form (H, R{r1 7→ w},halt[τ]) (and, moreover, Ψ;∅ `TAL w : τ wvalfor some Ψsuch that `TALH : Ψ).

Proof. Suppose P = (H, R, Ifull). Let T D be the derivation of`TAL P. The proof is by cases on the first instruction ofIfull.

Case halt T Dhas the form

`TALH : Ψ Ψ;∅ `TALR: Γ

Ψ;; Γ`TALr1:τ Ψ;; Γ`TALhalt[τ]

`TAL(H, R,halt[τ]) .

By ˆR Typing we may deduce that ˆR(r1) is defined and Ψ;∅ `TAL Rˆ(r1) :τ wval.

In other words,R=R0{r17→w}and Ψ;∅ `TALw:τ wval.

Case add,mul,sub T D has the form

`TALH: Ψ Ψ;∅ `TALR: Γ

Ψ;; Γ`TALrs:int Ψ;; Γ`TALv:int · · · Ψ;; Γ`TALarithprd, rs, v;I

`TAL(H, R, Ifull) .

By Canonical Forms,R(rs) andR(v) each represent integer literals. HenceP 7−→

(H, R{rd7→R(rs)pRˆ(v)}, I).

Case bnz T D has the form

`TALH : Ψ Ψ;∅ `TAL R: Γ

Ψ;; Γ`TALr:int Ψ;; Γ`TALv:[ ].Γ0 · · · Ψ;; Γ`TALbnzr, v;I

`TAL(H, R, Ifull) .

By Canonical Forms,R(r) is an integer literal and ˆR(v) =`[σ1, . . . , σn] withH(`) = code[α1, . . . , αn].Γ00.I0. IfR(r) = 0 then P 7−→(H, R, I). If R(r)6= 0 then P 7−→

(H, R, I0[~σ/~α]).

Case jmp T D has the form

`TALH : Ψ Ψ;∅ `TALR: Γ

Ψ;; Γ`TALv:[ ].Γ0 · · · Ψ;; Γ`TALjmpv

`TAL(H, R, Ifull) .

By Canonical Forms, ˆR(v) = `[σ1, . . . , σn] with H(`) = code[α1, . . . , αn].Γ00.I0. HenceP 7−→(H, R, I0[~σ/~α]).

Case ld T D has the form

`TALH : Ψ Ψ;∅ `TAL R: Γ

Ψ;; Γ`TALrs:0ϕ0, . . . , τn−ϕn−11 i · · ·

Ψ;; Γ`TALldrd, rs[i];I (1≤i < n)

`TAL(H, R, Ifull) .

By Canonical Forms, R(rs) = ` with H(`) = hw0, . . . , wn−1i. Hence P 7−→

(H, R{rd7→wi}, I).

Case malloc Suppose that Ifull is of the formmallocr[τ1, . . . , τn];I. Then P 7−→

(H{`7→ h?τ1, . . . ,?τni}, R{r7→`}, I) for some`6∈H. Case mov T D has the form

`TALH : Ψ Ψ;∅ `TAL R: Γ

Ψ;; Γ`TALv:τ · · · Ψ;; Γ`TALmovr, v;I

`TAL(H, R, Ifull) . By ˆR Typing, ˆR(v) is defined. HenceP 7−→(H, R{r7→Rˆ(v)}, I).

`TAL(H, R, Ifull) . By ˆR Typing, ˆR(v) is defined. HenceP 7−→(H, R{r7→Rˆ(v)}, I).

Documents relatifs