Module Environ
Unsafe environments. We define here a datatype for environments. Since typing is not yet defined, it is not possible to check the informations added in environments, and that is why we speak here of ``unsafe'' environments.
val build_lazy_val : lazy_val -> (Vmvalues.values * Names.Id.Set.t) -> unitval force_lazy_val : lazy_val -> (Vmvalues.values * Names.Id.Set.t) optionval dummy_lazy_val : unit -> lazy_val
type link_info=|Linked of string|LinkedInteractive of string|NotLinkedLinking information for the native compiler
type key= int CEphemeron.key option Stdlib.reftype constant_key= Declarations.constant_body * (link_info Stdlib.ref * key)type mind_key= Declarations.mutual_inductive_body * link_info Stdlib.reftype globalsglobals = constants + projections + inductive types + modules + module-types
type stratification={env_universes : UGraph.t;env_sprop_allowed : bool;env_universes_lbound : Univ.Level.t;env_engagement : Declarations.engagement;}type named_context_val= private{env_named_ctx : Constr.named_context;env_named_map : (Constr.named_declaration * lazy_val) Names.Id.Map.t;}type rel_context_val= private{env_rel_ctx : Constr.rel_context;env_rel_map : (Constr.rel_declaration * lazy_val) Range.t;}type env= private{env_globals : globals;env_named_context : named_context_val;env_rel_context : rel_context_val;env_nb_rel : int;env_stratification : stratification;env_typing_flags : Declarations.typing_flags;retroknowledge : Retroknowledge.retroknowledge;indirect_pterms : Opaqueproof.opaquetab;}
val oracle : env -> Conv_oracle.oracleval set_oracle : env -> Conv_oracle.oracle -> envval eq_named_context_val : named_context_val -> named_context_val -> boolval empty_env : envval universes : env -> UGraph.tval universes_lbound : env -> Univ.Level.tval set_universes_lbound : env -> Univ.Level.t -> envval rel_context : env -> Constr.rel_contextval named_context : env -> Constr.named_contextval named_context_val : env -> named_context_valval opaque_tables : env -> Opaqueproof.opaquetabval set_opaque_tables : env -> Opaqueproof.opaquetab -> envval engagement : env -> Declarations.engagementval typing_flags : env -> Declarations.typing_flagsval is_impredicative_set : env -> boolval type_in_type : env -> boolval deactivated_guard : env -> boolval indices_matter : env -> boolval check_template : env -> boolval is_impredicative_sort : env -> Sorts.t -> boolval is_impredicative_univ : env -> Univ.Universe.t -> boolval empty_context : env -> boolis the local context empty
Context of de Bruijn variables (rel_context)
val nb_rel : env -> intval push_rel : Constr.rel_declaration -> env -> envval push_rel_context : Constr.rel_context -> env -> envval push_rec_types : Constr.rec_declaration -> env -> envval lookup_rel : int -> env -> Constr.rel_declarationLooks up in the context of local vars referred by indice (
rel_context) raisesNot_foundif the index points out of the context
Recurrence on rel_context
val fold_rel_context : (env -> Constr.rel_declaration -> 'a -> 'a) -> env -> init:'a -> 'a
Context of variables (section variables and goal assumptions)
val named_context_of_val : named_context_val -> Constr.named_contextval val_of_named_context : Constr.named_context -> named_context_valval empty_named_context_val : named_context_valval ids_of_named_context_val : named_context_val -> Names.Id.Set.tval map_named_val : (Constr.constr -> Constr.constr) -> named_context_val -> named_context_valmap_named_val f ctxtapplyfto the body and the type of each declarations. *** /!\ ***f tshould be convertible with t
val push_named : Constr.named_declaration -> env -> envval push_named_context : Constr.named_context -> env -> envval push_named_context_val : Constr.named_declaration -> named_context_val -> named_context_val
val lookup_named : Names.variable -> env -> Constr.named_declarationval lookup_named_val : Names.variable -> env -> lazy_valval lookup_named_ctxt : Names.variable -> named_context_val -> Constr.named_declarationval evaluable_named : Names.variable -> env -> boolval named_type : Names.variable -> env -> Constr.typesval named_body : Names.variable -> env -> Constr.constr option
Recurrence on named_context: older declarations processed first
val fold_named_context : (env -> Constr.named_declaration -> 'a -> 'a) -> env -> init:'a -> 'aval fold_named_context_reverse : ('a -> Constr.named_declaration -> 'a) -> init:'a -> env -> 'aRecurrence on
named_contextstarting from younger decl
val reset_with_named_context : named_context_val -> env -> envThis forgets rel context and sets a new named context
val fold_constants : (Names.Constant.t -> Declarations.constant_body -> 'a -> 'a) -> env -> 'a -> 'aUseful for printing
Global constants
Add entries to global environment
val add_constant : Names.Constant.t -> Declarations.constant_body -> env -> envval add_constant_key : Names.Constant.t -> Declarations.constant_body -> link_info -> env -> envval lookup_constant_key : Names.Constant.t -> env -> constant_keyval lookup_constant : Names.Constant.t -> env -> Declarations.constant_bodyLooks up in the context of global constant names raises
Not_foundif the required path is not found
val evaluable_constant : Names.Constant.t -> env -> boolval polymorphic_constant : Names.Constant.t -> env -> boolNew-style polymorphism
val polymorphic_pconstant : Constr.pconstant -> env -> boolval type_in_type_constant : Names.Constant.t -> env -> bool
...
type const_evaluation_result=|NoBody|Opaque|IsPrimitive of CPrimitives.t
exceptionNotEvaluableConst of const_evaluation_result
val constant_type : env -> Names.Constant.t Univ.puniverses -> Constr.types Univ.constrainedval constant_value_and_type : env -> Names.Constant.t Univ.puniverses -> Constr.constr option * Constr.types * Univ.Constraint.tThe universe context associated to the constant, empty if not polymorphic
val constant_context : env -> Names.Constant.t -> Univ.AUContext.tThe universe context associated to the constant, empty if not polymorphic
val body_of_constant_body : env -> Declarations.constant_body -> (Constr.constr * Univ.AUContext.t) optionReturns the body of the constant if it has any, and the polymorphic context it lives in. For monomorphic constant, the latter is empty, and for polymorphic constants, the term contains De Bruijn universe variables that need to be instantiated.
val constant_value_in : env -> Names.Constant.t Univ.puniverses -> Constr.constrval constant_type_in : env -> Names.Constant.t Univ.puniverses -> Constr.typesval constant_opt_value_in : env -> Names.Constant.t Univ.puniverses -> Constr.constr optionval is_primitive : env -> Names.Constant.t -> bool
Primitive projections
val lookup_projection : Names.Projection.t -> env -> Constr.typesChecks that the number of parameters is correct.
val get_projection : env -> Names.inductive -> proj_arg:int -> Names.Projection.Repr.t optionval get_projections : env -> Names.inductive -> Names.Projection.Repr.t array optionval lookup_mind_key : Names.MutInd.t -> env -> mind_keyInductive types
val add_mind_key : Names.MutInd.t -> mind_key -> env -> envval add_mind : Names.MutInd.t -> Declarations.mutual_inductive_body -> env -> envval lookup_mind : Names.MutInd.t -> env -> Declarations.mutual_inductive_bodyLooks up in the context of global inductive names raises
Not_foundif the required path is not found
val mind_context : env -> Names.MutInd.t -> Univ.AUContext.tThe universe context associated to the inductive, empty if not polymorphic
val polymorphic_ind : Names.inductive -> env -> boolNew-style polymorphism
val polymorphic_pind : Constr.pinductive -> env -> boolval type_in_type_ind : Names.inductive -> env -> boolval template_polymorphic_ind : Names.inductive -> env -> boolOld-style polymorphism
val template_polymorphic_variables : Names.inductive -> env -> Univ.Level.t listval template_polymorphic_pind : Constr.pinductive -> env -> boolval template_checked_ind : Names.inductive -> env -> bool
Modules
val add_modtype : Declarations.module_type_body -> env -> envval shallow_add_module : Declarations.module_body -> env -> envshallow_add_moduledoes not add module components
val lookup_module : Names.ModPath.t -> env -> Declarations.module_bodyval lookup_modtype : Names.ModPath.t -> env -> Declarations.module_type_body
Universe constraints
val add_constraints : Univ.Constraint.t -> env -> envAdd universe constraints to the environment.
- raises UniverseInconsistency
.
val check_constraints : Univ.Constraint.t -> env -> boolCheck constraints are satifiable in the environment.
val push_context : ?strict:bool -> Univ.UContext.t -> env -> envval push_context_set : ?strict:bool -> Univ.ContextSet.t -> env -> envval push_constraints_to_env : 'a Univ.constrained -> env -> envval push_subgraph : Univ.ContextSet.t -> env -> envpush_subgraph univs envadds the universes and constraints inunivstoenvaspush_context_set ~strict:false univs env, and also checks that they do not imply new transitive constraints between pre-existing universes inenv.
val set_engagement : Declarations.engagement -> env -> envval set_typing_flags : Declarations.typing_flags -> env -> envval make_sprop_cumulative : env -> envval set_allow_sprop : bool -> env -> envval sprop_allowed : env -> boolval universes_of_global : env -> Names.GlobRef.t -> Univ.AUContext.t
Sets of referred section variables
global_vars_set env c returns the list of id's occurring either directly as Var id in c or indirectly as a section variable dependent in a global reference occurring in c
val global_vars_set : env -> Constr.constr -> Names.Id.Set.tval vars_of_global : env -> Names.GlobRef.t -> Names.Id.Set.tval really_needed : env -> Names.Id.Set.t -> Names.Id.Set.tclosure of the input id set w.r.t. dependency
val keep_hyps : env -> Names.Id.Set.t -> Constr.named_contextlike
really_neededbut computes a well ordered named context
Unsafe judgments.
We introduce here the pre-type of judgments, which is actually only a datatype to store a term with its type and the type of its type.
val on_judgment : ('a -> 'b) -> ('a, 'a) punsafe_judgment -> ('b, 'b) punsafe_judgmentval on_judgment_value : ('c -> 'c) -> ('c, 't) punsafe_judgment -> ('c, 't) punsafe_judgmentval on_judgment_type : ('t -> 't) -> ('c, 't) punsafe_judgment -> ('c, 't) punsafe_judgment
type unsafe_judgment= (Constr.constr, Constr.types) punsafe_judgment
val make_judge : 'constr -> 'types -> ('constr, 'types) punsafe_judgmentval j_val : ('constr, 'types) punsafe_judgment -> 'constrval j_type : ('constr, 'types) punsafe_judgment -> 'types
type 'types punsafe_type_judgment={utj_val : 'types;utj_type : Sorts.t;}type unsafe_type_judgment= Constr.types punsafe_type_judgment
val apply_to_hyp : named_context_val -> Names.variable -> (Constr.named_context -> Constr.named_declaration -> Constr.named_context -> Constr.named_declaration) -> named_context_valapply_to_hyp sign id fsplitsignintotail::(id,_,_)::headand returntail::(f head (id,_,_) (rev tail))::head. the value associated to id should not change
val remove_hyps : Names.Id.Set.t -> (Constr.named_declaration -> Constr.named_declaration) -> (lazy_val -> lazy_val) -> named_context_val -> named_context_valval is_polymorphic : env -> Names.GlobRef.t -> boolval is_template_polymorphic : env -> Names.GlobRef.t -> boolval get_template_polymorphic_variables : env -> Names.GlobRef.t -> Univ.Level.t listval is_template_checked : env -> Names.GlobRef.t -> boolval is_type_in_type : env -> Names.GlobRef.t -> boolval no_link_info : link_infoNative compiler
val set_retroknowledge : env -> Retroknowledge.retroknowledge -> envPrimitives