src/Pure/General/properties.ML
author wenzelm
Mon, 24 Feb 2020 20:57:29 +0100
changeset 71465 910a081cca74
parent 51665 cba83c9f72b9
child 77772 7e0d920b4e6e
permissions -rw-r--r--
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;

(*  Title:      Pure/General/properties.ML
    Author:     Makarius

Property lists.
*)

signature PROPERTIES =
sig
  type entry = string * string
  type T = entry list
  val entry_ord: entry ord
  val ord: T ord
  val defined: T -> string -> bool
  val get: T -> string -> string option
  val put: entry -> T -> T
  val remove: string -> T -> T
  val seconds: T -> string -> Time.time
end;

structure Properties: PROPERTIES =
struct

type entry = string * string;
type T = entry list;

val entry_ord = prod_ord string_ord string_ord;
val ord = list_ord entry_ord;

fun defined (props: T) name = AList.defined (op =) props name;
fun get (props: T) name = AList.lookup (op =) props name;
fun put entry (props: T) = AList.update (op =) entry props;
fun remove name (props: T) = AList.delete (op =) name props;

fun seconds props name =
  (case AList.lookup (op =) props name of
    NONE => Time.zeroTime
  | SOME s => Time.fromReal (the_default 0.0 (Real.fromString s)));

end;