Large.v 44.2 KB
Newer Older
Adam Chlipala's avatar
Adam Chlipala committed
1
(* Copyright (c) 2009-2012, Adam Chlipala
2 3 4 5 6 7 8 9 10
 * 
 * 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 *)
11 12
Require Import Arith.

13
Require Import CpdtTactics.
14 15 16 17 18

Set Implicit Arguments.
(* end hide *)


Adam Chlipala's avatar
Adam Chlipala committed
19 20 21
(** %\part{The Big Picture}

   \chapter{Proving in the Large}% *)
22

23
(** It is somewhat unfortunate that the term "theorem proving" looks so much like the word "theory."  Most researchers and practitioners in software assume that mechanized theorem proving is profoundly impractical.  Indeed, until recently, most advances in theorem proving for higher-order logics have been largely theoretical.  However, starting around the beginning of the 21st century, there was a surge in the use of proof assistants in serious verification efforts.  That line of work is still quite new, but I believe it is not too soon to distill some lessons on how to work effectively with large formal proofs.
24 25 26 27 28 29

   Thus, this chapter gives some tips for structuring and maintaining large Coq developments. *)


(** * Ltac Anti-Patterns *)

30
(** In this book, I have been following an unusual style, where proofs are not considered finished until they are %\index{fully automated proofs}%"fully automated," in a certain sense.  Each such theorem is proved by a single tactic.  Since Ltac is a Turing-complete programming language, it is not hard to squeeze arbitrary heuristics into single tactics, using operators like the semicolon to combine steps.  In contrast, most Ltac proofs "in the wild" consist of many steps, performed by individual tactics followed by periods.  Is it really worth drawing a distinction between proof steps terminated by semicolons and steps terminated by periods?
31

Adam Chlipala's avatar
Adam Chlipala committed
32
   I argue that this is, in fact, a very important distinction, with serious consequences for a majority of important verification domains.  The more uninteresting drudge work a proof domain involves, the more important it is to work to prove theorems with single tactics.  From an automation standpoint, single-tactic proofs can be extremely effective, and automation becomes more and more critical as proofs are populated by more uninteresting detail.  In this section, I will give some examples of the consequences of more common proof styles.
33 34 35 36 37 38 39 40 41 42 43 44 45 46 47 48 49 50 51

   As a running example, consider a basic language of arithmetic expressions, an interpreter for it, and a transformation that scales up every constant in an expression. *)

Inductive exp : Set :=
| Const : nat -> exp
| Plus : exp -> exp -> exp.

Fixpoint eval (e : exp) : nat :=
  match e with
    | Const n => n
    | Plus e1 e2 => eval e1 + eval e2
  end.

Fixpoint times (k : nat) (e : exp) : exp :=
  match e with
    | Const n => Const (k * n)
    | Plus e1 e2 => Plus (times k e1) (times k e2)
  end.

Adam Chlipala's avatar
Adam Chlipala committed
52
(** We can write a very manual proof that [times] really implements multiplication. *)
53 54 55 56 57 58 59 60 61 62 63 64 65 66

Theorem eval_times : forall k e,
  eval (times k e) = k * eval e.
  induction e.

  trivial.

  simpl.
  rewrite IHe1.
  rewrite IHe2.
  rewrite mult_plus_distr_l.
  trivial.
Qed.

Adam Chlipala's avatar
Adam Chlipala committed
67
(* begin thide *)
68
(** We use spaces to separate the two inductive cases, but note that these spaces have no real semantic content; Coq does not enforce that our spacing matches the real case structure of a proof.  The second case mentions automatically generated hypothesis names explicitly.  As a result, innocuous changes to the theorem statement can invalidate the proof. *)
69 70 71

Reset eval_times.

Adam Chlipala's avatar
Adam Chlipala committed
72
Theorem eval_times : forall k x,
73 74 75 76 77 78
  eval (times k x) = k * eval x.
  induction x.

  trivial.

  simpl.
79
(** %\vspace{-.15in}%[[
80
  rewrite IHe1.
81
]]
82

83
<<
84
Error: The reference IHe1 was not found in the current environment.
85
>>
86 87 88 89 90 91 92 93 94 95 96 97 98 99 100 101 102 103 104 105

  The inductive hypotheses are named [IHx1] and [IHx2] now, not [IHe1] and [IHe2]. *)

Abort.

(** We might decide to use a more explicit invocation of [induction] to give explicit binders for all of the names that we will reference later in the proof. *)

Theorem eval_times : forall k e,
  eval (times k e) = k * eval e.
  induction e as [ | ? IHe1 ? IHe2 ].

  trivial.

  simpl.
  rewrite IHe1.
  rewrite IHe2.
  rewrite mult_plus_distr_l.
  trivial.
Qed.

106
(** We pass %\index{tactics!induction}%[induction] an%\index{intro pattern}% _intro pattern_, using a [|] character to separate out instructions for the different inductive cases.  Within a case, we write [?] to ask Coq to generate a name automatically, and we write an explicit name to assign that name to the corresponding new variable.  It is apparent that, to use intro patterns to avoid proof brittleness, one needs to keep track of the seemingly unimportant facts of the orders in which variables are introduced.  Thus, the script keeps working if we replace [e] by [x], but it has become more cluttered.  Arguably, neither proof is particularly easy to follow.
107

Adam Chlipala's avatar
Adam Chlipala committed
108
   That category of complaint has to do with understanding proofs as static artifacts.  As with programming in general, with serious projects, it tends to be much more important to be able to support evolution of proofs as specifications change.  Unstructured proofs like the above examples can be very hard to update in concert with theorem statements.  For instance, consider how the last proof script plays out when we modify [times] to introduce a bug. *)
109 110 111 112 113 114 115 116 117 118 119 120 121 122 123 124

Reset times.

Fixpoint times (k : nat) (e : exp) : exp :=
  match e with
    | Const n => Const (1 + k * n)
    | Plus e1 e2 => Plus (times k e1) (times k e2)
  end.

Theorem eval_times : forall k e,
  eval (times k e) = k * eval e.
  induction e as [ | ? IHe1 ? IHe2 ].

  trivial.

  simpl.
125
(** %\vspace{-.15in}%[[
126
  rewrite IHe1.
127
]]
128

129
<<
130
Error: The reference IHe1 was not found in the current environment.
131
>>
132
  *)
133 134 135

Abort.

