Commit e56b9f3a authored by Adam Chlipala's avatar Adam Chlipala

Batch of changes based on proofreader feedback

parent 453fc375
......@@ -57,7 +57,7 @@ CoInductive evilStream := Nil.
(** The definition is surprisingly simple. Starting from the definition of [list], we just need to change the keyword [Inductive] to [CoInductive]. We could have left a [Nil] constructor in our definition, but we will leave it out to force all of our streams to be infinite.
How do we write down a stream constant? Obviously simple application of constructors is not good enough, since we could only denote finite objects that way. Rather, whereas recursive definitions were necessary to _use_ values of recursive inductive types effectively, here we find that we need%\index{co-recursive definitions}% _co-recursive definitions_ to _build_ values of co-inductive types effectively.
How do we write down a stream constant? Obviously, simple application of constructors is not good enough, since we could only denote finite objects that way. Rather, whereas recursive definitions were necessary to _use_ values of recursive inductive types effectively, here we find that we need%\index{co-recursive definitions}% _co-recursive definitions_ to _build_ values of co-inductive types effectively.
We can define a stream consisting only of zeroes.%\index{Vernacular commands!CoFixpoint}% *)
......@@ -191,11 +191,13 @@ CoFixpoint bad : stream nat := tl (Cons 0 bad).
Imagine that Coq had accepted our definition, and consider how we might evaluate [approx bad 1]. We would be trying to calculate the first element in the stream [bad]. However, it is not hard to see that the definition of [bad] "begs the question": unfolding the definition of [tl], we see that we essentially say "define [bad] to equal itself"! Of course such an equation admits no single well-defined solution, which does not fit well with the determinism of Gallina reduction.
Since Coq can be considered to check definitions after inlining and simplification of previously defined identifiers, the basic guardedness condition rules out our definition of [bad]. Such an inlining reduces [bad] to:
Coq's complete rule for co-recursive definitions includes not just the basic guardedness condition, but also a requirement about where co-recursive calls may occur. In particular, a co-recursive call must be a direct argument to a constructor, _nested only inside of other constructor calls or [fun] or [match] expressions_. In the definition of [bad], we erroneously nested the co-recursive call inside a call to [tl], and we nested inside a call to [interleave] in the definition of [map'].
Coq helps the user out a little by performing the guardedness check after using computation to simplify terms. For instance, any co-recursive function definition can be expanded by inserting extra calls to an identity function, and this change preserves guardedness. However, in other cases computational simplification can reveal why definitions are dangerous. Consider what happens when we inline the definition of [tl] in [bad]:
[[
CoFixpoint bad : stream nat := bad.
]]
This is the same looping definition we rejected earlier. A similar inlining process reveals the way that Coq saw our failed definition of [map']:
This is the same looping definition we rejected earlier. A similar inlining process reveals an alternate view on our failed definition of [map']:
[[
CoFixpoint map' (s : stream A) : stream B :=
match s with
......@@ -384,7 +386,7 @@ Theorem ones_eq' : stream_eq ones ones'.
*)
Abort.
(** The standard [auto] machinery sees that our goal matches an assumption and so applies that assumption, even though this violates guardedness. One usually starts a proof like this by [destruct]ing some parameter and running a custom tactic to figure out the first proof rule to apply for each case. Alternatively, there are tricks that can be played with "hiding" the co-inductive hypothesis.
(** The standard [auto] machinery sees that our goal matches an assumption and so applies that assumption, even though this violates guardedness. A correct proof strategy for a theorem like this usually starts by [destruct]ing some parameter and running a custom tactic to figure out the first proof rule to apply for each case. Alternatively, there are tricks that can be played with "hiding" the co-inductive hypothesis.
%\medskip%
......@@ -404,7 +406,7 @@ Definition hd A (s : stream A) : A :=
Section stream_eq_coind.
Variable A : Type.
Variable R : stream A -> stream A -> Prop.
(** This relation generalizes the theorem we want to prove, characterizing exactly which two arguments to [stream_eq] we want to consider. *)
(** This relation generalizes the theorem we want to prove, defining a set of pairs of streams that we must eventually prove contains the particular pair we care about. *)
Hypothesis Cons_case_hd : forall s1 s2, R s1 s2 -> hd s1 = hd s2.
Hypothesis Cons_case_tl : forall s1 s2, R s1 s2 -> R (tl s1) (tl s2).
......@@ -494,7 +496,7 @@ Eval simpl in approx fact_slow 5.
: list nat
]]
Now, to prove that the two versions are equivalent, it is helpful to prove (and add as a proof hint) a quick lemma about the computational behavior of [fact]. *)
Now, to prove that the two versions are equivalent, it is helpful to prove (and add as a proof hint) a quick lemma about the computational behavior of [fact]. (I intentionally skip explaining its proof at this point.) *)
(* begin thide *)
Lemma fact_def : forall x n,
......@@ -547,7 +549,7 @@ Qed.
(** We close the chapter with a quick motivating example for more complex uses of co-inductive types. We will define a co-inductive semantics for a simple imperative programming language and use that semantics to prove the correctness of a trivial optimization that removes spurious additions by 0. We follow the technique of%\index{co-inductive big-step operational semantics}% _co-inductive big-step operational semantics_ %\cite{BigStep}%.
We define a suggestive synonym for [nat], as we will consider programs with infinitely many variables, represented as [nat]s. *)
We define a suggestive synonym for [nat], as we will consider programs over infinitely many variables, represented as [nat]s. *)
Definition var := nat.
......@@ -578,7 +580,7 @@ Inductive cmd : Set :=
| Seq : cmd -> cmd -> cmd
| While : exp -> cmd -> cmd.
(** We could define an inductive relation to characterize the results of command evaluation. However, such a relation would not capture _nonterminating_ executions. With a co-inductive relation, we can capture both cases. The parameters of the relation are an initial state, a command, and a final state. A program that does not terminate in a particular initial state is related to _any_ final state. *)
(** We could define an inductive relation to characterize the results of command evaluation. However, such a relation would not capture _nonterminating_ executions. With a co-inductive relation, we can capture both cases. The parameters of the relation are an initial state, a command, and a final state. A program that does not terminate in a particular initial state is related to _any_ final state. For more realistic languages than this one, it is often possible for programs to _crash_, in which case a semantics would generally relate their executions to no final states; so relating safely non-terminating programs to all final states provides a crucial distinction. *)
CoInductive evalCmd : vars -> cmd -> vars -> Prop :=
| EvalAssign : forall vs v e, evalCmd vs (Assign v e) (set vs v (evalExp vs e))
......
This diff is collapsed.
This diff is collapsed.
......@@ -56,7 +56,7 @@ With that perspective in mind, this chapter is sort of a mirror image of the las
(** * Propositional Logic *)
(** Let us begin with a brief tour through the definitions of the connectives for propositional logic. We will work within a Coq section that provides us with a set of propositional variables. In Coq parlance, these are just terms of type [Prop.] *)
(** Let us begin with a brief tour through the definitions of the connectives for propositional logic. We will work within a Coq section that provides us with a set of propositional variables. In Coq parlance, these are just variables of type [Prop]. *)
Section Propositional.
Variables P Q R : Prop.
......@@ -101,7 +101,7 @@ We have also already seen the definition of [True]. For a demonstration of a lo
(* begin thide *)
intro.
(** At this point, we have an inconsistent hypothesis [2 + 2 = 5], so the specific conclusion is not important. We use the %\index{tactics!elimtype}%[elimtype] tactic to state a proposition, telling Coq that we wish to construct a proof of the new proposition and then prove the original goal by case analysis on the structure of the new auxiliary proof. Since [False] has no constructors, [elimtype False] simply leaves us with the obligation to prove [False]. *)
(** At this point, we have an inconsistent hypothesis [2 + 2 = 5], so the specific conclusion is not important. We use the %\index{tactics!elimtype}%[elimtype] tactic. For a full description of it, see the Coq manual. For our purposes, we only need the variant [elimtype False], which lets us replace any conclusion formula with [False], because any fact follows from an inconsistent context. *)
elimtype False.
(** [[
......@@ -294,7 +294,7 @@ subgoal 2 is
(* end thide *)
Qed.
(** Sometimes propositional reasoning forms important plumbing for the proof of a theorem, but we still need to apply some other smarts about, say, arithmetic. The tactic %\index{tactics!intuition}%[intuition] is a generalization of [tauto] that proves everything it can using propositional reasoning. When some goals remain, it uses propositional laws to simplify them as far as possible. Consider this example, which uses the list concatenation operator [++] from the standard library. *)
(** Sometimes propositional reasoning forms important plumbing for the proof of a theorem, but we still need to apply some other smarts about, say, arithmetic. The tactic %\index{tactics!intuition}%[intuition] is a generalization of [tauto] that proves everything it can using propositional reasoning. When some further facts must be established to finish the proof, [intuition] uses propositional laws to simplify them as far as possible. Consider this example, which uses the list concatenation operator %\coqdocnotation{%#<tt>#++#</tt>#%}% from the standard library. *)
Theorem arith_comm : forall ls1 ls2 : list nat,
length ls1 = length ls2 \/ length ls1 + length ls2 = 6
......@@ -350,10 +350,12 @@ End Propositional.
(** * What Does It Mean to Be Constructive? *)
(** One potential point of confusion in the presentation so far is the distinction between [bool] and [Prop]. The datatype [bool] is built from two values [true] and [false], while [Prop] is a more primitive type that includes among its members [True] and [False]. Why not collapse these two concepts into one, and why must there be more than two states of mathematical truth?
(** One potential point of confusion in the presentation so far is the distinction between [bool] and [Prop]. The datatype [bool] is built from two values [true] and [false], while [Prop] is a more primitive type that includes among its members [True] and [False]. Why not collapse these two concepts into one, and why must there be more than two states of mathematical truth, [True] and [False]?
The answer comes from the fact that Coq implements%\index{constructive logic}% _constructive_ or%\index{intuitionistic logic|see{constructive logic}}% _intuitionistic_ logic, in contrast to the%\index{classical logic}% _classical_ logic that you may be more familiar with. In constructive logic, classical tautologies like [~ ~ P -> P] and [P \/ ~ P] do not always hold. In general, we can only prove these tautologies when [P] is%\index{decidability}% _decidable_, in the sense of %\index{computability|see{decidability}}%computability theory. The Curry-Howard encoding that Coq uses for [or] allows us to extract either a proof of [P] or a proof of [~ P] from any proof of [P \/ ~ P]. Since our proofs are just functional programs which we can run, a general %\index{law of the excluded middle}%law of the excluded middle would give us a decision procedure for the halting problem, where the instantiations of [P] would be formulas like "this particular Turing machine halts."
A similar paradoxical situation would result if every proposition evaluated to either [True] or [False]. Evaluation in Coq is decidable, so we would be limiting ourselves to decidable propositions only.
Hence the distinction between [bool] and [Prop]. Programs of type [bool] are computational by construction; we can always run them to determine their results. Many [Prop]s are undecidable, and so we can write more expressive formulas with [Prop]s than with [bool]s, but the inevitable consequence is that we cannot simply "run a [Prop] to determine its truth."
Constructive logic lets us define all of the logical connectives in an aesthetically appealing way, with orthogonal inductive definitions. That is, each connective is defined independently using a simple, shared mechanism. Constructivity also enables a trick called%\index{program extraction}% _program extraction_, where we write programs by phrasing them as theorems to be proved. Since our proofs are just functional programs, we can extract executable programs from our final proofs, which we could not do as naturally with classical proofs.
......@@ -375,7 +377,9 @@ We will see more about Coq's program extraction facility in a later chapter. Ho
(Note that here, as always, each [forall] quantifier has the largest possible scope, so that the type of [ex_intro] could also be written [forall x : A, (P x -> ex P)].)
The family [ex] is parameterized by the type [A] that we quantify over, and by a predicate [P] over [A]s. We prove an existential by exhibiting some [x] of type [A], along with a proof of [P x]. As usual, there are tactics that save us from worrying about the low-level details most of the time. We use the equality operator [=], which, depending on the settings in which they learned logic, different people will say either is or is not part of first-order logic. For our purposes, it is. *)
The family [ex] is parameterized by the type [A] that we quantify over, and by a predicate [P] over [A]s. We prove an existential by exhibiting some [x] of type [A], along with a proof of [P x]. As usual, there are tactics that save us from worrying about the low-level details most of the time.
Here is an example of a theorem statement with existential quantification. We use the equality operator [=], which, depending on the settings in which they learned logic, different people will say either is or is not part of first-order logic. For our purposes, it is. *)
Theorem exist1 : exists x : nat, x + 1 = 2.
(* begin thide *)
......@@ -809,7 +813,7 @@ Qed.
(** We write the proof in a way that avoids the use of local variable or hypothesis names, using the %\index{tactics!match}%[match] tactic form to do pattern-matching on the goal. We use unification variables prefixed by question marks in the pattern, and we take advantage of the possibility to mention a unification variable twice in one pattern, to enforce equality between occurrences. The hint to rewrite with [plus_n_Sm] in a particular direction saves us from having to figure out the right place to apply that theorem, and we also take critical advantage of a new tactic, %\index{tactics!eauto}%[eauto].
The [crush] tactic uses the tactic [intuition], which, when it runs out of tricks to try using only propositional logic, by default tries the tactic [auto], which we saw in an earlier example. The [auto] tactic attempts %\index{Prolog}%Prolog-style logic programming, searching through all proof trees up to a certain depth that are built only out of hints that have been registered with [Hint] commands. (See Chapter 13 for a first-principles introduction to what we mean by "Prolog-style logic programming.") Compared to Prolog, [auto] places an important restriction: it never introduces new unification variables during search. That is, every time a rule is applied during proof search, all of its arguments must be deducible by studying the form of the goal. This restriction is relaxed for [eauto], at the cost of possibly exponentially greater running time. In this particular case, we know that [eauto] has only a small space of proofs to search, so it makes sense to run it. It is common in effectively automated Coq proofs to see a bag of standard tactics applied to pick off the "easy" subgoals, finishing with [eauto] to handle the tricky parts that can benefit from ad-hoc exhaustive search.
The [crush] tactic uses the tactic [intuition], which, when it runs out of tricks to try using only propositional logic, by default tries the tactic [auto], which we saw in an earlier example. For now, think of [eauto] as a potentially more expensive version of [auto] that considers more possible proofs; see Chapter 13 for more detail. The quick summary is that [eauto] considers applying a lemma even when the form of the current goal doesn not uniquely determine the values of all of the lemma's quantified variables.
The original theorem now follows trivially from our lemma. *)
......
......@@ -148,7 +148,7 @@ Fixpoint progDenote (p : prog) (s : stack) : option stack :=
(** ** Translation *)
(** Our compiler itself is now unsurprising. The list concatenation operator %\index{Gallina operators!++}%[++] comes from the Coq standard library. *)
(** Our compiler itself is now unsurprising. The list concatenation operator %\index{Gallina operators!++}\coqdocnotation{%#<tt>#++#</tt>#%}% comes from the Coq standard library. *)
Fixpoint compile (e : exp) : prog :=
match e with
......@@ -303,7 +303,7 @@ We only need to unfold the first occurrence of [progDenote] to prove the goal.
]]
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 [None (A:=stack)], which has an annotation specifying that the type of the term ought to be [option stack]. This is phrased as an explicit instantiation of a named type parameter [A] from the definition of [option].
This last [unfold] has left us with an anonymous recursive definition of [progDenote] (similarly to how [fun] or "lambda" constructs in general allow anonymous non-recursive functions), 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 [None (A:=stack)], which has an annotation specifying that the type of the term ought to be [option stack]. This is phrased as an explicit instantiation of a named type parameter [A] from the definition of [option].
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}%
*)
......@@ -390,7 +390,7 @@ We start out the same way as before, introducing new free variables and unfoldin
]]
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}% *)
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}% (Here and elsewhere, it is possible to tell the difference between inputs and outputs to Coq by periods at the ends of the inputs.) *)
Check app_assoc_reverse.
(** %\vspace{-.15in}%[[
......@@ -587,8 +587,7 @@ Definition tbinopDenote arg1 arg2 res (b : tbinop arg1 arg2 res)
(** 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}% _dependent pattern match_, where the necessary _type_ of each case body depends on the _value_ 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.
The same tricks suffice to define an expression denotation function in an unsurprising way:
*)
The same tricks suffice to define an expression denotation function in an unsurprising way. Note that the [type] arguments to the [TBinop] constructor must be included explicitly in pattern-matching, but here we write underscores because we do not need to refer to those arguments directly. *)
Fixpoint texpDenote t (e : texp t) : typeDenote t :=
match e with
......
......@@ -107,7 +107,7 @@ The term [zgtz pf] fails to type-check. Somehow the type checker has failed to
In this case, we must use a [return] annotation to declare the relationship between the _value_ of the [match] discriminee and the _type_ of the result. There is no annotation that lets us declare a relationship between the discriminee and the type of a variable that is already in scope; hence, we delay the binding of [pf], so that we can use the [return] annotation to express the needed relationship.
We are lucky that Coq's heuristics infer the [return] clause (specifically, [return n > 0 -> nat]) for us in this case. *)
We are lucky that Coq's heuristics infer the [return] clause (specifically, [return n > 0 -> nat]) for us in the definition of [pred_strong1], leading to the following elaborated code: *)
Definition pred_strong1' (n : nat) : n > 0 -> nat :=
match n return n > 0 -> nat with
......@@ -139,6 +139,8 @@ let pred_strong1 = function
(** The proof argument has disappeared! We get exactly the OCaml code we would have written manually. This is our first demonstration of the main technically interesting feature of Coq program extraction: proofs are erased systematically.
%\medskip%
We can reimplement our dependently typed [pred] based on%\index{subset types}% _subset types_, defined in the standard library with the type family %\index{Gallina terms!sig}%[sig]. *)
(* begin hide *)
......@@ -389,7 +391,7 @@ Inductive sumbool (A : Prop) (B : Prop) : Set :=
left : A -> {A} + {B} | right : B -> {A} + {B}
]]
We can define some notations to make working with [sumbool] more convenient. *)
Here, the constructors of [sumbool] have types written in terms of a registered notation for [sumbool], such that the result type of each constructor desugars to [sumbool A B]. We can define some notations of our own to make working with [sumbool] more convenient. *)
Notation "'Yes'" := (left _ _).
Notation "'No'" := (right _ _).
......@@ -638,7 +640,7 @@ Eval compute in pred_strong8 0.
(** * Monadic Notations *)
(** We can treat [maybe] like a monad%~\cite{Monads}\index{monad}\index{failure monad}%, in the same way that the Haskell <<Maybe>> type is interpreted as a failure monad. Our [maybe] has the wrong type to be a literal monad, but a "bind"-like notation will still be helpful. *)
(** We can treat [maybe] like a monad%~\cite{Monads}\index{monad}\index{failure monad}%, in the same way that the Haskell <<Maybe>> type is interpreted as a failure monad. Our [maybe] has the wrong type to be a literal monad, but a "bind"-like notation will still be helpful. %Note that the notation definition uses an ASCII \texttt{<-}, while later code uses (in this rendering) a nicer left arrow $\leftarrow$.% *)
Notation "x <- e1 ; e2" := (match e1 with
| Unknown => ??
......@@ -661,7 +663,7 @@ Definition doublePred : forall n1 n2 : nat, {{p | n1 = S (fst p) /\ n2 = S (snd
Defined.
(* end thide *)
(** We can build a [sumor] version of the "bind" notation and use it to write a similarly straightforward version of this function. %The operator rendered here as $\longleftarrow$ is noted in the source as a less-than character followed by two hyphens.% *)
(** We can build a [sumor] version of the "bind" notation and use it to write a similarly straightforward version of this function. %Again, the notation definition exposes the ASCII syntax with an operator \texttt{<-{}-}, while the later code uses a nicer long left arrow $\longleftarrow$.% *)
Notation "x <-- e1 ; e2" := (match e1 with
| inright _ => !!
......@@ -896,7 +898,7 @@ Definition typeCheck' : forall e : exp, {t : type | hasType e t} + {forall t, ~
[||TBool||]
end); clear F; crush' tt hasType; eauto.
(** We clear [F], the local name for the recursive function, to avoid strange proofs that refer to recursive calls that we never make. The [crush] variant %\index{tactics!crush'}%[crush'] helps us by performing automatic inversion on instances of the predicates specified in its second argument. Once we throw in [eauto] to apply [hasType_det] for us, we have discharged all the subgoals. *)
(** We clear [F], the local name for the recursive function, to avoid strange proofs that refer to recursive calls that we never make. Such a step is usually warranted when defining a recursive function with [refine]. The [crush] variant %\index{tactics!crush'}%[crush'] helps us by performing automatic inversion on instances of the predicates specified in its second argument. Once we throw in [eauto] to apply [hasType_det] for us, we have discharged all the subgoals. *)
(* end thide *)
......
......@@ -11,6 +11,14 @@
<webMaster>adam@chlipala.net</webMaster>
<docs>http://blogs.law.harvard.edu/tech/rss</docs>
<item>
<title>Batch of changes based on proofreader feedback</title>
<pubDate>Mon, 8 Oct 2012 16:04:09 EDT</pubDate>
<link>http://adam.chlipala.net/cpdt/</link>
<author>adamc@csail.mit.edu</author>
<description>Thanks to everyone who is helping with the final proofreading!</description>
</item>
<item>
<title>Batch of changes based on proofreader feedback</title>
<pubDate>Tue, 2 Oct 2012 11:34:17 EDT</pubDate>
......
Markdown is supported
0% or
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment