- val safe_basename_of_global : Names.GlobRef.t -> Names.Id.t
- val warning_axioms : unit -> unit
- val warning_opaques : bool -> unit
- val warning_ambiguous_name : ?loc:Loc.t -> (Libnames.qualid * Names.ModPath.t * Names.GlobRef.t) -> unit
- val warning_id : string -> unit
- val error_axiom_scheme : Names.GlobRef.t -> int -> 'a
- val error_constant : Names.GlobRef.t -> 'a
- val error_inductive : Names.GlobRef.t -> 'a
- val error_nb_cons : unit -> 'a
- val error_module_clash : Names.ModPath.t -> Names.ModPath.t -> 'a
- val error_no_module_expr : Names.ModPath.t -> 'a
- val error_singleton_become_prop : Names.Id.t -> Names.GlobRef.t option -> 'a
- val error_unknown_module : Libnames.qualid -> 'a
- val error_scheme : unit -> 'a
- val error_not_visible : Names.GlobRef.t -> 'a
- val error_MPfile_as_mod : Names.ModPath.t -> bool -> 'a
- val check_inside_module : unit -> unit
- val check_inside_section : unit -> unit
- val check_loaded_modfile : Names.ModPath.t -> unit
- val msg_of_implicit : Miniml.kill_reason -> string
- val err_or_warn_remaining_implicit : Miniml.kill_reason -> unit
- val info_file : string -> unit
- val occur_kn_in_ref : Names.MutInd.t -> Names.GlobRef.t -> bool
- val repr_of_r : Names.GlobRef.t -> Names.ModPath.t * Names.Label.t
- val modpath_of_r : Names.GlobRef.t -> Names.ModPath.t
- val label_of_r : Names.GlobRef.t -> Names.Label.t
- val base_mp : Names.ModPath.t -> Names.ModPath.t
- val is_modfile : Names.ModPath.t -> bool
- val string_of_modfile : Names.ModPath.t -> string
- val file_of_modfile : Names.ModPath.t -> string
- val is_toplevel : Names.ModPath.t -> bool
- val at_toplevel : Names.ModPath.t -> bool
- val mp_length : Names.ModPath.t -> int
- val prefixes_mp : Names.ModPath.t -> Names.MPset.t
- val common_prefix_from_list : Names.ModPath.t -> Names.ModPath.t list -> Names.ModPath.t option
- val get_nth_label_mp : int -> Names.ModPath.t -> Names.Label.t
- val labels_of_ref : Names.GlobRef.t -> Names.ModPath.t * Names.Label.t list
- val add_typedef : Names.Constant.t -> Opaqueproof.opaque Declarations.constant_body -> Miniml.ml_type -> unit
- val lookup_typedef : Names.Constant.t -> Opaqueproof.opaque Declarations.constant_body -> Miniml.ml_type option
- val add_cst_type : Names.Constant.t -> Opaqueproof.opaque Declarations.constant_body -> Miniml.ml_schema -> unit
- val lookup_cst_type : Names.Constant.t -> Opaqueproof.opaque Declarations.constant_body -> Miniml.ml_schema option
- val add_ind : Names.MutInd.t -> Declarations.mutual_inductive_body -> Miniml.ml_ind -> unit
- val lookup_ind : Names.MutInd.t -> Declarations.mutual_inductive_body -> Miniml.ml_ind option
- val add_inductive_kind : Names.MutInd.t -> Miniml.inductive_kind -> unit
- val is_coinductive : Names.GlobRef.t -> bool
- val is_coinductive_type : Miniml.ml_type -> bool
- val get_record_fields : Names.GlobRef.t -> Names.GlobRef.t option list
- val record_fields_of_type : Miniml.ml_type -> Names.GlobRef.t option list
- val add_recursors : Environ.env -> Names.MutInd.t -> unit
- val is_recursor : Names.GlobRef.t -> bool
- val add_projection : int -> Names.Constant.t -> Names.inductive -> unit
- val is_projection : Names.GlobRef.t -> bool
- val projection_arity : Names.GlobRef.t -> int
- val projection_info : Names.GlobRef.t -> Names.inductive * int
- val add_info_axiom : Names.GlobRef.t -> unit
- val remove_info_axiom : Names.GlobRef.t -> unit
- val add_log_axiom : Names.GlobRef.t -> unit
- val add_opaque : Names.GlobRef.t -> unit
- val remove_opaque : Names.GlobRef.t -> unit
- val reset_tables : unit -> unit
- val access_opaque : unit -> bool
- val auto_inline : unit -> bool
- val type_expand : unit -> bool
- val keep_singleton : unit -> bool
- type opt_flag-  = - {| opt_kill_dum : bool; |  | opt_fix_fun : bool; |  | opt_case_iot : bool; |  | opt_case_idr : bool; |  | opt_case_idg : bool; |  | opt_case_cst : bool; |  | opt_case_fun : bool; |  | opt_case_app : bool; |  | opt_let_app : bool; |  | opt_lin_let : bool; |  | opt_lin_beta : bool; |  
 - }
- val optims : unit -> opt_flag
- val conservative_types : unit -> bool
- type lang-  = | | Ocaml |  | | Haskell |  | | Scheme |  | | JSON |  
 
- val lang : unit -> lang
- val set_modular : bool -> unit
- val modular : unit -> bool
- val set_library : bool -> unit
- val library : unit -> bool
- val set_extrcompute : bool -> unit
- val is_extrcompute : unit -> bool
- val to_inline : Names.GlobRef.t -> bool
- val to_keep : Names.GlobRef.t -> bool
- val implicits_of_global : Names.GlobRef.t -> Int.Set.t
- val type_scheme_nb_args_hook : (Environ.env -> Constr.t -> int) Hook.t
- val is_custom : Names.GlobRef.t -> bool
- val is_inline_custom : Names.GlobRef.t -> bool
- val find_custom : Names.GlobRef.t -> string
- val find_type_custom : Names.GlobRef.t -> string list * string
- val is_custom_match : Miniml.ml_branch array -> bool
- val find_custom_match : Miniml.ml_branch array -> string
- type int_or_id-  =