package coq

  1. Overview
  2. Docs
Legend:
Library
Module
Module type
Parameter
Class
Class type
type vodigest =
  1. | Dvo_or_vi of Digest.t
  2. | Dvivo of Digest.t * Digest.t
val digest_match : actual:vodigest -> required:vodigest -> bool
Safe environments

Since we are now able to type terms, we can define an abstract type of safe environments, where objects are typed before being added.

We also provide functionality for modules : start_module, end_module, etc.

type safe_environment
type section_data
val empty_environment : safe_environment
val is_initial : safe_environment -> bool
val env_of_safe_env : safe_environment -> Environ.env
val sections_of_safe_env : safe_environment -> section_data Section.t option
val structure_body_of_safe_env : safe_environment -> Declarations.structure_body

The safe_environment state monad

type safe_transformer0 = safe_environment -> safe_environment
type 'a safe_transformer = safe_environment -> 'a * safe_environment
Stm machinery
type private_constants
val empty_private_constants : private_constants
val is_empty_private_constants : private_constants -> bool

concat_private e1 e2 adds the constants of e1 to e2, i.e. constants in e1 must be more recent than those of e2.

val push_private_constants : Environ.env -> private_constants -> Environ.env

Push the constants in the environment if not already there.

val universes_of_private : private_constants -> Univ.ContextSet.t
val is_curmod_library : safe_environment -> bool
val join_safe_environment : ?except:Future.UUIDSet.t -> safe_environment -> safe_environment
val is_joined_environment : safe_environment -> bool

Enriching a safe environment

Insertion of global axioms or definitions

type side_effect_declaration =
  1. | DefinitionEff : Entries.definition_entry -> side_effect_declaration
  2. | OpaqueEff : Constr.constr Entries.opaque_entry -> side_effect_declaration
type exported_private_constant = Names.Constant.t
val export_private_constants : private_constants -> exported_private_constant list safe_transformer

returns the main constant

Similar to add_constant but also returns a certificate

Adding an inductive type

Adding a module or a module type

Adding universe constraints

val push_context_set : strict:bool -> Univ.ContextSet.t -> safe_transformer0
val add_constraints : Univ.Constraint.t -> safe_transformer0

Setting the type theory flavor

val set_indices_matter : bool -> safe_transformer0
val set_share_reduction : bool -> safe_transformer0
val set_check_guarded : bool -> safe_transformer0
val set_check_positive : bool -> safe_transformer0
val set_check_universes : bool -> safe_transformer0
val set_VM : bool -> safe_transformer0
val set_native_compiler : bool -> safe_transformer0
val set_allow_sprop : bool -> safe_transformer0
val check_engagement : Environ.env -> Declarations.set_predicativity -> unit
Interactive section functions
val open_section : safe_transformer0
val close_section : safe_transformer0
val sections_are_opened : safe_environment -> bool

Insertion of local declarations (Local or Variables)

val push_named_assum : (Names.Id.t * Constr.types) -> safe_transformer0
val push_section_context : (Names.Name.t array * Univ.UContext.t) -> safe_transformer0

Add local universes to a polymorphic section

Interactive module functions
val allow_delayed_constants : bool ref

Traditional mode: check at end of module that no future was created.

The optional result type is given without its functorial part

val current_modpath : safe_environment -> Names.ModPath.t
val current_dirpath : safe_environment -> Names.DirPath.t
Libraries : loading and saving compilation units
type compiled_library
val module_of_library : compiled_library -> Declarations.module_body
val univs_of_library : compiled_library -> Univ.ContextSet.t
val export : ?except:Future.UUIDSet.t -> output_native_objects:bool -> safe_environment -> Names.DirPath.t -> Names.ModPath.t * compiled_library * Nativelib.native_library
Safe typing judgments
type judgment
val j_val : judgment -> Constr.constr
val j_type : judgment -> Constr.constr

The safe typing of a term returns a typing judgment.

Queries
val exists_objlabel : Names.Label.t -> safe_environment -> bool
val constant_of_delta_kn_senv : safe_environment -> Names.KerName.t -> Names.Constant.t
val mind_of_delta_kn_senv : safe_environment -> Names.KerName.t -> Names.MutInd.t
Retroknowledge / Native compiler
val register_inline : Names.Constant.t -> safe_transformer0
val register_inductive : Names.inductive -> 'a CPrimitives.prim_ind -> safe_transformer0
OCaml

Innovation. Community. Security.