Module Libobject

type substitutivity =
| Dispose
| Substitute
| Keep
| Anticipate
type object_name = Libnames.full_path * Names.KerName.t
type open_filter
type ('a, 'b) object_declaration = {
object_name : string;
object_stage : Summary.Stage.t;
cache_function : 'b -> unit;
load_function : int -> 'b -> unit;
open_function : open_filter -> int -> 'b -> unit;
classify_function : 'a -> substitutivity;
subst_function : (Mod_subst.substitution * 'a) -> 'a;
discharge_function : 'a -> 'a option;
rebuild_function : 'a -> 'a;
}
val unfiltered : open_filter
val make_filter : finite:bool -> string CAst.t list -> open_filter

Anomaly when the list is empty.

type category
val create_category : string -> category

Anomaly if called more than once for a given string.

val in_filter : cat:category option -> open_filter -> bool

On cat:None, returns whether the filter allows opening uncategorized objects.

On cat:(Some category), returns whether the filter allows opening objects in the given category.

val simple_open : ?⁠cat:category -> ('i -> 'a -> unit) -> open_filter -> 'i -> 'a -> unit

Combinator for making objects with simple category-based open behaviour. When cat:None, can be opened by Unfiltered, but also by Filtered with a negative set.

val filter_and : open_filter -> open_filter -> open_filter option

Returns None when the intersection is empty.

val filter_or : open_filter -> open_filter -> open_filter
val default_object : ?⁠stage:Summary.Stage.t -> string -> ('a'b) object_declaration
val ident_subst_function : (Mod_subst.substitution * 'a) -> 'a

the identity substitution function

...
module Dyn : Dyn.S
type obj = Dyn.t
type algebraic_objects =
| Objs of t list
| Ref of Names.ModPath.t * Mod_subst.substitution
and t =
| ModuleObject of Names.Id.t * substitutive_objects
| ModuleTypeObject of Names.Id.t * substitutive_objects
| IncludeObject of algebraic_objects
| KeepObject of Names.Id.t * t list
| ExportObject of {
mpl : (open_filter * Names.ModPath.t) list;
}
| AtomicObject of obj
and substitutive_objects = Names.MBId.t list * algebraic_objects
val declare_object : ('a'a) object_declaration -> 'a -> obj
val declare_object_full : ('aNametab.object_prefix * 'a) object_declaration -> 'a Dyn.tag
val declare_named_object_full : ('aobject_name * 'a) object_declaration -> (Names.Id.t * 'a) Dyn.tag
val declare_named_object : ('aobject_name * 'a) object_declaration -> Names.Id.t -> 'a -> obj
val declare_named_object_gen : ('aNametab.object_prefix * 'a) object_declaration -> 'a -> obj
val cache_object : (Nametab.object_prefix * obj) -> unit
val load_object : int -> (Nametab.object_prefix * obj) -> unit
val open_object : open_filter -> int -> (Nametab.object_prefix * obj) -> unit
val subst_object : (Mod_subst.substitution * obj) -> obj
val classify_object : obj -> substitutivity
val discharge_object : obj -> obj option
val rebuild_object : obj -> obj
val object_stage : obj -> Summary.Stage.t
val local_object : ?⁠stage:Summary.Stage.t -> string -> cache:('a -> unit) -> discharge:('a -> 'a option) -> ('a'a) object_declaration
val local_object_nodischarge : ?⁠stage:Summary.Stage.t -> string -> cache:('a -> unit) -> ('a'a) object_declaration
val global_object : ?⁠cat:category -> ?⁠stage:Summary.Stage.t -> string -> cache:('a -> unit) -> subst:((Mod_subst.substitution * 'a) -> 'a) option -> discharge:('a -> 'a option) -> ('a'a) object_declaration
val global_object_nodischarge : ?⁠cat:category -> ?⁠stage:Summary.Stage.t -> string -> cache:('a -> unit) -> subst:((Mod_subst.substitution * 'a) -> 'a) option -> ('a'a) object_declaration
val superglobal_object : ?⁠stage:Summary.Stage.t -> string -> cache:('a -> unit) -> subst:((Mod_subst.substitution * 'a) -> 'a) option -> discharge:('a -> 'a option) -> ('a'a) object_declaration
val superglobal_object_nodischarge : ?⁠stage:Summary.Stage.t -> string -> cache:('a -> unit) -> subst:((Mod_subst.substitution * 'a) -> 'a) option -> ('a'a) object_declaration
Debug
val dump : unit -> (int * string) list