CpdtTactics.v 10.6 KB
Newer Older
Adam Chlipala's avatar
Adam Chlipala committed
1
(* Copyright (c) 2008-2012, Adam Chlipala
2 3 4 5 6 7 8 9
 * 
 * This work is licensed under a
 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
 * Unported License.
 * The license text is available at:
 *   http://creativecommons.org/licenses/by-nc-nd/3.0/
 *)

10
Require Import Eqdep List.
11

Adam Chlipala's avatar
Adam Chlipala committed
12 13
Require Omega.

14 15
Set Implicit Arguments.

16

Adam Chlipala's avatar
Adam Chlipala committed
17
(** A version of [injection] that does some standard simplifications afterward: clear the hypothesis in question, bring the new facts above the double line, and attempt substitution for known variables. *)
18
Ltac inject H := injection H; clear H; intros; try subst.
19

Adam Chlipala's avatar
Adam Chlipala committed
20
(** Try calling tactic function [f] on all hypotheses, keeping the first application that doesn't fail. *)
21 22 23 24 25
Ltac appHyps f :=
  match goal with
    | [ H : _ |- _ ] => f H
  end.

Adam Chlipala's avatar
Adam Chlipala committed
26
(** Succeed iff [x] is in the list [ls], represented with left-associated nested tuples. *)
27 28 29
Ltac inList x ls :=
  match ls with
    | x => idtac
30
    | (_, x) => idtac
31 32 33
    | (?LS, _) => inList x LS
  end.

Adam Chlipala's avatar
Adam Chlipala committed
34
(** Try calling tactic function [f] on every element of tupled list [ls], keeping the first call not to fail. *)
35 36 37 38 39 40
Ltac app f ls :=
  match ls with
    | (?LS, ?X) => f X || app f LS || fail 1
    | _ => f ls
  end.

Adam Chlipala's avatar
Adam Chlipala committed
41
(** Run [f] on every element of [ls], not just the first that doesn't fail. *)
42 43 44 45 46 47 48
Ltac all f ls :=
  match ls with
    | (?LS, ?X) => f X; all f LS
    | (_, _) => fail 1
    | _ => f ls
  end.

Adam Chlipala's avatar
Adam Chlipala committed
49 50
(** Workhorse tactic to simplify hypotheses for a variety of proofs.
   * Argument [invOne] is a tuple-list of predicates for which we always do inversion automatically. *)
51
Ltac simplHyp invOne :=
Adam Chlipala's avatar
Adam Chlipala committed
52
  (** Helper function to do inversion on certain hypotheses, where [H] is the hypothesis and [F] its head symbol *)
53
  let invert H F :=
Adam Chlipala's avatar
Adam Chlipala committed
54 55 56 57 58 59 60
    (** We only proceed for those predicates in [invOne]. *)
    inList F invOne;
    (** This case covers an inversion that succeeds immediately, meaning no constructors of [F] applied. *)
      (inversion H; fail)
    (** Otherwise, we only proceed if inversion eliminates all but one constructor case. *)
      || (inversion H; [idtac]; clear H; try subst) in

Adam Chlipala's avatar
Adam Chlipala committed
61
  match goal with
Adam Chlipala's avatar
Adam Chlipala committed
62
    (** Eliminate all existential hypotheses. *)
63 64
    | [ H : ex _ |- _ ] => destruct H

Adam Chlipala's avatar
Adam Chlipala committed
65
    (** Find opportunities to take advantage of injectivity of data constructors, for several different arities. *)
66
    | [ H : ?F ?X = ?F ?Y |- ?G ] =>
Adam Chlipala's avatar
Adam Chlipala committed
67
      (** This first branch of the [||] fails the whole attempt iff the arguments of the constructor applications are already easy to prove equal. *)
68
      (assert (X = Y); [ assumption | fail 1 ])
Adam Chlipala's avatar
Adam Chlipala committed
69 70
      (** If we pass that filter, then we use injection on [H] and do some simplification as in [inject].
         * The odd-looking check of the goal form is to avoid cases where [injection] gives a more complex result because of dependent typing, which we aren't equipped to handle here. *)
71 72 73 74 75 76 77 78 79 80 81 82 83
      || (injection H;
        match goal with
          | [ |- X = Y -> G ] =>
            try clear H; intros; try subst
        end)
    | [ H : ?F ?X ?U = ?F ?Y ?V |- ?G ] =>
      (assert (X = Y); [ assumption
        | assert (U = V); [ assumption | fail 1 ] ])
      || (injection H;
        match goal with
          | [ |- U = V -> X = Y -> G ] =>
            try clear H; intros; try subst
        end)
84

Adam Chlipala's avatar
Adam Chlipala committed
85
    (** Consider some different arities of a predicate [F] in a hypothesis that we might want to invert. *)
86 87 88 89 90
    | [ H : ?F _ |- _ ] => invert H F
    | [ H : ?F _ _ |- _ ] => invert H F
    | [ H : ?F _ _ _ |- _ ] => invert H F
    | [ H : ?F _ _ _ _ |- _ ] => invert H F
    | [ H : ?F _ _ _ _ _ |- _ ] => invert H F
Adam Chlipala's avatar
Adam Chlipala committed
91

Adam Chlipala's avatar
Adam Chlipala committed
92
    (** Use an (axiom-dependent!) inversion principle for dependent pairs, from the standard library. *)
Adam Chlipala's avatar
Adam Chlipala committed
93
    | [ H : existT _ ?T _ = existT _ ?T _ |- _ ] => generalize (inj_pair2 _ _ _ _ _ H); clear H
Adam Chlipala's avatar
Adam Chlipala committed
94 95

    (** If we're not ready to use that principle yet, try the standard inversion, which often enables the previous rule. *)
Adam Chlipala's avatar
Adam Chlipala committed
96
    | [ H : existT _ _ _ = existT _ _ _ |- _ ] => inversion H; clear H
97

Adam Chlipala's avatar
Adam Chlipala committed
98
    (** Similar logic to the cases for constructor injectivity above, but specialized to [Some], since the above cases won't deal with polymorphic constructors. *)
99
    | [ H : Some _ = Some _ |- _ ] => injection H; clear H
Adam Chlipala's avatar
Adam Chlipala committed
100 101
  end.

Adam Chlipala's avatar
Adam Chlipala committed
102
(** Find some hypothesis to rewrite with, ensuring that [auto] proves all of the extra subgoals added by [rewrite]. *)
103 104
Ltac rewriteHyp :=
  match goal with
Adam Chlipala's avatar
Adam Chlipala committed
105
    | [ H : _ |- _ ] => rewrite H by solve [ auto ]
106 107
  end.

Adam Chlipala's avatar
Adam Chlipala committed
108
(** Combine [autorewrite] with automatic hypothesis rewrites. *)
109 110
Ltac rewriterP := repeat (rewriteHyp; autorewrite with core in *).
Ltac rewriter := autorewrite with core in *; rewriterP.
111

Adam Chlipala's avatar
Adam Chlipala committed
112
(** This one is just so darned useful, let's add it as a hint here. *)
113
Hint Rewrite app_ass.
114

Adam Chlipala's avatar
Adam Chlipala committed
115
(** Devious marker predicate to use for encoding state within proof goals *)
116 117
Definition done (T : Type) (x : T) := True.

Adam Chlipala's avatar
Adam Chlipala committed
118 119
(** Try a new instantiation of a universally quantified fact, proved by [e].
   * [trace] is an accumulator recording which instantiations we choose. *)
120
Ltac inster e trace :=
Adam Chlipala's avatar
Adam Chlipala committed
121
  (** Does [e] have any quantifiers left? *)
122 123
  match type of e with
    | forall x : _, _ =>
Adam Chlipala's avatar
Adam Chlipala committed
124
      (** Yes, so let's pick the first context variable of the right type. *)
125 126 127 128 129 130
      match goal with
        | [ H : _ |- _ ] =>
          inster (e H) (trace, H)
        | _ => fail 2
      end
    | _ =>
Adam Chlipala's avatar
Adam Chlipala committed
131
      (** No more quantifiers, so now we check if the trace we computed was already used. *)
132 133
      match trace with
        | (_, _) =>
Adam Chlipala's avatar
Adam Chlipala committed
134
          (** We only reach this case if the trace is nonempty, ensuring that [inster] fails if no progress can be made. *)
135
          match goal with
Adam Chlipala's avatar
Adam Chlipala committed
136 137 138
            | [ H : done (trace, _) |- _ ] =>
              (** Uh oh, found a record of this trace in the context!  Abort to backtrack to try another trace. *)
              fail 1
139
            | _ =>
Adam Chlipala's avatar
Adam Chlipala committed
140
              (** What is the type of the proof [e] now? *)
141 142 143
              let T := type of e in
                match type of T with
                  | Prop =>
Adam Chlipala's avatar
Adam Chlipala committed
144
                    (** [e] should be thought of as a proof, so let's add it to the context, and also add a new marker hypothesis recording our choice of trace. *)
145
                    generalize e; intro;
Adam Chlipala's avatar
Adam Chlipala committed
146
                      assert (done (trace, tt)) by constructor
147
                  | _ =>
Adam Chlipala's avatar
Adam Chlipala committed
148
                    (** [e] is something beside a proof.  Better make sure no element of our current trace was generated by a previous call to [inster], or we might get stuck in an infinite loop!  (We store previous [inster] terms in second positions of tuples used as arguments to [done] in hypotheses.  Proofs instantiated by [inster] merely use [tt] in such positions.) *)
149 150 151 152 153
                    all ltac:(fun X =>
                      match goal with
                        | [ H : done (_, X) |- _ ] => fail 1
                        | _ => idtac
                      end) trace;
Adam Chlipala's avatar
Adam Chlipala committed
154
                    (** Pick a new name for our new instantiation. *)
155
                    let i := fresh "i" in (pose (i := e);
Adam Chlipala's avatar
Adam Chlipala committed
156
                      assert (done (trace, i)) by constructor)
157 158 159 160 161
                end
          end
      end
  end.

Adam Chlipala's avatar
Adam Chlipala committed
162
(** After a round of application with the above, we will have a lot of junk [done] markers to clean up; hence this tactic. *)
163 164 165 166 167
Ltac un_done :=
  repeat match goal with
           | [ H : done _ |- _ ] => clear H
         end.

168 169
Require Import JMeq.

Adam Chlipala's avatar
Adam Chlipala committed
170 171 172
(** A more parameterized version of the famous [crush].  Extra arguments are:
   * - A tuple-list of lemmas we try [inster]-ing 
   * - A tuple-list of predicates we try inversion for *)
173
Ltac crush' lemmas invOne :=
Adam Chlipala's avatar
Adam Chlipala committed
174 175 176 177 178 179 180 181 182 183 184 185 186 187 188 189 190 191 192 193 194 195 196 197 198 199 200 201 202 203 204
  (** A useful combination of standard automation *)
  let sintuition := simpl in *; intuition; try subst;
    repeat (simplHyp invOne; intuition; try subst); try congruence in

  (** A fancier version of [rewriter] from above, which uses [crush'] to discharge side conditions *)
  let rewriter := autorewrite with core in *;
    repeat (match goal with
              | [ H : ?P |- _ ] =>
                match P with
                  | context[JMeq] => fail 1 (** JMeq is too fancy to deal with here. *)
                  | _ => rewrite H by crush' lemmas invOne
                end
            end; autorewrite with core in *) in

  (** Now the main sequence of heuristics: *)
    (sintuition; rewriter;
      match lemmas with
        | false => idtac (** No lemmas?  Nothing to do here *)
        | _ =>
          (** Try a loop of instantiating lemmas... *)
          repeat ((app ltac:(fun L => inster L L) lemmas
          (** ...or instantiating hypotheses... *)
            || appHyps ltac:(fun L => inster L L));
          (** ...and then simplifying hypotheses. *)
          repeat (simplHyp invOne; intuition)); un_done
      end;
      sintuition; rewriter; sintuition;
      (** End with a last attempt to prove an arithmetic fact with [omega], or prove any sort of fact in a context that is contradictory by reasoning that [omega] can do. *)
      try omega; try (elimtype False; omega)).

(** [crush] instantiates [crush'] with the simplest possible parameters. *)
205
Ltac crush := crush' false fail.
206

Adam Chlipala's avatar
Adam Chlipala committed
207 208
(** * Wrap Program's [dependent destruction] in a slightly more pleasant form *)

209
Require Import Program.Equality.
210

Adam Chlipala's avatar
Adam Chlipala committed
211 212
(** Run [dependent destruction] on [E] and look for opportunities to simplify the result.
   The weird introduction of [x] helps get around limitations of [dependent destruction], in terms of which sorts of arguments it will accept (e.g., variables bound to hypotheses within Ltac [match]es). *)
213
Ltac dep_destruct E :=
214 215 216
  let x := fresh "x" in
    remember E as x; simpl in x; dependent destruction x;
      try match goal with
217
            | [ H : _ = E |- _ ] => try rewrite <- H in *; clear H
218
          end.
Adam Chlipala's avatar
Adam Chlipala committed
219

Adam Chlipala's avatar
Adam Chlipala committed
220
(** Nuke all hypotheses that we can get away with, without invalidating the goal statement. *)
Adam Chlipala's avatar
Adam Chlipala committed
221 222 223 224
Ltac clear_all :=
  repeat match goal with
           | [ H : _ |- _ ] => clear H
         end.
Adam Chlipala's avatar
Adam Chlipala committed
225

Adam Chlipala's avatar
Adam Chlipala committed
226 227
(** Instantiate a quantifier in a hypothesis [H] with value [v], or, if [v] doesn't have the right type, with a new unification variable.
   * Also prove the lefthand sides of any implications that this exposes, simplifying [H] to leave out those implications. *)
Adam Chlipala's avatar
Adam Chlipala committed
228
Ltac guess v H :=
Adam Chlipala's avatar
Adam Chlipala committed
229 230 231 232 233 234
  repeat match type of H with
           | forall x : ?T, _ =>
             match type of T with
               | Prop =>
                 (let H' := fresh "H'" in
                   assert (H' : T); [
Adam Chlipala's avatar
Adam Chlipala committed
235
                     solve [ eauto 6 ]
Adam Chlipala's avatar
Adam Chlipala committed
236
                     | specialize (H H'); clear H' ])
Adam Chlipala's avatar
Adam Chlipala committed
237 238
                 || fail 1
               | _ =>
Adam Chlipala's avatar
Adam Chlipala committed
239
                 specialize (H v)
Adam Chlipala's avatar
Adam Chlipala committed
240
                 || let x := fresh "x" in
Adam Chlipala's avatar
Adam Chlipala committed
241
                   evar (x : T);
Adam Chlipala's avatar
Adam Chlipala committed
242 243
                   let x' := eval unfold x in x in
                     clear x; specialize (H x')
Adam Chlipala's avatar
Adam Chlipala committed
244 245 246
             end
         end.

Adam Chlipala's avatar
Adam Chlipala committed
247
(** Version of [guess] that leaves the original [H] intact *)
Adam Chlipala's avatar
Adam Chlipala committed
248
Ltac guessKeep v H :=
Adam Chlipala's avatar
Adam Chlipala committed
249
  let H' := fresh "H'" in
Adam Chlipala's avatar
Adam Chlipala committed
250
    generalize H; intro H'; guess v H'.