Module Tac2dyn.Val
Toplevel values
val create : string -> 'a tag
create n
returns a tag describing a type calledn
.create
raises an exception ifn
is already registered. Type names are hashed, socreate
may raise even if no type with the exact same name was registered due to a collision.
val anonymous : int -> 'a tag
anonymous i
returns a tag describing ani
-th anonymous type. Ifanonymous
is not used together withcreate
,max_int
anonymous types are available.anonymous
raises an exception ifi
is already registered.
val eq : 'a tag -> 'b tag -> ('a, 'b) CSig.eq option
eq t1 t2
returnsSome witness
ift1
is the same ast2
,None
otherwise.
val repr : 'a tag -> string
repr tag
returns the name of the type represented bytag
.
val dump : unit -> (int * string) list
dump ()
returns a list of (tag, name) pairs for every type tag registered in thisDyn.Make
instance.
val name : string -> any option
name n
returnsSome t
where t is a boxed tag previously registered withcreate n
, orNone
if there is no such tag.
module Map : functor (Value : Dyn.ValueS) -> Dyn.MapS with type 'a key = 'a tag and type 'a value = 'a Value.t
Map from type tags to values parameterized by the tag type
module HMap : functor (V1 : Dyn.ValueS) -> functor (V2 : Dyn.ValueS) -> sig ... end
module Easy : sig ... end