aboutsummaryrefslogtreecommitdiffhomepage
path: root/theories
diff options
context:
space:
mode:
authorGravatar msozeau <msozeau@85f007b7-540e-0410-9357-904b9bb8a0f7>2008-01-30 04:21:51 +0000
committerGravatar msozeau <msozeau@85f007b7-540e-0410-9357-904b9bb8a0f7>2008-01-30 04:21:51 +0000
commit2b9f73c7e86ac718c0ce4c47d6a24ffc2d01499d (patch)
tree4036fe3c992e406c790d165b17424c3530653277 /theories
parent90b063be6b3c23a54e1d59974e09ee14f2941338 (diff)
Work on dependent induction tactic and friends, finish the test-suite example
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@10487 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'theories')
-rw-r--r--theories/Program/Equality.v42
-rw-r--r--theories/Program/Tactics.v24
2 files changed, 48 insertions, 18 deletions
diff --git a/theories/Program/Equality.v b/theories/Program/Equality.v
index 85ed18891..c570aa983 100644
--- a/theories/Program/Equality.v
+++ b/theories/Program/Equality.v
@@ -1,3 +1,4 @@
+(* -*- coq-prog-args: ("-emacs-U") -*- *)
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
@@ -13,6 +14,8 @@
Require Export ProofIrrelevance.
Require Export JMeq.
+Require Import Coq.Program.Tactics.
+
(** Notation for heterogenous equality. *)
Notation " [ x : X ] = [ y : Y ] " := (@JMeq X x Y y) (at level 0, X at next level, Y at next level).
@@ -170,27 +173,36 @@ Ltac simplify_eqs :=
simpl ; simpl_eqs ; clear_eq_ctx ; clear_refl_eqs ;
try subst ; simpl ; repeat simpl_uip ; rewrite_refl_id.
-(** A tactic to remove trivial equality guards in hypotheses. *)
+(** A tactic that tries to remove trivial equality guards in induction hypotheses coming
+ from [dependent induction]/[generalize_eqs] invocations. *)
Ltac simpl_IH_eq H :=
- let tac H' := clear H ; rename H' into H in
- let H' := fresh "H" in
- match type of H with
- | JMeq _ _ -> _ =>
- assert (H' := H (JMeq_refl _)) ; tac H'
- | _ = _ -> _ =>
- assert (H' := H (refl_equal _)) ; tac H'
- end.
+ match type of H with
+ | JMeq _ _ -> _ =>
+ refine_hyp (H (JMeq_refl _))
+ | ?x = _ -> _ =>
+ refine_hyp (H (refl_equal x))
+ | _ -> ?x = _ -> _ =>
+ refine_hyp (H _ (refl_equal x))
+ | _ -> _ -> ?x = _ -> _ =>
+ refine_hyp (H _ _ (refl_equal x))
+ | _ -> _ -> _ -> ?x = _ -> _ =>
+ refine_hyp (H _ _ _ (refl_equal x))
+ | _ -> _ -> _ -> _ -> ?x = _ -> _ =>
+ refine_hyp (H _ _ _ _ (refl_equal x))
+ | _ -> _ -> _ -> _ -> _ -> ?x = _ -> _ =>
+ refine_hyp (H _ _ _ _ _ (refl_equal x))
+ end.
Ltac simpl_IH_eqs H := repeat simpl_IH_eq H.
-Ltac simpl_IHs_eqs :=
+Ltac do_simpl_IHs_eqs :=
match goal with
- | [ H : JMeq _ _ -> _ |- _ ] => simpl_IH_eqs H
- | [ H : _ = _ -> _ |- _ ] => simpl_IH_eqs H
+ | [ H : context [ JMeq _ _ -> _ ] |- _ ] => progress (simpl_IH_eqs H)
+ | [ H : context [ _ = _ -> _ ] |- _ ] => progress (simpl_IH_eqs H)
end.
-Require Import Coq.Program.Tactics.
+Ltac simpl_IHs_eqs := repeat do_simpl_IHs_eqs.
(** The following tactics allow to do induction on an already instantiated inductive predicate
by first generalizing it and adding the proper equalities to the context, in a maner similar to
@@ -198,10 +210,10 @@ Require Import Coq.Program.Tactics.
Tactic Notation "dependent" "induction" ident(H) :=
generalize_eqs H ; clear H ; (intros until 1 || intros until H) ;
- induction H ; intros ; subst* ; try discriminates ; try simpl_IHs_eqs.
+ induction H ; intros ; subst* ; try discriminates ; simpl_IHs_eqs.
(** This tactic also generalizes the goal by the given variables before the induction. *)
Tactic Notation "dependent" "induction" ident(H) "generalizing" ne_hyp_list(l) :=
generalize_eqs H ; clear H ; (intros until 1 || intros until H) ;
- generalize l ; clear l ; induction H ; intros ; subst* ; try discriminates ; try simpl_IHs_eqs.
+ generalize l ; clear l ; induction H ; intros ; subst* ; try discriminates ; simpl_IHs_eqs.
diff --git a/theories/Program/Tactics.v b/theories/Program/Tactics.v
index ad03de2f4..bb06f37b5 100644
--- a/theories/Program/Tactics.v
+++ b/theories/Program/Tactics.v
@@ -8,7 +8,8 @@
(*i $Id$ i*)
-(** This module implements various tactics used to simplify the goals produced by Program. *)
+(** This module implements various tactics used to simplify the goals produced by Program,
+ which are also generally useful. *)
(** Destructs one pair, without care regarding naming. *)
@@ -166,6 +167,23 @@ Ltac add_hypothesis H' p :=
end
end.
+Tactic Notation "pose" constr(c) "as" ident(H) := assert(H:=c).
+
+(** A tactic to refine an hypothesis by supplying some of its arguments. *)
+
+Ltac refine_hyp c :=
+ let H' := fresh "H" in
+ let tac H := assert(H' := c) ; clear H ; rename H' into H in
+ match c with
+ | ?H _ => tac H
+ | ?H _ _ => tac H
+ | ?H _ _ _ => tac H
+ | ?H _ _ _ _ => tac H
+ | ?H _ _ _ _ _ => tac H
+ | ?H _ _ _ _ _ _ => tac H
+ | ?H _ _ _ _ _ _ _ => tac H
+ | ?H _ _ _ _ _ _ _ _ => tac H
+ end.
(** The default simplification tactic used by Program is defined by [program_simpl], sometimes [auto with *]
is overkill and slows things down, better rebind using [Obligations Tactic := tac] in this case,
@@ -175,6 +193,6 @@ Ltac program_simplify :=
simpl ; intros ; destruct_conjs ; simpl proj1_sig in * ; subst* ; try autoinjection ; try discriminates ;
try (solve [ red ; intros ; destruct_conjs ; try autoinjection ; discriminates ]).
-Ltac default_program_simpl := program_simplify ; auto with *.
+Ltac program_simpl := program_simplify ; auto with *.
-Ltac program_simpl := default_program_simpl.
+Ltac obligations_tactic := program_simpl.