From b359ef0ffad7fd1fc0e4db99fc1e38a1389802bc Mon Sep 17 00:00:00 2001 From: letouzey Date: Wed, 2 Nov 2011 18:59:57 +0000 Subject: Add type annotations around all calls to Libobject.declare_object These annotations are purely optional, but could be quite helpful when trying to understand the code, and in particular trying to trace which which data-structure may end in the libobject part of a vo. By the way, we performed some code simplifications : - in Library, a part of the REQUIRE objects was unused. - in Declaremods, we removed some checks that were marked as useless, this allows to slightly simplify the stored objects. To investigate someday : in recordops, the RECMETHODS is storing some evar_maps. This is ok for the moment, but might not be in the future (cf previous commit on auto hints). This RECMETHODS was not detected by my earlier tests : not used in the stdlib ? git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@14627 85f007b7-540e-0410-9357-904b9bb8a0f7 --- library/declare.ml | 14 +++-- library/declaremods.ml | 143 ++++++++++++++++++++++--------------------------- library/goptions.ml | 2 +- library/heads.ml | 4 +- library/impargs.ml | 6 ++- library/lib.ml | 2 +- library/library.ml | 13 +++-- 7 files changed, 94 insertions(+), 90 deletions(-) (limited to 'library') diff --git a/library/declare.ml b/library/declare.ml index 10dc35b95..0afd4bd9e 100644 --- a/library/declare.ml +++ b/library/declare.ml @@ -78,7 +78,10 @@ let discharge_variable (_,o) = match o with | Inr (id,_) -> Some (Inl (variable_constraints id)) | Inl _ -> Some o -let inVariable = +type variable_obj = + (Univ.constraints, identifier * variable_declaration) union + +let inVariable : variable_obj -> obj = declare_object { (default_object "VARIABLE") with cache_function = cache_variable; discharge_function = discharge_variable; @@ -151,7 +154,10 @@ let dummy_constant (ce,_,mk) = dummy_constant_entry,[],mk let classify_constant cst = Substitute (dummy_constant cst) -let inConstant = +type constant_obj = + global_declaration * Dischargedhypsmap.discharged_hyps * logical_kind + +let inConstant : constant_obj -> obj = declare_object { (default_object "CONSTANT") with cache_function = cache_constant; load_function = load_constant; @@ -246,7 +252,9 @@ let dummy_inductive_entry (_,m) = ([],{ mind_entry_finite = true; mind_entry_inds = List.map dummy_one_inductive_entry m.mind_entry_inds }) -let inInductive = +type inductive_obj = Dischargedhypsmap.discharged_hyps * mutual_inductive_entry + +let inInductive : inductive_obj -> obj = declare_object {(default_object "INDUCTIVE") with cache_function = cache_inductive; load_function = load_inductive; diff --git a/library/declaremods.ml b/library/declaremods.ml index 95d2310ce..90d4245a4 100644 --- a/library/declaremods.ml +++ b/library/declaremods.ml @@ -179,14 +179,21 @@ let check_sub mtb sub_mtb_l = environment. *) let check_subtypes mp sub_mtb_l = - let mb = Global.lookup_module mp in + let mb = + try Global.lookup_module mp + with Not_found -> assert false + in let mtb = Modops.module_type_of_module None mb in check_sub mtb sub_mtb_l (* Same for module type [mp] *) let check_subtypes_mt mp sub_mtb_l = - check_sub (Global.lookup_modtype mp) sub_mtb_l + let mtb = + try Global.lookup_modtype mp + with Not_found -> assert false + in + check_sub mtb sub_mtb_l (* Create a functor type entry *) @@ -274,42 +281,26 @@ let conv_names_do_module exists what iter_objects i functions can be called only once (and "end_mod*" set the flag to false then) *) -let cache_module ((sp,kn),(entry,substobjs)) = +let cache_module ((sp,kn),substobjs) = let dir,mp = dir_of_sp sp, mp_of_kn kn in do_module false "cache" load_objects 1 dir mp substobjs [] - -(* TODO: This check is not essential *) -let check_empty s = function - | None -> () - | Some _ -> - anomaly ("We should never have full info in " ^ s^"!") - (* When this function is called the module itself is already in the environment. This function loads its objects only *) -let load_module i (oname,(entry,substobjs)) = - (* TODO: This check is not essential *) - check_empty "load_module" entry; +let load_module i (oname,substobjs) = conv_names_do_module false "load" load_objects i oname substobjs - -let open_module i (oname,(entry,substobjs)) = - (* TODO: This check is not essential *) - check_empty "open_module" entry; +let open_module i (oname,substobjs) = conv_names_do_module true "open" open_objects i oname substobjs +let subst_module (subst,(mbids,mp,objs)) = + (mbids,subst_mp subst mp, subst_objects subst objs) +let classify_module substobjs = Substitute substobjs -let subst_module (subst,(entry,(mbids,mp,objs))) = - check_empty "subst_module" entry; - (None,(mbids,subst_mp subst mp, subst_objects subst objs)) - - -let classify_module (_,substobjs) = - Substitute (None,substobjs) - -let (in_module,out_module) = +let (in_module : substitutive_objects -> obj), + (out_module : obj -> substitutive_objects) = declare_object_full {(default_object "MODULE") with cache_function = cache_module; load_function = load_module; @@ -337,7 +328,7 @@ let open_keep i ((sp,kn),seg) = let dirpath,mp = dir_of_sp sp, mp_of_kn kn in open_objects i (dirpath,(mp,empty_dirpath)) seg -let in_modkeep = +let in_modkeep : lib_objects -> obj = declare_object {(default_object "MODULE KEEP OBJECTS") with cache_function = cache_keep; load_function = load_keep; @@ -369,6 +360,7 @@ let _ = Summary.declare_summary "MODTYPE-INFO" let cache_modtype ((sp,kn),(entry,modtypeobjs,sub_mty_l)) = let mp = mp_of_kn kn in + (* We enrich the global environment *) let _ = match entry with | None -> @@ -392,7 +384,7 @@ let cache_modtype ((sp,kn),(entry,modtypeobjs,sub_mty_l)) = let load_modtype i ((sp,kn),(entry,modtypeobjs,_)) = - check_empty "load_modtype" entry; + assert (entry = None); if Nametab.exists_modtype sp then errorlabstrm "cache_modtype" @@ -404,7 +396,7 @@ let load_modtype i ((sp,kn),(entry,modtypeobjs,_)) = let open_modtype i ((sp,kn),(entry,_,_)) = - check_empty "open_modtype" entry; + assert (entry = None); if try Nametab.locate_modtype (qualid_of_path sp) <> (mp_of_kn kn) @@ -416,15 +408,18 @@ let open_modtype i ((sp,kn),(entry,_,_)) = Nametab.push_modtype (Nametab.Exactly i) sp (mp_of_kn kn) let subst_modtype (subst,(entry,(mbids,mp,objs),_)) = - check_empty "subst_modtype" entry; + assert (entry = None); (entry,(mbids,subst_mp subst mp,subst_objects subst objs),[]) - let classify_modtype (_,substobjs,_) = Substitute (None,substobjs,[]) +type modtype_obj = + (module_struct_entry * Entries.inline) option (* will be None in vo *) + * substitutive_objects + * module_type_body list -let in_modtype = +let in_modtype : modtype_obj -> obj = declare_object {(default_object "MODULE TYPE") with cache_function = cache_modtype; open_function = open_modtype; @@ -432,35 +427,27 @@ let in_modtype = subst_function = subst_modtype; classify_function = classify_modtype } -let rec replace_module_object idl ( mbids, mp, lib_stack) (mbids2,mp2,objs) mp1= - if mbids<>[] then - error "Unexpected functor objects" - else - let rec replace_idl = function - | _,[] -> [] - | id::idl,(id',obj)::tail when id = id' -> - if object_tag obj = "MODULE" then - (match idl with - [] -> (id, in_module - (None,(mbids,(MPdot(mp,label_of_id id)),subst_objects - (map_mp mp1 (MPdot(mp,label_of_id id)) empty_delta_resolver) objs)))::tail - | _ -> - let (_,substobjs) = out_module obj in - let substobjs' = replace_module_object idl substobjs - (mbids2,mp2,objs) mp in - (id, in_module (None,substobjs'))::tail - ) - else error "MODULE expected!" - | idl,lobj::tail -> lobj::replace_idl (idl,tail) - in - (mbids, mp, replace_idl (idl,lib_stack)) +let rec replace_module_object idl (mbids,mp,lib_stack) (mbids2,mp2,objs) mp1 = + if mbids<>[] then anomaly "Unexpected functor objects"; + let rec replace_idl = function + | _,[] -> [] + | id::idl,(id',obj)::tail when id = id' -> + if object_tag obj <> "MODULE" then anomaly "MODULE expected!"; + let substobjs = + if idl = [] then + let mp' = MPdot(mp, label_of_id id) in + mbids, mp', subst_objects (map_mp mp1 mp' empty_delta_resolver) objs + else + replace_module_object idl (out_module obj) (mbids2,mp2,objs) mp + in + (id, in_module substobjs)::tail + | idl,lobj::tail -> lobj::replace_idl (idl,tail) + in + (mbids, mp, replace_idl (idl,lib_stack)) -let discr_resolver mb = - match mb.mod_type with - SEBstruct _ -> - Some mb.mod_delta - | _ -> (*case mp is a functor *) - None +let discr_resolver mb = match mb.mod_type with + | SEBstruct _ -> Some mb.mod_delta + | _ -> None (* when mp is a functor *) (* Small function to avoid module typing during substobjs retrivial *) let rec get_objs_modtype_application env = function @@ -622,7 +609,7 @@ let end_module () = | Some mp_from,(mbids,_,objs) -> (mbids,mp,subst_objects (map_mp mp_from mp resolver) objs) in - let node = in_module (None,substobjs) in + let node = in_module substobjs in let objects = if keep = [] || mbids <> [] then special@[node] (* no keep objects or we are defining a functor *) @@ -859,7 +846,7 @@ let declare_module_ interp_modtype interp_modexpr id args res mexpr_o fs = reset_scope_subst (); ignore (add_leaf id - (in_module (Some (entry), substobjs))); + (in_module substobjs)); mmp (* Include *) @@ -887,32 +874,33 @@ let lift_oname (sp,kn) = let dir,_ = Libnames.repr_path sp in (dir,mp) -let cache_include (oname,((me,is_mod),(mbis,mp1,objs))) = +let cache_include (oname,(me,(mbis,mp1,objs))) = let dir,mp1 = lift_oname oname in let prefix = (dir,(mp1,empty_dirpath)) in load_objects 1 prefix objs; - open_objects 1 prefix objs - -let load_include i (oname,((me,is_mod),(mbis,mp1,objs))) = + open_objects 1 prefix objs + +let load_include i (oname,(me,(mbis,mp1,objs))) = let dir,mp1 = lift_oname oname in let prefix = (dir,(mp1,empty_dirpath)) in load_objects i prefix objs - - -let open_include i (oname,((me,is_mod),(mbis,mp1,objs))) = + +let open_include i (oname,(me,(mbis,mp1,objs))) = let dir,mp1 = lift_oname oname in let prefix = (dir,(mp1,empty_dirpath)) in open_objects i prefix objs - -let subst_include (subst,((me,is_mod),substobj)) = + +let subst_include (subst,(me,substobj)) = let (mbids,mp,objs) = substobj in let substobjs = (mbids,subst_mp subst mp,subst_objects subst objs) in - ((subst_inc_expr subst me,is_mod),substobjs) - -let classify_include ((me,is_mod),substobjs) = - Substitute ((me,is_mod),substobjs) + (subst_inc_expr subst me,substobjs) + +let classify_include (me,substobjs) = Substitute (me,substobjs) + +type include_obj = module_struct_entry * substitutive_objects -let (in_include,out_include) = +let (in_include : include_obj -> obj), + (out_include : obj -> include_obj) = declare_object_full {(default_object "INCLUDE") with cache_function = cache_include; load_function = load_include; @@ -987,8 +975,7 @@ let declare_one_include_inner annot (me,is_mod) = let substobjs = (mbids,mp1, subst_objects (map_mp mp mp1 resolver) objs) in reset_scope_subst (); - ignore (add_leaf id - (in_include ((me,is_mod), substobjs))) + ignore (add_leaf id (in_include (me, substobjs))) let declare_one_include interp_struct (me_ast,annot) = declare_one_include_inner annot diff --git a/library/goptions.ml b/library/goptions.ml index fbce68fad..7c522022f 100644 --- a/library/goptions.ml +++ b/library/goptions.ml @@ -94,7 +94,7 @@ module MakeTable = if p' == p then obj else (f,p') in - let inGo = + let inGo : option_mark * A.t -> obj = Libobject.declare_object {(Libobject.default_object nick) with Libobject.load_function = load_options; Libobject.open_function = load_options; diff --git a/library/heads.ml b/library/heads.ml index 813031558..9f9f1ca25 100644 --- a/library/heads.ml +++ b/library/heads.ml @@ -167,7 +167,9 @@ let discharge_head (_,(ref,k)) = let rebuild_head (ref,k) = (ref, compute_head ref) -let inHead = +type head_obj = evaluable_global_reference * head_approximation + +let inHead : head_obj -> obj = declare_object {(default_object "HEAD") with cache_function = cache_head; load_function = load_head; diff --git a/library/impargs.ml b/library/impargs.ml index 3f60bf96a..2d5ffe9ab 100644 --- a/library/impargs.ml +++ b/library/impargs.ml @@ -548,7 +548,11 @@ let rebuild_implicits (req,l) = let classify_implicits (req,_ as obj) = if req = ImplLocal then Dispose else Substitute obj -let inImplicits = +type implicits_obj = + implicit_discharge_request * + (global_reference * implicits_list list) list + +let inImplicits : implicits_obj -> obj = declare_object {(default_object "IMPLICITS") with cache_function = cache_implicits; load_function = load_implicits; diff --git a/library/lib.ml b/library/lib.ml index 23d334a5d..7554fd0bb 100644 --- a/library/lib.ml +++ b/library/lib.ml @@ -539,7 +539,7 @@ let close_section () = (*****************) (* Backtracking. *) -let (inLabel,outLabel) = +let (inLabel : int -> obj), (outLabel : obj -> int) = declare_object_full {(default_object "DOT") with classify_function = (fun _ -> Dispose)} diff --git a/library/library.ml b/library/library.ml index 9c602adbb..071c7a46c 100644 --- a/library/library.ml +++ b/library/library.ml @@ -300,7 +300,7 @@ let subst_import (_,o) = o let classify_import (_,export as obj) = if export then Substitute obj else Dispose -let in_import = +let in_import : dir_path * bool -> obj = declare_object {(default_object "IMPORT LIBRARY") with cache_function = cache_import; open_function = open_import; @@ -498,7 +498,7 @@ let rec_intern_library_from_file idopt f = type library_reference = dir_path list * bool option -let register_library (dir,m) = +let register_library m = Declaremods.register_library m.library_name m.library_compiled @@ -526,7 +526,9 @@ let discharge_require (_,o) = Some o (* open_function is never called from here because an Anticipate object *) -let in_require = +type require_obj = library_t list * dir_path list * bool option + +let in_require : require_obj -> obj = declare_object {(default_object "REQUIRE") with cache_function = cache_require; load_function = load_require; @@ -541,7 +543,8 @@ let xml_require = ref (fun d -> ()) let set_xml_require f = xml_require := f let require_library_from_dirpath modrefl export = - let needed = List.rev (List.fold_left rec_intern_library [] modrefl) in + let needed = List.fold_left rec_intern_library [] modrefl in + let needed = List.rev_map snd needed in let modrefl = List.map fst modrefl in if Lib.is_module_or_modtype () then begin @@ -561,7 +564,7 @@ let require_library qidl export = let require_library_from_file idopt file export = let modref,needed = rec_intern_library_from_file idopt file in - let needed = List.rev needed in + let needed = List.rev_map snd needed in if Lib.is_module_or_modtype () then begin add_anonymous_leaf (in_require (needed,[modref],None)); Option.iter (fun exp -> add_anonymous_leaf (in_import (modref,exp))) -- cgit v1.2.3