Skip to content

Commit

Permalink
Fixing #4782 (a typing error not captured when dealing with bindings).
Browse files Browse the repository at this point in the history
Trying to now catch all unification errors, but without a clear view
at whether some errors could be tolerated at the point of checking the
type of the binding.
  • Loading branch information
herbelin committed Jun 11, 2016
1 parent d095d43 commit f9695eb
Show file tree
Hide file tree
Showing 2 changed files with 9 additions and 1 deletion.
2 changes: 1 addition & 1 deletion proofs/clenv.ml
Original file line number Diff line number Diff line change
Expand Up @@ -459,7 +459,7 @@ let clenv_unify_binding_type clenv c t u =
let evd,c = w_coerce_to_type (cl_env clenv) clenv.evd c t u in
TypeProcessed, { clenv with evd = evd }, c
with
| PretypeError (_,_,ActualTypeNotCoercible (_,_,NotClean _)) as e ->
| PretypeError (_,_,ActualTypeNotCoercible _) as e ->
raise e
| e when precatchable_exception e ->
TypeNotProcessed, clenv, c
Expand Down
8 changes: 8 additions & 0 deletions test-suite/bugs/closed/4782.v
Original file line number Diff line number Diff line change
@@ -0,0 +1,8 @@
(* About typing of with bindings *)

Record r : Type := mk_r { type : Type; cond : type -> Prop }.

Inductive p : Prop := consp : forall (e : r) (x : type e), cond e x -> p.

Goal p.
Fail apply consp with (fun _ : bool => mk_r unit (fun x => True)) nil.

0 comments on commit f9695eb

Please sign in to comment.