package coq
- 
        
          Context of de Bruijn variables (rel_context)
- 
        
          Recurrence on rel_context
- Context of variables (section variables and goal assumptions)
- 
        
          Recurrence on named_context: older declarations processed first
- Global constants
- Add entries to global environment
- ...
- Primitive projections
- Name quotients
- Modules
- Universe constraints
- Sets of referred section variables
- Unsafe judgments.
Install
    
    dune-project
 Dependency
Authors
Maintainers
Sources
sha256=b1501d686c21836302191ae30f610cca57fb309214c126518ca009363ad2cd3c
    
    
  doc/coq-core.kernel/Environ/index.html
Module EnvironSource
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.
Environments have the following components:
- a context for de Bruijn variables
- a context for de Bruijn variables vm values
- a context for section variables and goal assumptions
- a context for section variables and goal assumptions vm values
- a context for global constants and axioms
- a context for inductive definitions
- a set of universe constraints
- a flag telling if Set is, can be, or cannot be set impredicative
Linking information for the native compiler
type stratification = {- env_universes : UGraph.t;
- env_sprop_allowed : bool;
- env_universes_lbound : UGraph.Bound.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;(*- Identifier-indexed version of *)- env_named_ctx
- env_named_var : Constr.t list;(*- List of identifiers in *)- env_named_ctx, in the same order, including let-ins. This is not used in the kernel, but is critical to preserve sharing of evar instances in the proof engine.
}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.t;
- 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;
}Context of de Bruijn variables (rel_context)
Looks up in the context of local vars referred by indice (rel_context) raises Not_found if the index points out of the context
Recurrence on rel_context
Context of variables (section variables and goal assumptions)
val map_named_val : 
  (Constr.named_declaration -> Constr.named_declaration) ->
  named_context_val ->
  named_context_valmap_named_val f ctxt apply f to the body and the type of each declarations. *** /!\ *** f t should be convertible with t, and preserve the name
val push_named_context_val : 
  Constr.named_declaration ->
  named_context_val ->
  named_context_valLooks up in the context of local vars referred by names (named_context) raises Not_found if the Id.t is not found
Recurrence on named_context: older declarations processed first
val fold_named_context : 
  (env -> Constr.named_declaration -> 'a -> 'a) ->
  env ->
  init:'a ->
  'aval match_named_context_val : 
  named_context_val ->
  (Constr.named_declaration * lazy_val * named_context_val) optionval fold_named_context_reverse : 
  ('a -> Constr.named_declaration -> 'a) ->
  init:'a ->
  env ->
  'aRecurrence on named_context starting from younger decl
This forgets rel context and sets a new named context
This removes the n last declarations from the rel context
val fold_constants : 
  (Names.Constant.t ->
    Opaqueproof.opaque Declarations.constant_body ->
    'a ->
    'a) ->
  env ->
  'a ->
  'aUseful for printing
val fold_inductives : 
  (Names.MutInd.t -> Declarations.mutual_inductive_body -> 'a -> 'a) ->
  env ->
  'a ->
  'aGlobal constants
Add entries to global environment
val add_constant : 
  Names.Constant.t ->
  Opaqueproof.opaque Declarations.constant_body ->
  env ->
  envval add_constant_key : 
  Names.Constant.t ->
  Opaqueproof.opaque Declarations.constant_body ->
  link_info ->
  env ->
  envval lookup_constant : 
  Names.Constant.t ->
  env ->
  Opaqueproof.opaque Declarations.constant_bodyLooks up in the context of global constant names raises an anomaly if the required path is not found
New-style polymorphism
...
constant_value env c raises NotEvaluableConst Opaque if c is opaque, NotEvaluableConst NoBody if it has no body, NotEvaluableConst IsProj if c is a projection, NotEvaluableConst (IsPrimitive p) if c is primitive p and an anomaly if it does not exist in env
type const_evaluation_result = - | NoBody
- | Opaque
- | IsPrimitive of Univ.Instance.t * CPrimitives.t
val 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
The universe context associated to the constant, empty if not polymorphic
Primitive projections
Checks that the number of parameters is correct.
Inductive types
Looks up in the context of global inductive names raises an anomaly if the required path is not found
The universe context associated to the inductive, empty if not polymorphic
New-style polymorphism
Old-style polymorphism
Name quotients
Modules
shallow_add_module does not add module components
Universe constraints
Add universe constraints to the environment.
Check constraints are satifiable in the environment.
push_context ?(strict=false) ctx env pushes the universe context to the environment.
push_context_set ?(strict=false) ctx env pushes the universe context set to the environment. It does not fail even if one of the universes is already declared.
push_subgraph univs env adds the universes and constraints in univs to env as push_context_set ~strict:false univs env, and also checks that they do not imply new transitive constraints between pre-existing universes in env.
update_typing_flags ?typing_flags may update env with optional typing flags
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
closure of the input id set w.r.t. dependency
like really_needed but 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 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 f split sign into tail::(id,_,_)::head and return tail::(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_valPrimitives
- 
        
          Context of de Bruijn variables (rel_context)
- 
        
          Recurrence on rel_context
- Context of variables (section variables and goal assumptions)
- 
        
          Recurrence on named_context: older declarations processed first
- Global constants
- Add entries to global environment
- ...
- Primitive projections
- Name quotients
- Modules
- Universe constraints
- Sets of referred section variables
- Unsafe judgments.