annotate src/Match.v @ 135:091583baf345

Through completer'
author Adam Chlipala <adamc@hcoop.net>
date Sun, 26 Oct 2008 14:11:41 -0400
parents f9d8f33c9b46
children d6b1e9de7fc1
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@134 88 [[
adamc@133 89 User error: Bound head variable
adamc@134 90 ]]
adamc@133 91
adamc@134 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@134 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.
adamc@134 95
adamc@134 96 We have used [Hint Rewrite] in many examples so far. [crush] uses these hints by calling [autorewrite]. Our rewrite hints have taken the form [Hint Rewrite lemma : cpdt], adding them to the [cpdt] rewrite database. This is because, in contrast to [auto], [autorewrite] has no default database. Thus, we set the convention that [crush] uses the [cpdt] database.
adamc@134 97
adamc@134 98 This example shows a direct use of [autorewrite]. *)
adamc@134 99
adamc@134 100 Section autorewrite.
adamc@134 101 Variable A : Set.
adamc@134 102 Variable f : A -> A.
adamc@134 103
adamc@134 104 Hypothesis f_f : forall x, f (f x) = f x.
adamc@134 105
adamc@134 106 Hint Rewrite f_f : my_db.
adamc@134 107
adamc@134 108 Lemma f_f_f : forall x, f (f (f x)) = f x.
adamc@134 109 intros; autorewrite with my_db; reflexivity.
adamc@134 110 Qed.
adamc@134 111
adamc@134 112 (** 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: *)
adamc@134 113
adamc@134 114 Section garden_path.
adamc@134 115 Variable g : A -> A.
adamc@134 116 Hypothesis f_g : forall x, f x = g x.
adamc@134 117 Hint Rewrite f_g : my_db.
adamc@134 118
adamc@134 119 Lemma f_f_f' : forall x, f (f (f x)) = f x.
adamc@134 120 intros; autorewrite with my_db.
adamc@134 121 (** [[
adamc@134 122
adamc@134 123 ============================
adamc@134 124 g (g (g x)) = g x
adamc@134 125 ]] *)
adamc@134 126 Abort.
adamc@134 127
adamc@134 128 (** 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. *)
adamc@134 129
adamc@134 130 Reset garden_path.
adamc@134 131
adamc@134 132 (** [autorewrite] works with quantified equalities that include additional premises, but we must be careful to avoid similar incorrect rewritings. *)
adamc@134 133
adamc@134 134 Section garden_path.
adamc@134 135 Variable P : A -> Prop.
adamc@134 136 Variable g : A -> A.
adamc@134 137 Hypothesis f_g : forall x, P x -> f x = g x.
adamc@134 138 Hint Rewrite f_g : my_db.
adamc@134 139
adamc@134 140 Lemma f_f_f' : forall x, f (f (f x)) = f x.
adamc@134 141 intros; autorewrite with my_db.
adamc@134 142 (** [[
adamc@134 143
adamc@134 144 ============================
adamc@134 145 g (g (g x)) = g x
adamc@134 146
adamc@134 147 subgoal 2 is:
adamc@134 148 P x
adamc@134 149 subgoal 3 is:
adamc@134 150 P (f x)
adamc@134 151 subgoal 4 is:
adamc@134 152 P (f x)
adamc@134 153 ]] *)
adamc@134 154 Abort.
adamc@134 155
adamc@134 156 (** The inappropriate rule fired the same three times as before, even though we know we will not be able to prove the premises. *)
adamc@134 157
adamc@134 158 Reset garden_path.
adamc@134 159
adamc@134 160 (** Our final, successful, attempt uses an extra argument to [Hint Rewrite] that specifies a tactic to apply to generated premises. *)
adamc@134 161
adamc@134 162 Section garden_path.
adamc@134 163 Variable P : A -> Prop.
adamc@134 164 Variable g : A -> A.
adamc@134 165 Hypothesis f_g : forall x, P x -> f x = g x.
adamc@134 166 Hint Rewrite f_g using assumption : my_db.
adamc@134 167
adamc@134 168 Lemma f_f_f' : forall x, f (f (f x)) = f x.
adamc@134 169 intros; autorewrite with my_db; reflexivity.
adamc@134 170 Qed.
adamc@134 171
adamc@134 172 (** [autorewrite] will still use [f_g] when the generated premise is among our assumptions. *)
adamc@134 173
adamc@134 174 Lemma f_f_f_g : forall x, P x -> f (f x) = g x.
adamc@134 175 intros; autorewrite with my_db; reflexivity.
adamc@134 176 Qed.
adamc@134 177 End garden_path.
adamc@134 178
adamc@134 179 (** It can also be useful to use the [autorewrite with db in *] form, which does rewriting in hypotheses, as well as in the conclusion. *)
adamc@134 180
adamc@134 181 Lemma in_star : forall x y, f (f (f (f x))) = f (f y)
adamc@134 182 -> f x = f (f (f y)).
adamc@134 183 intros; autorewrite with my_db in *; assumption.
adamc@134 184 Qed.
adamc@134 185
adamc@134 186 End autorewrite.
adamc@135 187
adamc@135 188
adamc@135 189 (** * Ltac Programming Basics *)
adamc@135 190
adamc@135 191 (** We have already seen many examples of Ltac programs. In the rest of this chapter, we attempt to give a more principled introduction to the important features and design patterns.
adamc@135 192
adamc@135 193 One common use for [match] tactics is identification of subjects for case analysis, as we see in this tactic definition. *)
adamc@135 194
adamc@135 195 Ltac find_if :=
adamc@135 196 match goal with
adamc@135 197 | [ |- if ?X then _ else _ ] => destruct X
adamc@135 198 end.
adamc@135 199
adamc@135 200 (** The tactic checks if the conclusion is an [if], [destruct]ing the test expression if so. Certain classes of theorem are trivial to prove automatically with such a tactic. *)
adamc@135 201
adamc@135 202 Theorem hmm : forall (a b c : bool),
adamc@135 203 if a
adamc@135 204 then if b
adamc@135 205 then True
adamc@135 206 else True
adamc@135 207 else if c
adamc@135 208 then True
adamc@135 209 else True.
adamc@135 210 intros; repeat find_if; constructor.
adamc@135 211 Qed.
adamc@135 212
adamc@135 213 (** The [repeat] that we use here is called a %\textit{%#<i>#tactical#</i>#%}%, or tactic combinator. The behavior of [repeat t] is to loop through running [t], running [t] on all generated subgoals, running [t] on %\textit{%#<i>#their#</i>#%}% generated subgoals, and so on. When [t] fails at any point in this search tree, that particular subgoal is left to be handled by later tactics. Thus, it is important never to use [repeat] with a tactic that always succeeds.
adamc@135 214
adamc@135 215 Another very useful Ltac building block is %\textit{%#<i>#context patterns#</i>#%}%. *)
adamc@135 216
adamc@135 217 Ltac find_if_inside :=
adamc@135 218 match goal with
adamc@135 219 | [ |- context[if ?X then _ else _] ] => destruct X
adamc@135 220 end.
adamc@135 221
adamc@135 222 (** The behavior of this tactic is to find any subterm of the conclusion that is an [if] and then [destruct] the test expression. This version subsumes [find_if]. *)
adamc@135 223
adamc@135 224 Theorem hmm' : forall (a b c : bool),
adamc@135 225 if a
adamc@135 226 then if b
adamc@135 227 then True
adamc@135 228 else True
adamc@135 229 else if c
adamc@135 230 then True
adamc@135 231 else True.
adamc@135 232 intros; repeat find_if_inside; constructor.
adamc@135 233 Qed.
adamc@135 234
adamc@135 235 (** We can also use [find_if_inside] to prove goals that [find_if] does not simplify sufficiently. *)
adamc@135 236
adamc@135 237 Theorem duh2 : forall (a b : bool),
adamc@135 238 (if a then 42 else 42) = (if b then 42 else 42).
adamc@135 239 intros; repeat find_if_inside; reflexivity.
adamc@135 240 Qed.
adamc@135 241
adamc@135 242 (** Many decision procedures can be coded in Ltac via "[repeat match] loops." For instance, we can implement a subset of the functionality of [tauto]. *)
adamc@135 243
adamc@135 244 Ltac my_tauto :=
adamc@135 245 repeat match goal with
adamc@135 246 | [ H : ?P |- ?P ] => exact H
adamc@135 247
adamc@135 248 | [ |- True ] => constructor
adamc@135 249 | [ |- _ /\ _ ] => constructor
adamc@135 250 | [ |- _ -> _ ] => intro
adamc@135 251
adamc@135 252 | [ H : False |- _ ] => destruct H
adamc@135 253 | [ H : _ /\ _ |- _ ] => destruct H
adamc@135 254 | [ H : _ \/ _ |- _ ] => destruct H
adamc@135 255
adamc@135 256 | [ H1 : ?P -> ?Q, H2 : ?P |- _ ] =>
adamc@135 257 let H := fresh "H" in
adamc@135 258 generalize (H1 H2); clear H1; intro H
adamc@135 259 end.
adamc@135 260
adamc@135 261 (** Since [match] patterns can share unification variables between hypothesis and conclusion patterns, it is easy to figure out when the conclusion matches a hypothesis. The [exact] tactic solves a goal completely when given a proof term of the proper type.
adamc@135 262
adamc@135 263 It is also trivial to implement the "introduction rules" for a few of the connectives. Implementing elimination rules is only a little more work, since we must bind a name for a hypothesis to [destruct].
adamc@135 264
adamc@135 265 The last rule implements modus ponens. The most interesting part is the use of the Ltac-level [let] with a [fresh] expression. [fresh] takes in a name base and returns a fresh hypothesis variable based on that name. We use the new name variable [H] as the name we assign to the result of modus ponens. The use of [generalize] changes our conclusion to be an implication from [Q]. We clear the original hypothesis and move [Q] into the context with name [H]. *)
adamc@135 266
adamc@135 267 Section propositional.
adamc@135 268 Variables P Q R : Prop.
adamc@135 269
adamc@135 270 Theorem and_comm : (P \/ Q \/ False) /\ (P -> Q) -> True /\ Q.
adamc@135 271 my_tauto.
adamc@135 272 Qed.
adamc@135 273 End propositional.
adamc@135 274
adamc@135 275 (** It was relatively easy to implement modus ponens, because we do not lose information by clearing every implication that we use. If we want to implement a similarly-complete procedure for quantifier instantiation, we need a way to ensure that a particular proposition is not already included among our hypotheses. To do that effectively, we first need to learn a bit more about the semantics of [match].
adamc@135 276
adamc@135 277 It is tempting to assume that [match] works like it does in ML. In fact, there are a few critical differences in its behavior. One is that we may include arbitrary expressions in patterns, instead of being restricted to variables and constructors. Another is that the same variable may appear multiple times, inducing an implicit equality constraint.
adamc@135 278
adamc@135 279 There is a related pair of two other differences that are much more important than the others. [match] has a %\textit{%#<i>#backtracking semantics for failure#</i>#%}%. In ML, pattern matching works by finding the first pattern to match and then executing its body. If the body raises an exception, then the overall match raises the same exception. In Coq, failures in case bodies instead trigger continued search through the list of cases.
adamc@135 280
adamc@135 281 For instance, this (unnecessarily verbose) proof script works: *)
adamc@135 282
adamc@135 283 Theorem m1 : True.
adamc@135 284 match goal with
adamc@135 285 | [ |- _ ] => intro
adamc@135 286 | [ |- True ] => constructor
adamc@135 287 end.
adamc@135 288 Qed.
adamc@135 289
adamc@135 290 (** The first case matches trivially, but its body tactic fails, since the conclusion does not begin with a quantifier or implication. In a similar ML match, that would mean that the whole pattern-match fails. In Coq, we backtrack and try the next pattern, which also matches. Its body tactic succeeds, so the overall tactic succeeds as well.
adamc@135 291
adamc@135 292 The example shows how failure can move to a different pattern within a [match]. Failure can also trigger an attempt to find %\textit{%#<i>#a different way of matching a single pattern#</i>#%}%. Consider another example: *)
adamc@135 293
adamc@135 294 Theorem m2 : forall P Q R : Prop, P -> Q -> R -> Q.
adamc@135 295 intros; match goal with
adamc@135 296 | [ H : _ |- _ ] => pose H
adamc@135 297 end.
adamc@135 298 (** [[
adamc@135 299
adamc@135 300 r := H1 : R
adamc@135 301 ============================
adamc@135 302 Q
adamc@135 303 ]]
adamc@135 304
adamc@135 305 By applying [pose], a convenient debugging tool for "leaking information out of [match]es," we see that this [match] first tries binding [H] to [H1], which cannot be used to prove [Q]. Nonetheless, the following variation on the tactic succeeds at proving the goal: *)
adamc@135 306
adamc@135 307 match goal with
adamc@135 308 | [ H : _ |- _ ] => exact H
adamc@135 309 end.
adamc@135 310 Qed.
adamc@135 311
adamc@135 312 (** The tactic first unifies [H] with [H1], as before, but [exact H] fails in that case, so the tactic engine searches for more possible values of [H]. Eventually, it arrives at the correct value, so that [exact H] and the overall tactic succeed. *)
adamc@135 313
adamc@135 314 (** Now we are equipped to implement a tactic for checking that a proposition is not among our hypotheses: *)
adamc@135 315
adamc@135 316 Ltac notHyp P :=
adamc@135 317 match goal with
adamc@135 318 | [ _ : P |- _ ] => fail 1
adamc@135 319 | _ =>
adamc@135 320 match P with
adamc@135 321 | ?P1 /\ ?P2 => first [ notHyp P1 | notHyp P2 | fail 2 ]
adamc@135 322 | _ => idtac
adamc@135 323 end
adamc@135 324 end.
adamc@135 325
adamc@135 326 (** We use the equality checking that is built into pattern-matching to see if there is a hypothesis that matches the proposition exactly. If so, we use the [fail] tactic. Without arguments, [fail] signals normal tactic failure, as you might expect. When [fail] is passed an argument [n], [n] is used to count outwards through the enclosing cases of backtracking search. In this case, [fail 1] says "fail not just in this pattern-matching branch, but for the whole [match]." The second case will never be tried when the [fail 1] is reached.
adamc@135 327
adamc@135 328 This second case, used when [P] matches no hypothesis, checks if [P] is a conjunction. Other simplifications may have split conjunctions into their component formulas, so we need to check that at least one of those components is also not represented. To achieve this, we apply the [first] tactical, which takes a list of tactics and continues down the list until one of them does not fail. The [fail 2] at the end says to [fail] both the [first] and the [match] wrapped around it.
adamc@135 329
adamc@135 330 The body of the [?P1 /\ ?P2] case guarantees that, if it is reached, we either succeed completely or fail completely. Thus, if we reach the wildcard case, [P] is not a conjunction. We use [idtac], a tactic that would be silly to apply on its own, since its effect is to succeed at doing nothing. Nonetheless, [idtac] is a useful placeholder for cases like what we see here.
adamc@135 331
adamc@135 332 With the non-presence check implemented, it is easy to build a tactic that takes as input a proof term and adds its conclusion as a new hypothesis, only if that conclusion is not already present, failing otherwise. *)
adamc@135 333
adamc@135 334 Ltac extend pf :=
adamc@135 335 let t := type of pf in
adamc@135 336 notHyp t; generalize pf; intro.
adamc@135 337
adamc@135 338 (** We see the useful [type of] operator of Ltac. This operator could not be implemented in Gallina, but it is easy to support in Ltac. We end up with [t] bound to the type of [pf]. We check that [t] is not already present. If so, we use a [generalize]/[intro] combo to add a new hypothesis proved by [pf].
adamc@135 339
adamc@135 340 With these tactics defined, we can write a tactic [completer] for adding to the context all consequences of a set of simple first-order formulas. *)
adamc@135 341
adamc@135 342 Ltac completer :=
adamc@135 343 repeat match goal with
adamc@135 344 | [ |- _ /\ _ ] => constructor
adamc@135 345 | [ H : _ /\ _ |- _ ] => destruct H
adamc@135 346 | [ H : ?P -> ?Q, H' : ?P |- _ ] =>
adamc@135 347 generalize (H H'); clear H; intro H
adamc@135 348 | [ |- forall x, _ ] => intro
adamc@135 349
adamc@135 350 | [ H : forall x, ?P x -> _, H' : ?P ?X |- _ ] =>
adamc@135 351 extend (H X H')
adamc@135 352 end.
adamc@135 353
adamc@135 354 (** We use the same kind of conjunction and implication handling as previously. Note that, since [->] is the special non-dependent case of [forall], the fourth rule handles [intro] for implications, too.
adamc@135 355
adamc@135 356 In the fifth rule, when we find a [forall] fact [H] with a premise matching one of our hypotheses, we add the appropriate instantiation of [H]'s conclusion, if we have not already added it.
adamc@135 357
adamc@135 358 We can check that [completer] is working properly: *)
adamc@135 359
adamc@135 360 Section firstorder.
adamc@135 361 Variable A : Set.
adamc@135 362 Variables P Q R S : A -> Prop.
adamc@135 363
adamc@135 364 Hypothesis H1 : forall x, P x -> Q x /\ R x.
adamc@135 365 Hypothesis H2 : forall x, R x -> S x.
adamc@135 366
adamc@135 367 Theorem fo : forall x, P x -> S x.
adamc@135 368 completer.
adamc@135 369 (** [[
adamc@135 370
adamc@135 371 x : A
adamc@135 372 H : P x
adamc@135 373 H0 : Q x
adamc@135 374 H3 : R x
adamc@135 375 H4 : S x
adamc@135 376 ============================
adamc@135 377 S x
adamc@135 378 ]] *)
adamc@135 379
adamc@135 380 assumption.
adamc@135 381 Qed.
adamc@135 382 End firstorder.
adamc@135 383
adamc@135 384 (** We narrowly avoided a subtle pitfall in our definition of [completer]. Let us try another definition that even seems preferable to the original, to the untrained eye. *)
adamc@135 385
adamc@135 386 Ltac completer' :=
adamc@135 387 repeat match goal with
adamc@135 388 | [ |- _ /\ _ ] => constructor
adamc@135 389 | [ H : _ /\ _ |- _ ] => destruct H
adamc@135 390 | [ H : ?P -> _, H' : ?P |- _ ] =>
adamc@135 391 generalize (H H'); clear H; intro H
adamc@135 392 | [ |- forall x, _ ] => intro
adamc@135 393
adamc@135 394 | [ H : forall x, ?P x -> _, H' : ?P ?X |- _ ] =>
adamc@135 395 extend (H X H')
adamc@135 396 end.
adamc@135 397
adamc@135 398 (** The only difference is in the modus ponens rule, where we have replaced an unused unification variable [?Q] with a wildcard. Let us try our example again with this version: *)
adamc@135 399
adamc@135 400 Section firstorder'.
adamc@135 401 Variable A : Set.
adamc@135 402 Variables P Q R S : A -> Prop.
adamc@135 403
adamc@135 404 Hypothesis H1 : forall x, P x -> Q x /\ R x.
adamc@135 405 Hypothesis H2 : forall x, R x -> S x.
adamc@135 406
adamc@135 407 Theorem fo' : forall x, P x -> S x.
adamc@135 408 (** [[
adamc@135 409
adamc@135 410 completer'.
adamc@135 411
adamc@135 412 Coq loops forever at this point. What went wrong? *)
adamc@135 413 Abort.
adamc@135 414 End firstorder'.