Module Evd
This file defines the pervasive unification state used everywhere in Coq tactic engine. It is very low-level and most of the functions exported here are irrelevant to the standard API user. Consider using Evarutil, Sigma or Proofview instead.
A unification state (of type evar_map) is primarily a finite mapping from existential variables to records containing the type of the evar (evar_concl), the context under which it was introduced (evar_hyps) and its definition (evar_body). evar_extra is used to add any other kind of information.
It also contains conversion constraints, debugging information and information about meta variables.
type econstrtype etypes= econstr
Existential variables and unification states
Evar filters
module Filter : sig ... endmodule Abstraction : sig ... endEvar infos
type evar_body=|Evar_empty|Evar_defined of econstrtype evar_info={evar_concl : econstr;Type of the evar.
evar_hyps : Environ.named_context_val;Context of the evar.
evar_body : evar_body;Optional content of the evar.
evar_filter : Filter.t;Boolean mask over
evar_hyps. Should have the same length. When filtered out, the corresponding variable is not allowed to occur in the solutionevar_abstract_arguments : Abstraction.t;Boolean information over
evar_hyps, telling if an hypothesis instance can be imitated or should stay abstract in HO unification problems and inversion (seesecond_order_matching_with_argsfor its use).evar_source : Evar_kinds.t Loc.located;Information about the evar.
evar_candidates : econstr list option;List of possible solutions when known that it is a finite list
}
val make_evar : Environ.named_context_val -> etypes -> evar_infoval evar_concl : evar_info -> econstrval evar_context : evar_info -> (econstr, etypes) Context.Named.ptval evar_filtered_context : evar_info -> (econstr, etypes) Context.Named.ptval evar_hyps : evar_info -> Environ.named_context_valval evar_filtered_hyps : evar_info -> Environ.named_context_valval evar_body : evar_info -> evar_bodyval evar_candidates : evar_info -> Constr.constr list optionval evar_filter : evar_info -> Filter.tval evar_env : evar_info -> Environ.envval evar_filtered_env : evar_info -> Environ.envval map_evar_body : (econstr -> econstr) -> evar_body -> evar_bodyval map_evar_info : (econstr -> econstr) -> evar_info -> evar_info
Unification state
*
type evar_mapType of unification state. Essentially a bunch of state-passing data needed to handle incremental term construction.
val empty : evar_mapThe empty evar map.
val from_env : Environ.env -> evar_mapThe empty evar map with given universe context, taking its initial universes from env.
val is_empty : evar_map -> boolWhether an evarmap is empty.
val has_undefined : evar_map -> boolhas_undefined sigmaistrueif and only if there are uninstantiated evars insigma.
val new_evar : evar_map -> ?name:Names.Id.t -> ?typeclass_candidate:bool -> evar_info -> evar_map * Evar.tCreates a fresh evar mapping to the given information.
val add : evar_map -> Evar.t -> evar_info -> evar_mapadd sigma ev infoaddsevwith evar infoinfoin sigma. Precondition: ev must not preexist insigma.
val find_undefined : evar_map -> Evar.t -> evar_infoSame as
findbut restricted to undefined evars. For efficiency reasons.
val fold : (Evar.t -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'aApply a function to all evars and their associated info in an evarmap.
val fold_undefined : (Evar.t -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'aSame as
fold, but restricted to undefined evars. For efficiency reasons.
val raw_map : (Evar.t -> evar_info -> evar_info) -> evar_map -> evar_mapApply the given function to all evars in the map. Beware: this function expects the argument function to preserve the kind of
evar_body, i.e. it must sendEvar_emptytoEvar_emptyandEvar_defined cto someEvar_defined c'.
val raw_map_undefined : (Evar.t -> evar_info -> evar_info) -> evar_map -> evar_mapSame as
raw_map, but restricted to undefined evars. For efficiency reasons.
val define : Evar.t -> econstr -> evar_map -> evar_mapSet the body of an evar to the given constr. It is expected that:
- The evar is already present in the evarmap.
- The evar is not defined in the evarmap yet.
- All the evars present in the constr should be present in the evar map.
val define_with_evar : Evar.t -> econstr -> evar_map -> evar_mapSame as
define ev body evd, except the body must be an existential variableev'. This additionally makesev'inherit theobligationandtypeclassflags ofev.
val cmap : (econstr -> econstr) -> evar_map -> evar_mapMap the function on all terms in the evar map.
val add_constraints : evar_map -> Univ.Constraint.t -> evar_mapAdd universe constraints in an evar map.
val undefined_map : evar_map -> evar_info Evar.Map.tAccess the undefined evar mapping directly.
val drop_all_defined : evar_map -> evar_mapval is_maybe_typeclass_hook : (evar_map -> Constr.constr -> bool) Hook.t
Instantiating partial terms
val existential_value : evar_map -> econstr Constr.pexistential -> econstrexistential_value sigma evraisesNotInstantiatedEvarifevhas no body andNot_foundif it does not exist insigma
val existential_value0 : evar_map -> Constr.existential -> Constr.constrval existential_type : evar_map -> econstr Constr.pexistential -> etypesval existential_type0 : evar_map -> Constr.existential -> Constr.typesval existential_opt_value : evar_map -> econstr Constr.pexistential -> econstr optionSame as
existential_valuebut returns an option instead of raising an exception.
val existential_opt_value0 : evar_map -> Constr.existential -> Constr.constr optionval evar_instance_array : (Constr.named_declaration -> 'a -> bool) -> evar_info -> 'a array -> (Names.Id.t * 'a) listval instantiate_evar_array : evar_info -> econstr -> econstr array -> econstrval evars_reset_evd : ?with_conv_pbs:bool -> ?with_univs:bool -> evar_map -> evar_map -> evar_mapspiwack: this function seems to somewhat break the abstraction.
Misc
val restrict : Evar.t -> Filter.t -> ?candidates:econstr list -> ?src:Evar_kinds.t Loc.located -> evar_map -> evar_map * Evar.tRestrict an undefined evar into a new evar by filtering context and possibly limiting the instances to a set of candidates (candidates are filtered according to the filter)
val is_restricted_evar : evar_map -> Evar.t -> Evar.t optionTell if an evar comes from restriction of another evar, and if yes, which
val set_typeclass_evars : evar_map -> Evar.Set.t -> evar_mapMark the given set of evars as available for resolution.
Precondition: they should indeed refer to undefined typeclass evars.
val is_typeclass_evar : evar_map -> Evar.t -> boolIs the evar declared resolvable for typeclass resolution
val downcast : Evar.t -> etypes -> evar_map -> evar_mapChange the type of an undefined evar to a new type assumed to be a subtype of its current type; subtyping must be ensured by caller
val evar_source : Evar.t -> evar_map -> Evar_kinds.t Loc.locatedConvenience function. Wrapper around
findto recover the source of an evar in a given evar map.
val evar_ident : Evar.t -> evar_map -> Names.Id.t optionval rename : Evar.t -> Names.Id.t -> evar_map -> evar_mapval evar_key : Names.Id.t -> evar_map -> Evar.tval evar_source_of_meta : Constr.metavariable -> evar_map -> Evar_kinds.t Loc.locatedval dependent_evar_ident : Evar.t -> evar_map -> Names.Id.t
Side-effects
val emit_side_effects : Safe_typing.private_constants -> evar_map -> evar_mapPush a side-effect into the evar map.
val eval_side_effects : evar_map -> Safe_typing.private_constantsReturn the effects contained in the evar map.
Future goals
val declare_future_goal : ?tag:goal_kind -> Evar.t -> evar_map -> evar_mapAdds an existential variable to the list of future goals. For internal uses only.
val declare_principal_goal : ?tag:goal_kind -> Evar.t -> evar_map -> evar_mapAdds an existential variable to the list of future goals and make it principal. Only one existential variable can be made principal, an error is raised otherwise. For internal uses only.
val future_goals : evar_map -> Evar.t listRetrieves the list of future goals. Used by the
refineprimitive of the tactic engine.
val principal_future_goal : evar_map -> Evar.t optionRetrieves the name of the principal existential variable if there is one. Used by the
refineprimitive of the tactic engine.
val save_future_goals : evar_map -> future_goalsRetrieves the list of future goals including the principal future goal. Used by the
refineprimitive of the tactic engine.
val reset_future_goals : evar_map -> evar_mapClears the list of future goals (as well as the principal future goal). Used by the
refineprimitive of the tactic engine.
val restore_future_goals : evar_map -> future_goals -> evar_mapSets the future goals (including the principal future goal) to a previous value. Intended to be used after a local list of future goals has been consumed. Used by the
refineprimitive of the tactic engine.
val fold_future_goals : (evar_map -> Evar.t -> evar_map) -> evar_map -> future_goals -> evar_mapFold future goals
val map_filter_future_goals : (Evar.t -> Evar.t option) -> future_goals -> future_goalsApplies a function on the future goals
val filter_future_goals : (Evar.t -> bool) -> future_goals -> future_goalsApplies a filter on the future goals
val dispatch_future_goals : future_goals -> Evar.t list * Evar.t list * Evar.t list * Evar.t optionReturns the future_goals dispatched into regular, shelved, given_up goals; last argument is the goal tagged as principal if any
val extract_given_up_future_goals : future_goals -> Evar.t list * Evar.t listAn ad hoc variant for Proof.proof; not for general use
val shelve_on_future_goals : Evar.t list -> future_goals -> future_goalsPush goals on the shelve of future goals
Sort variables
Evar maps also keep track of the universe constraints defined at a given point. This section defines the relevant manipulation functions.
val add_universe_constraints : evar_map -> UnivProblem.Set.t -> evar_mapAdd the given universe unification constraints to the evar map.
- raises UniversesDiffer
in case a first-order unification fails.
- raises UniverseInconsistency
.
Extra data
Evar maps can contain arbitrary data, allowing to use an extensible state. As evar maps are theoretically used in a strict state-passing style, such additional data should be passed along transparently. Some old and bug-prone code tends to drop them nonetheless, so you should keep cautious.
Enriching with evar maps
type 'a sigma={it : 'a;The base object.
sigma : evar_map;The added unification state.
}The type constructor
'a sigmaadds an evar map to an object of type'a.
The state monad with state an evar map
Meta machinery
These functions are almost deprecated. They were used before the introduction of the full-fledged evar calculus. In an ideal world, they should be removed. Alas, some parts of the code still use them. Do not use in newly-written code.
module Metaset : Util.Set.S with type S.elt = Constr.metavariablemodule Metamap : Util.Map.ExtS with type key = Constr.metavariable and module Set := Metasettype 'a freelisted={rebus : 'a;freemetas : Metaset.t;}
val metavars_of : econstr -> Metaset.tval mk_freelisted : econstr -> econstr freelistedval map_fl : ('a -> 'b) -> 'a freelisted -> 'b freelisted
val eq_instance_constraint : instance_constraint -> instance_constraint -> bool
type instance_status= instance_constraint * instance_typing_status
type clbinding=|Cltyp of Names.Name.t * econstr freelisted|Clval of Names.Name.t * econstr freelisted * instance_status * econstr freelistedtype conv_pb= Reduction.conv_pbUnification constraints
type evar_constraint= conv_pb * Environ.env * econstr * econstr
val add_conv_pb : ?tail:bool -> evar_constraint -> evar_map -> evar_mapThe following two functions are for internal use only, see
Evarutil.add_unification_pbfor a safe interface.
val conv_pbs : evar_map -> evar_constraint listval extract_changed_conv_pbs : evar_map -> (Evar.Set.t -> evar_constraint -> bool) -> evar_map * evar_constraint listval extract_all_conv_pbs : evar_map -> evar_map * evar_constraint listval loc_of_conv_pb : evar_map -> evar_constraint -> Loc.t option
val evars_of_term : Constr.constr -> Evar.Set.tincluding evars in instances of evars
val evars_of_named_context : (econstr, etypes) Context.Named.pt -> Evar.Set.tval evars_of_filtered_evar_info : evar_info -> Evar.Set.tval meta_list : evar_map -> (Constr.metavariable * clbinding) listMetas
val meta_defined : evar_map -> Constr.metavariable -> boolval meta_value : evar_map -> Constr.metavariable -> econstrmeta_fvalueraisesNot_foundif meta not in map orAnomalyif meta has no value
val meta_fvalue : evar_map -> Constr.metavariable -> econstr freelisted * instance_statusval meta_opt_fvalue : evar_map -> Constr.metavariable -> (econstr freelisted * instance_status) optionval meta_type : evar_map -> Constr.metavariable -> etypesval meta_type0 : evar_map -> Constr.metavariable -> Constr.typesval meta_ftype : evar_map -> Constr.metavariable -> etypes freelistedval meta_name : evar_map -> Constr.metavariable -> Names.Name.tval meta_declare : Constr.metavariable -> etypes -> ?name:Names.Name.t -> evar_map -> evar_mapval meta_assign : Constr.metavariable -> (econstr * instance_status) -> evar_map -> evar_mapval meta_reassign : Constr.metavariable -> (econstr * instance_status) -> evar_map -> evar_mapval clear_metas : evar_map -> evar_mapval meta_merge : ?with_univs:bool -> evar_map -> evar_map -> evar_mapmeta_merge evd1 evd2returnsevd2extended with the metas ofevd1
val undefined_metas : evar_map -> Constr.metavariable listval map_metas_fvalue : (econstr -> econstr) -> evar_map -> evar_mapval map_metas : (econstr -> econstr) -> evar_map -> evar_map
type metabinding= Constr.metavariable * econstr * instance_status
val retract_coercible_metas : evar_map -> metabinding list * evar_map
FIXME: Nothing to do here
type rigid= UState.rigid=|UnivRigid|UnivFlexible of boolIs substitution by an algebraic ok?
type 'a in_evar_universe_context= 'a * UState.t
val restrict_universe_context : evar_map -> Univ.LSet.t -> evar_mapval universe_of_name : evar_map -> Names.Id.t -> Univ.Level.tRaises Not_found if not a name for a universe in this map.
val universe_binders : evar_map -> UnivNames.universe_bindersval new_univ_level_variable : ?loc:Loc.t -> ?name:Names.Id.t -> rigid -> evar_map -> evar_map * Univ.Level.tval new_univ_variable : ?loc:Loc.t -> ?name:Names.Id.t -> rigid -> evar_map -> evar_map * Univ.Universe.tval new_sort_variable : ?loc:Loc.t -> ?name:Names.Id.t -> rigid -> evar_map -> evar_map * Sorts.tval add_global_univ : evar_map -> Univ.Level.t -> evar_mapval universe_rigidity : evar_map -> Univ.Level.t -> rigidval make_flexible_variable : evar_map -> algebraic:bool -> Univ.Level.t -> evar_mapSee
UState.make_flexible_variable
val make_nonalgebraic_variable : evar_map -> Univ.Level.t -> evar_mapSee
UState.make_nonalgebraic_variable.
val is_sort_variable : evar_map -> Sorts.t -> Univ.Level.t optionis_sort_variable evm sreturnsSome uorNoneifsis not a local sort variable declared inevm
val is_flexible_level : evar_map -> Univ.Level.t -> boolval normalize_universe : evar_map -> Univ.Universe.t -> Univ.Universe.tval normalize_universe_instance : evar_map -> Univ.Instance.t -> Univ.Instance.tval set_leq_sort : Environ.env -> evar_map -> Sorts.t -> Sorts.t -> evar_mapval set_eq_sort : Environ.env -> evar_map -> Sorts.t -> Sorts.t -> evar_mapval set_eq_level : evar_map -> Univ.Level.t -> Univ.Level.t -> evar_mapval set_leq_level : evar_map -> Univ.Level.t -> Univ.Level.t -> evar_mapval set_eq_instances : ?flex:bool -> evar_map -> Univ.Instance.t -> Univ.Instance.t -> evar_mapval check_eq : evar_map -> Univ.Universe.t -> Univ.Universe.t -> boolval check_leq : evar_map -> Univ.Universe.t -> Univ.Universe.t -> boolval check_constraints : evar_map -> Univ.Constraint.t -> boolval evar_universe_context : evar_map -> UState.tval universe_context_set : evar_map -> Univ.ContextSet.tval universe_subst : evar_map -> UnivSubst.universe_opt_substval universes : evar_map -> UGraph.tval to_universe_context : evar_map -> Univ.UContext.tto_universe_context evmextracts the local universes and constraints ofevmand orders the universes the same asUniv.ContextSet.to_context.
val univ_entry : poly:bool -> evar_map -> Entries.universes_entryval const_univ_entry : poly:bool -> evar_map -> Entries.universes_entryval check_univ_decl : poly:bool -> evar_map -> UState.universe_decl -> Entries.universes_entryval merge_universe_context : evar_map -> UState.t -> evar_mapval set_universe_context : evar_map -> UState.t -> evar_mapval merge_context_set : ?loc:Loc.t -> ?sideff:bool -> rigid -> evar_map -> Univ.ContextSet.t -> evar_mapval merge_universe_subst : evar_map -> UnivSubst.universe_opt_subst -> evar_mapval with_context_set : ?loc:Loc.t -> rigid -> evar_map -> 'a Univ.in_universe_context_set -> evar_map * 'aval nf_univ_variables : evar_map -> evar_map * Univ.universe_substval fix_undefined_variables : evar_map -> evar_mapval refresh_undefined_universes : evar_map -> evar_map * Univ.universe_level_substval minimize_universes : evar_map -> evar_mapUniverse minimization
val update_sigma_env : evar_map -> Environ.env -> evar_map
val fresh_sort_in_family : ?loc:Loc.t -> ?rigid:rigid -> evar_map -> Sorts.family -> evar_map * Sorts.tval fresh_constant_instance : ?loc:Loc.t -> Environ.env -> evar_map -> Names.Constant.t -> evar_map * Constr.pconstantval fresh_inductive_instance : ?loc:Loc.t -> Environ.env -> evar_map -> Names.inductive -> evar_map * Constr.pinductiveval fresh_constructor_instance : ?loc:Loc.t -> Environ.env -> evar_map -> Names.constructor -> evar_map * Constr.pconstructorval fresh_global : ?loc:Loc.t -> ?rigid:rigid -> ?names:Univ.Instance.t -> Environ.env -> evar_map -> Names.GlobRef.t -> evar_map * econstr
Summary names
val evar_counter_summary_tag : int Summary.Dyn.tagval create_evar_defs : evar_map -> evar_mapDeprecated functions
module MiniEConstr : sig ... endUse this module only to bootstrap EConstr