Skip to content

Commit

Permalink
Backport PR #12759: [vernac] refine check for unresolved evars
Browse files Browse the repository at this point in the history
  • Loading branch information
Zimmi48 committed Sep 8, 2020
2 parents c5af27f + 64ea7c0 commit f5f684f
Show file tree
Hide file tree
Showing 4 changed files with 43 additions and 10 deletions.
20 changes: 16 additions & 4 deletions test-suite/output/RecordMissingField.out
Original file line number Diff line number Diff line change
@@ -1,4 +1,16 @@
File "stdin", line 8, characters 5-22:
Error: Cannot infer field y2p of record point2d in environment:
p : point2d

The command has indeed failed with message:
The following term contains unresolved implicit arguments:
(fun p : point2d => {| x2p := x2p p + 1; y2p := ?y2p |})
More precisely:
- ?y2p: Cannot infer field y2p of record point2d in environment:
p : point2d
The command has indeed failed with message:
The following term contains unresolved implicit arguments:
(fun p : point2d => {| x2p := x2p p + (fun n : nat => ?n) 1; y2p := ?y2p |})
More precisely:
- ?n: Cannot infer this placeholder of type "nat" in
environment:
p : point2d
n : nat
- ?y2p: Cannot infer field y2p of record point2d in environment:
p : point2d
8 changes: 6 additions & 2 deletions test-suite/output/RecordMissingField.v
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,10 @@ should contain missing field, and the inferred type of the record **)

Record point2d := mkPoint { x2p: nat; y2p: nat }.


Definition increment_x (p: point2d) : point2d :=
Fail Definition increment_x (p: point2d) : point2d :=
{| x2p := x2p p + 1; |}.

(* Here there is also an unresolved implicit, which should give an
understadable error as well *)
Fail Definition increment_x (p: point2d) : point2d :=
{| x2p := x2p p + (fun n => _) 1; |}.
22 changes: 20 additions & 2 deletions vernac/declare.ml
Original file line number Diff line number Diff line change
Expand Up @@ -1011,13 +1011,31 @@ let declare_assumption ?fix_exn ~name ~scope ~hook ~impargs ~uctx pe =
Exninfo.iraise exn

(* Preparing proof entries *)
let error_unresolved_evars env sigma t evars =
let pr_unresolved_evar e =
hov 2 (str"- " ++ Printer.pr_existential_key sigma e ++ str ": " ++
Himsg.explain_pretype_error env sigma
(Pretype_errors.UnsolvableImplicit (e,None)))
in
CErrors.user_err (hov 0 begin
str "The following term contains unresolved implicit arguments:"++ fnl () ++
str " " ++ Printer.pr_econstr_env env sigma t ++ fnl () ++
str "More precisely: " ++ fnl () ++
v 0 (prlist_with_sep cut pr_unresolved_evar (Evar.Set.elements evars))
end)

let check_evars_are_solved env sigma t =
let t = EConstr.of_constr t in
let evars = Evarutil.undefined_evars_of_term sigma t in
if not (Evar.Set.is_empty evars) then error_unresolved_evars env sigma t evars

let prepare_definition ?opaque ?inline ?fix_exn ~poly ~udecl ~types ~body sigma =
let env = Global.env () in
Pretyping.check_evars_are_solved ~program_mode:false env sigma;
let sigma, (body, types) = Evarutil.finalize ~abort_on_undefined_evars:true
let sigma, (body, types) = Evarutil.finalize ~abort_on_undefined_evars:false
sigma (fun nf -> nf body, Option.map nf types)
in
Option.iter (check_evars_are_solved env sigma) types;
check_evars_are_solved env sigma body;
let univs = Evd.check_univ_decl ~poly sigma udecl in
let entry = definition_entry ?fix_exn ?opaque ?inline ?types ~univs body in
let uctx = Evd.evar_universe_context sigma in
Expand Down
3 changes: 1 addition & 2 deletions vernac/declare.mli
Original file line number Diff line number Diff line change
Expand Up @@ -323,8 +323,7 @@ val declare_entry
normalized w.r.t. the passed [evar_map] [sigma]. Universes should
be handled properly, including minimization and restriction. Note
that [sigma] is checked for unresolved evars, thus you should be
careful not to submit open terms or evar maps with stale,
unresolved existentials *)
careful not to submit open terms *)
val declare_definition
: name:Id.t
-> scope:locality
Expand Down

0 comments on commit f5f684f

Please sign in to comment.