136
(** Can you spot what went wrong, without stepping through the script step-by-step?  The problem is that [trivial] never fails.  Originally, [trivial] had been succeeding in proving an equality that follows by reflexivity.  Our change to [times] leads to a case where that equality is no longer true.  The invocation [trivial] happily leaves the false equality in place, and we continue on to the span of tactics intended for the second inductive case.  Unfortunately, those tactics end up being applied to the _first_ case instead.
Adam Chlipala's avatar
Adam Chlipala committed
137

138
   The problem with [trivial] could be "solved" by writing, e.g., [solve [ trivial ]] instead, so that an error is signaled early on if something unexpected happens.  However, the root problem is that the syntax of a tactic invocation does not imply how many subgoals it produces.  Much more confusing instances of this problem are possible.  For example, if a lemma [L] is modified to take an extra hypothesis, then uses of [apply L] will generate more subgoals than before.  Old unstructured proof scripts will become hopelessly jumbled, with tactics applied to inappropriate subgoals.  Because of the lack of structure, there is usually relatively little to be gleaned from knowledge of the precise point in a proof script where an error is raised. *)
139 140 141 142 143 144 145 146 147

Reset times.

Fixpoint times (k : nat) (e : exp) : exp :=
  match e with
    | Const n => Const (k * n)
    | Plus e1 e2 => Plus (times k e1) (times k e2)
  end.

Adam Chlipala's avatar
Adam Chlipala committed
148
(** Many real developments try to make essentially unstructured proofs look structured by applying careful indentation conventions, idempotent case-marker tactics included solely to serve as documentation, and so on.  All of these strategies suffer from the same kind of failure of abstraction that was just demonstrated.  I like to say that if you find yourself caring about indentation in a proof script, it is a sign that the script is structured poorly.
149 150 151 152 153 154 155 156 157 158

   We can rewrite the current proof with a single tactic. *)

Theorem eval_times : forall k e,
  eval (times k e) = k * eval e.
  induction e as [ | ? IHe1 ? IHe2 ]; [
    trivial
    | simpl; rewrite IHe1; rewrite IHe2; rewrite mult_plus_distr_l; trivial ].
Qed.

159
(** We use the form of the semicolon operator that allows a different tactic to be specified for each generated subgoal.  This is an improvement in robustness of the script.  We no longer need to worry about tactics from one case being applied to a different case.  Still, the proof script is not especially readable.  Probably most readers would not find it helpful in explaining why the theorem is true.  The same could be said for scripts which use the%\index{bullets}% _bullets_ or curly braces provided by Coq 8.4, which allow code like the above to be stepped through interactively, with periods in place of the semicolons, while representing proof structure in a way that is enforced by Coq.  Interactive replay of scripts becomes easier, but readability is not really helped.
160 161 162 163 164 165 166 167 168 169 170 171 172 173 174 175 176 177 178 179 180 181 182 183 184 185

   The situation gets worse in considering extensions to the theorem we want to prove.  Let us add multiplication nodes to our [exp] type and see how the proof fares. *)

Reset exp.

Inductive exp : Set :=
| Const : nat -> exp
| Plus : exp -> exp -> exp
| Mult : exp -> exp -> exp.

Fixpoint eval (e : exp) : nat :=
  match e with
    | Const n => n
    | Plus e1 e2 => eval e1 + eval e2
    | Mult e1 e2 => eval e1 * eval e2
  end.

Fixpoint times (k : nat) (e : exp) : exp :=
  match e with
    | Const n => Const (k * n)
    | Plus e1 e2 => Plus (times k e1) (times k e2)
    | Mult e1 e2 => Mult (times k e1) e2
  end.

Theorem eval_times : forall k e,
  eval (times k e) = k * eval e.
186
(** %\vspace{-.25in}%[[
187 188 189
  induction e as [ | ? IHe1 ? IHe2 ]; [
    trivial
    | simpl; rewrite IHe1; rewrite IHe2; rewrite mult_plus_distr_l; trivial ].
190
]]
191

192
<<
193
Error: Expects a disjunctive pattern with 3 branches.
194
>>
195
  *)
196 197 198 199 200 201 202 203 204 205 206 207 208 209 210 211
Abort.

(** Unsurprisingly, the old proof fails, because it explicitly says that there are two inductive cases.  To update the script, we must, at a minimum, remember the order in which the inductive cases are generated, so that we can insert the new case in the appropriate place.  Even then, it will be painful to add the case, because we cannot walk through proof steps interactively when they occur inside an explicit set of cases. *)

Theorem eval_times : forall k e,
  eval (times k e) = k * eval e.
  induction e as [ | ? IHe1 ? IHe2 | ? IHe1 ? IHe2 ]; [
    trivial
    | simpl; rewrite IHe1; rewrite IHe2; rewrite mult_plus_distr_l; trivial
    | simpl; rewrite IHe1; rewrite mult_assoc; trivial ].
Qed.

(** Now we are in a position to see how much nicer is the style of proof that we have followed in most of this book. *)

Reset eval_times.

212
Hint Rewrite mult_plus_distr_l.
Adam Chlipala's avatar
Adam Chlipala committed
213

214 215 216 217
Theorem eval_times : forall k e,
  eval (times k e) = k * eval e.
  induction e; crush.
Qed.
Adam Chlipala's avatar
Adam Chlipala committed
218
(* end thide *)
219

Adam Chlipala's avatar
Adam Chlipala committed
220 221
(** This style is motivated by a hard truth: one person's manual proof script is almost always mostly inscrutable to most everyone else.  I claim that step-by-step formal proofs are a poor way of conveying information.  Thus, we had might as well cut out the steps and automate as much as possible.

222
   What about the illustrative value of proofs?  Most informal proofs are read to convey the big ideas of proofs.  How can reading [induction e; crush] convey any big ideas?  My position is that any ideas that standard automation can find are not very big after all, and the _real_ big ideas should be expressed through lemmas that are added as hints.
Adam Chlipala's avatar
Adam Chlipala committed
223 224 225 226 227 228 229 230 231 232 233 234 235 236 237 238 239 240 241 242 243 244 245

   An example should help illustrate what I mean.  Consider this function, which rewrites an expression using associativity of addition and multiplication. *)

