LogicProg.v 34.2 KB
Newer Older
1
(* Copyright (c) 2011-2012, Adam Chlipala
Adam Chlipala's avatar
Adam Chlipala committed
2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19 20 21 22 23
 * 
 * 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 CpdtTactics.

Set Implicit Arguments.

(* end hide *)

(** %\part{Proof Engineering}

   \chapter{Proof Search by Logic Programming}% *)

24 25 26
(** The Curry-Howard correspondence tells us that proving is %``%#"#just#"#%''% programming, but the pragmatics of the two activities are very different.  Generally we care about properties of a program besides its type, but the same is not true about proofs.  Any proof of a theorem will do just as well.  As a result, automated proof search is conceptually simpler than automated programming.

   The paradigm of %\index{logic programming}%logic programming, as embodied in languages like %\index{Prolog}%Prolog, is a good match for proof search in higher-order logic.  This chapter introduces the details, attempting to avoid any dependence on past logic programming experience. *)
Adam Chlipala's avatar
Adam Chlipala committed
27 28 29 30


(** * Introducing Logic Programming *)

31 32
(** Recall the definition of addition from the standard library. *)

Adam Chlipala's avatar
Adam Chlipala committed
33
Print plus.
34 35 36 37 38 39 40 41 42 43
(** %\vspace{-.15in}%[[
plus = 
fix plus (n m : nat) : nat := match n with
                              | 0 => m
                              | S p => S (plus p m)
                              end

]]

This is a recursive definition, in the style of functional programming.  We might also follow the style of logic programming, which corresponds to the inductive relations we have defined in previous chapters. *)
Adam Chlipala's avatar
Adam Chlipala committed
44 45 46 47 48 49

Inductive plusR : nat -> nat -> nat -> Prop :=
| PlusO : forall m, plusR O m m
| PlusS : forall n m r, plusR n m r
  -> plusR (S n) m (S r).

50 51
(** Intuitively, a fact [plusR n m r] only holds when [plus n m = r].  It is not hard to prove this correspondence formally. *)

Adam Chlipala's avatar
Adam Chlipala committed
52 53 54 55 56 57 58 59 60 61 62 63 64 65 66 67 68 69 70
(* begin thide *)
Hint Constructors plusR.
(* end thide *)

Theorem plus_plusR : forall n m,
  plusR n m (n + m).
(* begin thide *)
  induction n; crush.
Qed.
(* end thide *)

Theorem plusR_plus : forall n m r,
  plusR n m r
  -> r = n + m.
(* begin thide *)
  induction 1; crush.
Qed.
(* end thide *)

71 72
(** With the functional definition of [plus], simple equalities about arithmetic follow by computation. *)

Adam Chlipala's avatar
Adam Chlipala committed
73 74 75 76 77 78
Example four_plus_three : 4 + 3 = 7.
(* begin thide *)
  reflexivity.
Qed.
(* end thide *)

79 80 81 82 83 84 85 86
Print four_plus_three.
(** %\vspace{-.15in}%[[
four_plus_three = eq_refl
]]

With the relational definition, the same equalities take more steps to prove, but the process is completely mechanical.  For example, consider this simple-minded manual proof search strategy.  The steps with error messages shown afterward will be omitted from the final script.
*)

Adam Chlipala's avatar
Adam Chlipala committed
87 88
Example four_plus_three' : plusR 4 3 7.
(* begin thide *)
89 90 91 92 93 94 95 96 97 98 99 100 101 102 103 104 105 106 107 108 109 110 111 112 113 114 115 116 117 118 119 120 121 122 123 124 125
(** %\vspace{-.2in}%[[
  apply PlusO.
]]
%\vspace{-.2in}%
<<
Error: Impossible to unify "plusR 0 ?24 ?24" with "plusR 4 3 7".
>> *)
  apply PlusS.
(** %\vspace{-.2in}%[[
  apply PlusO.
]]
%\vspace{-.2in}%
<<
Error: Impossible to unify "plusR 0 ?25 ?25" with "plusR 3 3 6".
>> *)
  apply PlusS.
(** %\vspace{-.2in}%[[
  apply PlusO.
]]
%\vspace{-.2in}%
<<
Error: Impossible to unify "plusR 0 ?26 ?26" with "plusR 2 3 5".
>> *)
  apply PlusS.
(** %\vspace{-.2in}%[[
  apply PlusO.
]]
%\vspace{-.2in}%
<<
Error: Impossible to unify "plusR 0 ?27 ?27" with "plusR 1 3 4".
>> *)
  apply PlusS.
  apply PlusO.

(** At this point the proof is completed.  It is no doubt clear that a simple procedure could find all proofs of this kind for us.  We are just exploring all possible proof trees, built from the two candidate steps [apply PlusO] and [apply PlusS].  The built-in tactic %\index{tactics!auto}%[auto] does exactly that, since above we used [Hint Constructors] to register the two candidate proof steps as hints. *)

Restart.
Adam Chlipala's avatar
Adam Chlipala committed
126 127 128 129
  auto.
Qed.
(* end thide *)

130 131 132 133 134 135 136 137 138
Print four_plus_three'.
(** %\vspace{-.15in}%[[
four_plus_three' = PlusS (PlusS (PlusS (PlusS (PlusO 3))))
]]
*)

(** Let us try the same approach on a slightly more complex goal. *)

Example five_plus_three : plusR 5 3 8.
Adam Chlipala's avatar
Adam Chlipala committed
139
(* begin thide *)
140 141 142 143
  auto.

(** This time, [auto] is not enough to make any progress.  Since even a single candidate step may lead to an infinite space of possible proof trees, [auto] is parameterized on the maximum depth of trees to consider.  The default depth is 5, and it turns out that we need depth 6 to prove the goal. *)

Adam Chlipala's avatar
Adam Chlipala committed
144
  auto 6.
145 146 147

(** Sometimes it is useful to see a description of the proof tree that [auto] finds, with the %\index{tactics!info}%[info] tactical. *)

Adam Chlipala's avatar
Adam Chlipala committed
148 149
Restart.
  info auto 6.
150 151 152 153 154
(** %\vspace{-.15in}%[[
 == apply PlusS; apply PlusS; apply PlusS; apply PlusS; 
    apply PlusS; apply PlusO.
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
155 156 157
Qed.
(* end thide *)

158
(** The two key components of logic programming are %\index{backtracking}\emph{%#<i>#backtracking#</i>#%}% and %\index{unification}\emph{%#<i>#unification#</i>#%}%.  To see these techniques in action, consider this further silly example.  Here our candidate proof steps will be reflexivity and quantifier instantiation. *)
Adam Chlipala's avatar
Adam Chlipala committed
159 160 161

Example seven_minus_three : exists x, x + 3 = 7.
(* begin thide *)
162 163 164 165 166 167 168 169 170 171 172 173 174 175 176 177 178 179
(** For explanatory purposes, let us simulate a user with minimal understanding of arithmetic.  We start by choosing an instantiation for the quantifier.  Recall that [ex_intro] is the constructor for existentially quantified formulas. *)

  apply ex_intro with 0.
(** %\vspace{-.2in}%[[
  reflexivity.
]]
%\vspace{-.2in}%
<<
  Error: Impossible to unify "7" with "0 + 3".
>>

This seems to be a dead end.  Let us %\emph{%#<i>#backtrack#</i>#%}% to the point where we ran [apply] and make a better alternate choice.
*)

Restart.
  apply ex_intro with 4.
  reflexivity.
Qed.
Adam Chlipala's avatar
Adam Chlipala committed
180 181
(* end thide *)

182 183 184 185
(** The above was a fairly tame example of backtracking.  In general, any node in an under-construction proof tree may be the destination of backtracking an arbitrarily large number of times, as different candidate proof steps are found not to lead to full proof trees, within the depth bound passed to [auto].

   Next we demonstrate unification, which will be easier when we switch to the relational formulation of addition. *)

Adam Chlipala's avatar
Adam Chlipala committed
186 187
Example seven_minus_three' : exists x, plusR x 3 7.
(* begin thide *)
188 189 190 191 192 193 194 195 196 197 198 199 200 201 202 203 204 205 206 207 208 209 210 211 212 213 214 215 216
(** We could attempt to guess the quantifier instantiation manually as before, but here there is no need.  Instead of [apply], we use %\index{tactics!eapply}%[eapply] instead, which proceeds with placeholder %\index{unification variable}\emph{%#<i>#unification variables#</i>#%}% standing in for those parameters we wish to postpone guessing. *)

  eapply ex_intro.
(** [[
1 subgoal
  
  ============================
   plusR ?70 3 7
]]

Now we can finish the proof with the right applications of [plusR]'s constructors.  Note that new unification variables are being generated to stand for new unknowns. *)

  apply PlusS.
(** [[
  ============================
   plusR ?71 3 6
]]
*)
  apply PlusS. apply PlusS. apply PlusS.
(** [[
  ============================
   plusR ?74 3 3
]]
*)
  apply PlusO.

(** The [auto] tactic will not perform these sorts of steps that introduce unification variables, but the %\index{tactics!eauto}%[eauto] tactic will.  It is helpful to work with two separate tactics, because proof search in the [eauto] style can uncover many more potential proof trees and hence take much longer to run. *)

Restart.
Adam Chlipala's avatar
Adam Chlipala committed
217
  info eauto 6.
218 219 220 221 222
(** %\vspace{-.15in}%[[
 == eapply ex_intro; apply PlusS; apply PlusS; 
    apply PlusS; apply PlusS; apply PlusO.
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
223 224 225
Qed.
(* end thide *)

226 227
(** This proof gives us our first example where logic programming simplifies proof search compared to functional programming.  In general, functional programs are only meant to be run in a single direction; a function has disjoint sets of inputs and outputs.  In the last example, we effectively ran a logic program backwards, deducing an input that gives rise to a certain output.  The same works for deducing an unknown value of the other input. *)

Adam Chlipala's avatar
Adam Chlipala committed
228 229
Example seven_minus_four' : exists x, plusR 4 x 7.
(* begin thide *)
230
  eauto 6.
Adam Chlipala's avatar
Adam Chlipala committed
231 232 233
Qed.
(* end thide *)

234 235
(** By proving the right auxiliary facts, we can reason about specific functional programs in the same way as we did above for a logic program.  Let us prove that the constructors of [plusR] have natural interpretations as lemmas about [plus].  We can find the first such lemma already proved in the standard library, using the %\index{Vernacular commands!SearchRewrite}%[SearchRewrite] command to find a library function proving an equality whose lefthand or righthand side matches a pattern with wildcards. *)

Adam Chlipala's avatar
Adam Chlipala committed
236 237
(* begin thide *)
SearchRewrite (O + _).
238 239 240 241 242
(** %\vspace{-.15in}%[[
plus_O_n: forall n : nat, 0 + n = n
]]

The command %\index{Vernacular commands!Hint Immediate}%[Hint Immediate] asks [auto] and [eauto] to consider this lemma as a candidate step for any leaf of a proof tree. *)
Adam Chlipala's avatar
Adam Chlipala committed
243 244 245

Hint Immediate plus_O_n.

246 247
(** The counterpart to [PlusS] we will prove ourselves. *)

Adam Chlipala's avatar
Adam Chlipala committed
248 249 250 251 252 253
Lemma plusS : forall n m r,
  n + m = r
  -> S n + m = S r.
  crush.
Qed.

254 255
(** The command %\index{Vernacular commands!Hint Resolve}%[Hint Resolve] adds a new candidate proof step, to be attempted at any level of a proof tree, not just at leaves. *)

Adam Chlipala's avatar
Adam Chlipala committed
256 257 258
Hint Resolve plusS.
(* end thide *)

259 260 261
(** Now that we have registered the proper hints, we can replicate our previous examples with the normal, functional addition [plus]. *)

Example seven_minus_three'' : exists x, x + 3 = 7.
Adam Chlipala's avatar
Adam Chlipala committed
262
(* begin thide *)
263
  eauto 6.
Adam Chlipala's avatar
Adam Chlipala committed
264 265 266 267 268
Qed.
(* end thide *)

Example seven_minus_four : exists x, 4 + x = 7.
(* begin thide *)
269
  eauto 6.
Adam Chlipala's avatar
Adam Chlipala committed
270 271 272
Qed.
(* end thide *)

273 274 275
(** This new hint database is far from a complete decision procedure, as we see in a further example that [eauto] does not finish. *)

Example seven_minus_four_zero : exists x, 4 + x + 0 = 7.
Adam Chlipala's avatar
Adam Chlipala committed
276 277 278 279 280
(* begin thide *)
  eauto 6.
Abort.
(* end thide *)

281 282
(** A further lemma will be helpful. *)

Adam Chlipala's avatar
Adam Chlipala committed
283 284 285 286 287 288 289 290 291 292
(* begin thide *)
Lemma plusO : forall n m,
  n = m
  -> n + 0 = m.
  crush.
Qed.

Hint Resolve plusO.
(* end thide *)

293 294
(** Note that, if we consider the inputs to [plus] as the inputs of a corresponding logic program, the new rule [plusO] introduces an ambiguity.  For instance, a sum [0 + 0] would match both of [plus_O_n] and [plusO], depending on which operand we focus on.  This ambiguity may increase the number of potential search trees, slowing proof search, but semantically it presents no problems, and in fact it leads to an automated proof of the present example. *)

Adam Chlipala's avatar
Adam Chlipala committed
295 296
Example seven_minus_four_zero : exists x, 4 + x + 0 = 7.
(* begin thide *)
297
  eauto 7.
Adam Chlipala's avatar
Adam Chlipala committed
298 299 300
Qed.
(* end thide *)

301 302
(** Just how much damage can be done by adding hints that grow the space of possible proof trees?  A classic gotcha comes from unrestricted use of transitivity, as embodied in this library theorem about equality: *)

Adam Chlipala's avatar
Adam Chlipala committed
303
Check eq_trans.
304 305 306 307 308 309 310
(** %\vspace{-.15in}%[[
eq_trans
     : forall (A : Type) (x y z : A), x = y -> y = z -> x = z
]]
*)

(** Hints are scoped over sections, so let us enter a section to contain the effects of an unfortunate hint choice. *)
Adam Chlipala's avatar
Adam Chlipala committed
311 312 313 314

Section slow.
  Hint Resolve eq_trans.

315 316
  (** The following fact is false, but that does not stop [eauto] from taking a very long time to search for proofs of it.  We use the handy %\index{Vernacular commands!Time}%[Time] command to measure how long a proof step takes to run.  None of the following steps make any progress. *)

Adam Chlipala's avatar
Adam Chlipala committed
317 318
  Example three_minus_four_zero : exists x, 1 + x = 0.
    Time eauto 1.
319 320 321 322 323 324
(** %\vspace{-.15in}%
<<
Finished transaction in 0. secs (0.u,0.s)
>>
*)

Adam Chlipala's avatar
Adam Chlipala committed
325
    Time eauto 2.
326 327 328 329 330 331
(** %\vspace{-.15in}%
<<
Finished transaction in 0. secs (0.u,0.s)
>>
*)

Adam Chlipala's avatar
Adam Chlipala committed
332
    Time eauto 3.
333 334 335 336 337 338
(** %\vspace{-.15in}%
<<
Finished transaction in 0. secs (0.008u,0.s)
>>
*)

Adam Chlipala's avatar
Adam Chlipala committed
339
    Time eauto 4.
340 341 342 343 344 345
(** %\vspace{-.15in}%
<<
Finished transaction in 0. secs (0.068005u,0.004s)
>>
*)

Adam Chlipala's avatar
Adam Chlipala committed
346
    Time eauto 5.
347 348 349 350 351 352 353 354
(** %\vspace{-.15in}%
<<
Finished transaction in 2. secs (1.92012u,0.044003s)
>>
*)

(** We see worrying exponential growth in running time, and the %\index{tactics!debug}%[debug] tactical helps us see where [eauto] is wasting its time, outputting a trace of every proof step that is attempted.  The rule [eq_trans] applies at every node of a proof tree, and [eauto] tries all such positions. *)

Adam Chlipala's avatar
Adam Chlipala committed
355
    debug eauto 3.
356 357 358 359 360 361 362 363 364 365 366 367 368 369 370 371 372 373 374 375 376 377 378 379
(** [[
1 depth=3 
1.1 depth=2 eapply ex_intro
1.1.1 depth=1 apply plusO
1.1.1.1 depth=0 eapply eq_trans
1.1.2 depth=1 eapply eq_trans
1.1.2.1 depth=1 apply plus_n_O
1.1.2.1.1 depth=0 apply plusO
1.1.2.1.2 depth=0 eapply eq_trans
1.1.2.2 depth=1 apply @eq_refl
1.1.2.2.1 depth=0 apply plusO
1.1.2.2.2 depth=0 eapply eq_trans
1.1.2.3 depth=1 apply eq_add_S ; trivial
1.1.2.3.1 depth=0 apply plusO
1.1.2.3.2 depth=0 eapply eq_trans
1.1.2.4 depth=1 apply eq_sym ; trivial
1.1.2.4.1 depth=0 eapply eq_trans
1.1.2.5 depth=0 apply plusO
1.1.2.6 depth=0 apply plusS
1.1.2.7 depth=0 apply f_equal (A:=nat)
1.1.2.8 depth=0 apply f_equal2 (A1:=nat) (A2:=nat)
1.1.2.9 depth=0 eapply eq_trans
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
380 381 382
  Abort.
End slow.

383 384
(** Sometimes, though, transitivity is just what is needed to get a proof to go through automatically with [eauto].  For those cases, we can use named %\index{hint databases}\emph{%#<i>#hint databases#</i>#%}% to segragate hints into different groups that may be called on as needed.  Here we put [eq_trans] into the database [slow]. *)

Adam Chlipala's avatar
Adam Chlipala committed
385 386 387 388 389 390
(* begin thide *)
Hint Resolve eq_trans : slow.
(* end thide *)

Example three_minus_four_zero : exists x, 1 + x = 0.
(* begin thide *)
391 392 393 394 395 396 397 398
  Time eauto.
(** %\vspace{-.15in}%
<<
Finished transaction in 0. secs (0.004u,0.s)
>>

This [eauto] fails to prove the goal, but it least it takes substantially less than the 2 seconds required above! *)

Adam Chlipala's avatar
Adam Chlipala committed
399 400 401
Abort.
(* end thide *)

402 403
(** One simple example from before runs in the same amount of time, avoiding pollution by transivity. *)

Adam Chlipala's avatar
Adam Chlipala committed
404 405
Example seven_minus_three_again : exists x, x + 3 = 7.
(* begin thide *)
406 407 408 409 410 411 412
  Time eauto 6.
(** %\vspace{-.15in}%
<<
Finished transaction in 0. secs (0.004001u,0.s)
>>
%\vspace{-.2in}% *)

Adam Chlipala's avatar
Adam Chlipala committed
413 414 415
Qed.
(* end thide *)

416 417
(** When we %\emph{%#<i>#do#</i>#%}% need transitivity, we ask for it explicitly. *)

Adam Chlipala's avatar
Adam Chlipala committed
418 419 420 421 422
Example needs_trans : forall x y, 1 + x = y
  -> y = 2
  -> exists z, z + x = 3.
(* begin thide *)
  info eauto with slow.
423 424 425 426 427 428 429 430
(** %\vspace{-.2in}%[[
 == intro x; intro y; intro H; intro H0; simple eapply ex_intro; 
    apply plusS; simple eapply eq_trans.
    exact H.
    
    exact H0.
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
431 432 433
Qed.
(* end thide *)

434 435
(** The [info] trace shows that [eq_trans] was used in just the position where it is needed to complete the proof.  We also see that [auto] and [eauto] always perform [intro] steps without counting them toward the bound on proof tree depth. *)

Adam Chlipala's avatar
Adam Chlipala committed
436 437 438

(** * Searching for Underconstrained Values *)

439 440
(** Recall the definition of the list length function. *)

Adam Chlipala's avatar
Adam Chlipala committed
441
Print length.
442 443 444 445 446 447 448 449 450 451 452
(** %\vspace{-.15in}%[[
length = 
fun A : Type =>
fix length (l : list A) : nat :=
  match l with
  | nil => 0
  | _ :: l' => S (length l')
  end
]]

This function is easy to reason about in the forward direction, computing output from input. *)
Adam Chlipala's avatar
Adam Chlipala committed
453 454 455 456 457 458

Example length_1_2 : length (1 :: 2 :: nil) = 2.
  auto.
Qed.

Print length_1_2.
459 460 461 462 463
(** %\vspace{-.15in}%[[
length_1_2 = eq_refl
]]

As in the last section, we will prove some lemmas to recast [length] in logic programming style, to help us compute inputs from outputs. *)
Adam Chlipala's avatar
Adam Chlipala committed
464 465 466 467 468 469 470 471 472 473 474 475 476 477 478

(* begin thide *)
Theorem length_O : forall A, length (nil (A := A)) = O.
  crush.
Qed.

Theorem length_S : forall A (h : A) t n,
  length t = n
  -> length (h :: t) = S n.
  crush.
Qed.

Hint Resolve length_O length_S.
(* end thide *)

479 480
(** Let us apply these hints to prove that a [list nat] of length 2 exists. *)

Adam Chlipala's avatar
Adam Chlipala committed
481 482 483
Example length_is_2 : exists ls : list nat, length ls = 2.
(* begin thide *)
  eauto.
484 485 486 487 488 489 490 491 492
(** <<
No more subgoals but non-instantiated existential variables:
Existential 1 = ?20249 : [ |- nat]
Existential 2 = ?20252 : [ |- nat]
>>

Coq complains that we finished the proof without determining the values of some unification variables created during proof search.  The error message may seem a bit silly, since %\emph{%#<i>#any#</i>#%}% value of type [nat] (for instance, 0) can be plugged in for either variable!  However, for more complex types, finding their inhabitants may be as complex as theorem-proving in general.

The %\index{Vernacular commands!Show Proof}%[Show Proof] command shows exactly which proof term [eauto] has found, with the undetermined unification variables appearing explicitly where they are used. *)
Adam Chlipala's avatar
Adam Chlipala committed
493 494

  Show Proof.
495 496 497 498 499 500 501 502
(** <<
Proof: ex_intro (fun ls : list nat => length ls = 2)
         (?20249 :: ?20252 :: nil)
         (length_S ?20249 (?20252 :: nil)
            (length_S ?20252 nil (length_O nat)))
>>
%\vspace{-.2in}% *)

Adam Chlipala's avatar
Adam Chlipala committed
503 504 505
Abort.
(* end thide *)

506 507
(** We see that the two unification variables stand for the two elements of the list.  Indeed, list length is independent of data values.  Paradoxically, we can make the proof search process easier by constraining the list further, so that proof search naturally locates appropriate data elements by unification.  The library predicate [Forall] will be helpful. *)

Adam Chlipala's avatar
Adam Chlipala committed
508
Print Forall.
509 510 511 512 513 514 515
(** %\vspace{-.15in}%[[
Inductive Forall (A : Type) (P : A -> Prop) : list A -> Prop :=
    Forall_nil : Forall P nil
  | Forall_cons : forall (x : A) (l : list A),
                  P x -> Forall P l -> Forall P (x :: l)
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
516 517 518 519 520 521 522 523

Example length_is_2 : exists ls : list nat, length ls = 2
  /\ Forall (fun n => n >= 1) ls.
(* begin thide *)
  eauto 9.
Qed.
(* end thide *)

524 525 526 527 528 529 530 531 532 533 534 535 536 537 538 539
(** We can see which list [eauto] found by printing the proof term. *)

Print length_is_2.
(** %\vspace{-.15in}%[[
length_is_2 = 
ex_intro
  (fun ls : list nat => length ls = 2 /\ Forall (fun n : nat => n >= 1) ls)
  (1 :: 1 :: nil)
  (conj (length_S 1 (1 :: nil) (length_S 1 nil (length_O nat)))
     (Forall_cons 1 (le_n 1)
        (Forall_cons 1 (le_n 1) (Forall_nil (fun n : nat => n >= 1)))))
]]
*)

(** Let us try one more, fancier example.  First, we use a standard high-order function to define a function for summing all data elements of a list. *)

Adam Chlipala's avatar
Adam Chlipala committed
540 541
Definition sum := fold_right plus O.

542 543
(** Another basic lemma will be helpful to guide proof search. *)

Adam Chlipala's avatar
Adam Chlipala committed
544 545 546 547 548 549 550 551 552
(* begin thide *)
Lemma plusO' : forall n m,
  n = m
  -> 0 + n = m.
  crush.
Qed.

Hint Resolve plusO'.

553 554
(** Finally, we meet %\index{Vernacular commands!Hint Extern}%[Hint Extern], the command to register a custom hint.  That is, we provide a pattern to match against goals during proof search.  Whenever the pattern matches, a tactic (given to the right of an arrow [=>]) is attempted.  Below, the number [1] gives a priority for this step.  Lower priorities are tried before higher priorities, which can have a significant effect on proof search time. *)

Adam Chlipala's avatar
Adam Chlipala committed
555 556 557
Hint Extern 1 (sum _ = _) => simpl.
(* end thide *)

558 559
(** Now we can find a length-2 list whose sum is 0. *)

Adam Chlipala's avatar
Adam Chlipala committed
560 561 562 563 564 565 566
Example length_and_sum : exists ls : list nat, length ls = 2
  /\ sum ls = O.
(* begin thide *)
  eauto 7.
Qed.
(* end thide *)

567
(* begin hide *)
Adam Chlipala's avatar
Adam Chlipala committed
568
Print length_and_sum.
569 570 571
(* end hide *)

(** Printing the proof term shows the unsurprising list that is found.  Here is an example where it is less obvious which list will be used.  Can you guess which list [eauto] will choose? *)
Adam Chlipala's avatar
Adam Chlipala committed
572 573 574 575 576 577 578 579

Example length_and_sum' : exists ls : list nat, length ls = 5
  /\ sum ls = 42.
(* begin thide *)
  eauto 15.
Qed.
(* end thide *)

580
(* begin hide *)
Adam Chlipala's avatar
Adam Chlipala committed
581
Print length_and_sum'.
582 583 584
(* end hide *)

(** We will give away part of the answer and say that the above list is less interesting than we would like, because it contains too many zeroes.  A further constraint forces a different solution for a smaller instance of the problem. *)
Adam Chlipala's avatar
Adam Chlipala committed
585 586 587 588 589 590 591 592 593

Example length_and_sum'' : exists ls : list nat, length ls = 2
  /\ sum ls = 3
  /\ Forall (fun n => n <> 0) ls.
(* begin thide *)
  eauto 11.
Qed.
(* end thide *)

594
(* begin hide *)
Adam Chlipala's avatar
Adam Chlipala committed
595
Print length_and_sum''.
596 597 598
(* end hide *)

(** We could continue through exercises of this kind, but even more interesting than finding lists automatically is finding %\emph{%#<i>#programs#</i>#%}% automatically. *)
Adam Chlipala's avatar
Adam Chlipala committed
599 600 601 602


(** * Synthesizing Programs *)

603 604
(** Here is a simple syntax type for arithmetic expressions, similar to those we have used several times before in the book.  In this case, we allow expressions to mention exactly one distinguished variable. *)

Adam Chlipala's avatar
Adam Chlipala committed
605 606 607 608 609
Inductive exp : Set :=
| Const : nat -> exp
| Var : exp
| Plus : exp -> exp -> exp.

610 611
(** An inductive relation specifies the semantics of an expression, relating a variable value and an expression to the expression value. *)

Adam Chlipala's avatar
Adam Chlipala committed
612 613 614 615 616 617 618 619 620 621 622
Inductive eval (var : nat) : exp -> nat -> Prop :=
| EvalConst : forall n, eval var (Const n) n
| EvalVar : eval var Var var
| EvalPlus : forall e1 e2 n1 n2, eval var e1 n1
  -> eval var e2 n2
  -> eval var (Plus e1 e2) (n1 + n2).

(* begin thide *)
Hint Constructors eval.
(* end thide *)

623 624
(** We can use [auto] to execute the semantics for specific expressions. *)

Adam Chlipala's avatar
Adam Chlipala committed
625 626 627 628 629 630
Example eval1 : forall var, eval var (Plus Var (Plus (Const 8) Var)) (var + (8 + var)).
(* begin thide *)
  auto.
Qed.
(* end thide *)

631 632
(** Unfortunately, just the constructors of [eval] are not enough to prove theorems like the following, which depends on an arithmetic identity. *)

Adam Chlipala's avatar
Adam Chlipala committed
633 634 635 636 637 638
Example eval1' : forall var, eval var (Plus Var (Plus (Const 8) Var)) (2 * var + 8).
(* begin thide *)
  eauto.
Abort.
(* end thide *)

639 640
(** To help prove [eval1'], we prove an alternate version of [EvalPlus] that inserts an extra equality premise. *)

Adam Chlipala's avatar
Adam Chlipala committed
641 642 643 644 645 646 647 648 649 650
(* begin thide *)
Theorem EvalPlus' : forall var e1 e2 n1 n2 n, eval var e1 n1
  -> eval var e2 n2
  -> n1 + n2 = n
  -> eval var (Plus e1 e2) n.
  crush.
Qed.

Hint Resolve EvalPlus'.

651 652
(** Further, we instruct [eauto] to apply %\index{tactics!omega}%[omega], a standard tactic that provides a complete decision procedure for quantifier-free linear arithmetic.  Via [Hint Extern], we ask for use of [omega] on any equality goal.  The [abstract] tactical generates a new lemma for every such successful proof, so that, in the final proof term, the lemma may be referenced in place of dropping in the full proof of the arithmetic equality. *)

Adam Chlipala's avatar
Adam Chlipala committed
653 654 655
Hint Extern 1 (_ = _) => abstract omega.
(* end thide *)

656 657
(** Now we can return to [eval1'] and prove it automatically. *)

Adam Chlipala's avatar
Adam Chlipala committed
658 659 660 661 662 663 664
Example eval1' : forall var, eval var (Plus Var (Plus (Const 8) Var)) (2 * var + 8).
(* begin thide *)
  eauto.
Qed.
(* end thide *)

Print eval1'.
665 666 667 668 669 670 671 672 673 674 675 676 677
(** %\vspace{-.15in}%[[
eval1' = 
fun var : nat =>
EvalPlus' (EvalVar var) (EvalPlus (EvalConst var 8) (EvalVar var))
  (eval1'_subproof var)
     : forall var : nat,
       eval var (Plus Var (Plus (Const 8) Var)) (2 * var + 8)
]]
*)

(** The lemma [eval1'_subproof] was generated by [abstract omega].

   Now we are ready to take advantage of logic programming's flexibility by searching for a program (arithmetic expression) that always evaluates to a particular symbolic value. *)
Adam Chlipala's avatar
Adam Chlipala committed
678 679 680 681 682 683 684 685

Example synthesize1 : exists e, forall var, eval var e (var + 7).
(* begin thide *)
  eauto.
Qed.
(* end thide *)

Print synthesize1.
686 687 688 689 690 691 692 693 694
(** %\vspace{-.15in}%[[
synthesize1 = 
ex_intro (fun e : exp => forall var : nat, eval var e (var + 7))
  (Plus Var (Const 7))
  (fun var : nat => EvalPlus (EvalVar var) (EvalConst var 7))
]]
*)

(** Here are two more examples showing off our program synthesis abilities. *)
Adam Chlipala's avatar
Adam Chlipala committed
695 696 697 698 699 700 701

Example synthesize2 : exists e, forall var, eval var e (2 * var + 8).
(* begin thide *)
  eauto.
Qed.
(* end thide *)

702
(* begin hide *)
Adam Chlipala's avatar
Adam Chlipala committed
703
Print synthesize2.
704
(* end hide *)
Adam Chlipala's avatar
Adam Chlipala committed
705 706 707 708 709 710 711

Example synthesize3 : exists e, forall var, eval var e (3 * var + 42).
(* begin thide *)
  eauto.
Qed.
(* end thide *)

712
(* begin hide *)
Adam Chlipala's avatar
Adam Chlipala committed
713
Print synthesize3.
714 715 716 717 718
(* end hide *)

(** These examples show linear expressions over the variable [var].  Any such expression is equivalent to [k * var + n] for some [k] and [n].  It is probably not so surprising that we can prove that any expression's semantics is equivalent to some such linear expression, but it is tedious to prove such a fact manually.  To finish this section, we will use [eauto] to complete the proof, finding [k] and [n] values automatically.

   We prove a series of lemmas and add them as hints.  We have alternate [eval] constructor lemmas and some facts about arithmetic. *)
Adam Chlipala's avatar
Adam Chlipala committed
719 720 721 722 723 724 725 726 727 728 729 730 731 732 733 734 735 736 737 738 739 740 741 742 743 744 745 746 747 748 749 750 751 752 753 754 755

(* begin thide *)
Theorem EvalConst' : forall var n m, n = m
  -> eval var (Const n) m.
  crush.
Qed.

Hint Resolve EvalConst'.

Theorem zero_times : forall n m r,
  r = m
  -> r = 0 * n + m.
  crush.
Qed.

Hint Resolve zero_times.

Theorem EvalVar' : forall var n,
  var = n
  -> eval var Var n.
  crush.
Qed.

Hint Resolve EvalVar'.

Theorem plus_0 : forall n r,
  r = n
  -> r = n + 0.
  crush.
Qed.

Theorem times_1 : forall n, n = 1 * n.
  crush.
Qed.

Hint Resolve plus_0 times_1.

756 757
(** We finish with one more arithmetic lemma that is particularly specialized to this theorem.  This fact happens to follow by the axioms of the %\emph{%#<i>#ring#</i>#%}% algebraic structure, so, since the naturals form a ring, we can use the built-in tactic %\index{tactics!ring}%[ring]. *)

Adam Chlipala's avatar
Adam Chlipala committed
758 759 760 761 762 763 764 765 766
Require Import Arith Ring.

Theorem combine : forall x k1 k2 n1 n2,
  (k1 * x + n1) + (k2 * x + n2) = (k1 + k2) * x + (n1 + n2).
  intros; ring.
Qed.

Hint Resolve combine.

767 768
(** Our choice of hints is cheating, to an extent, by telegraphing the procedure for choosing values of [k] and [n].  Nonetheless, with these lemmas in place, we achieve an automated proof without explicitly orchestrating the lemmas' composition. *)

Adam Chlipala's avatar
Adam Chlipala committed
769 770 771 772 773
Theorem linear : forall e, exists k, exists n,
  forall var, eval var e (k * var + n).
  induction e; crush; eauto.
Qed.

774
(* begin hide *)
Adam Chlipala's avatar
Adam Chlipala committed
775
Print linear.
776
(* end hide *)
Adam Chlipala's avatar
Adam Chlipala committed
777 778
(* end thide *)

779 780
(** By printing the proof term, it is possible to see the procedure that is used to choose the constants for each input term. *)

Adam Chlipala's avatar
Adam Chlipala committed
781 782 783

(** * More on [auto] Hints *)

784
(** Let us stop at this point and take stock of the possibilities for [auto] and [eauto] hints.  Hints are contained within %\textit{%#<i>#hint databases#</i>#%}%, which we have seen extended in many examples so far.  When no hint database is specified, a default database is used.  Hints in the default database are always used by [auto] or [eauto].  The chance to extend hint databases imperatively is important, because, in Ltac programming, we cannot create %``%#"#global variables#"#%''% whose values can be extended seamlessly by different modules in different source files.  We have seen the advantages of hints so far, where [crush] can be defined once and for all, while still automatically applying the hints we add throughout developments.  In fact, [crush] is defined in terms of [auto], which explains how we achieve this extensibility.  Other user-defined tactics can take similar advantage of [auto] and [eauto].
Adam Chlipala's avatar
Adam Chlipala committed
785

786
The basic hints for [auto] and [eauto] are %\index{Vernacular commands!Hint Immediate}%[Hint Immediate lemma], asking to try solving a goal immediately by applying a lemma and discharging any hypotheses with a single proof step each; %\index{Vernacular commands!Hint Resolve}%[Resolve lemma], which does the same but may add new premises that are themselves to be subjects of nested proof search; %\index{Vernacular commands!Hint Constructors}%[Constructors type], which acts like [Resolve] applied to every constructor of an inductive type; and %\index{Vernacular commands!Hint Unfold}%[Unfold ident], which tries unfolding [ident] when it appears at the head of a proof goal.  Each of these [Hint] commands may be used with a suffix, as in [Hint Resolve lemma : my_db].  This adds the hint only to the specified database, so that it would only be used by, for instance, [auto with my_db].  An additional argument to [auto] specifies the maximum depth of proof trees to search in depth-first order, as in [auto 8] or [auto 8 with my_db].  The default depth is 5.
Adam Chlipala's avatar
Adam Chlipala committed
787

788
All of these [Hint] commands can be issued alternatively with a more primitive hint kind, [Extern].  A few more examples of [Hint Extern] should illustrate more of the possibilities. *)
Adam Chlipala's avatar
Adam Chlipala committed
789 790 791 792 793 794 795 796 797

Theorem bool_neq : true <> false.
(* begin thide *)
  auto.

  (** [crush] would have discharged this goal, but the default hint database for [auto] contains no hint that applies. *)

Abort.

798
(** It is hard to come up with a [bool]-specific hint that is not just a restatement of the theorem we mean to prove.  Luckily, a simpler form suffices, by appealing to the built-in tactic %\index{tactics!congruence}%[congruence], a complete procedure for the theory of equality, uninterpreted functions, and datatype constructors. *)
Adam Chlipala's avatar
Adam Chlipala committed
799 800 801 802 803 804 805 806

Hint Extern 1 (_ <> _) => congruence.

Theorem bool_neq : true <> false.
  auto.
Qed.
(* end thide *)

807
(** [Extern] hints may be implemented with the full Ltac language.  This example shows a case where a hint uses a [match]. *)
Adam Chlipala's avatar
Adam Chlipala committed
808 809 810 811 812 813 814 815 816 817 818

Section forall_and.
  Variable A : Set.
  Variables P Q : A -> Prop.

  Hypothesis both : forall x, P x /\ Q x.

  Theorem forall_and : forall z, P z.
(* begin thide *)
    crush.

819
    (** The [crush] invocation makes no progress beyond what [intros] would have accomplished.  An [auto] invocation will not apply the hypothesis [both] to prove the goal, because the conclusion of [both] does not unify with the conclusion of the goal.  However, we can teach [auto] to handle this kind of goal. *)
Adam Chlipala's avatar
Adam Chlipala committed
820 821 822 823 824 825 826 827 828 829

    Hint Extern 1 (P ?X) =>
      match goal with
        | [ H : forall x, P x /\ _ |- _ ] => apply (proj1 (H X))
      end.

    auto.
  Qed.
(* end thide *)

830
  (** We see that an [Extern] pattern may bind unification variables that we use in the associated tactic.  The function [proj1] is from the standard library, for extracting a proof of [R] from a proof of [R /\ S]. *)
Adam Chlipala's avatar
Adam Chlipala committed
831 832 833 834 835 836 837 838 839

End forall_and.

(** After our success on this example, we might get more ambitious and seek to generalize the hint to all possible predicates [P].
   [[
  Hint Extern 1 (?P ?X) =>
    match goal with
      | [ H : forall x, P x /\ _ |- _ ] => apply (proj1 (H X))
    end.
840 841
]]
<<
Adam Chlipala's avatar
Adam Chlipala committed
842
User error: Bound head variable
843 844 845
>>

   Coq's [auto] hint databases work as tables mapping %\index{head symbol}\textit{%#<i>#head symbols#</i>#%}% to lists of tactics to try.  Because of this, the constant head of an [Extern] pattern must be determinable statically.  In our first [Extern] hint, the head symbol was [not], since [x <> y] desugars to [not (eq x y)]; and, in the second example, the head symbol was [P].
Adam Chlipala's avatar
Adam Chlipala committed
846

847
   Fortunately, a more basic form of [Hint Extern] also applies.  We may simply leave out the pattern to the left of the [=>], incorporating the corresponding logic into the Ltac script. *)
Adam Chlipala's avatar
Adam Chlipala committed
848

849 850 851 852 853 854
Hint Extern 1 =>
  match goal with
    | [ H : forall x, ?P x /\ _ |- ?P ?X ] => apply (proj1 (H X))
  end.

(** Be forewarned that a [Hint Extern] of this kind will be applied at %\emph{%#<i>#every#</i>#%}% node of a proof tree, so an expensive Ltac script may slow proof search significantly. *)
Adam Chlipala's avatar
Adam Chlipala committed
855 856 857 858


(** * Rewrite Hints *)

859
(** Another dimension of extensibility with hints is rewriting with quantified equalities.  We have used the associated command %\index{Vernacular commands!Hint Rewrite}%[Hint Rewrite] in many examples so far.  The [crush] tactic uses these hints by calling the built-in tactic %\index{tactics!autorewrite}%[autorewrite].  Our rewrite hints have taken the form [Hint Rewrite lemma], which by default adds them to the default hint database [core]; but alternate hint databases may also be specified just as with, e.g., [Hint Resolve].
Adam Chlipala's avatar
Adam Chlipala committed
860

861
   The next example shows a direct use of [autorewrite].  Note that, while [Hint Rewrite] uses a default database, [autorewrite] requires that a database be named. *)
Adam Chlipala's avatar
Adam Chlipala committed
862 863 864 865 866 867 868

Section autorewrite.
  Variable A : Set.
  Variable f : A -> A.

  Hypothesis f_f : forall x, f (f x) = f x.

869
  Hint Rewrite f_f.
Adam Chlipala's avatar
Adam Chlipala committed
870 871

  Lemma f_f_f : forall x, f (f (f x)) = f x.
872
    intros; autorewrite with core; reflexivity.
Adam Chlipala's avatar
Adam Chlipala committed
873 874 875 876 877 878 879
  Qed.

  (** There are a few ways in which [autorewrite] can lead to trouble when insufficient care is taken in choosing hints.  First, the set of hints may define a nonterminating rewrite system, in which case invocations to [autorewrite] may not terminate.  Second, we may add hints that %``%#"#lead [autorewrite] down the wrong path.#"#%''%  For instance: *)

  Section garden_path.
    Variable g : A -> A.
    Hypothesis f_g : forall x, f x = g x.
880
    Hint Rewrite f_g.
Adam Chlipala's avatar
Adam Chlipala committed
881 882

    Lemma f_f_f' : forall x, f (f (f x)) = f x.
883
      intros; autorewrite with core.
Adam Chlipala's avatar
Adam Chlipala committed
884 885 886 887 888 889 890 891 892 893 894 895
      (** [[
============================
 g (g (g x)) = g x
          ]]
          *)

    Abort.

    (** Our new hint was used to rewrite the goal into a form where the old hint could no longer be applied.  This %``%#"#non-monotonicity#"#%''% of rewrite hints contrasts with the situation for [auto], where new hints may slow down proof search but can never %``%#"#break#"#%''% old proofs.  The key difference is that [auto] either solves a goal or makes no changes to it, while [autorewrite] may change goals without solving them.  The situation for [eauto] is slightly more complicated, as changes to hint databases may change the proof found for a particular goal, and that proof may influence the settings of unification variables that appear elsewhere in the proof state. *)

  Reset garden_path.

896
  (** The [autorewrite] tactic also works with quantified equalities that include additional premises, but we must be careful to avoid similar incorrect rewritings. *)
Adam Chlipala's avatar
Adam Chlipala committed
897 898 899 900 901

  Section garden_path.
    Variable P : A -> Prop.
    Variable g : A -> A.
    Hypothesis f_g : forall x, P x -> f x = g x.
902
    Hint Rewrite f_g.
Adam Chlipala's avatar
Adam Chlipala committed
903 904

    Lemma f_f_f' : forall x, f (f (f x)) = f x.
905
      intros; autorewrite with core.
Adam Chlipala's avatar
Adam Chlipala committed
906 907 908 909 910 911 912 913 914 915 916 917 918 919 920 921 922 923 924 925 926 927 928 929 930 931 932
      (** [[

  ============================
   g (g (g x)) = g x

subgoal 2 is:
 P x
subgoal 3 is:
 P (f x)
subgoal 4 is:
 P (f x)
          ]]
          *)

    Abort.

    (** The inappropriate rule fired the same three times as before, even though we know we will not be able to prove the premises. *)

  Reset garden_path.

  (** Our final, successful, attempt uses an extra argument to [Hint Rewrite] that specifies a tactic to apply to generated premises.  Such a hint is only used when the tactic succeeds for all premises, possibly leaving further subgoals for some premises. *)

  Section garden_path.
    Variable P : A -> Prop.
    Variable g : A -> A.
    Hypothesis f_g : forall x, P x -> f x = g x.
(* begin thide *)
933
    Hint Rewrite f_g using assumption.
Adam Chlipala's avatar
Adam Chlipala committed
934 935 936 937
(* end thide *)

    Lemma f_f_f' : forall x, f (f (f x)) = f x.
(* begin thide *)
938
      intros; autorewrite with core; reflexivity.
Adam Chlipala's avatar
Adam Chlipala committed
939 940 941
    Qed.
(* end thide *)

942
    (** We may still use [autorewrite] to apply [f_g] when the generated premise is among our assumptions. *)
Adam Chlipala's avatar
Adam Chlipala committed
943 944 945

    Lemma f_f_f_g : forall x, P x -> f (f x) = g x.
(* begin thide *)
946
      intros; autorewrite with core; reflexivity.
Adam Chlipala's avatar
Adam Chlipala committed
947 948 949 950 951 952
(* end thide *)
    Qed.
  End garden_path.

  (** remove printing * *)

953
  (** It can also be useful to apply the [autorewrite with db in *] form, which does rewriting in hypotheses, as well as in the conclusion. *)
Adam Chlipala's avatar
Adam Chlipala committed
954 955 956 957 958 959

  (** printing * $*$ *)

  Lemma in_star : forall x y, f (f (f (f x))) = f (f y)
    -> f x = f (f (f y)).
(* begin thide *)
960
    intros; autorewrite with core in *; assumption.
Adam Chlipala's avatar
Adam Chlipala committed
961 962 963 964
(* end thide *)
  Qed.

End autorewrite.
965 966

(** Many proofs can be automated in pleasantly modular ways with deft combination of [auto] and [autorewrite]. *)