aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Matthieu Sozeau <mattam@mattam.org>2016-01-07 17:20:51 +0100
committerGravatar Matthieu Sozeau <mattam@mattam.org>2016-01-07 17:23:07 +0100
commite309d8119cd82bdf1216751fb076d438782fb60f (patch)
tree6758e811c1520ce40b9f9ddf11574ca0c72805d6
parent905e3dd364e8be20771c39393e7e114f2e4b8cd8 (diff)
Fix bug #4480: progress was not checked for setoid_rewrite.
Also ensure we stay compatible with 8.4: progress could now be made simply because of beta redexes in the goal.
-rw-r--r--tactics/rewrite.ml15
-rw-r--r--test-suite/bugs/closed/4480.v12
2 files changed, 20 insertions, 7 deletions
diff --git a/tactics/rewrite.ml b/tactics/rewrite.ml
index a230ea251..6d61879e8 100644
--- a/tactics/rewrite.ml
+++ b/tactics/rewrite.ml
@@ -1526,7 +1526,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma strat clause =
let treat sigma res =
match res with
| None -> newfail 0 (str "Nothing to rewrite")
- | Some None -> Proofview.tclUNIT ()
+ | Some None -> newfail 0 (str"Failed to progress")
| Some (Some res) ->
let (undef, prf, newt) = res in
let fold ev _ accu = if Evd.mem sigma ev then accu else ev :: accu in
@@ -1596,12 +1596,13 @@ let tactic_init_setoid () =
(** Setoid rewriting when called with "rewrite_strat" *)
let cl_rewrite_clause_strat strat clause =
tclTHEN (tactic_init_setoid ())
- (fun gl ->
- try Proofview.V82.of_tactic (cl_rewrite_clause_newtac strat clause) gl
- with RewriteFailure e ->
- errorlabstrm "" (str"setoid rewrite failed: " ++ e)
- | Refiner.FailError (n, pp) ->
- tclFAIL n (str"setoid rewrite failed: " ++ Lazy.force pp) gl)
+ (tclWEAK_PROGRESS
+ (fun gl ->
+ try Proofview.V82.of_tactic (cl_rewrite_clause_newtac strat clause) gl
+ with RewriteFailure e ->
+ errorlabstrm "" (str"setoid rewrite failed: " ++ e)
+ | Refiner.FailError (n, pp) ->
+ tclFAIL n (str"setoid rewrite failed: " ++ Lazy.force pp) gl))
(** Setoid rewriting when called with "setoid_rewrite" *)
let cl_rewrite_clause l left2right occs clause gl =
diff --git a/test-suite/bugs/closed/4480.v b/test-suite/bugs/closed/4480.v
new file mode 100644
index 000000000..08a86330f
--- /dev/null
+++ b/test-suite/bugs/closed/4480.v
@@ -0,0 +1,12 @@
+Require Import Setoid.
+
+Definition proj (P Q : Prop) := P.
+
+Lemma foo (P : Prop) : proj P P = P.
+Admitted.
+Lemma trueI : True <-> True.
+Admitted.
+Goal True.
+ Fail setoid_rewrite foo.
+ Fail setoid_rewrite trueI.
+ \ No newline at end of file