Fixpoint reassoc (e : exp) : exp :=
  match e with
    | Const _ => e
    | Plus e1 e2 =>
      let e1' := reassoc e1 in
      let e2' := reassoc e2 in
        match e2' with
          | Plus e21 e22 => Plus (Plus e1' e21) e22
          | _ => Plus e1' e2'
        end
    | Mult e1 e2 =>
      let e1' := reassoc e1 in
      let e2' := reassoc e2 in
        match e2' with
          | Mult e21 e22 => Mult (Mult e1' e21) e22
          | _ => Mult e1' e2'
        end
  end.

Theorem reassoc_correct : forall e, eval (reassoc e) = eval e.
Adam Chlipala's avatar
Adam Chlipala committed
246
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
247 248
  induction e; crush;
    match goal with
249
      | [ |- context[match ?E with Const _ => _ | _ => _ end] ] =>
Adam Chlipala's avatar
Adam Chlipala committed
250 251 252 253 254 255 256 257 258 259
        destruct E; crush
    end.

  (** One subgoal remains:
     [[
  IHe2 : eval e3 * eval e4 = eval e2
  ============================
   eval e1 * eval e3 * eval e4 = eval e1 * eval e2
   ]]

260
   The [crush] tactic does not know how to finish this goal.  We could finish the proof manually. *)
Adam Chlipala's avatar
Adam Chlipala committed
261 262 263 264 265 266 267 268 269 270 271 272 273 274 275 276

  rewrite <- IHe2; crush.

  (** However, the proof would be easier to understand and maintain if we separated this insight into a separate lemma. *)

Abort.

Lemma rewr : forall a b c d, b * c = d -> a * b * c = a * d.
  crush.
Qed.

Hint Resolve rewr.

Theorem reassoc_correct : forall e, eval (reassoc e) = eval e.
  induction e; crush;
    match goal with
277
      | [ |- context[match ?E with Const _ => _ | _ => _ end] ] =>
Adam Chlipala's avatar
Adam Chlipala committed
278 279 280
        destruct E; crush
    end.
Qed.
Adam Chlipala's avatar
Adam Chlipala committed
281
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
282 283 284

(** In the limit, a complicated inductive proof might rely on one hint for each inductive case.  The lemma for each hint could restate the associated case.  Compared to manual proof scripts, we arrive at more readable results.  Scripts no longer need to depend on the order in which cases are generated.  The lemmas are easier to digest separately than are fragments of tactic code, since lemma statements include complete proof contexts.  Such contexts can only be extracted from monolithic manual proofs by stepping through scripts interactively.

Adam Chlipala's avatar
Adam Chlipala committed
285 286
   The more common situation is that a large induction has several easy cases that automation makes short work of.  In the remaining cases, automation performs some standard simplification.  Among these cases, some may require quite involved proofs; such a case may deserve a hint lemma of its own, where the lemma statement may copy the simplified version of the case.  Alternatively, the proof script for the main theorem may be extended with some automation code targeted at the specific case.  Even such targeted scripting is more desirable than manual proving, because it may be read and understood without knowledge of a proof's hierarchical structure, case ordering, or name binding structure.

287
   A competing alternative to the common style of Coq tactics is the%\index{declarative proof scripts}% _declarative_ style, most frequently associated today with the %\index{Isar}%Isar%~\cite{Isar}% language.  A declarative proof script is very explicit about subgoal structure and introduction of local names, aiming for human readability.  The coding of proof automation is taken to be outside the scope of the proof language, an assumption related to the idea that it is not worth building new automation for each serious theorem.  I have shown in this book many examples of theorem-specific automation, which I believe is crucial for scaling to significant results.  Declarative proof scripts make it easier to read scripts to modify them for theorem statement changes, but the alternate%\index{adaptive proof scripts}% _adaptive_ style from this book allows use of the _same_ scripts for many versions of a theorem.
Adam Chlipala's avatar
Adam Chlipala committed
288 289 290

   Perhaps I am a pessimist for thinking that fully formal proofs will inevitably consist of details that are uninteresting to people, but it is my preference to focus on conveying proof-specific details through choice of lemmas.  Additionally, adaptive Ltac scripts contain bits of automation that can be understood in isolation.  For instance, in a big [repeat match] loop, each case can generally be digested separately, which is a big contrast from trying to understand the hierarchical structure of a script in a more common style.  Adaptive scripts rely on variable binding, but generally only over very small scopes, whereas understanding a traditional script requires tracking the identities of local variables potentially across pages of code.

291
   One might also wonder why it makes sense to prove all theorems automatically (in the sense of adaptive proof scripts) but not construct all programs automatically.  My view there is that _program synthesis_ is a very useful idea that deserves broader application!  In practice, there are difficult obstacles in the way of finding a program automatically from its specification.  A typical specification is not exhaustive in its description of program properties.  For instance, details of performance on particular machine architectures are often omitted.  As a result, a synthesized program may be correct in some sense while suffering from deficiencies in other senses.  Program synthesis research will continue to come up with ways of dealing with this problem, but the situation for theorem proving is fundamentally different.  Following mathematical practice, the only property of a formal proof that we care about is which theorem it proves, and it is trivial to check this property automatically.  In other words, with a simple criterion for what makes a proof acceptable, automatic search is straightforward.  Of course, in practice we also care about understandability of proofs to facilitate long-term maintenance, and that is just what the techniques outlined above are meant to support, and the next section gives some related advice. *)
Adam Chlipala's avatar
Adam Chlipala committed
292

293

Adam Chlipala's avatar
Adam Chlipala committed
294 295
(** * Debugging and Maintaining Automation *)

296
(** Fully automated proofs are desirable because they open up possibilities for automatic adaptation to changes of specification.  A well-engineered script within a narrow domain can survive many changes to the formulation of the problem it solves.  Still, as we are working with higher-order logic, most theorems fall within no obvious decidable theories.  It is inevitable that most long-lived automated proofs will need updating.
Adam Chlipala's avatar
Adam Chlipala committed
297

298
   Before we are ready to update our proofs, we need to write them in the first place.  While fully automated scripts are most robust to changes of specification, it is hard to write every new proof directly in that form.  Instead, it is useful to begin a theorem with exploratory proving and then gradually refine it into a suitable automated form.
Adam Chlipala's avatar
Adam Chlipala committed
299

Adam Chlipala's avatar
Adam Chlipala committed
300
   Consider this theorem from Chapter 8, which we begin by proving in a mostly manual way, invoking [crush] after each step to discharge any low-hanging fruit.  Our manual effort involves choosing which expressions to case-analyze on. *)
Adam Chlipala's avatar
Adam Chlipala committed
301 302 303 304 305 306

(* begin hide *)
Require Import MoreDep.
(* end hide *)

Theorem cfold_correct : forall t (e : exp t), expDenote e = expDenote (cfold e).
Adam Chlipala's avatar
Adam Chlipala committed
307
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
308 309 310 311 312 313 314 315 316 317 318 319 320 321 322 323 324 325 326 327 328 329 330 331 332 333 334 335 336 337
  induction e; crush.

  dep_destruct (cfold e1); crush.
  dep_destruct (cfold e2); crush.

  dep_destruct (cfold e1); crush.
  dep_destruct (cfold e2); crush.

  dep_destruct (cfold e1); crush.
  dep_destruct (cfold e2); crush.

  dep_destruct (cfold e1); crush.
  dep_destruct (expDenote e1); crush.

  dep_destruct (cfold e); crush.

  dep_destruct (cfold e); crush.
Qed.

(** In this complete proof, it is hard to avoid noticing a pattern.  We rework the proof, abstracting over the patterns we find. *)

Reset cfold_correct.

Theorem cfold_correct : forall t (e : exp t), expDenote e = expDenote (cfold e).
  induction e; crush.

  (** The expression we want to destruct here turns out to be the discriminee of a [match], and we can easily enough write a tactic that destructs all such expressions. *)

  Ltac t :=
    repeat (match goal with
338
              | [ |- context[match ?E with NConst _ => _ | _ => _ end] ] =>
Adam Chlipala's avatar
Adam Chlipala committed
339 340 341 342 343 344 345 346 347 348 349 350 351 352 353 354 355 356 357 358 359 360 361 362
                dep_destruct E
            end; crush).

  t.

  (** This tactic invocation discharges the whole case.  It does the same on the next two cases, but it gets stuck on the fourth case. *)

  t.

  t.

  t.

  (** The subgoal's conclusion is:
     [[
    ============================
   (if expDenote e1 then expDenote (cfold e2) else expDenote (cfold e3)) =
   expDenote (if expDenote e1 then cfold e2 else cfold e3)
   ]]

   We need to expand our [t] tactic to handle this case. *)

  Ltac t' :=
    repeat (match goal with
363
              | [ |- context[match ?E with NConst _ => _ | _ => _ end] ] =>
Adam Chlipala's avatar
Adam Chlipala committed
364 365 366 367 368 369 370 371 372 373 374 375 376 377
                dep_destruct E
              | [ |- (if ?E then _ else _) = _ ] => destruct E
            end; crush).

  t'.

  (** Now the goal is discharged, but [t'] has no effect on the next subgoal. *)

  t'.

  (** A final revision of [t] finishes the proof. *)

  Ltac t'' :=
    repeat (match goal with
378
              | [ |- context[match ?E with NConst _ => _ | _ => _ end] ] =>
Adam Chlipala's avatar
Adam Chlipala committed
379 380 381 382 383 384 385 386 387 388 389 390
                dep_destruct E
              | [ |- (if ?E then _ else _) = _ ] => destruct E
              | [ |- context[match pairOut ?E with Some _ => _
                               | None => _ end] ] =>
                dep_destruct E
            end; crush).

  t''.

  t''.
Qed.

391
(** We can take the final tactic and move it into the initial part of the proof script, arriving at a nicely automated proof. *)
Adam Chlipala's avatar
Adam Chlipala committed
392 393 394 395 396 397

Reset t.

Theorem cfold_correct : forall t (e : exp t), expDenote e = expDenote (cfold e).
  induction e; crush;
    repeat (match goal with
398
              | [ |- context[match ?E with NConst _ => _ | _ => _ end] ] =>
Adam Chlipala's avatar
Adam Chlipala committed
399 400 401 402 403 404 405
                dep_destruct E
              | [ |- (if ?E then _ else _) = _ ] => destruct E
              | [ |- context[match pairOut ?E with Some _ => _
                               | None => _ end] ] =>
                dep_destruct E
            end; crush).
Qed.
Adam Chlipala's avatar
Adam Chlipala committed
406
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
407

408
(** Even after we put together nice automated proofs, we must deal with specification changes that can invalidate them.  It is not generally possible to step through single-tactic proofs interactively.  There is a command %\index{Vernacular commands!Debug On}%[Debug On] that lets us step through points in tactic execution, but the debugger tends to make counterintuitive choices of which points we would like to stop at, and per-point output is quite verbose, so most Coq users do not find this debugging mode very helpful.  How are we to understand what has broken in a script that used to work?
409 410 411 412 413 414 415 416 417 418

   An example helps demonstrate a useful approach.  Consider what would have happened in our proof of [reassoc_correct] if we had first added an unfortunate rewriting hint. *)

Reset reassoc_correct.

Theorem confounder : forall e1 e2 e3,
  eval e1 * eval e2 * eval e3 = eval e1 * (eval e2 + 1 - 1) * eval e3.
  crush.
Qed.

419
Hint Rewrite confounder.
420 421

Theorem reassoc_correct : forall e, eval (reassoc e) = eval e.
Adam Chlipala's avatar
Adam Chlipala committed
422
(* begin thide *)
423 424
  induction e; crush;
    match goal with
425
      | [ |- context[match ?E with Const _ => _ | _ => _ end] ] =>
426 427 428 429 430 431 432 433 434 435
        destruct E; crush
    end.

  (** One subgoal remains:
     
     [[
  ============================
   eval e1 * (eval e3 + 1 - 1) * eval e4 = eval e1 * eval e2
   ]]

436
   The poorly chosen rewrite rule fired, changing the goal to a form where another hint no longer applies.  Imagine that we are in the middle of a large development with many hints.  How would we diagnose the problem?  First, we might not be sure which case of the inductive proof has gone wrong.  It is useful to separate out our automation procedure and apply it manually. *)
437 438 439 440

  Restart.

  Ltac t := crush; match goal with
441
                     | [ |- context[match ?E with Const _ => _ | _ => _ end] ] =>
442 443 444 445 446
                       destruct E; crush
                   end.

  induction e.

Adam Chlipala's avatar
Adam Chlipala committed
447
  (** Since we see the subgoals before any simplification occurs, it is clear that this is the case for constants.  Our [t] makes short work of it. *)
448 449 450 451 452 453 454 455 456 457 458
  
  t.

  (** The next subgoal, for addition, is also discharged without trouble. *)

  t.

  (** The final subgoal is for multiplication, and it is here that we get stuck in the proof state summarized above. *)

  t.

459
  (** What is [t] doing to get us to this point?  The %\index{tactics!info}%[info] command can help us answer this kind of question.  (As of this writing, [info] is no longer functioning in the most recent Coq release, but I hope it returns.) *)
460 461 462

  Undo.
  info t.
463

464
  (* begin hide *)
465
  (* begin thide *)
466
  Definition eir := eq_ind_r.
467
  (* end thide *)
468 469
  (* end hide *)

470
  (** %\vspace{-.15in}%[[
471 472
 == simpl in *; intuition; subst; autorewrite with core in *; 
    simpl in *; intuition; subst; autorewrite with core in *; 
473 474 475 476 477
    simpl in *; intuition; subst; destruct (reassoc e2).
    simpl in *; intuition.
    
    simpl in *; intuition.
    
478
    simpl in *; intuition; subst; autorewrite with core in *;
479 480 481
       refine (eq_ind_r
                 (fun n : nat =>
                  n * (eval e3 + 1 - 1) * eval e4 = eval e1 * eval e2) _ IHe1);
482 483
       autorewrite with core in *; simpl in *; intuition; 
    subst; autorewrite with core in *; simpl in *; 
484 485 486 487 488 489 490 491 492 493
    intuition; subst.
 
    ]]

    A detailed trace of [t]'s execution appears.  Since we are using the very general [crush] tactic, many of these steps have no effect and only occur as instances of a more general strategy.  We can copy-and-paste the details to see where things go wrong. *)

  Undo.

  (** We arbitrarily split the script into chunks.  The first few seem not to do any harm. *)

494 495
  simpl in *; intuition; subst; autorewrite with core in *.
  simpl in *; intuition; subst; autorewrite with core in *.
496 497 498 499 500 501
  simpl in *; intuition; subst; destruct (reassoc e2).
  simpl in *; intuition.
  simpl in *; intuition.

  (** The next step is revealed as the culprit, bringing us to the final unproved subgoal. *)

502
  simpl in *; intuition; subst; autorewrite with core in *.
503 504 505 506 507 508 509 510

  (** We can split the steps further to assign blame. *)

  Undo.

  simpl in *.
  intuition.
  subst.
511
  autorewrite with core in *.
512 513 514 515 516

  (** It was the final of these four tactics that made the rewrite.  We can find out exactly what happened.  The [info] command presents hierarchical views of proof steps, and we can zoom down to a lower level of detail by applying [info] to one of the steps that appeared in the original trace. *)

  Undo.

517
  info autorewrite with core in *.
518
  (** %\vspace{-.15in}%[[
519 520 521 522 523 524 525
 == refine (eq_ind_r (fun n : nat => n = eval e1 * eval e2) _
              (confounder (reassoc e1) e3 e4)).
      ]]

      The way a rewrite is displayed is somewhat baroque, but we can see that theorem [confounder] is the final culprit.  At this point, we could remove that hint, prove an alternate version of the key lemma [rewr], or come up with some other remedy.  Fixing this kind of problem tends to be relatively easy once the problem is revealed. *)

Abort.
Adam Chlipala's avatar
Adam Chlipala committed
526
(* end thide *)
527

Adam Chlipala's avatar
Adam Chlipala committed
528 529 530 531
(** Sometimes a change to a development has undesirable performance consequences, even if it does not prevent any old proof scripts from completing.  If the performance consequences are severe enough, the proof scripts can be considered broken for practical purposes.

   Here is one example of a performance surprise. *)

Adam Chlipala's avatar
Adam Chlipala committed
532 533 534
Section slow.
  Hint Resolve trans_eq.

Adam Chlipala's avatar
Adam Chlipala committed
535 536
  (** The central element of the problem is the addition of transitivity as a hint.  With transitivity available, it is easy for proof search to wind up exploring exponential search spaces.  We also add a few other arbitrary variables and hypotheses, designed to lead to trouble later. *)

Adam Chlipala's avatar
Adam Chlipala committed
537 538 539 540 541 542 543
  Variable A : Set.
  Variables P Q R S : A -> A -> Prop.
  Variable f : A -> A.

  Hypothesis H1 : forall x y, P x y -> Q x y -> R x y -> f x = f y.
  Hypothesis H2 : forall x y, S x y -> R x y.

544
  (** We prove a simple lemma very quickly, using the %\index{Vernacular commands!Time}%[Time] command to measure exactly how quickly. *)
Adam Chlipala's avatar
Adam Chlipala committed
545

Adam Chlipala's avatar
Adam Chlipala committed
546
  Lemma slow : forall x y, P x y -> Q x y -> S x y -> f x = f y.
Adam Chlipala's avatar
Adam Chlipala committed
547
    Time eauto 6.
548
(** <<
Adam Chlipala's avatar
Adam Chlipala committed
549
Finished transaction in 0. secs (0.068004u,0.s)
550
>>
551
*)
Adam Chlipala's avatar
Adam Chlipala committed
552

Adam Chlipala's avatar
Adam Chlipala committed
553 554
  Qed.

Adam Chlipala's avatar
Adam Chlipala committed
555 556
  (** Now we add a different hypothesis, which is innocent enough; in fact, it is even provable as a theorem. *)

Adam Chlipala's avatar
Adam Chlipala committed
557 558 559
  Hypothesis H3 : forall x y, x = y -> f x = f y.

  Lemma slow' : forall x y, P x y -> Q x y -> S x y -> f x = f y.
Adam Chlipala's avatar
Adam Chlipala committed
560
    Time eauto 6.
561
    (** <<
Adam Chlipala's avatar
Adam Chlipala committed
562
Finished transaction in 2. secs (1.264079u,0.s)
563
      >>
564
      %\vspace{-.15in}%Why has the search time gone up so much?  The [info] command is not much help, since it only shows the result of search, not all of the paths that turned out to be worthless. *)
Adam Chlipala's avatar
Adam Chlipala committed
565

Adam Chlipala's avatar
Adam Chlipala committed
566
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
567 568
    Restart.
    info eauto 6.
569
    (** %\vspace{-.15in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
570 571
 == intro x; intro y; intro H; intro H0; intro H4;
       simple eapply trans_eq.
572
    simple apply eq_refl.
Adam Chlipala's avatar
Adam Chlipala committed
573 574
    
    simple eapply trans_eq.
575
    simple apply eq_refl.
Adam Chlipala's avatar
Adam Chlipala committed
576 577
    
    simple eapply trans_eq.
578
    simple apply eq_refl.
Adam Chlipala's avatar
Adam Chlipala committed
579 580 581 582 583 584 585 586 587
    
    simple apply H1.
    eexact H.
    
    eexact H0.
    
    simple apply H2; eexact H4.
    ]]

588
    This output does not tell us why proof search takes so long, but it does provide a clue that would be useful if we had forgotten that we added transitivity as a hint.  The [eauto] tactic is applying depth-first search, and the proof script where the real action is ends up buried inside a chain of pointless invocations of transitivity, where each invocation uses reflexivity to discharge one subgoal.  Each increment to the depth argument to [eauto] adds another silly use of transitivity.  This wasted proof effort only adds linear time overhead, as long as proof search never makes false steps.  No false steps were made before we added the new hypothesis, but somehow the addition made possible a new faulty path.  To understand which paths we enabled, we can use the %\index{tactics!debug}%[debug] command. *)
Adam Chlipala's avatar
Adam Chlipala committed
589 590 591 592

    Restart.
    debug eauto 6.

593
    (* begin hide *)
594
    (* begin thide *)
595
    Definition deeeebug := (@eq_refl, @sym_eq).
596
    (* end thide *)
597 598
    (* end hide *)

Adam Chlipala's avatar
Adam Chlipala committed
599 600 601 602 603 604 605 606 607 608
    (** The output is a large proof tree.  The beginning of the tree is enough to reveal what is happening:
       [[
1 depth=6 
1.1 depth=6 intro
1.1.1 depth=6 intro
1.1.1.1 depth=6 intro
1.1.1.1.1 depth=6 intro
1.1.1.1.1.1 depth=6 intro
1.1.1.1.1.1.1 depth=5 apply H3
1.1.1.1.1.1.1.1 depth=4 eapply trans_eq
609
1.1.1.1.1.1.1.1.1 depth=4 apply eq_refl
Adam Chlipala's avatar
Adam Chlipala committed
610
1.1.1.1.1.1.1.1.1.1 depth=3 eapply trans_eq
611
1.1.1.1.1.1.1.1.1.1.1 depth=3 apply eq_refl
Adam Chlipala's avatar
Adam Chlipala committed
612
1.1.1.1.1.1.1.1.1.1.1.1 depth=2 eapply trans_eq
613
1.1.1.1.1.1.1.1.1.1.1.1.1 depth=2 apply eq_refl
Adam Chlipala's avatar
Adam Chlipala committed
614
1.1.1.1.1.1.1.1.1.1.1.1.1.1 depth=1 eapply trans_eq
615
1.1.1.1.1.1.1.1.1.1.1.1.1.1.1 depth=1 apply eq_refl
Adam Chlipala's avatar
Adam Chlipala committed
616 617 618 619 620 621
1.1.1.1.1.1.1.1.1.1.1.1.1.1.1.1 depth=0 eapply trans_eq
1.1.1.1.1.1.1.1.1.1.1.1.1.1.2 depth=1 apply sym_eq ; trivial
1.1.1.1.1.1.1.1.1.1.1.1.1.1.2.1 depth=0 eapply trans_eq
1.1.1.1.1.1.1.1.1.1.1.1.1.1.3 depth=0 eapply trans_eq
1.1.1.1.1.1.1.1.1.1.1.1.2 depth=2 apply sym_eq ; trivial
1.1.1.1.1.1.1.1.1.1.1.1.2.1 depth=1 eapply trans_eq
622
1.1.1.1.1.1.1.1.1.1.1.1.2.1.1 depth=1 apply eq_refl
Adam Chlipala's avatar
Adam Chlipala committed
623 624 625 626 627 628
1.1.1.1.1.1.1.1.1.1.1.1.2.1.1.1 depth=0 eapply trans_eq
1.1.1.1.1.1.1.1.1.1.1.1.2.1.2 depth=1 apply sym_eq ; trivial
1.1.1.1.1.1.1.1.1.1.1.1.2.1.2.1 depth=0 eapply trans_eq
1.1.1.1.1.1.1.1.1.1.1.1.2.1.3 depth=0 eapply trans_eq
       ]]

629
       The first choice [eauto] makes is to apply [H3], since [H3] has the fewest hypotheses of all of the hypotheses and hints that match.  However, it turns out that the single hypothesis generated is unprovable.  That does not stop [eauto] from trying to prove it with an exponentially sized tree of applications of transitivity, reflexivity, and symmetry of equality.  It is the children of the initial [apply H3] that account for all of the noticeable time in proof execution.  In a more realistic development, we might use this output of [debug] to realize that adding transitivity as a hint was a bad idea. *)
Adam Chlipala's avatar
Adam Chlipala committed
630

Adam Chlipala's avatar
Adam Chlipala committed
631
  Qed.
Adam Chlipala's avatar
Adam Chlipala committed
632
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
633 634
End slow.

Adam Chlipala's avatar
Adam Chlipala committed
635 636
(** As aggravating as the above situation may be, there is greater aggravation to be had from importing library modules with commands like %\index{Vernacular commands!Require Import}%[Require Import].  Such a command imports not just the Gallina terms from a module, but also all the hints for [auto], [eauto], and [autorewrite].  Some very recent versions of Coq include mechanisms for removing hints from databases, but the proper solution is to be very conservative in exporting hints from modules.  Consider putting hints in named databases, so that they may be used only when called upon explicitly, as demonstrated in Chapter 13.

637
It is also easy to end up with a proof script that uses too much memory.  As tactics run, they avoid generating proof terms, since serious proof search will consider many possible avenues, and we do not want to build proof terms for subproofs that end up unused.  Instead, tactic execution maintains%\index{thunks}% _thunks_ (suspended computations, represented with closures), such that a tactic's proof-producing thunk is only executed when we run %\index{Vernacular commands!Qed}%[Qed].  These thunks can use up large amounts of space, such that a proof script exhausts available memory, even when we know that we could have used much less memory by forcing some thunks earlier.
Adam Chlipala's avatar
Adam Chlipala committed
638

639
   The %\index{tactics!abstract}%[abstract] tactical helps us force thunks by proving some subgoals as their own lemmas.  For instance, a proof [induction x; crush] can in many cases be made to use significantly less peak memory by changing it to [induction x; abstract crush].  The main limitation of [abstract] is that it can only be applied to subgoals that are proved completely, with no undetermined unification variables in their initial states.  Still, many large automated proofs can realize vast memory savings via [abstract]. *)
Adam Chlipala's avatar
Adam Chlipala committed
640

Adam Chlipala's avatar
Adam Chlipala committed
641

642 643
(** * Modules *)

644
(** Last chapter's examples of proof by reflection demonstrate opportunities for implementing abstract proof strategies with stronger formal guarantees than can be had with Ltac scripting.  Coq's _module system_ provides another tool for more rigorous development of generic theorems.  This feature is inspired by the module systems found in Standard ML%~\cite{modules}% and OCaml, and the discussion that follows assumes familiarity with the basics of one of those systems.
645

646
   ML modules facilitate the grouping of %\index{abstract type}%abstract types with operations over those types.  Moreover, there is support for%\index{functor}% _functors_, which are functions from modules to modules.  A canonical example of a functor is one that builds a data structure implementation from a module that describes a domain of keys and its associated comparison operations.
647

648
   When we add modules to a base language with dependent types, it becomes possible to use modules and functors to formalize kinds of reasoning that are common in algebra.  For instance, the following module signature captures the essence of the algebraic structure known as a group.  A group consists of a carrier set [G], an associative binary operation [f], a left identity element [id] for [f], and an operation [i] that is a left inverse for [f].%\index{Vernacular commands!Module Type}% *)
649

650 651 652
Module Type GROUP.
  Parameter G : Set.
  Parameter f : G -> G -> G.
653
  Parameter id : G.
654 655 656
  Parameter i : G -> G.

  Axiom assoc : forall a b c, f (f a b) c = f a (f b c).
657 658
  Axiom ident : forall a, f id a = a.
  Axiom inverse : forall a, f (i a) a = id.
659 660
End GROUP.

661
(** Many useful theorems hold of arbitrary groups.  We capture some such theorem statements in another module signature.%\index{Vernacular commands!Declare Module}% *)
662

663 664 665
Module Type GROUP_THEOREMS.
  Declare Module M : GROUP.

666
  Axiom ident' : forall a, M.f a M.id = a.
667

668
  Axiom inverse' : forall a, M.f a (M.i a) = M.id.
669

670
  Axiom unique_ident : forall id', (forall a, M.f id' a = a) -> id' = M.id.
671 672
End GROUP_THEOREMS.

Adam Chlipala's avatar
Adam Chlipala committed
673 674 675
(** We implement generic proofs of these theorems with a functor, whose input is an arbitrary group [M]. %\index{Vernacular commands!Module}% *)

Module GroupProofs (M : GROUP) : GROUP_THEOREMS with Module M := M.
676
  (** As in ML, Coq provides multiple options for ascribing signatures to modules.  Here we use just the colon operator, which implements%\index{opaque ascription}% _opaque ascription_, hiding all details of the module not exposed by the signature.  Another option is%\index{transparent ascription}% _transparent ascription_ via the [<:] operator, which checks for signature compatibility without hiding implementation details.  Here we stick with opaque ascription but employ the [with] operation to add more detail to a signature, exposing just those implementation details that we need to.  For instance, here we expose the underlying group representation set and operator definitions.  Without such a refinement, we would get an output module proving theorems about some unknown group, which is not very useful.  Also note that opaque ascription can in Coq have some undesirable consequences without analogues in ML, since not just the types but also the _definitions_ of identifiers have significance in type checking and theorem proving. *)
677

678
  Module M := M.
Adam Chlipala's avatar
Adam Chlipala committed
679
  (** To ensure that the module we are building meets the [GROUP_THEOREMS] signature, we add an extra local name for [M], the functor argument. *)
680 681

  Import M.
Adam Chlipala's avatar
Adam Chlipala committed
682 683 684
  (** It would be inconvenient to repeat the prefix [M.] everywhere in our theorem statements and proofs, so we bring all the identifiers of [M] into the local scope unqualified.

     Now we are ready to prove the three theorems.  The proofs are completely manual, which may seem ironic given the content of the previous sections!  This illustrates another lesson, which is that short proof scripts that change infrequently may be worth leaving unautomated.  It would take some effort to build suitable generic automation for these theorems about groups, so I stick with manual proof scripts to avoid distracting us from the main message of the section.  We take the proofs from the Wikipedia page on elementary group theory. *)
685

686
  Theorem inverse' : forall a, f a (i a) = id.
687 688 689 690 691 692 693 694 695 696 697
    intro.
    rewrite <- (ident (f a (i a))).
    rewrite <- (inverse (f a (i a))) at 1.
    rewrite assoc.
    rewrite assoc.
    rewrite <- (assoc (i a) a (i a)).
    rewrite inverse.
    rewrite ident.
    apply inverse.
  Qed.

698
  Theorem ident' : forall a, f a id = a.
699 700 701 702 703 704 705
    intro.
    rewrite <- (inverse a).
    rewrite <- assoc.
    rewrite inverse'.
    apply ident.
  Qed.

706
  Theorem unique_ident : forall id', (forall a, M.f id' a = a) -> id' = M.id.
707
    intros.
708
    rewrite <- (H id).
709 710 711
    symmetry.
    apply ident'.
  Qed.
Adam Chlipala's avatar
Adam Chlipala committed
712
End GroupProofs.
Adam Chlipala's avatar
Adam Chlipala committed
713

714 715
(** We can show that the integers with [+] form a group. *)

Adam Chlipala's avatar
Adam Chlipala committed
716 717 718 719 720 721
Require Import ZArith.
Open Scope Z_scope.

Module Int.
  Definition G := Z.
  Definition f x y := x + y.
722
  Definition id := 0.
Adam Chlipala's avatar
Adam Chlipala committed
723 724 725 726 727
  Definition i x := -x.

  Theorem assoc : forall a b c, f (f a b) c = f a (f b c).
    unfold f; crush.
  Qed.
728 729
  Theorem ident : forall a, f id a = a.
    unfold f, id; crush.
Adam Chlipala's avatar
Adam Chlipala committed
730
  Qed.
731 732
  Theorem inverse : forall a, f (i a) a = id.
    unfold f, i, id; crush.
Adam Chlipala's avatar
Adam Chlipala committed
733 734 735
  Qed.
End Int.

736 737
(** Next, we can produce integer-specific versions of the generic group theorems. *)

Adam Chlipala's avatar
Adam Chlipala committed
738
Module IntProofs := GroupProofs(Int).
Adam Chlipala's avatar
Adam Chlipala committed
739

Adam Chlipala's avatar
Adam Chlipala committed
740
Check IntProofs.unique_ident.
741
(** %\vspace{-.15in}% [[
Adam Chlipala's avatar
Adam Chlipala committed
742
  IntProofs.unique_ident
743
     : forall e' : Int.G, (forall a : Int.G, Int.f e' a = a) -> e' = Int.e
744
     ]]
745 746

Projections like [Int.G] are known to be definitionally equal to the concrete values we have assigned to them, so the above theorem yields as a trivial corollary the following more natural restatement: *)
Adam Chlipala's avatar
Adam Chlipala committed
747

748
Theorem unique_ident : forall id', (forall a, id' + a = a) -> id' = 0.
Adam Chlipala's avatar
Adam Chlipala committed
749
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
750
  exact IntProofs.unique_ident.
