changeset 26:65314ca099ed

Start of Inductive Types
author Adam Chlipala <adamc@hcoop.net>
date Mon, 08 Sep 2008 14:19:50 -0400
parents 26ad686e68f2
children 8788249c7d3a
files Makefile src/InductiveTypes.v src/Intro.v src/StackMachine.v
diffstat 4 files changed, 171 insertions(+), 17 deletions(-) [+]
line wrap: on
line diff
--- a/Makefile	Mon Sep 08 12:42:57 2008 -0400
+++ b/Makefile	Mon Sep 08 14:19:50 2008 -0400
@@ -1,6 +1,6 @@
 MODULES_NODOC := Tactics
 MODULES_PROSE := Intro
-MODULES_CODE  := StackMachine
+MODULES_CODE  := StackMachine InductiveTypes
 MODULES_DOC   := $(MODULES_PROSE) $(MODULES_CODE)
 MODULES       := $(MODULES_NODOC) $(MODULES_DOC)
 VS            := $(MODULES:%=src/%.v)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/InductiveTypes.v	Mon Sep 08 14:19:50 2008 -0400
@@ -0,0 +1,152 @@
+(* Copyright (c) 2008, Adam Chlipala
+ * 
+ * This work is licensed under a
+ * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
+ * Unported License.
+ * The license text is available at:
+ *   http://creativecommons.org/licenses/by-nc-nd/3.0/
+ *)
+
+(* begin hide *)
+Require Import List.
+
+Require Import Tactics.
+
+Set Implicit Arguments.
+(* end hide *)
+
+
+(** %\chapter{Inductive Types}% *)
+
+(** In a sense, CIC is built from just two relatively straightforward features: function types and inductive types.  From this modest foundation, we can prove effectively all of the theorems of math and carry out effectively all program verifications, with enough effort expended.  This chapter introduces induction and recursion in Coq and shares some "design patterns" for overcoming common pitfalls with them. *)
+
+
+(** * Enumerations *)
+
+(** Coq inductive types generalize the algebraic datatypes found in Haskell and ML.  Confusingly enough, inductive types also generalize generalized algebraic datatypes (GADTs), by adding the possibility for type dependency.  Even so, it is worth backing up from the examples of the last chapter and going over basic, algebraic datatype uses of inductive datatypes, because the chance to prove things about the values of these types adds new wrinkles beyond usual practice in Haskell and ML.
+
+The singleton type [unit] is an inductive type: *)
+
+Inductive unit : Set :=
+  | tt.
+
+(** This vernacular command defines a new inductive type [unit] whose only value is [tt], as we can see by checking the types of the two identifiers: *)
+
+Check unit.
+(** [[
+
+  unit : Set
+]] *)
+Check tt.
+(** [[
+
+  tt : unit
+]] *)
+
+(** We can prove that [unit] is a genuine singleton type. *)
+
+Theorem unit_singleton : forall x : unit, x = tt.
+(** The important thing about an inductive type is, unsurprisingly, that you can do induction over its values, and induction is the key to proving this theorem.  We ask to proceed by induction on the variable [x]. *)
+  induction x.
+(** The goal changes to: [[
+
+ tt = tt
+]] *)
+  (** ...which we can discharge trivially. *)
+  reflexivity.
+Qed.
+
+(** It seems kind of odd to write a proof by induction with no inductive hypotheses.  We could have arrived at the same result by beginning the proof with: [[
+
+  destruct x.
+...which corresponds to "proof by case analysis" in classical math.  For non-recursive inductive types, the two tactics will always have identical behavior.  Often case analysis is sufficient, even in proofs about recursive types, and it is nice to avoid introducing unneeded induction hypotheses.
+
+What exactly %\textit{%#<i>#is#</i>#%}% the induction principle for [unit]?  We can ask Coq: *)
+
+Check unit_ind.
+(** [[
+
+unit_ind : forall P : unit -> Prop, P tt -> forall u : unit, P u
+]]
+
+Every [Inductive] command defining a type [T] also defines an induction principle named [T_ind].  Coq follows the Curry-Howard correspondence and includes the ingredients of programming and proving in the same single syntactic class.  Thus, our type, operations over it, and principles for reasoning about it all live in the same language and are described by the same type system.  The key to telling what is a program and what is a proof lies in the distinction between the type [Prop], which appears in our induction principle; and the type [Set], which we have seen a few times already.
+
+The convention goes like this: [Set] is the type of normal types, and the values of such types are programs.  [Prop] is the type of logical propositions, and the values of such types are proofs.  Thus, an induction principle has a type that shows us that it is a function for building proofs.
+
+Specifically, [unit_ind] quantifies over a predicate [P] over [unit] values.  If we can present a proof that [P] holds of [tt], then we are rewarded with a proof that [P] holds for any value [u] of type [unit].  In our last proof, the predicate was [(fun u : unit => u = tt)].
+
+%\medskip%
+
+We can define an inductive type even simpler than [unit]: *)
+
+Inductive Empty_set : Set := .
+
+(** [Empty_set] has no elements.  We can prove fun theorems about it: *)
+
+Theorem the_sky_is_falling : forall x : Empty_set, 2 + 2 = 5.
+  destruct 1.
+Qed.
+
+(** Because [Empty_set] has no elements, the fact of having an element of this type implies anything.  We use [destruct 1] instead of [destruct x] in the proof because unused quantified variables are relegated to being referred to by number.  (There is a good reason for this, related to the unity of quantifiers and dependent function types.)
+
+We can see the induction principle that made this proof so easy: *)
+
+Check Empty_set_ind.
+(** [[
+
+Empty_set_ind : forall (P : Empty_set -> Prop) (e : Empty_set), P e
+]]
+
+In other words, any predicate over values from the empty set holds vacuously of every such element.  In the last proof, we chose the predicate [(fun _ : Empty_set => 2 + 2 = 5)].
+
+We can also apply this get-out-of-jail-free card programmatically.  Here is a lazy way of converting values of [Empty_set] to values of [unit]: *)
+
+Definition e2u (e : Empty_set) : unit := match e with end.
+
+(** We employ [match] pattern matching as in the last chapter.  Since we match on a value whose type has no constructors, there is no need to provide any branches.
+
+%\medskip%
+
+Moving up the ladder of complexity, we can define the booleans: *)
+
+Inductive bool : Set :=
+| true
+| false.
+
+(** We can use less vacuous pattern matching to define boolean negation. *)
+
+Definition not (b : bool) : bool :=
+  match b with
+    | true => false
+    | false => true
+  end.
+
+(** We might want to prove that [not] is its own inverse operation. *)
+
+Theorem not_inverse : forall b : bool, not (not b) = b.
+  destruct b.
+
+  (** After we case analyze on [b], we are left with one subgoal for each constructor of [bool].
+
+[[
+
+2 subgoals
+  
+  ============================
+   not (not true) = true
+]]
+
+[[
+subgoal 2 is:
+ not (not false) = false
+]]
+
+The first subgoal follows by Coq's rules of computation, so we can dispatch it easily: *)
+
+  reflexivity.
+
+(** Likewise for the second subgoal, so we can restart the proof and give a very compact justification. *)
+
+Restart.
+  destruct b; reflexivity.
+Qed.
--- a/src/Intro.v	Mon Sep 08 12:42:57 2008 -0400
+++ b/src/Intro.v	Mon Sep 08 14:19:50 2008 -0400
@@ -185,6 +185,8 @@
 \hline
 Some Quick Examples & \texttt{StackMachine.v} \\
 \hline
+Inductive Types & \texttt{InductiveTypes.v} \\
+\hline
 \end{tabular} \end{center}
 
 % *)
--- a/src/StackMachine.v	Mon Sep 08 12:42:57 2008 -0400
+++ b/src/StackMachine.v	Mon Sep 08 14:19:50 2008 -0400
@@ -205,7 +205,7 @@
 
 (** 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: *)
 
-Theorem compileCorrect : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
+Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
 (* begin hide *)
 Abort.
 (* end hide *)
@@ -214,7 +214,7 @@
 (** 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 %\textit{%#<i>#strengthening the induction hypothesis#</i>#%}%.  We do that by proving an auxiliary lemma:
 *)
 
-Lemma compileCorrect' : forall e p s, progDenote (compile e ++ p) s = progDenote p (expDenote e :: s).
+Lemma compile_correct' : forall e p s, progDenote (compile e ++ p) s = progDenote p (expDenote e :: s).
 
 (** After the period in the [Lemma] command, we are in %\textit{%#<i>#the interactive proof-editing mode#</i>#%}%.  We find ourselves staring at this ominous screen of text:
 
@@ -468,7 +468,7 @@
 
 Abort.
 
-Lemma compileCorrect' : forall e s p, progDenote (compile e ++ p) s =
+Lemma compile_correct' : forall e s p, progDenote (compile e ++ p) s =
   progDenote p (expDenote e :: s).
   induction e; crush.
 Qed.
@@ -477,7 +477,7 @@
 
 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. *)
 
-Theorem compileCorrect : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
+Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
   intros.
 
 (** [[
@@ -487,7 +487,7 @@
    progDenote (compile e) nil = Some (expDenote e :: nil)
 ]]
 
-At this point, we want to massage the lefthand side to match the statement of [compileCorrect'].  A theorem from the standard library is useful: *)
+At this point, we want to massage the lefthand side to match the statement of [compile_correct'].  A theorem from the standard library is useful: *)
 
 Check app_nil_end.
 
@@ -510,7 +510,7 @@
 
 Now we can apply the lemma. *)
 
-  rewrite compileCorrect'.
+  rewrite compile_correct'.
 
 (** [[
 
@@ -802,7 +802,7 @@
 
 (** We can state a correctness theorem similar to the last one. *)
 
-Theorem tcompileCorrect : forall t (e : texp t), tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
+Theorem tcompile_correct : forall t (e : texp t), tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
 (* begin hide *)
 Abort.
 (* end hide *)
@@ -810,12 +810,12 @@
 
 (** 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: *)
 
-Lemma tcompileCorrect' : forall t (e : texp t)
+Lemma tcompile_correct' : forall t (e : texp t)
   ts (s : vstack ts),
   tprogDenote (tcompile e ts) s
   = (texpDenote e, s).
 
-(** While lemma [compileCorrect'] 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.
+(** 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.
 
    Let us try to prove this theorem in the same way that we settled on in the last section. *)
 
@@ -836,7 +836,7 @@
 *)
 Abort.
 
-Lemma tconcatCorrect : forall ts ts' ts'' (p : tprog ts ts') (p' : tprog ts' ts'')
+Lemma tconcat_correct : forall ts ts' ts'' (p : tprog ts ts') (p' : tprog ts' ts'')
   (s : vstack ts),
   tprogDenote (tconcat p p') s
   = tprogDenote p' (tprogDenote p s).
@@ -845,13 +845,13 @@
 
 (** This one goes through completely automatically.
 
-Some code behind the scenes registers [app_ass] for use by [crush].  We must register [tconcatCorrect] similarly to get the same effect: *)
+Some code behind the scenes registers [app_ass] for use by [crush].  We must register [tconcat_correct] similarly to get the same effect: *)
 
-Hint Rewrite tconcatCorrect : cpdt.
+Hint Rewrite tconcat_correct : cpdt.
 
-(** 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 [tcompileCorrect'], proving it automatically this time. *)
+(** 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. *)
 
-Lemma tcompileCorrect' : forall t (e : texp t)
+Lemma tcompile_correct' : forall t (e : texp t)
   ts (s : vstack ts),
   tprogDenote (tcompile e ts) s
   = (texpDenote e, s).
@@ -860,9 +860,9 @@
 
 (** We can register this main lemma as another hint, allowing us to prove the final theorem trivially. *)
 
-Hint Rewrite tcompileCorrect' : cpdt.
+Hint Rewrite tcompile_correct' : cpdt.
 
-Theorem tcompileCorrect : forall t (e : texp t), tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
+Theorem tcompile_correct : forall t (e : texp t), tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
   crush.
 Qed.
 (* end thide *)