annotate src/Match.v @ 133:28ef7f0da085

auto
author Adam Chlipala <adamc@hcoop.net>
date Sun, 26 Oct 2008 10:28:00 -0400
parents bc1f7d3687e7
children f9d8f33c9b46
rev   line source
adamc@132 1 (* Copyright (c) 2008, Adam Chlipala
adamc@132 2 *
adamc@132 3 * This work is licensed under a
adamc@132 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@132 5 * Unported License.
adamc@132 6 * The license text is available at:
adamc@132 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@132 8 *)
adamc@132 9
adamc@132 10 (* begin hide *)
adamc@132 11 Require Import List.
adamc@132 12
adamc@132 13 Require Import Tactics.
adamc@132 14
adamc@132 15 Set Implicit Arguments.
adamc@132 16 (* end hide *)
adamc@132 17
adamc@132 18
adamc@132 19 (** %\part{Proof Engineering}
adamc@132 20
adamc@132 21 \chapter{Proof Search in Ltac}% *)
adamc@132 22
adamc@132 23 (** We have seen many examples of proof automation so far. This chapter aims to give a principled presentation of the features of Ltac, focusing in particular on the Ltac [match] construct, which supports a novel approach to backtracking search. First, though, we will run through some useful automation tactics that are built into Coq. They are described in detail in the manual, so we only outline what is possible. *)
adamc@132 24
adamc@132 25 (** * Some Built-In Automation Tactics *)
adamc@132 26
adamc@132 27 (** A number of tactics are called repeatedly by [crush]. [intuition] simplifies propositional structure of goals. [congruence] applies the rules of equality and congruence closure, plus properties of constructors of inductive types. The [omega] tactic provides a complete decision procedure for a theory that is called quantifier-free linear arithmetic or Presburger arithmetic, depending on whom you ask. That is, [omega] proves any goal that follows from looking only at parts of that goal that can be interpreted as propositional formulas whose atomic formulas are basic comparison operations on natural numbers or integers.
adamc@132 28
adamc@132 29 The [ring] tactic solves goals by appealing to the axioms of rings or semi-rings (as in algebra), depending on the type involved. Coq developments may declare new types to be parts of rings and semi-rings by proving the associated axioms. There is a simlar tactic [field] for simplifying values in fields by conversion to fractions over rings. Both [ring] and [field] can only solve goals that are equalities. The [fourier] tactic uses Fourier's method to prove inequalities over real numbers, which are axiomatized in the Coq standard library.
adamc@132 30
adamc@133 31 The %\textit{%#<i>#setoid#</i>#%}% facility makes it possible to register new equivalence relations to be understood by tactics like [rewrite]. For instance, [Prop] is registered as a setoid with the equivalence relation "if and only if." The ability to register new setoids can be very useful in proofs of a kind common in math, where all reasoning is done after "modding out by a relation." *)
adamc@132 32
adamc@132 33
adamc@133 34 (** * Hint Databases *)
adamc@132 35
adamc@133 36 (** Another class of built-in tactics includes [auto], [eauto], and [autorewrite]. These are based on %\textit{%#<i>#hint databases#</i>#%}%, which we have seen extended in many examples so far. These tactics are 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.
adamc@133 37
adamc@133 38 The basic hints for [auto] and [eauto] are [Hint Immediate lemma], asking to try solving a goal immediately by applying the premise-free lemma; [Resolve lemma], which does the same but may add new premises that are themselves to be subjects of proof search; [Constructor type], which acts like [Resolve] applied to every constructor of an inductive type; and [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.
adamc@133 39
adamc@133 40 All of these [Hint] commands can be issued alternatively with a more primitive hint kind, [Extern]. A few examples should do best to explain how [Hint Extern] works. *)
adamc@133 41
adamc@133 42 Theorem bool_neq : true <> false.
adamc@133 43 auto.
adamc@133 44 (** [crush] would have discharged this goal, but the default hint database for [auto] contains no hint that applies. *)
adamc@133 45 Abort.
adamc@133 46
adamc@133 47 (** 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. *)
adamc@133 48
adamc@133 49 Hint Extern 1 (_ <> _) => congruence.
adamc@133 50
adamc@133 51 Theorem bool_neq : true <> false.
adamc@133 52 auto.
adamc@133 53 Qed.
adamc@133 54
adamc@133 55 (** Our hint says: "whenever the conclusion matches the pattern [_ <> _], try applying [congruence]." The [1] is a cost for this rule. During proof search, whenever multiple rules apply, rules are tried in increasing cost order, so it pays to assign high costs to relatively expensive [Extern] hints.
adamc@133 56
adamc@133 57 [Extern] hints may be implemented with the full Ltac language. This example shows a case where a hint uses a [match]. *)
adamc@133 58
adamc@133 59 Section forall_and.
adamc@133 60 Variable A : Set.
adamc@133 61 Variables P Q : A -> Prop.
adamc@133 62
adamc@133 63 Hypothesis both : forall x, P x /\ Q x.
adamc@133 64
adamc@133 65 Theorem forall_and : forall z, P z.
adamc@133 66 crush.
adamc@133 67 (** [crush] makes no progress beyond what [intros] would have accomplished. [auto] 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. *)
adamc@133 68
adamc@133 69 Hint Extern 1 (P ?X) =>
adamc@133 70 match goal with
adamc@133 71 | [ H : forall x, P x /\ _ |- _ ] => apply (proj1 (H X))
adamc@133 72 end.
adamc@133 73
adamc@133 74 auto.
adamc@133 75 Qed.
adamc@133 76
adamc@133 77 (** We see that an [Extern] pattern may bind unification variables that we use in the associated tactic. [proj1] is a function from the standard library for extracting a proof of [R] from a proof of [R /\ S]. *)
adamc@133 78 End forall_and.
adamc@133 79
adamc@133 80 (** After our success on this example, we might get more ambitious and seek to generalize the hint to all possible predicates [P].
adamc@133 81
adamc@133 82 [[
adamc@133 83 Hint Extern 1 (?P ?X) =>
adamc@133 84 match goal with
adamc@133 85 | [ H : forall x, ?P x /\ _ |- _ ] => apply (proj1 (H X))
adamc@133 86 end.
adamc@133 87
adamc@133 88 [[
adamc@133 89 User error: Bound head variable
adamc@133 90 ]]
adamc@133 91
adamc@133 92 Coq's [auto] hint databases work as tables mapping %\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].
adamc@133 93
adamc@133 94 This restriction on [Extern] hints is the main limitation of the [auto] mechanism, preventing us from using it for general context simplifications that are not keyed off of the form of the conclusion. This is perhaps just as well, since we can often code more efficient tactics with specialized Ltac programs, and we will see how in later sections of the chapter. *)