aboutsummaryrefslogtreecommitdiff
path: root/kernel/typeops.ml
diff options
context:
space:
mode:
authorherbelin2013-02-28 07:33:37 +0000
committerherbelin2013-02-28 07:33:37 +0000
commitc43f9128237ac16fa0d7741744e3944ca72e7475 (patch)
tree9d36fa1cb3364d16c00e506d786822303b6a027c /kernel/typeops.ml
parent87bd13c7a6552f33782e0e69ef705b356a2cf741 (diff)
More informative error when a global reference is used in a context of
local variables which is a different from the one of its definition. E.g.: Section A. Variable n:nat. Definition c:=n. Goal True. clear n. Check c. [I'm however unsure that "n" should not continue to be accessible via some global (qualified) name, even after the "clear n".] git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@16256 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'kernel/typeops.ml')
-rw-r--r--kernel/typeops.ml33
1 files changed, 10 insertions, 23 deletions
diff --git a/kernel/typeops.ml b/kernel/typeops.ml
index efc7a4ee8b..18e6fec791 100644
--- a/kernel/typeops.ml
+++ b/kernel/typeops.ml
@@ -91,33 +91,20 @@ let judge_of_variable env id =
(* Management of context of variables. *)
-(* Checks if a context of variable can be instantiated by the
+(* Checks if a context of variables can be instantiated by the
variables of the current env *)
(* TODO: check order? *)
-let check_hyps_inclusion env sign =
+let check_hyps_inclusion env c sign =
Sign.fold_named_context
(fun (id,_,ty1) () ->
- let ty2 = named_type id env in
- if not (eq_constr ty2 ty1) then
- error "types do not match")
+ try
+ let ty2 = named_type id env in
+ if not (eq_constr ty2 ty1) then raise Exit
+ with Not_found | Exit ->
+ error_reference_variables env id c)
sign
~init:()
-
-let check_args env c hyps =
- try check_hyps_inclusion env hyps
- with UserError _ | Not_found ->
- error_reference_variables env c
-
-
-(* Checks if the given context of variables [hyps] is included in the
- current context of [env]. *)
-(*
-let check_hyps id env hyps =
- let hyps' = named_context env in
- if not (hyps_inclusion env hyps hyps') then
- error_reference_variables env id
-*)
(* Instantiation of terms on real arguments. *)
(* Make a type polymorphic if an arity *)
@@ -162,7 +149,7 @@ let type_of_constant env cst =
let judge_of_constant_knowing_parameters env cst jl =
let c = mkConst cst in
let cb = lookup_constant cst env in
- let _ = check_args env c cb.const_hyps in
+ let _ = check_hyps_inclusion env c cb.const_hyps in
let paramstyp = Array.map (fun j -> j.uj_type) jl in
let t = type_of_constant_knowing_parameters env cb.const_type paramstyp in
make_judge c t
@@ -308,7 +295,7 @@ let judge_of_cast env cj k tj =
let judge_of_inductive_knowing_parameters env ind jl =
let c = mkInd ind in
let (mib,mip) = lookup_mind_specif env ind in
- check_args env c mib.mind_hyps;
+ check_hyps_inclusion env c mib.mind_hyps;
let paramstyp = Array.map (fun j -> j.uj_type) jl in
let t = Inductive.type_of_inductive_knowing_parameters env mip paramstyp in
make_judge c t
@@ -323,7 +310,7 @@ let judge_of_constructor env c =
let _ =
let ((kn,_),_) = c in
let mib = lookup_mind kn env in
- check_args env constr mib.mind_hyps in
+ check_hyps_inclusion env constr mib.mind_hyps in
let specif = lookup_mind_specif env (inductive_of_constructor c) in
make_judge constr (type_of_constructor c specif)