aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--kernel/type_errors.ml2
-rw-r--r--kernel/type_errors.mli2
-rw-r--r--pretyping/pretype_errors.ml4
-rw-r--r--pretyping/pretype_errors.mli2
-rw-r--r--toplevel/himsg.ml6
5 files changed, 8 insertions, 8 deletions
diff --git a/kernel/type_errors.ml b/kernel/type_errors.ml
index 39223c726..269e30fca 100644
--- a/kernel/type_errors.ml
+++ b/kernel/type_errors.ml
@@ -57,7 +57,7 @@ type type_error =
| BadPattern of constructor * constr
| BadConstructor of constructor * inductive
| WrongNumargConstructor of constructor_path * int
- | WrongPredicateArity of constr * int * int
+ | WrongPredicateArity of constr * constr * constr
| NeedsInversion of constr * constr
(* Relocation error *)
| QualidNotFound of qualid
diff --git a/kernel/type_errors.mli b/kernel/type_errors.mli
index 670a9f5de..22c6bff99 100644
--- a/kernel/type_errors.mli
+++ b/kernel/type_errors.mli
@@ -61,7 +61,7 @@ type type_error =
| BadPattern of constructor * constr
| BadConstructor of constructor * inductive
| WrongNumargConstructor of constructor_path * int
- | WrongPredicateArity of constr * int * int
+ | WrongPredicateArity of constr * constr * constr
| NeedsInversion of constr * constr
(* Relocation error *)
| QualidNotFound of qualid
diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml
index 3a6897d98..dbc717fc0 100644
--- a/pretyping/pretype_errors.ml
+++ b/pretyping/pretype_errors.ml
@@ -49,8 +49,8 @@ let error_bad_constructor_loc loc k cstr ind =
let error_wrong_numarg_constructor_loc loc k c n =
raise_pretype_error (loc, k, Global.env(), WrongNumargConstructor (c,n))
-let error_wrong_predicate_arity_loc loc k env c n1 n2 =
- raise_pretype_error (loc, k, env, WrongPredicateArity (c,n1,n2))
+let error_wrong_predicate_arity_loc loc env c n1 n2 =
+ raise_pretype_error (loc, CCI, env, WrongPredicateArity (c,n1,n2))
let error_needs_inversion k env x t =
raise (TypeError (k, env, NeedsInversion (x,t)))
diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli
index 7d93dbb18..88f5889eb 100644
--- a/pretyping/pretype_errors.mli
+++ b/pretyping/pretype_errors.mli
@@ -53,7 +53,7 @@ val error_wrong_numarg_constructor_loc :
loc -> path_kind -> constructor_path -> int -> 'b
val error_wrong_predicate_arity_loc :
- loc -> path_kind -> env -> constr -> int -> int -> 'b
+ loc -> env -> constr -> constr -> constr -> 'b
val error_needs_inversion : path_kind -> env -> constr -> constr -> 'a
diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml
index 57b2c2ac8..cdd9b4ad8 100644
--- a/toplevel/himsg.ml
+++ b/toplevel/himsg.ml
@@ -308,9 +308,9 @@ let explain_wrong_numarg_of_constructor k ctx cstr n =
let explain_wrong_predicate_arity k ctx pred nondep_arity dep_arity=
let pp = prterm_env ctx pred in
[<'sTR "The elimination predicate " ; pp; 'cUT;
- 'sTR "should be of arity " ;
- 'iNT nondep_arity ; 'sTR " (for non dependent case) or " ;
- 'iNT dep_arity ; 'sTR " (for dependent case).">]
+ 'sTR "should be of arity" ; 'sPC;
+ prterm_env ctx nondep_arity ; 'sPC; 'sTR "(for non dependent case) or" ;
+ 'sPC; prterm_env ctx dep_arity ; 'sPC; 'sTR "(for dependent case).">]
let explain_needs_inversion k ctx x t =
let px = prterm_env ctx x in