Adam Chlipala's avatar
Adam Chlipala committed
751
Qed.
Adam Chlipala's avatar
Adam Chlipala committed
752
(* end thide *)
753

754
(** As in ML, the module system provides an effective way to structure large developments.  Unlike in ML, Coq modules add no expressiveness; we can implement any module as an inhabitant of a dependent record type.  It is the second-class nature of modules that makes them easier to use than dependent records in many case.  Because modules may only be used in quite restricted ways, it is easier to support convenient module coding through special commands and editing modes, as the above example demonstrates.  An isomorphic implementation with records would have suffered from lack of such conveniences as module subtyping and importation of the fields of a module.  On the other hand, all module values must be determined statically, so modules may not be computed, e.g., within the definitions of normal functions, based on particular function parameters. *)
Adam Chlipala's avatar
Adam Chlipala committed
755 756 757 758


(** * Build Processes *)

759
(* begin hide *)
760
(* begin thide *)
761 762 763 764 765 766 767 768 769 770 771 772 773 774
Module Lib.
  Module A.
  End A.
  Module B.
  End B.
  Module C.
  End C.
End Lib.
Module Client.
  Module D.
  End D.
  Module E.
  End E.
End Client.
775
(* end thide *)
776 777
(* end hide *)

Adam Chlipala's avatar
Adam Chlipala committed
778 779
(** As in software development, large Coq projects are much more manageable when split across multiple files and when decomposed into libraries.  Coq and Proof General provide very good support for these activities.

780
   Consider a library that we will name [Lib], housed in directory <<LIB>> and split between files <<A.v>>, <<B.v>>, and <<C.v>>.  A simple %\index{Makefile}%Makefile will compile the library, relying on the standard Coq tool %\index{coq\_makefile}%<<coq_makefile>> to do the hard work.
Adam Chlipala's avatar
Adam Chlipala committed
781 782 783 784 785 786 787 788

<<
MODULES := A B C
VS      := $(MODULES:%=%.v)

.PHONY: coq clean

coq: Makefile.coq
789
        $(MAKE) -f Makefile.coq
Adam Chlipala's avatar
Adam Chlipala committed
790 791 792 793 794

Makefile.coq: Makefile $(VS)
        coq_makefile -R . Lib $(VS) -o Makefile.coq

clean:: Makefile.coq
795
        $(MAKE) -f Makefile.coq clean
Adam Chlipala's avatar
Adam Chlipala committed
796 797 798
        rm -f Makefile.coq
>>

799
   The Makefile begins by defining a variable <<VS>> holding the list of filenames to be included in the project.  The primary target is <<coq>>, which depends on the construction of an auxiliary Makefile called <<Makefile.coq>>.  Another rule explains how to build that file.  We call <<coq_makefile>>, using the <<-R>> flag to specify that files in the current directory should be considered to belong to the library [Lib].  This Makefile will build a compiled version of each module, such that <<X.v>> is compiled into <<X.vo>>.
Adam Chlipala's avatar
Adam Chlipala committed
800

801
   Now code in <<B.v>> may refer to definitions in <<A.v>> after running
Adam Chlipala's avatar
Adam Chlipala committed
802 803 804
   [[
Require Import Lib.A.
   ]]
805
   %\vspace{-.15in}%Library [Lib] is presented as a module, containing a submodule [A], which contains the definitions from <<A.v>>.  These are genuine modules in the sense of Coq's module system, and they may be passed to functors and so on.
Adam Chlipala's avatar
Adam Chlipala committed
806

807
   The command [Require Import] is a convenient combination of two more primitive commands.  The %\index{Vernacular commands!Require}%[Require] command finds the <<.vo>> file containing the named module, ensuring that the module is loaded into memory.  The %\index{Vernacular commands!Import}%[Import] command loads all top-level definitions of the named module into the current namespace, and it may be used with local modules that do not have corresponding <<.vo>> files.  Another command, %\index{Vernacular commands!Load}%[Load], is for inserting the contents of a named file verbatim.  It is generally better to use the module-based commands, since they avoid rerunning proof scripts, and they facilitate reorganization of directory structure without the need to change code.
Adam Chlipala's avatar
Adam Chlipala committed
808

809
   Now we would like to use our library from a different development, called [Client] and found in directory <<CLIENT>>, which has its own Makefile.
Adam Chlipala's avatar
Adam Chlipala committed
810 811 812 813 814 815 816 817

<<
MODULES := D E
VS      := $(MODULES:%=%.v)

.PHONY: coq clean

coq: Makefile.coq
818
        $(MAKE) -f Makefile.coq
Adam Chlipala's avatar
Adam Chlipala committed
819 820 821 822 823

Makefile.coq: Makefile $(VS)
        coq_makefile -R LIB Lib -R . Client $(VS) -o Makefile.coq

clean:: Makefile.coq
824
        $(MAKE) -f Makefile.coq clean
Adam Chlipala's avatar
Adam Chlipala committed
825 826 827
        rm -f Makefile.coq
>>

828
   We change the <<coq_makefile>> call to indicate where the library [Lib] is found.  Now <<D.v>> and <<E.v>> can refer to definitions from [Lib] module [A] after running
Adam Chlipala's avatar
Adam Chlipala committed
829 830 831
   [[
Require Import Lib.A.
   ]]
832
   %\vspace{-.15in}\noindent{}%and <<E.v>> can refer to definitions from <<D.v>> by running
Adam Chlipala's avatar
Adam Chlipala committed
833 834 835
   [[
Require Import Client.D.
   ]]
836
   %\vspace{-.15in}%It can be useful to split a library into several files, but it is also inconvenient for client code to import library modules individually.  We can get the best of both worlds by, for example, adding an extra source file <<Lib.v>> to [Lib]'s directory and Makefile, where that file contains just this line:%\index{Vernacular commands!Require Export}%
Adam Chlipala's avatar
Adam Chlipala committed
837 838 839
   [[
Require Export Lib.A Lib.B Lib.C.
   ]]
840
   %\vspace{-.15in}%Now client code can import all definitions from all of [Lib]'s modules simply by running
Adam Chlipala's avatar
Adam Chlipala committed
841 842 843
   [[
Require Import Lib.
   ]]
844
   %\vspace{-.15in}%The two Makefiles above share a lot of code, so, in practice, it is useful to define a common Makefile that is included by multiple library-specific Makefiles.
Adam Chlipala's avatar
Adam Chlipala committed
845 846 847

   %\medskip%

848
   The remaining ingredient is the proper way of editing library code files in Proof General.  Recall this snippet of <<.emacs>> code from Chapter 2, which tells Proof General where to find the library associated with this book.
Adam Chlipala's avatar
Adam Chlipala committed
849 850 851 852 853 854 855 856 857 858 859 860 861 862 863 864 865 866 867 868

<<
(custom-set-variables
  ...
  '(coq-prog-args '("-I" "/path/to/cpdt/src"))
  ...
)
>>

   To do interactive editing of our current example, we just need to change the flags to point to the right places.

<<
(custom-set-variables
  ...
; '(coq-prog-args '("-I" "/path/to/cpdt/src"))
  '(coq-prog-args '("-R" "LIB" "Lib" "-R" "CLIENT" "Client"))
  ...
)
>>

869
   When working on multiple projects, it is useful to leave multiple versions of this setting in your <<.emacs>> file, commenting out all but one of them at any moment in time.  To switch between projects, change the commenting structure and restart Emacs.
870

871
   Alternatively, we can revisit the directory-local settings approach and write the following into a file <<.dir-locals.el>> in <<CLIENT>>:
872 873 874 875 876 877

<<
((coq-mode . ((coq-prog-args .
  ("-emacs-U" "-R" "LIB" "Lib" "-R" "CLIENT" "Client")))))
>>
*)