aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--pretyping/pretyping.ml2
-rw-r--r--test-suite/success/evars.v5
2 files changed, 6 insertions, 1 deletions
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 947469ca0..e68a25a87 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -1118,7 +1118,7 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update =
with Not_found ->
try
let (n,_,t') = lookup_rel_id id (rel_context env) in
- if is_conv env.ExtraEnv.env !evdref t t' then mkRel n, update else raise Not_found
+ if is_conv env.ExtraEnv.env !evdref t (lift n t') then mkRel n, update else raise Not_found
with Not_found ->
try
let t' = env |> lookup_named id |> NamedDecl.get_type in
diff --git a/test-suite/success/evars.v b/test-suite/success/evars.v
index 5b13f35d5..253b48e4d 100644
--- a/test-suite/success/evars.v
+++ b/test-suite/success/evars.v
@@ -421,3 +421,8 @@ Goal exists n : nat, n = n -> True.
eexists.
set (H := _ = _).
Abort.
+
+(* Check interpretation of default evar instance in pretyping *)
+(* (reported as bug #7356) *)
+
+Check fun (P : nat -> Prop) (x:nat) (h:P x) => exist _ ?[z] (h : P ?z).