comparison src/Predicates.v @ 50:a21eb02cc7c6

Recursive predicates
author Adam Chlipala <adamc@hcoop.net>
date Sun, 28 Sep 2008 11:57:15 -0400
parents 827d7e8a7d9e
children 9ceee967b1fc ce0c9d481ee3
comparison
equal deleted inserted replaced
49:827d7e8a7d9e 50:a21eb02cc7c6
438 438
439 (* EX: Define an inductive type capturing when a list has exactly two elements. Prove that your predicate does not hold of the empty list, and prove that, whenever it holds of a list, the length of that list is two. *) 439 (* EX: Define an inductive type capturing when a list has exactly two elements. Prove that your predicate does not hold of the empty list, and prove that, whenever it holds of a list, the length of that list is two. *)
440 440
441 (* end hide *) 441 (* end hide *)
442 442
443
444 (** * Recursive Predicates *)
445
446 (** We have already seen all of the ingredients we need to build interesting recursive predicates, like this predicate capturing even-ness. *)
447
448 Inductive even : nat -> Prop :=
449 | EvenO : even O
450 | EvenSS : forall n, even n -> even (S (S n)).
451
452 (** Think of [even] as another judgment defined by natural deduction rules. [EvenO] is a rule with nothing above the line and [even O] below the line, and [EvenSS] is a rule with [even n] above the line and [even (S (S n))] below.
453
454 The proof techniques of the last section are easily adapted. *)
455
456 Theorem even_0 : even 0.
457 constructor.
458 Qed.
459
460 Theorem even_4 : even 4.
461 constructor; constructor; constructor.
462 Qed.
463
464 (** It is not hard to see that sequences of constructor applications like the above can get tedious. We can avoid them using Coq's hint facility. *)
465
466 Hint Constructors even.
467
468 Theorem even_4' : even 4.
469 auto.
470 Qed.
471
472 Theorem even_1_contra : even 1 -> False.
473 inversion 1.
474 Qed.
475
476 Theorem even_3_contra : even 3 -> False.
477 inversion 1.
478 (** [[
479
480 H : even 3
481 n : nat
482 H1 : even 1
483 H0 : n = 1
484 ============================
485 False
486 ]] *)
487
488 (** [inversion] can be a little overzealous at times, as we can see here with the introduction of the unused variable [n] and an equality hypothesis about it. For more complicated predicates, though, adding such assumptions is critical to dealing with the undecidability of general inversion. *)
489
490 inversion H1.
491 Qed.
492
493 (** We can also do inductive proofs about [even]. *)
494
495 Theorem even_plus : forall n m, even n -> even m -> even (n + m).
496 (** It seems a reasonable first choice to proceed by induction on [n]. *)
497 induction n; crush.
498 (** [[
499
500 n : nat
501 IHn : forall m : nat, even n -> even m -> even (n + m)
502 m : nat
503 H : even (S n)
504 H0 : even m
505 ============================
506 even (S (n + m))
507 ]] *)
508
509 (** We will need to use the hypotheses [H] and [H0] somehow. The most natural choice is to invert [H]. *)
510
511 inversion H.
512 (** [[
513
514 n : nat
515 IHn : forall m : nat, even n -> even m -> even (n + m)
516 m : nat
517 H : even (S n)
518 H0 : even m
519 n0 : nat
520 H2 : even n0
521 H1 : S n0 = n
522 ============================
523 even (S (S n0 + m))
524 ]] *)
525
526 (** Simplifying the conclusion brings us to a point where we can apply a constructor. *)
527 simpl.
528 (** [[
529
530 ============================
531 even (S (S (n0 + m)))
532 ]] *)
533
534 constructor.
535 (** [[
536
537 ============================
538 even (n0 + m)
539 ]] *)
540
541 (** At this point, we would like to apply the inductive hypothesis, which is: *)
542 (** [[
543
544 IHn : forall m : nat, even n -> even m -> even (n + m)
545 ]] *)
546
547 (** Unfortunately, the goal mentions [n0] where it would need to mention [n] to match [IHn]. We could keep looking for a way to finish this proof from here, but it turns out that we can make our lives much easier by changing our basic strategy. Instead of inducting on the structure of [n], we should induct %\textit{%#<i>#on the structure of one of the [even] proofs#</i>#%}%. This technique is commonly called %\textit{%#<i>#rule induction#</i>#%}% in programming language semantics. In the setting of Coq, we have already seen how predicates are defined using the same inductive type mechanism as datatypes, so the fundamental unity of rule induction with "normal" induction is apparent. *)
548
549 Restart.
550
551 induction 1.
552 (** [[
553
554 m : nat
555 ============================
556 even m -> even (0 + m)
557
558 subgoal 2 is:
559 even m -> even (S (S n) + m)
560 ]] *)
561
562 (** The first case is easily discharged by [crush], based on the hint we added earlier to try the constructors of [even]. *)
563
564 crush.
565
566 (** Now we focus on the second case: *)
567 intro.
568
569 (** [[
570
571 m : nat
572 n : nat
573 H : even n
574 IHeven : even m -> even (n + m)
575 H0 : even m
576 ============================
577 even (S (S n) + m)
578 ]] *)
579
580 (** We simplify and apply a constructor, as in our last proof attempt. *)
581
582 simpl; constructor.
583 (** [[
584
585 ============================
586 even (n + m)
587 ]] *)
588
589 (** Now we have an exact match with our inductive hypothesis, and the remainder of the proof is trivial. *)
590
591 apply IHeven; assumption.
592
593 (** In fact, [crush] can handle all of the details of the proof once we declare the induction strategy. *)
594
595 Restart.
596 induction 1; crush.
597 Qed.
598
599 (** Induction on recursive predicates has similar pitfalls to those we encountered with inversion in the last section. *)
600
601 Theorem even_contra : forall n, even (S (n + n)) -> False.
602 induction 1.
603 (** [[
604
605 n : nat
606 ============================
607 False
608
609 subgoal 2 is:
610 False
611 ]] *)
612
613 (** We are already sunk trying to prove the first subgoal, since the argument to [even] was replaced by a fresh variable internally. This time, we find it easiest to prove this theorem by way of a lemma. Instead of trusting [induction] to replace expressions with fresh variables, we do it ourselves, explicitly adding the appropriate equalities as new assumptions. *)
614 Abort.
615
616 Lemma even_contra' : forall n', even n' -> forall n, n' = S (n + n) -> False.
617 induction 1; crush.
618
619 (** At this point, it is useful to consider all cases of [n] and [n0] being zero or nonzero. Only one of these cases has any trickyness to it. *)
620 destruct n; destruct n0; crush.
621
622 (** [[
623
624 n : nat
625 H : even (S n)
626 IHeven : forall n0 : nat, S n = S (n0 + n0) -> False
627 n0 : nat
628 H0 : S n = n0 + S n0
629 ============================
630 False
631 ]] *)
632
633 (** At this point it is useful to use a theorem from the standard library, which we also proved with a different name in the last chapter. *)
634 Check plus_n_Sm.
635 (** [[
636
637 plus_n_Sm
638 : forall n m : nat, S (n + m) = n + S m
639 ]] *)
640
641 rewrite <- plus_n_Sm in H0.
642
643 (** The induction hypothesis lets us complete the proof. *)
644 apply IHeven with n0; assumption.
645
646 (** As usual, we can rewrite the proof to avoid referencing any locally-generated names, which makes our proof script more robust to changes in the theorem statement. *)
647 Restart.
648 Hint Rewrite <- plus_n_Sm : cpdt.
649
650 induction 1; crush;
651 match goal with
652 | [ H : S ?N = ?N0 + ?N0 |- _ ] => destruct N; destruct N0
653 end; crush; eauto.
654 Qed.
655
656 (** We write the proof in a way that avoids the use of local variable or hypothesis names, using the [match] tactic form to do pattern-matching on the goal. We use unification variables prefixed by question marks in the pattern, and we take advantage of the possibility to mention a unification variable twice in one pattern, to enforce equality between occurrences. The hint to rewrite with [plus_n_Sm] in a particular direction saves us from having to figure out the right place to apply that theorem, and we also take critical advantage of a new tactic, [eauto].
657
658 [crush] uses the tactic [intuition], which, when it runs out of tricks to try using only propositional logic, by default tries the tactic [auto], which we saw in an earlier example. [auto] attempts Prolog-style logic programming, searching through all proof trees up to a certain depth that are built only out of hints that have been registered with [Hint] commands. Compared to Prolog, [auto] places an important restriction: it never introduces new unification variables during search. That is, every time a rule is applied during proof search, all of its arguments must be deducible by studying the form of the goal. [eauto] relaxes this restriction, at the cost of possibly exponentially greater running time. In this particular case, we know that [eauto] has only a small space of proofs to search, so it makes sense to run it. It is common in effectively-automated Coq proofs to see a bag of standard tactics applied to pick off the "easy" subgoals, finishing off with [eauto] to handle the tricky parts that can benefit from ad-hoc exhaustive search.
659
660 The original theorem now follows trivially from our lemma. *)
661
662 Theorem even_contra : forall n, even (S (n + n)) -> False.
663 intros; apply even_contra' with (S (n + n)) n; trivial.
664 Qed.