aboutsummaryrefslogtreecommitdiffhomepage
path: root/checker/safe_typing.ml
diff options
context:
space:
mode:
authorGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2015-07-10 01:13:59 +0200
committerGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2015-07-10 19:18:41 +0200
commit9c732a5c878bac2592cb397aca3d17cfefdcd023 (patch)
tree7defb39c88bdf0d163ca323955d11f1a50d2367d /checker/safe_typing.ml
parent591e7e484d544e958595a0fb784336ae050a9c74 (diff)
Option -type-in-type: added support in checker and making it contaminating
in vo files (this was not done yet in 24d0027f0 and 090fffa57b). Reused field "engagement" to carry information about both impredicativity of set and type in type. For the record: maybe some further checks to do around the sort of the inductive types in coqchk?
Diffstat (limited to 'checker/safe_typing.ml')
-rw-r--r--checker/safe_typing.ml22
1 files changed, 15 insertions, 7 deletions
diff --git a/checker/safe_typing.ml b/checker/safe_typing.ml
index 810d6e0b6..dd9482313 100644
--- a/checker/safe_typing.ml
+++ b/checker/safe_typing.ml
@@ -32,13 +32,21 @@ let full_add_module dp mb univs digest =
let env = Modops.add_module mb env in
genv := add_digest env dp digest
-(* Check that the engagement expected by a library matches the initial one *)
-let check_engagement env c =
- match engagement env, c with
- | Some ImpredicativeSet, Some ImpredicativeSet -> ()
- | _, None -> ()
- | _, Some ImpredicativeSet ->
- error "Needs option -impredicative-set"
+(* Check that the engagement expected by a library extends the initial one *)
+let check_engagement env (expected_impredicative_set,expected_type_in_type) =
+ let impredicative_set,type_in_type = Environ.engagement env in
+ begin
+ match impredicative_set, expected_impredicative_set with
+ | PredicativeSet, ImpredicativeSet ->
+ Errors.error "Needs option -impredicative-set."
+ | _ -> ()
+ end;
+ begin
+ match type_in_type, expected_type_in_type with
+ | StratifiedType, TypeInType ->
+ Errors.error "Needs option -type-in-type."
+ | _ -> ()
+ end
(* Libraries = Compiled modules *)