comparison src/StackMachine.v @ 312:495153a41819

Pass through second half of StackMachine
author Adam Chlipala <adam@chlipala.net>
date Thu, 01 Sep 2011 11:32:15 -0400
parents 4cb3ba8604bc
children d5787b70cf48
comparison
equal deleted inserted replaced
311:4cb3ba8604bc 312:495153a41819
6 * The license text is available at: 6 * The license text is available at:
7 * http://creativecommons.org/licenses/by-nc-nd/3.0/ 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
8 *) 8 *)
9 9
10 (* begin hide *) 10 (* begin hide *)
11 Require Import Arith Bool List. 11 Require Import Bool Arith List.
12 12
13 Require Import Tactics. 13 Require Import Tactics.
14 14
15 Set Implicit Arguments. 15 Set Implicit Arguments.
16 (* end hide *) 16 (* end hide *)
19 (** %\chapter{Some Quick Examples}% *) 19 (** %\chapter{Some Quick Examples}% *)
20 20
21 21
22 (** 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. This chapter is not meant to give full explanations of the features that are employed. Rather, it is meant more as an advertisement of what is possible. Later chapters will introduce all of the concepts in bottom-up fashion. 22 (** 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. This chapter is not meant to give full explanations of the features that are employed. Rather, it is meant more as an advertisement of what is possible. Later chapters will introduce all of the concepts in bottom-up fashion.
23 23
24 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 two lines %\index{Vernacular commands!Require}%[Require Import] #<span class="inlinecode"><span class="id" type="var">#%\coqdocconstructor{%Arith%}%#</span></span># [Bool] #<span class="inlinecode"><span class="id" type="var">#%\coqdocconstructor{%List%}%#</span></span># [Tactics.] and %\index{Vernacular commands!Set Implicit Arguments}%[Set Implicit] #<span class="inlinecode"><span class="id" type="keyword">#%\coqdockw{%Arguments%}%#</span></span>#[.] at the start of the file, to match some code hidden in this rendering of the chapter source. In general, similar commands will be hidden in the book rendering of each chapter's source code, so you will need to insert them in from-scratch replayings of the code that is presented. To be more specific, every chapter begins with some imports of other modules, followed by [Set Implicit] #<span class="inlinecode"><span class="id" type="keyword">#%\coqdockw{%Arguments%}%#</span></span>#[.], where the latter affects the default behavior of definitions regarding type inference. 24 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 two lines %\index{Vernacular commands!Require}%[Require Import Bool] #<span class="inlinecode"><span class="id" type="var">#%\coqdocconstructor{%Arith%}%#</span></span># #<span class="inlinecode"><span class="id" type="var">#%\coqdocconstructor{%List%}%#</span></span># [Tactics.] and %\index{Vernacular commands!Set Implicit Arguments}%[Set Implicit] #<span class="inlinecode"><span class="id" type="keyword">#%\coqdockw{%Arguments%}%#</span></span>#[.] at the start of the file, to match some code hidden in this rendering of the chapter source. In general, similar commands will be hidden in the book rendering of each chapter's source code, so you will need to insert them in from-scratch replayings of the code that is presented. To be more specific, every chapter begins with some imports of other modules, followed by [Set Implicit] #<span class="inlinecode"><span class="id" type="keyword">#%\coqdockw{%Arguments%}%#</span></span>#[.], where the latter affects the default behavior of definitions regarding type inference.
25 *) 25 *)
26 26
27 27
28 (** * Arithmetic Expressions Over Natural Numbers *) 28 (** * Arithmetic Expressions Over Natural Numbers *)
29 29
499 499
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. 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.
501 501
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. 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.
503 503
504 The %\index{Vernacular commands!Qed}%[Qed] command checks that the proof is finished and, if so, saves it. 504 The %\index{Vernacular commands!Qed}%[Qed] command checks that the proof is finished and, if so, saves it. The tactic commands we have written above are an example of a %\emph{%#<i>#proof script#</i>#%}%, or a series of Ltac programs; while [Qed] uses the result of the script to generate a %\emph{%#<i>#proof term#</i>#%}%, a well-typed term of Gallina. To believe that a theorem is true, we only need to trust that the (relatively simple) checker for proof terms is correct; the use of proof scripts is immaterial. Part I of this book will introduce the principles behind encoding all proofs as terms of Gallina.
505 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}% *) 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}% *)
507 507
508 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).
509 intros. 509 intros.
573 | TEq : forall t, tbinop t t Bool 573 | TEq : forall t, tbinop t t Bool
574 | TLt : tbinop Nat Nat Bool. 574 | TLt : tbinop Nat Nat Bool.
575 575
576 (** The definition of [tbinop] is different from [binop] in an important way. Where we declared that [binop] has type [Set], here we declare that [tbinop] has type [type -> type -> type -> Set]. We define [tbinop] as an %\emph{%#<i>#indexed type family#</i>#%}%. Indexed inductive types are at the heart of Coq's expressive power; almost everything else of interest is defined in terms of them. 576 (** The definition of [tbinop] is different from [binop] in an important way. Where we declared that [binop] has type [Set], here we declare that [tbinop] has type [type -> type -> type -> Set]. We define [tbinop] as an %\emph{%#<i>#indexed type family#</i>#%}%. Indexed inductive types are at the heart of Coq's expressive power; almost everything else of interest is defined in terms of them.
577 577
578 The inuitive explanation of [tbinop] is that a [tbinop t1 t2 t] is a binary operator whose operands should have types [t1] and [t2], and whose result has type [t]. For instance, constructor [TLt] (for less-than comparison of numbers) is assigned type [tbinop Nat Nat Bool], meaning the operator's arguments are naturals and its result is boolean. The type of [TEq] introduces a small bit of additional complication via polymorphism: we want to allow equality comparison of any two values of any type, as long as they have the %\emph{%#<i>#same#</i>#%}% type.
579
578 ML and Haskell have indexed algebraic datatypes. For instance, their list types are indexed by the type of data that the list carries. However, compared to Coq, ML and Haskell 98 place two important restrictions on datatype definitions. 580 ML and Haskell have indexed algebraic datatypes. For instance, their list types are indexed by the type of data that the list carries. However, compared to Coq, ML and Haskell 98 place two important restrictions on datatype definitions.
579 581
580 First, the indices of the range of each data constructor must be type variables bound at the top level of the datatype definition. There is no way to do what we did here, where we, for instance, say that [TPlus] is a constructor building a [tbinop] whose indices are all fixed at [Nat]. %\emph{%#<i>#Generalized algebraic datatypes (GADTs)#</i>#%}% are a popular feature in GHC Haskell and other languages that removes this first restriction. 582 First, the indices of the range of each data constructor must be type variables bound at the top level of the datatype definition. There is no way to do what we did here, where we, for instance, say that [TPlus] is a constructor building a [tbinop] whose indices are all fixed at [Nat]. %\index{generalized algebraic datatypes}\index{GADTs|see{generalized algebraic datatypes}}\emph{%#<i>#Generalized algebraic datatypes (GADTs)#</i>#%}~\cite{GADT}% are a popular feature in %\index{GHC Haskell}%GHC Haskell and other languages that removes this first restriction.
581 583
582 The second restriction is not lifted by GADTs. In ML and Haskell, indices of types must be types and may not be %\emph{%#<i>#expressions#</i>#%}%. In Coq, types may be indexed by arbitrary Gallina terms. Type indices can live in the same universe as programs, and we can compute with them just like regular programs. Haskell supports a hobbled form of computation in type indices based on multi-parameter type classes, and recent extensions like type functions bring Haskell programming even closer to %``%#"#real#"#%''% functional programming with types, but, without dependent typing, there must always be a gap between how one programs with types and how one programs normally. 584 The second restriction is not lifted by GADTs. In ML and Haskell, indices of types must be types and may not be %\emph{%#<i>#expressions#</i>#%}%. In Coq, types may be indexed by arbitrary Gallina terms. Type indices can live in the same universe as programs, and we can compute with them just like regular programs. Haskell supports a hobbled form of computation in type indices based on %\index{Haskell}%multi-parameter type classes, and recent extensions like type functions bring Haskell programming even closer to %``%#"#real#"#%''% functional programming with types, but, without dependent typing, there must always be a gap between how one programs with types and how one programs normally.
583 *) 585 *)
584 586
585 (** We can define a similar type family for typed expressions. *) 587 (** We can define a similar type family for typed expressions, where a term of type [texp t] can be assigned object language type [t]. (It is conventional in the world of interactive theorem proving to call the language of the proof assistant the %\index{meta language}\emph{%#<i>#meta language#</i>#%}% and a language being formalized the %\index{object language}\emph{%#<i>#object language#</i>#%}%.) *)
586 588
587 Inductive texp : type -> Set := 589 Inductive texp : type -> Set :=
588 | TNConst : nat -> texp Nat 590 | TNConst : nat -> texp Nat
589 | TBConst : bool -> texp Bool 591 | TBConst : bool -> texp Bool
590 | TBinop : forall arg1 arg2 res, tbinop arg1 arg2 res -> texp arg1 -> texp arg2 -> texp res. 592 | TBinop : forall t1 t2 t, tbinop t1 t2 t -> texp t1 -> texp t2 -> texp t.
591 593
592 (** Thanks to our use of dependent types, every well-typed [texp] represents a well-typed source expression, by construction. This turns out to be very convenient for many things we might want to do with expressions. For instance, it is easy to adapt our interpreter approach to defining semantics. We start by defining a function mapping the types of our languages into Coq types: *) 594 (** Thanks to our use of dependent types, every well-typed [texp] represents a well-typed source expression, by construction. This turns out to be very convenient for many things we might want to do with expressions. For instance, it is easy to adapt our interpreter approach to defining semantics. We start by defining a function mapping the types of our languages into Coq types: *)
593 595
594 Definition typeDenote (t : type) : Set := 596 Definition typeDenote (t : type) : Set :=
595 match t with 597 match t with
596 | Nat => nat 598 | Nat => nat
597 | Bool => bool 599 | Bool => bool
598 end. 600 end.
599 601
600 (** It can take a few moments to come to terms with the fact that [Set], the type of types of programs, is itself a first-class type, and that we can write functions that return [Set]s. Past that wrinkle, the definition of [typeDenote] is trivial, relying on the [nat] and [bool] types from the Coq standard library. 602 (** It can take a few moments to come to terms with the fact that [Set], the type of types of programs, is itself a first-class type, and that we can write functions that return [Set]s. Past that wrinkle, the definition of [typeDenote] is trivial, relying on the [nat] and [bool] types from the Coq standard library. We can interpret binary operators by relying on standard-library equality test functions [eqb] and [beq_nat] for booleans and naturals, respectively, along with a less-than test [leb]: *)
601 603
602 We need to define one auxiliary function, implementing a boolean binary %``%#"#less-than#"#%''% operator, which only appears in the standard library with a type fancier than what we are prepared to deal with here. The code is entirely standard and ML-like, with the one caveat being that the Coq [nat] type uses a unary representation, where [O] is zero and [S n] is the successor of [n].
603 *)
604
605 Fixpoint lessThan (n1 n2 : nat) : bool :=
606 match n1, n2 with
607 | O, S _ => true
608 | S n1', S n2' => lessThan n1' n2'
609 | _, _ => false
610 end.
611
612 (** Now we can interpret binary operators, relying on standard-library equality test functions [eqb] and [beq_nat] for booleans and naturals, respectively: *)
613
614 Definition tbinopDenote arg1 arg2 res (b : tbinop arg1 arg2 res)
615 : typeDenote arg1 -> typeDenote arg2 -> typeDenote res :=
616 match b in (tbinop arg1 arg2 res)
617 return (typeDenote arg1 -> typeDenote arg2 -> typeDenote res) with
618 | TPlus => plus
619 | TTimes => mult
620 | TEq Nat => beq_nat
621 | TEq Bool => eqb
622 | TLt => lessThan
623 end.
624
625 (** This function has just a few differences from the denotation functions we saw earlier. First, [tbinop] is an indexed type, so its indices become additional arguments to [tbinopDenote]. Second, we need to perform a genuine %\emph{%#<i>#dependent pattern match#</i>#%}% to come up with a definition of this function that type-checks. In each branch of the [match], we need to use branch-specific information about the indices to [tbinop]. General type inference that takes such information into account is undecidable, so it is often necessary to write annotations, like we see above on the line with [match].
626
627 The [in] annotation restates the type of the term being case-analyzed. Though we use the same names for the indices as we use in the type of the original argument binder, these are actually fresh variables, and they are %\emph{%#<i>#binding occurrences#</i>#%}%. Their scope is the [return] clause. That is, [arg1], [arg2], and [res] are new bound variables bound only within the return clause [typeDenote arg1 -> typeDenote arg2 -> typeDenote res]. By being explicit about the functional relationship between the type indices and the match result, we regain decidable type inference.
628
629 In fact, recent Coq versions use some heuristics that can save us the trouble of writing [match] annotations, and those heuristics get the job done in this case. We can get away with writing just: *)
630
631 (* begin hide *)
632 Reset tbinopDenote.
633 (* end hide *)
634 Definition tbinopDenote arg1 arg2 res (b : tbinop arg1 arg2 res) 604 Definition tbinopDenote arg1 arg2 res (b : tbinop arg1 arg2 res)
635 : typeDenote arg1 -> typeDenote arg2 -> typeDenote res := 605 : typeDenote arg1 -> typeDenote arg2 -> typeDenote res :=
636 match b with 606 match b with
637 | TPlus => plus 607 | TPlus => plus
638 | TTimes => mult 608 | TTimes => mult
639 | TEq Nat => beq_nat 609 | TEq Nat => beq_nat
640 | TEq Bool => eqb 610 | TEq Bool => eqb
641 | TLt => lessThan 611 | TLt => leb
642 end. 612 end.
643 613
644 (** 614 (** This function has just a few differences from the denotation functions we saw earlier. First, [tbinop] is an indexed type, so its indices become additional arguments to [tbinopDenote]. Second, we need to perform a genuine %\index{dependent pattern matching}\emph{%#<i>#dependent pattern match#</i>#%}%, where the necessary %\emph{%#<i>#type#</i>#%}% of each case body depends on the %\emph{%#<i>#value#</i>#%}% that has been matched. At this early stage, we will not go into detail on the many subtle aspects of Gallina that support dependent pattern-matching, but the subject is central to Part II of the book.
615
645 The same tricks suffice to define an expression denotation function in an unsurprising way: 616 The same tricks suffice to define an expression denotation function in an unsurprising way:
646 *) 617 *)
647 618
648 Fixpoint texpDenote t (e : texp t) : typeDenote t := 619 Fixpoint texpDenote t (e : texp t) : typeDenote t :=
649 match e with 620 match e with
658 (** [= 42 : typeDenote Nat] *) 629 (** [= 42 : typeDenote Nat] *)
659 630
660 Eval simpl in texpDenote (TBConst true). 631 Eval simpl in texpDenote (TBConst true).
661 (** [= true : typeDenote Bool] *) 632 (** [= true : typeDenote Bool] *)
662 633
663 Eval simpl in texpDenote (TBinop TTimes (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)). 634 Eval simpl in texpDenote (TBinop TTimes (TBinop TPlus (TNConst 2) (TNConst 2))
635 (TNConst 7)).
664 (** [= 28 : typeDenote Nat] *) 636 (** [= 28 : typeDenote Nat] *)
665 637
666 Eval simpl in texpDenote (TBinop (TEq Nat) (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)). 638 Eval simpl in texpDenote (TBinop (TEq Nat) (TBinop TPlus (TNConst 2) (TNConst 2))
667 (** [= false : typeDenote Bool] *) 639 (TNConst 7)).
668 640 (** [= ] %\coqdocconstructor{%#<tt>#false#</tt>#%}% [ : typeDenote Bool] *)
669 Eval simpl in texpDenote (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)). 641
642 Eval simpl in texpDenote (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2))
643 (TNConst 7)).
670 (** [= true : typeDenote Bool] *) 644 (** [= true : typeDenote Bool] *)
671 645
672 646
673 (** ** Target Language *) 647 (** ** Target Language *)
674 648
683 (** Any stack classified by a [tstack] must have exactly as many elements, and each stack element must have the type found in the same position of the stack type. 657 (** Any stack classified by a [tstack] must have exactly as many elements, and each stack element must have the type found in the same position of the stack type.
684 658
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. *) 659 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. *)
686 660
687 Inductive tinstr : tstack -> tstack -> Set := 661 Inductive tinstr : tstack -> tstack -> Set :=
688 | TINConst : forall s, nat -> tinstr s (Nat :: s) 662 | TiNConst : forall s, nat -> tinstr s (Nat :: s)
689 | TIBConst : forall s, bool -> tinstr s (Bool :: s) 663 | TiBConst : forall s, bool -> tinstr s (Bool :: s)
690 | TiBinop : forall arg1 arg2 res s, 664 | TiBinop : forall arg1 arg2 res s,
691 tbinop arg1 arg2 res 665 tbinop arg1 arg2 res
692 -> tinstr (arg1 :: arg2 :: s) (res :: s). 666 -> tinstr (arg1 :: arg2 :: s) (res :: s).
693 667
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. *) 668 (** 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. *)
706 match ts with 680 match ts with
707 | nil => unit 681 | nil => unit
708 | t :: ts' => typeDenote t * vstack ts' 682 | t :: ts' => typeDenote t * vstack ts'
709 end%type. 683 end%type.
710 684
711 (** This is another [Set]-valued function. This time it is recursive, which is perfectly valid, since [Set] is not treated specially in determining which functions may be written. We say that the value stack of an empty stack type is any value of type [unit], which has just a single value, [tt]. A nonempty stack type leads to a value stack that is a pair, whose first element has the proper type and whose second element follows the representation for the remainder of the stack type. We write [%type] so that Coq knows to interpret [*] as Cartesian product rather than multiplication. 685 (** This is another [Set]-valued function. This time it is recursive, which is perfectly valid, since [Set] is not treated specially in determining which functions may be written. We say that the value stack of an empty stack type is any value of type [unit], which has just a single value, [tt]. A nonempty stack type leads to a value stack that is a pair, whose first element has the proper type and whose second element follows the representation for the remainder of the stack type. We write [%]%\index{notation scopes}\coqdocvar{%#<tt>#type#</tt>#%}% as an instruction to Coq's extensible parser. In particular, this directive applies to the whole [match] expression, which we ask to be parsed as though it were a type, so that the operator [*] is interpreted as Cartesian product instead of, say, multiplication. (Note that this use of %\coqdocvar{%#<tt>#type#</tt>#%}% has no connection to the inductive type [type] that we have defined.)
712 686
713 This idea of programming with types can take a while to internalize, but it enables a very simple definition of instruction denotation. Our definition is like what you might expect from a Lisp-like version of ML that ignored type information. Nonetheless, the fact that [tinstrDenote] passes the type-checker guarantees that our stack machine programs can never go wrong. *) 687 This idea of programming with types can take a while to internalize, but it enables a very simple definition of instruction denotation. Our definition is like what you might expect from a Lisp-like version of ML that ignored type information. Nonetheless, the fact that [tinstrDenote] passes the type-checker guarantees that our stack machine programs can never go wrong. We use a special form of [let] to destructure a multi-level tuple. *)
714 688
715 Definition tinstrDenote ts ts' (i : tinstr ts ts') : vstack ts -> vstack ts' := 689 Definition tinstrDenote ts ts' (i : tinstr ts ts') : vstack ts -> vstack ts' :=
716 match i with 690 match i with
717 | TINConst _ n => fun s => (n, s) 691 | TiNConst _ n => fun s => (n, s)
718 | TIBConst _ b => fun s => (b, s) 692 | TiBConst _ b => fun s => (b, s)
719 | TiBinop _ _ _ _ b => fun s => 693 | TiBinop _ _ _ _ b => fun s =>
720 match s with 694 let '(arg1, (arg2, s')) := s in
721 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s') 695 ((tbinopDenote b) arg1 arg2, s')
722 end
723 end. 696 end.
724 697
725 (** Why do we choose to use an anonymous function to bind the initial stack in every case of the [match]? Consider this well-intentioned but invalid alternative version: 698 (** Why do we choose to use an anonymous function to bind the initial stack in every case of the [match]? Consider this well-intentioned but invalid alternative version:
726
727 [[ 699 [[
728 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' := 700 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' :=
729 match i with 701 match i with
730 | TINConst _ n => (n, s) 702 | TiNConst _ n => (n, s)
731 | TIBConst _ b => (b, s) 703 | TiBConst _ b => (b, s)
732 | TiBinop _ _ _ _ b => 704 | TiBinop _ _ _ _ b =>
733 match s with 705 let '(arg1, (arg2, s')) := s in
734 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s') 706 ((tbinopDenote b) arg1 arg2, s')
735 end
736 end. 707 end.
737 708
738 ]] 709 ]]
739 710
740 The Coq type-checker complains that: 711 The Coq type-checker complains that:
741 712
742 [[ 713 <<
743 The term "(n, s)" has type "(nat * vstack ts)%type" 714 The term "(n, s)" has type "(nat * vstack ts)%type"
744 while it is expected to have type "vstack ?119". 715 while it is expected to have type "vstack ?119".
745 716 >>
746 ]] 717
747 718 This and other mysteries of Coq dependent typing we postpone until Part II of the book. The upshot of our later discussion is that it is often useful to push inside of [match] branches those function parameters whose types depend on the type of the value being matched. Our later, more complete treatement of Gallina's typing rules will explain why this helps.
748 The text [?119] stands for a unification variable. We can try to help Coq figure out the value of this variable with an explicit annotation on our [match] expression.
749
750 [[
751 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' :=
752 match i in tinstr ts ts' return vstack ts' with
753 | TINConst _ n => (n, s)
754 | TIBConst _ b => (b, s)
755 | TiBinop _ _ _ _ b =>
756 match s with
757 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s')
758 end
759 end.
760
761 ]]
762
763 Now the error message changes.
764
765 [[
766 The term "(n, s)" has type "(nat * vstack ts)%type"
767 while it is expected to have type "vstack (Nat :: t)".
768
769 ]]
770
771 Recall from our earlier discussion of [match] annotations that we write the annotations to express to the type-checker the relationship between the type indices of the case object and the result type of the [match]. Coq chooses to assign to the wildcard [_] after [TINConst] the name [t], and the type error is telling us that the type checker cannot prove that [t] is the same as [ts]. By moving [s] out of the [match], we lose the ability to express, with [in] and [return] clauses, the relationship between the shared index [ts] of [s] and [i].
772
773 There %\emph{%#<i>#are#</i>#%}% reasonably general ways of getting around this problem without pushing binders inside [match]es. However, the alternatives are significantly more involved, and the technique we use here is almost certainly the best choice, whenever it applies.
774
775 *) 719 *)
776 720
777 (** We finish the semantics with a straightforward definition of program denotation. *) 721 (** We finish the semantics with a straightforward definition of program denotation. *)
778 722
779 Fixpoint tprogDenote ts ts' (p : tprog ts ts') : vstack ts -> vstack ts' := 723 Fixpoint tprogDenote ts ts' (p : tprog ts ts') : vstack ts -> vstack ts' :=
795 739
796 (** With that function in place, the compilation is defined very similarly to how it was before, modulo the use of dependent typing. *) 740 (** With that function in place, the compilation is defined very similarly to how it was before, modulo the use of dependent typing. *)
797 741
798 Fixpoint tcompile t (e : texp t) (ts : tstack) : tprog ts (t :: ts) := 742 Fixpoint tcompile t (e : texp t) (ts : tstack) : tprog ts (t :: ts) :=
799 match e with 743 match e with
800 | TNConst n => TCons (TINConst _ n) (TNil _) 744 | TNConst n => TCons (TiNConst _ n) (TNil _)
801 | TBConst b => TCons (TIBConst _ b) (TNil _) 745 | TBConst b => TCons (TiBConst _ b) (TNil _)
802 | TBinop _ _ _ b e1 e2 => tconcat (tcompile e2 _) 746 | TBinop _ _ _ b e1 e2 => tconcat (tcompile e2 _)
803 (tconcat (tcompile e1 _) (TCons (TiBinop _ b) (TNil _))) 747 (tconcat (tcompile e1 _) (TCons (TiBinop _ b) (TNil _)))
804 end. 748 end.
805 749
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. 750 (** 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.
811 (** [[ 755 (** [[
812 tcompile = 756 tcompile =
813 fix tcompile (t : type) (e : texp t) (ts : tstack) {struct e} : 757 fix tcompile (t : type) (e : texp t) (ts : tstack) {struct e} :
814 tprog ts (t :: ts) := 758 tprog ts (t :: ts) :=
815 match e in (texp t0) return (tprog ts (t0 :: ts)) with 759 match e in (texp t0) return (tprog ts (t0 :: ts)) with
816 | TNConst n => TCons (TINConst ts n) (TNil (Nat :: ts)) 760 | TNConst n => TCons (TiNConst ts n) (TNil (Nat :: ts))
817 | TBConst b => TCons (TIBConst ts b) (TNil (Bool :: ts)) 761 | TBConst b => TCons (TiBConst ts b) (TNil (Bool :: ts))
818 | TBinop arg1 arg2 res b e1 e2 => 762 | TBinop arg1 arg2 res b e1 e2 =>
819 tconcat (tcompile arg2 e2 ts) 763 tconcat (tcompile arg2 e2 ts)
820 (tconcat (tcompile arg1 e1 (arg2 :: ts)) 764 (tconcat (tcompile arg1 e1 (arg2 :: ts))
821 (TCons (TiBinop ts b) (TNil (res :: ts)))) 765 (TCons (TiBinop ts b) (TNil (res :: ts))))
822 end 766 end
826 770
827 771
828 (** We can check that the compiler generates programs that behave appropriately on our sample programs from above: *) 772 (** We can check that the compiler generates programs that behave appropriately on our sample programs from above: *)
829 773
830 Eval simpl in tprogDenote (tcompile (TNConst 42) nil) tt. 774 Eval simpl in tprogDenote (tcompile (TNConst 42) nil) tt.
831 (** [= (42, tt) : vstack (Nat :: nil)] *) 775 (** [= (42, tt) : vstack (][Nat :: nil)] *)
832 776
833 Eval simpl in tprogDenote (tcompile (TBConst true) nil) tt. 777 Eval simpl in tprogDenote (tcompile (TBConst true) nil) tt.
834 (** [= (true, tt) : vstack (Bool :: nil)] *) 778 (** [= (][true][, tt) : vstack (][Bool :: nil)] *)
835 779
836 Eval simpl in tprogDenote (tcompile (TBinop TTimes (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)) nil) tt. 780 Eval simpl in tprogDenote (tcompile (TBinop TTimes (TBinop TPlus (TNConst 2)
837 (** [= (28, tt) : vstack (Nat :: nil)] *) 781 (TNConst 2)) (TNConst 7)) nil) tt.
838 782 (** [= (28, tt) : vstack (][Nat :: nil)] *)
839 Eval simpl in tprogDenote (tcompile (TBinop (TEq Nat) (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)) nil) tt. 783
840 (** [= (false, tt) : vstack (Bool :: nil)] *) 784 Eval simpl in tprogDenote (tcompile (TBinop (TEq Nat) (TBinop TPlus (TNConst 2)
841 785 (TNConst 2)) (TNConst 7)) nil) tt.
842 Eval simpl in tprogDenote (tcompile (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)) nil) tt. 786 (** [= (]%\coqdocconstructor{%#<tt>#false#</tt>#%}%[, tt) : vstack (][Bool :: nil)] *)
843 (** [= (true, tt) : vstack (Bool :: nil)] *) 787
788 Eval simpl in tprogDenote (tcompile (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2))
789 (TNConst 7)) nil) tt.
790 (** [= (][true][, tt) : vstack (][Bool :: nil)] *)
844 791
845 792
846 (** ** Translation Correctness *) 793 (** ** Translation Correctness *)
847 794
848 (** We can state a correctness theorem similar to the last one. *) 795 (** We can state a correctness theorem similar to the last one. *)
852 (* begin hide *) 799 (* begin hide *)
853 Abort. 800 Abort.
854 (* end hide *) 801 (* end hide *)
855 (* begin thide *) 802 (* begin thide *)
856 803
857 (** Again, we need to strengthen the theorem statement so that the induction will go through. This time, I will develop an alternative approach to this kind of proof, stating the key lemma as: *) 804 (** Again, we need to strengthen the theorem statement so that the induction will go through. This time, to provide an excuse to demonstrate different tactics, I will develop an alternative approach to this kind of proof, stating the key lemma as: *)
858 805
859 Lemma tcompile_correct' : forall t (e : texp t) ts (s : vstack ts), 806 Lemma tcompile_correct' : forall t (e : texp t) ts (s : vstack ts),
860 tprogDenote (tcompile e ts) s = (texpDenote e, s). 807 tprogDenote (tcompile e ts) s = (texpDenote e, s).
861 808
862 (** While lemma [compile_correct'] quantified over a program that is the %``%#"#continuation#"#%''% for the expression we are considering, here we avoid drawing in any extra syntactic elements. In addition to the source expression and its type, we also quantify over an initial stack type and a stack compatible with it. Running the compilation of the program starting from that stack, we should arrive at a stack that differs only in having the program's denotation pushed onto it. 809 (** While lemma [compile_correct'] quantified over a program that is the %``%#"#continuation#"#%''% for the expression we are considering, here we avoid drawing in any extra syntactic elements. In addition to the source expression and its type, we also quantify over an initial stack type and a stack compatible with it. Running the compilation of the program starting from that stack, we should arrive at a stack that differs only in having the program's denotation pushed onto it.
874 (TCons (TiBinop ts t) (TNil (res :: ts))))) s = 821 (TCons (TiBinop ts t) (TNil (res :: ts))))) s =
875 (tbinopDenote t (texpDenote e1) (texpDenote e2), s) 822 (tbinopDenote t (texpDenote e1) (texpDenote e2), s)
876 823
877 ]] 824 ]]
878 825
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]. 826 We need an analogue to the [app_assoc_reverse] theorem that we used to rewrite the goal in the last section. We can abort this proof and prove such a lemma about [tconcat].
880 *) 827 *)
881 828
882 Abort. 829 Abort.
883 830
884 Lemma tconcat_correct : forall ts ts' ts'' (p : tprog ts ts') (p' : tprog ts' ts'') 831 Lemma tconcat_correct : forall ts ts' ts'' (p : tprog ts ts') (p' : tprog ts' ts'')
888 induction p; crush. 835 induction p; crush.
889 Qed. 836 Qed.
890 837
891 (** This one goes through completely automatically. 838 (** This one goes through completely automatically.
892 839
893 Some code behind the scenes registers [app_ass] for use by [crush]. We must register [tconcat_correct] similarly to get the same effect: *) 840 Some code behind the scenes registers [app_assoc_reverse] for use by [crush]. We must register [tconcat_correct] similarly to get the same effect:%\index{Verncular commands!Hint Rewrite}% *)
894 841
842 (* begin hide *)
895 Hint Rewrite tconcat_correct : cpdt. 843 Hint Rewrite tconcat_correct : cpdt.
896 844 (* end hide *)
897 (** We ask that the lemma be used for left-to-right rewriting, and we ask for the hint to be added to the hint database called [cpdt], which is the database used by [crush]. Now we are ready to return to [tcompile_correct'], proving it automatically this time. *) 845 (** %\noindent%[Hint] %\coqdockw{%#<tt>#Rewrite#</tt>#%}% [tconcat_correct : cpdt.] *)
846
847 (** Here we meet the pervasive concept of a %\emph{%#<i>#hint#</i>#%}%. Many proofs can be found through exhaustive enumerations of combinations of possible proof steps; hints provide the set of steps to consider. The tactic [crush] is applying such brute force search for us silently, and it will consider more possibilities as we add more hints. This particular hint asks that the lemma be used for left-to-right rewriting, and we ask for the hint to be added to the hint database called [cpdt], which is the database used by [crush]. In general, segragating hints into different databases is helpful to control the performance of proof search, in cases where domain knowledge allows us to narrow the set of proof steps to be considered in brute force search. Part III of this book considers such pragmatic aspects of proof search in much more detail.
848
849 Now we are ready to return to [tcompile_correct'], proving it automatically this time. *)
898 850
899 Lemma tcompile_correct' : forall t (e : texp t) ts (s : vstack ts), 851 Lemma tcompile_correct' : forall t (e : texp t) ts (s : vstack ts),
900 tprogDenote (tcompile e ts) s = (texpDenote e, s). 852 tprogDenote (tcompile e ts) s = (texpDenote e, s).
901 induction e; crush. 853 induction e; crush.
902 Qed. 854 Qed.
903 855
904 (** We can register this main lemma as another hint, allowing us to prove the final theorem trivially. *) 856 (** We can register this main lemma as another hint, allowing us to prove the final theorem trivially. *)
905 857
858 (* begin hide *)
906 Hint Rewrite tcompile_correct' : cpdt. 859 Hint Rewrite tcompile_correct' : cpdt.
860 (* end hide *)
861 (** %\noindent%[Hint ]%\coqdockw{%#<tt>#Rewrite#</tt>#%}%[ tcompile_correct' : cpdt.] *)
907 862
908 Theorem tcompile_correct : forall t (e : texp t), 863 Theorem tcompile_correct : forall t (e : texp t),
909 tprogDenote (tcompile e nil) tt = (texpDenote e, tt). 864 tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
910 crush. 865 crush.
911 Qed. 866 Qed.
912 (* end thide *) 867 (* end thide *)
868
869 (** It is probably worth emphasizing that we are doing more than building mathematical models. Our compilers are functional programs that can be executed efficiently. One strategy for doing so is based on %\index{program extraction}\emph{%#<i>#program extraction#</i>#%}%, which generates OCaml code from Coq developments. For instance, we run a command to output the OCaml version of [tcompile]:%\index{Vernacular commands!Extraction}% *)
870
871 (* begin hide *)
872 Extraction tcompile.
873 (* end hide *)
874 (** %\noindent\coqdockw{%#<tt>#Extraction#</tt>#%}%[ tcompile.] *)
875
876 (** <<
877 let rec tcompile t e ts =
878 match e with
879 | TNConst n ->
880 TCons (ts, (Cons (Nat, ts)), (Cons (Nat, ts)), (TiNConst (ts, n)), (TNil
881 (Cons (Nat, ts))))
882 | TBConst b ->
883 TCons (ts, (Cons (Bool, ts)), (Cons (Bool, ts)), (TiBConst (ts, b)),
884 (TNil (Cons (Bool, ts))))
885 | TBinop (t1, t2, t0, b, e1, e2) ->
886 tconcat ts (Cons (t2, ts)) (Cons (t0, ts)) (tcompile t2 e2 ts)
887 (tconcat (Cons (t2, ts)) (Cons (t1, (Cons (t2, ts)))) (Cons (t0, ts))
888 (tcompile t1 e1 (Cons (t2, ts))) (TCons ((Cons (t1, (Cons (t2,
889 ts)))), (Cons (t0, ts)), (Cons (t0, ts)), (TiBinop (t1, t2, t0, ts,
890 b)), (TNil (Cons (t0, ts))))))
891 >>
892
893 We can compile this code with the usual OCaml compiler and obtain an executable program with halfway decent performance.
894
895 This chapter has been a whirlwind tour through two examples of the style of Coq development that I advocate. Parts II and III of the book focus on the key elements of that style, namely dependent types and scripted proof automation, respectively. Before we get there, we will spend some time in Part I on more standard foundational material. Part I may still be of interest to seasoned Coq hackers, since I follow the highly automated proof style even at that early stage. *)