aboutsummaryrefslogtreecommitdiffhomepage
path: root/kernel
diff options
context:
space:
mode:
authorGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2006-10-17 16:39:46 +0000
committerGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2006-10-17 16:39:46 +0000
commit1853b9a3fa77f29f41bc3cf131d691a7690e258b (patch)
tree7e3c697f2204988d44d95b8b1fd47e7d3da2f6ae /kernel
parentb967f487c538a119a51a95f3669b5f6937a69357 (diff)
Clarification des contraintes sur le contexte de paramètres des
inductifs dans le test de sous-typage (exigence du même nombre d'arguments uniformes attendus mais pas d'exigence spéciale sur les définitions locales du contexte à partir du moment où les types et constructeurs sont convertibles quand généralisés par rapport au contexte de paramètres) git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@9247 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'kernel')
-rw-r--r--kernel/subtyping.ml9
1 files changed, 6 insertions, 3 deletions
diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml
index 0acd37da4..c94219316 100644
--- a/kernel/subtyping.ml
+++ b/kernel/subtyping.ml
@@ -94,7 +94,7 @@ let check_inductive cst env msid1 l info1 mib2 spec2 =
(* listrec ignored *)
(* finite done *)
(* nparams done *)
- (* params_ctxt done *)
+ (* params_ctxt done because part of the inductive types *)
(* Don't check the sort of the type if polymorphic *)
let cst = check_conv cst conv env (type_of_inductive (mib1,p1)) (type_of_inductive (mib2,p2))
in
@@ -114,9 +114,12 @@ let check_inductive cst env msid1 l info1 mib2 spec2 =
assert (Array.length mib1.mind_packets >= 1
&& Array.length mib2.mind_packets >= 1);
- (* TODO: we should allow renaming of parameters at least ! *)
+ (* Check that the expected numbers of uniform parameters are the same *)
+ (* No need to check the contexts of parameters: it is checked *)
+ (* at the time of checking the inductive arities in check_packet. *)
+ (* Notice that we don't expect the local definitions to match: only *)
+ (* the inductive types and constructors types have to be convertible *)
check (fun mib -> mib.mind_nparams);
- check (fun mib -> mib.mind_params_ctxt);
begin
match mib2.mind_equiv with