Module Library

val require_library_from_dirpath : lib_resolver:(Names.DirPath.t -> CUnix.physical_path) -> (Names.DirPath.t * string) list -> bool option -> unit
...

Require = load in the environment + open (if the optional boolean is not None); mark also for export if the boolean is Some true

Start the compilation of a library
type seg_sum

Segments of a library

type seg_lib
type seg_univ = Univ.ContextSet.t * bool
type seg_proofs = Opaqueproof.opaque_proofterm array
type ('document, 'counters) todo_proofs =
| ProofsTodoNone
| ProofsTodoSomeEmpty of Future.UUIDSet.t
| ProofsTodoSome of Future.UUIDSet.t * ((Future.UUID.t'document) Stateid.request * bool) list * 'counters
val save_library_to : ('document'counters) todo_proofs -> output_native_objects:bool -> Names.DirPath.t -> string -> Opaqueproof.opaquetab -> unit
val load_library_todo : CUnix.physical_path -> seg_sum * seg_lib * seg_univ * 'tasks * seg_proofs
val save_library_raw : string -> seg_sum -> seg_lib -> seg_univ -> seg_proofs -> unit
Interrogate the status of libraries
val library_is_loaded : Names.DirPath.t -> bool
  • Tell if a library is loaded
val loaded_libraries : unit -> Names.DirPath.t list
  • Tell which libraries are loaded
val library_full_filename : Names.DirPath.t -> string
  • Return the full filename of a loaded library.
val overwrite_library_filenames : string -> unit
  • Overwrite the filename of all libraries (used when restoring a state)
val native_name_from_filename : string -> string
Native compiler.
val indirect_accessor : Opaqueproof.indirect_accessor
Opaque accessors
val intern_state : string -> unit

Low-level state overwriting, not very safe

val extern_state : string -> unit