Mercurial > cpdt > repo
comparison src/Interps.v @ 255:05250878e4ca
Prosified Interps
author | Adam Chlipala <adamc@hcoop.net> |
---|---|
date | Wed, 16 Dec 2009 13:50:24 -0500 |
parents | 2a34c4dc6a10 |
children | 2c88fc1dbe33 |
comparison
equal
deleted
inserted
replaced
254:e3c3b7ef5901 | 255:05250878e4ca |
---|---|
16 (* end hide *) | 16 (* end hide *) |
17 | 17 |
18 | 18 |
19 (** %\chapter{Type-Theoretic Interpreters}% *) | 19 (** %\chapter{Type-Theoretic Interpreters}% *) |
20 | 20 |
21 (** TODO: Prose for this chapter *) | 21 (** Throughout this book, we have given semantics for programming languages via executable interpreters written in Gallina. PHOAS is quite compatible with that model, when we want to formalize many of the wide variety of interesting non-Turing-complete programming languages. Most such languages have very straightforward elaborations into Gallina. In this chapter, we show how to extend our past approach to higher-order languages encoded with PHOAS, and we show how simple program transformations may be proved correct with respect to these elaborative semantics. *) |
22 | 22 |
23 | 23 |
24 (** * Simply-Typed Lambda Calculus *) | 24 (** * Simply-Typed Lambda Calculus *) |
25 | |
26 (** We begin with a copy of last chapter's encoding of the syntax of simply-typed lambda calculus with natural numbers and addition. The primes at the ends of constructor names are gone, since here our primary subject is [exp]s instead of [Exp]s. *) | |
25 | 27 |
26 Module STLC. | 28 Module STLC. |
27 Inductive type : Type := | 29 Inductive type : Type := |
28 | Nat : type | 30 | Nat : type |
29 | Arrow : type -> type -> type. | 31 | Arrow : type -> type -> type. |
56 Implicit Arguments Const [var]. | 58 Implicit Arguments Const [var]. |
57 Implicit Arguments Plus [var]. | 59 Implicit Arguments Plus [var]. |
58 Implicit Arguments App [var t1 t2]. | 60 Implicit Arguments App [var t1 t2]. |
59 Implicit Arguments Abs [var t1 t2]. | 61 Implicit Arguments Abs [var t1 t2]. |
60 | 62 |
63 (** The definitions that follow will be easier to read if we define some parsing notations for the constructors. *) | |
64 | |
61 Notation "# v" := (Var v) (at level 70). | 65 Notation "# v" := (Var v) (at level 70). |
62 | 66 |
63 Notation "^ n" := (Const n) (at level 70). | 67 Notation "^ n" := (Const n) (at level 70). |
64 Infix "+^" := Plus (left associativity, at level 79). | 68 Infix "+^" := Plus (left associativity, at level 79). |
65 | 69 |
66 Infix "@" := App (left associativity, at level 77). | 70 Infix "@" := App (left associativity, at level 77). |
67 Notation "\ x , e" := (Abs (fun x => e)) (at level 78). | 71 Notation "\ x , e" := (Abs (fun x => e)) (at level 78). |
68 Notation "\ ! , e" := (Abs (fun _ => e)) (at level 78). | 72 Notation "\ ! , e" := (Abs (fun _ => e)) (at level 78). |
69 | 73 |
70 Definition zero : Exp Nat := fun _ => ^0. | 74 (** A few examples will be useful for testing the functions we will write. *) |
71 Definition one : Exp Nat := fun _ => ^1. | 75 |
72 Definition zpo : Exp Nat := fun _ => zero _ +^ one _. | 76 Example zero : Exp Nat := fun _ => ^0. |
73 Definition ident : Exp (Nat --> Nat) := fun _ => \x, #x. | 77 Example one : Exp Nat := fun _ => ^1. |
74 Definition app_ident : Exp Nat := fun _ => ident _ @ zpo _. | 78 Example zpo : Exp Nat := fun _ => zero _ +^ one _. |
75 Definition app : Exp ((Nat --> Nat) --> Nat --> Nat) := fun _ => | 79 Example ident : Exp (Nat --> Nat) := fun _ => \x, #x. |
76 \f, \x, #f @ #x. | 80 Example app_ident : Exp Nat := fun _ => ident _ @ zpo _. |
77 Definition app_ident' : Exp Nat := fun _ => app _ @ ident _ @ zpo _. | 81 Example app : Exp ((Nat --> Nat) --> Nat --> Nat) := fun _ => \f, \x, #f @ #x. |
82 Example app_ident' : Exp Nat := fun _ => app _ @ ident _ @ zpo _. | |
78 | 83 |
79 (* EX: Define an interpreter for [Exp]s. *) | 84 (* EX: Define an interpreter for [Exp]s. *) |
80 | 85 |
81 (* begin thide *) | 86 (* begin thide *) |
87 (** To write our interpreter, we must first interpret object language types as meta language types. *) | |
88 | |
82 Fixpoint typeDenote (t : type) : Set := | 89 Fixpoint typeDenote (t : type) : Set := |
83 match t with | 90 match t with |
84 | Nat => nat | 91 | Nat => nat |
85 | t1 --> t2 => typeDenote t1 -> typeDenote t2 | 92 | t1 --> t2 => typeDenote t1 -> typeDenote t2 |
86 end. | 93 end. |
94 | |
95 (** The crucial trick of the expression interpreter is to represent variables using the [typeDenote] function. Due to limitations in Coq's syntax extension system, we cannot take advantage of some of our notations when they appear in patterns, so, to be consistent, in patterns we avoid notations altogether. *) | |
87 | 96 |
88 Fixpoint expDenote t (e : exp typeDenote t) : typeDenote t := | 97 Fixpoint expDenote t (e : exp typeDenote t) : typeDenote t := |
89 match e with | 98 match e with |
90 | Var _ v => v | 99 | Var _ v => v |
91 | 100 |
97 end. | 106 end. |
98 | 107 |
99 Definition ExpDenote t (e : Exp t) := expDenote (e _). | 108 Definition ExpDenote t (e : Exp t) := expDenote (e _). |
100 (* end thide *) | 109 (* end thide *) |
101 | 110 |
111 (** Some tests establish that [ExpDenote] produces Gallina terms like we might write manually. *) | |
112 | |
102 Eval compute in ExpDenote zero. | 113 Eval compute in ExpDenote zero. |
114 (** %\vspace{-.15in}% [[ | |
115 = 0 | |
116 : typeDenote Nat | |
117 ]] *) | |
118 | |
103 Eval compute in ExpDenote one. | 119 Eval compute in ExpDenote one. |
120 (** %\vspace{-.15in}% [[ | |
121 = 1 | |
122 : typeDenote Nat | |
123 ]] *) | |
124 | |
104 Eval compute in ExpDenote zpo. | 125 Eval compute in ExpDenote zpo. |
126 (** %\vspace{-.15in}% [[ | |
127 = 1 | |
128 : typeDenote Nat | |
129 ]] *) | |
130 | |
105 Eval compute in ExpDenote ident. | 131 Eval compute in ExpDenote ident. |
132 (** %\vspace{-.15in}% [[ | |
133 = fun x : nat => x | |
134 : typeDenote (Nat --> Nat) | |
135 ]] *) | |
136 | |
106 Eval compute in ExpDenote app_ident. | 137 Eval compute in ExpDenote app_ident. |
138 (** %\vspace{-.15in}% [[ | |
139 = 1 | |
140 : typeDenote Nat | |
141 ]] *) | |
142 | |
107 Eval compute in ExpDenote app. | 143 Eval compute in ExpDenote app. |
144 (** %\vspace{-.15in}% [[ | |
145 = fun (x : nat -> nat) (x0 : nat) => x x0 | |
146 : typeDenote ((Nat --> Nat) --> Nat --> Nat) | |
147 ]] *) | |
148 | |
108 Eval compute in ExpDenote app_ident'. | 149 Eval compute in ExpDenote app_ident'. |
150 (** %\vspace{-.15in}% [[ | |
151 = 1 | |
152 : typeDenote Nat | |
153 ]] *) | |
154 | |
109 | 155 |
110 (* EX: Define a constant-folding function for [Exp]s. *) | 156 (* EX: Define a constant-folding function for [Exp]s. *) |
157 | |
158 (** We can update to the higher-order case our common example of a constant folding function. The workhorse function [cfold] is parameterized to apply to an [exp] that uses any [var] type. An output of [cfold] uses the same [var] type as the input. As in the definition of [expDenote], we cannot use most of our notations in patterns, but we use them freely to make the bodies of [match] cases easier to read. *) | |
111 | 159 |
112 (* begin thide *) | 160 (* begin thide *) |
113 Section cfold. | 161 Section cfold. |
114 Variable var : type -> Type. | 162 Variable var : type -> Type. |
115 | 163 |
125 | Const n1, Const n2 => ^(n1 + n2) | 173 | Const n1, Const n2 => ^(n1 + n2) |
126 | _, _ => e1' +^ e2' | 174 | _, _ => e1' +^ e2' |
127 end | 175 end |
128 | 176 |
129 | App _ _ e1 e2 => cfold e1 @ cfold e2 | 177 | App _ _ e1 e2 => cfold e1 @ cfold e2 |
130 | Abs _ _ e' => Abs (fun x => cfold (e' x)) | 178 | Abs _ _ e' => \x, cfold (e' x) |
131 end. | 179 end. |
132 End cfold. | 180 End cfold. |
133 | 181 |
134 Definition Cfold t (E : Exp t) : Exp t := fun _ => cfold (E _). | 182 Definition Cfold t (E : Exp t) : Exp t := fun _ => cfold (E _). |
135 (* end thide *) | 183 (* end thide *) |
136 | 184 |
137 (* EX: Prove the correctness of constant-folding. *) | 185 (* EX: Prove the correctness of constant-folding. *) |
186 | |
187 (** Now we would like to prove the correctness of [Cfold], which follows from a simple inductive lemma about [cfold]. *) | |
138 | 188 |
139 (* begin thide *) | 189 (* begin thide *) |
140 Lemma cfold_correct : forall t (e : exp _ t), | 190 Lemma cfold_correct : forall t (e : exp _ t), |
141 expDenote (cfold e) = expDenote e. | 191 expDenote (cfold e) = expDenote e. |
142 induction e; crush; try (ext_eq; crush); | 192 induction e; crush; try (ext_eq; crush); |
152 (* end thide *) | 202 (* end thide *) |
153 End STLC. | 203 End STLC. |
154 | 204 |
155 | 205 |
156 (** * Adding Products and Sums *) | 206 (** * Adding Products and Sums *) |
207 | |
208 (** The example is easily adapted to support products and sums, the basis of non-recursive datatypes in ML and Haskell. *) | |
157 | 209 |
158 Module PSLC. | 210 Module PSLC. |
159 (* EX: Extend the development with products and sums. *) | 211 (* EX: Extend the development with products and sums. *) |
160 | 212 |
161 (* begin thide *) | 213 (* begin thide *) |
237 Notation "#2 e" := (Snd e) (at level 75). | 289 Notation "#2 e" := (Snd e) (at level 75). |
238 | 290 |
239 Notation "'case' e 'of' x => e1 | y => e2" := (SumCase e (fun x => e1) (fun y => e2)) | 291 Notation "'case' e 'of' x => e1 | y => e2" := (SumCase e (fun x => e1) (fun y => e2)) |
240 (at level 79). | 292 (at level 79). |
241 | 293 |
242 Definition swap : Exp (Nat ** Nat --> Nat ** Nat) := fun _ => | 294 (** A few examples can be defined easily, using the notations above. *) |
243 \p, [#2 #p, #1 #p]. | 295 |
244 Definition zo : Exp (Nat ** Nat) := fun _ => [^0, ^1]. | 296 Example swap : Exp (Nat ** Nat --> Nat ** Nat) := fun _ => \p, [#2 #p, #1 #p]. |
245 Definition swap_zo : Exp (Nat ** Nat) := fun _ => swap _ @ zo _. | 297 Example zo : Exp (Nat ** Nat) := fun _ => [^0, ^1]. |
246 | 298 Example swap_zo : Exp (Nat ** Nat) := fun _ => swap _ @ zo _. |
247 Definition natOut : Exp (Nat ++ Nat --> Nat) := fun _ => | 299 |
300 Example natOut : Exp (Nat ++ Nat --> Nat) := fun _ => | |
248 \s, case #s of x => #x | y => #y +^ #y. | 301 \s, case #s of x => #x | y => #y +^ #y. |
249 Definition ns1 : Exp (Nat ++ Nat) := fun _ => Inl (^3). | 302 Example ns1 : Exp (Nat ++ Nat) := fun _ => Inl (^3). |
250 Definition ns2 : Exp (Nat ++ Nat) := fun _ => Inr (^5). | 303 Example ns2 : Exp (Nat ++ Nat) := fun _ => Inr (^5). |
251 Definition natOut_ns1 : Exp Nat := fun _ => natOut _ @ ns1 _. | 304 Example natOut_ns1 : Exp Nat := fun _ => natOut _ @ ns1 _. |
252 Definition natOut_ns2 : Exp Nat := fun _ => natOut _ @ ns2 _. | 305 Example natOut_ns2 : Exp Nat := fun _ => natOut _ @ ns2 _. |
306 | |
307 (** The semantics adapts without incident. *) | |
253 | 308 |
254 (* begin thide *) | 309 (* begin thide *) |
255 Fixpoint typeDenote (t : type) : Set := | 310 Fixpoint typeDenote (t : type) : Set := |
256 match t with | 311 match t with |
257 | Nat => nat | 312 | Nat => nat |
285 | 340 |
286 Definition ExpDenote t (e : Exp t) := expDenote (e _). | 341 Definition ExpDenote t (e : Exp t) := expDenote (e _). |
287 (* end thide *) | 342 (* end thide *) |
288 | 343 |
289 Eval compute in ExpDenote swap. | 344 Eval compute in ExpDenote swap. |
345 (** %\vspace{-.15in}% [[ | |
346 = fun x : nat * nat => (let (_, y) := x in y, let (x0, _) := x in x0) | |
347 : typeDenote (Nat ** Nat --> Nat ** Nat) | |
348 ]] *) | |
349 | |
290 Eval compute in ExpDenote zo. | 350 Eval compute in ExpDenote zo. |
351 (** %\vspace{-.15in}% [[ | |
352 = (0, 1) | |
353 : typeDenote (Nat ** Nat) | |
354 ]] *) | |
355 | |
291 Eval compute in ExpDenote swap_zo. | 356 Eval compute in ExpDenote swap_zo. |
292 | 357 (** %\vspace{-.15in}% [[ |
293 Eval compute in ExpDenote natOut. | 358 = (1, 0) |
359 : typeDenote (Nat ** Nat) | |
360 ]] *) | |
361 | |
362 Eval cbv beta iota delta -[plus] in ExpDenote natOut. | |
363 (** %\vspace{-.15in}% [[ | |
364 = fun x : nat + nat => match x with | |
365 | inl v => v | |
366 | inr v => v + v | |
367 end | |
368 : typeDenote (Nat ++ Nat --> Nat) | |
369 ]] *) | |
370 | |
294 Eval compute in ExpDenote ns1. | 371 Eval compute in ExpDenote ns1. |
372 (** %\vspace{-.15in}% [[ | |
373 = inl nat 3 | |
374 : typeDenote (Nat ++ Nat) | |
375 ]] *) | |
376 | |
295 Eval compute in ExpDenote ns2. | 377 Eval compute in ExpDenote ns2. |
378 (** %\vspace{-.15in}% [[ | |
379 = inr nat 5 | |
380 : typeDenote (Nat ++ Nat) | |
381 ]] *) | |
382 | |
296 Eval compute in ExpDenote natOut_ns1. | 383 Eval compute in ExpDenote natOut_ns1. |
384 (** %\vspace{-.15in}% [[ | |
385 = 3 | |
386 : typeDenote Nat | |
387 ]] *) | |
388 | |
297 Eval compute in ExpDenote natOut_ns2. | 389 Eval compute in ExpDenote natOut_ns2. |
390 (** %\vspace{-.15in}% [[ | |
391 = 10 | |
392 : typeDenote Nat | |
393 ]] *) | |
394 | |
395 (** We adapt the [cfold] function using the same basic dependent-types trick that we applied in an earlier chapter to a very similar function for a language without variables. *) | |
298 | 396 |
299 (* begin thide *) | 397 (* begin thide *) |
300 Section cfold. | 398 Section cfold. |
301 Variable var : type -> Type. | 399 Variable var : type -> Type. |
302 | 400 |
305 | t1 ** t2 => option (exp var t1 * exp var t2) | 403 | t1 ** t2 => option (exp var t1 * exp var t2) |
306 | _ => unit | 404 | _ => unit |
307 end. | 405 end. |
308 | 406 |
309 Definition pairOutDefault (t : type) : pairOutType t := | 407 Definition pairOutDefault (t : type) : pairOutType t := |
310 match t return pairOutType t with | 408 match t with |
311 | _ ** _ => None | 409 | _ ** _ => None |
312 | _ => tt | 410 | _ => tt |
313 end. | 411 end. |
314 | 412 |
315 Definition pairOut t1 t2 (e : exp var (t1 ** t2)) | 413 Definition pairOut t1 t2 (e : exp var (t1 ** t2)) |
331 | Const n1, Const n2 => ^(n1 + n2) | 429 | Const n1, Const n2 => ^(n1 + n2) |
332 | _, _ => e1' +^ e2' | 430 | _, _ => e1' +^ e2' |
333 end | 431 end |
334 | 432 |
335 | App _ _ e1 e2 => cfold e1 @ cfold e2 | 433 | App _ _ e1 e2 => cfold e1 @ cfold e2 |
336 | Abs _ _ e' => Abs (fun x => cfold (e' x)) | 434 | Abs _ _ e' => \x, cfold (e' x) |
337 | 435 |
338 | Pair _ _ e1 e2 => [cfold e1, cfold e2] | 436 | Pair _ _ e1 e2 => [cfold e1, cfold e2] |
339 | Fst t1 _ e' => | 437 | Fst t1 _ e' => |
340 let e'' := cfold e' in | 438 let e'' := cfold e' in |
341 match pairOut e'' with | 439 match pairOut e'' with |
358 end. | 456 end. |
359 End cfold. | 457 End cfold. |
360 | 458 |
361 Definition Cfold t (E : Exp t) : Exp t := fun _ => cfold (E _). | 459 Definition Cfold t (E : Exp t) : Exp t := fun _ => cfold (E _). |
362 | 460 |
461 (** The proofs are almost as straightforward as before. We first establish two simple theorems about pairs and their projections. *) | |
462 | |
363 Section pairs. | 463 Section pairs. |
364 Variables A B : Type. | 464 Variables A B : Type. |
365 | 465 |
366 Variable v1 : A. | 466 Variable v1 : A. |
367 Variable v2 : B. | 467 Variable v2 : B. |
375 destruct v; crush. | 475 destruct v; crush. |
376 Qed. | 476 Qed. |
377 End pairs. | 477 End pairs. |
378 | 478 |
379 Hint Resolve pair_eta1 pair_eta2. | 479 Hint Resolve pair_eta1 pair_eta2. |
480 | |
481 (** To the proof script for the main lemma, we add just one more [match] case, detecting when case analysis is appropriate on discriminees of matches over sum types. *) | |
380 | 482 |
381 Lemma cfold_correct : forall t (e : exp _ t), | 483 Lemma cfold_correct : forall t (e : exp _ t), |
382 expDenote (cfold e) = expDenote e. | 484 expDenote (cfold e) = expDenote e. |
383 induction e; crush; try (ext_eq; crush); | 485 induction e; crush; try (ext_eq; crush); |
384 repeat (match goal with | 486 repeat (match goal with |