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 universes_entry
=
|
Monomorphic_entry of Univ.ContextSet.t
|
Polymorphic_entry of UVars.UContext.t
type t
Type of universe unification states. They allow the incremental building of universe constraints during an interactive proof.
Constructors
val empty : t
val make : lbound:UGraph.Bound.t -> UGraph.t -> t
val make_with_initial_binders : lbound:UGraph.Bound.t -> UGraph.t -> Names.lident list -> t
val from_env : ?binders:Names.lident list -> Environ.env -> t
Main entry point at the beginning of a declaration declaring the binding names as rigid universes.
val of_names : (UnivNames.universe_binders * UnivNames.rev_binders) -> t
Main entry point when only names matter, e.g. for printing.
val of_context_set : UnivGen.sort_context_set -> t
Main entry point when starting from the instance of a global reference, e.g. when building a scheme.
Projections and other destructors
val context_set : t -> Univ.ContextSet.t
The local context of the state, i.e. a set of bound variables together with their associated constraints.
val sort_context_set : t -> UnivGen.sort_context_set
type universe_opt_subst
= UnivFlex.t
val subst : t -> UnivFlex.t
The local universes that are unification variables
val is_algebraic : Univ.Level.t -> t -> bool
Can this universe be instantiated with an algebraic universe (ie it appears in inferred types only).
val constraints : t -> Univ.Constraints.t
Shorthand for
context_set
composed withContextSet
.constraints.
val context : t -> UVars.UContext.t
Shorthand for
context_set
withContext_set
.to_context.
type named_universes_entry
= universes_entry * UnivNames.universe_binders
val univ_entry : poly:bool -> t -> named_universes_entry
Pick from
context
orcontext_set
based onpoly
.
val universe_binders : t -> UnivNames.universe_binders
Return local names of universes.
val nf_qvar : t -> Sorts.QVar.t -> Sorts.Quality.t
Returns the normal form of the sort variable.
val nf_quality : t -> Sorts.Quality.t -> Sorts.Quality.t
val nf_instance : t -> UVars.Instance.t -> UVars.Instance.t
val nf_level : t -> Univ.Level.t -> Univ.Level.t
Must not be allowed to be algebraic
val nf_universe : t -> Univ.Universe.t -> Univ.Universe.t
val nf_sort : t -> Sorts.t -> Sorts.t
Returns the normal form of the sort.
val nf_relevance : t -> Sorts.relevance -> Sorts.relevance
Returns the normal form of the relevance.
Constraints handling
val add_constraints : t -> Univ.Constraints.t -> t
- raises UniversesDiffer
when universes differ
val add_universe_constraints : t -> UnivProblem.Set.t -> t
- raises UniversesDiffer
when universes differ
val check_qconstraints : t -> Sorts.QConstraints.t -> bool
val check_universe_constraints : t -> UnivProblem.Set.t -> bool
val add_quconstraints : t -> Sorts.QUConstraints.t -> t
Names
val quality_of_name : t -> Names.Id.t -> Sorts.QVar.t
val universe_of_name : t -> Names.Id.t -> Univ.Level.t
Retrieve the universe associated to the name.
val name_level : Univ.Level.t -> Names.Id.t -> t -> t
Gives a name to the level (making it a binder). Asserts the name is not already used by a level
Unification
val restrict_universe_context : lbound:UGraph.Bound.t -> Univ.ContextSet.t -> Univ.Level.Set.t -> Univ.ContextSet.t
restrict_universe_context lbound (univs,csts) keep
restrictsunivs
to the universes inkeep
. The constraintscsts
are adjusted so that transitive constraints between remaining universes (those inkeep
and those not inunivs
) are preserved.
val restrict : t -> Univ.Level.Set.t -> t
restrict uctx ctx
restricts the local universes ofuctx
toctx
extended by local named universes and side effect universes (fromdemote_seff_univs
). Transitive constraints between retained universes are preserved.
val restrict_even_binders : t -> Univ.Level.Set.t -> t
restrict_even_binders uctx ctx
restricts the local universes ofuctx
toctx
extended by side effect universes (fromdemote_seff_univs
). Transitive constraints between retained universes are preserved.
val univ_rigid : rigid
val univ_flexible : rigid
val univ_flexible_alg : rigid
val merge : ?loc:Loc.t -> sideff:bool -> rigid -> t -> Univ.ContextSet.t -> t
val merge_sort_variables : ?loc:Loc.t -> sideff:bool -> t -> Sorts.QVar.Set.t -> t
val merge_sort_context : ?loc:Loc.t -> sideff:bool -> rigid -> t -> UnivGen.sort_context_set -> t
val emit_side_effects : Safe_typing.private_constants -> t -> t
val demote_global_univs : Environ.env -> t -> t
Removes from the uctx_local part of the UState the universes and constraints that are present in the universe graph in the input env (supposedly the global ones)
val demote_seff_univs : Univ.Level.Set.t -> t -> t
Mark 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_sort_variable : ?loc:Loc.t -> ?name:Names.Id.t -> t -> t * Sorts.QVar.t
Declare a new local sort.
val new_univ_variable : ?loc:Loc.t -> rigid -> Names.Id.t option -> t -> t * Univ.Level.t
Declare a new local universe; use rigid if a global or bound universe; use flexible for a universe existential variable; use univ_flexible_alg for a universe existential variable allowed to be instantiated with an algebraic universe
val add_global_univ : t -> Univ.Level.t -> t
val make_nonalgebraic_variable : t -> Univ.Level.t -> t
cf UnivFlex
val normalize_variables : t -> t
val constrain_variables : Univ.Level.Set.t -> t -> t
val fix_undefined_variables : t -> t
cf UnivFlex
type ('a, 'b, 'c) gen_universe_decl
=
{
univdecl_qualities : 'a;
univdecl_instance : 'b;
univdecl_extensible_instance : bool;
univdecl_constraints : 'c;
univdecl_extensible_constraints : bool;
}
type universe_decl
= (Sorts.QVar.t list, Univ.Level.t list, Univ.Constraints.t) gen_universe_decl
val default_univ_decl : universe_decl
val check_univ_decl : poly:bool -> t -> universe_decl -> named_universes_entry
check_univ_decl ctx decl
If non extensible in
decl
, check that the local universes (resp. universe constraints) inctx
are implied bydecl
.Return a
universes_entry
containing the local universes ofctx
and their constraints.When polymorphic, the universes corresponding to
decl.univdecl_instance
come first in the order defined by that list.
val check_mono_univ_decl : t -> universe_decl -> Univ.ContextSet.t
TODO: Document me
Pretty-printing
val pr_uctx_level : t -> Univ.Level.t -> Pp.t
val pr_uctx_qvar : t -> Sorts.QVar.t -> Pp.t
val qualid_of_level : t -> Univ.Level.t -> Libnames.qualid option
val id_of_level : t -> Univ.Level.t -> Names.Id.t option
Only looks in the local names, not in the nametab.
val id_of_qvar : t -> Sorts.QVar.t -> Names.Id.t option
val pr_weak : (Univ.Level.t -> Pp.t) -> t -> Pp.t
val pr_sort_opt_subst : t -> Pp.t
module Internal : sig ... end