comparison src/LogicProg.v @ 373:b13f76abc724

Prose for second LogicProg section
author Adam Chlipala <adam@chlipala.net>
date Mon, 26 Mar 2012 16:04:52 -0400
parents 3c039c72eb40
children f3146d40c2a1
comparison
equal deleted inserted replaced
372:3c039c72eb40 373:b13f76abc724
434 (** 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. *) 434 (** 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. *)
435 435
436 436
437 (** * Searching for Underconstrained Values *) 437 (** * Searching for Underconstrained Values *)
438 438
439 (** Recall the definition of the list length function. *)
440
439 Print length. 441 Print length.
442 (** %\vspace{-.15in}%[[
443 length =
444 fun A : Type =>
445 fix length (l : list A) : nat :=
446 match l with
447 | nil => 0
448 | _ :: l' => S (length l')
449 end
450 ]]
451
452 This function is easy to reason about in the forward direction, computing output from input. *)
440 453
441 Example length_1_2 : length (1 :: 2 :: nil) = 2. 454 Example length_1_2 : length (1 :: 2 :: nil) = 2.
442 auto. 455 auto.
443 Qed. 456 Qed.
444 457
445 Print length_1_2. 458 Print length_1_2.
459 (** %\vspace{-.15in}%[[
460 length_1_2 = eq_refl
461 ]]
462
463 As in the last section, we will prove some lemmas to recast [length] in logic programming style, to help us compute inputs from outputs. *)
446 464
447 (* begin thide *) 465 (* begin thide *)
448 Theorem length_O : forall A, length (nil (A := A)) = O. 466 Theorem length_O : forall A, length (nil (A := A)) = O.
449 crush. 467 crush.
450 Qed. 468 Qed.
456 Qed. 474 Qed.
457 475
458 Hint Resolve length_O length_S. 476 Hint Resolve length_O length_S.
459 (* end thide *) 477 (* end thide *)
460 478
479 (** Let us apply these hints to prove that a [list nat] of length 2 exists. *)
480
461 Example length_is_2 : exists ls : list nat, length ls = 2. 481 Example length_is_2 : exists ls : list nat, length ls = 2.
462 (* begin thide *) 482 (* begin thide *)
463 eauto. 483 eauto.
484 (** <<
485 No more subgoals but non-instantiated existential variables:
486 Existential 1 = ?20249 : [ |- nat]
487 Existential 2 = ?20252 : [ |- nat]
488 >>
489
490 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.
491
492 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. *)
464 493
465 Show Proof. 494 Show Proof.
495 (** <<
496 Proof: ex_intro (fun ls : list nat => length ls = 2)
497 (?20249 :: ?20252 :: nil)
498 (length_S ?20249 (?20252 :: nil)
499 (length_S ?20252 nil (length_O nat)))
500 >>
501 %\vspace{-.2in}% *)
502
466 Abort. 503 Abort.
467 (* end thide *) 504 (* end thide *)
468 505
506 (** 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. *)
507
469 Print Forall. 508 Print Forall.
509 (** %\vspace{-.15in}%[[
510 Inductive Forall (A : Type) (P : A -> Prop) : list A -> Prop :=
511 Forall_nil : Forall P nil
512 | Forall_cons : forall (x : A) (l : list A),
513 P x -> Forall P l -> Forall P (x :: l)
514 ]]
515 *)
470 516
471 Example length_is_2 : exists ls : list nat, length ls = 2 517 Example length_is_2 : exists ls : list nat, length ls = 2
472 /\ Forall (fun n => n >= 1) ls. 518 /\ Forall (fun n => n >= 1) ls.
473 (* begin thide *) 519 (* begin thide *)
474 eauto 9. 520 eauto 9.
475 Qed. 521 Qed.
476 (* end thide *) 522 (* end thide *)
477 523
524 (** We can see which list [eauto] found by printing the proof term. *)
525
526 Print length_is_2.
527 (** %\vspace{-.15in}%[[
528 length_is_2 =
529 ex_intro
530 (fun ls : list nat => length ls = 2 /\ Forall (fun n : nat => n >= 1) ls)
531 (1 :: 1 :: nil)
532 (conj (length_S 1 (1 :: nil) (length_S 1 nil (length_O nat)))
533 (Forall_cons 1 (le_n 1)
534 (Forall_cons 1 (le_n 1) (Forall_nil (fun n : nat => n >= 1)))))
535 ]]
536 *)
537
538 (** 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. *)
539
478 Definition sum := fold_right plus O. 540 Definition sum := fold_right plus O.
541
542 (** Another basic lemma will be helpful to guide proof search. *)
479 543
480 (* begin thide *) 544 (* begin thide *)
481 Lemma plusO' : forall n m, 545 Lemma plusO' : forall n m,
482 n = m 546 n = m
483 -> 0 + n = m. 547 -> 0 + n = m.
484 crush. 548 crush.
485 Qed. 549 Qed.
486 550
487 Hint Resolve plusO'. 551 Hint Resolve plusO'.
488 552
553 (** 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. *)
554
489 Hint Extern 1 (sum _ = _) => simpl. 555 Hint Extern 1 (sum _ = _) => simpl.
490 (* end thide *) 556 (* end thide *)
557
558 (** Now we can find a length-2 list whose sum is 0. *)
491 559
492 Example length_and_sum : exists ls : list nat, length ls = 2 560 Example length_and_sum : exists ls : list nat, length ls = 2
493 /\ sum ls = O. 561 /\ sum ls = O.
494 (* begin thide *) 562 (* begin thide *)
495 eauto 7. 563 eauto 7.
496 Qed. 564 Qed.
497 (* end thide *) 565 (* end thide *)
498 566
567 (* begin hide *)
499 Print length_and_sum. 568 Print length_and_sum.
569 (* end hide *)
570
571 (** 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? *)
500 572
501 Example length_and_sum' : exists ls : list nat, length ls = 5 573 Example length_and_sum' : exists ls : list nat, length ls = 5
502 /\ sum ls = 42. 574 /\ sum ls = 42.
503 (* begin thide *) 575 (* begin thide *)
504 eauto 15. 576 eauto 15.
505 Qed. 577 Qed.
506 (* end thide *) 578 (* end thide *)
507 579
580 (* begin hide *)
508 Print length_and_sum'. 581 Print length_and_sum'.
582 (* end hide *)
583
584 (** 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. *)
509 585
510 Example length_and_sum'' : exists ls : list nat, length ls = 2 586 Example length_and_sum'' : exists ls : list nat, length ls = 2
511 /\ sum ls = 3 587 /\ sum ls = 3
512 /\ Forall (fun n => n <> 0) ls. 588 /\ Forall (fun n => n <> 0) ls.
513 (* begin thide *) 589 (* begin thide *)
514 eauto 11. 590 eauto 11.
515 Qed. 591 Qed.
516 (* end thide *) 592 (* end thide *)
517 593
594 (* begin hide *)
518 Print length_and_sum''. 595 Print length_and_sum''.
596 (* end hide *)
597
598 (** We could continue through exercises of this kind, but even more interesting than finding lists automatically is finding %\emph{%#<i>#programs#</i>#%}% automatically. *)
519 599
520 600
521 (** * Synthesizing Programs *) 601 (** * Synthesizing Programs *)
522 602
523 Inductive exp : Set := 603 Inductive exp : Set :=