Mercurial > cpdt > repo
comparison src/StackMachine.v @ 311:4cb3ba8604bc
Pass through first half of StackMachine, along with some reorganization of the build process
author | Adam Chlipala <adam@chlipala.net> |
---|---|
date | Mon, 29 Aug 2011 15:31:06 -0400 |
parents | d2cb78f54454 |
children | 495153a41819 |
comparison
equal
deleted
inserted
replaced
310:cf67b7d6ceac | 311:4cb3ba8604bc |
---|---|
29 | 29 |
30 (** We will begin with that staple of compiler textbooks, arithmetic expressions over a single type of numbers. *) | 30 (** We will begin with that staple of compiler textbooks, arithmetic expressions over a single type of numbers. *) |
31 | 31 |
32 (** ** Source Language *) | 32 (** ** Source Language *) |
33 | 33 |
34 (** We begin with the syntax of the source language. *) | 34 (** We begin with the syntax of the source language.%\index{Vernacular commands!Inductive}% *) |
35 | 35 |
36 Inductive binop : Set := Plus | Times. | 36 Inductive binop : Set := Plus | Times. |
37 | 37 |
38 (** 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. *) | 38 (** Our first line of Coq code should be unsurprising to ML and Haskell programmers. We define an %\index{algebraic datatypes}%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 %\index{Gallina terms!Set}%[: 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. *) |
39 | 39 |
40 Inductive exp : Set := | 40 Inductive exp : Set := |
41 | Const : nat -> exp | 41 | Const : nat -> exp |
42 | Binop : binop -> exp -> exp -> exp. | 42 | Binop : binop -> exp -> exp -> exp. |
43 | 43 |
44 (** 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. | 44 (** 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. |
45 | 45 |
46 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>#%}%, the inverted 'A' symbol for %\texttt{%#<tt>#forall#</tt>#%}%, and the Cartesian product 'X' for %\texttt{%#<tt>#*#</tt>#%}%. When in doubt about the ASCII version of a symbol, you can consult the chapter source code. | 46 A note for readers following along in the PDF version: %\index{coqdoc}%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>#%}%, the inverted %`%#'#A' symbol for %\texttt{%#<tt>#forall#</tt>#%}%, and the Cartesian product %`%#'#X' for %\texttt{%#<tt>#*#</tt>#%}%. When in doubt about the ASCII version of a symbol, you can consult the chapter source code. |
47 | 47 |
48 %\medskip% | 48 %\medskip% |
49 | 49 |
50 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.) *) | 50 Now we are ready to say what these programs mean. We will do this by writing an %\index{interpreters}%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.)%\index{Vernacular commands!Definition}% *) |
51 | 51 |
52 Definition binopDenote (b : binop) : nat -> nat -> nat := | 52 Definition binopDenote (b : binop) : nat -> nat -> nat := |
53 match b with | 53 match b with |
54 | Plus => plus | 54 | Plus => plus |
55 | Times => mult | 55 | Times => mult |
75 | Times => mult | 75 | Times => mult |
76 end. | 76 end. |
77 | 77 |
78 ]] | 78 ]] |
79 | 79 |
80 Languages like Haskell and ML have a convenient %\emph{%#<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. | 80 Languages like Haskell and ML have a convenient %\index{principal typing}\index{type inference}\emph{%#<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. |
81 | 81 |
82 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 %\emph{%#<i>#Calculus of Inductive Constructions (CIC)#</i>#%}%, which is an extension of the older %\emph{%#<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 %\emph{%#<i>#strong normalization#</i>#%}%, meaning that every program (and, more importantly, every proof term) terminates; and %\emph{%#<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. | 82 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 %\index{Calculus of Inductive Constructions}\index{CIC|see{Calculus of Inductive Constructions}}\emph{%#<i>#Calculus of Inductive Constructions (CIC)#</i>#%}~\cite{CIC}%, which is an extension of the older %\index{Calculus of Constructions}\index{CoC|see{Calculus of Constructions}}\emph{%#<i>#Calculus of Constructions (CoC)#</i>#%}~\cite{CoC}%. 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 %\index{strong normalization}\emph{%#<i>#strong normalization#</i>#%}~\cite{CIC}%, meaning that every program (and, more importantly, every proof term) terminates; and %\index{relative consistency}\emph{%#<i>#relative consistency#</i>#%}~\cite{SetsInTypes}% with systems like versions of %\index{Zermelo-Fraenkel set theory}%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. |
83 | 83 |
84 Coq is actually based on an extension of CIC called %\emph{%#<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 internally 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. | 84 Coq is actually based on an extension of CIC called %\index{Gallina}\emph{%#<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 internally 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. |
85 | 85 |
86 Next, there is %\emph{%#<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. | 86 Next, there is %\index{Ltac}\emph{%#<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. |
87 | 87 |
88 Finally, commands like [Inductive] and [Definition] are part of %\emph{%#<i>#the vernacular#</i>#%}%, which includes all sorts of useful queries and requests to the Coq system. Every Coq source file is a series of vernacular commands, where many command forms take arguments that are Gallina or Ltac programs. (Actually, Coq source files are more like %\emph{%#<i>#trees#</i>#%}% of vernacular commands, thanks to various nested scoping constructs.) | 88 Finally, commands like [Inductive] and [Definition] are part of %\index{Vernacular commands}\emph{%#<i>#the Vernacular#</i>#%}%, which includes all sorts of useful queries and requests to the Coq system. Every Coq source file is a series of vernacular commands, where many command forms take arguments that are Gallina or Ltac programs. (Actually, Coq source files are more like %\emph{%#<i>#trees#</i>#%}% of vernacular commands, thanks to various nested scoping constructs.) |
89 | 89 |
90 %\medskip% | 90 %\medskip% |
91 | 91 |
92 We can give a simple definition of the meaning of an expression: *) | 92 We can give a simple definition of the meaning of an expression:%\index{Vernacular commands!Fixpoint}% *) |
93 | 93 |
94 Fixpoint expDenote (e : exp) : nat := | 94 Fixpoint expDenote (e : exp) : nat := |
95 match e with | 95 match e with |
96 | Const n => n | 96 | Const n => n |
97 | Binop b e1 e2 => (binopDenote b) (expDenote e1) (expDenote e2) | 97 | Binop b e1 e2 => (binopDenote b) (expDenote e1) (expDenote e2) |
98 end. | 98 end. |
99 | 99 |
100 (** We declare explicitly that this is a recursive definition, using the keyword [Fixpoint]. The rest should be old hat for functional programmers. *) | 100 (** We declare explicitly that this is a recursive definition, using the keyword [Fixpoint]. The rest should be old hat for functional programmers. *) |
101 | 101 |
102 (** It is convenient to be able to test definitions before starting to prove things about them. We can verify that our semantics is sensible by evaluating some sample uses. *) | 102 (** It is convenient to be able to test definitions before starting to prove things about them. We can verify that our semantics is sensible by evaluating some sample uses, using the command %\index{Vernacular commands!Eval}%[Eval]. This command takes an argument expressing a %\index{reduction strategy}\emph{%#<i>#reduction strategy#</i>#%}%, or an %``%#"#order of evaluation.#"#%''% Unlike with ML, which hardcodes an %\emph{%#<i>#eager#</i>#%}% reduction strategy, or Haskell, which hardcodes a %\emph{%#<i>#lazy#</i>#%}% strategy, in Coq we are free to choose between these and many other orders of evaluation, because all Coq programs terminate. In fact, Coq silently checked %\index{termination checking}%termination of our [Fixpoint] definition above, using a simple heuristic based on monotonically decreasing size of arguments across recursive calls. |
103 | |
104 To return to our test evaluations, we run the [Eval] command using the [simpl] evaluation strategy, whose definition is best postponed until we have learned more about Coq's foundations, but which usually gets the job done. *) | |
103 | 105 |
104 Eval simpl in expDenote (Const 42). | 106 Eval simpl in expDenote (Const 42). |
105 (** [= 42 : nat] *) | 107 (** [= 42 : nat] *) |
106 | 108 |
107 Eval simpl in expDenote (Binop Plus (Const 2) (Const 2)). | 109 Eval simpl in expDenote (Binop Plus (Const 2) (Const 2)). |
113 (** ** Target Language *) | 115 (** ** Target Language *) |
114 | 116 |
115 (** We will compile our source programs onto a simple stack machine, whose syntax is: *) | 117 (** We will compile our source programs onto a simple stack machine, whose syntax is: *) |
116 | 118 |
117 Inductive instr : Set := | 119 Inductive instr : Set := |
118 | IConst : nat -> instr | 120 | iConst : nat -> instr |
119 | IBinop : binop -> instr. | 121 | iBinop : binop -> instr. |
120 | 122 |
121 Definition prog := list instr. | 123 Definition prog := list instr. |
122 Definition stack := list nat. | 124 Definition stack := list nat. |
123 | 125 |
124 (** An instruction either pushes a constant onto the stack or pops two arguments, applies a binary operator to them, and pushes the result onto the stack. A program is a list of instructions, and a stack is a list of natural numbers. | 126 (** An instruction either pushes a constant onto the stack or pops two arguments, applies a binary operator to them, and pushes the result onto the stack. A program is a list of instructions, and a stack is a list of natural numbers. |
125 | 127 |
126 We can give instructions meanings as functions from stacks to optional stacks, where running an instruction results in [None] in case of a stack underflow and results in [Some s'] when the result of execution is the new stack [s']. [::] is the %``%#"#list cons#"#%''% operator from the Coq standard library. *) | 128 We can give instructions meanings as functions from stacks to optional stacks, where running an instruction results in [None] in case of a stack underflow and results in [Some s'] when the result of execution is the new stack [s']. %\index{Gallina operators!::}%The infix operator [::] is %``%#"#list cons#"#%''% from the Coq standard library.%\index{Gallina terms!option}% *) |
127 | 129 |
128 Definition instrDenote (i : instr) (s : stack) : option stack := | 130 Definition instrDenote (i : instr) (s : stack) : option stack := |
129 match i with | 131 match i with |
130 | IConst n => Some (n :: s) | 132 | iConst n => Some (n :: s) |
131 | IBinop b => | 133 | iBinop b => |
132 match s with | 134 match s with |
133 | arg1 :: arg2 :: s' => Some ((binopDenote b) arg1 arg2 :: s') | 135 | arg1 :: arg2 :: s' => Some ((binopDenote b) arg1 arg2 :: s') |
134 | _ => None | 136 | _ => None |
135 end | 137 end |
136 end. | 138 end. |
137 | 139 |
138 (** With [instrDenote] defined, it is easy to define a function [progDenote], which iterates application of [instrDenote] through a whole program. | 140 (** With [instrDenote] defined, it is easy to define a function [progDenote], which iterates application of [instrDenote] through a whole program. *) |
139 | |
140 [[ | |
141 Fixpoint progDenote (p : prog) (s : stack) {struct p} : option stack := | |
142 match p with | |
143 | nil => Some s | |
144 | i :: p' => | |
145 match instrDenote i s with | |
146 | None => None | |
147 | Some s' => progDenote p' s' | |
148 end | |
149 end. | |
150 | |
151 ]] | |
152 | |
153 There is one interesting difference compared to our previous example of a [Fixpoint]. This recursive function takes two arguments, [p] and [s]. It is critical for the soundness of Coq that every program terminate, so a shallow syntactic termination check is imposed on every recursive function definition. One of the function parameters must be designated to decrease monotonically across recursive calls. That is, every recursive call must use a version of that argument that has been pulled out of the current argument by some number of [match] expressions. [expDenote] has only one argument, so we did not need to specify which of its arguments decreases. For [progDenote], we resolve the ambiguity by writing [{struct p}] to indicate that argument [p] decreases structurally. | |
154 | |
155 Recent versions of Coq will also infer a termination argument, so that we may write simply: *) | |
156 | 141 |
157 Fixpoint progDenote (p : prog) (s : stack) : option stack := | 142 Fixpoint progDenote (p : prog) (s : stack) : option stack := |
158 match p with | 143 match p with |
159 | nil => Some s | 144 | nil => Some s |
160 | i :: p' => | 145 | i :: p' => |
165 end. | 150 end. |
166 | 151 |
167 | 152 |
168 (** ** Translation *) | 153 (** ** Translation *) |
169 | 154 |
170 (** Our compiler itself is now unsurprising. [++] is the list concatenation operator from the Coq standard library. *) | 155 (** Our compiler itself is now unsurprising. The list concatenation operator %\index{Gallina operators!++}%[++] comes from the Coq standard library. *) |
171 | 156 |
172 Fixpoint compile (e : exp) : prog := | 157 Fixpoint compile (e : exp) : prog := |
173 match e with | 158 match e with |
174 | Const n => IConst n :: nil | 159 | Const n => iConst n :: nil |
175 | Binop b e1 e2 => compile e2 ++ compile e1 ++ IBinop b :: nil | 160 | Binop b e1 e2 => compile e2 ++ compile e1 ++ iBinop b :: nil |
176 end. | 161 end. |
177 | 162 |
178 | 163 |
179 (** Before we set about proving that this compiler is correct, we can try a few test runs, using our sample programs from earlier. *) | 164 (** Before we set about proving that this compiler is correct, we can try a few test runs, using our sample programs from earlier. *) |
180 | 165 |
181 Eval simpl in compile (Const 42). | 166 Eval simpl in compile (Const 42). |
182 (** [= IConst 42 :: nil : prog] *) | 167 (** [= iConst 42 :: nil : prog] *) |
183 | 168 |
184 Eval simpl in compile (Binop Plus (Const 2) (Const 2)). | 169 Eval simpl in compile (Binop Plus (Const 2) (Const 2)). |
185 (** [= IConst 2 :: IConst 2 :: IBinop Plus :: nil : prog] *) | 170 (** [= iConst 2 :: iConst 2 :: iBinop Plus :: nil : prog] *) |
186 | 171 |
187 Eval simpl in compile (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)). | 172 Eval simpl in compile (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)). |
188 (** [= IConst 7 :: IConst 2 :: IConst 2 :: IBinop Plus :: IBinop Times :: nil : prog] *) | 173 (** [= iConst 7 :: iConst 2 :: iConst 2 :: iBinop Plus :: iBinop Times :: nil : prog] *) |
189 | 174 |
190 (** We can also run our compiled programs and check that they give the right results. *) | 175 (** We can also run our compiled programs and check that they give the right results. *) |
191 | 176 |
192 Eval simpl in progDenote (compile (Const 42)) nil. | 177 Eval simpl in progDenote (compile (Const 42)) nil. |
193 (** [= Some (42 :: nil) : option stack] *) | 178 (** [= Some (42 :: nil) : option stack] *) |
194 | 179 |
195 Eval simpl in progDenote (compile (Binop Plus (Const 2) (Const 2))) nil. | 180 Eval simpl in progDenote (compile (Binop Plus (Const 2) (Const 2))) nil. |
196 (** [= Some (4 :: nil) : option stack] *) | 181 (** [= Some (4 :: nil) : option stack] *) |
197 | 182 |
198 Eval simpl in progDenote (compile (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7))) nil. | 183 Eval simpl in progDenote (compile (Binop Times (Binop Plus (Const 2) (Const 2)) |
184 (Const 7))) nil. | |
199 (** [= Some (28 :: nil) : option stack] *) | 185 (** [= Some (28 :: nil) : option stack] *) |
200 | 186 |
201 | 187 |
202 (** ** Translation Correctness *) | 188 (** ** Translation Correctness *) |
203 | 189 |
204 (** We are ready to prove that our compiler is implemented correctly. We can use a new vernacular command [Theorem] to start a correctness proof, in terms of the semantics we defined earlier: *) | 190 (** We are ready to prove that our compiler is implemented correctly. We can use a new vernacular command [Theorem] to start a correctness proof, in terms of the semantics we defined earlier:%\index{Vernacular commands!Theorem}% *) |
205 | 191 |
206 Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil). | 192 Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil). |
207 (* begin hide *) | 193 (* begin hide *) |
208 Abort. | 194 Abort. |
209 (* end hide *) | 195 (* end hide *) |
210 (* begin thide *) | 196 (* begin thide *) |
211 | 197 |
212 (** Though a pencil-and-paper proof might clock out at this point, writing %``%#"#by a routine induction on [e],#"#%''% it turns out not to make sense to attack this proof directly. We need to use the standard trick of %\emph{%#<i>#strengthening the induction hypothesis#</i>#%}%. We do that by proving an auxiliary lemma: | 198 (** Though a pencil-and-paper proof might clock out at this point, writing %``%#"#by a routine induction on [e],#"#%''% it turns out not to make sense to attack this proof directly. We need to use the standard trick of %\index{strengthening the induction hypothesis}\emph{%#<i>#strengthening the induction hypothesis#</i>#%}%. We do that by proving an auxiliary lemma, using the command [Lemma] that is a synonym for [Theorem], conventionally used for less important theorems that appear in the proofs of primary theorems.%\index{Vernacular commands!Lemma}% *) |
213 *) | |
214 | 199 |
215 Lemma compile_correct' : forall e p s, | 200 Lemma compile_correct' : forall e p s, |
216 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s). | 201 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s). |
217 | 202 |
218 (** After the period in the [Lemma] command, we are in %\emph{%#<i>#the interactive proof-editing mode#</i>#%}%. We find ourselves staring at this ominous screen of text: | 203 (** After the period in the [Lemma] command, we are in %\index{interactive proof-editing mode}\emph{%#<i>#the interactive proof-editing mode#</i>#%}%. We find ourselves staring at this ominous screen of text: |
219 | 204 |
220 [[ | 205 [[ |
221 1 subgoal | 206 1 subgoal |
222 | 207 |
223 ============================ | 208 ============================ |
224 forall (e : exp) (p : list instr) (s : stack), | 209 forall (e : exp) (p : list instr) (s : stack), |
225 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s) | 210 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s) |
226 | 211 |
227 ]] | 212 ]] |
228 | 213 |
229 Coq seems to be restating the lemma for us. What we are seeing is a limited case of a more general protocol for describing where we are in a proof. We are told that we have a single subgoal. In general, during a proof, we can have many pending subgoals, each of which is a logical proposition to prove. Subgoals can be proved in any order, but it usually works best to prove them in the order that Coq chooses. | 214 Coq seems to be restating the lemma for us. What we are seeing is a limited case of a more general protocol for describing where we are in a proof. We are told that we have a single subgoal. In general, during a proof, we can have many pending %\index{subgoals}%subgoals, each of which is a logical proposition to prove. Subgoals can be proved in any order, but it usually works best to prove them in the order that Coq chooses. |
230 | 215 |
231 Next in the output, we see our single subgoal described in full detail. There is a double-dashed line, above which would be our free variables and hypotheses, if we had any. Below the line is the conclusion, which, in general, is to be proved from the hypotheses. | 216 Next in the output, we see our single subgoal described in full detail. There is a double-dashed line, above which would be our free variables and %\index{hypotheses}%hypotheses, if we had any. Below the line is the %\index{conclusion}%conclusion, which, in general, is to be proved from the hypotheses. |
232 | 217 |
233 We manipulate the proof state by running commands called %\emph{%#<i>#tactics#</i>#%}%. Let us start out by running one of the most important tactics: | 218 We manipulate the proof state by running commands called %\index{tactics}\emph{%#<i>#tactics#</i>#%}%. Let us start out by running one of the most important tactics:%\index{tactics!induction}% |
234 *) | 219 *) |
235 | 220 |
236 induction e. | 221 induction e. |
237 | 222 |
238 (** We declare that this proof will proceed by induction on the structure of the expression [e]. This swaps out our initial subgoal for two new subgoals, one for each case of the inductive proof: | 223 (** We declare that this proof will proceed by induction on the structure of the expression [e]. This swaps out our initial subgoal for two new subgoals, one for each case of the inductive proof: |
239 | 224 |
240 [[ | 225 %\vspace{.1in} \noindent 2 \coqdockw{subgoals}\vspace{-.1in}%#<tt>2 subgoals</tt># |
241 2 subgoals | 226 |
242 | 227 [[ |
243 n : nat | 228 n : nat |
244 ============================ | 229 ============================ |
245 forall (s : stack) (p : list instr), | 230 forall (s : stack) (p : list instr), |
246 progDenote (compile (Const n) ++ p) s = | 231 progDenote (compile (Const n) ++ p) s = |
247 progDenote p (expDenote (Const n) :: s) | 232 progDenote p (expDenote (Const n) :: s) |
248 ]] | 233 ]] |
249 [[ | 234 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt># |
250 subgoal 2 is: | 235 [[ |
251 forall (s : stack) (p : list instr), | 236 forall (s : stack) (p : list instr), |
252 progDenote (compile (Binop b e1 e2) ++ p) s = | 237 progDenote (compile (Binop b e1 e2) ++ p) s = |
253 progDenote p (expDenote (Binop b e1 e2) :: s) | 238 progDenote p (expDenote (Binop b e1 e2) :: s) |
254 | 239 |
255 ]] | 240 ]] |
256 | 241 |
257 The first and current subgoal is displayed with the double-dashed line below free variables and hypotheses, while later subgoals are only summarized with their conclusions. We see an example of a free variable in the first subgoal; [n] is a free variable of type [nat]. The conclusion is the original theorem statement where [e] has been replaced by [Const n]. In a similar manner, the second case has [e] replaced by a generalized invocation of the [Binop] expression constructor. We can see that proving both cases corresponds to a standard proof by structural induction. | 242 The first and current subgoal is displayed with the double-dashed line below free variables and hypotheses, while later subgoals are only summarized with their conclusions. We see an example of a %\index{free variable}%free variable in the first subgoal; [n] is a free variable of type [nat]. The conclusion is the original theorem statement where [e] has been replaced by [Const n]. In a similar manner, the second case has [e] replaced by a generalized invocation of the [Binop] expression constructor. We can see that proving both cases corresponds to a standard proof by %\index{structural induction}%structural induction. |
258 | 243 |
259 We begin the first case with another very common tactic. | 244 We begin the first case with another very common tactic.%\index{tactics!intros}% |
260 *) | 245 *) |
261 | 246 |
262 intros. | 247 intros. |
263 | 248 |
264 (** The current subgoal changes to: | 249 (** The current subgoal changes to: |
273 | 258 |
274 ]] | 259 ]] |
275 | 260 |
276 We see that [intros] changes [forall]-bound variables at the beginning of a goal into free variables. | 261 We see that [intros] changes [forall]-bound variables at the beginning of a goal into free variables. |
277 | 262 |
278 To progress further, we need to use the definitions of some of the functions appearing in the goal. The [unfold] tactic replaces an identifier with its definition. | 263 To progress further, we need to use the definitions of some of the functions appearing in the goal. The [unfold] tactic replaces an identifier with its definition.%\index{tactics!unfold}% |
279 *) | 264 *) |
280 | 265 |
266 (* begin hide *) | |
281 unfold compile. | 267 unfold compile. |
268 (* end hide *) | |
269 (** %\coqdockw{unfold} \coqdocdefinition{compile}.%#<tt>unfold compile.</tt># *) | |
282 (** [[ | 270 (** [[ |
283 n : nat | 271 n : nat |
284 s : stack | 272 s : stack |
285 p : list instr | 273 p : list instr |
286 ============================ | 274 ============================ |
287 progDenote ((IConst n :: nil) ++ p) s = | 275 progDenote ((iConst n :: nil) ++ p) s = |
288 progDenote p (expDenote (Const n) :: s) | 276 progDenote p (expDenote (Const n) :: s) |
289 | 277 |
290 ]] | 278 ]] |
291 *) | 279 *) |
292 | 280 |
281 (* begin hide *) | |
293 unfold expDenote. | 282 unfold expDenote. |
294 | 283 (* end hide *) |
284 (** %\coqdockw{unfold} \coqdocdefinition{expDenote}.%#<tt>unfold expDenote.</tt># *) | |
295 (** [[ | 285 (** [[ |
296 n : nat | 286 n : nat |
297 s : stack | 287 s : stack |
298 p : list instr | 288 p : list instr |
299 ============================ | 289 ============================ |
300 progDenote ((IConst n :: nil) ++ p) s = progDenote p (n :: s) | 290 progDenote ((iConst n :: nil) ++ p) s = progDenote p (n :: s) |
301 | 291 |
302 ]] | 292 ]] |
303 | 293 |
304 We only need to unfold the first occurrence of [progDenote] to prove the goal: *) | 294 We only need to unfold the first occurrence of [progDenote] to prove the goal. An [at] clause used with [unfold] specifies a particular occurrence of an identifier to unfold, where we count occurrences from left to right.%\index{tactics!unfold}% *) |
305 | 295 |
296 (* begin hide *) | |
306 unfold progDenote at 1. | 297 unfold progDenote at 1. |
307 | 298 (* end hide *) |
308 (** [[ | 299 (** %\coqdockw{unfold} \coqdocdefinition{progDenote} \coqdoctac{at} 1.%#<tt>unfold progDenote at 1.</tt># *) |
309 | 300 (** [[ |
310 n : nat | 301 n : nat |
311 s : stack | 302 s : stack |
312 p : list instr | 303 p : list instr |
313 ============================ | 304 ============================ |
314 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} : | 305 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} : |
318 | i :: p' => | 309 | i :: p' => |
319 match instrDenote i s0 with | 310 match instrDenote i s0 with |
320 | Some s' => progDenote p' s' | 311 | Some s' => progDenote p' s' |
321 | None => None (A:=stack) | 312 | None => None (A:=stack) |
322 end | 313 end |
323 end) ((IConst n :: nil) ++ p) s = | 314 end) ((iConst n :: nil) ++ p) s = |
324 progDenote p (n :: s) | 315 progDenote p (n :: s) |
325 | 316 |
326 ]] | 317 ]] |
327 | 318 |
328 This last [unfold] has left us with an anonymous fixpoint version of [progDenote], which will generally happen when unfolding recursive definitions. Fortunately, in this case, we can eliminate such complications right away, since the structure of the argument [(IConst n :: nil) ++ p] is known, allowing us to simplify the internal pattern match with the [simpl] tactic: | 319 This last [unfold] has left us with an anonymous fixpoint version of [progDenote], which will generally happen when unfolding recursive definitions. Note that Coq has automatically renamed the [fix] arguments [p] and [s] to [p0] and [s0], to avoid clashes with our local free variables. There is also a subterm %\coqdocconstructor{None} (\coqdocvar{A}:=\coqdocdefinition{stack})%#<tt>None (A:=stack)</tt>#, which has an annotation specifying that the type of the term ought to be [option A]. This is phrased as an explicit instantiation of a named type parameter [A] from the definition of [option]. |
320 | |
321 Fortunately, in this case, we can eliminate the complications of anonymous recursion right away, since the structure of the argument ([iConst n :: nil) ++ p] is known, allowing us to simplify the internal pattern match with the [simpl] tactic, which applies the same reduction strategy that we used earlier with [Eval] (and whose details we still postpone).%\index{tactics!simpl}% | |
329 *) | 322 *) |
330 | 323 |
331 simpl. | 324 simpl. |
332 (** [[ | 325 (** [[ |
333 n : nat | 326 n : nat |
345 end | 338 end |
346 end) p (n :: s) = progDenote p (n :: s) | 339 end) p (n :: s) = progDenote p (n :: s) |
347 | 340 |
348 ]] | 341 ]] |
349 | 342 |
350 Now we can unexpand the definition of [progDenote]: | 343 Now we can unexpand the definition of [progDenote]:%\index{tactics!fold}% |
351 *) | 344 *) |
352 | 345 |
346 (* begin hide *) | |
353 fold progDenote. | 347 fold progDenote. |
354 | 348 (* end hide *) |
355 (** [[ | 349 (** %\coqdockw{fold} \coqdocdefinition{progDenote}.%#<tt>fold progDenote.</tt># *) |
356 | 350 (** [[ |
357 n : nat | 351 n : nat |
358 s : stack | 352 s : stack |
359 p : list instr | 353 p : list instr |
360 ============================ | 354 ============================ |
361 progDenote p (n :: s) = progDenote p (n :: s) | 355 progDenote p (n :: s) = progDenote p (n :: s) |
362 | 356 |
363 ]] | 357 ]] |
364 | 358 |
365 It looks like we are at the end of this case, since we have a trivial equality. Indeed, a single tactic finishes us off: | 359 It looks like we are at the end of this case, since we have a trivial equality. Indeed, a single tactic finishes us off:%\index{tactics!reflexivity}% |
366 *) | 360 *) |
367 | 361 |
368 reflexivity. | 362 reflexivity. |
369 | 363 |
370 (** On to the second inductive case: | 364 (** On to the second inductive case: |
382 progDenote (compile (Binop b e1 e2) ++ p) s = | 376 progDenote (compile (Binop b e1 e2) ++ p) s = |
383 progDenote p (expDenote (Binop b e1 e2) :: s) | 377 progDenote p (expDenote (Binop b e1 e2) :: s) |
384 | 378 |
385 ]] | 379 ]] |
386 | 380 |
387 We see our first example of hypotheses above the double-dashed line. They are the inductive hypotheses [IHe1] and [IHe2] corresponding to the subterms [e1] and [e2], respectively. | 381 We see our first example of %\index{hypotheses}%hypotheses above the double-dashed line. They are the inductive hypotheses [IHe1] and [IHe2] corresponding to the subterms [e1] and [e2], respectively. |
388 | 382 |
389 We start out the same way as before, introducing new free variables and unfolding and folding the appropriate definitions. The seemingly frivolous [unfold]/[fold] pairs are actually accomplishing useful work, because [unfold] will sometimes perform easy simplifications. *) | 383 We start out the same way as before, introducing new free variables and unfolding and folding the appropriate definitions. The seemingly frivolous [unfold]/%\coqdockw{fold}%#<tt>fold</tt># pairs are actually accomplishing useful work, because [unfold] will sometimes perform easy simplifications. %\index{tactics!intros}\index{tactics!unfold}\index{tactics!fold}% *) |
390 | 384 |
391 intros. | 385 intros. |
386 (* begin hide *) | |
392 unfold compile. | 387 unfold compile. |
393 fold compile. | 388 fold compile. |
394 unfold expDenote. | 389 unfold expDenote. |
395 fold expDenote. | 390 fold expDenote. |
391 (* end hide *) | |
392 (** %\coqdockw{unfold} \coqdocdefinition{compile}. | |
393 \coqdockw{fold} \coqdocdefinition{compile}. | |
394 \coqdockw{unfold} \coqdocdefinition{expDenote}. | |
395 \coqdockw{fold} \coqdocdefinition{expDenote}.% | |
396 #<tt>unfold compile. fold compile. unfold expDenote. fold expDenote.</tt># *) | |
396 | 397 |
397 (** Now we arrive at a point where the tactics we have seen so far are insufficient. No further definition unfoldings get us anywhere, so we will need to try something different. | 398 (** Now we arrive at a point where the tactics we have seen so far are insufficient. No further definition unfoldings get us anywhere, so we will need to try something different. |
398 | 399 |
399 [[ | 400 [[ |
400 b : binop | 401 b : binop |
405 IHe2 : forall (s : stack) (p : list instr), | 406 IHe2 : forall (s : stack) (p : list instr), |
406 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s) | 407 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s) |
407 s : stack | 408 s : stack |
408 p : list instr | 409 p : list instr |
409 ============================ | 410 ============================ |
410 progDenote ((compile e2 ++ compile e1 ++ IBinop b :: nil) ++ p) s = | 411 progDenote ((compile e2 ++ compile e1 ++ iBinop b :: nil) ++ p) s = |
411 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s) | 412 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s) |
412 | 413 |
413 ]] | 414 ]] |
414 | 415 |
415 What we need is the associative law of list concatenation, which is available as a theorem [app_ass] in the standard library. *) | 416 What we need is the associative law of list concatenation, which is available as a theorem [app_assoc_reverse] in the standard library.%\index{Vernacular commands!Check}% *) |
416 | 417 |
417 Check app_ass. | 418 Check app_assoc. |
418 (** [[ | 419 |
419 app_ass | 420 (** [[ |
421 app_assoc_reverse | |
420 : forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n | 422 : forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n |
421 | 423 |
422 ]] | 424 ]] |
423 | 425 |
424 If we did not already know the name of the theorem, we could use the [SearchRewrite] command to find it, based on a pattern that we would like to rewrite: *) | 426 If we did not already know the name of the theorem, we could use the %\index{Vernacular commands!SearchRewrite}\coqdockw{%#<tt>#SearchRewrite#</tt>#%}% command to find it, based on a pattern that we would like to rewrite: *) |
425 | 427 |
428 (* begin hide *) | |
426 SearchRewrite ((_ ++ _) ++ _). | 429 SearchRewrite ((_ ++ _) ++ _). |
427 (** [[ | 430 (* end hide *) |
428 app_ass: forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n | 431 (** %\coqdockw{%#<tt>#SearchRewrite#</tt>#%}% [((_ ++ _) ++ _).] *) |
429 ass_app: forall (A : Type) (l m n : list A), l ++ m ++ n = (l ++ m) ++ n | 432 (** [[ |
430 | 433 app_assoc_reverse: |
431 ]] | 434 forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n |
432 | 435 ]] |
433 We use it to perform a rewrite: *) | 436 %\vspace{-.25in}% |
434 | 437 [[ |
435 rewrite app_ass. | 438 app_assoc: forall (A : Type) (l m n : list A), l ++ m ++ n = (l ++ m) ++ n |
439 | |
440 ]] | |
441 | |
442 We use [app_assoc_reverse] to perform a rewrite: %\index{tactics!rewrite}% *) | |
443 | |
444 rewrite app_assoc_reverse. | |
436 | 445 |
437 (** changing the conclusion to: | 446 (** changing the conclusion to: |
438 | 447 |
439 [[ | 448 [[ |
440 progDenote (compile e2 ++ (compile e1 ++ IBinop b :: nil) ++ p) s = | 449 progDenote (compile e2 ++ (compile e1 ++ iBinop b :: nil) ++ p) s = |
441 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s) | 450 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s) |
442 | 451 |
443 ]] | 452 ]] |
444 | 453 |
445 Now we can notice that the lefthand side of the equality matches the lefthand side of the second inductive hypothesis, so we can rewrite with that hypothesis, too: *) | 454 Now we can notice that the lefthand side of the equality matches the lefthand side of the second inductive hypothesis, so we can rewrite with that hypothesis, too.%\index{tactics!rewrite}% *) |
446 | 455 |
447 rewrite IHe2. | 456 rewrite IHe2. |
448 (** [[ | 457 (** [[ |
449 progDenote ((compile e1 ++ IBinop b :: nil) ++ p) (expDenote e2 :: s) = | 458 progDenote ((compile e1 ++ iBinop b :: nil) ++ p) (expDenote e2 :: s) = |
450 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s) | 459 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s) |
451 | 460 |
452 ]] | 461 ]] |
453 | 462 |
454 The same process lets us apply the remaining hypothesis. *) | 463 The same process lets us apply the remaining hypothesis.%\index{tactics!rewrite}% *) |
455 | 464 |
456 rewrite app_ass. | 465 rewrite app_assoc_reverse. |
457 rewrite IHe1. | 466 rewrite IHe1. |
458 (** [[ | 467 (** [[ |
459 progDenote ((IBinop b :: nil) ++ p) (expDenote e1 :: expDenote e2 :: s) = | 468 progDenote ((iBinop b :: nil) ++ p) (expDenote e1 :: expDenote e2 :: s) = |
460 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s) | 469 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s) |
461 | 470 |
462 ]] | 471 ]] |
463 | 472 |
464 Now we can apply a similar sequence of tactics to that that ended the proof of the first case. | 473 Now we can apply a similar sequence of tactics to the one that ended the proof of the first case.%\index{tactics!unfold}\index{tactics!simpl}\index{tactics!fold}\index{tactics!reflexivity}% |
465 *) | 474 *) |
466 | 475 |
476 (* begin hide *) | |
467 unfold progDenote at 1. | 477 unfold progDenote at 1. |
468 simpl. | 478 simpl. |
469 fold progDenote. | 479 fold progDenote. |
470 reflexivity. | 480 reflexivity. |
471 | 481 (* end hide *) |
472 (** And the proof is completed, as indicated by the message: | 482 (** %\coqdockw{unfold} \coqdocdefinition{progDenote} \coqdoctac{at} 1. \coqdockw{simpl}. \coqdockw{fold} \coqdocdefinition{progDenote}. \coqdockw{reflexivity}.%#<tt>unfold progDenote at 1. simpl. fold progDenote. reflexivity.</tt># *) |
473 | 483 |
474 [[ | 484 (** And the proof is completed, as indicated by the message: *) |
475 Proof completed. | 485 |
476 | 486 (** %\coqdockw{Proof} \coqdockw{completed}.%#<tt>Proof completed.</tt># *) |
477 ]] | 487 |
478 | 488 (** And there lies our first proof. Already, even for simple theorems like this, the final proof script is unstructured and not very enlightening to readers. If we extend this approach to more serious theorems, we arrive at the unreadable proof scripts that are the favorite complaints of opponents of tactic-based proving. Fortunately, Coq has rich support for scripted automation, and we can take advantage of such a scripted tactic (defined elsewhere) to make short work of this lemma. We abort the old proof attempt and start again.%\index{Vernacular commands!Abort}% |
479 And there lies our first proof. Already, even for simple theorems like this, the final proof script is unstructured and not very enlightening to readers. If we extend this approach to more serious theorems, we arrive at the unreadable proof scripts that are the favorite complaints of opponents of tactic-based proving. Fortunately, Coq has rich support for scripted automation, and we can take advantage of such a scripted tactic (defined elsewhere) to make short work of this lemma. We abort the old proof attempt and start again. | |
480 *) | 489 *) |
481 | 490 |
482 Abort. | 491 Abort. |
492 | |
493 (** %\index{tactics!induction}\index{tactics!crush}% *) | |
483 | 494 |
484 Lemma compile_correct' : forall e s p, progDenote (compile e ++ p) s = | 495 Lemma compile_correct' : forall e s p, progDenote (compile e ++ p) s = |
485 progDenote p (expDenote e :: s). | 496 progDenote p (expDenote e :: s). |
486 induction e; crush. | 497 induction e; crush. |
487 Qed. | 498 Qed. |
488 | 499 |
489 (** We need only to state the basic inductive proof scheme and call a tactic that automates the tedious reasoning in between. In contrast to the period tactic terminator from our last proof, the semicolon tactic separator supports structured, compositional proofs. The tactic [t1; t2] has the effect of running [t1] and then running [t2] on each remaining subgoal. The semicolon is one of the most fundamental building blocks of effective proof automation. The period terminator is very useful for exploratory proving, where you need to see intermediate proof states, but final proofs of any serious complexity should have just one period, terminating a single compound tactic that probably uses semicolons. | 500 (** We need only to state the basic inductive proof scheme and call a tactic that automates the tedious reasoning in between. In contrast to the period tactic terminator from our last proof, the %\index{tacticals!semicolon}%semicolon tactic separator supports structured, compositional proofs. The tactic [t1; t2] has the effect of running [t1] and then running [t2] on each remaining subgoal. The semicolon is one of the most fundamental building blocks of effective proof automation. The period terminator is very useful for exploratory proving, where you need to see intermediate proof states, but final proofs of any serious complexity should have just one period, terminating a single compound tactic that probably uses semicolons. |
490 | 501 |
491 The [crush] tactic comes from the library associated with this book and is not part of the Coq standard library. The book's library contains a number of other tactics that are especially helpful in highly-automated proofs. | 502 The [crush] tactic comes from the library associated with this book and is not part of the Coq standard library. The book's library contains a number of other tactics that are especially helpful in highly-automated proofs. |
492 | 503 |
493 The proof of our main theorem is now easy. We prove it with four period-terminated tactics, though separating them with semicolons would work as well; the version here is easier to step through. *) | 504 The %\index{Vernacular commands!Qed}%[Qed] command checks that the proof is finished and, if so, saves it. |
505 | |
506 The proof of our main theorem is now easy. We prove it with four period-terminated tactics, though separating them with semicolons would work as well; the version here is easier to step through.%\index{tactics!intros}% *) | |
494 | 507 |
495 Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil). | 508 Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil). |
496 intros. | 509 intros. |
497 (** [[ | 510 (** [[ |
498 e : exp | 511 e : exp |
506 Check app_nil_end. | 519 Check app_nil_end. |
507 (** [[ | 520 (** [[ |
508 app_nil_end | 521 app_nil_end |
509 : forall (A : Type) (l : list A), l = l ++ nil | 522 : forall (A : Type) (l : list A), l = l ++ nil |
510 ]] | 523 ]] |
511 *) | 524 %\index{tactics!rewrite}% *) |
512 | 525 |
513 rewrite (app_nil_end (compile e)). | 526 rewrite (app_nil_end (compile e)). |
514 | 527 |
515 (** This time, we explicitly specify the value of the variable [l] from the theorem statement, since multiple expressions of list type appear in the conclusion. [rewrite] might choose the wrong place to rewrite if we did not specify which we want. | 528 (** This time, we explicitly specify the value of the variable [l] from the theorem statement, since multiple expressions of list type appear in the conclusion. [rewrite] might choose the wrong place to rewrite if we did not specify which we want. |
516 | 529 |
519 ============================ | 532 ============================ |
520 progDenote (compile e ++ nil) nil = Some (expDenote e :: nil) | 533 progDenote (compile e ++ nil) nil = Some (expDenote e :: nil) |
521 | 534 |
522 ]] | 535 ]] |
523 | 536 |
524 Now we can apply the lemma. *) | 537 Now we can apply the lemma.%\index{tactics!rewrite}% *) |
525 | 538 |
526 rewrite compile_correct'. | 539 rewrite compile_correct'. |
527 (** [[ | 540 (** [[ |
528 e : exp | 541 e : exp |
529 ============================ | 542 ============================ |
530 progDenote nil (expDenote e :: nil) = Some (expDenote e :: nil) | 543 progDenote nil (expDenote e :: nil) = Some (expDenote e :: nil) |
531 | 544 |
532 ]] | 545 ]] |
533 | 546 |
534 We are almost done. The lefthand and righthand sides can be seen to match by simple symbolic evaluation. That means we are in luck, because Coq identifies any pair of terms as equal whenever they normalize to the same result by symbolic evaluation. By the definition of [progDenote], that is the case here, but we do not need to worry about such details. A simple invocation of [reflexivity] does the normalization and checks that the two results are syntactically equal. *) | 547 We are almost done. The lefthand and righthand sides can be seen to match by simple symbolic evaluation. That means we are in luck, because Coq identifies any pair of terms as equal whenever they normalize to the same result by symbolic evaluation. By the definition of [progDenote], that is the case here, but we do not need to worry about such details. A simple invocation of %\index{tactics!reflexivity}%[reflexivity] does the normalization and checks that the two results are syntactically equal.%\index{tactics!reflexivity}% *) |
535 | 548 |
536 reflexivity. | 549 reflexivity. |
537 Qed. | 550 Qed. |
538 (* end thide *) | 551 (* end thide *) |
552 | |
553 (** This proof can be shortened and made automated, but we leave that as an exercise for the reader. *) | |
539 | 554 |
540 | 555 |
541 (** * Typed Expressions *) | 556 (** * Typed Expressions *) |
542 | 557 |
543 (** In this section, we will build on the initial example by adding additional expression forms that depend on static typing of terms for safety. *) | 558 (** In this section, we will build on the initial example by adding additional expression forms that depend on static typing of terms for safety. *) |
670 We can define instructions in terms of stack types, where every instruction's type tells us what initial stack type it expects and what final stack type it will produce. *) | 685 We can define instructions in terms of stack types, where every instruction's type tells us what initial stack type it expects and what final stack type it will produce. *) |
671 | 686 |
672 Inductive tinstr : tstack -> tstack -> Set := | 687 Inductive tinstr : tstack -> tstack -> Set := |
673 | TINConst : forall s, nat -> tinstr s (Nat :: s) | 688 | TINConst : forall s, nat -> tinstr s (Nat :: s) |
674 | TIBConst : forall s, bool -> tinstr s (Bool :: s) | 689 | TIBConst : forall s, bool -> tinstr s (Bool :: s) |
675 | TIBinop : forall arg1 arg2 res s, | 690 | TiBinop : forall arg1 arg2 res s, |
676 tbinop arg1 arg2 res | 691 tbinop arg1 arg2 res |
677 -> tinstr (arg1 :: arg2 :: s) (res :: s). | 692 -> tinstr (arg1 :: arg2 :: s) (res :: s). |
678 | 693 |
679 (** Stack machine programs must be a similar inductive family, since, if we again used the [list] type family, we would not be able to guarantee that intermediate stack types match within a program. *) | 694 (** Stack machine programs must be a similar inductive family, since, if we again used the [list] type family, we would not be able to guarantee that intermediate stack types match within a program. *) |
680 | 695 |
699 | 714 |
700 Definition tinstrDenote ts ts' (i : tinstr ts ts') : vstack ts -> vstack ts' := | 715 Definition tinstrDenote ts ts' (i : tinstr ts ts') : vstack ts -> vstack ts' := |
701 match i with | 716 match i with |
702 | TINConst _ n => fun s => (n, s) | 717 | TINConst _ n => fun s => (n, s) |
703 | TIBConst _ b => fun s => (b, s) | 718 | TIBConst _ b => fun s => (b, s) |
704 | TIBinop _ _ _ _ b => fun s => | 719 | TiBinop _ _ _ _ b => fun s => |
705 match s with | 720 match s with |
706 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s') | 721 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s') |
707 end | 722 end |
708 end. | 723 end. |
709 | 724 |
712 [[ | 727 [[ |
713 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' := | 728 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' := |
714 match i with | 729 match i with |
715 | TINConst _ n => (n, s) | 730 | TINConst _ n => (n, s) |
716 | TIBConst _ b => (b, s) | 731 | TIBConst _ b => (b, s) |
717 | TIBinop _ _ _ _ b => | 732 | TiBinop _ _ _ _ b => |
718 match s with | 733 match s with |
719 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s') | 734 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s') |
720 end | 735 end |
721 end. | 736 end. |
722 | 737 |
735 [[ | 750 [[ |
736 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' := | 751 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' := |
737 match i in tinstr ts ts' return vstack ts' with | 752 match i in tinstr ts ts' return vstack ts' with |
738 | TINConst _ n => (n, s) | 753 | TINConst _ n => (n, s) |
739 | TIBConst _ b => (b, s) | 754 | TIBConst _ b => (b, s) |
740 | TIBinop _ _ _ _ b => | 755 | TiBinop _ _ _ _ b => |
741 match s with | 756 match s with |
742 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s') | 757 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s') |
743 end | 758 end |
744 end. | 759 end. |
745 | 760 |
783 Fixpoint tcompile t (e : texp t) (ts : tstack) : tprog ts (t :: ts) := | 798 Fixpoint tcompile t (e : texp t) (ts : tstack) : tprog ts (t :: ts) := |
784 match e with | 799 match e with |
785 | TNConst n => TCons (TINConst _ n) (TNil _) | 800 | TNConst n => TCons (TINConst _ n) (TNil _) |
786 | TBConst b => TCons (TIBConst _ b) (TNil _) | 801 | TBConst b => TCons (TIBConst _ b) (TNil _) |
787 | TBinop _ _ _ b e1 e2 => tconcat (tcompile e2 _) | 802 | TBinop _ _ _ b e1 e2 => tconcat (tcompile e2 _) |
788 (tconcat (tcompile e1 _) (TCons (TIBinop _ b) (TNil _))) | 803 (tconcat (tcompile e1 _) (TCons (TiBinop _ b) (TNil _))) |
789 end. | 804 end. |
790 | 805 |
791 (** One interesting feature of the definition is the underscores appearing to the right of [=>] arrows. Haskell and ML programmers are quite familiar with compilers that infer type parameters to polymorphic values. In Coq, it is possible to go even further and ask the system to infer arbitrary terms, by writing underscores in place of specific values. You may have noticed that we have been calling functions without specifying all of their arguments. For instance, the recursive calls here to [tcompile] omit the [t] argument. Coq's %\emph{%#<i>#implicit argument#</i>#%}% mechanism automatically inserts underscores for arguments that it will probably be able to infer. Inference of such values is far from complete, though; generally, it only works in cases similar to those encountered with polymorphic type instantiation in Haskell and ML. | 806 (** One interesting feature of the definition is the underscores appearing to the right of [=>] arrows. Haskell and ML programmers are quite familiar with compilers that infer type parameters to polymorphic values. In Coq, it is possible to go even further and ask the system to infer arbitrary terms, by writing underscores in place of specific values. You may have noticed that we have been calling functions without specifying all of their arguments. For instance, the recursive calls here to [tcompile] omit the [t] argument. Coq's %\emph{%#<i>#implicit argument#</i>#%}% mechanism automatically inserts underscores for arguments that it will probably be able to infer. Inference of such values is far from complete, though; generally, it only works in cases similar to those encountered with polymorphic type instantiation in Haskell and ML. |
792 | 807 |
793 The underscores here are being filled in with stack types. That is, the Coq type inferencer is, in a sense, inferring something about the flow of control in the translated programs. We can take a look at exactly which values are filled in: *) | 808 The underscores here are being filled in with stack types. That is, the Coq type inferencer is, in a sense, inferring something about the flow of control in the translated programs. We can take a look at exactly which values are filled in: *) |
801 | TNConst n => TCons (TINConst ts n) (TNil (Nat :: ts)) | 816 | TNConst n => TCons (TINConst ts n) (TNil (Nat :: ts)) |
802 | TBConst b => TCons (TIBConst ts b) (TNil (Bool :: ts)) | 817 | TBConst b => TCons (TIBConst ts b) (TNil (Bool :: ts)) |
803 | TBinop arg1 arg2 res b e1 e2 => | 818 | TBinop arg1 arg2 res b e1 e2 => |
804 tconcat (tcompile arg2 e2 ts) | 819 tconcat (tcompile arg2 e2 ts) |
805 (tconcat (tcompile arg1 e1 (arg2 :: ts)) | 820 (tconcat (tcompile arg1 e1 (arg2 :: ts)) |
806 (TCons (TIBinop ts b) (TNil (res :: ts)))) | 821 (TCons (TiBinop ts b) (TNil (res :: ts)))) |
807 end | 822 end |
808 : forall t : type, texp t -> forall ts : tstack, tprog ts (t :: ts) | 823 : forall t : type, texp t -> forall ts : tstack, tprog ts (t :: ts) |
809 ]] | 824 ]] |
810 *) | 825 *) |
811 | 826 |
854 | 869 |
855 [[ | 870 [[ |
856 tprogDenote | 871 tprogDenote |
857 (tconcat (tcompile e2 ts) | 872 (tconcat (tcompile e2 ts) |
858 (tconcat (tcompile e1 (arg2 :: ts)) | 873 (tconcat (tcompile e1 (arg2 :: ts)) |
859 (TCons (TIBinop ts t) (TNil (res :: ts))))) s = | 874 (TCons (TiBinop ts t) (TNil (res :: ts))))) s = |
860 (tbinopDenote t (texpDenote e1) (texpDenote e2), s) | 875 (tbinopDenote t (texpDenote e1) (texpDenote e2), s) |
861 | 876 |
862 ]] | 877 ]] |
863 | 878 |
864 We need an analogue to the [app_ass] theorem that we used to rewrite the goal in the last section. We can abort this proof and prove such a lemma about [tconcat]. | 879 We need an analogue to the [app_ass] theorem that we used to rewrite the goal in the last section. We can abort this proof and prove such a lemma about [tconcat]. |