annotate src/Firstorder.v @ 161:6087a32b1926

Firstorder clean-ups after class
author Adam Chlipala <adamc@hcoop.net>
date Tue, 04 Nov 2008 11:23:20 -0500
parents 2022e3f2aa26
children 02569049397b
rev   line source
adamc@152 1 (* Copyright (c) 2008, Adam Chlipala
adamc@152 2 *
adamc@152 3 * This work is licensed under a
adamc@152 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@152 5 * Unported License.
adamc@152 6 * The license text is available at:
adamc@152 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@152 8 *)
adamc@152 9
adamc@152 10 (* begin hide *)
adamc@156 11 Require Import Arith String List.
adamc@152 12
adamc@152 13 Require Import Tactics.
adamc@152 14
adamc@152 15 Set Implicit Arguments.
adamc@152 16 (* end hide *)
adamc@152 17
adamc@152 18
adamc@153 19 (** %\part{Formalizing Programming Languages and Compilers}
adamc@152 20
adamc@153 21 \chapter{First-Order Abstract Syntax}% *)
adamc@152 22
adamc@152 23 (** TODO: Prose for this chapter *)
adamc@152 24
adamc@152 25
adamc@152 26 (** * Concrete Binding *)
adamc@152 27
adamc@152 28 Module Concrete.
adamc@152 29
adamc@152 30 Definition var := string.
adamc@152 31 Definition var_eq := string_dec.
adamc@152 32
adamc@152 33 Inductive exp : Set :=
adamc@152 34 | Const : bool -> exp
adamc@152 35 | Var : var -> exp
adamc@152 36 | App : exp -> exp -> exp
adamc@152 37 | Abs : var -> exp -> exp.
adamc@152 38
adamc@152 39 Inductive type : Set :=
adamc@152 40 | Bool : type
adamc@152 41 | Arrow : type -> type -> type.
adamc@152 42
adamc@152 43 Infix "-->" := Arrow (right associativity, at level 60).
adamc@152 44
adamc@152 45 Definition ctx := list (var * type).
adamc@152 46
adamc@152 47 Reserved Notation "G |-v x : t" (no associativity, at level 90, x at next level).
adamc@152 48
adamc@152 49 Inductive lookup : ctx -> var -> type -> Prop :=
adamc@152 50 | First : forall x t G,
adamc@152 51 (x, t) :: G |-v x : t
adamc@152 52 | Next : forall x t x' t' G,
adamc@152 53 x <> x'
adamc@152 54 -> G |-v x : t
adamc@152 55 -> (x', t') :: G |-v x : t
adamc@152 56
adamc@152 57 where "G |-v x : t" := (lookup G x t).
adamc@152 58
adamc@152 59 Hint Constructors lookup.
adamc@152 60
adamc@152 61 Reserved Notation "G |-e e : t" (no associativity, at level 90, e at next level).
adamc@152 62
adamc@152 63 Inductive hasType : ctx -> exp -> type -> Prop :=
adamc@152 64 | TConst : forall G b,
adamc@152 65 G |-e Const b : Bool
adamc@152 66 | TVar : forall G v t,
adamc@152 67 G |-v v : t
adamc@152 68 -> G |-e Var v : t
adamc@152 69 | TApp : forall G e1 e2 dom ran,
adamc@152 70 G |-e e1 : dom --> ran
adamc@152 71 -> G |-e e2 : dom
adamc@152 72 -> G |-e App e1 e2 : ran
adamc@152 73 | TAbs : forall G x e' dom ran,
adamc@152 74 (x, dom) :: G |-e e' : ran
adamc@152 75 -> G |-e Abs x e' : dom --> ran
adamc@152 76
adamc@152 77 where "G |-e e : t" := (hasType G e t).
adamc@152 78
adamc@152 79 Hint Constructors hasType.
adamc@152 80
adamc@155 81 Lemma weaken_lookup : forall x t G' G1,
adamc@155 82 G1 |-v x : t
adamc@155 83 -> G1 ++ G' |-v x : t.
adamc@152 84 induction G1 as [ | [x' t'] tl ]; crush;
adamc@152 85 match goal with
adamc@152 86 | [ H : _ |-v _ : _ |- _ ] => inversion H; crush
adamc@152 87 end.
adamc@152 88 Qed.
adamc@152 89
adamc@152 90 Hint Resolve weaken_lookup.
adamc@152 91
adamc@152 92 Theorem weaken_hasType' : forall G' G e t,
adamc@152 93 G |-e e : t
adamc@155 94 -> G ++ G' |-e e : t.
adamc@152 95 induction 1; crush; eauto.
adamc@152 96 Qed.
adamc@152 97
adamc@155 98 Theorem weaken_hasType : forall e t,
adamc@155 99 nil |-e e : t
adamc@155 100 -> forall G', G' |-e e : t.
adamc@155 101 intros; change G' with (nil ++ G');
adamc@152 102 eapply weaken_hasType'; eauto.
adamc@152 103 Qed.
adamc@152 104
adamc@161 105 Hint Resolve weaken_hasType.
adamc@152 106
adamc@152 107 Section subst.
adamc@152 108 Variable x : var.
adamc@152 109 Variable e1 : exp.
adamc@152 110
adamc@152 111 Fixpoint subst (e2 : exp) : exp :=
adamc@152 112 match e2 with
adamc@152 113 | Const b => Const b
adamc@152 114 | Var x' =>
adamc@152 115 if var_eq x' x
adamc@152 116 then e1
adamc@152 117 else Var x'
adamc@152 118 | App e1 e2 => App (subst e1) (subst e2)
adamc@152 119 | Abs x' e' =>
adamc@152 120 Abs x' (if var_eq x' x
adamc@152 121 then e'
adamc@152 122 else subst e')
adamc@152 123 end.
adamc@152 124
adamc@152 125 Variable xt : type.
adamc@152 126 Hypothesis Ht' : nil |-e e1 : xt.
adamc@152 127
adamc@157 128 Notation "x # G" := (forall t' : type, In (x, t') G -> False) (no associativity, at level 90).
adamc@152 129
adamc@157 130 Lemma subst_lookup' : forall x' t,
adamc@152 131 x <> x'
adamc@155 132 -> forall G1, G1 ++ (x, xt) :: nil |-v x' : t
adamc@155 133 -> G1 |-v x' : t.
adamc@152 134 induction G1 as [ | [x'' t'] tl ]; crush;
adamc@152 135 match goal with
adamc@152 136 | [ H : _ |-v _ : _ |- _ ] => inversion H
adamc@152 137 end; crush.
adamc@152 138 Qed.
adamc@152 139
adamc@157 140 Hint Resolve subst_lookup'.
adamc@152 141
adamc@157 142 Lemma subst_lookup : forall x' t G1,
adamc@157 143 x' # G1
adamc@155 144 -> G1 ++ (x, xt) :: nil |-v x' : t
adamc@155 145 -> t = xt.
adamc@152 146 induction G1 as [ | [x'' t'] tl ]; crush; eauto;
adamc@152 147 match goal with
adamc@152 148 | [ H : _ |-v _ : _ |- _ ] => inversion H
adamc@155 149 end; crush; elimtype False; eauto;
adamc@155 150 match goal with
adamc@155 151 | [ H : nil |-v _ : _ |- _ ] => inversion H
adamc@155 152 end.
adamc@152 153 Qed.
adamc@152 154
adamc@157 155 Implicit Arguments subst_lookup [x' t G1].
adamc@152 156
adamc@155 157 Lemma shadow_lookup : forall v t t' G1,
adamc@152 158 G1 |-v x : t'
adamc@155 159 -> G1 ++ (x, xt) :: nil |-v v : t
adamc@155 160 -> G1 |-v v : t.
adamc@152 161 induction G1 as [ | [x'' t''] tl ]; crush;
adamc@152 162 match goal with
adamc@152 163 | [ H : nil |-v _ : _ |- _ ] => inversion H
adamc@152 164 | [ H1 : _ |-v _ : _, H2 : _ |-v _ : _ |- _ ] =>
adamc@152 165 inversion H1; crush; inversion H2; crush
adamc@152 166 end.
adamc@152 167 Qed.
adamc@152 168
adamc@155 169 Lemma shadow_hasType' : forall G e t,
adamc@152 170 G |-e e : t
adamc@155 171 -> forall G1, G = G1 ++ (x, xt) :: nil
adamc@152 172 -> forall t'', G1 |-v x : t''
adamc@155 173 -> G1 |-e e : t.
adamc@152 174 Hint Resolve shadow_lookup.
adamc@152 175
adamc@152 176 induction 1; crush; eauto;
adamc@152 177 match goal with
adamc@152 178 | [ H : (?x0, _) :: _ ++ (x, _) :: _ |-e _ : _ |- _ ] =>
adamc@152 179 destruct (var_eq x0 x); subst; eauto
adamc@152 180 end.
adamc@155 181 Qed.
adamc@152 182
adamc@155 183 Lemma shadow_hasType : forall G1 e t t'',
adamc@155 184 G1 ++ (x, xt) :: nil |-e e : t
adamc@152 185 -> G1 |-v x : t''
adamc@155 186 -> G1 |-e e : t.
adamc@152 187 intros; eapply shadow_hasType'; eauto.
adamc@152 188 Qed.
adamc@152 189
adamc@152 190 Hint Resolve shadow_hasType.
adamc@152 191
adamc@157 192 Lemma disjoint_cons : forall x x' t (G : ctx),
adamc@157 193 x # G
adamc@157 194 -> x' <> x
adamc@157 195 -> x # (x', t) :: G.
adamc@157 196 firstorder;
adamc@157 197 match goal with
adamc@157 198 | [ H : (_, _) = (_, _) |- _ ] => injection H
adamc@157 199 end; crush.
adamc@157 200 Qed.
adamc@157 201
adamc@157 202 Hint Resolve disjoint_cons.
adamc@157 203
adamc@152 204 Theorem subst_hasType : forall G e2 t,
adamc@152 205 G |-e e2 : t
adamc@155 206 -> forall G1, G = G1 ++ (x, xt) :: nil
adamc@157 207 -> x # G1
adamc@155 208 -> G1 |-e subst e2 : t.
adamc@152 209 induction 1; crush;
adamc@152 210 try match goal with
adamc@152 211 | [ |- context[if ?E then _ else _] ] => destruct E
adamc@152 212 end; crush; eauto 6;
adamc@152 213 match goal with
adamc@157 214 | [ H1 : x # _, H2 : _ |-v x : _ |- _ ] =>
adamc@157 215 rewrite (subst_lookup H1 H2)
adamc@152 216 end; crush.
adamc@152 217 Qed.
adamc@152 218
adamc@152 219 Theorem subst_hasType_closed : forall e2 t,
adamc@152 220 (x, xt) :: nil |-e e2 : t
adamc@152 221 -> nil |-e subst e2 : t.
adamc@155 222 intros; eapply subst_hasType; eauto.
adamc@152 223 Qed.
adamc@152 224 End subst.
adamc@152 225
adamc@154 226 Hint Resolve subst_hasType_closed.
adamc@154 227
adamc@154 228 Notation "[ x ~> e1 ] e2" := (subst x e1 e2) (no associativity, at level 80).
adamc@154 229
adamc@154 230 Inductive val : exp -> Prop :=
adamc@154 231 | VConst : forall b, val (Const b)
adamc@154 232 | VAbs : forall x e, val (Abs x e).
adamc@154 233
adamc@154 234 Hint Constructors val.
adamc@154 235
adamc@154 236 Reserved Notation "e1 ==> e2" (no associativity, at level 90).
adamc@154 237
adamc@154 238 Inductive step : exp -> exp -> Prop :=
adamc@154 239 | Beta : forall x e1 e2,
adamc@161 240 val e2
adamc@161 241 -> App (Abs x e1) e2 ==> [x ~> e2] e1
adamc@154 242 | Cong1 : forall e1 e2 e1',
adamc@154 243 e1 ==> e1'
adamc@154 244 -> App e1 e2 ==> App e1' e2
adamc@154 245 | Cong2 : forall e1 e2 e2',
adamc@154 246 val e1
adamc@154 247 -> e2 ==> e2'
adamc@154 248 -> App e1 e2 ==> App e1 e2'
adamc@154 249
adamc@154 250 where "e1 ==> e2" := (step e1 e2).
adamc@154 251
adamc@154 252 Hint Constructors step.
adamc@154 253
adamc@155 254 Lemma progress' : forall G e t, G |-e e : t
adamc@154 255 -> G = nil
adamc@154 256 -> val e \/ exists e', e ==> e'.
adamc@154 257 induction 1; crush; eauto;
adamc@154 258 try match goal with
adamc@154 259 | [ H : _ |-e _ : _ --> _ |- _ ] => inversion H
adamc@154 260 end;
adamc@154 261 repeat match goal with
adamc@154 262 | [ H : _ |- _ ] => solve [ inversion H; crush; eauto ]
adamc@154 263 end.
adamc@154 264 Qed.
adamc@154 265
adamc@155 266 Theorem progress : forall e t, nil |-e e : t
adamc@155 267 -> val e \/ exists e', e ==> e'.
adamc@155 268 intros; eapply progress'; eauto.
adamc@155 269 Qed.
adamc@155 270
adamc@155 271 Lemma preservation' : forall G e t, G |-e e : t
adamc@154 272 -> G = nil
adamc@154 273 -> forall e', e ==> e'
adamc@155 274 -> nil |-e e' : t.
adamc@154 275 induction 1; inversion 2; crush; eauto;
adamc@154 276 match goal with
adamc@154 277 | [ H : _ |-e Abs _ _ : _ |- _ ] => inversion H
adamc@154 278 end; eauto.
adamc@154 279 Qed.
adamc@154 280
adamc@155 281 Theorem preservation : forall e t, nil |-e e : t
adamc@155 282 -> forall e', e ==> e'
adamc@155 283 -> nil |-e e' : t.
adamc@155 284 intros; eapply preservation'; eauto.
adamc@155 285 Qed.
adamc@155 286
adamc@152 287 End Concrete.
adamc@156 288
adamc@156 289
adamc@156 290 (** * De Bruijn Indices *)
adamc@156 291
adamc@156 292 Module DeBruijn.
adamc@156 293
adamc@156 294 Definition var := nat.
adamc@156 295 Definition var_eq := eq_nat_dec.
adamc@156 296
adamc@156 297 Inductive exp : Set :=
adamc@156 298 | Const : bool -> exp
adamc@156 299 | Var : var -> exp
adamc@156 300 | App : exp -> exp -> exp
adamc@156 301 | Abs : exp -> exp.
adamc@156 302
adamc@156 303 Inductive type : Set :=
adamc@156 304 | Bool : type
adamc@156 305 | Arrow : type -> type -> type.
adamc@156 306
adamc@156 307 Infix "-->" := Arrow (right associativity, at level 60).
adamc@156 308
adamc@156 309 Definition ctx := list type.
adamc@156 310
adamc@156 311 Reserved Notation "G |-v x : t" (no associativity, at level 90, x at next level).
adamc@156 312
adamc@156 313 Inductive lookup : ctx -> var -> type -> Prop :=
adamc@156 314 | First : forall t G,
adamc@156 315 t :: G |-v O : t
adamc@156 316 | Next : forall x t t' G,
adamc@156 317 G |-v x : t
adamc@156 318 -> t' :: G |-v S x : t
adamc@156 319
adamc@156 320 where "G |-v x : t" := (lookup G x t).
adamc@156 321
adamc@156 322 Hint Constructors lookup.
adamc@156 323
adamc@156 324 Reserved Notation "G |-e e : t" (no associativity, at level 90, e at next level).
adamc@156 325
adamc@156 326 Inductive hasType : ctx -> exp -> type -> Prop :=
adamc@156 327 | TConst : forall G b,
adamc@156 328 G |-e Const b : Bool
adamc@156 329 | TVar : forall G v t,
adamc@156 330 G |-v v : t
adamc@156 331 -> G |-e Var v : t
adamc@156 332 | TApp : forall G e1 e2 dom ran,
adamc@156 333 G |-e e1 : dom --> ran
adamc@156 334 -> G |-e e2 : dom
adamc@156 335 -> G |-e App e1 e2 : ran
adamc@156 336 | TAbs : forall G e' dom ran,
adamc@156 337 dom :: G |-e e' : ran
adamc@156 338 -> G |-e Abs e' : dom --> ran
adamc@156 339
adamc@156 340 where "G |-e e : t" := (hasType G e t).
adamc@156 341
adamc@156 342 Hint Constructors hasType.
adamc@156 343
adamc@156 344 Lemma weaken_lookup : forall G' v t G,
adamc@156 345 G |-v v : t
adamc@156 346 -> G ++ G' |-v v : t.
adamc@156 347 induction 1; crush.
adamc@156 348 Qed.
adamc@156 349
adamc@156 350 Hint Resolve weaken_lookup.
adamc@156 351
adamc@156 352 Theorem weaken_hasType' : forall G' G e t,
adamc@156 353 G |-e e : t
adamc@156 354 -> G ++ G' |-e e : t.
adamc@156 355 induction 1; crush; eauto.
adamc@156 356 Qed.
adamc@156 357
adamc@156 358 Theorem weaken_hasType : forall e t,
adamc@156 359 nil |-e e : t
adamc@156 360 -> forall G', G' |-e e : t.
adamc@156 361 intros; change G' with (nil ++ G');
adamc@156 362 eapply weaken_hasType'; eauto.
adamc@156 363 Qed.
adamc@156 364
adamc@161 365 Hint Resolve weaken_hasType.
adamc@156 366
adamc@156 367 Section subst.
adamc@156 368 Variable e1 : exp.
adamc@156 369
adamc@156 370 Fixpoint subst (x : var) (e2 : exp) : exp :=
adamc@156 371 match e2 with
adamc@156 372 | Const b => Const b
adamc@156 373 | Var x' =>
adamc@156 374 if var_eq x' x
adamc@156 375 then e1
adamc@156 376 else Var x'
adamc@156 377 | App e1 e2 => App (subst x e1) (subst x e2)
adamc@156 378 | Abs e' => Abs (subst (S x) e')
adamc@156 379 end.
adamc@156 380
adamc@156 381 Variable xt : type.
adamc@156 382
adamc@156 383 Lemma subst_eq : forall t G1,
adamc@156 384 G1 ++ xt :: nil |-v length G1 : t
adamc@156 385 -> t = xt.
adamc@156 386 induction G1; inversion 1; crush.
adamc@156 387 Qed.
adamc@156 388
adamc@156 389 Implicit Arguments subst_eq [t G1].
adamc@156 390
adamc@156 391 Lemma subst_eq' : forall t G1 x,
adamc@156 392 G1 ++ xt :: nil |-v x : t
adamc@156 393 -> x <> length G1
adamc@156 394 -> G1 |-v x : t.
adamc@156 395 induction G1; inversion 1; crush;
adamc@156 396 match goal with
adamc@156 397 | [ H : nil |-v _ : _ |- _ ] => inversion H
adamc@156 398 end.
adamc@156 399 Qed.
adamc@156 400
adamc@156 401 Hint Resolve subst_eq'.
adamc@156 402
adamc@156 403 Lemma subst_neq : forall v t G1,
adamc@156 404 G1 ++ xt :: nil |-v v : t
adamc@156 405 -> v <> length G1
adamc@156 406 -> G1 |-e Var v : t.
adamc@156 407 induction G1; inversion 1; crush.
adamc@156 408 Qed.
adamc@156 409
adamc@156 410 Hint Resolve subst_neq.
adamc@156 411
adamc@156 412 Hypothesis Ht' : nil |-e e1 : xt.
adamc@156 413
adamc@156 414 Lemma hasType_push : forall dom G1 e' ran,
adamc@156 415 dom :: G1 |-e subst (length (dom :: G1)) e' : ran
adamc@156 416 -> dom :: G1 |-e subst (S (length G1)) e' : ran.
adamc@156 417 trivial.
adamc@156 418 Qed.
adamc@156 419
adamc@156 420 Hint Resolve hasType_push.
adamc@156 421
adamc@156 422 Theorem subst_hasType : forall G e2 t,
adamc@156 423 G |-e e2 : t
adamc@156 424 -> forall G1, G = G1 ++ xt :: nil
adamc@156 425 -> G1 |-e subst (length G1) e2 : t.
adamc@156 426 induction 1; crush;
adamc@156 427 try match goal with
adamc@156 428 | [ |- context[if ?E then _ else _] ] => destruct E
adamc@156 429 end; crush; eauto 6;
adamc@156 430 try match goal with
adamc@156 431 | [ H : _ |-v _ : _ |- _ ] =>
adamc@156 432 rewrite (subst_eq H)
adamc@156 433 end; crush.
adamc@156 434 Qed.
adamc@156 435
adamc@156 436 Theorem subst_hasType_closed : forall e2 t,
adamc@156 437 xt :: nil |-e e2 : t
adamc@156 438 -> nil |-e subst O e2 : t.
adamc@156 439 intros; change O with (length (@nil type)); eapply subst_hasType; eauto.
adamc@156 440 Qed.
adamc@156 441 End subst.
adamc@156 442
adamc@156 443 Hint Resolve subst_hasType_closed.
adamc@156 444
adamc@156 445 Notation "[ x ~> e1 ] e2" := (subst e1 x e2) (no associativity, at level 80).
adamc@156 446
adamc@156 447 Inductive val : exp -> Prop :=
adamc@156 448 | VConst : forall b, val (Const b)
adamc@156 449 | VAbs : forall e, val (Abs e).
adamc@156 450
adamc@156 451 Hint Constructors val.
adamc@156 452
adamc@156 453 Reserved Notation "e1 ==> e2" (no associativity, at level 90).
adamc@156 454
adamc@156 455 Inductive step : exp -> exp -> Prop :=
adamc@156 456 | Beta : forall e1 e2,
adamc@161 457 val e2
adamc@161 458 -> App (Abs e1) e2 ==> [O ~> e2] e1
adamc@156 459 | Cong1 : forall e1 e2 e1',
adamc@156 460 e1 ==> e1'
adamc@156 461 -> App e1 e2 ==> App e1' e2
adamc@156 462 | Cong2 : forall e1 e2 e2',
adamc@156 463 val e1
adamc@156 464 -> e2 ==> e2'
adamc@156 465 -> App e1 e2 ==> App e1 e2'
adamc@156 466
adamc@156 467 where "e1 ==> e2" := (step e1 e2).
adamc@156 468
adamc@156 469 Hint Constructors step.
adamc@156 470
adamc@156 471 Lemma progress' : forall G e t, G |-e e : t
adamc@156 472 -> G = nil
adamc@156 473 -> val e \/ exists e', e ==> e'.
adamc@156 474 induction 1; crush; eauto;
adamc@156 475 try match goal with
adamc@156 476 | [ H : _ |-e _ : _ --> _ |- _ ] => inversion H
adamc@156 477 end;
adamc@156 478 repeat match goal with
adamc@156 479 | [ H : _ |- _ ] => solve [ inversion H; crush; eauto ]
adamc@156 480 end.
adamc@156 481 Qed.
adamc@156 482
adamc@156 483 Theorem progress : forall e t, nil |-e e : t
adamc@156 484 -> val e \/ exists e', e ==> e'.
adamc@156 485 intros; eapply progress'; eauto.
adamc@156 486 Qed.
adamc@156 487
adamc@156 488 Lemma preservation' : forall G e t, G |-e e : t
adamc@156 489 -> G = nil
adamc@156 490 -> forall e', e ==> e'
adamc@156 491 -> nil |-e e' : t.
adamc@156 492 induction 1; inversion 2; crush; eauto;
adamc@156 493 match goal with
adamc@156 494 | [ H : _ |-e Abs _ : _ |- _ ] => inversion H
adamc@156 495 end; eauto.
adamc@156 496 Qed.
adamc@156 497
adamc@156 498 Theorem preservation : forall e t, nil |-e e : t
adamc@156 499 -> forall e', e ==> e'
adamc@156 500 -> nil |-e e' : t.
adamc@156 501 intros; eapply preservation'; eauto.
adamc@156 502 Qed.
adamc@156 503
adamc@156 504 End DeBruijn.