aboutsummaryrefslogtreecommitdiffhomepage
path: root/theories/Program
diff options
context:
space:
mode:
authorGravatar msozeau <msozeau@85f007b7-540e-0410-9357-904b9bb8a0f7>2011-10-17 09:42:55 +0000
committerGravatar msozeau <msozeau@85f007b7-540e-0410-9357-904b9bb8a0f7>2011-10-17 09:42:55 +0000
commit0ec22639868500d1b3a3755f64d6f002775530a8 (patch)
treeb40603f6bfa70416137f6fd95ad697850a483252 /theories/Program
parentffaa47e6890d7aada9be1576edf26399ae9e1b50 (diff)
Fix bug #2456 and wrong unfolding of lets in the goal due to [unfold] doing zeta reductions.
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@14564 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'theories/Program')
-rw-r--r--theories/Program/Equality.v18
1 files changed, 9 insertions, 9 deletions
diff --git a/theories/Program/Equality.v b/theories/Program/Equality.v
index c41b8bffd..06ff7cd10 100644
--- a/theories/Program/Equality.v
+++ b/theories/Program/Equality.v
@@ -28,7 +28,7 @@ Hint Extern 10 => is_ground_goal ; progress exfalso : exfalso.
Definition block {A : Type} (a : A) := a.
Ltac block_goal := match goal with [ |- ?T ] => change (block T) end.
-Ltac unblock_goal := unfold block in *.
+Ltac unblock_goal := cbv beta delta [block].
(** Notation for heterogenous equality. *)
@@ -411,26 +411,26 @@ Ltac depind id := do_depind ltac:(fun hyp => do_ind hyp) id.
(** A variant where generalized variables should be given by the user. *)
-Ltac do_depelim' tac H :=
- (try intros until H) ; block_goal ; generalize_eqs H ; tac H ; simplify_dep_elim ;
+Ltac do_depelim' rev tac H :=
+ (try intros until H) ; block_goal ; rev H ; generalize_eqs H ; tac H ; simplify_dep_elim ;
simplify_IH_hyps ; unblock_goal.
(** Calls [destruct] on the generalized hypothesis, results should be similar to inversion.
By default, we don't try to generalize the hyp by its variable indices. *)
Tactic Notation "dependent" "destruction" ident(H) :=
- do_depelim' ltac:(fun hyp => do_case hyp) H.
+ do_depelim' ltac:(fun hyp => idtac) ltac:(fun hyp => do_case hyp) H.
Tactic Notation "dependent" "destruction" ident(H) "using" constr(c) :=
- do_depelim' ltac:(fun hyp => destruct hyp using c) H.
+ do_depelim' ltac:(fun hyp => idtac) ltac:(fun hyp => destruct hyp using c) H.
(** This tactic also generalizes the goal by the given variables before the elimination. *)
Tactic Notation "dependent" "destruction" ident(H) "generalizing" ne_hyp_list(l) :=
- do_depelim' ltac:(fun hyp => revert l ; do_case hyp) H.
+ do_depelim' ltac:(fun hyp => revert l) ltac:(fun hyp => do_case hyp) H.
Tactic Notation "dependent" "destruction" ident(H) "generalizing" ne_hyp_list(l) "using" constr(c) :=
- do_depelim' ltac:(fun hyp => revert l ; destruct hyp using c) H.
+ do_depelim' ltac:(fun hyp => revert l) ltac:(fun hyp => destruct hyp using c) H.
(** Then we have wrappers for usual calls to induction. One can customize the induction tactic by
writting another wrapper calling do_depelim. We suppose the hyp has to be generalized before
@@ -445,7 +445,7 @@ Tactic Notation "dependent" "induction" ident(H) "using" constr(c) :=
(** This tactic also generalizes the goal by the given variables before the induction. *)
Tactic Notation "dependent" "induction" ident(H) "generalizing" ne_hyp_list(l) :=
- do_depelim' ltac:(fun hyp => generalize l ; clear l ; do_ind hyp) H.
+ do_depelim' ltac:(fun hyp => revert l) ltac:(fun hyp => do_ind hyp) H.
Tactic Notation "dependent" "induction" ident(H) "generalizing" ne_hyp_list(l) "using" constr(c) :=
- do_depelim' ltac:(fun hyp => generalize l ; clear l ; induction hyp using c) H.
+ do_depelim' ltac:(fun hyp => revert l) ltac:(fun hyp => induction hyp using c) H.