Mercurial > cpdt > repo
annotate src/Tactics.v @ 50:a21eb02cc7c6
Recursive predicates
author | Adam Chlipala <adamc@hcoop.net> |
---|---|
date | Sun, 28 Sep 2008 11:57:15 -0400 |
parents | 95e24b629ad9 |
children | 9ceee967b1fc |
rev | line source |
---|---|
adamc@2 | 1 (* Copyright (c) 2008, Adam Chlipala |
adamc@2 | 2 * |
adamc@2 | 3 * This work is licensed under a |
adamc@2 | 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0 |
adamc@2 | 5 * Unported License. |
adamc@2 | 6 * The license text is available at: |
adamc@2 | 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/ |
adamc@2 | 8 *) |
adamc@2 | 9 |
adamc@2 | 10 Require Import List. |
adamc@2 | 11 |
adamc@38 | 12 Require Omega. |
adamc@38 | 13 |
adamc@2 | 14 |
adamc@50 | 15 Ltac simplHyp := |
adamc@50 | 16 match goal with |
adamc@50 | 17 | [ H : S _ = S _ |- _ ] => injection H; clear H; intros; subst |
adamc@50 | 18 end. |
adamc@50 | 19 |
adamc@2 | 20 Ltac rewriteHyp := |
adamc@2 | 21 match goal with |
adamc@2 | 22 | [ H : _ |- _ ] => rewrite H |
adamc@2 | 23 end. |
adamc@2 | 24 |
adamc@2 | 25 Ltac rewriterP := repeat (rewriteHyp; autorewrite with cpdt in *). |
adamc@2 | 26 |
adamc@2 | 27 Ltac rewriter := autorewrite with cpdt in *; rewriterP. |
adamc@2 | 28 |
adamc@2 | 29 Hint Rewrite app_ass : cpdt. |
adamc@2 | 30 |
adamc@50 | 31 Ltac sintuition := simpl in *; intuition; try simplHyp. |
adamc@2 | 32 |
adamc@38 | 33 Ltac crush := sintuition; rewriter; sintuition; try omega. |