aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--tactics/rewrite.ml18
-rw-r--r--test-suite/bugs/closed/4754.v35
2 files changed, 48 insertions, 5 deletions
diff --git a/tactics/rewrite.ml b/tactics/rewrite.ml
index c7cfc4dc7..0892dc9e6 100644
--- a/tactics/rewrite.ml
+++ b/tactics/rewrite.ml
@@ -424,14 +424,22 @@ let split_head = function
hd :: tl -> hd, tl
| [] -> assert(false)
+let eq_pb (ty, env, x, y as pb) (ty', env', x', y' as pb') =
+ pb == pb' || (ty == ty' && Constr.equal x x' && Constr.equal y y')
+
+let problem_inclusion x y =
+ List.for_all (fun pb -> List.exists (fun pb' -> eq_pb pb pb') y) x
+
let evd_convertible env evd x y =
try
- let evd = Evarconv.the_conv_x env x y evd in
(* Unfortunately, the_conv_x might say they are unifiable even if some
- unsolvable constraints remain, so we check them here *)
- let evd = Evarconv.consider_remaining_unif_problems env evd in
- let () = Evarconv.check_problems_are_solved env evd in
- Some evd
+ unsolvable constraints remain, so we check that this unification
+ does not introduce any new problem. *)
+ let _, pbs = Evd.extract_all_conv_pbs evd in
+ let evd' = Evarconv.the_conv_x env x y evd in
+ let _, pbs' = Evd.extract_all_conv_pbs evd' in
+ if evd' == evd || problem_inclusion pbs' pbs then Some evd'
+ else None
with e when Errors.noncritical e -> None
let convertible env evd x y =
diff --git a/test-suite/bugs/closed/4754.v b/test-suite/bugs/closed/4754.v
new file mode 100644
index 000000000..5bb3cd1be
--- /dev/null
+++ b/test-suite/bugs/closed/4754.v
@@ -0,0 +1,35 @@
+
+Require Import Coq.Setoids.Setoid Coq.Classes.Morphisms.
+Definition f (v : option nat) := match v with
+ | Some k => Some k
+ | None => None
+ end.
+
+Axioms F G : (option nat -> option nat) -> Prop.
+Axiom FG : forall f, f None = None -> F f = G f.
+
+Axiom admit : forall {T}, T.
+
+Existing Instance eq_Reflexive.
+
+Global Instance foo (A := nat)
+ : Proper ((pointwise_relation _ eq)
+ ==> eq ==> forall_relation (fun _ => Basics.flip Basics.impl))
+ (@option_rect A (fun _ => Prop)) | 0.
+exact admit.
+Qed.
+
+Global Instance bar (A := nat)
+ : Proper ((pointwise_relation _ eq)
+ ==> eq ==> eq ==> Basics.flip Basics.impl)
+ (@option_rect A (fun _ => Prop)) | 0.
+exact admit.
+Qed.
+
+Goal forall k, option_rect (fun _ => Prop) (fun v : nat => v = v /\ F f) True k.
+Proof.
+ intro.
+ pose proof (_ : (Proper (_ ==> eq ==> _) and)).
+ setoid_rewrite (FG _ _); [ | reflexivity.. ].
+ Undo.
+ setoid_rewrite (FG _ eq_refl). (* Error: Tactic failure: setoid rewrite failed: Nothing to rewrite. in 8.5 *) Admitted. \ No newline at end of file