aboutsummaryrefslogtreecommitdiff
path: root/pretyping
diff options
context:
space:
mode:
authorherbelin2013-02-28 07:33:41 +0000
committerherbelin2013-02-28 07:33:41 +0000
commit9a431a1b43dead6d692c942986a25f0f8986465a (patch)
tree5d11bfeee35709817dc664168fa3be1dc42f5cb1 /pretyping
parentc43f9128237ac16fa0d7741744e3944ca72e7475 (diff)
Repairing r16205: errors raised by check_evar_instance were no longer
trapped by solve_simple_eqn. git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@16257 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'pretyping')
-rw-r--r--pretyping/evarconv.ml5
-rw-r--r--pretyping/evarsolve.ml16
-rw-r--r--pretyping/evarsolve.mli7
-rw-r--r--pretyping/pretype_errors.ml2
-rw-r--r--pretyping/pretype_errors.mli2
5 files changed, 17 insertions, 15 deletions
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index 1258e4a09e..6d6aa6d5d6 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -854,8 +854,9 @@ let rec solve_unconstrained_evars_with_canditates evd =
match reconsider_conv_pbs conv_algo evd with
| Success evd -> solve_unconstrained_evars_with_canditates evd
| UnifFailure _ -> aux l
- with e when Pretype_errors.precatchable_exception e ->
- aux l in
+ with
+ | IllTypedInstance _ as e
+ | e when Pretype_errors.precatchable_exception e -> aux l in
(* List.rev is there to favor most dependent solutions *)
(* and favor progress when used with the refine tactics *)
let evd = aux (List.rev l) in
diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml
index f29d982c2d..01878024b3 100644
--- a/pretyping/evarsolve.ml
+++ b/pretyping/evarsolve.ml
@@ -1034,6 +1034,8 @@ type conv_fun =
type conv_fun_bool =
env -> evar_map -> conv_pb -> constr -> constr -> bool
+exception IllTypedInstance of env * types * types
+
let check_evar_instance evd evk1 body conv_algo =
let evi = Evd.find evd evk1 in
let evenv = evar_env evi in
@@ -1044,10 +1046,7 @@ let check_evar_instance evd evk1 body conv_algo =
in
match conv_algo evenv evd Reduction.CUMUL ty evi.evar_concl with
| Success evd -> evd
- | UnifFailure (evd,d) ->
- (* TODO: use the error? *)
- user_err_loc (fst (evar_source evk1 evd),"",
- str "Unable to find a well-typed instantiation")
+ | UnifFailure _ -> raise (IllTypedInstance (evenv,ty,evi.evar_concl))
(* Solve pbs ?e[t1..tn] = ?e[u1..un] which arise often in fixpoint
* definitions. We try to unify the ti with the ui pairwise. The pairs
@@ -1128,7 +1127,6 @@ exception NotInvertibleUsingOurAlgorithm of constr
exception NotEnoughInformationToProgress of (Id.t * evar_projection) list
exception OccurCheckIn of evar_map * constr
exception MetaOccurInBodyInternal
-exception InstanceNotSameTypeInternal
let rec invert_definition conv_algo choose env evd (evk,argsv as ev) rhs =
let aliases = make_alias_map env in
@@ -1310,8 +1308,8 @@ and evar_define conv_algo ?(choose=false) env evd (evk,argsv as ev) rhs =
str "----> " ++ int ev ++ str " := " ++
print_constr body);
raise e in*)
- let evd' = Evd.define evk body evd' in
- check_evar_instance evd' evk body conv_algo
+ let evd' = check_evar_instance evd' evk body conv_algo in
+ Evd.define evk body evd'
with
| NotEnoughInformationToProgress sols ->
postpone_non_unique_projection env evd ev sols rhs
@@ -1391,6 +1389,6 @@ let solve_simple_eqn conv_algo ?(choose=false) env evd (pbty,(evk1,args1 as ev1)
UnifFailure (evd,OccurCheck (evk1,rhs))
| MetaOccurInBodyInternal ->
UnifFailure (evd,MetaOccurInBody evk1)
- | InstanceNotSameTypeInternal ->
- UnifFailure (evd,InstanceNotSameType evk1)
+ | IllTypedInstance (env,t,u) ->
+ UnifFailure (evd,InstanceNotSameType (evk1,env,t,u))
diff --git a/pretyping/evarsolve.mli b/pretyping/evarsolve.mli
index e34332c9b0..3e769e02de 100644
--- a/pretyping/evarsolve.mli
+++ b/pretyping/evarsolve.mli
@@ -54,5 +54,8 @@ val is_unification_pattern : env * int -> evar_map -> constr -> constr list ->
val solve_pattern_eqn : env -> constr list -> constr -> constr
-val check_evar_instance : evar_map -> existential_key -> constr -> conv_fun ->
- evar_map
+exception IllTypedInstance of env * types * types
+
+(* May raise IllTypedInstance if types are not convertible *)
+val check_evar_instance :
+ evar_map -> existential_key -> constr -> conv_fun -> evar_map
diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml
index 91d24ec698..72cc6502a2 100644
--- a/pretyping/pretype_errors.ml
+++ b/pretyping/pretype_errors.ml
@@ -22,7 +22,7 @@ type unification_error =
| NoCanonicalStructure
| ConversionFailed of env * constr * constr
| MetaOccurInBody of existential_key
- | InstanceNotSameType of existential_key
+ | InstanceNotSameType of existential_key * env * types * types
| UnifUnivInconsistency
type pretype_error =
diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli
index ee79651316..69994531d3 100644
--- a/pretyping/pretype_errors.mli
+++ b/pretyping/pretype_errors.mli
@@ -25,7 +25,7 @@ type unification_error =
| NoCanonicalStructure
| ConversionFailed of env * constr * constr
| MetaOccurInBody of existential_key
- | InstanceNotSameType of existential_key
+ | InstanceNotSameType of existential_key * env * types * types
| UnifUnivInconsistency
type pretype_error =