Lib
Lib: record of operations, backtrack, low-level sections
This module provides a general mechanism to keep a trace of all operations and to backtrack (undo) those operations. It provides also the section mechanism (at a low level; discharge is not known at this step).
type export = (export_flag * Libobject.open_filter) option
val make_oname : Nametab.object_prefix -> Names.Id.t -> Libobject.object_name
val make_foname : Names.Id.t -> Libobject.object_name
val oname_prefix : Libobject.object_name -> Nametab.object_prefix
type 'summary node =
| CompilingLibrary of Nametab.object_prefix |
| OpenedModule of is_type * export * Nametab.object_prefix * 'summary |
| OpenedSection of Nametab.object_prefix * 'summary |
val node_prefix : 'summary node -> Nametab.object_prefix
Extract the object_prefix
component. Note that it is the prefix of the objects *inside* this node, eg in Module M.
we have OpenedModule
with prefix containing M
.
type 'summary library_segment = ('summary node * Libobject.t list) list
Adding operations (which call the cache
method, and getting the current list of operations (most recent ones coming first).
val add_leaf : Libobject.obj -> unit
val add_discharged_leaf : Libobject.discharged_obj -> unit
The function contents
gives access to the current entire segment
val contents : unit -> Summary.Interp.frozen library_segment
val prefix : unit -> Nametab.object_prefix
User-side names
cwd()
is (prefix()).obj_dir
current_mp()
is (prefix()).obj_mp
Inside a library A.B module M section S, we have
make_path (resp make_path_except_section) uses cwd (resp cwd_except_section) make_kn uses current_mp
val cwd : unit -> Names.DirPath.t
val cwd_except_section : unit -> Names.DirPath.t
val current_dirpath : bool -> Names.DirPath.t
val make_path : Names.Id.t -> Libnames.full_path
val make_path_except_section : Names.Id.t -> Libnames.full_path
val current_mp : unit -> Names.ModPath.t
Kernel-side names
val make_kn : Names.Id.t -> Names.KerName.t
type discharged_item =
| DischargedExport of Libobject.ExportObj.t |
| DischargedLeaf of Libobject.discharged_obj |
type classified_objects = {
substobjs : Libobject.t list; |
keepobjs : Libobject.t list; |
escapeobjs : Libobject.t list; |
anticipateobjs : Libobject.t list; |
}
module type StagedLibS = sig ... end
The StagedLibS
abstraction describes operations and traversal for Lib at a given stage.
We provide two instances of StagedLibS
, corresponding to the Synterp and Interp stages.
module Synterp : StagedLibS with type summary = Summary.Synterp.frozen
module Interp : StagedLibS with type summary = Summary.Interp.frozen
val start_compilation : Names.DirPath.t -> Names.ModPath.t -> unit
type compilation_result = {
info : Library_info.t; |
synterp_objects : classified_objects; |
interp_objects : classified_objects; |
}
val end_compilation : Names.DirPath.t -> compilation_result
Finalize the compilation of a library.
val library_dp : unit -> Names.DirPath.t
The function library_dp
returns the DirPath.t
of the current compiling library (or default_library
)
val split_modpath : Names.ModPath.t -> Names.DirPath.t * Names.Id.t list
Extract the library part of a name even if in a section
val library_part : Names.GlobRef.t -> Names.DirPath.t
val section_segment_of_constant : Names.Constant.t -> Cooking.cooking_info
Section management for discharge
val section_segment_of_inductive : Names.MutInd.t -> Cooking.cooking_info
val section_segment_of_reference : Names.GlobRef.t -> Cooking.cooking_info
val section_instance : Names.GlobRef.t -> Constr.t array
val is_in_section : Names.GlobRef.t -> bool
val discharge_proj_repr : Names.Projection.Repr.t -> Names.Projection.Repr.t
discharge_proj_repr p
discharges projection p
if the associated record was defined in the current section. If that is not the case, it returns p
unchanged.
Compatibility layer