annotate src/Hoas.v @ 222:13620dfd5f97

Port Hoas
author Adam Chlipala <adamc@hcoop.net>
date Mon, 16 Nov 2009 12:08:18 -0500
parents 8905f28ffeef
children 3c4ed57c9907
rev   line source
adamc@222 1 (* Copyright (c) 2008-2009, Adam Chlipala
adamc@158 2 *
adamc@158 3 * This work is licensed under a
adamc@158 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@158 5 * Unported License.
adamc@158 6 * The license text is available at:
adamc@158 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@158 8 *)
adamc@158 9
adamc@158 10 (* begin hide *)
adamc@168 11 Require Import Eqdep String List.
adamc@158 12
adamc@168 13 Require Import Axioms Tactics.
adamc@158 14
adamc@158 15 Set Implicit Arguments.
adamc@158 16 (* end hide *)
adamc@158 17
adamc@158 18
adamc@158 19 (** %\chapter{Higher-Order Abstract Syntax}% *)
adamc@158 20
adamc@158 21 (** TODO: Prose for this chapter *)
adamc@158 22
adamc@158 23
adamc@158 24 (** * Parametric Higher-Order Abstract Syntax *)
adamc@158 25
adamc@158 26 Inductive type : Type :=
adamc@159 27 | Nat : type
adamc@158 28 | Arrow : type -> type -> type.
adamc@158 29
adamc@158 30 Infix "-->" := Arrow (right associativity, at level 60).
adamc@158 31
adamc@158 32 Section exp.
adamc@158 33 Variable var : type -> Type.
adamc@158 34
adamc@158 35 Inductive exp : type -> Type :=
adamc@159 36 | Const' : nat -> exp Nat
adamc@159 37 | Plus' : exp Nat -> exp Nat -> exp Nat
adamc@159 38
adamc@158 39 | Var : forall t, var t -> exp t
adamc@158 40 | App' : forall dom ran, exp (dom --> ran) -> exp dom -> exp ran
adamc@158 41 | Abs' : forall dom ran, (var dom -> exp ran) -> exp (dom --> ran).
adamc@158 42 End exp.
adamc@158 43
adamc@158 44 Implicit Arguments Const' [var].
adamc@158 45 Implicit Arguments Var [var t].
adamc@158 46 Implicit Arguments Abs' [var dom ran].
adamc@158 47
adamc@158 48 Definition Exp t := forall var, exp var t.
adamc@169 49 (* begin thide *)
adamc@158 50 Definition Exp1 t1 t2 := forall var, var t1 -> exp var t2.
adamc@158 51
adamc@159 52 Definition Const (n : nat) : Exp Nat :=
adamc@159 53 fun _ => Const' n.
adamc@159 54 Definition Plus (E1 E2 : Exp Nat) : Exp Nat :=
adamc@159 55 fun _ => Plus' (E1 _) (E2 _).
adamc@158 56 Definition App dom ran (F : Exp (dom --> ran)) (X : Exp dom) : Exp ran :=
adamc@158 57 fun _ => App' (F _) (X _).
adamc@158 58 Definition Abs dom ran (B : Exp1 dom ran) : Exp (dom --> ran) :=
adamc@158 59 fun _ => Abs' (B _).
adamc@169 60 (* end thide *)
adamc@169 61
adamc@169 62 (* EX: Define appropriate shorthands, so that these definitions type-check. *)
adamc@158 63
adamc@166 64 Definition zero := Const 0.
adamc@166 65 Definition one := Const 1.
adamc@166 66 Definition one_again := Plus zero one.
adamc@166 67 Definition ident : Exp (Nat --> Nat) := Abs (fun _ X => Var X).
adamc@166 68 Definition app_ident := App ident one_again.
adamc@166 69 Definition app : Exp ((Nat --> Nat) --> Nat --> Nat) := fun _ =>
adamc@166 70 Abs' (fun f => Abs' (fun x => App' (Var f) (Var x))).
adamc@166 71 Definition app_ident' := App (App app ident) one_again.
adamc@166 72
adamc@169 73 (* EX: Define a function to count the number of variable occurrences in an [Exp]. *)
adamc@169 74
adamc@169 75 (* begin thide *)
adamc@222 76 Fixpoint countVars t (e : exp (fun _ => unit) t) : nat :=
adamc@166 77 match e with
adamc@166 78 | Const' _ => 0
adamc@166 79 | Plus' e1 e2 => countVars e1 + countVars e2
adamc@166 80 | Var _ _ => 1
adamc@166 81 | App' _ _ e1 e2 => countVars e1 + countVars e2
adamc@166 82 | Abs' _ _ e' => countVars (e' tt)
adamc@166 83 end.
adamc@166 84
adamc@166 85 Definition CountVars t (E : Exp t) : nat := countVars (E _).
adamc@169 86 (* end thide *)
adamc@166 87
adamc@166 88 Eval compute in CountVars zero.
adamc@166 89 Eval compute in CountVars one.
adamc@166 90 Eval compute in CountVars one_again.
adamc@166 91 Eval compute in CountVars ident.
adamc@166 92 Eval compute in CountVars app_ident.
adamc@166 93 Eval compute in CountVars app.
adamc@166 94 Eval compute in CountVars app_ident'.
adamc@166 95
adamc@169 96 (* EX: Define a function to count the number of occurrences of a single distinguished variable. *)
adamc@169 97
adamc@169 98 (* begin thide *)
adamc@222 99 Fixpoint countOne t (e : exp (fun _ => bool) t) : nat :=
adamc@166 100 match e with
adamc@166 101 | Const' _ => 0
adamc@166 102 | Plus' e1 e2 => countOne e1 + countOne e2
adamc@166 103 | Var _ true => 1
adamc@166 104 | Var _ false => 0
adamc@166 105 | App' _ _ e1 e2 => countOne e1 + countOne e2
adamc@166 106 | Abs' _ _ e' => countOne (e' false)
adamc@166 107 end.
adamc@166 108
adamc@166 109 Definition CountOne t1 t2 (E : Exp1 t1 t2) : nat :=
adamc@166 110 countOne (E _ true).
adamc@169 111 (* end thide *)
adamc@166 112
adamc@166 113 Definition ident1 : Exp1 Nat Nat := fun _ X => Var X.
adamc@166 114 Definition add_self : Exp1 Nat Nat := fun _ X => Plus' (Var X) (Var X).
adamc@166 115 Definition app_zero : Exp1 (Nat --> Nat) Nat := fun _ X => App' (Var X) (Const' 0).
adamc@166 116 Definition app_ident1 : Exp1 Nat Nat := fun _ X => App' (Abs' (fun Y => Var Y)) (Var X).
adamc@166 117
adamc@166 118 Eval compute in CountOne ident1.
adamc@166 119 Eval compute in CountOne add_self.
adamc@166 120 Eval compute in CountOne app_zero.
adamc@166 121 Eval compute in CountOne app_ident1.
adamc@166 122
adamc@169 123 (* EX: Define a function to pretty-print [Exp]s as strings. *)
adamc@169 124
adamc@169 125 (* begin thide *)
adamc@166 126 Section ToString.
adamc@166 127 Open Scope string_scope.
adamc@166 128
adamc@166 129 Fixpoint natToString (n : nat) : string :=
adamc@166 130 match n with
adamc@166 131 | O => "O"
adamc@166 132 | S n' => "S(" ++ natToString n' ++ ")"
adamc@166 133 end.
adamc@166 134
adamc@222 135 Fixpoint toString t (e : exp (fun _ => string) t) (cur : string) : string * string :=
adamc@166 136 match e with
adamc@166 137 | Const' n => (cur, natToString n)
adamc@166 138 | Plus' e1 e2 =>
adamc@166 139 let (cur', s1) := toString e1 cur in
adamc@166 140 let (cur'', s2) := toString e2 cur' in
adamc@166 141 (cur'', "(" ++ s1 ++ ") + (" ++ s2 ++ ")")
adamc@166 142 | Var _ s => (cur, s)
adamc@166 143 | App' _ _ e1 e2 =>
adamc@166 144 let (cur', s1) := toString e1 cur in
adamc@166 145 let (cur'', s2) := toString e2 cur' in
adamc@166 146 (cur'', "(" ++ s1 ++ ") (" ++ s2 ++ ")")
adamc@166 147 | Abs' _ _ e' =>
adamc@166 148 let (cur', s) := toString (e' cur) (cur ++ "'") in
adamc@166 149 (cur', "(\" ++ cur ++ ", " ++ s ++ ")")
adamc@166 150 end.
adamc@166 151
adamc@166 152 Definition ToString t (E : Exp t) : string := snd (toString (E _) "x").
adamc@166 153 End ToString.
adamc@169 154 (* end thide *)
adamc@166 155
adamc@166 156 Eval compute in ToString zero.
adamc@166 157 Eval compute in ToString one.
adamc@166 158 Eval compute in ToString one_again.
adamc@166 159 Eval compute in ToString ident.
adamc@166 160 Eval compute in ToString app_ident.
adamc@166 161 Eval compute in ToString app.
adamc@166 162 Eval compute in ToString app_ident'.
adamc@166 163
adamc@169 164 (* EX: Define a substitution function. *)
adamc@169 165
adamc@169 166 (* begin thide *)
adamc@158 167 Section flatten.
adamc@158 168 Variable var : type -> Type.
adamc@158 169
adamc@222 170 Fixpoint flatten t (e : exp (exp var) t) : exp var t :=
adamc@222 171 match e with
adamc@159 172 | Const' n => Const' n
adamc@159 173 | Plus' e1 e2 => Plus' (flatten e1) (flatten e2)
adamc@158 174 | Var _ e' => e'
adamc@158 175 | App' _ _ e1 e2 => App' (flatten e1) (flatten e2)
adamc@158 176 | Abs' _ _ e' => Abs' (fun x => flatten (e' (Var x)))
adamc@158 177 end.
adamc@158 178 End flatten.
adamc@158 179
adamc@158 180 Definition Subst t1 t2 (E1 : Exp t1) (E2 : Exp1 t1 t2) : Exp t2 := fun _ =>
adamc@158 181 flatten (E2 _ (E1 _)).
adamc@169 182 (* end thide *)
adamc@158 183
adamc@166 184 Eval compute in Subst one ident1.
adamc@166 185 Eval compute in Subst one add_self.
adamc@166 186 Eval compute in Subst ident app_zero.
adamc@166 187 Eval compute in Subst one app_ident1.
adamc@166 188
adamc@158 189
adamc@158 190 (** * A Type Soundness Proof *)
adamc@158 191
adamc@158 192 Reserved Notation "E1 ==> E2" (no associativity, at level 90).
adamc@158 193
adamc@158 194 Inductive Val : forall t, Exp t -> Prop :=
adamc@159 195 | VConst : forall n, Val (Const n)
adamc@158 196 | VAbs : forall dom ran (B : Exp1 dom ran), Val (Abs B).
adamc@158 197
adamc@158 198 Hint Constructors Val.
adamc@158 199
adamc@162 200 Inductive Ctx : type -> type -> Type :=
adamc@162 201 | AppCong1 : forall (dom ran : type),
adamc@162 202 Exp dom -> Ctx (dom --> ran) ran
adamc@162 203 | AppCong2 : forall (dom ran : type),
adamc@162 204 Exp (dom --> ran) -> Ctx dom ran
adamc@162 205 | PlusCong1 : Exp Nat -> Ctx Nat Nat
adamc@162 206 | PlusCong2 : Exp Nat -> Ctx Nat Nat.
adamc@162 207
adamc@162 208 Inductive isCtx : forall t1 t2, Ctx t1 t2 -> Prop :=
adamc@162 209 | IsApp1 : forall dom ran (X : Exp dom), isCtx (AppCong1 ran X)
adamc@162 210 | IsApp2 : forall dom ran (F : Exp (dom --> ran)), Val F -> isCtx (AppCong2 F)
adamc@162 211 | IsPlus1 : forall E2, isCtx (PlusCong1 E2)
adamc@162 212 | IsPlus2 : forall E1, Val E1 -> isCtx (PlusCong2 E1).
adamc@162 213
adamc@162 214 Definition plug t1 t2 (C : Ctx t1 t2) : Exp t1 -> Exp t2 :=
adamc@222 215 match C with
adamc@162 216 | AppCong1 _ _ X => fun F => App F X
adamc@162 217 | AppCong2 _ _ F => fun X => App F X
adamc@162 218 | PlusCong1 E2 => fun E1 => Plus E1 E2
adamc@162 219 | PlusCong2 E1 => fun E2 => Plus E1 E2
adamc@162 220 end.
adamc@162 221
adamc@162 222 Infix "@" := plug (no associativity, at level 60).
adamc@162 223
adamc@158 224 Inductive Step : forall t, Exp t -> Exp t -> Prop :=
adamc@158 225 | Beta : forall dom ran (B : Exp1 dom ran) (X : Exp dom),
adamc@160 226 Val X
adamc@160 227 -> App (Abs B) X ==> Subst X B
adamc@159 228 | Sum : forall n1 n2,
adamc@159 229 Plus (Const n1) (Const n2) ==> Const (n1 + n2)
adamc@162 230 | Cong : forall t t' (C : Ctx t t') E E' E1,
adamc@162 231 isCtx C
adamc@162 232 -> E1 = C @ E
adamc@162 233 -> E ==> E'
adamc@162 234 -> E1 ==> C @ E'
adamc@159 235
adamc@158 236 where "E1 ==> E2" := (Step E1 E2).
adamc@158 237
adamc@162 238 Hint Constructors isCtx Step.
adamc@158 239
adamc@169 240 (* EX: Prove type soundness. *)
adamc@169 241
adamc@169 242 (* begin thide *)
adamc@158 243 Inductive Closed : forall t, Exp t -> Prop :=
adamc@164 244 | CConst : forall n,
adamc@164 245 Closed (Const n)
adamc@159 246 | CPlus : forall E1 E2,
adamc@159 247 Closed E1
adamc@159 248 -> Closed E2
adamc@159 249 -> Closed (Plus E1 E2)
adamc@158 250 | CApp : forall dom ran (E1 : Exp (dom --> ran)) E2,
adamc@158 251 Closed E1
adamc@158 252 -> Closed E2
adamc@158 253 -> Closed (App E1 E2)
adamc@158 254 | CAbs : forall dom ran (E1 : Exp1 dom ran),
adamc@158 255 Closed (Abs E1).
adamc@158 256
adamc@158 257 Axiom closed : forall t (E : Exp t), Closed E.
adamc@158 258
adamc@160 259 Ltac my_crush' :=
adamc@167 260 crush;
adamc@158 261 repeat (match goal with
adamc@158 262 | [ H : _ |- _ ] => generalize (inj_pairT2 _ _ _ _ _ H); clear H
adamc@167 263 end; crush).
adamc@160 264
adamc@162 265 Hint Extern 1 (_ = _ @ _) => simpl.
adamc@162 266
adamc@158 267 Lemma progress' : forall t (E : Exp t),
adamc@158 268 Closed E
adamc@158 269 -> Val E \/ exists E', E ==> E'.
adamc@158 270 induction 1; crush;
adamc@159 271 repeat match goal with
adamc@167 272 | [ H : Val _ |- _ ] => inversion H; []; clear H; my_crush'
adamc@162 273 end; eauto 6.
adamc@158 274 Qed.
adamc@158 275
adamc@158 276 Theorem progress : forall t (E : Exp t),
adamc@158 277 Val E \/ exists E', E ==> E'.
adamc@158 278 intros; apply progress'; apply closed.
adamc@158 279 Qed.
adamc@169 280 (* end thide *)
adamc@158 281
adamc@160 282
adamc@160 283 (** * Big-Step Semantics *)
adamc@160 284
adamc@160 285 Reserved Notation "E1 ===> E2" (no associativity, at level 90).
adamc@160 286
adamc@160 287 Inductive BigStep : forall t, Exp t -> Exp t -> Prop :=
adamc@160 288 | SConst : forall n,
adamc@160 289 Const n ===> Const n
adamc@160 290 | SPlus : forall E1 E2 n1 n2,
adamc@160 291 E1 ===> Const n1
adamc@160 292 -> E2 ===> Const n2
adamc@160 293 -> Plus E1 E2 ===> Const (n1 + n2)
adamc@160 294
adamc@160 295 | SApp : forall dom ran (E1 : Exp (dom --> ran)) E2 B V2 V,
adamc@160 296 E1 ===> Abs B
adamc@160 297 -> E2 ===> V2
adamc@160 298 -> Subst V2 B ===> V
adamc@160 299 -> App E1 E2 ===> V
adamc@160 300 | SAbs : forall dom ran (B : Exp1 dom ran),
adamc@160 301 Abs B ===> Abs B
adamc@160 302
adamc@160 303 where "E1 ===> E2" := (BigStep E1 E2).
adamc@160 304
adamc@160 305 Hint Constructors BigStep.
adamc@160 306
adamc@169 307 (* EX: Prove the equivalence of the small- and big-step semantics. *)
adamc@169 308
adamc@169 309 (* begin thide *)
adamc@160 310 Reserved Notation "E1 ==>* E2" (no associativity, at level 90).
adamc@160 311
adamc@160 312 Inductive MultiStep : forall t, Exp t -> Exp t -> Prop :=
adamc@160 313 | Done : forall t (E : Exp t), E ==>* E
adamc@160 314 | OneStep : forall t (E E' E'' : Exp t),
adamc@160 315 E ==> E'
adamc@160 316 -> E' ==>* E''
adamc@160 317 -> E ==>* E''
adamc@160 318
adamc@160 319 where "E1 ==>* E2" := (MultiStep E1 E2).
adamc@160 320
adamc@160 321 Hint Constructors MultiStep.
adamc@160 322
adamc@160 323 Theorem MultiStep_trans : forall t (E1 E2 E3 : Exp t),
adamc@160 324 E1 ==>* E2
adamc@160 325 -> E2 ==>* E3
adamc@160 326 -> E1 ==>* E3.
adamc@160 327 induction 1; eauto.
adamc@160 328 Qed.
adamc@160 329
adamc@160 330 Theorem Big_Val : forall t (E V : Exp t),
adamc@160 331 E ===> V
adamc@160 332 -> Val V.
adamc@160 333 induction 1; crush.
adamc@160 334 Qed.
adamc@160 335
adamc@160 336 Theorem Val_Big : forall t (V : Exp t),
adamc@160 337 Val V
adamc@160 338 -> V ===> V.
adamc@160 339 destruct 1; crush.
adamc@160 340 Qed.
adamc@160 341
adamc@160 342 Hint Resolve Big_Val Val_Big.
adamc@160 343
adamc@162 344 Lemma Multi_Cong : forall t t' (C : Ctx t t'),
adamc@162 345 isCtx C
adamc@162 346 -> forall E E', E ==>* E'
adamc@162 347 -> C @ E ==>* C @ E'.
adamc@160 348 induction 2; crush; eauto.
adamc@160 349 Qed.
adamc@160 350
adamc@162 351 Lemma Multi_Cong' : forall t t' (C : Ctx t t') E1 E2 E E',
adamc@162 352 isCtx C
adamc@162 353 -> E1 = C @ E
adamc@162 354 -> E2 = C @ E'
adamc@162 355 -> E ==>* E'
adamc@162 356 -> E1 ==>* E2.
adamc@162 357 crush; apply Multi_Cong; auto.
adamc@162 358 Qed.
adamc@162 359
adamc@162 360 Hint Resolve Multi_Cong'.
adamc@162 361
adamc@162 362 Ltac mtrans E :=
adamc@162 363 match goal with
adamc@162 364 | [ |- E ==>* _ ] => fail 1
adamc@162 365 | _ => apply MultiStep_trans with E; [ solve [ eauto ] | eauto ]
adamc@162 366 end.
adamc@160 367
adamc@160 368 Theorem Big_Multi : forall t (E V : Exp t),
adamc@160 369 E ===> V
adamc@160 370 -> E ==>* V.
adamc@162 371 induction 1; crush; eauto;
adamc@162 372 repeat match goal with
adamc@162 373 | [ n1 : _, E2 : _ |- _ ] => mtrans (Plus (Const n1) E2)
adamc@162 374 | [ n1 : _, n2 : _ |- _ ] => mtrans (Plus (Const n1) (Const n2))
adamc@162 375 | [ B : _, E2 : _ |- _ ] => mtrans (App (Abs B) E2)
adamc@162 376 end.
adamc@160 377 Qed.
adamc@160 378
adamc@160 379 Lemma Big_Val' : forall t (V1 V2 : Exp t),
adamc@160 380 Val V2
adamc@160 381 -> V1 = V2
adamc@160 382 -> V1 ===> V2.
adamc@160 383 crush.
adamc@160 384 Qed.
adamc@160 385
adamc@160 386 Hint Resolve Big_Val'.
adamc@160 387
adamc@167 388 Ltac equate_conj F G :=
adamc@167 389 match constr:(F, G) with
adamc@167 390 | (_ ?x1, _ ?x2) => constr:(x1 = x2)
adamc@167 391 | (_ ?x1 ?y1, _ ?x2 ?y2) => constr:(x1 = x2 /\ y1 = y2)
adamc@167 392 | (_ ?x1 ?y1 ?z1, _ ?x2 ?y2 ?z2) => constr:(x1 = x2 /\ y1 = y2 /\ z1 = z2)
adamc@222 393 | (_ ?x1 ?y1 ?z1 ?u1, _ ?x2 ?y2 ?z2 ?u2) =>
adamc@222 394 constr:(x1 = x2 /\ y1 = y2 /\ z1 = z2 /\ u1 = u2)
adamc@222 395 | (_ ?x1 ?y1 ?z1 ?u1 ?v1, _ ?x2 ?y2 ?z2 ?u2 ?v2) =>
adamc@222 396 constr:(x1 = x2 /\ y1 = y2 /\ z1 = z2 /\ u1 = u2 /\ v1 = v2)
adamc@167 397 end.
adamc@167 398
adamc@167 399 Ltac my_crush :=
adamc@167 400 my_crush';
adamc@167 401 repeat (match goal with
adamc@167 402 | [ H : ?F = ?G |- _ ] =>
adamc@167 403 (let H' := fresh "H'" in
adamc@167 404 assert (H' : F (fun _ => unit) = G (fun _ => unit)); [ congruence
adamc@167 405 | discriminate || injection H'; clear H' ];
adamc@167 406 my_crush';
adamc@167 407 repeat match goal with
adamc@167 408 | [ H : context[fun _ => unit] |- _ ] => clear H
adamc@167 409 end;
adamc@167 410 match type of H with
adamc@167 411 | ?F = ?G =>
adamc@167 412 let ec := equate_conj F G in
adamc@167 413 let var := fresh "var" in
adamc@167 414 assert ec; [ intuition; unfold Exp; apply ext_eq; intro var;
adamc@167 415 assert (H' : F var = G var); try congruence;
adamc@167 416 match type of H' with
adamc@167 417 | ?X = ?Y =>
adamc@167 418 let X := eval hnf in X in
adamc@167 419 let Y := eval hnf in Y in
adamc@167 420 change (X = Y) in H'
adamc@167 421 end; injection H'; my_crush'; tauto
adamc@167 422 | intuition; subst ]
adamc@167 423 end);
adamc@167 424 clear H
adamc@167 425 end; my_crush');
adamc@167 426 my_crush'.
adamc@167 427
adamc@160 428 Lemma Multi_Big' : forall t (E E' : Exp t),
adamc@160 429 E ==> E'
adamc@160 430 -> forall E'', E' ===> E''
adamc@160 431 -> E ===> E''.
adamc@160 432 induction 1; crush; eauto;
adamc@160 433 match goal with
adamc@160 434 | [ H : _ ===> _ |- _ ] => inversion H; my_crush; eauto
adamc@162 435 end;
adamc@162 436 match goal with
adamc@162 437 | [ H : isCtx _ |- _ ] => inversion H; my_crush; eauto
adamc@160 438 end.
adamc@160 439 Qed.
adamc@160 440
adamc@160 441 Hint Resolve Multi_Big'.
adamc@160 442
adamc@160 443 Theorem Multi_Big : forall t (E V : Exp t),
adamc@160 444 E ==>* V
adamc@160 445 -> Val V
adamc@160 446 -> E ===> V.
adamc@160 447 induction 1; crush; eauto.
adamc@160 448 Qed.
adamc@169 449 (* end thide *)