src/Tools/Metis/src/Name.sig
author wenzelm
Mon, 15 Oct 2012 15:28:56 +0200
changeset 49855 b3110dec1a32
parent 42102 fcfd07f122d4
child 72004 913162a47d9f
permissions -rw-r--r--
updated to polyml-5.5.0 which reduces chance of HOL-IMP failure (although it is hard to reproduce anyway);

(* ========================================================================= *)
(* NAMES                                                                     *)
(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
(* ========================================================================= *)

signature Name =
sig

(* ------------------------------------------------------------------------- *)
(* A type of names.                                                          *)
(* ------------------------------------------------------------------------- *)

type name

(* ------------------------------------------------------------------------- *)
(* A total ordering.                                                         *)
(* ------------------------------------------------------------------------- *)

val compare : name * name -> order

val equal : name -> name -> bool

(* ------------------------------------------------------------------------- *)
(* Fresh names.                                                              *)
(* ------------------------------------------------------------------------- *)

val newName : unit -> name

val newNames : int -> name list

val variantPrime : {avoid : name -> bool} -> name -> name

val variantNum : {avoid : name -> bool} -> name -> name

(* ------------------------------------------------------------------------- *)
(* Parsing and pretty printing.                                              *)
(* ------------------------------------------------------------------------- *)

val pp : name Print.pp

val toString : name -> string

val fromString : string -> name

end