comparison src/Match.v @ 140:8cce3247496b

Prose for matcher
author Adam Chlipala <adamc@hcoop.net>
date Mon, 27 Oct 2008 10:09:10 -0400
parents a9e90bacbd16
children ce4cc7fa9b2b
comparison
equal deleted inserted replaced
139:a9e90bacbd16 140:8cce3247496b
583 Theorem test_inster2 : forall x y, x <> y -> P x -> Q (f y) -> Q (f x). 583 Theorem test_inster2 : forall x y, x <> y -> P x -> Q (f y) -> Q (f x).
584 intros; inster 3. 584 intros; inster 3.
585 Qed. 585 Qed.
586 End test_inster. 586 End test_inster.
587 587
588 (** The style employed in the definition of [inster] can seem very counterintuitive to functional programmers. Usually, functional programs accumulate state changes in explicit arguments to recursive functions. In Ltac, the state of the current subgoal is always implicit. Nonetheless, in contrast to general imperative programming, it is easy to undo any changes to this state, and indeed such "undoing" happens automatically at failures within [match]es. In this way, Ltac programming is similar to programming in Haskell with a stateful failure monad that supports a composition operator along the lines of the [first] tactical.
589
590 Functional programming purists may react indignantly to the suggestion of programming this way. Nonetheless, as with other kinds of "monadic programming," many problems are much simpler to solve with Ltac than they would be with explicit, pure proof manipulation in ML or Haskell. To demonstrate, we will write a basic simplification procedure for logical implications.
591
592 This procedure is inspired by one for separation logic, where conjuncts in formulas are thought of as "resources," such that we lose no completeness by "crossing out" equal conjuncts on the two sides of an implication. This process is complicated by the fact that, for reasons of modularity, our formulas can have arbitrary nested tree structure (branching at conjunctions) and may include existential quantifiers. It is helpful for the matching process to "go under" quantifiers and in fact decide how to instantiate existential quantifiers in the conclusion.
593
594 To distinguish the implications that our tactic handles from the implications that will show up as "plumbing" in various lemmas, we define a wrapper definition, a notation, and a tactic. *)
588 595
589 Definition imp (P1 P2 : Prop) := P1 -> P2. 596 Definition imp (P1 P2 : Prop) := P1 -> P2.
590
591 Infix "-->" := imp (no associativity, at level 95). 597 Infix "-->" := imp (no associativity, at level 95).
592
593 Ltac imp := unfold imp; firstorder. 598 Ltac imp := unfold imp; firstorder.
599
600 (** These lemmas about [imp] will be useful in the tactic that we will write. *)
594 601
595 Theorem and_True_prem : forall P Q, 602 Theorem and_True_prem : forall P Q,
596 (P /\ True --> Q) 603 (P /\ True --> Q)
597 -> (P --> Q). 604 -> (P --> Q).
598 imp. 605 imp.
637 Theorem comm_conc : forall P Q R, 644 Theorem comm_conc : forall P Q R,
638 (R --> P /\ Q) 645 (R --> P /\ Q)
639 -> (R --> Q /\ P). 646 -> (R --> Q /\ P).
640 imp. 647 imp.
641 Qed. 648 Qed.
649
650 (** The first order of business in crafting our [matcher] tactic will be auxiliary support for searching through formula trees. The [search_prem] tactic implements running its tactic argument [tac] on every subformula of an [imp] premise. As it traverses a tree, [search_prem] applies some of the above lemmas to rewrite the goal to bring different subformulas to the head of the goal. That is, for every subformula [P] of the implication premise, we want [P] to "have a turn," where the premise is rearranged into the form [P /\ Q] for some [Q]. The tactic [tac] should expect to see a goal in this form and focus its attention on the first conjunct of the premise. *)
642 651
643 Ltac search_prem tac := 652 Ltac search_prem tac :=
644 let rec search P := 653 let rec search P :=
645 tac 654 tac
646 || (apply and_True_prem; tac) 655 || (apply and_True_prem; tac)
653 | [ |- ?P /\ _ --> _ ] => search P 662 | [ |- ?P /\ _ --> _ ] => search P
654 | [ |- _ /\ ?P --> _ ] => apply comm_prem; search P 663 | [ |- _ /\ ?P --> _ ] => apply comm_prem; search P
655 | [ |- _ --> _ ] => progress (tac || (apply and_True_prem; tac)) 664 | [ |- _ --> _ ] => progress (tac || (apply and_True_prem; tac))
656 end. 665 end.
657 666
667 (** To understand how [search_prem] works, we turn first to the final [match]. If the premise begins with a conjunction, we call the [search] procedure on each of the conjuncts, or only the first conjunct, if that already yields a case where [tac] does not fail. [search P] expects and maintains the invariant that the premise is of the form [P /\ Q] for some [Q]. We pass [P] explicitly as a kind of decreasing induction measure, to avoid looping forever when [tac] always fails. The second [match] case calls a commutativity lemma to realize this invariant, before passing control to [search]. The final [match] case tries applying [tac] directly and then, if that fails, changes the form of the goal by adding an extraneous [True] conjunct and calls [tac] again.
668
669 [search] itself tries the same tricks as in the last case of the final [match]. Additionally, if neither works, it checks if [P] is a conjunction. If so, it calls itself recursively on each conjunct, first applying associativity lemmas to maintain the goal-form invariant.
670
671 We will also want a dual function [search_conc], which does tree search through an [imp] conclusion. *)
672
658 Ltac search_conc tac := 673 Ltac search_conc tac :=
659 let rec search P := 674 let rec search P :=
660 tac 675 tac
661 || (apply and_True_conc; tac) 676 || (apply and_True_conc; tac)
662 || match P with 677 || match P with
668 | [ |- _ --> ?P /\ _ ] => search P 683 | [ |- _ --> ?P /\ _ ] => search P
669 | [ |- _ --> _ /\ ?P ] => apply comm_conc; search P 684 | [ |- _ --> _ /\ ?P ] => apply comm_conc; search P
670 | [ |- _ --> _ ] => progress (tac || (apply and_True_conc; tac)) 685 | [ |- _ --> _ ] => progress (tac || (apply and_True_conc; tac))
671 end. 686 end.
672 687
688 (** Now we can prove a number of lemmas that are suitable for application by our search tactics. A lemma that is meant to handle a premise should have the form [P /\ Q --> R] for some interesting [P], and a lemma that is meant to handle a conclusion should have the form [P --> Q /\ R] for some interesting [Q]. *)
689
673 Theorem False_prem : forall P Q, 690 Theorem False_prem : forall P Q,
674 False /\ P --> Q. 691 False /\ P --> Q.
675 imp. 692 imp.
676 Qed. 693 Qed.
677 694
697 (Q --> P x /\ R) 714 (Q --> P x /\ R)
698 -> (Q --> ex P /\ R). 715 -> (Q --> ex P /\ R).
699 imp. 716 imp.
700 Qed. 717 Qed.
701 718
719 (** We will also want a "base case" lemma for finishing proofs where cancelation has removed every constituent of the conclusion. *)
720
702 Theorem imp_True : forall P, 721 Theorem imp_True : forall P,
703 P --> True. 722 P --> True.
704 imp. 723 imp.
705 Qed. 724 Qed.
725
726 (** Our final [matcher] tactic is now straightforward. First, we [intros] all variables into scope. Then we attempt simple premise simplifications, finishing the proof upon finding [False] and eliminating any existential quantifiers that we find. After that, we search through the conclusion. We remove [True] conjuncts, remove existential quantifiers by introducing unification variables for their bound variables, and search for matching premises to cancel. Finally, when no more progress is made, we see if the goal has become trivial and can be solved by [imp_True]. *)
706 727
707 Ltac matcher := 728 Ltac matcher :=
708 intros; 729 intros;
709 repeat search_prem ltac:(apply False_prem || (apply ex_prem; intro)); 730 repeat search_prem ltac:(apply False_prem || (apply ex_prem; intro));
710 repeat search_conc ltac:(apply True_conc || eapply ex_conc || search_prem ltac:(apply Match)); 731 repeat search_conc ltac:(apply True_conc || eapply ex_conc
711 try apply imp_True. 732 || search_prem ltac:(apply Match));
733 try apply imp_True.
734
735 (** Our tactic succeeds at proving a simple example. *)
712 736
713 Theorem t2 : forall P Q : Prop, 737 Theorem t2 : forall P Q : Prop,
714 Q /\ (P /\ False) /\ P --> P /\ Q. 738 Q /\ (P /\ False) /\ P --> P /\ Q.
715 matcher. 739 matcher.
716 Qed. 740 Qed.
717 741
742 (** In the generated proof, we find a trace of the workings of the search tactics. *)
743
744 Print t2.
745 (** [[
746
747 t2 =
748 fun P Q : Prop =>
749 comm_prem (assoc_prem1 (assoc_prem2 (False_prem (P:=P /\ P /\ Q) (P /\ Q))))
750 : forall P Q : Prop, Q /\ (P /\ False) /\ P --> P /\ Q
751 ]] *)
752
753 (** We can also see that [matcher] is well-suited for cases where some human intervention is needed after the automation finishes. *)
754
718 Theorem t3 : forall P Q R : Prop, 755 Theorem t3 : forall P Q R : Prop,
719 P /\ Q --> Q /\ R /\ P. 756 P /\ Q --> Q /\ R /\ P.
720 matcher. 757 matcher.
758 (** [[
759
760 ============================
761 True --> R
762 ]]
763
764 [matcher] canceled those conjuncts that it was able to cancel, leaving a simplified subgoal for us, much as [intuition] does. *)
721 Abort. 765 Abort.
766
767 (** [matcher] even succeeds at guessing quantifier instantiations. It is the unification that occurs in uses of the [Match] lemma that does the real work here. *)
722 768
723 Theorem t4 : forall (P : nat -> Prop) Q, (exists x, P x /\ Q) --> Q /\ (exists x, P x). 769 Theorem t4 : forall (P : nat -> Prop) Q, (exists x, P x /\ Q) --> Q /\ (exists x, P x).
724 matcher. 770 matcher.
725 Qed. 771 Qed.
726 772
773 Print t4.
774
775 (** [[
776
777 t4 =
778 fun (P : nat -> Prop) (Q : Prop) =>
779 and_True_prem
780 (ex_prem (P:=fun x : nat => P x /\ Q)
781 (fun x : nat =>
782 assoc_prem2
783 (Match (P:=Q)
784 (and_True_conc
785 (ex_conc (fun x0 : nat => P x0) x
786 (Match (P:=P x) (imp_True (P:=True))))))))
787 : forall (P : nat -> Prop) (Q : Prop),
788 (exists x : nat, P x /\ Q) --> Q /\ (exists x : nat, P x)
789 ]] *)
790
791