Module UGraph
val set_type_in_type : bool -> t -> t
val type_in_type : t -> bool
When
type_in_type
, functions adding constraints do not fail and may instead ignore inconsistent constraints.Checking functions such as
check_leq
always returntrue
.
type 'a check_function
= t -> 'a -> 'a -> bool
val check_leq : Univ.Universe.t check_function
val check_eq : Univ.Universe.t check_function
val check_eq_level : Univ.Level.t check_function
val initial_universes : t
The initial graph of universes: Prop < Set
val initial_universes_with : t -> t
Initial universes, but keeping options such as type in type from the argument.
val check_eq_instances : Univ.Instance.t check_function
Check equality of instances w.r.t. a universe graph
...
val enforce_constraint : Univ.univ_constraint -> t -> t
val merge_constraints : Univ.Constraint.t -> t -> t
val check_constraint : t -> Univ.univ_constraint -> bool
val check_constraints : Univ.Constraint.t -> t -> bool
val enforce_leq_alg : Univ.Universe.t -> Univ.Universe.t -> t -> Univ.Constraint.t * t
Picks an arbitrary set of constraints sufficient to ensure
u <= v
.
module Bound : sig ... end
val add_universe : Univ.Level.t -> lbound:Bound.t -> strict:bool -> t -> t
val add_universe_unconstrained : Univ.Level.t -> t -> t
Add a universe without (Prop,Set) <= u
exception
UndeclaredLevel of Univ.Level.t
Check that the universe levels are declared. Otherwise
- raises UndeclaredLevel
l for the first undeclared level found.
val check_declared_universes : t -> Univ.LSet.t -> unit
val empty_universes : t
The empty graph of universes
val constraints_of_universes : t -> Univ.Constraint.t * Univ.LSet.t list
constraints_of_universes g
returnscsts
andpartition
wherecsts
are the non-Eq constraints andpartition
is the partition of the universes into equivalence classes.
val choose : (Univ.Level.t -> bool) -> t -> Univ.Level.t -> Univ.Level.t option
choose p g u
picks a universe verifyingp
and equal tou
ing
.
val constraints_for : kept:Univ.LSet.t -> t -> Univ.Constraint.t
constraints_for ~kept g
returns the constraints about the universeskept
ing
up to transitivity.eg if
g
isa <= b <= c
thenconstraints_for ~kept:{a, c} g
isa <= c
.
val domain : t -> Univ.LSet.t
Known universes
val check_subtype : lbound:Bound.t -> Univ.AUContext.t check_function
check_subtype univ ctx1 ctx2
checks whetherctx2
is an instance ofctx1
.
Dumping
type node
=
|
Alias of Univ.Level.t
|
Node of bool Univ.LMap.t
Nodes v s.t. u < v (true) or u <= v (false)
val repr : t -> node Univ.LMap.t
Pretty-printing of universes.
val pr_universes : (Univ.Level.t -> Pp.t) -> node Univ.LMap.t -> Pp.t
val check_universes_invariants : t -> unit
Debugging