annotate book/src/StackMachine.v @ 9:c2e8e9c20643

Stack Machine initial source language
author Adam Chlipala <adamc@hcoop.net>
date Mon, 01 Sep 2008 11:18:58 -0400
parents f913d32a49e4
children d8363f2a3bef
rev   line source
adamc@2 1 (* Copyright (c) 2008, Adam Chlipala
adamc@2 2 *
adamc@2 3 * This work is licensed under a
adamc@2 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@2 5 * Unported License.
adamc@2 6 * The license text is available at:
adamc@2 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@2 8 *)
adamc@2 9
adamc@3 10 (* begin hide *)
adamc@2 11 Require Import List.
adamc@2 12
adamc@2 13 Require Import Tactics.
adamc@3 14 (* end hide *)
adamc@2 15
adamc@2 16
adamc@9 17 (** I will start off by jumping right in to a fully-worked set of examples, building certified compilers from increasingly complicated source languages to stack machines. We will meet a few useful tactics and see how they can be used in manual proofs, and we will also see how easily these proofs can be automated instead. I assume that you have installed Coq and Proof General.
adamc@9 18
adamc@9 19 As always, you can step through the source file %\texttt{%#<tt>#StackMachine.v#</tt>#%}% for this chapter interactively in Proof General. Alternatively, to get a feel for the whole lifecycle of creating a Coq development, you can enter the pieces of source code in this chapter in a new %\texttt{%#<tt>#.v#</tt>#%}% file in an Emacs buffer. If you do the latter, include a line [Require Import List Tactics] at the start of the file, to match some code hidden from the chapter source, and be sure to run the Coq binary %\texttt{%#<tt>#coqtop#</tt>#%}% with the command-line argument %\texttt{%#<tt>#-I SRC#</tt>#%}%, where %\texttt{%#<tt>#SRC#</tt>#%}% is the path to a directory containing the source for this book. In either case, if you have installed Proof General properly, it should start automatically when you visit a %\texttt{%#<tt>#.v#</tt>#%}% buffer in Emacs. *)
adamc@9 20
adamc@9 21
adamc@2 22 (** * Arithmetic expressions over natural numbers *)
adamc@2 23
adamc@9 24 (** We will begin with that staple of compiler textbooks, arithemtic expressions over a single type of numbers. *)
adamc@9 25
adamc@9 26 (** ** Source language *)
adamc@9 27
adamc@9 28 (** We begin with the syntax of the source language. *)
adamc@2 29
adamc@4 30 Inductive binop : Set := Plus | Times.
adamc@2 31
adamc@9 32 (** Our first line of Coq code should be unsurprising to ML and Haskell programmers. We define an algebraic datatype [binop] to stand for the binary operators of our source language. There are just two wrinkles compared to ML and Haskell. First, we use the keyword [Inductive], in place of %\texttt{%#<tt>#data#</tt>#%}%, %\texttt{%#<tt>#datatype#</tt>#%}%, or %\texttt{%#<tt>#type#</tt>#%}%. This is not just a trivial surface syntax difference; inductive types in Coq are much more expressive than garden variety algebraic datatypes, essentially enabling us to encode all of mathematics, though we begin humbly in this chapter. Second, there is the [: Set] fragment, which declares that we are defining a datatype that should be thought of as a constituent of programs. Later, we will see other options for defining datatypes in the universe of proofs or in an infinite hierarchy of universes, encompassing both programs and proofs, that is useful in higher-order constructions. *)
adamc@9 33
adamc@4 34 Inductive exp : Set :=
adamc@4 35 | Const : nat -> exp
adamc@4 36 | Binop : binop -> exp -> exp -> exp.
adamc@2 37
adamc@9 38 (** Now we define the type of arithmetic expressions. We write that a constant may be built from one argument, a natural number; and a binary operation may be built from a choice of operator and two operand expressions.
adamc@9 39
adamc@9 40 A note for readers following along in the PDF version: coqdoc supports pretty-printing of tokens in LaTeX or HTML. Where you see a right arrow character, the source contains the ASCII text %\texttt{%#<tt>#->#</tt>#%}%. Other examples of this substitution appearing in this chapter are a double right arrow for %\texttt{%#<tt>#=>#</tt>#%}% and the inverted 'A' symbol for %\texttt{%#<tt>#forall#</tt>#%}%. When in doubt about the ASCII version of a symbol, you can consult the chapter source code.
adamc@9 41
adamc@9 42 %\medskip%
adamc@9 43
adamc@9 44 Now we are ready to say what these programs mean. We will do this by writing an interpreter that can be thought of as a trivial operational or denotational semantics. (If you are not familiar with these semantic techniques, no need to worry; we will stick to "common sense" constructions.) *)
adamc@9 45
adamc@4 46 Definition binopDenote (b : binop) : nat -> nat -> nat :=
adamc@4 47 match b with
adamc@4 48 | Plus => plus
adamc@4 49 | Times => mult
adamc@4 50 end.
adamc@2 51
adamc@9 52 (** The meaning of a binary operator is a binary function over naturals, defined with pattern-matching notation analogous to the %\texttt{%#<tt>#case#</tt>#%}% and %\texttt{%#<tt>#match#</tt>#%}% of ML and Haskell, and referring to the functions [plus] and [mult] from the Coq standard library. The keyword [Definition] is Coq's all-purpose notation for binding a term of the programming language to a name, with some associated syntactic sugar, like the notation we see here for defining a function. That sugar could be expanded to yield this definition:
adamc@9 53
adamc@9 54 [[
adamc@9 55 Definition binopDenote : binop -> nat -> nat -> nat := fun (b : binop) =>
adamc@9 56 match b with
adamc@9 57 | Plus => plus
adamc@9 58 | Times => mult
adamc@9 59 end.
adamc@9 60
adamc@9 61 In this example, we could also omit all of the type annotations, arriving at:
adamc@9 62
adamc@9 63 [[
adamc@9 64 Definition binopDenote := fun b =>
adamc@9 65 match b with
adamc@9 66 | Plus => plus
adamc@9 67 | Times => mult
adamc@9 68 end.
adamc@9 69
adamc@9 70 Languages like Haskell and ML have a convenient %\textit{%#<i>#principal typing#</i>#%}% property, which gives us strong guarantees about how effective type inference will be. Unfortunately, Coq's type system is so expressive that any kind of "complete" type inference is impossible, and the task even seems to be hard heuristically in practice. Nonetheless, Coq includes some very helpful heuristics, many of them copying the workings of Haskell and ML type-checkers for programs that fall in simple fragments of Coq's language.
adamc@9 71
adamc@9 72 This is as good a time as any to mention the preponderance of different languages associated with Coq. The theoretical foundation of Coq is a formal system called the %\textit{%#<i>#Calculus of Inductive Constructions (CIC)#</i>#%}%, which is an extension of the older %\textit{%#<i>#Calculus of Constructions (CoC)#</i>#%}%. CIC is quite a spartan foundation, which is helpful for proving metatheory but not so helpful for real development. Still, it is nice to know that it has been proved that CIC enjoys properties like %\textit{%#<i>#strong normalization#</i>#%}%, meaning that every program (and, more importantly, every proof term) terminates; and %\textit{%#<i>#relative consistency#</i>#%}% with systems like versions of Zermelo Fraenkel set theory, which roughly means that you can believe that Coq proofs mean that the corresponding propositions are "really true," if you believe in set theory.
adamc@9 73
adamc@9 74 Coq is actually based on an extension of CIC called %\textit{%#<i>#Gallina#</i>#%}%. The text after the [:=] and before the period in the last code example is a term of Gallina. Gallina adds many useful features that are not compiled internalluy to more primitive CIC features. The important metatheorems about CIC have not been extended to the full breadth of these features, but most Coq users do not seem to lose much sleep over this omission.
adamc@9 75
adamc@9 76 Commands like [Inductive] and [Definition] are part of %\textit{%#<i>#the vernacular#</i>#%}%, which includes all sorts of useful queries and requests to the Coq system.
adamc@9 77
adamc@9 78 Finally, there is %\textit{%#<i>#Ltac#</i>#%}%, Coq's domain-specific language for writing proofs and decision procedures. We will see some basic examples of Ltac later in this chapter, and much of this book is devoted to more involved Ltac examples.
adamc@9 79
adamc@9 80 %\medskip%
adamc@9 81
adamc@9 82 We can give a simple definition of the meaning of an expression: *)
adamc@9 83
adamc@4 84 Fixpoint expDenote (e : exp) : nat :=
adamc@4 85 match e with
adamc@4 86 | Const n => n
adamc@4 87 | Binop b e1 e2 => (binopDenote b) (expDenote e1) (expDenote e2)
adamc@4 88 end.
adamc@2 89
adamc@9 90 (** We declare explicitly that this is a recursive definition, using the keyword [Fixpoint]. The rest should be old hat for functional programmers. *)
adamc@2 91
adamc@9 92
adamc@9 93 (** ** Target language *)
adamc@2 94
adamc@4 95 Inductive instr : Set :=
adamc@4 96 | IConst : nat -> instr
adamc@4 97 | IBinop : binop -> instr.
adamc@2 98
adamc@4 99 Definition prog := list instr.
adamc@4 100 Definition stack := list nat.
adamc@2 101
adamc@4 102 Definition instrDenote (i : instr) (s : stack) : option stack :=
adamc@4 103 match i with
adamc@4 104 | IConst n => Some (n :: s)
adamc@4 105 | IBinop b =>
adamc@4 106 match s with
adamc@4 107 | arg1 :: arg2 :: s' => Some ((binopDenote b) arg1 arg2 :: s')
adamc@4 108 | _ => None
adamc@4 109 end
adamc@4 110 end.
adamc@2 111
adamc@4 112 Fixpoint progDenote (p : prog) (s : stack) {struct p} : option stack :=
adamc@4 113 match p with
adamc@4 114 | nil => Some s
adamc@4 115 | i :: p' =>
adamc@4 116 match instrDenote i s with
adamc@4 117 | None => None
adamc@4 118 | Some s' => progDenote p' s'
adamc@4 119 end
adamc@4 120 end.
adamc@2 121
adamc@2 122
adamc@9 123 (** ** Translation *)
adamc@2 124
adamc@4 125 Fixpoint compile (e : exp) : prog :=
adamc@4 126 match e with
adamc@4 127 | Const n => IConst n :: nil
adamc@4 128 | Binop b e1 e2 => compile e2 ++ compile e1 ++ IBinop b :: nil
adamc@4 129 end.
adamc@2 130
adamc@2 131
adamc@9 132 (** ** Translation correctness *)
adamc@2 133
adamc@4 134 Lemma compileCorrect' : forall e s p, progDenote (compile e ++ p) s =
adamc@4 135 progDenote p (expDenote e :: s).
adamc@4 136 induction e.
adamc@2 137
adamc@4 138 intros.
adamc@4 139 unfold compile.
adamc@4 140 unfold expDenote.
adamc@4 141 simpl.
adamc@4 142 reflexivity.
adamc@2 143
adamc@4 144 intros.
adamc@4 145 unfold compile.
adamc@4 146 fold compile.
adamc@4 147 unfold expDenote.
adamc@4 148 fold expDenote.
adamc@4 149 rewrite app_ass.
adamc@4 150 rewrite IHe2.
adamc@4 151 rewrite app_ass.
adamc@4 152 rewrite IHe1.
adamc@4 153 simpl.
adamc@4 154 reflexivity.
adamc@4 155 Abort.
adamc@2 156
adamc@4 157 Lemma compileCorrect' : forall e s p, progDenote (compile e ++ p) s =
adamc@4 158 progDenote p (expDenote e :: s).
adamc@4 159 induction e; crush.
adamc@4 160 Qed.
adamc@2 161
adamc@4 162 Theorem compileCorrect : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
adamc@4 163 intro.
adamc@4 164 rewrite (app_nil_end (compile e)).
adamc@4 165 rewrite compileCorrect'.
adamc@4 166 reflexivity.
adamc@4 167 Qed.