Skip to content

Commit

Permalink
Fix coq#9344, coq#9348: incorrect unsafe to_constr in vnorm
Browse files Browse the repository at this point in the history
  • Loading branch information
SkySkimmer authored and maximedenes committed Apr 29, 2019
1 parent 05c5c3a commit af3673b
Show file tree
Hide file tree
Showing 3 changed files with 5 additions and 1 deletion.
2 changes: 1 addition & 1 deletion pretyping/vnorm.ml
Original file line number Diff line number Diff line change
Expand Up @@ -202,7 +202,7 @@ and nf_univ_args ~nb_univs mk env sigma stk =
and nf_evar env sigma evk stk =
let evi = try Evd.find sigma evk with Not_found -> assert false in
let hyps = Environ.named_context_of_val (Evd.evar_filtered_hyps evi) in
let concl = EConstr.Unsafe.to_constr @@ Evd.evar_concl evi in
let concl = EConstr.to_constr ~abort_on_undefined_evars:false sigma @@ Evd.evar_concl evi in
if List.is_empty hyps then
nf_stk env sigma (mkEvar (evk, [||])) concl stk
else match stk with
Expand Down
1 change: 1 addition & 0 deletions test-suite/bugs/closed/bug_9344.v
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
Compute _ I.
3 changes: 3 additions & 0 deletions test-suite/bugs/closed/bug_9348.v
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
Set Primitive Projections.
Record r {A} := R {f : A -> A}.
Compute f _ I.

0 comments on commit af3673b

Please sign in to comment.