annotate src/Hoas.v @ 159:8b2b652ab0ee

Add Plus
author Adam Chlipala <adamc@hcoop.net>
date Mon, 03 Nov 2008 09:50:22 -0500
parents fabfaa93c9ea
children 56e205f966cc
rev   line source
adamc@158 1 (* Copyright (c) 2008, 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@158 11 Require Import Arith Eqdep String List.
adamc@158 12
adamc@158 13 Require Import 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@158 49 Definition Exp1 t1 t2 := forall var, var t1 -> exp var t2.
adamc@158 50
adamc@159 51 Definition Const (n : nat) : Exp Nat :=
adamc@159 52 fun _ => Const' n.
adamc@159 53 Definition Plus (E1 E2 : Exp Nat) : Exp Nat :=
adamc@159 54 fun _ => Plus' (E1 _) (E2 _).
adamc@158 55 Definition App dom ran (F : Exp (dom --> ran)) (X : Exp dom) : Exp ran :=
adamc@158 56 fun _ => App' (F _) (X _).
adamc@158 57 Definition Abs dom ran (B : Exp1 dom ran) : Exp (dom --> ran) :=
adamc@158 58 fun _ => Abs' (B _).
adamc@158 59
adamc@158 60 Section flatten.
adamc@158 61 Variable var : type -> Type.
adamc@158 62
adamc@158 63 Fixpoint flatten t (e : exp (exp var) t) {struct e} : exp var t :=
adamc@158 64 match e in exp _ t return exp _ t with
adamc@159 65 | Const' n => Const' n
adamc@159 66 | Plus' e1 e2 => Plus' (flatten e1) (flatten e2)
adamc@158 67 | Var _ e' => e'
adamc@158 68 | App' _ _ e1 e2 => App' (flatten e1) (flatten e2)
adamc@158 69 | Abs' _ _ e' => Abs' (fun x => flatten (e' (Var x)))
adamc@158 70 end.
adamc@158 71 End flatten.
adamc@158 72
adamc@158 73 Definition Subst t1 t2 (E1 : Exp t1) (E2 : Exp1 t1 t2) : Exp t2 := fun _ =>
adamc@158 74 flatten (E2 _ (E1 _)).
adamc@158 75
adamc@158 76
adamc@158 77 (** * A Type Soundness Proof *)
adamc@158 78
adamc@158 79 Reserved Notation "E1 ==> E2" (no associativity, at level 90).
adamc@158 80
adamc@158 81 Inductive Val : forall t, Exp t -> Prop :=
adamc@159 82 | VConst : forall n, Val (Const n)
adamc@158 83 | VAbs : forall dom ran (B : Exp1 dom ran), Val (Abs B).
adamc@158 84
adamc@158 85 Hint Constructors Val.
adamc@158 86
adamc@158 87 Inductive Step : forall t, Exp t -> Exp t -> Prop :=
adamc@158 88 | Beta : forall dom ran (B : Exp1 dom ran) (X : Exp dom),
adamc@158 89 App (Abs B) X ==> Subst X B
adamc@159 90 | AppCong1 : forall dom ran (F : Exp (dom --> ran)) (X : Exp dom) F',
adamc@158 91 F ==> F'
adamc@158 92 -> App F X ==> App F' X
adamc@159 93 | AppCong2 : forall dom ran (F : Exp (dom --> ran)) (X : Exp dom) X',
adamc@158 94 Val F
adamc@158 95 -> X ==> X'
adamc@158 96 -> App F X ==> App F X'
adamc@158 97
adamc@159 98 | Sum : forall n1 n2,
adamc@159 99 Plus (Const n1) (Const n2) ==> Const (n1 + n2)
adamc@159 100 | PlusCong1 : forall E1 E2 E1',
adamc@159 101 E1 ==> E1'
adamc@159 102 -> Plus E1 E2 ==> Plus E1' E2
adamc@159 103 | PlusCong2 : forall E1 E2 E2',
adamc@159 104 E2 ==> E2'
adamc@159 105 -> Plus E1 E2 ==> Plus E1 E2'
adamc@159 106
adamc@158 107 where "E1 ==> E2" := (Step E1 E2).
adamc@158 108
adamc@158 109 Hint Constructors Step.
adamc@158 110
adamc@158 111 Inductive Closed : forall t, Exp t -> Prop :=
adamc@158 112 | CConst : forall b,
adamc@158 113 Closed (Const b)
adamc@159 114 | CPlus : forall E1 E2,
adamc@159 115 Closed E1
adamc@159 116 -> Closed E2
adamc@159 117 -> Closed (Plus E1 E2)
adamc@158 118 | CApp : forall dom ran (E1 : Exp (dom --> ran)) E2,
adamc@158 119 Closed E1
adamc@158 120 -> Closed E2
adamc@158 121 -> Closed (App E1 E2)
adamc@158 122 | CAbs : forall dom ran (E1 : Exp1 dom ran),
adamc@158 123 Closed (Abs E1).
adamc@158 124
adamc@158 125 Axiom closed : forall t (E : Exp t), Closed E.
adamc@158 126
adamc@158 127 Ltac my_crush :=
adamc@158 128 crush;
adamc@158 129 repeat (match goal with
adamc@158 130 | [ H : _ |- _ ] => generalize (inj_pairT2 _ _ _ _ _ H); clear H
adamc@158 131 end; crush).
adamc@158 132
adamc@158 133 Lemma progress' : forall t (E : Exp t),
adamc@158 134 Closed E
adamc@158 135 -> Val E \/ exists E', E ==> E'.
adamc@158 136 induction 1; crush;
adamc@159 137 repeat match goal with
adamc@159 138 | [ H : Val _ |- _ ] => inversion H; []; clear H; my_crush
adamc@159 139 end; eauto.
adamc@158 140 Qed.
adamc@158 141
adamc@158 142 Theorem progress : forall t (E : Exp t),
adamc@158 143 Val E \/ exists E', E ==> E'.
adamc@158 144 intros; apply progress'; apply closed.
adamc@158 145 Qed.
adamc@158 146