Module Globnames

type global_reference = Names.GlobRef.t =
| VarRef of Names.variable

A reference to the section-context.

| ConstRef of Names.Constant.t

A reference to the environment.

| IndRef of Names.inductive

A reference to an inductive type.

| ConstructRef of Names.constructor

A reference to a constructor of an inductive type.

Global reference is a kernel side type for all references together
val isVarRef : Names.GlobRef.t -> bool
val isConstRef : Names.GlobRef.t -> bool
val isIndRef : Names.GlobRef.t -> bool
val isConstructRef : Names.GlobRef.t -> bool
val eq_gr : Names.GlobRef.t -> Names.GlobRef.t -> bool
val canonical_gr : Names.GlobRef.t -> Names.GlobRef.t
val destVarRef : Names.GlobRef.t -> Names.variable
val destConstRef : Names.GlobRef.t -> Names.Constant.t
val destIndRef : Names.GlobRef.t -> Names.inductive
val destConstructRef : Names.GlobRef.t -> Names.constructor
val is_global : Names.GlobRef.t -> Constr.constr -> bool
val subst_constructor : Mod_subst.substitution -> Names.constructor -> Names.constructor
val subst_global : Mod_subst.substitution -> Names.GlobRef.t -> Names.GlobRef.t * Constr.constr Univ.univ_abstracted option
val subst_global_reference : Mod_subst.substitution -> Names.GlobRef.t -> Names.GlobRef.t
val printable_constr_of_global : Names.GlobRef.t -> Constr.constr

This constr is not safe to be typechecked, universe polymorphism is not handled here: just use for printing

val global_of_constr : Constr.constr -> Names.GlobRef.t

Turn a construction denoting a global reference into a global reference; raise Not_found if not a global reference

module RefOrdered = Names.GlobRef.Ordered
module RefOrdered_env = Names.GlobRef.Ordered_env
module Refset = Names.GlobRef.Set
module Refmap = Names.GlobRef.Map
module Refset_env = Names.GlobRef.Set_env
module Refmap_env = Names.GlobRef.Map_env
Extended global references
type syndef_name = Names.KerName.t
type extended_global_reference =
| TrueGlobal of Names.GlobRef.t
| SynDef of syndef_name
module ExtRefOrdered : sig ... end
type global_reference_or_constr =
| IsGlobal of Names.GlobRef.t
| IsConstr of Constr.constr