Names
This file defines a lot of different notions of names used pervasively in the kernel as well as in other places. The essential datatypes exported by this API are:
module Id : sig ... end
Representation and operations on identifiers.
module Name : sig ... end
Representation and operations on identifiers that are allowed to be anonymous (i.e. "_" in concrete syntax).
type variable = Id.t
type module_ident = Id.t
module ModIdset : Util.Set.S with type elt = module_ident
module ModIdmap : Util.Map.ExtS with type key = module_ident and module Set := ModIdset
module DirPath : sig ... end
module DPset : Util.Set.S with type elt = DirPath.t
module Label : sig ... end
module MBId : sig ... end
module MBIset : Util.Set.S with type elt = MBId.t
module ModPath : sig ... end
module MPset : Util.Set.S with type elt = ModPath.t
module KerName : sig ... end
module KNpred : Predicate.S with type elt = KerName.t
module type EqType = sig ... end
module type QNameS = sig ... end
module Constant : sig ... end
module Cpred : Predicate.S with type elt = Constant.t
The *_env
modules consider an order on user part of names the others consider an order on canonical part of names
module Cset : CSig.SetS with type elt = Constant.t
module Cset_env : CSig.SetS with type elt = Constant.t
module Cmap : Util.Map.ExtS with type key = Constant.t and module Set := Cset
A map whose keys are constants (values of the Constant.t
type). Keys are ordered wrt. "canonical form" of the constant.
module Cmap_env : Util.Map.ExtS with type key = Constant.t and module Set := Cset_env
A map whose keys are constants (values of the Constant.t
type). Keys are ordered wrt. "user form" of the constant.
module MutInd : sig ... end
module Mindmap_env : CMap.ExtS with type key = MutInd.t
module Ind : sig ... end
type inductive = Ind.t
module Construct : sig ... end
type constructor = Construct.t
module Indset : CSet.S with type elt = inductive
module Constrset : CSet.S with type elt = constructor
module Indset_env : CSet.S with type elt = inductive
module Constrset_env : CSet.S with type elt = constructor
module Indmap_env : CMap.ExtS with type key = inductive and module Set := Indset_env
module Constrmap_env : CMap.ExtS with type key = constructor and module Set := Constrset_env
val constr_modpath : constructor -> ModPath.t
val ith_constructor_of_inductive : inductive -> int -> constructor
val inductive_of_constructor : constructor -> inductive
val index_of_constructor : constructor -> int
val ind_hash : inductive -> int
val ind_user_hash : inductive -> int
val ind_syntactic_hash : inductive -> int
val eq_constructor : constructor -> constructor -> bool
val eq_user_constructor : constructor -> constructor -> bool
val eq_syntactic_constructor : constructor -> constructor -> bool
val constructor_ord : constructor -> constructor -> int
val constructor_hash : constructor -> int
val constructor_user_ord : constructor -> constructor -> int
val constructor_user_hash : constructor -> int
val constructor_syntactic_ord : constructor -> constructor -> int
val constructor_syntactic_hash : constructor -> int
val hcons_con : Constant.t -> Constant.t
val hcons_construct : constructor -> constructor
index in the rel_context
part of environment starting by the end, inverse of de Bruijn indice
val eq_constant_key : Constant.t -> Constant.t -> bool
equalities on constant and inductive names (for the checker)
module Projection : sig ... end
module GlobRef : sig ... end
Located identifiers and objects with syntax.
type lstring = string CAst.t