Module System

Coqtop specific system utilities
Directories
type unix_path = string
type file_kind =
| FileDir of unix_path * string
| FileRegular of string
val (//) : unix_path -> string -> unix_path
val exists_dir : unix_path -> bool
val exclude_directory : unix_path -> unit
val process_directory : (file_kind -> unit) -> unix_path -> unit
val process_subdirectories : (unix_path -> string -> unit) -> unix_path -> unit
Files and load paths
val all_subdirs : unix_path:string -> (CUnix.physical_path * string list) list
val is_in_path : CUnix.load_path -> string -> bool
val is_in_system_path : string -> bool
val where_in_path : ?⁠warn:bool -> CUnix.load_path -> string -> CUnix.physical_path * string
val find_file_in_path : ?⁠warn:bool -> CUnix.load_path -> string -> CUnix.physical_path * string
val trust_file_cache : bool Stdlib.ref

trust_file_cache indicates whether we trust the underlying mapped file-system not to change along the execution of Coq. This assumption greatly speds up file search, but it is often inconvenient in interactive mode

val file_exists_respecting_case : string -> string -> bool
I/O functions
type magic_number_error = {
filename : string;
actual : int32;
expected : int32;
}
exception Bad_magic_number of magic_number_error
exception Bad_version_number of magic_number_error
val raw_extern_state : int -> string -> Stdlib.out_channel
val raw_intern_state : int -> string -> Stdlib.in_channel
val extern_state : int -> string -> 'a -> unit
val intern_state : int -> string -> 'a
val with_magic_number_check : ('a -> 'b) -> 'a -> 'b
val marshal_out : Stdlib.out_channel -> 'a -> unit
val marshal_in : string -> Stdlib.in_channel -> 'a
val check_caml_version : caml:string -> file:string -> unit
Time stamps.
type time
val get_time : unit -> time
val time_difference : time -> time -> float

in seconds

val fmt_time_difference : time -> time -> Pp.t
val with_time : batch:bool -> header:Pp.t -> ('a -> 'b) -> 'a -> 'b
val get_toplevel_path : ?⁠byte:bool -> string -> string

get_toplevel_path program builds a complete path to the executable denoted by program. This involves:

  • locating the directory: we don't rely on PATH as to make calls to /foo/bin/coqtop chose the right /foo/bin/coqproofworker
  • adding the proper suffixes: .opt/.byte depending on the current mode, + .exe if in windows.

Note that this function doesn't check that the executable actually exists. This is left back to caller, as well as the choice of fallback strategy. We could add a fallback strategy here but it is better not to as in most cases if this function fails to construct the right name you want you execution to fail rather than fall into choosing some random binary from the system-wide installation of Coq.