Module UState
This file defines universe unification states which are part of evarmaps. Most of the API below is reexported in Evd. Consider using higher-level primitives when needed.
type tType of universe unification states. They allow the incremental building of universe constraints during an interactive proof.
Constructors
val empty : tval make : lbound:Univ.Level.t -> UGraph.t -> tval make_with_initial_binders : lbound:Univ.Level.t -> UGraph.t -> Names.lident list -> tval is_empty : t -> boolval union : t -> t -> tval of_context_set : Univ.ContextSet.t -> tval of_binders : UnivNames.universe_binders -> tval universe_binders : t -> UnivNames.universe_binders
Projections
val context_set : t -> Univ.ContextSet.tThe local context of the state, i.e. a set of bound variables together with their associated constraints.
val subst : t -> UnivSubst.universe_opt_substThe local universes that are unification variables
val algebraics : t -> Univ.LSet.tThe subset of unification variables that can be instantiated with algebraic universes as they appear in inferred types only.
val constraints : t -> Univ.Constraint.tShorthand for
context_setcomposed withContextSet.constraints.
val context : t -> Univ.UContext.tShorthand for
context_setwithContext_set.to_context.
val univ_entry : poly:bool -> t -> Entries.universes_entryPick from
contextorcontext_setbased onpoly.
Constraints handling
val drop_weak_constraints : bool Stdlib.refval add_constraints : t -> Univ.Constraint.t -> t- raises UniversesDiffer
when universes differ
val add_universe_constraints : t -> UnivProblem.Set.t -> t- raises UniversesDiffer
when universes differ
Names
val universe_of_name : t -> Names.Id.t -> Univ.Level.tRetrieve the universe associated to the name.
Unification
val restrict_universe_context : lbound:Univ.Level.t -> Univ.ContextSet.t -> Univ.LSet.t -> Univ.ContextSet.trestrict_universe_context lbound (univs,csts) keeprestrictsunivsto the universes inkeep. The constraintscstsare adjusted so that transitive constraints between remaining universes (those inkeepand those not inunivs) are preserved.
val restrict : t -> Univ.LSet.t -> trestrict uctx ctxrestricts the local universes ofuctxtoctxextended by local named universes and side effect universes (fromdemote_seff_univs). Transitive constraints between retained universes are preserved.
val univ_rigid : rigidval univ_flexible : rigidval univ_flexible_alg : rigidval merge : ?loc:Loc.t -> sideff:bool -> rigid -> t -> Univ.ContextSet.t -> tval merge_subst : t -> UnivSubst.universe_opt_subst -> tval emit_side_effects : Safe_typing.private_constants -> t -> tval demote_seff_univs : Univ.LSet.t -> t -> tMark the universes as not local any more, because they have been globally declared by some side effect. You should be using emit_side_effects instead.
val new_univ_variable : ?loc:Loc.t -> rigid -> Names.Id.t option -> t -> t * Univ.Level.tval add_global_univ : t -> Univ.Level.t -> tval make_flexible_variable : t -> algebraic:bool -> Univ.Level.t -> tmake_flexible_variable g algebraic lTurn the variable
lflexible, and algebraic ifalgebraicis true andlcan be. That is if there are no strict upper constraints onland and it does not appear in the instance of any non-algebraic universe. Otherwise the variable is just made flexible.If
lis already algebraic it will remain so even withalgebraic:false.
val make_nonalgebraic_variable : t -> Univ.Level.t -> tMake the level non algebraic. Undefined behaviour on already-defined algebraics.
val make_flexible_nonalgebraic : t -> tTurn all undefined flexible algebraic variables into simply flexible ones. Can be used in case the variables might appear in universe instances (typically for polymorphic program obligations).
val is_sort_variable : t -> Sorts.t -> Univ.Level.t optionval normalize_variables : t -> Univ.universe_subst * tval constrain_variables : Univ.LSet.t -> t -> tval abstract_undefined_variables : t -> tval fix_undefined_variables : t -> tval refresh_undefined_univ_variables : t -> t * Univ.universe_level_substval minimize : t -> tUniverse minimization
type ('a, 'b) gen_universe_decl={univdecl_instance : 'a;univdecl_extensible_instance : bool;univdecl_constraints : 'b;univdecl_extensible_constraints : bool;}type universe_decl= (Names.lident list, Univ.Constraint.t) gen_universe_decl
val default_univ_decl : universe_declval check_univ_decl : poly:bool -> t -> universe_decl -> Entries.universes_entrycheck_univ_decl ctx declIf non extensible in
decl, check that the local universes (resp. universe constraints) inctxare implied bydecl.Return a
Entries.constant_universes_entrycontaining the local universes ofctxand their constraints.When polymorphic, the universes corresponding to
decl.univdecl_instancecome first in the order defined by that list.
val check_mono_univ_decl : t -> universe_decl -> Univ.ContextSet.t
TODO: Document me
val update_sigma_env : t -> Environ.env -> t
Pretty-printing
val pr_uctx_level : t -> Univ.Level.t -> Pp.tval qualid_of_level : t -> Univ.Level.t -> Libnames.qualid optionval pr_weak : (Univ.Level.t -> Pp.t) -> t -> Pp.t