--- a/doc-src/Nitpick/nitpick.tex Mon Sep 13 16:44:20 2010 +0200
+++ b/doc-src/Nitpick/nitpick.tex Mon Sep 13 21:24:10 2010 +0200
@@ -121,11 +121,10 @@
in Isabelle's \texttt{src/HOL/Nitpick\_Examples/Manual\_Nits.thy} theory.
Throughout this manual, we will explicitly invoke the \textbf{nitpick} command.
-Nitpick also provides an automatic mode that can be enabled using the
-``Auto Nitpick'' option from the ``Isabelle'' menu in Proof General. In this
-mode, Nitpick is run on every newly entered theorem, much like Auto Quickcheck.
-The collective time limit for Auto Nitpick and Auto Quickcheck can be set using
-the ``Auto Counterexample Time Limit'' option.
+Nitpick also provides an automatic mode that can be enabled via the ``Auto
+Nitpick'' option from the ``Isabelle'' menu in Proof General. In this mode,
+Nitpick is run on every newly entered theorem. The time limit for Auto Nitpick
+and other automatic tools can be set using the ``Auto Tools Time Limit'' option.
\newbox\boxA
\setbox\boxA=\hbox{\texttt{nospam}}
@@ -1883,14 +1882,15 @@
You can instruct Nitpick to run automatically on newly entered theorems by
enabling the ``Auto Nitpick'' option from the ``Isabelle'' menu in Proof
-General. For automatic runs, \textit{user\_axioms} (\S\ref{mode-of-operation})
-and \textit{assms} (\S\ref{mode-of-operation}) are implicitly enabled,
-\textit{blocking} (\S\ref{mode-of-operation}), \textit{verbose}
-(\S\ref{output-format}), and \textit{debug} (\S\ref{output-format}) are
-disabled, \textit{max\_potential} (\S\ref{output-format}) is taken to be 0, and
-\textit{timeout} (\S\ref{timeouts}) is superseded by the ``Auto Counterexample
-Time Limit'' in Proof General's ``Isabelle'' menu. Nitpick's output is also more
-concise.
+General. For automatic runs, \textit{user\_axioms} (\S\ref{mode-of-operation}),
+\textit{assms} (\S\ref{mode-of-operation}), and \textit{mono}
+(\S\ref{scope-of-search}) are implicitly enabled, \textit{blocking}
+(\S\ref{mode-of-operation}), \textit{verbose} (\S\ref{output-format}), and
+\textit{debug} (\S\ref{output-format}) are disabled, \textit{max\_threads}
+(\S\ref{optimizations}) is taken to be 1, \textit{max\_potential}
+(\S\ref{output-format}) is taken to be 0, and \textit{timeout}
+(\S\ref{timeouts}) is superseded by the ``Auto Tools Time Limit'' in
+Proof General's ``Isabelle'' menu. Nitpick's output is also more concise.
The number of options can be overwhelming at first glance. Do not let that worry
you: Nitpick's defaults have been chosen so that it almost always does the right
@@ -2171,7 +2171,8 @@
scopes and finitizing types. If the option is set to \textit{smart}, Nitpick
performs a monotonicity check on the type. Setting this option to \textit{true}
can reduce the number of scopes tried, but it can also diminish the chance of
-finding a counterexample, as demonstrated in \S\ref{scope-monotonicity}.
+finding a counterexample, as demonstrated in \S\ref{scope-monotonicity}. The
+option is implicitly set to \textit{true} for automatic runs.
\nopagebreak
{\small See also \textit{card} (\S\ref{scope-of-search}),
@@ -2527,7 +2528,7 @@
\opdefault{max\_threads}{int}{0}
Specifies the maximum number of threads to use in Kodkod. If this option is set
to 0, Kodkod will compute an appropriate value based on the number of processor
-cores available.
+cores available. The option is implicitly set to 1 for automatic runs.
\nopagebreak
{\small See also \textit{batch\_size} (\S\ref{optimizations}) and
--- a/doc-src/Sledgehammer/sledgehammer.tex Mon Sep 13 16:44:20 2010 +0200
+++ b/doc-src/Sledgehammer/sledgehammer.tex Mon Sep 13 21:24:10 2010 +0200
@@ -97,6 +97,13 @@
inferences going through the kernel. Thus its results are correct by
construction.
+In this manual, we will explicitly invoke the \textbf{sledgehammer} command.
+Sledgehammer also provides an automatic mode that can be enabled via the
+``Auto Sledgehammer'' option from the ``Isabelle'' menu in Proof General. In
+this mode, Sledgehammer is run on every newly entered theorem. The time limit
+for Auto Sledgehammer and other automatic tools can be set using the ``Auto
+Tools Time Limit'' option.
+
\newbox\boxA
\setbox\boxA=\hbox{\texttt{nospam}}
@@ -248,10 +255,10 @@
you may get better results if you first simplify the problem to remove
higher-order features.
-Note that problems can be easy for auto and difficult for ATPs, but the reverse
-is also true, so don't be discouraged if your first attempts fail. Because the
-system refers to all theorems known to Isabelle, it is particularly suitable
-when your goal has a short proof from lemmas that you don't know about.
+Note that problems can be easy for \textit{auto} and difficult for ATPs, but the
+reverse is also true, so don't be discouraged if your first attempts fail.
+Because the system refers to all theorems known to Isabelle, it is particularly
+suitable when your goal has a short proof from lemmas that you don't know about.
\section{Command Syntax}
\label{command-syntax}
@@ -319,12 +326,21 @@
through the relevance filter. It may be of the form ``(\textit{facts})'', where
\textit{facts} is a space-separated list of Isabelle facts (theorems, local
assumptions, etc.), in which case the relevance filter is bypassed and the given
-facts are used. It may also be of the form (\textit{add}:\ \textit{facts}$_1$),
-(\textit{del}:\ \textit{facts}$_2$), or (\textit{add}:\ \textit{facts}$_1$\
-\textit{del}:\ \textit{facts}$_2$), where the relevance filter is instructed to
+facts are used. It may also be of the form ``(\textit{add}:\ \textit{facts}$_1$)'',
+``(\textit{del}:\ \textit{facts}$_2$)'', or ``(\textit{add}:\ \textit{facts}$_1$\
+\textit{del}:\ \textit{facts}$_2$)'', where the relevance filter is instructed to
proceed as usual except that it should consider \textit{facts}$_1$
highly-relevant and \textit{facts}$_2$ fully irrelevant.
+You can instruct Sledgehammer to run automatically on newly entered theorems by
+enabling the ``Auto Sledgehammer'' option from the ``Isabelle'' menu in Proof
+General. For automatic runs, only the first ATP set using \textit{atps}
+(\S\ref{mode-of-operation}) is considered, \textit{verbose}
+(\S\ref{output-format}) and \textit{debug} (\S\ref{output-format}) are disabled,
+fewer facts are passed to the ATP, and \textit{timeout} (\S\ref{timeouts}) is
+superseded by the ``Auto Tools Time Limit'' in Proof General's ``Isabelle''
+menu. Sledgehammer's output is also more concise.
+
\section{Option Reference}
\label{option-reference}
@@ -418,7 +434,7 @@
\opnodefault{atp}{string}
Alias for \textit{atps}.
-\opdefault{timeout}{time}{$\mathbf{60}$ s}
+\opdefault{timeout}{time}{$\mathbf{30}$ s}
Specifies the maximum amount of time that the ATPs should spend searching for a
proof. For historical reasons, the default value of this option can be
overridden using the option ``Sledgehammer: Time Limit'' from the ``Isabelle''
--- a/src/HOL/HOL.thy Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/HOL.thy Mon Sep 13 21:24:10 2010 +0200
@@ -2006,6 +2006,10 @@
*} "solve goal by normalization"
+subsection {* Try *}
+
+setup {* Try.setup *}
+
subsection {* Counterexample Search Units *}
subsubsection {* Quickcheck *}
--- a/src/HOL/IsaMakefile Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/IsaMakefile Mon Sep 13 21:24:10 2010 +0200
@@ -122,8 +122,8 @@
$(SRC)/Tools/IsaPlanner/rw_tools.ML \
$(SRC)/Tools/IsaPlanner/zipper.ML \
$(SRC)/Tools/atomize_elim.ML \
- $(SRC)/Tools/auto_counterexample.ML \
$(SRC)/Tools/auto_solve.ML \
+ $(SRC)/Tools/auto_tools.ML \
$(SRC)/Tools/coherent.ML \
$(SRC)/Tools/cong_tac.ML \
$(SRC)/Tools/eqsubst.ML \
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Mon Sep 13 21:24:10 2010 +0200
@@ -24,6 +24,8 @@
datatype sh_data = ShData of {
calls: int,
success: int,
+ nontriv_calls: int,
+ nontriv_success: int,
lemmas: int,
max_lems: int,
time_isa: int,
@@ -33,11 +35,13 @@
datatype me_data = MeData of {
calls: int,
success: int,
+ nontriv_calls: int,
+ nontriv_success: int,
proofs: int,
time: int,
timeout: int,
lemmas: int * int * int,
- posns: Position.T list
+ posns: (Position.T * bool) list
}
datatype min_data = MinData of {
@@ -46,29 +50,35 @@
}
fun make_sh_data
- (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail) =
- ShData{calls=calls, success=success, lemmas=lemmas, max_lems=max_lems,
+ (calls,success,nontriv_calls,nontriv_success,lemmas,max_lems,time_isa,
+ time_atp,time_atp_fail) =
+ ShData{calls=calls, success=success, nontriv_calls=nontriv_calls,
+ nontriv_success=nontriv_success, lemmas=lemmas, max_lems=max_lems,
time_isa=time_isa, time_atp=time_atp, time_atp_fail=time_atp_fail}
fun make_min_data (succs, ab_ratios) =
MinData{succs=succs, ab_ratios=ab_ratios}
-fun make_me_data (calls,success,proofs,time,timeout,lemmas,posns) =
- MeData{calls=calls, success=success, proofs=proofs, time=time,
+fun make_me_data (calls,success,nontriv_calls,nontriv_success,proofs,time,
+ timeout,lemmas,posns) =
+ MeData{calls=calls, success=success, nontriv_calls=nontriv_calls,
+ nontriv_success=nontriv_success, proofs=proofs, time=time,
timeout=timeout, lemmas=lemmas, posns=posns}
-val empty_sh_data = make_sh_data (0, 0, 0, 0, 0, 0, 0)
+val empty_sh_data = make_sh_data (0, 0, 0, 0, 0, 0, 0, 0, 0)
val empty_min_data = make_min_data (0, 0)
-val empty_me_data = make_me_data (0, 0, 0, 0, 0, (0,0,0), [])
+val empty_me_data = make_me_data (0, 0, 0, 0, 0, 0, 0, (0,0,0), [])
-fun tuple_of_sh_data (ShData {calls, success, lemmas, max_lems, time_isa,
- time_atp, time_atp_fail}) = (calls, success, lemmas, max_lems, time_isa,
- time_atp, time_atp_fail)
+fun tuple_of_sh_data (ShData {calls, success, nontriv_calls, nontriv_success,
+ lemmas, max_lems, time_isa,
+ time_atp, time_atp_fail}) = (calls, success, nontriv_calls, nontriv_success,
+ lemmas, max_lems, time_isa, time_atp, time_atp_fail)
fun tuple_of_min_data (MinData {succs, ab_ratios}) = (succs, ab_ratios)
-fun tuple_of_me_data (MeData {calls, success, proofs, time, timeout, lemmas,
- posns}) = (calls, success, proofs, time, timeout, lemmas, posns)
+fun tuple_of_me_data (MeData {calls, success, nontriv_calls, nontriv_success,
+ proofs, time, timeout, lemmas, posns}) = (calls, success, nontriv_calls,
+ nontriv_success, proofs, time, timeout, lemmas, posns)
datatype metis = Unminimized | Minimized | UnminimizedFT | MinimizedFT
@@ -117,32 +127,40 @@
fun inc_max (n:int) (s,sos,m) = (s+n, sos + n*n, Int.max(m,n));
val inc_sh_calls = map_sh_data
- (fn (calls, success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
- => (calls + 1, success, lemmas,max_lems, time_isa, time_atp, time_atp_fail))
+ (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
+ => (calls + 1, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_atp, time_atp_fail))
val inc_sh_success = map_sh_data
- (fn (calls, success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
- => (calls, success + 1, lemmas,max_lems, time_isa, time_atp, time_atp_fail))
+ (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
+ => (calls, success + 1, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_atp, time_atp_fail))
+
+val inc_sh_nontriv_calls = map_sh_data
+ (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
+ => (calls, success, nontriv_calls + 1, nontriv_success, lemmas, max_lems, time_isa, time_atp, time_atp_fail))
+
+val inc_sh_nontriv_success = map_sh_data
+ (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
+ => (calls, success, nontriv_calls, nontriv_success + 1, lemmas,max_lems, time_isa, time_atp, time_atp_fail))
fun inc_sh_lemmas n = map_sh_data
- (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
- => (calls,success,lemmas+n,max_lems,time_isa,time_atp,time_atp_fail))
+ (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp,time_atp_fail)
+ => (calls,success,nontriv_calls, nontriv_success, lemmas+n,max_lems,time_isa,time_atp,time_atp_fail))
fun inc_sh_max_lems n = map_sh_data
- (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
- => (calls,success,lemmas,Int.max(max_lems,n),time_isa,time_atp,time_atp_fail))
+ (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp,time_atp_fail)
+ => (calls,success,nontriv_calls, nontriv_success, lemmas,Int.max(max_lems,n),time_isa,time_atp,time_atp_fail))
fun inc_sh_time_isa t = map_sh_data
- (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
- => (calls,success,lemmas,max_lems,time_isa + t,time_atp,time_atp_fail))
+ (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp,time_atp_fail)
+ => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa + t,time_atp,time_atp_fail))
fun inc_sh_time_atp t = map_sh_data
- (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
- => (calls,success,lemmas,max_lems,time_isa,time_atp + t,time_atp_fail))
+ (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp,time_atp_fail)
+ => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp + t,time_atp_fail))
fun inc_sh_time_atp_fail t = map_sh_data
- (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
- => (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail + t))
+ (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp,time_atp_fail)
+ => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp,time_atp_fail + t))
val inc_min_succs = map_min_data
(fn (succs,ab_ratios) => (succs+1, ab_ratios))
@@ -151,32 +169,40 @@
(fn (succs, ab_ratios) => (succs, ab_ratios+r))
val inc_metis_calls = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
- => (calls + 1, success, proofs, time, timeout, lemmas,posns))
+ (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
+ => (calls + 1, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns))
val inc_metis_success = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
- => (calls, success + 1, proofs, time, timeout, lemmas,posns))
+ (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
+ => (calls, success + 1, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns))
+
+val inc_metis_nontriv_calls = map_me_data
+ (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
+ => (calls, success, nontriv_calls + 1, nontriv_success, proofs, time, timeout, lemmas,posns))
+
+val inc_metis_nontriv_success = map_me_data
+ (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
+ => (calls, success, nontriv_calls, nontriv_success + 1, proofs, time, timeout, lemmas,posns))
val inc_metis_proofs = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
- => (calls, success, proofs + 1, time, timeout, lemmas,posns))
+ (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
+ => (calls, success, nontriv_calls, nontriv_success, proofs + 1, time, timeout, lemmas,posns))
fun inc_metis_time m t = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
- => (calls, success, proofs, time + t, timeout, lemmas,posns)) m
+ (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
+ => (calls, success, nontriv_calls, nontriv_success, proofs, time + t, timeout, lemmas,posns)) m
val inc_metis_timeout = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
- => (calls, success, proofs, time, timeout + 1, lemmas,posns))
+ (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
+ => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout + 1, lemmas,posns))
fun inc_metis_lemmas m n = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
- => (calls, success, proofs, time, timeout, inc_max n lemmas, posns)) m
+ (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
+ => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, inc_max n lemmas, posns)) m
fun inc_metis_posns m pos = map_me_data
- (fn (calls,success,proofs,time,timeout,lemmas,posns)
- => (calls, success, proofs, time, timeout, lemmas, pos::posns)) m
+ (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
+ => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas, pos::posns)) m
local
@@ -188,12 +214,14 @@
if n > 0 then (Real.fromInt t / 1000.0) / Real.fromInt n else 0.0
fun log_sh_data log
- (calls, success, lemmas, max_lems, time_isa, time_atp, time_atp_fail) =
+ (calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_atp, time_atp_fail) =
(log ("Total number of sledgehammer calls: " ^ str calls);
log ("Number of successful sledgehammer calls: " ^ str success);
log ("Number of sledgehammer lemmas: " ^ str lemmas);
log ("Max number of sledgehammer lemmas: " ^ str max_lems);
log ("Success rate: " ^ percentage success calls ^ "%");
+ log ("Total number of nontrivial sledgehammer calls: " ^ str nontriv_calls);
+ log ("Number of successful nontrivial sledgehammer calls: " ^ str nontriv_success);
log ("Total time for sledgehammer calls (Isabelle): " ^ str3 (time time_isa));
log ("Total time for successful sledgehammer calls (ATP): " ^ str3 (time time_atp));
log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time time_atp_fail));
@@ -206,18 +234,25 @@
)
-fun str_of_pos pos =
+fun str_of_pos (pos, triv) =
let val str0 = string_of_int o the_default 0
- in str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos) end
+ in
+ str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos) ^
+ (if triv then "[T]" else "")
+ end
fun log_me_data log tag sh_calls (metis_calls, metis_success,
- metis_proofs, metis_time, metis_timeout, (lemmas, lems_sos, lems_max),
+ metis_nontriv_calls, metis_nontriv_success,
+ metis_proofs, metis_time, metis_timeout, (lemmas, lems_sos, lems_max),
metis_posns) =
(log ("Total number of " ^ tag ^ "metis calls: " ^ str metis_calls);
log ("Number of successful " ^ tag ^ "metis calls: " ^ str metis_success ^
" (proof: " ^ str metis_proofs ^ ")");
log ("Number of " ^ tag ^ "metis timeouts: " ^ str metis_timeout);
log ("Success rate: " ^ percentage metis_success sh_calls ^ "%");
+ log ("Total number of nontrivial " ^ tag ^ "metis calls: " ^ str metis_nontriv_calls);
+ log ("Number of successful nontrivial " ^ tag ^ "metis calls: " ^ str metis_nontriv_success ^
+ " (proof: " ^ str metis_proofs ^ ")");
log ("Number of successful " ^ tag ^ "metis lemmas: " ^ str lemmas);
log ("SOS of successful " ^ tag ^ "metis lemmas: " ^ str lems_sos);
log ("Max number of successful " ^ tag ^ "metis lemmas: " ^ str lems_max);
@@ -304,11 +339,12 @@
val {context = ctxt, facts = chained_ths, goal} = Proof.goal st
val thy = ProofContext.theory_of ctxt
val i = 1
- val change_dir = (fn SOME d => Config.put Sledgehammer.dest_dir d | _ => I)
- val ctxt' =
- ctxt
- |> change_dir dir
- |> Config.put Sledgehammer.measure_run_time true
+ fun change_dir (SOME dir) = Config.put Sledgehammer.dest_dir dir
+ | change_dir NONE = I
+ val st' =
+ st |> Proof.map_context
+ (change_dir dir
+ #> Config.put Sledgehammer.measure_run_time true)
val params as {full_types, relevance_thresholds, max_relevant, ...} =
Sledgehammer_Isar.default_params thy
[("timeout", Int.toString timeout ^ " s")]
@@ -320,7 +356,7 @@
(the_default default_max_relevant max_relevant)
relevance_override chained_ths hyp_ts concl_t
val problem =
- {ctxt = ctxt', goal = goal, subgoal = i,
+ {state = st', goal = goal, subgoal = i,
axioms = map (Sledgehammer_Translate.prepare_axiom ctxt) axioms}
val time_limit =
(case hard_timeout of
@@ -353,9 +389,11 @@
in
-fun run_sledgehammer args named_thms id ({pre=st, log, ...}: Mirabelle.run_args) =
+fun run_sledgehammer trivial args named_thms id ({pre=st, log, ...}: Mirabelle.run_args) =
let
+ val triv_str = if trivial then "[T] " else ""
val _ = change_data id inc_sh_calls
+ val _ = if trivial then () else change_data id inc_sh_nontriv_calls
val (prover_name, prover) = get_atp (Proof.theory_of st) args
val dir = AList.lookup (op =) args keepK
val timeout = Mirabelle.get_int_setting args (prover_timeoutK, 30)
@@ -371,19 +409,20 @@
SOME ((name, loc), thms_of_name (Proof.context_of st) name)
in
change_data id inc_sh_success;
+ if trivial then () else change_data id inc_sh_nontriv_success;
change_data id (inc_sh_lemmas (length names));
change_data id (inc_sh_max_lems (length names));
change_data id (inc_sh_time_isa time_isa);
change_data id (inc_sh_time_atp time_atp);
named_thms := SOME (map_filter get_thms names);
- log (sh_tag id ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
+ log (sh_tag id ^ triv_str ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
string_of_int time_atp ^ ") [" ^ prover_name ^ "]:\n" ^ msg)
end
| SH_FAIL (time_isa, time_atp) =>
let
val _ = change_data id (inc_sh_time_isa time_isa)
val _ = change_data id (inc_sh_time_atp_fail time_atp)
- in log (sh_tag id ^ "failed: " ^ msg) end
+ in log (sh_tag id ^ triv_str ^ "failed: " ^ msg) end
| SH_ERROR => log (sh_tag id ^ "failed: " ^ msg)
end
@@ -421,7 +460,7 @@
end
-fun run_metis full m name named_thms id
+fun run_metis trivial full m name named_thms id
({pre=st, timeout, log, pos, ...}: Mirabelle.run_args) =
let
fun metis thms ctxt =
@@ -431,9 +470,10 @@
fun with_time (false, t) = "failed (" ^ string_of_int t ^ ")"
| with_time (true, t) = (change_data id (inc_metis_success m);
+ if trivial then () else change_data id (inc_metis_nontriv_success m);
change_data id (inc_metis_lemmas m (length named_thms));
change_data id (inc_metis_time m t);
- change_data id (inc_metis_posns m pos);
+ change_data id (inc_metis_posns m (pos, trivial));
if name = "proof" then change_data id (inc_metis_proofs m) else ();
"succeeded (" ^ string_of_int t ^ ")")
fun timed_metis thms =
@@ -444,6 +484,7 @@
val _ = log separator
val _ = change_data id (inc_metis_calls m)
+ val _ = if trivial then () else change_data id (inc_metis_nontriv_calls m)
in
maps snd named_thms
|> timed_metis
@@ -451,6 +492,8 @@
|> snd
end
+val try_timeout = Time.fromSeconds 30
+
fun sledgehammer_action args id (st as {pre, name, ...}: Mirabelle.run_args) =
let val goal = Thm.major_prem_of (#goal (Proof.goal pre)) in
if can Logic.dest_conjunction goal orelse can Logic.dest_equals goal
@@ -460,22 +503,24 @@
Unsynchronized.ref (NONE : ((string * locality) * thm list) list option)
val minimize = AList.defined (op =) args minimizeK
val metis_ft = AList.defined (op =) args metis_ftK
+
+ val trivial = Try.invoke_try (SOME try_timeout) pre
fun apply_metis m1 m2 =
if metis_ft
then
if not (Mirabelle.catch_result metis_tag false
- (run_metis false m1 name (these (!named_thms))) id st)
+ (run_metis trivial false m1 name (these (!named_thms))) id st)
then
(Mirabelle.catch_result metis_tag false
- (run_metis true m2 name (these (!named_thms))) id st; ())
+ (run_metis trivial true m2 name (these (!named_thms))) id st; ())
else ()
else
(Mirabelle.catch_result metis_tag false
- (run_metis false m1 name (these (!named_thms))) id st; ())
+ (run_metis trivial false m1 name (these (!named_thms))) id st; ())
in
change_data id (set_mini minimize);
- Mirabelle.catch sh_tag (run_sledgehammer args named_thms) id st;
+ Mirabelle.catch sh_tag (run_sledgehammer trivial args named_thms) id st;
if is_some (!named_thms)
then
(apply_metis Unminimized UnminimizedFT;
--- a/src/HOL/Mutabelle/MutabelleExtra.thy Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Mutabelle/MutabelleExtra.thy Mon Sep 13 21:24:10 2010 +0200
@@ -29,7 +29,7 @@
nitpick_params [timeout = 5 s, sat_solver = MiniSat, no_overlord, verbose, card = 1-5, iter = 1,2,4,8,12]
refute_params [maxtime = 10, minsize = 1, maxsize = 5, satsolver = jerusat]
*)
-ML {* Auto_Counterexample.time_limit := 10 *}
+ML {* Auto_Tools.time_limit := 10 *}
text {* Uncomment the following ML code to check the counterexample generation with all theorems of Complex_Main. *}
@@ -53,4 +53,4 @@
ML {* Output.priority_fn := old_pr *}
ML {* Output.warning_fn := old_wa *}
-end
\ No newline at end of file
+end
--- a/src/HOL/Mutabelle/mutabelle_extra.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML Mon Sep 13 21:24:10 2010 +0200
@@ -94,13 +94,13 @@
(map_types (inst_type insts) (Mutabelle.freeze t));
fun invoke_quickcheck quickcheck_generator thy t =
- TimeLimit.timeLimit (Time.fromSeconds (! Auto_Counterexample.time_limit))
+ TimeLimit.timeLimit (Time.fromSeconds (!Auto_Tools.time_limit))
(fn _ =>
case Quickcheck.gen_test_term (ProofContext.init_global thy) true (SOME quickcheck_generator)
size iterations (preprocess thy [] t) of
(NONE, (time_report, report)) => (NoCex, (time_report, report))
| (SOME _, (time_report, report)) => (GenuineCex, (time_report, report))) ()
- handle TimeLimit.TimeOut => (Timeout, ([("timelimit", !Auto_Counterexample.time_limit)], NONE))
+ handle TimeLimit.TimeOut => (Timeout, ([("timelimit", !Auto_Tools.time_limit)], NONE))
fun quickcheck_mtd quickcheck_generator =
("quickcheck_" ^ quickcheck_generator, invoke_quickcheck quickcheck_generator)
--- a/src/HOL/Sledgehammer.thy Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Sledgehammer.thy Mon Sep 13 21:24:10 2010 +0200
@@ -102,6 +102,7 @@
use "Tools/Sledgehammer/metis_clauses.ML"
use "Tools/Sledgehammer/metis_tactics.ML"
+setup Metis_Tactics.setup
use "Tools/Sledgehammer/sledgehammer_util.ML"
use "Tools/Sledgehammer/sledgehammer_filter.ML"
@@ -111,6 +112,6 @@
setup Sledgehammer.setup
use "Tools/Sledgehammer/sledgehammer_minimize.ML"
use "Tools/Sledgehammer/sledgehammer_isar.ML"
-setup Metis_Tactics.setup
+setup Sledgehammer_Isar.setup
end
--- a/src/HOL/Tools/ATP/atp_systems.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/ATP/atp_systems.ML Mon Sep 13 21:24:10 2010 +0200
@@ -122,10 +122,6 @@
priority ("Available ATPs: " ^
commas (sort_strings (Symtab.keys (Data.get thy))) ^ ".")
-fun available_atps thy =
- priority ("Available ATPs: " ^
- commas (sort_strings (Symtab.keys (Data.get thy))) ^ ".")
-
fun to_secs bonus time = (Time.toMilliseconds time + bonus + 999) div 1000
(* E prover *)
@@ -303,9 +299,11 @@
fun maybe_remote (name, config) =
name |> not (is_installed config) ? remotify_name
+(* The first prover of the list is used by Auto Sledgehammer. Because of the low
+ timeout, it makes sense to put SPASS first. *)
fun default_atps_param_value () =
- space_implode " " ([maybe_remote e] @
- (if is_installed (snd spass) then [fst spass] else []) @
+ space_implode " " ((if is_installed (snd spass) then [fst spass] else []) @
+ [maybe_remote e] @
[if forall (is_installed o snd) [e, spass] then
remotify_name (fst vampire)
else
--- a/src/HOL/Tools/Nitpick/kodkod.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Nitpick/kodkod.ML Mon Sep 13 21:24:10 2010 +0200
@@ -1126,7 +1126,7 @@
| NONE =>
let val outcome = do_solve () in
(case outcome of
- Normal (ps, js, "") =>
+ Normal (_, _, "") =>
Synchronized.change cached_outcome
(K (SOME ((max_solutions, problems), outcome)))
| _ => ());
--- a/src/HOL/Tools/Nitpick/nitpick.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick.ML Mon Sep 13 21:24:10 2010 +0200
@@ -144,7 +144,7 @@
rel_table: nut NameTable.table,
unsound: bool,
scope: scope}
-
+
type rich_problem = KK.problem * problem_extension
fun pretties_for_formulas _ _ [] = []
@@ -224,7 +224,9 @@
fun pprint_v f = () |> verbose ? pprint o f
fun pprint_d f = () |> debug ? pprint o f
val print = pprint o curry Pretty.blk 0 o pstrs
+(*
val print_g = pprint o Pretty.str
+*)
val print_m = pprint_m o K o plazy
val print_v = pprint_v o K o plazy
@@ -552,9 +554,8 @@
(plain_bounds @ sel_bounds) formula,
main_j0 |> bits > 0 ? Integer.add (bits + 1))
val (built_in_bounds, built_in_axioms) =
- bounds_and_axioms_for_built_in_rels_in_formulas debug ofs
- univ_card nat_card int_card main_j0
- (formula :: declarative_axioms)
+ bounds_and_axioms_for_built_in_rels_in_formulas debug univ_card
+ nat_card int_card main_j0 (formula :: declarative_axioms)
val bounds = built_in_bounds @ plain_bounds @ sel_bounds
|> not debug ? merge_bounds
val axioms = built_in_axioms @ declarative_axioms
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Mon Sep 13 21:24:10 2010 +0200
@@ -595,9 +595,13 @@
else case Typedef.get_info ctxt s of
(* When several entries are returned, it shouldn't matter much which one
we take (according to Florian Haftmann). *)
+ (* The "Logic.varifyT_global" calls are a temporary hack because these
+ types's type variables sometimes clash with locally fixed type variables.
+ Remove these calls once "Typedef" is fully localized. *)
({abs_type, rep_type, Abs_name, Rep_name, ...},
{set_def, Rep, Abs_inverse, Rep_inverse, ...}) :: _ =>
- SOME {abs_type = abs_type, rep_type = rep_type, Abs_name = Abs_name,
+ SOME {abs_type = Logic.varifyT_global abs_type,
+ rep_type = Logic.varifyT_global rep_type, Abs_name = Abs_name,
Rep_name = Rep_name, set_def = set_def, prop_of_Rep = prop_of Rep,
set_name = set_prefix ^ s, Abs_inverse = SOME Abs_inverse,
Rep_inverse = SOME Rep_inverse}
@@ -776,9 +780,11 @@
| NONE => raise TERM ("Nitpick_HOL.mate_of_rep_fun", [Const x]))
| mate_of_rep_fun _ x = raise TERM ("Nitpick_HOL.mate_of_rep_fun", [Const x])
fun rep_type_for_quot_type thy (T as Type (s, _)) =
- let val {qtyp, rtyp, ...} = Quotient_Info.quotdata_lookup thy s in
- instantiate_type thy qtyp T rtyp
- end
+ let val {qtyp, rtyp, ...} = Quotient_Info.quotdata_lookup thy s in
+ instantiate_type thy qtyp T rtyp
+ end
+ | rep_type_for_quot_type _ T =
+ raise TYPE ("Nitpick_HOL.rep_type_for_quot_type", [T], [])
fun equiv_relation_for_quot_type thy (Type (s, Ts)) =
let
val {qtyp, equiv_rel, equiv_thm, ...} =
@@ -1079,7 +1085,7 @@
| _ => t
fun coerce_bound_0_in_term hol_ctxt new_T old_T =
old_T <> new_T ? coerce_bound_no (coerce_term hol_ctxt [new_T] old_T new_T) 0
-and coerce_term (hol_ctxt as {ctxt, stds, fast_descrs, ...}) Ts new_T old_T t =
+and coerce_term (hol_ctxt as {ctxt, stds, ...}) Ts new_T old_T t =
if old_T = new_T then
t
else
@@ -1567,7 +1573,7 @@
| [_, (@{const True}, head_t2)] => head_t2
| [_, (@{const False}, head_t2)] => @{const Not} $ head_t2
| _ => raise BAD ("Nitpick_HOL.optimized_case_def", "impossible cases")
- else
+ else
@{const True} |> fold_rev (add_constr_case res_T) cases
else
fst (hd cases) |> fold_rev (add_constr_case res_T) (tl cases)
@@ -1748,7 +1754,7 @@
val rep_T = domain_type (domain_type T)
val eps_fun = Const (@{const_name Eps},
(rep_T --> bool_T) --> rep_T)
- val normal_fun =
+ val normal_fun =
Const (quot_normal_name_for_type ctxt abs_T,
rep_T --> rep_T)
val abs_fun = Const (@{const_name Quot}, rep_T --> abs_T)
@@ -1956,7 +1962,7 @@
|> partial ? cons (HOLogic.mk_Trueprop (equiv_rel $ sel_a_t $ sel_a_t))
end
-fun codatatype_bisim_axioms (hol_ctxt as {thy, ctxt, stds, ...}) T =
+fun codatatype_bisim_axioms (hol_ctxt as {ctxt, stds, ...}) T =
let
val xs = datatype_constrs hol_ctxt T
val set_T = T --> bool_T
@@ -2163,7 +2169,7 @@
(disjuncts_of body)
val base_body = nonrecs |> List.foldl s_disj @{const False}
val step_body = recs |> map (repair_rec j)
- |> List.foldl s_disj @{const False}
+ |> List.foldl s_disj @{const False}
in
(list_abs (tl xs, incr_bv (~1, j, base_body))
|> ap_n_split (length arg_Ts) tuple_T bool_T,
--- a/src/HOL/Tools/Nitpick/nitpick_isar.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_isar.ML Mon Sep 13 21:24:10 2010 +0200
@@ -10,7 +10,7 @@
sig
type params = Nitpick.params
- val auto: bool Unsynchronized.ref
+ val auto : bool Unsynchronized.ref
val default_params : theory -> (string * string) list -> params
val setup : theory -> theory
end;
@@ -24,12 +24,16 @@
open Nitpick_Nut
open Nitpick
-val auto = Unsynchronized.ref false;
+val auto = Unsynchronized.ref false
+
+(* Maximum number of scopes for Auto Nitpick. Be frugal since it has to share
+ its time slot with several other automatic tools. *)
+val max_auto_scopes = 6
val _ =
ProofGeneralPgip.add_preference Preferences.category_tracing
(Preferences.bool_pref auto "auto-nitpick"
- "Whether to run Nitpick automatically.")
+ "Run Nitpick automatically.")
type raw_param = string * string list
@@ -67,7 +71,6 @@
("show_datatypes", "false"),
("show_consts", "false"),
("format", "1"),
- ("atoms", ""),
("max_potential", "1"),
("max_genuine", "1"),
("check_potential", "false"),
@@ -101,7 +104,7 @@
fun is_known_raw_param s =
AList.defined (op =) default_default_params s orelse
AList.defined (op =) negated_params s orelse
- member (op =) ["max", "show_all", "whack", "eval", "expect"] s orelse
+ member (op =) ["max", "show_all", "whack", "atoms", "eval", "expect"] s orelse
exists (fn p => String.isPrefix (p ^ " ") s)
["card", "max", "iter", "box", "dont_box", "finitize", "dont_finitize",
"mono", "non_mono", "std", "non_std", "wf", "non_wf", "format",
@@ -109,7 +112,7 @@
fun check_raw_param (s, _) =
if is_known_raw_param s then ()
- else error ("Unknown parameter: " ^ quote s ^ ".")
+ else error ("Unknown parameter: " ^ quote s ^ ".")
fun unnegate_param_name name =
case AList.lookup (op =) negated_params name of
@@ -212,8 +215,8 @@
lookup_assigns read prefix "" (space_explode " ")
fun lookup_time name =
case lookup name of
- NONE => NONE
- | SOME s => parse_time_option name s
+ SOME s => parse_time_option name s
+ | NONE => NONE
val read_type_polymorphic =
Syntax.read_typ ctxt #> Logic.mk_type
#> singleton (Variable.polymorphic ctxt) #> Logic.dest_type
@@ -223,17 +226,19 @@
AList.lookup (op =) raw_params #> these #> map read_term_polymorphic
val read_const_polymorphic = read_term_polymorphic #> dest_Const
val cards_assigns = lookup_ints_assigns read_type_polymorphic "card" 1
+ |> auto ? map (apsnd (take max_auto_scopes))
val maxes_assigns = lookup_ints_assigns read_const_polymorphic "max" ~1
val iters_assigns = lookup_ints_assigns read_const_polymorphic "iter" 0
val bitss = lookup_int_seq "bits" 1
val bisim_depths = lookup_int_seq "bisim_depth" ~1
val boxes = lookup_bool_option_assigns read_type_polymorphic "box"
val finitizes = lookup_bool_option_assigns read_type_polymorphic "finitize"
- val monos = lookup_bool_option_assigns read_type_polymorphic "mono"
+ val monos = if auto then [(NONE, SOME true)]
+ else lookup_bool_option_assigns read_type_polymorphic "mono"
val stds = lookup_bool_assigns read_type_polymorphic "std"
val wfs = lookup_bool_option_assigns read_const_polymorphic "wf"
val sat_solver = lookup_string "sat_solver"
- val blocking = not auto andalso lookup_bool "blocking"
+ val blocking = auto orelse lookup_bool "blocking"
val falsify = lookup_bool "falsify"
val debug = not auto andalso lookup_bool "debug"
val verbose = debug orelse (not auto andalso lookup_bool "verbose")
@@ -252,7 +257,7 @@
val kodkod_sym_break = lookup_int "kodkod_sym_break"
val timeout = if auto then NONE else lookup_time "timeout"
val tac_timeout = lookup_time "tac_timeout"
- val max_threads = Int.max (0, lookup_int "max_threads")
+ val max_threads = if auto then 1 else Int.max (0, lookup_int "max_threads")
val show_datatypes = debug orelse lookup_bool "show_datatypes"
val show_consts = debug orelse lookup_bool "show_consts"
val formats = lookup_ints_assigns read_term_polymorphic "format" 0
@@ -346,10 +351,7 @@
else handle_exceptions ctxt)
(fn (_, state) => pick_nits_in_subgoal state params auto i step
|>> curry (op =) "genuine")
- in
- if auto orelse blocking then go ()
- else (Toplevel.thread true (tap go); (false, state)) (* FIXME no threads in user-space *)
- end
+ in if blocking then go () else Future.fork (tap go) |> K (false, state) end
fun nitpick_trans (params, i) =
Toplevel.keep (fn st =>
@@ -362,7 +364,7 @@
fun nitpick_params_trans params =
Toplevel.theory
(fold set_default_raw_param params
- #> tap (fn thy =>
+ #> tap (fn thy =>
writeln ("Default parameters for Nitpick:\n" ^
(case rev (default_raw_params thy) of
[] => "none"
@@ -385,6 +387,6 @@
fun auto_nitpick state =
if not (!auto) then (false, state) else pick_nits [] true 1 0 state
-val setup = Auto_Counterexample.register_tool ("nitpick", auto_nitpick)
+val setup = Auto_Tools.register_tool ("nitpick", auto_nitpick)
end;
--- a/src/HOL/Tools/Nitpick/nitpick_kodkod.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML Mon Sep 13 21:24:10 2010 +0200
@@ -23,7 +23,7 @@
val sequential_int_bounds : int -> Kodkod.int_bound list
val pow_of_two_int_bounds : int -> int -> Kodkod.int_bound list
val bounds_and_axioms_for_built_in_rels_in_formulas :
- bool -> int Typtab.table -> int -> int -> int -> int -> Kodkod.formula list
+ bool -> int -> int -> int -> int -> Kodkod.formula list
-> Kodkod.bound list * Kodkod.formula list
val bound_for_plain_rel : Proof.context -> bool -> nut -> Kodkod.bound
val bound_for_sel_rel :
@@ -130,7 +130,7 @@
fun built_in_rels_in_formulas formulas =
let
- fun rel_expr_func (KK.Rel (x as (n, j))) =
+ fun rel_expr_func (KK.Rel (x as (_, j))) =
(j < 0 andalso x <> unsigned_bit_word_sel_rel andalso
x <> signed_bit_word_sel_rel)
? insert (op =) x
@@ -204,7 +204,7 @@
else if m = 0 orelse n = 0 then (0, 1)
else let val p = isa_zgcd (m, n) in (isa_div (m, p), isa_div (n, p)) end
-fun tabulate_built_in_rel debug ofs univ_card nat_card int_card j0
+fun tabulate_built_in_rel debug univ_card nat_card int_card j0
(x as (n, _)) =
(check_arity "" univ_card n;
if x = not3_rel then
@@ -245,7 +245,7 @@
else
raise ARG ("Nitpick_Kodkod.tabulate_built_in_rel", "unknown relation"))
-fun bound_for_built_in_rel debug ofs univ_card nat_card int_card main_j0
+fun bound_for_built_in_rel debug univ_card nat_card int_card main_j0
(x as (n, j)) =
if n = 2 andalso j <= suc_rels_base then
let val (y as (k, j0), tabulate) = atom_seq_for_suc_rel x in
@@ -258,8 +258,8 @@
end
else
let
- val (nick, ts) = tabulate_built_in_rel debug ofs univ_card nat_card
- int_card main_j0 x
+ val (nick, ts) = tabulate_built_in_rel debug univ_card nat_card int_card
+ main_j0 x
in ([(x, nick)], [KK.TupleSet ts]) end
fun axiom_for_built_in_rel (x as (n, j)) =
@@ -274,13 +274,13 @@
end
else
NONE
-fun bounds_and_axioms_for_built_in_rels_in_formulas debug ofs univ_card nat_card
+fun bounds_and_axioms_for_built_in_rels_in_formulas debug univ_card nat_card
int_card main_j0 formulas =
let val rels = built_in_rels_in_formulas formulas in
- (map (bound_for_built_in_rel debug ofs univ_card nat_card int_card main_j0)
+ (map (bound_for_built_in_rel debug univ_card nat_card int_card main_j0)
rels,
map_filter axiom_for_built_in_rel rels)
- end
+ end
fun bound_comment ctxt debug nick T R =
short_name nick ^
@@ -741,7 +741,7 @@
KK.Iden)
(* Cycle breaking in the bounds takes care of singly recursive datatypes, hence
the first equation. *)
-fun acyclicity_axioms_for_datatypes kk [_] = []
+fun acyclicity_axioms_for_datatypes _ [_] = []
| acyclicity_axioms_for_datatypes kk nfas =
maps (fn nfa => map (acyclicity_axiom_for_datatype kk nfa o fst) nfa) nfas
@@ -800,11 +800,11 @@
| NONE => false
fun sym_break_axioms_for_constr_pair hol_ctxt binarize
- (kk as {kk_all, kk_or, kk_iff, kk_implies, kk_and, kk_some, kk_subset,
- kk_intersect, kk_join, kk_project, ...}) rel_table nfas dtypes
+ (kk as {kk_all, kk_or, kk_implies, kk_and, kk_some, kk_intersect,
+ kk_join, ...}) rel_table nfas dtypes
(constr_ord,
({const = const1 as (_, T1), delta = delta1, epsilon = epsilon1, ...},
- {const = const2 as (_, T2), delta = delta2, epsilon = epsilon2, ...})
+ {const = const2 as (_, _), delta = delta2, epsilon = epsilon2, ...})
: constr_spec * constr_spec) =
let
val dataT = body_type T1
@@ -1418,8 +1418,8 @@
kk_join (to_rep (Func (unopt_rep R, Opt bool_atom_R)) u1) true_atom
else
to_rep (Func (R, Formula Neut)) u1
- | Op1 (First, T, R, u1) => to_nth_pair_sel 0 T R u1
- | Op1 (Second, T, R, u1) => to_nth_pair_sel 1 T R u1
+ | Op1 (First, _, R, u1) => to_nth_pair_sel 0 R u1
+ | Op1 (Second, _, R, u1) => to_nth_pair_sel 1 R u1
| Op1 (Cast, _, R, u1) =>
((case rep_of u1 of
Formula _ =>
@@ -1685,7 +1685,7 @@
| to_expr_assign (RelReg (j, _, R)) u =
KK.AssignRelReg ((arity_of_rep R, j), to_r u)
| to_expr_assign u1 _ = raise NUT ("Nitpick_Kodkod.to_expr_assign", [u1])
- and to_atom (x as (k, j0)) u =
+ and to_atom (x as (_, j0)) u =
case rep_of u of
Formula _ => atom_from_formula kk j0 (to_f u)
| R => atom_from_rel_expr kk x R (to_r u)
@@ -1727,7 +1727,7 @@
rel_expr_from_rel_expr kk new_R old_R
(kk_project_seq r j0 (arity_of_rep old_R))
and to_product Rs us = fold1 kk_product (map (uncurry to_opt) (Rs ~~ us))
- and to_nth_pair_sel n res_T res_R u =
+ and to_nth_pair_sel n res_R u =
case u of
Tuple (_, _, us) => to_rep res_R (nth us n)
| _ => let
@@ -1791,7 +1791,7 @@
[KK.IntEq (KK.IntReg 2,
oper (KK.IntReg 0) (KK.IntReg 1))]))))
end
- and to_apply (R as Formula _) func_u arg_u =
+ and to_apply (R as Formula _) _ _ =
raise REP ("Nitpick_Kodkod.to_apply", [R])
| to_apply res_R func_u arg_u =
case unopt_rep (rep_of func_u) of
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML Mon Sep 13 21:24:10 2010 +0200
@@ -904,18 +904,6 @@
fun term_for_rep maybe_opt unfold =
reconstruct_term maybe_opt unfold pool wacky_names scope atomss
sel_names rel_table bounds
- fun nth_value_of_type card T n =
- let
- fun aux unfold = term_for_rep true unfold T T (Atom (card, 0)) [[n]]
- in
- case aux false of
- t as Const (s, _) =>
- if String.isPrefix (cyclic_const_prefix ()) s then
- HOLogic.mk_eq (t, aux true)
- else
- t
- | t => t
- end
val all_values =
all_values_of_type pool wacky_names scope atomss sel_names rel_table
bounds
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML Mon Sep 13 21:24:10 2010 +0200
@@ -253,7 +253,7 @@
MAlpha
else case T of
Type (@{type_name fun}, [T1, T2]) =>
- MFun (fresh_mfun_for_fun_type mdata false T1 T2)
+ MFun (fresh_mfun_for_fun_type mdata all_minus T1 T2)
| Type (@{type_name prod}, [T1, T2]) => MPair (pairself do_type (T1, T2))
| Type (z as (s, _)) =>
if could_exist_alpha_sub_mtype ctxt alpha_T T then
@@ -549,8 +549,7 @@
consts = consts}
handle List.Empty => initial_gamma (* FIXME: needed? *)
-fun consider_term (mdata as {hol_ctxt as {thy, ctxt, stds, fast_descrs,
- def_table, ...},
+fun consider_term (mdata as {hol_ctxt = {thy, ctxt, stds, fast_descrs, ...},
alpha_T, max_fresh, ...}) =
let
fun is_enough_eta_expanded t =
@@ -617,7 +616,7 @@
accum ||> add_mtypes_equal expected_bound_M (mtype_of_mterm bound_m)
|> do_term body_t ||> apfst pop_bound
val bound_M = mtype_of_mterm bound_m
- val (M1, a, M2) = dest_MFun bound_M
+ val (M1, a, _) = dest_MFun bound_M
in
(MApp (MRaw (t0, MFun (bound_M, S Minus, bool_M)),
MAbs (abs_s, abs_T, M1, a,
@@ -626,8 +625,7 @@
MApp (bound_m, MRaw (Bound 0, M1))),
body_m))), accum)
end
- and do_term t (accum as (gamma as {bound_Ts, bound_Ms, frees, consts},
- cset)) =
+ and do_term t (accum as ({bound_Ts, bound_Ms, frees, consts}, cset)) =
(trace_msg (fn () => " \<Gamma> \<turnstile> " ^
Syntax.string_of_term ctxt t ^ " : _?");
case t of
@@ -660,7 +658,7 @@
|>> curry3 MFun bool_M (S Minus)
| @{const_name Pair} => do_pair_constr T accum
| @{const_name fst} => do_nth_pair_sel 0 T accum
- | @{const_name snd} => do_nth_pair_sel 1 T accum
+ | @{const_name snd} => do_nth_pair_sel 1 T accum
| @{const_name Id} =>
(MFun (mtype_for (domain_type T), S Minus, bool_M), accum)
| @{const_name converse} =>
@@ -787,8 +785,6 @@
val (m2, accum) = do_term t2 accum
in
let
- val T11 = domain_type (fastype_of1 (bound_Ts, t1))
- val T2 = fastype_of1 (bound_Ts, t2)
val M11 = mtype_of_mterm m1 |> dest_MFun |> #1
val M2 = mtype_of_mterm m2
in (MApp (m1, m2), accum ||> add_is_sub_mtype M2 M11) end
@@ -867,7 +863,7 @@
end
| Const (x as (@{const_name All}, _)) $ Abs (s1, T1, t1) =>
do_quantifier x s1 T1 t1
- | Const (x0 as (s0 as @{const_name Ex}, T0))
+ | Const (x0 as (@{const_name Ex}, T0))
$ (t1 as Abs (s1, T1, t1')) =>
(case sn of
Plus => do_quantifier x0 s1 T1 t1'
@@ -894,7 +890,7 @@
in
(MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2),
accum)
- end
+ end
else
do_term t accum
| _ => do_term t accum)
@@ -1105,7 +1101,7 @@
val (T1, T2) = pairself (curry fastype_of1 new_Ts) (t1, t2)
val (t1', T2') =
case T1 of
- Type (s, [T11, T12]) =>
+ Type (s, [T11, T12]) =>
(if s = @{type_name fin_fun} then
select_nth_constr_arg ctxt stds (fin_fun_constr T11 T12) t1
0 (T11 --> T12)
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML Mon Sep 13 21:24:10 2010 +0200
@@ -653,7 +653,8 @@
FreeName _ => true
| Op1 (SingletonSet, _, _, _) => true
| Op1 (Converse, _, _, u1) => is_fully_representable_set u1
- | Op2 (oper, _, _, u1, u2) =>
+ | Op2 (Lambda, _, _, _, Cst (False, _, _)) => true
+ | Op2 (oper, _, _, u1, _) =>
if oper = Apply then
case u1 of
ConstName (s, _, _) =>
@@ -661,7 +662,6 @@
| _ => false
else
false
- | Op2 (Lambda, _, _, _, Cst (False, _, _)) => true
| _ => false
fun rep_for_abs_fun scope T =
--- a/src/HOL/Tools/Nitpick/nitpick_preproc.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_preproc.ML Mon Sep 13 21:24:10 2010 +0200
@@ -982,7 +982,7 @@
else if s = @{const_name TYPE} then
accum
else case def_of_const thy def_table x of
- SOME def =>
+ SOME _ =>
fold (add_eq_axiom depth) (equational_fun_axioms hol_ctxt x)
accum
| NONE =>
--- a/src/HOL/Tools/Nitpick/nitpick_rep.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_rep.ML Mon Sep 13 21:24:10 2010 +0200
@@ -238,15 +238,14 @@
Vect (card_of_type card_assigns T1, (best_one_rep_for_type scope T2))
| best_one_rep_for_type scope (Type (@{type_name prod}, Ts)) =
Struct (map (best_one_rep_for_type scope) Ts)
- | best_one_rep_for_type {card_assigns, datatypes, ofs, ...} T =
+ | best_one_rep_for_type {card_assigns, ofs, ...} T =
Atom (card_of_type card_assigns T, offset_of_type ofs T)
fun best_opt_set_rep_for_type scope (Type (@{type_name fun}, [T1, T2])) =
Func (best_one_rep_for_type scope T1, best_opt_set_rep_for_type scope T2)
| best_opt_set_rep_for_type (scope as {ofs, ...}) T =
opt_rep ofs T (best_one_rep_for_type scope T)
-fun best_non_opt_set_rep_for_type (scope as {ofs, ...})
- (Type (@{type_name fun}, [T1, T2])) =
+fun best_non_opt_set_rep_for_type scope (Type (@{type_name fun}, [T1, T2])) =
(case (best_one_rep_for_type scope T1,
best_non_opt_set_rep_for_type scope T2) of
(R1, Atom (2, _)) => Func (R1, Formula Neut)
--- a/src/HOL/Tools/Sledgehammer/metis_clauses.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/metis_clauses.ML Mon Sep 13 21:24:10 2010 +0200
@@ -270,7 +270,7 @@
val tvars = gen_TVars (length args)
val tvars_srts = ListPair.zip (tvars, args)
in
- ArityClause {name = name,
+ ArityClause {name = name,
conclLit = TConsLit (`make_type_class cls,
`make_fixed_type_const tcons,
tvars ~~ tvars),
--- a/src/HOL/Tools/Sledgehammer/metis_tactics.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/metis_tactics.ML Mon Sep 13 21:24:10 2010 +0200
@@ -735,7 +735,7 @@
in (mode, add_type_thm (type_ext thy (map prop_of (cls @ ths))) lmap) end
fun refute cls =
- Metis.Resolution.loop (Metis.Resolution.new Metis.Resolution.default cls);
+ Metis.Resolution.loop (Metis.Resolution.new Metis.Resolution.default {axioms = cls, conjecture = []});
fun is_false t = t aconv (HOLogic.mk_Trueprop HOLogic.false_const);
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML Mon Sep 13 21:24:10 2010 +0200
@@ -28,7 +28,7 @@
timeout: Time.time,
expect: string}
type problem =
- {ctxt: Proof.context,
+ {state: Proof.state,
goal: thm,
subgoal: int,
axioms: (term * ((string * locality) * fol_formula) option) list}
@@ -52,8 +52,8 @@
val messages: int option -> unit
val get_prover_fun : theory -> string -> prover
val run_sledgehammer :
- params -> int -> relevance_override -> (string -> minimize_command)
- -> Proof.state -> unit
+ params -> bool -> int -> relevance_override -> (string -> minimize_command)
+ -> Proof.state -> bool * Proof.state
val setup : theory -> theory
end;
@@ -97,7 +97,7 @@
expect: string}
type problem =
- {ctxt: Proof.context,
+ {state: Proof.state,
goal: thm,
subgoal: int,
axioms: (term * ((string * locality) * fol_formula) option) list}
@@ -230,18 +230,18 @@
(* generic TPTP-based provers *)
-fun prover_fun atp_name
+fun prover_fun auto atp_name
{exec, required_execs, arguments, has_incomplete_mode, proof_delims,
known_failures, default_max_relevant, explicit_forall,
use_conjecture_for_hypotheses}
({debug, verbose, overlord, full_types, explicit_apply,
max_relevant, isar_proof, isar_shrink_factor, timeout, ...} : params)
- minimize_command ({ctxt, goal, subgoal, axioms} : problem) =
+ minimize_command ({state, goal, subgoal, axioms} : problem) =
let
+ val ctxt = Proof.context_of state
val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal
val max_relevant = the_default default_max_relevant max_relevant
val axioms = take max_relevant axioms
- (* path to unique problem file *)
val the_dest_dir = if overlord then getenv "ISABELLE_HOME_USER"
else Config.get ctxt dest_dir
val the_problem_prefix = Config.get ctxt problem_prefix
@@ -285,7 +285,7 @@
| [] =>
if File.exists command then
let
- fun do_run complete timeout =
+ fun run complete timeout =
let
val command = command_line complete timeout probfile
val ((output, msecs), res_code) =
@@ -309,17 +309,17 @@
val conjecture_shape =
conjecture_offset + 1 upto conjecture_offset + length hyp_ts + 1
|> map single
+ val run_twice = has_incomplete_mode andalso not auto
val timer = Timer.startRealTimer ()
val result =
- do_run false (if has_incomplete_mode then
- Time.fromMilliseconds
+ run false (if run_twice then
+ Time.fromMilliseconds
(2 * Time.toMilliseconds timeout div 3)
- else
- timeout)
- |> has_incomplete_mode
+ else
+ timeout)
+ |> run_twice
? (fn (_, msecs0, _, SOME _) =>
- do_run true
- (Time.- (timeout, Timer.checkRealTimer timer))
+ run true (Time.- (timeout, Timer.checkRealTimer timer))
|> (fn (output, msecs, proof, outcome) =>
(output, msecs0 + msecs, proof, outcome))
| result => result)
@@ -343,12 +343,15 @@
repair_conjecture_shape_and_theorem_names output conjecture_shape
axiom_names
val important_message = extract_important_message output
+ val banner = if auto then "Sledgehammer found a proof"
+ else "Try this command"
val (message, used_thm_names) =
case outcome of
NONE =>
proof_text isar_proof
(pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
- (full_types, minimize_command, proof, axiom_names, goal, subgoal)
+ (banner, full_types, minimize_command, proof, axiom_names, goal,
+ subgoal)
|>> (fn message =>
message ^ (if verbose then
"\nReal CPU time: " ^ string_of_int msecs ^ " ms."
@@ -367,14 +370,14 @@
conjecture_shape = conjecture_shape}
end
-fun get_prover_fun thy name = prover_fun name (get_prover thy name)
+fun get_prover_fun thy name = prover_fun false name (get_prover thy name)
-fun run_prover ctxt (params as {blocking, verbose, max_relevant, timeout,
- expect, ...})
- i n relevance_override minimize_command
- (problem as {goal, axioms, ...})
+fun run_prover (params as {blocking, verbose, max_relevant, timeout, expect,
+ ...})
+ auto i n minimize_command (problem as {state, goal, axioms, ...})
(prover as {default_max_relevant, ...}, atp_name) =
let
+ val ctxt = Proof.context_of state
val birth_time = Time.now ()
val death_time = Time.+ (birth_time, timeout)
val max_relevant = the_default default_max_relevant max_relevant
@@ -390,72 +393,91 @@
""
else
"\n" ^ Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i)))
- fun run () =
+ fun go () =
let
val (outcome_code, message) =
- prover_fun atp_name prover params (minimize_command atp_name) problem
+ prover_fun auto atp_name prover params (minimize_command atp_name)
+ problem
|> (fn {outcome, message, ...} =>
(if is_some outcome then "none" else "some", message))
handle ERROR message => ("unknown", "Error: " ^ message ^ "\n")
| exn => ("unknown", "Internal error:\n" ^
ML_Compiler.exn_message exn ^ "\n")
- in
- if expect = "" orelse outcome_code = expect then
- ()
- else if blocking then
- error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
- else
- warning ("Unexpected outcome: " ^ quote outcome_code ^ ".");
- message
+ val _ =
+ if expect = "" orelse outcome_code = expect then
+ ()
+ else if blocking then
+ error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
+ else
+ warning ("Unexpected outcome: " ^ quote outcome_code ^ ".");
+ in (outcome_code = "some", message) end
+ in
+ if auto then
+ let val (success, message) = TimeLimit.timeLimit timeout go () in
+ (success, state |> success ? Proof.goal_message (fn () =>
+ Pretty.chunks [Pretty.str "", Pretty.mark Markup.hilite
+ (Pretty.str message)]))
end
- in
- if blocking then priority (desc ^ "\n" ^ TimeLimit.timeLimit timeout run ())
- else Async_Manager.launch das_Tool birth_time death_time desc run
+ else if blocking then
+ let val (success, message) = TimeLimit.timeLimit timeout go () in
+ priority (desc ^ "\n" ^ message); (success, state)
+ end
+ else
+ (Async_Manager.launch das_Tool birth_time death_time desc (snd o go);
+ (false, state))
end
+val auto_max_relevant_divisor = 2
+
fun run_sledgehammer (params as {blocking, verbose, atps, full_types,
relevance_thresholds, max_relevant, ...})
- i relevance_override minimize_command state =
- if null atps then error "No ATP is set."
- else case subgoal_count state of
- 0 => priority "No subgoal!"
- | n =>
- let
- val thy = Proof.theory_of state
- val timer = Timer.startRealTimer ()
- val _ = () |> not blocking ? kill_atps
- val _ = priority "Sledgehammering..."
- val provers = map (`(get_prover thy)) atps
- fun run () =
- let
- val {context = ctxt, facts = chained_ths, goal} = Proof.goal state
- val (_, hyp_ts, concl_t) = strip_subgoal goal i
- val max_max_relevant =
- case max_relevant of
- SOME n => n
- | NONE => fold (Integer.max o #default_max_relevant o fst)
- provers 0
- val axioms =
- relevant_facts ctxt full_types relevance_thresholds
- max_max_relevant relevance_override chained_ths
- hyp_ts concl_t
- val problem =
- {ctxt = ctxt, goal = goal, subgoal = i,
- axioms = map (prepare_axiom ctxt) axioms}
- val num_axioms = length axioms
- val _ =
- (if blocking then Par_List.map else map)
- (run_prover ctxt params i n relevance_override
- minimize_command problem) provers
- in
- if verbose andalso blocking then
- priority ("Total time: " ^
- signed_string_of_int (Time.toMilliseconds
- (Timer.checkRealTimer timer)) ^ " ms.")
- else
- ()
- end
- in if blocking then run () else Future.fork run |> K () end
+ auto i relevance_override minimize_command state =
+ if null atps then
+ error "No ATP is set."
+ else case subgoal_count state of
+ 0 => (priority "No subgoal!"; (false, state))
+ | n =>
+ let
+ val thy = Proof.theory_of state
+ val timer = Timer.startRealTimer ()
+ val _ = () |> not blocking ? kill_atps
+ val _ = if auto then () else priority "Sledgehammering..."
+ val provers = map (`(get_prover thy)) atps
+ fun go () =
+ let
+ val {context = ctxt, facts = chained_ths, goal} = Proof.goal state
+ val (_, hyp_ts, concl_t) = strip_subgoal goal i
+ val max_max_relevant =
+ case max_relevant of
+ SOME n => n
+ | NONE =>
+ 0 |> fold (Integer.max o #default_max_relevant o fst) provers
+ |> auto ? (fn n => n div auto_max_relevant_divisor)
+ val axioms =
+ relevant_facts ctxt full_types relevance_thresholds
+ max_max_relevant relevance_override chained_ths
+ hyp_ts concl_t
+ val problem =
+ {state = state, goal = goal, subgoal = i,
+ axioms = map (prepare_axiom ctxt) axioms}
+ val run_prover = run_prover params auto i n minimize_command problem
+ in
+ if auto then
+ fold (fn prover => fn (true, state) => (true, state)
+ | (false, _) => run_prover prover)
+ provers (false, state)
+ else
+ (if blocking then Par_List.map else map) run_prover provers
+ |> tap (fn _ =>
+ if verbose then
+ priority ("Total time: " ^
+ signed_string_of_int (Time.toMilliseconds
+ (Timer.checkRealTimer timer)) ^ " ms.")
+ else
+ ())
+ |> exists fst |> rpair state
+ end
+ in if blocking then go () else Future.fork (tap go) |> K (false, state) end
val setup =
dest_dir_setup
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Mon Sep 13 21:24:10 2010 +0200
@@ -8,10 +8,12 @@
sig
type params = Sledgehammer.params
- val atps: string Unsynchronized.ref
- val timeout: int Unsynchronized.ref
- val full_types: bool Unsynchronized.ref
+ val auto : bool Unsynchronized.ref
+ val atps : string Unsynchronized.ref
+ val timeout : int Unsynchronized.ref
+ val full_types : bool Unsynchronized.ref
val default_params : theory -> (string * string) list -> params
+ val setup : theory -> theory
end;
structure Sledgehammer_Isar : SLEDGEHAMMER_ISAR =
@@ -22,11 +24,19 @@
open Sledgehammer
open Sledgehammer_Minimize
+val auto = Unsynchronized.ref false
+
+val _ =
+ ProofGeneralPgip.add_preference Preferences.category_tracing
+ (Preferences.bool_pref auto "auto-sledgehammer"
+ "Run Sledgehammer automatically.")
+
(** Sledgehammer commands **)
-fun add_to_relevance_override ns : relevance_override =
+val no_relevance_override = {add = [], del = [], only = false}
+fun add_relevance_override ns : relevance_override =
{add = ns, del = [], only = false}
-fun del_from_relevance_override ns : relevance_override =
+fun del_relevance_override ns : relevance_override =
{add = [], del = ns, only = false}
fun only_relevance_override ns : relevance_override =
{add = ns, del = [], only = true}
@@ -40,7 +50,7 @@
(*** parameters ***)
val atps = Unsynchronized.ref ""
-val timeout = Unsynchronized.ref 60
+val timeout = Unsynchronized.ref 30
val full_types = Unsynchronized.ref false
val _ =
@@ -132,7 +142,7 @@
val infinity_time_in_secs = 60 * 60 * 24 * 365
val the_timeout = the_default (Time.fromSeconds infinity_time_in_secs)
-fun extract_params default_params override_params =
+fun extract_params auto default_params override_params =
let
val override_params = map unalias_raw_param override_params
val raw_params = rev override_params @ rev default_params
@@ -143,11 +153,10 @@
SOME s => parse_bool_option option name s
| NONE => default_value
val lookup_bool = the o general_lookup_bool false (SOME false)
- val lookup_bool_option = general_lookup_bool true NONE
fun lookup_time name =
- the_timeout (case lookup name of
- NONE => NONE
- | SOME s => parse_time_option name s)
+ case lookup name of
+ SOME s => parse_time_option name s
+ | NONE => NONE
fun lookup_int name =
case lookup name of
NONE => 0
@@ -167,11 +176,11 @@
case lookup name of
SOME "smart" => NONE
| _ => SOME (lookup_int name)
- val blocking = lookup_bool "blocking"
- val debug = lookup_bool "debug"
- val verbose = debug orelse lookup_bool "verbose"
+ val blocking = auto orelse lookup_bool "blocking"
+ val debug = not auto andalso lookup_bool "debug"
+ val verbose = debug orelse (not auto andalso lookup_bool "verbose")
val overlord = lookup_bool "overlord"
- val atps = lookup_string "atps" |> space_explode " "
+ val atps = lookup_string "atps" |> space_explode " " |> auto ? single o hd
val full_types = lookup_bool "full_types"
val explicit_apply = lookup_bool "explicit_apply"
val relevance_thresholds =
@@ -180,7 +189,7 @@
val max_relevant = lookup_int_option "max_relevant"
val isar_proof = lookup_bool "isar_proof"
val isar_shrink_factor = Int.max (1, lookup_int "isar_shrink_factor")
- val timeout = lookup_time "timeout"
+ val timeout = (if auto then NONE else lookup_time "timeout") |> the_timeout
val expect = lookup_string "expect"
in
{blocking = blocking, debug = debug, verbose = verbose, overlord = overlord,
@@ -190,8 +199,8 @@
timeout = timeout, expect = expect}
end
-fun get_params thy = extract_params (default_raw_params thy)
-fun default_params thy = get_params thy o map (apsnd single)
+fun get_params auto thy = extract_params auto (default_raw_params thy)
+fun default_params thy = get_params false thy o map (apsnd single)
(* Sledgehammer the given subgoal *)
@@ -225,11 +234,13 @@
in
if subcommand = runN then
let val i = the_default 1 opt_i in
- run_sledgehammer (get_params thy override_params) i relevance_override
- (minimize_command override_params i) state
+ run_sledgehammer (get_params false thy override_params) false i
+ relevance_override (minimize_command override_params i)
+ state
+ |> K ()
end
else if subcommand = minimizeN then
- run_minimize (get_params thy override_params) (the_default 1 opt_i)
+ run_minimize (get_params false thy override_params) (the_default 1 opt_i)
(#add relevance_override) state
else if subcommand = messagesN then
messages opt_i
@@ -254,7 +265,7 @@
fun sledgehammer_params_trans params =
Toplevel.theory
(fold set_default_raw_param params
- #> tap (fn thy =>
+ #> tap (fn thy =>
writeln ("Default parameters for Sledgehammer:\n" ^
(case rev (default_raw_params thy) of
[] => "none"
@@ -270,14 +281,13 @@
val parse_fact_refs =
Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) Parse_Spec.xthm)
val parse_relevance_chunk =
- (Args.add |-- Args.colon |-- parse_fact_refs >> add_to_relevance_override)
- || (Args.del |-- Args.colon |-- parse_fact_refs
- >> del_from_relevance_override)
+ (Args.add |-- Args.colon |-- parse_fact_refs >> add_relevance_override)
+ || (Args.del |-- Args.colon |-- parse_fact_refs >> del_relevance_override)
|| (parse_fact_refs >> only_relevance_override)
val parse_relevance_override =
Scan.optional (Args.parens (Scan.repeat parse_relevance_chunk
>> merge_relevance_overrides))
- (add_to_relevance_override [])
+ no_relevance_override
val parse_sledgehammer_command =
(Scan.optional Parse.short_ident runN -- parse_params
-- parse_relevance_override -- Scan.option Parse.nat) #>> sledgehammer_trans
@@ -293,4 +303,15 @@
"set and display the default parameters for Sledgehammer" Keyword.thy_decl
parse_sledgehammer_params_command
+fun auto_sledgehammer state =
+ if not (!auto) then
+ (false, state)
+ else
+ let val thy = Proof.theory_of state in
+ run_sledgehammer (get_params true thy []) true 1 no_relevance_override
+ (minimize_command [] 1) state
+ end
+
+val setup = Auto_Tools.register_tool ("sledgehammer", auto_sledgehammer)
+
end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML Mon Sep 13 21:24:10 2010 +0200
@@ -60,7 +60,7 @@
val axioms = maps (fn (n, ths) => map (pair n) ths) axioms
val {context = ctxt, goal, ...} = Proof.goal state
val problem =
- {ctxt = ctxt, goal = goal, subgoal = subgoal,
+ {state = state, goal = goal, subgoal = subgoal,
axioms = map (prepare_axiom ctxt) axioms}
val result as {outcome, used_thm_names, ...} = prover params (K "") problem
in
@@ -81,7 +81,7 @@
fun sublinear_minimize _ [] p = p
| sublinear_minimize test (x :: xs) (seen, result) =
case test (xs @ seen) of
- result as {outcome = NONE, proof, used_thm_names, ...} : prover_result =>
+ result as {outcome = NONE, used_thm_names, ...} : prover_result =>
sublinear_minimize test (filter_used_facts used_thm_names xs)
(filter_used_facts used_thm_names seen, result)
| _ => sublinear_minimize test xs (x :: seen, result)
@@ -94,7 +94,7 @@
fun minimize_theorems {atps = [], ...} _ _ _ _ = error "No ATP is set."
| minimize_theorems (params as {debug, atps = atp :: _, full_types,
isar_proof, isar_shrink_factor, timeout, ...})
- i n state axioms =
+ i (_ : int) state axioms =
let
val thy = Proof.theory_of state
val prover = get_prover_fun thy atp
@@ -130,7 +130,8 @@
(SOME min_thms,
proof_text isar_proof
(pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
- (full_types, K "", proof, axiom_names, goal, i) |> fst)
+ ("Minimized command", full_types, K "", proof, axiom_names, goal,
+ i) |> fst)
end
| {outcome = SOME TimedOut, ...} =>
(NONE, "Timeout: You can increase the time limit using the \"timeout\" \
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML Mon Sep 13 21:24:10 2010 +0200
@@ -11,8 +11,8 @@
type locality = Sledgehammer_Filter.locality
type minimize_command = string list -> string
type metis_params =
- bool * minimize_command * string * (string * locality) list vector * thm
- * int
+ string * bool * minimize_command * string * (string * locality) list vector
+ * thm * int
type isar_params =
string Symtab.table * bool * int * Proof.context * int list list
type text_result = string * (string * locality) list
@@ -33,7 +33,8 @@
type minimize_command = string list -> string
type metis_params =
- bool * minimize_command * string * (string * locality) list vector * thm * int
+ string * bool * minimize_command * string * (string * locality) list vector
+ * thm * int
type isar_params =
string Symtab.table * bool * int * Proof.context * int list list
type text_result = string * (string * locality) list
@@ -614,8 +615,8 @@
"(" ^ metis_name full_types ^ " " ^ space_implode " " ss ^ ")"
fun metis_command full_types i n (ls, ss) =
metis_using ls ^ metis_apply i n ^ metis_call full_types ss
-fun metis_line full_types i n ss =
- "Try this command: " ^
+fun metis_line banner full_types i n ss =
+ banner ^ ": " ^
Markup.markup Markup.sendback (metis_command full_types i n ([], ss)) ^ "."
fun minimize_line _ [] = ""
| minimize_line minimize_command ss =
@@ -630,13 +631,13 @@
#> List.partition (curry (op =) Chained o snd)
#> pairself (sort_distinct (string_ord o pairself fst))
-fun metis_proof_text (full_types, minimize_command, atp_proof, axiom_names,
- goal, i) =
+fun metis_proof_text (banner, full_types, minimize_command, atp_proof,
+ axiom_names, goal, i) =
let
val (chained_lemmas, other_lemmas) = used_facts axiom_names atp_proof
val n = Logic.count_prems (prop_of goal)
in
- (metis_line full_types i n (map fst other_lemmas) ^
+ (metis_line banner full_types i n (map fst other_lemmas) ^
minimize_line minimize_command (map fst (other_lemmas @ chained_lemmas)),
other_lemmas @ chained_lemmas)
end
@@ -1003,7 +1004,7 @@
in do_proof end
fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
- (other_params as (full_types, _, atp_proof, axiom_names,
+ (other_params as (_, full_types, _, atp_proof, axiom_names,
goal, i)) =
let
val (params, hyp_ts, concl_t) = strip_subgoal goal i
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML Mon Sep 13 21:24:10 2010 +0200
@@ -25,7 +25,7 @@
val strip_subgoal : thm -> int -> (string * typ) list * term list * term
val reserved_isar_keyword_table : unit -> unit Symtab.table
end;
-
+
structure Sledgehammer_Util : SLEDGEHAMMER_UTIL =
struct
--- a/src/HOL/Tools/meson.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/meson.ML Mon Sep 13 21:24:10 2010 +0200
@@ -11,7 +11,6 @@
val term_pair_of: indexname * (typ * 'a) -> term * 'a
val flexflex_first_order: thm -> thm
val size_of_subgoals: thm -> int
- val estimated_num_clauses: Proof.context -> int -> term -> int
val has_too_many_clauses: Proof.context -> term -> bool
val make_cnf: thm list -> thm -> Proof.context -> thm list * Proof.context
val finish_cnf: thm list -> thm list
@@ -94,7 +93,6 @@
(** First-order Resolution **)
-fun typ_pair_of (ix, (sort,ty)) = (TVar (ix,sort), ty);
fun term_pair_of (ix, (ty,t)) = (Var (ix,ty), t);
(*FIXME: currently does not "rename variables apart"*)
@@ -117,7 +115,7 @@
[] => th
| pairs =>
let val thy = theory_of_thm th
- val (tyenv, tenv) =
+ val (_, tenv) =
fold (Pattern.first_order_match thy) pairs (Vartab.empty, Vartab.empty)
val t_pairs = map term_pair_of (Vartab.dest tenv)
val th' = Thm.instantiate ([], map (pairself (cterm_of thy)) t_pairs) th
@@ -154,7 +152,7 @@
(*Are any of the logical connectives in "bs" present in the term?*)
fun has_conns bs =
- let fun has (Const(a,_)) = false
+ let fun has (Const _) = false
| has (Const(@{const_name Trueprop},_) $ p) = has p
| has (Const(@{const_name Not},_) $ p) = has p
| has (Const(@{const_name HOL.disj},_) $ p $ q) = member (op =) bs @{const_name HOL.disj} orelse has p orelse has q
@@ -238,7 +236,7 @@
(*** Removal of duplicate literals ***)
(*Forward proof, passing extra assumptions as theorems to the tactic*)
-fun forward_res2 ctxt nf hyps st =
+fun forward_res2 nf hyps st =
case Seq.pull
(REPEAT
(Misc_Legacy.METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1)
@@ -250,7 +248,7 @@
rls (initially []) accumulates assumptions of the form P==>False*)
fun nodups_aux ctxt rls th = nodups_aux ctxt rls (th RS disj_assoc)
handle THM _ => tryres(th,rls)
- handle THM _ => tryres(forward_res2 ctxt (nodups_aux ctxt) rls (th RS disj_forward2),
+ handle THM _ => tryres(forward_res2 (nodups_aux ctxt) rls (th RS disj_forward2),
[disj_FalseD1, disj_FalseD2, asm_rl])
handle THM _ => th;
@@ -263,7 +261,7 @@
(*** The basic CNF transformation ***)
-fun estimated_num_clauses ctxt bound t =
+fun estimated_num_clauses bound t =
let
fun sum x y = if x < bound andalso y < bound then x+y else bound
fun prod x y = if x < bound andalso y < bound then x*y else bound
@@ -294,7 +292,7 @@
fun has_too_many_clauses ctxt t =
let val max_cl = Config.get ctxt max_clauses in
- estimated_num_clauses ctxt (max_cl + 1) t > max_cl
+ estimated_num_clauses (max_cl + 1) t > max_cl
end
(*Replaces universally quantified variables by FREE variables -- because
@@ -448,7 +446,7 @@
(*Generate Horn clauses for all contrapositives of a clause. The input, th,
is a HOL disjunction.*)
fun add_contras crules th hcs =
- let fun rots (0,th) = hcs
+ let fun rots (0,_) = hcs
| rots (k,th) = zero_var_indexes (make_horn crules th) ::
rots(k-1, assoc_right (th RS disj_comm))
in case nliterals(prop_of th) of
@@ -639,7 +637,7 @@
(*This version does only one inference per call;
having only one eq_assume_tac speeds it up!*)
fun prolog_step_tac' horns =
- let val (horn0s, hornps) = (*0 subgoals vs 1 or more*)
+ let val (horn0s, _) = (*0 subgoals vs 1 or more*)
take_prefix Thm.no_prems horns
val nrtac = net_resolve_tac horns
in fn i => eq_assume_tac i ORELSE
--- a/src/HOL/Tools/try.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/HOL/Tools/try.ML Mon Sep 13 21:24:10 2010 +0200
@@ -6,24 +6,36 @@
signature TRY =
sig
- val invoke_try : Proof.state -> unit
+ val auto : bool Unsynchronized.ref
+ val invoke_try : Time.time option -> Proof.state -> bool
+ val setup : theory -> theory
end;
structure Try : TRY =
struct
-val timeout = Time.fromSeconds 5
+val auto = Unsynchronized.ref false
+
+val _ =
+ ProofGeneralPgip.add_preference Preferences.category_tracing
+ (setmp_noncritical auto true (fn () =>
+ Preferences.bool_pref auto
+ "auto-try" "Try standard proof methods.") ());
-fun can_apply pre post tac st =
+val default_timeout = Time.fromSeconds 5
+
+fun can_apply timeout_opt pre post tac st =
let val {goal, ...} = Proof.goal st in
- case TimeLimit.timeLimit timeout (Seq.pull o tac) (pre st) of
+ case (case timeout_opt of
+ SOME timeout => TimeLimit.timeLimit timeout
+ | NONE => fn f => fn x => f x) (Seq.pull o tac) (pre st) of
SOME (x, _) => nprems_of (post x) < nprems_of goal
| NONE => false
end
-fun do_generic command pre post apply st =
+fun do_generic timeout_opt command pre post apply st =
let val timer = Timer.startRealTimer () in
- if can_apply pre post apply st then
+ if can_apply timeout_opt pre post apply st then
SOME (command, Time.toMilliseconds (Timer.checkRealTimer timer))
else
NONE
@@ -37,8 +49,8 @@
Method.apply (named_method thy name) ctxt []
end
-fun do_named_method name st =
- do_generic name (#goal o Proof.goal) snd
+fun do_named_method name timeout_opt st =
+ do_generic timeout_opt name (#goal o Proof.goal) snd
(apply_named_method name (Proof.context_of st)) st
fun apply_named_method_on_first_goal name ctxt =
@@ -46,8 +58,9 @@
Proof.refine (Method.SelectGoals (1, Method.Basic (named_method thy name)))
end
-fun do_named_method_on_first_goal name st =
- do_generic (name ^ (if nprems_of (#goal (Proof.goal st)) > 1 then "[1]"
+fun do_named_method_on_first_goal name timeout_opt st =
+ do_generic timeout_opt
+ (name ^ (if nprems_of (#goal (Proof.goal st)) > 1 then "[1]"
else "")) I (#goal o Proof.goal)
(apply_named_method_on_first_goal name (Proof.context_of st)) st
@@ -60,22 +73,42 @@
fun time_string (s, ms) = s ^ ": " ^ string_of_int ms ^ " ms"
-fun invoke_try st =
- case do_methods |> Par_List.map (fn f => f st)
+fun do_try auto timeout_opt st =
+ case do_methods |> Par_List.map (fn f => f timeout_opt st)
|> map_filter I |> sort (int_ord o pairself snd) of
- [] => writeln "No proof found."
+ [] => (if auto then () else writeln "No proof found."; (false, st))
| xs as (s, _) :: _ =>
- priority ("Try this command: " ^
+ let
+ val xs = xs |> map swap |> AList.coalesce (op =)
+ |> map (swap o apsnd commas)
+ val message =
+ (if auto then "Auto Try found a proof" else "Try this command") ^ ": " ^
Markup.markup Markup.sendback
((if nprems_of (#goal (Proof.goal st)) = 1 then "by" else "apply") ^
" " ^ s) ^
- ".\n(" ^ space_implode "; " (map time_string xs) ^ ")\n")
+ ".\n(" ^ space_implode "; " (map time_string xs) ^ ")\n"
+ in
+ (true, st |> (if auto then
+ Proof.goal_message
+ (fn () => Pretty.chunks [Pretty.str "",
+ Pretty.markup Markup.hilite
+ [Pretty.str message]])
+ else
+ tap (fn _ => priority message)))
+ end
+
+val invoke_try = fst oo do_try false
val tryN = "try"
val _ =
Outer_Syntax.improper_command tryN
"try a combination of proof methods" Keyword.diag
- (Scan.succeed (Toplevel.keep (invoke_try o Toplevel.proof_of)))
+ (Scan.succeed (Toplevel.keep (K () o do_try false (SOME default_timeout)
+ o Toplevel.proof_of)))
+
+fun auto_try st = if not (!auto) then (false, st) else do_try true NONE st
+
+val setup = Auto_Tools.register_tool (tryN, auto_try)
end;
--- a/src/Tools/Code_Generator.thy Mon Sep 13 16:44:20 2010 +0200
+++ b/src/Tools/Code_Generator.thy Mon Sep 13 21:24:10 2010 +0200
@@ -8,8 +8,8 @@
imports Pure
uses
"~~/src/Tools/cache_io.ML"
+ "~~/src/Tools/auto_tools.ML"
"~~/src/Tools/auto_solve.ML"
- "~~/src/Tools/auto_counterexample.ML"
"~~/src/Tools/quickcheck.ML"
"~~/src/Tools/value.ML"
"~~/src/Tools/Code/code_preproc.ML"
@@ -26,7 +26,8 @@
begin
setup {*
- Code_Preproc.setup
+ Auto_Solve.setup
+ #> Code_Preproc.setup
#> Code_Simp.setup
#> Code_ML.setup
#> Code_Haskell.setup
--- a/src/Tools/Metis/make-metis Mon Sep 13 16:44:20 2010 +0200
+++ b/src/Tools/Metis/make-metis Mon Sep 13 21:24:10 2010 +0200
@@ -29,19 +29,19 @@
if [ "$(basename "$FILE" .sig)" != "$FILE" ]
then
echo -e "$FILE (global)" >&2
- "$THIS/scripts/mlpp" -c isabelle "src/$FILE" | \
+ "$THIS/scripts/mlpp" -c polyml "src/$FILE" | \
perl -p -e 's/\b([A-Za-z]+\.[A-Za-z]+)/Metis.\1/g;' -e 's/\bfrag\b/Metis.frag/;' | \
perl -p -e 's/\bref\b/Unsynchronized.ref/g;'
elif fgrep -q functor "src/$FILE"
then
- "$THIS/scripts/mlpp" -c isabelle "src/$FILE" | \
+ "$THIS/scripts/mlpp" -c polyml "src/$FILE" | \
perl -p -e 's/ union / op union /g;' -e 's/ subset / op subset /g;' | \
perl -p -e 's/\bref\b/Unsynchronized.ref/g;'
else
echo -e "$FILE (local)" >&2
echo "structure Metis = struct open Metis"
cat < "metis_env.ML"
- "$THIS/scripts/mlpp" -c isabelle "src/$FILE" | \
+ "$THIS/scripts/mlpp" -c polyml "src/$FILE" | \
perl -p -e 's/\bref\b/Unsynchronized.ref/g;'
echo "end;"
fi
--- a/src/Tools/Metis/metis.ML Mon Sep 13 16:44:20 2010 +0200
+++ b/src/Tools/Metis/metis.ML Mon Sep 13 21:24:10 2010 +0200
@@ -1,3 +1,19 @@
+(*
+ This file was generated by the "make-metis" script. A few changes were done
+ manually on the script's output; these are marked as follows:
+
+ MODIFIED by Jasmin Blanchette
+
+ Some of these changes are needed so that the ML files compiles at all. Others
+ are old tweaks by Lawrence C. Paulson that are needed, if nothing else, for
+ backward compatibility. The BSD License is used with Joe Hurd's kind
+ permission. Extract from a September 13, 2010 email written by Joe Hurd:
+
+ I hereby give permission to the Isabelle team to release Metis as part
+ of Isabelle, with the Metis code covered under the Isabelle BSD
+ license.
+*)
+
(******************************************************************)
(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
(* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *)
@@ -8,11 +24,95 @@
structure Metis = struct structure Word = Word structure Array = Array end;
+(**** Original file: Random.sig ****)
+
+(* Title: Tools/random_word.ML
+ Author: Makarius
+
+Simple generator for pseudo-random numbers, using unboxed word
+arithmetic only. Unprotected concurrency introduces some true
+randomness.
+*)
+
+signature Random =
+sig
+
+ val nextWord : unit -> word
+
+ val nextBool : unit -> bool
+
+ val nextInt : int -> int (* k -> [0,k) *)
+
+ val nextReal : unit -> real (* () -> [0,1) *)
+
+end;
+
+(**** Original file: Random.sml ****)
+
+structure Metis = struct open Metis
+(* Metis-specific ML environment *)
+nonfix ++ -- RL;
+val explode = String.explode;
+val implode = String.implode;
+val print = TextIO.print;
+val foldl = List.foldl;
+val foldr = List.foldr;
+
+(* Title: Tools/random_word.ML
+ Author: Makarius
+
+Simple generator for pseudo-random numbers, using unboxed word
+arithmetic only. Unprotected concurrency introduces some true
+randomness.
+*)
+
+structure Random :> Random =
+struct
+
+(* random words: 0w0 <= result <= max_word *)
+
+(*minimum length of unboxed words on all supported ML platforms*)
+val _ = Word.wordSize >= 30
+ orelse raise Fail ("Bad platform word size");
+
+val max_word = 0wx3FFFFFFF;
+val top_bit = 0wx20000000;
+
+(*multiplier according to Borosh and Niederreiter (for modulus = 2^30),
+ see http://random.mat.sbg.ac.at/~charly/server/server.html*)
+val a = 0w777138309;
+fun step x = Word.andb (a * x + 0w1, max_word);
+
+fun change r f = r := f (!r);
+local val rand = (*Unsynchronized.*)Unsynchronized.ref 0w1
+in fun nextWord () = ((*Unsynchronized.*)change rand step; ! rand) end;
+
+(*NB: higher bits are more random than lower ones*)
+fun nextBool () = Word.andb (nextWord (), top_bit) = 0w0;
+
+
+(* random integers: 0 <= result < k *)
+
+val max_int = Word.toInt max_word;
+
+fun nextInt k =
+ if k <= 0 orelse k > max_int then raise Fail ("next_int: out of range")
+ else if k = max_int then Word.toInt (nextWord ())
+ else Word.toInt (Word.mod (nextWord (), Word.fromInt k));
+
+(* random reals: 0.0 <= result < 1.0 *)
+
+val scaling = real max_int + 1.0;
+fun nextReal () = real (Word.toInt (nextWord ())) / scaling;
+
+end;
+end;
+
(**** Original file: Portable.sig ****)
(* ========================================================================= *)
(* ML SPECIFIC FUNCTIONS *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Portable =
@@ -37,12 +137,6 @@
val time : ('a -> 'b) -> 'a -> 'b
(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing) *)
-(* ------------------------------------------------------------------------- *)
-
-val CRITICAL: (unit -> 'a) -> 'a
-
-(* ------------------------------------------------------------------------- *)
(* Generating random values. *)
(* ------------------------------------------------------------------------- *)
@@ -56,11 +150,11 @@
end
-(**** Original file: PortableIsabelle.sml ****)
+(**** Original file: PortablePolyml.sml ****)
structure Metis = struct open Metis
(* Metis-specific ML environment *)
-nonfix ++ -- RL mem;
+nonfix ++ -- RL;
val explode = String.explode;
val implode = String.implode;
val print = TextIO.print;
@@ -68,7 +162,8 @@
val foldr = List.foldr;
(* ========================================================================= *)
-(* Isabelle ML SPECIFIC FUNCTIONS *)
+(* POLY/ML SPECIFIC FUNCTIONS *)
+(* Copyright (c) 2008 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Portable :> Portable =
@@ -78,36 +173,65 @@
(* The ML implementation. *)
(* ------------------------------------------------------------------------- *)
-val ml = ml_system;
+val ml = "polyml";
(* ------------------------------------------------------------------------- *)
(* Pointer equality using the run-time system. *)
(* ------------------------------------------------------------------------- *)
-val pointerEqual = pointer_eq;
-
-(* ------------------------------------------------------------------------- *)
-(* Timing function applications a la Mosml.time. *)
-(* ------------------------------------------------------------------------- *)
-
-val time = timeap;
-
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing) *)
-(* ------------------------------------------------------------------------- *)
-
-fun CRITICAL e = NAMED_CRITICAL "metis" e;
+fun pointerEqual (x : 'a, y : 'a) = PolyML.pointerEq(x,y);
+
+(* ------------------------------------------------------------------------- *)
+(* Timing function applications. *)
+(* ------------------------------------------------------------------------- *)
+
+fun time f x =
+ let
+ fun p t =
+ let
+ val s = Time.fmt 3 t
+ in
+ case size (List.last (String.fields (fn x => x = #".") s)) of
+ 3 => s
+ | 2 => s ^ "0"
+ | 1 => s ^ "00"
+ | _ => raise Fail "Portable.time"
+ end
+
+ val c = Timer.startCPUTimer ()
+
+ val r = Timer.startRealTimer ()
+
+ fun pt () =
+ let
+ val {usr,sys} = Timer.checkCPUTimer c
+ val real = Timer.checkRealTimer r
+ in
+ print
+ ("User: " ^ p usr ^ " System: " ^ p sys ^
+ " Real: " ^ p real ^ "\n")
+ end
+
+ val y = f x handle e => (pt (); raise e)
+
+ val () = pt ()
+ in
+ y
+ end;
(* ------------------------------------------------------------------------- *)
(* Generating random values. *)
(* ------------------------------------------------------------------------- *)
-val randomWord = Random_Word.next_word;
-val randomBool = Random_Word.next_bool;
-fun randomInt n = Random_Word.next_int 0 (n - 1);
-val randomReal = Random_Word.next_real;
-
-end;
+val randomWord = Random.nextWord;
+
+val randomBool = Random.nextBool;
+
+val randomInt = Random.nextInt;
+
+val randomReal = Random.nextReal;
+
+end
(* ------------------------------------------------------------------------- *)
(* Quotations a la Moscow ML. *)
@@ -116,868 +240,11 @@
datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a;
end;
-(**** Original file: PP.sig ****)
-
-(* ========================================================================= *)
-(* PP -- pretty-printing -- from the SML/NJ library *)
-(* *)
-(* Modified for Moscow ML from SML/NJ Library version 0.2 *)
-(* *)
-(* COPYRIGHT (c) 1992 by AT&T Bell Laboratories. *)
-(* *)
-(* STANDARD ML OF NEW JERSEY COPYRIGHT NOTICE, LICENSE AND DISCLAIMER. *)
-(* *)
-(* Permission to use, copy, modify, and distribute this software and its *)
-(* documentation for any purpose and without fee is hereby granted, *)
-(* provided that the above copyright notice appear in all copies and that *)
-(* both the copyright notice and this permission notice and warranty *)
-(* disclaimer appear in supporting documentation, and that the name of *)
-(* AT&T Bell Laboratories or any AT&T entity not be used in advertising *)
-(* or publicity pertaining to distribution of the software without *)
-(* specific, written prior permission. *)
-(* *)
-(* AT&T disclaims all warranties with regard to this software, including *)
-(* all implied warranties of merchantability and fitness. In no event *)
-(* shall AT&T be liable for any special, indirect or consequential *)
-(* damages or any damages whatsoever resulting from loss of use, data or *)
-(* profits, whether in an action of contract, negligence or other *)
-(* tortious action, arising out of or in connection with the use or *)
-(* performance of this software. *)
-(* ========================================================================= *)
-
-signature PP =
-sig
-
- type ppstream
-
- type ppconsumer =
- {consumer : string -> unit,
- linewidth : int,
- flush : unit -> unit}
-
- datatype break_style =
- CONSISTENT
- | INCONSISTENT
-
- val mk_ppstream : ppconsumer -> ppstream
-
- val dest_ppstream : ppstream -> ppconsumer
-
- val add_break : ppstream -> int * int -> unit
-
- val add_newline : ppstream -> unit
-
- val add_string : ppstream -> string -> unit
-
- val begin_block : ppstream -> break_style -> int -> unit
-
- val end_block : ppstream -> unit
-
- val clear_ppstream : ppstream -> unit
-
- val flush_ppstream : ppstream -> unit
-
- val with_pp : ppconsumer -> (ppstream -> unit) -> unit
-
- val pp_to_string : int -> (ppstream -> 'a -> unit) -> 'a -> string
-
-end
-
-(*
- This structure provides tools for creating customized Oppen-style
- pretty-printers, based on the type ppstream. A ppstream is an
- output stream that contains prettyprinting commands. The commands
- are placed in the stream by various function calls listed below.
-
- There following primitives add commands to the stream:
- begin_block, end_block, add_string, add_break, and add_newline.
- All calls to add_string, add_break, and add_newline must happen
- between a pair of calls to begin_block and end_block must be
- properly nested dynamically. All calls to begin_block and
- end_block must be properly nested (dynamically).
-
- [ppconsumer] is the type of sinks for pretty-printing. A value of
- type ppconsumer is a record
- { consumer : string -> unit,
- linewidth : int,
- flush : unit -> unit }
- of a string consumer, a specified linewidth, and a flush function
- which is called whenever flush_ppstream is called.
-
- A prettyprinter can be called outright to print a value. In
- addition, a prettyprinter for a base type or nullary datatype ty
- can be installed in the top-level system. Then the installed
- prettyprinter will be invoked automatically whenever a value of
- type ty is to be printed.
-
- [break_style] is the type of line break styles for blocks:
-
- [CONSISTENT] specifies that if any line break occurs inside the
- block, then all indicated line breaks occur.
-
- [INCONSISTENT] specifies that breaks will be inserted to only to
- avoid overfull lines.
-
- [mk_ppstream {consumer, linewidth, flush}] creates a new ppstream
- which invokes the consumer to output text, putting at most
- linewidth characters on each line.
-
- [dest_ppstream ppstrm] extracts the linewidth, flush function, and
- consumer from a ppstream.
-
- [add_break ppstrm (size, offset)] notifies the pretty-printer that
- a line break is possible at this point.
- * When the current block style is CONSISTENT:
- ** if the entire block fits on the remainder of the line, then
- output size spaces; else
- ** increase the current indentation by the block offset;
- further indent every item of the block by offset, and add
- one newline at every add_break in the block.
- * When the current block style is INCONSISTENT:
- ** if the next component of the block fits on the remainder of
- the line, then output size spaces; else
- ** issue a newline and indent to the current indentation level
- plus the block offset plus the offset.
-
- [add_newline ppstrm] issues a newline.
-
- [add_string ppstrm str] outputs the string str to the ppstream.
-
- [begin_block ppstrm style blockoffset] begins a new block and
- level of indentation, with the given style and block offset.
-
- [end_block ppstrm] closes the current block.
-
- [clear_ppstream ppstrm] restarts the stream, without affecting the
- underlying consumer.
-
- [flush_ppstream ppstrm] executes any remaining commands in the
- ppstream (that is, flushes currently accumulated output to the
- consumer associated with ppstrm); executes the flush function
- associated with the consumer; and calls clear_ppstream.
-
- [with_pp consumer f] makes a new ppstream from the consumer and
- applies f (which can be thought of as a producer) to that
- ppstream, then flushed the ppstream and returns the value of f.
-
- [pp_to_string linewidth printit x] constructs a new ppstream
- ppstrm whose consumer accumulates the output in a string s. Then
- evaluates (printit ppstrm x) and finally returns the string s.
-
-
- Example 1: A simple prettyprinter for Booleans:
-
- load "PP";
- fun ppbool pps d =
- let open PP
- in
- begin_block pps INCONSISTENT 6;
- add_string pps (if d then "right" else "wrong");
- end_block pps
- end;
-
- Now one may define a ppstream to print to, and exercise it:
-
- val ppstrm = Metis.PP.mk_ppstream {consumer =
- fn s => Metis.TextIO.output(Metis.TextIO.stdOut, s),
- linewidth = 72,
- flush =
- fn () => Metis.TextIO.flushOut Metis.TextIO.stdOut};
-
- fun ppb b = (ppbool ppstrm b; Metis.PP.flush_ppstream ppstrm);
-
- - ppb false;
- wrong> val it = () : unit
-
- The prettyprinter may also be installed in the toplevel system;
- then it will be used to print all expressions of type bool
- subsequently computed:
-
- - installPP ppbool;
- > val it = () : unit
- - 1=0;
- > val it = wrong : bool
- - 1=1;
- > val it = right : bool
-
- See library Meta for a description of installPP.
-
-
- Example 2: Prettyprinting simple expressions (examples/pretty/Metis.ppexpr.sml):
-
- datatype expr =
- Cst of int
- | Neg of expr
- | Plus of expr * expr
-
- fun ppexpr pps e0 =
- let open PP
- fun ppe (Cst i) = add_string pps (Metis.Int.toString i)
- | ppe (Neg e) = (add_string pps "~"; ppe e)
- | ppe (Plus(e1, e2)) = (begin_block pps CONSISTENT 0;
- add_string pps "(";
- ppe e1;
- add_string pps " + ";
- add_break pps (0, 1);
- ppe e2;
- add_string pps ")";
- end_block pps)
- in
- begin_block pps INCONSISTENT 0;
- ppe e0;
- end_block pps
- end
-
- val _ = installPP ppexpr;
-
- (* Some example values: *)
-
- val e1 = Cst 1;
- val e2 = Cst 2;
- val e3 = Plus(e1, Neg e2);
- val e4 = Plus(Neg e3, e3);
- val e5 = Plus(Neg e4, e4);
- val e6 = Plus(e5, e5);
- val e7 = Plus(e6, e6);
- val e8 =
- Plus(e3, Plus(e3, Plus(e3, Plus(e3, Plus(e3, Plus(e3, e7))))));
-*)
-
-(**** Original file: PP.sml ****)
-
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL mem;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
-(* ========================================================================= *)
-(* PP -- pretty-printing -- from the SML/NJ library *)
-(* *)
-(* Modified for Moscow ML from SML/NJ Library version 0.2 *)
-(* *)
-(* COPYRIGHT (c) 1992 by AT&T Bell Laboratories. *)
-(* *)
-(* STANDARD ML OF NEW JERSEY COPYRIGHT NOTICE, LICENSE AND DISCLAIMER. *)
-(* *)
-(* Permission to use, copy, modify, and distribute this software and its *)
-(* documentation for any purpose and without fee is hereby granted, *)
-(* provided that the above copyright notice appear in all copies and that *)
-(* both the copyright notice and this permission notice and warranty *)
-(* disclaimer appear in supporting documentation, and that the name of *)
-(* AT&T Bell Laboratories or any AT&T entity not be used in advertising *)
-(* or publicity pertaining to distribution of the software without *)
-(* specific, written prior permission. *)
-(* *)
-(* AT&T disclaims all warranties with regard to this software, including *)
-(* all implied warranties of merchantability and fitness. In no event *)
-(* shall AT&T be liable for any special, indirect or consequential *)
-(* damages or any damages whatsoever resulting from loss of use, data or *)
-(* profits, whether in an action of contract, negligence or other *)
-(* tortious action, arising out of or in connection with the use or *)
-(* performance of this software. *)
-(* ========================================================================= *)
-
-structure PP :> PP =
-struct
-
-open Array
-infix 9 sub
-
-(* the queue library, formerly in unit Ppqueue *)
-
-datatype Qend = Qback | Qfront
-
-exception QUEUE_FULL
-exception QUEUE_EMPTY
-exception REQUESTED_QUEUE_SIZE_TOO_SMALL
-
-local
- fun ++ i n = (i + 1) mod n
- fun -- i n = (i - 1) mod n
-in
-
-abstype 'a queue = QUEUE of {elems: 'a array, (* the contents *)
- front: int Unsynchronized.ref,
- back: int Unsynchronized.ref,
- size: int} (* fixed size of element array *)
-with
-
- fun is_empty (QUEUE{front=Unsynchronized.ref ~1, back=Unsynchronized.ref ~1,...}) = true
- | is_empty _ = false
-
- fun mk_queue n init_val =
- if (n < 2)
- then raise REQUESTED_QUEUE_SIZE_TOO_SMALL
- else QUEUE{elems=array(n, init_val), front=Unsynchronized.ref ~1, back=Unsynchronized.ref ~1, size=n}
-
- fun clear_queue (QUEUE{front,back,...}) = (front := ~1; back := ~1)
-
- fun queue_at Qfront (QUEUE{elems,front,...}) = elems sub !front
- | queue_at Qback (QUEUE{elems,back,...}) = elems sub !back
-
- fun en_queue Qfront item (Q as QUEUE{elems,front,back,size}) =
- if (is_empty Q)
- then (front := 0; back := 0;
- update(elems,0,item))
- else let val i = --(!front) size
- in if (i = !back)
- then raise QUEUE_FULL
- else (update(elems,i,item); front := i)
- end
- | en_queue Qback item (Q as QUEUE{elems,front,back,size}) =
- if (is_empty Q)
- then (front := 0; back := 0;
- update(elems,0,item))
- else let val i = ++(!back) size
- in if (i = !front)
- then raise QUEUE_FULL
- else (update(elems,i,item); back := i)
- end
-
- fun de_queue Qfront (Q as QUEUE{front,back,size,...}) =
- if (!front = !back) (* unitary queue *)
- then clear_queue Q
- else front := ++(!front) size
- | de_queue Qback (Q as QUEUE{front,back,size,...}) =
- if (!front = !back)
- then clear_queue Q
- else back := --(!back) size
-
-end (* abstype queue *)
-end (* local *)
-
-
-val magic: 'a -> 'a = fn x => x
-
-(* exception PP_FAIL of string *)
-
-datatype break_style = CONSISTENT | INCONSISTENT
-
-datatype break_info
- = FITS
- | PACK_ONTO_LINE of int
- | ONE_PER_LINE of int
-
-(* Some global values *)
-val INFINITY = 999999
-
-abstype indent_stack = Istack of break_info list Unsynchronized.ref
-with
- fun mk_indent_stack() = Istack (Unsynchronized.ref([]:break_info list))
- fun clear_indent_stack (Istack stk) = (stk := ([]:break_info list))
- fun top (Istack stk) =
- case !stk
- of nil => raise Fail "PP-error: top: badly formed block"
- | x::_ => x
- fun push (x,(Istack stk)) = stk := x::(!stk)
- fun pop (Istack stk) =
- case !stk
- of nil => raise Fail "PP-error: pop: badly formed block"
- | _::rest => stk := rest
-end
-
-(* The delim_stack is used to compute the size of blocks. It is
- a stack of indices into the token buffer. The indices only point to
- BBs, Es, and BRs. We push BBs and Es onto the stack until a BR
- is encountered. Then we compute sizes and pop. When we encounter
- a BR in the middle of a block, we compute the Distance_to_next_break
- of the previous BR in the block, if there was one.
-
- We need to be able to delete from the bottom of the delim_stack, so
- we use a queue, treated with a stack discipline, i.e., we only add
- items at the head of the queue, but can delete from the front or
- back of the queue.
-*)
-abstype delim_stack = Dstack of int queue
-with
- fun new_delim_stack i = Dstack(mk_queue i ~1)
- fun reset_delim_stack (Dstack q) = clear_queue q
-
- fun pop_delim_stack (Dstack d) = de_queue Qfront d
- fun pop_bottom_delim_stack (Dstack d) = de_queue Qback d
-
- fun push_delim_stack(i,Dstack d) = en_queue Qfront i d
- fun top_delim_stack (Dstack d) = queue_at Qfront d
- fun bottom_delim_stack (Dstack d) = queue_at Qback d
- fun delim_stack_is_empty (Dstack d) = is_empty d
-end
-
-
-type block_info = { Block_size : int Unsynchronized.ref,
- Block_offset : int,
- How_to_indent : break_style }
-
-
-(* Distance_to_next_break includes Number_of_blanks. Break_offset is
- a local offset for the break. BB represents a sequence of contiguous
- Begins. E represents a sequence of contiguous Ends.
-*)
-datatype pp_token
- = S of {String : string, Length : int}
- | BB of {Pblocks : block_info list Unsynchronized.ref, (* Processed *)
- Ublocks : block_info list Unsynchronized.ref} (* Unprocessed *)
- | E of {Pend : int Unsynchronized.ref, Uend : int Unsynchronized.ref}
- | BR of {Distance_to_next_break : int Unsynchronized.ref,
- Number_of_blanks : int,
- Break_offset : int}
-
-
-(* The initial values in the token buffer *)
-val initial_token_value = S{String = "", Length = 0}
-
-(* type ppstream = General.ppstream; *)
-datatype ppstream_ =
- PPS of
- {consumer : string -> unit,
- linewidth : int,
- flush : unit -> unit,
- the_token_buffer : pp_token array,
- the_delim_stack : delim_stack,
- the_indent_stack : indent_stack,
- ++ : int Unsynchronized.ref -> unit, (* increment circular buffer index *)
- space_left : int Unsynchronized.ref, (* remaining columns on page *)
- left_index : int Unsynchronized.ref, (* insertion index *)
- right_index : int Unsynchronized.ref, (* output index *)
- left_sum : int Unsynchronized.ref, (* size of strings and spaces inserted *)
- right_sum : int Unsynchronized.ref} (* size of strings and spaces printed *)
-
-type ppstream = ppstream_
-
-type ppconsumer = {consumer : string -> unit,
- linewidth : int,
- flush : unit -> unit}
-
-fun mk_ppstream {consumer,linewidth,flush} =
- if (linewidth<5)
- then raise Fail "PP-error: linewidth too_small"
- else let val buf_size = 3*linewidth
- in magic(
- PPS{consumer = consumer,
- linewidth = linewidth,
- flush = flush,
- the_token_buffer = array(buf_size, initial_token_value),
- the_delim_stack = new_delim_stack buf_size,
- the_indent_stack = mk_indent_stack (),
- ++ = fn i => i := ((!i + 1) mod buf_size),
- space_left = Unsynchronized.ref linewidth,
- left_index = Unsynchronized.ref 0, right_index = Unsynchronized.ref 0,
- left_sum = Unsynchronized.ref 0, right_sum = Unsynchronized.ref 0}
- ) : ppstream
- end
-
-fun dest_ppstream(pps : ppstream) =
- let val PPS{consumer,linewidth,flush, ...} = magic pps
- in {consumer=consumer,linewidth=linewidth,flush=flush} end
-
-local
- val space = " "
- fun mk_space (0,s) = String.concat s
- | mk_space (n,s) = mk_space((n-1), (space::s))
- val space_table = Vector.tabulate(100, fn i => mk_space(i,[]))
- fun nspaces n = Vector.sub(space_table, n)
- handle General.Subscript =>
- if n < 0
- then ""
- else let val n2 = n div 2
- val n2_spaces = nspaces n2
- val extra = if (n = (2*n2)) then "" else space
- in String.concat [n2_spaces, n2_spaces, extra]
- end
-in
- fun cr_indent (ofn, i) = ofn ("\n"^(nspaces i))
- fun indent (ofn,i) = ofn (nspaces i)
-end
-
-
-(* Print a the first member of a contiguous sequence of Begins. If there
- are "processed" Begins, then take the first off the list. If there are
- no processed Begins, take the last member off the "unprocessed" list.
- This works because the unprocessed list is treated as a stack, the
- processed list as a FIFO queue. How can an item on the unprocessed list
- be printable? Because of what goes on in add_string. See there for details.
-*)
-
-fun print_BB (_,{Pblocks = Unsynchronized.ref [], Ublocks = Unsynchronized.ref []}) =
- raise Fail "PP-error: print_BB"
- | print_BB (PPS{the_indent_stack,linewidth,space_left=Unsynchronized.ref sp_left,...},
- {Pblocks as Unsynchronized.ref({How_to_indent=CONSISTENT,Block_size,
- Block_offset}::rst),
- Ublocks=Unsynchronized.ref[]}) =
- (push ((if (!Block_size > sp_left)
- then ONE_PER_LINE (linewidth - (sp_left - Block_offset))
- else FITS),
- the_indent_stack);
- Pblocks := rst)
- | print_BB(PPS{the_indent_stack,linewidth,space_left=Unsynchronized.ref sp_left,...},
- {Pblocks as Unsynchronized.ref({Block_size,Block_offset,...}::rst),Ublocks=Unsynchronized.ref[]}) =
- (push ((if (!Block_size > sp_left)
- then PACK_ONTO_LINE (linewidth - (sp_left - Block_offset))
- else FITS),
- the_indent_stack);
- Pblocks := rst)
- | print_BB (PPS{the_indent_stack, linewidth, space_left=Unsynchronized.ref sp_left,...},
- {Ublocks,...}) =
- let fun pr_end_Ublock [{How_to_indent=CONSISTENT,Block_size,Block_offset}] l =
- (push ((if (!Block_size > sp_left)
- then ONE_PER_LINE (linewidth - (sp_left - Block_offset))
- else FITS),
- the_indent_stack);
- List.rev l)
- | pr_end_Ublock [{Block_size,Block_offset,...}] l =
- (push ((if (!Block_size > sp_left)
- then PACK_ONTO_LINE (linewidth - (sp_left - Block_offset))
- else FITS),
- the_indent_stack);
- List.rev l)
- | pr_end_Ublock (a::rst) l = pr_end_Ublock rst (a::l)
- | pr_end_Ublock _ _ =
- raise Fail "PP-error: print_BB: internal error"
- in Ublocks := pr_end_Ublock(!Ublocks) []
- end
-
-
-(* Uend should always be 0 when print_E is called. *)
-fun print_E (_,{Pend = Unsynchronized.ref 0, Uend = Unsynchronized.ref 0}) =
- raise Fail "PP-error: print_E"
- | print_E (istack,{Pend, ...}) =
- let fun pop_n_times 0 = ()
- | pop_n_times n = (pop istack; pop_n_times(n-1))
- in pop_n_times(!Pend); Pend := 0
- end
-
-
-(* "cursor" is how many spaces across the page we are. *)
-
-fun print_token(PPS{consumer,space_left,...}, S{String,Length}) =
- (consumer String;
- space_left := (!space_left) - Length)
- | print_token(ppstrm,BB b) = print_BB(ppstrm,b)
- | print_token(PPS{the_indent_stack,...},E e) =
- print_E (the_indent_stack,e)
- | print_token (PPS{the_indent_stack,space_left,consumer,linewidth,...},
- BR{Distance_to_next_break,Number_of_blanks,Break_offset}) =
- (case (top the_indent_stack)
- of FITS =>
- (space_left := (!space_left) - Number_of_blanks;
- indent (consumer,Number_of_blanks))
- | (ONE_PER_LINE cursor) =>
- let val new_cursor = cursor + Break_offset
- in space_left := linewidth - new_cursor;
- cr_indent (consumer,new_cursor)
- end
- | (PACK_ONTO_LINE cursor) =>
- if (!Distance_to_next_break > (!space_left))
- then let val new_cursor = cursor + Break_offset
- in space_left := linewidth - new_cursor;
- cr_indent(consumer,new_cursor)
- end
- else (space_left := !space_left - Number_of_blanks;
- indent (consumer,Number_of_blanks)))
-
-
-fun clear_ppstream(pps : ppstream) =
- let val PPS{the_token_buffer, the_delim_stack,
- the_indent_stack,left_sum, right_sum,
- left_index, right_index,space_left,linewidth,...}
- = magic pps
- val buf_size = 3*linewidth
- fun set i =
- if (i = buf_size)
- then ()
- else (update(the_token_buffer,i,initial_token_value);
- set (i+1))
- in set 0;
- clear_indent_stack the_indent_stack;
- reset_delim_stack the_delim_stack;
- left_sum := 0; right_sum := 0;
- left_index := 0; right_index := 0;
- space_left := linewidth
- end
-
-
-(* Move insertion head to right unless adding a BB and already at a BB,
- or unless adding an E and already at an E.
-*)
-fun BB_inc_right_index(PPS{the_token_buffer, right_index, ++,...})=
- case (the_token_buffer sub (!right_index))
- of (BB _) => ()
- | _ => ++right_index
-
-fun E_inc_right_index(PPS{the_token_buffer,right_index, ++,...})=
- case (the_token_buffer sub (!right_index))
- of (E _) => ()
- | _ => ++right_index
-
-
-fun pointers_coincide(PPS{left_index,right_index,the_token_buffer,...}) =
- (!left_index = !right_index) andalso
- (case (the_token_buffer sub (!left_index))
- of (BB {Pblocks = Unsynchronized.ref [], Ublocks = Unsynchronized.ref []}) => true
- | (BB _) => false
- | (E {Pend = Unsynchronized.ref 0, Uend = Unsynchronized.ref 0}) => true
- | (E _) => false
- | _ => true)
-
-fun advance_left (ppstrm as PPS{consumer,left_index,left_sum,
- the_token_buffer,++,...},
- instr) =
- let val NEG = ~1
- val POS = 0
- fun inc_left_sum (BR{Number_of_blanks, ...}) =
- left_sum := (!left_sum) + Number_of_blanks
- | inc_left_sum (S{Length, ...}) = left_sum := (!left_sum) + Length
- | inc_left_sum _ = ()
-
- fun last_size [{Block_size, ...}:block_info] = !Block_size
- | last_size (_::rst) = last_size rst
- | last_size _ = raise Fail "PP-error: last_size: internal error"
- fun token_size (S{Length, ...}) = Length
- | token_size (BB b) =
- (case b
- of {Pblocks = Unsynchronized.ref [], Ublocks = Unsynchronized.ref []} =>
- raise Fail "PP-error: BB_size"
- | {Pblocks as Unsynchronized.ref(_::_),Ublocks=Unsynchronized.ref[]} => POS
- | {Ublocks, ...} => last_size (!Ublocks))
- | token_size (E{Pend = Unsynchronized.ref 0, Uend = Unsynchronized.ref 0}) =
- raise Fail "PP-error: token_size.E"
- | token_size (E{Pend = Unsynchronized.ref 0, ...}) = NEG
- | token_size (E _) = POS
- | token_size (BR {Distance_to_next_break, ...}) = !Distance_to_next_break
- fun loop (instr) =
- if (token_size instr < 0) (* synchronization point; cannot advance *)
- then ()
- else (print_token(ppstrm,instr);
- inc_left_sum instr;
- if (pointers_coincide ppstrm)
- then ()
- else (* increment left index *)
-
- (* When this is evaluated, we know that the left_index has not yet
- caught up to the right_index. If we are at a BB or an E, we can
- increment left_index if there is no work to be done, i.e., all Begins
- or Ends have been dealt with. Also, we should do some housekeeping and
- clear the buffer at left_index, otherwise we can get errors when
- left_index catches up to right_index and we reset the indices to 0.
- (We might find ourselves adding a BB to an "old" BB, with the result
- that the index is not pushed onto the delim_stack. This can lead to
- mangled output.)
- *)
- (case (the_token_buffer sub (!left_index))
- of (BB {Pblocks = Unsynchronized.ref [], Ublocks = Unsynchronized.ref []}) =>
- (update(the_token_buffer,!left_index,
- initial_token_value);
- ++left_index)
- | (BB _) => ()
- | (E {Pend = Unsynchronized.ref 0, Uend = Unsynchronized.ref 0}) =>
- (update(the_token_buffer,!left_index,
- initial_token_value);
- ++left_index)
- | (E _) => ()
- | _ => ++left_index;
- loop (the_token_buffer sub (!left_index))))
- in loop instr
- end
-
-
-fun begin_block (pps : ppstream) style offset =
- let val ppstrm = magic pps : ppstream_
- val PPS{the_token_buffer, the_delim_stack,left_index,
- left_sum, right_index, right_sum,...}
- = ppstrm
- in
- (if (delim_stack_is_empty the_delim_stack)
- then (left_index := 0;
- left_sum := 1;
- right_index := 0;
- right_sum := 1)
- else BB_inc_right_index ppstrm;
- case (the_token_buffer sub (!right_index))
- of (BB {Ublocks, ...}) =>
- Ublocks := {Block_size = Unsynchronized.ref (~(!right_sum)),
- Block_offset = offset,
- How_to_indent = style}::(!Ublocks)
- | _ => (update(the_token_buffer, !right_index,
- BB{Pblocks = Unsynchronized.ref [],
- Ublocks = Unsynchronized.ref [{Block_size = Unsynchronized.ref (~(!right_sum)),
- Block_offset = offset,
- How_to_indent = style}]});
- push_delim_stack (!right_index, the_delim_stack)))
- end
-
-fun end_block(pps : ppstream) =
- let val ppstrm = magic pps : ppstream_
- val PPS{the_token_buffer,the_delim_stack,right_index,...}
- = ppstrm
- in
- if (delim_stack_is_empty the_delim_stack)
- then print_token(ppstrm,(E{Pend = Unsynchronized.ref 1, Uend = Unsynchronized.ref 0}))
- else (E_inc_right_index ppstrm;
- case (the_token_buffer sub (!right_index))
- of (E{Uend, ...}) => Uend := !Uend + 1
- | _ => (update(the_token_buffer,!right_index,
- E{Uend = Unsynchronized.ref 1, Pend = Unsynchronized.ref 0});
- push_delim_stack (!right_index, the_delim_stack)))
- end
-
-local
- fun check_delim_stack(PPS{the_token_buffer,the_delim_stack,right_sum,...}) =
- let fun check k =
- if (delim_stack_is_empty the_delim_stack)
- then ()
- else case(the_token_buffer sub (top_delim_stack the_delim_stack))
- of (BB{Ublocks as Unsynchronized.ref ((b as {Block_size, ...})::rst),
- Pblocks}) =>
- if (k>0)
- then (Block_size := !right_sum + !Block_size;
- Pblocks := b :: (!Pblocks);
- Ublocks := rst;
- if (List.length rst = 0)
- then pop_delim_stack the_delim_stack
- else ();
- check(k-1))
- else ()
- | (E{Pend,Uend}) =>
- (Pend := (!Pend) + (!Uend);
- Uend := 0;
- pop_delim_stack the_delim_stack;
- check(k + !Pend))
- | (BR{Distance_to_next_break, ...}) =>
- (Distance_to_next_break :=
- !right_sum + !Distance_to_next_break;
- pop_delim_stack the_delim_stack;
- if (k>0)
- then check k
- else ())
- | _ => raise Fail "PP-error: check_delim_stack.catchall"
- in check 0
- end
-in
-
- fun add_break (pps : ppstream) (n, break_offset) =
- let val ppstrm = magic pps : ppstream_
- val PPS{the_token_buffer,the_delim_stack,left_index,
- right_index,left_sum,right_sum, ++, ...}
- = ppstrm
- in
- (if (delim_stack_is_empty the_delim_stack)
- then (left_index := 0; right_index := 0;
- left_sum := 1; right_sum := 1)
- else ++right_index;
- update(the_token_buffer, !right_index,
- BR{Distance_to_next_break = Unsynchronized.ref (~(!right_sum)),
- Number_of_blanks = n,
- Break_offset = break_offset});
- check_delim_stack ppstrm;
- right_sum := (!right_sum) + n;
- push_delim_stack (!right_index,the_delim_stack))
- end
-
- fun flush_ppstream0(pps : ppstream) =
- let val ppstrm = magic pps : ppstream_
- val PPS{the_delim_stack,the_token_buffer, flush, left_index,...}
- = ppstrm
- in
- (if (delim_stack_is_empty the_delim_stack)
- then ()
- else (check_delim_stack ppstrm;
- advance_left(ppstrm, the_token_buffer sub (!left_index)));
- flush())
- end
-
-end (* local *)
-
-
-fun flush_ppstream ppstrm =
- (flush_ppstream0 ppstrm;
- clear_ppstream ppstrm)
-
-fun add_string (pps : ppstream) s =
- let val ppstrm = magic pps : ppstream_
- val PPS{the_token_buffer,the_delim_stack,consumer,
- right_index,right_sum,left_sum,
- left_index,space_left,++,...}
- = ppstrm
- fun fnl [{Block_size, ...}:block_info] = Block_size := INFINITY
- | fnl (_::rst) = fnl rst
- | fnl _ = raise Fail "PP-error: fnl: internal error"
-
- fun set(dstack,BB{Ublocks as Unsynchronized.ref[{Block_size,...}:block_info],...}) =
- (pop_bottom_delim_stack dstack;
- Block_size := INFINITY)
- | set (_,BB {Ublocks = Unsynchronized.ref(_::rst), ...}) = fnl rst
- | set (dstack, E{Pend,Uend}) =
- (Pend := (!Pend) + (!Uend);
- Uend := 0;
- pop_bottom_delim_stack dstack)
- | set (dstack,BR{Distance_to_next_break,...}) =
- (pop_bottom_delim_stack dstack;
- Distance_to_next_break := INFINITY)
- | set _ = raise (Fail "PP-error: add_string.set")
-
- fun check_stream () =
- if ((!right_sum - !left_sum) > !space_left)
- then if (delim_stack_is_empty the_delim_stack)
- then ()
- else let val i = bottom_delim_stack the_delim_stack
- in if (!left_index = i)
- then set (the_delim_stack, the_token_buffer sub i)
- else ();
- advance_left(ppstrm,
- the_token_buffer sub (!left_index));
- if (pointers_coincide ppstrm)
- then ()
- else check_stream ()
- end
- else ()
-
- val slen = String.size s
- val S_token = S{String = s, Length = slen}
-
- in if (delim_stack_is_empty the_delim_stack)
- then print_token(ppstrm,S_token)
- else (++right_index;
- update(the_token_buffer, !right_index, S_token);
- right_sum := (!right_sum)+slen;
- check_stream ())
- end
-
-
-(* Derived form. The +2 is for peace of mind *)
-fun add_newline (pps : ppstream) =
- let val PPS{linewidth, ...} = magic pps
- in add_break pps (linewidth+2,0) end
-
-(* Derived form. Builds a ppstream, sends pretty printing commands called in
- f to the ppstream, then flushes ppstream.
-*)
-
-fun with_pp ppconsumer ppfn =
- let val ppstrm = mk_ppstream ppconsumer
- in ppfn ppstrm;
- flush_ppstream0 ppstrm
- end
- handle Fail msg =>
- (TextIO.print (">>>> Pretty-printer failure: " ^ msg ^ "\n"))
-
-fun pp_to_string linewidth ppfn ob =
- let val l = Unsynchronized.ref ([]:string list)
- fun attach s = l := (s::(!l))
- in with_pp {consumer = attach, linewidth=linewidth, flush = fn()=>()}
- (fn ppstrm => ppfn ppstrm ob);
- String.concat(List.rev(!l))
- end
-end
-end;
-
(**** Original file: Useful.sig ****)
(* ========================================================================= *)
(* ML UTILITY FUNCTIONS *)
-(* Copyright (c) 2001-2005 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Useful =
@@ -991,8 +258,6 @@
exception Bug of string
-val partial : exn -> ('a -> 'b option) -> 'a -> 'b
-
val total : ('a -> 'b) -> 'a -> 'b option
val can : ('a -> 'b) -> 'a -> bool
@@ -1023,10 +288,6 @@
val exp : ('a * 'a -> 'a) -> 'a -> int -> 'a -> 'a
-val equal : ''a -> ''a -> bool
-
-val notEqual : ''a -> ''a -> bool
-
(* ------------------------------------------------------------------------- *)
(* Pairs. *)
(* ------------------------------------------------------------------------- *)
@@ -1060,6 +321,33 @@
val mwhile : ('a -> bool) -> ('a -> 's -> 'a * 's) -> 'a -> 's -> 'a * 's
(* ------------------------------------------------------------------------- *)
+(* Equality. *)
+(* ------------------------------------------------------------------------- *)
+
+val equal : ''a -> ''a -> bool
+
+val notEqual : ''a -> ''a -> bool
+
+val listEqual : ('a -> 'a -> bool) -> 'a list -> 'a list -> bool
+
+(* ------------------------------------------------------------------------- *)
+(* Comparisons. *)
+(* ------------------------------------------------------------------------- *)
+
+val mapCompare : ('a -> 'b) -> ('b * 'b -> order) -> 'a * 'a -> order
+
+val revCompare : ('a * 'a -> order) -> 'a * 'a -> order
+
+val prodCompare :
+ ('a * 'a -> order) -> ('b * 'b -> order) -> ('a * 'b) * ('a * 'b) -> order
+
+val lexCompare : ('a * 'a -> order) -> 'a list * 'a list -> order
+
+val optionCompare : ('a * 'a -> order) -> 'a option * 'a option -> order
+
+val boolCompare : bool * bool -> order (* false < true *)
+
+(* ------------------------------------------------------------------------- *)
(* Lists: note we count elements from 0. *)
(* ------------------------------------------------------------------------- *)
@@ -1073,15 +361,11 @@
val first : ('a -> 'b option) -> 'a list -> 'b option
-val index : ('a -> bool) -> 'a list -> int option
-
val maps : ('a -> 's -> 'b * 's) -> 'a list -> 's -> 'b list * 's
val mapsPartial : ('a -> 's -> 'b option * 's) -> 'a list -> 's -> 'b list * 's
-val enumerate : 'a list -> (int * 'a) list
-
-val zipwith : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list
+val zipWith : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list
val zip : 'a list -> 'b list -> ('a * 'b) list
@@ -1091,6 +375,24 @@
val cart : 'a list -> 'b list -> ('a * 'b) list
+val takeWhile : ('a -> bool) -> 'a list -> 'a list
+
+val dropWhile : ('a -> bool) -> 'a list -> 'a list
+
+val divideWhile : ('a -> bool) -> 'a list -> 'a list * 'a list
+
+val groups : ('a * 's -> bool * 's) -> 's -> 'a list -> 'a list list
+
+val groupsBy : ('a * 'a -> bool) -> 'a list -> 'a list list
+
+val groupsByFst : (''a * 'b) list -> (''a * 'b list) list
+
+val groupsOf : int -> 'a list -> 'a list list
+
+val index : ('a -> bool) -> 'a list -> int option
+
+val enumerate : 'a list -> (int * 'a) list
+
val divide : 'a list -> int -> 'a list * 'a list (* Subscript *)
val revDivide : 'a list -> int -> 'a list * 'a list (* Subscript *)
@@ -1122,23 +424,6 @@
val distinct : ''a list -> bool
(* ------------------------------------------------------------------------- *)
-(* Comparisons. *)
-(* ------------------------------------------------------------------------- *)
-
-val mapCompare : ('a -> 'b) -> ('b * 'b -> order) -> 'a * 'a -> order
-
-val revCompare : ('a * 'a -> order) -> 'a * 'a -> order
-
-val prodCompare :
- ('a * 'a -> order) -> ('b * 'b -> order) -> ('a * 'b) * ('a * 'b) -> order
-
-val lexCompare : ('a * 'a -> order) -> 'a list * 'a list -> order
-
-val optionCompare : ('a * 'a -> order) -> 'a option * 'a option -> order
-
-val boolCompare : bool * bool -> order (* true < false *)
-
-(* ------------------------------------------------------------------------- *)
(* Sorting and searching. *)
(* ------------------------------------------------------------------------- *)
@@ -1186,12 +471,24 @@
val split : string -> string -> string list
+val capitalize : string -> string
+
val mkPrefix : string -> string -> string
val destPrefix : string -> string -> string
val isPrefix : string -> string -> bool
+val stripPrefix : (char -> bool) -> string -> string
+
+val mkSuffix : string -> string -> string
+
+val destSuffix : string -> string -> string
+
+val isSuffix : string -> string -> bool
+
+val stripSuffix : (char -> bool) -> string -> string
+
(* ------------------------------------------------------------------------- *)
(* Tables. *)
(* ------------------------------------------------------------------------- *)
@@ -1248,9 +545,11 @@
val date : unit -> string
+val readDirectory : {directory : string} -> {filename : string} list
+
val readTextFile : {filename : string} -> string
-val writeTextFile : {filename : string, contents : string} -> unit
+val writeTextFile : {contents : string, filename : string} -> unit
(* ------------------------------------------------------------------------- *)
(* Profiling and error reporting. *)
@@ -1258,6 +557,8 @@
val try : ('a -> 'b) -> 'a -> 'b
+val chat : string -> unit
+
val warn : string -> unit
val die : string -> 'exit
@@ -1274,7 +575,7 @@
structure Metis = struct open Metis
(* Metis-specific ML environment *)
-nonfix ++ -- RL mem;
+nonfix ++ -- RL;
val explode = String.explode;
val implode = String.implode;
val print = TextIO.print;
@@ -1283,43 +584,62 @@
(* ========================================================================= *)
(* ML UTILITY FUNCTIONS *)
-(* Copyright (c) 2001-2004 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Useful :> Useful =
struct
(* ------------------------------------------------------------------------- *)
-(* Exceptions *)
+(* Exceptions. *)
(* ------------------------------------------------------------------------- *)
exception Error of string;
exception Bug of string;
-fun errorToString (Error message) = "\nError: " ^ message ^ "\n"
- | errorToString _ = raise Bug "errorToString: not an Error exception";
-
-fun bugToString (Bug message) = "\nBug: " ^ message ^ "\n"
- | bugToString _ = raise Bug "bugToString: not a Bug exception";
+fun errorToStringOption err =
+ case err of
+ Error message => SOME ("Error: " ^ message)
+ | _ => NONE;
+
+(*mlton
+val () = MLton.Exn.addExnMessager errorToStringOption;
+*)
+
+fun errorToString err =
+ case errorToStringOption err of
+ SOME s => "\n" ^ s ^ "\n"
+ | NONE => raise Bug "errorToString: not an Error exception";
+
+fun bugToStringOption err =
+ case err of
+ Bug message => SOME ("Bug: " ^ message)
+ | _ => NONE;
+
+(*mlton
+val () = MLton.Exn.addExnMessager bugToStringOption;
+*)
+
+fun bugToString err =
+ case bugToStringOption err of
+ SOME s => "\n" ^ s ^ "\n"
+ | NONE => raise Bug "bugToString: not a Bug exception";
fun total f x = SOME (f x) handle Error _ => NONE;
fun can f = Option.isSome o total f;
-fun partial (e as Error _) f x = (case f x of SOME y => y | NONE => raise e)
- | partial _ _ _ = raise Bug "partial: must take an Error exception";
-
-(* ------------------------------------------------------------------------- *)
-(* Tracing *)
+(* ------------------------------------------------------------------------- *)
+(* Tracing. *)
(* ------------------------------------------------------------------------- *)
val tracePrint = Unsynchronized.ref print;
-fun trace message = !tracePrint message;
-
-(* ------------------------------------------------------------------------- *)
-(* Combinators *)
+fun trace mesg = !tracePrint mesg;
+
+(* ------------------------------------------------------------------------- *)
+(* Combinators. *)
(* ------------------------------------------------------------------------- *)
fun C f x y = f y x;
@@ -1343,12 +663,8 @@
f
end;
-val equal = fn x => fn y => x = y;
-
-val notEqual = fn x => fn y => x <> y;
-
-(* ------------------------------------------------------------------------- *)
-(* Pairs *)
+(* ------------------------------------------------------------------------- *)
+(* Pairs. *)
(* ------------------------------------------------------------------------- *)
fun fst (x,_) = x;
@@ -1366,7 +682,7 @@
val op## = fn (f,g) => fn (x,y) => (f x, g y);
(* ------------------------------------------------------------------------- *)
-(* State transformers *)
+(* State transformers. *)
(* ------------------------------------------------------------------------- *)
val unit : 'a -> 's -> 'a * 's = pair;
@@ -1380,118 +696,21 @@
fun mwhile c b = let fun f a = if c a then bind (b a) f else unit a in f end;
(* ------------------------------------------------------------------------- *)
-(* Lists *)
-(* ------------------------------------------------------------------------- *)
-
-fun cons x y = x :: y;
-
-fun hdTl l = (hd l, tl l);
-
-fun append xs ys = xs @ ys;
-
-fun singleton a = [a];
-
-fun first f [] = NONE
- | first f (x :: xs) = (case f x of NONE => first f xs | s => s);
-
-fun index p =
- let
- fun idx _ [] = NONE
- | idx n (x :: xs) = if p x then SOME n else idx (n + 1) xs
- in
- idx 0
- end;
-
-fun maps (_ : 'a -> 's -> 'b * 's) [] = unit []
- | maps f (x :: xs) =
- bind (f x) (fn y => bind (maps f xs) (fn ys => unit (y :: ys)));
-
-fun mapsPartial (_ : 'a -> 's -> 'b option * 's) [] = unit []
- | mapsPartial f (x :: xs) =
- bind
- (f x)
- (fn yo =>
- bind
- (mapsPartial f xs)
- (fn ys => unit (case yo of NONE => ys | SOME y => y :: ys)));
-
-fun enumerate l = fst (maps (fn x => fn m => ((m, x), m + 1)) l 0);
-
-fun zipwith f =
- let
- fun z l [] [] = l
- | z l (x :: xs) (y :: ys) = z (f x y :: l) xs ys
- | z _ _ _ = raise Error "zipwith: lists different lengths";
- in
- fn xs => fn ys => rev (z [] xs ys)
- end;
-
-fun zip xs ys = zipwith pair xs ys;
-
-fun unzip ab =
- foldl (fn ((x, y), (xs, ys)) => (x :: xs, y :: ys)) ([], []) (rev ab);
-
-fun cartwith f =
- let
- fun aux _ res _ [] = res
- | aux xsCopy res [] (y :: yt) = aux xsCopy res xsCopy yt
- | aux xsCopy res (x :: xt) (ys as y :: _) =
- aux xsCopy (f x y :: res) xt ys
- in
- fn xs => fn ys =>
- let val xs' = rev xs in aux xs' [] xs' (rev ys) end
- end;
-
-fun cart xs ys = cartwith pair xs ys;
-
-local
- fun revDiv acc l 0 = (acc,l)
- | revDiv _ [] _ = raise Subscript
- | revDiv acc (h :: t) n = revDiv (h :: acc) t (n - 1);
-in
- fun revDivide l = revDiv [] l;
-end;
-
-fun divide l n = let val (a,b) = revDivide l n in (rev a, b) end;
-
-fun updateNth (n,x) l =
- let
- val (a,b) = revDivide l n
- in
- case b of [] => raise Subscript | _ :: t => List.revAppend (a, x :: t)
- end;
-
-fun deleteNth n l =
- let
- val (a,b) = revDivide l n
- in
- case b of [] => raise Subscript | _ :: t => List.revAppend (a,t)
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Sets implemented with lists *)
-(* ------------------------------------------------------------------------- *)
-
-fun mem x = List.exists (equal x);
-
-fun insert x s = if mem x s then s else x :: s;
-
-fun delete x s = List.filter (not o equal x) s;
-
-fun setify s = rev (foldl (fn (v,x) => if mem v x then x else v :: x) [] s);
-
-fun union s t = foldl (fn (v,x) => if mem v t then x else v :: x) t (rev s);
-
-fun intersect s t =
- foldl (fn (v,x) => if mem v t then v :: x else x) [] (rev s);
-
-fun difference s t =
- foldl (fn (v,x) => if mem v t then x else v :: x) [] (rev s);
-
-fun subset s t = List.all (fn x => mem x t) s;
-
-fun distinct [] = true
- | distinct (x :: rest) = not (mem x rest) andalso distinct rest;
+(* Equality. *)
+(* ------------------------------------------------------------------------- *)
+
+val equal = fn x => fn y => x = y;
+
+val notEqual = fn x => fn y => x <> y;
+
+fun listEqual xEq =
+ let
+ fun xsEq [] [] = true
+ | xsEq (x1 :: xs1) (x2 :: xs2) = xEq x1 x2 andalso xsEq xs1 xs2
+ | xsEq _ _ = false
+ in
+ xsEq
+ end;
(* ------------------------------------------------------------------------- *)
(* Comparisons. *)
@@ -1527,11 +746,196 @@
| optionCompare _ (_,NONE) = GREATER
| optionCompare cmp (SOME x, SOME y) = cmp (x,y);
-fun boolCompare (true,false) = LESS
- | boolCompare (false,true) = GREATER
+fun boolCompare (false,true) = LESS
+ | boolCompare (true,false) = GREATER
| boolCompare _ = EQUAL;
(* ------------------------------------------------------------------------- *)
+(* Lists. *)
+(* ------------------------------------------------------------------------- *)
+
+fun cons x y = x :: y;
+
+fun hdTl l = (hd l, tl l);
+
+fun append xs ys = xs @ ys;
+
+fun singleton a = [a];
+
+fun first f [] = NONE
+ | first f (x :: xs) = (case f x of NONE => first f xs | s => s);
+
+fun maps (_ : 'a -> 's -> 'b * 's) [] = unit []
+ | maps f (x :: xs) =
+ bind (f x) (fn y => bind (maps f xs) (fn ys => unit (y :: ys)));
+
+fun mapsPartial (_ : 'a -> 's -> 'b option * 's) [] = unit []
+ | mapsPartial f (x :: xs) =
+ bind
+ (f x)
+ (fn yo =>
+ bind
+ (mapsPartial f xs)
+ (fn ys => unit (case yo of NONE => ys | SOME y => y :: ys)));
+
+fun zipWith f =
+ let
+ fun z l [] [] = l
+ | z l (x :: xs) (y :: ys) = z (f x y :: l) xs ys
+ | z _ _ _ = raise Error "zipWith: lists different lengths";
+ in
+ fn xs => fn ys => rev (z [] xs ys)
+ end;
+
+fun zip xs ys = zipWith pair xs ys;
+
+fun unzip ab =
+ foldl (fn ((x, y), (xs, ys)) => (x :: xs, y :: ys)) ([], []) (rev ab);
+
+fun cartwith f =
+ let
+ fun aux _ res _ [] = res
+ | aux xsCopy res [] (y :: yt) = aux xsCopy res xsCopy yt
+ | aux xsCopy res (x :: xt) (ys as y :: _) =
+ aux xsCopy (f x y :: res) xt ys
+ in
+ fn xs => fn ys =>
+ let val xs' = rev xs in aux xs' [] xs' (rev ys) end
+ end;
+
+fun cart xs ys = cartwith pair xs ys;
+
+fun takeWhile p =
+ let
+ fun f acc [] = rev acc
+ | f acc (x :: xs) = if p x then f (x :: acc) xs else rev acc
+ in
+ f []
+ end;
+
+fun dropWhile p =
+ let
+ fun f [] = []
+ | f (l as x :: xs) = if p x then f xs else l
+ in
+ f
+ end;
+
+fun divideWhile p =
+ let
+ fun f acc [] = (rev acc, [])
+ | f acc (l as x :: xs) = if p x then f (x :: acc) xs else (rev acc, l)
+ in
+ f []
+ end;
+
+fun groups f =
+ let
+ fun group acc row x l =
+ case l of
+ [] =>
+ let
+ val acc = if null row then acc else rev row :: acc
+ in
+ rev acc
+ end
+ | h :: t =>
+ let
+ val (eor,x) = f (h,x)
+ in
+ if eor then group (rev row :: acc) [h] x t
+ else group acc (h :: row) x t
+ end
+ in
+ group [] []
+ end;
+
+fun groupsBy eq =
+ let
+ fun f (x_y as (x,_)) = (not (eq x_y), x)
+ in
+ fn [] => []
+ | h :: t =>
+ case groups f h t of
+ [] => [[h]]
+ | hs :: ts => (h :: hs) :: ts
+ end;
+
+local
+ fun fstEq ((x,_),(y,_)) = x = y;
+
+ fun collapse l = (fst (hd l), map snd l);
+in
+ fun groupsByFst l = map collapse (groupsBy fstEq l);
+end;
+
+fun groupsOf n =
+ let
+ fun f (_,i) = if i = 1 then (true,n) else (false, i - 1)
+ in
+ groups f (n + 1)
+ end;
+
+fun index p =
+ let
+ fun idx _ [] = NONE
+ | idx n (x :: xs) = if p x then SOME n else idx (n + 1) xs
+ in
+ idx 0
+ end;
+
+fun enumerate l = fst (maps (fn x => fn m => ((m, x), m + 1)) l 0);
+
+local
+ fun revDiv acc l 0 = (acc,l)
+ | revDiv _ [] _ = raise Subscript
+ | revDiv acc (h :: t) n = revDiv (h :: acc) t (n - 1);
+in
+ fun revDivide l = revDiv [] l;
+end;
+
+fun divide l n = let val (a,b) = revDivide l n in (rev a, b) end;
+
+fun updateNth (n,x) l =
+ let
+ val (a,b) = revDivide l n
+ in
+ case b of [] => raise Subscript | _ :: t => List.revAppend (a, x :: t)
+ end;
+
+fun deleteNth n l =
+ let
+ val (a,b) = revDivide l n
+ in
+ case b of [] => raise Subscript | _ :: t => List.revAppend (a,t)
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Sets implemented with lists. *)
+(* ------------------------------------------------------------------------- *)
+
+fun mem x = List.exists (equal x);
+
+fun insert x s = if mem x s then s else x :: s;
+
+fun delete x s = List.filter (not o equal x) s;
+
+fun setify s = rev (foldl (fn (v,x) => if mem v x then x else v :: x) [] s);
+
+fun union s t = foldl (fn (v,x) => if mem v t then x else v :: x) t (rev s);
+
+fun intersect s t =
+ foldl (fn (v,x) => if mem v t then v :: x else x) [] (rev s);
+
+fun difference s t =
+ foldl (fn (v,x) => if mem v t then x else v :: x) [] (rev s);
+
+fun subset s t = List.all (fn x => mem x t) s;
+
+fun distinct [] = true
+ | distinct (x :: rest) = not (mem x rest) andalso distinct rest;
+
+(* ------------------------------------------------------------------------- *)
(* Sorting and searching. *)
(* ------------------------------------------------------------------------- *)
@@ -1584,7 +988,7 @@
| l as [_] => l
| h :: t => mergePairs (findRuns [] h [] t)
end;
-
+
fun sortMap _ _ [] = []
| sortMap _ _ (l as [_]) = l
| sortMap f cmp xs =
@@ -1622,49 +1026,49 @@
end;
local
- fun both f g n = f n andalso g n;
-
- fun next f = let fun nx x = if f x then x else nx (x + 1) in nx end;
-
- fun looking res 0 _ _ = rev res
- | looking res n f x =
- let
- val p = next f x
- val res' = p :: res
- val f' = both f (not o divides p)
- in
- looking res' (n - 1) f' (p + 1)
- end;
-
- fun calcPrimes n = looking [] n (K true) 2
-
- val primesList = Unsynchronized.ref (calcPrimes 10);
-in
- fun primes n = CRITICAL (fn () =>
- if length (!primesList) <= n then List.take (!primesList,n)
+ fun calcPrimes ps n i =
+ if List.exists (fn p => divides p i) ps then calcPrimes ps n (i + 1)
else
let
- val l = calcPrimes n
- val () = primesList := l
- in
- l
- end);
-
- fun primesUpTo n = CRITICAL (fn () =>
- let
- fun f k [] =
- let
- val l = calcPrimes (2 * k)
- val () = primesList := l
- in
- f k (List.drop (l,k))
- end
- | f k (p :: ps) =
- if p <= n then f (k + 1) ps else List.take (!primesList, k)
- in
- f 0 (!primesList)
- end);
-end;
+ val ps = ps @ [i]
+ and n = n - 1
+ in
+ if n = 0 then ps else calcPrimes ps n (i + 1)
+ end;
+
+ val primesList = Unsynchronized.ref [2];
+in
+ fun primes n =
+ let
+ val Unsynchronized.ref ps = primesList
+
+ val k = n - length ps
+ in
+ if k <= 0 then List.take (ps,n)
+ else
+ let
+ val ps = calcPrimes ps k (List.last ps + 1)
+
+ val () = primesList := ps
+ in
+ ps
+ end
+ end;
+end;
+
+fun primesUpTo n =
+ let
+ fun f k =
+ let
+ val l = primes k
+
+ val p = List.last l
+ in
+ if p < n then f (2 * k) else takeWhile (fn j => j <= n) l
+ end
+ in
+ f 8
+ end;
(* ------------------------------------------------------------------------- *)
(* Strings. *)
@@ -1732,7 +1136,8 @@
val trim = implode o chop o rev o chop o rev o explode;
end;
-fun join _ [] = "" | join s (h :: t) = foldl (fn (x,y) => y ^ s ^ x) h t;
+fun join _ [] = ""
+ | join s (h :: t) = foldl (fn (x,y) => y ^ s ^ x) h t;
local
fun match [] l = SOME l
@@ -1755,23 +1160,58 @@
end;
end;
-(***
-fun pluralize {singular,plural} = fn 1 => singular | _ => plural;
-***)
+fun capitalize s =
+ if s = "" then s
+ else str (Char.toUpper (String.sub (s,0))) ^ String.extract (s,1,NONE);
fun mkPrefix p s = p ^ s;
fun destPrefix p =
let
- fun check s = String.isPrefix p s orelse raise Error "destPrefix"
+ fun check s =
+ if String.isPrefix p s then ()
+ else raise Error "destPrefix"
val sizeP = size p
in
- fn s => (check s; String.extract (s,sizeP,NONE))
+ fn s =>
+ let
+ val () = check s
+ in
+ String.extract (s,sizeP,NONE)
+ end
end;
fun isPrefix p = can (destPrefix p);
+fun stripPrefix pred s =
+ Substring.string (Substring.dropl pred (Substring.full s));
+
+fun mkSuffix p s = s ^ p;
+
+fun destSuffix p =
+ let
+ fun check s =
+ if String.isSuffix p s then ()
+ else raise Error "destSuffix"
+
+ val sizeP = size p
+ in
+ fn s =>
+ let
+ val () = check s
+
+ val sizeS = size s
+ in
+ String.substring (s, 0, sizeS - sizeP)
+ end
+ end;
+
+fun isSuffix p = can (destSuffix p);
+
+fun stripSuffix pred s =
+ Substring.string (Substring.dropr pred (Substring.full s));
+
(* ------------------------------------------------------------------------- *)
(* Tables. *)
(* ------------------------------------------------------------------------- *)
@@ -1790,13 +1230,20 @@
else padding ^ entry ^ row
end
in
- zipwith pad column
- end;
-
-fun alignTable [] rows = map (K "") rows
- | alignTable [{leftAlign = true, padChar = #" "}] rows = map hd rows
- | alignTable (align :: aligns) rows =
- alignColumn align (map hd rows) (alignTable aligns (map tl rows));
+ zipWith pad column
+ end;
+
+local
+ fun alignTab aligns rows =
+ case aligns of
+ [] => map (K "") rows
+ | [{leftAlign = true, padChar = #" "}] => map hd rows
+ | align :: aligns =>
+ alignColumn align (map hd rows) (alignTab aligns (map tl rows));
+in
+ fun alignTable aligns rows =
+ if null rows then [] else alignTab aligns rows;
+end;
(* ------------------------------------------------------------------------- *)
(* Reals. *)
@@ -1839,22 +1286,22 @@
local
val generator = Unsynchronized.ref 0
in
- fun newInt () = CRITICAL (fn () =>
+ fun newInt () =
let
val n = !generator
val () = generator := n + 1
in
n
- end);
+ end;
fun newInts 0 = []
- | newInts k = CRITICAL (fn () =>
+ | newInts k =
let
val n = !generator
val () = generator := n + k
in
interval n k
- end);
+ end;
end;
fun withRef (r,new) f x =
@@ -1884,17 +1331,43 @@
fun date () = Date.fmt "%d/%m/%Y" (Date.fromTimeLocal (Time.now ()));
+fun readDirectory {directory = dir} =
+ let
+ val dirStrm = OS.FileSys.openDir dir
+
+ fun readAll acc =
+ case OS.FileSys.readDir dirStrm of
+ NONE => acc
+ | SOME file =>
+ let
+ val filename = OS.Path.joinDirFile {dir = dir, file = file}
+
+ val acc = {filename = filename} :: acc
+ in
+ readAll acc
+ end
+
+ val filenames = readAll []
+
+ val () = OS.FileSys.closeDir dirStrm
+ in
+ rev filenames
+ end;
+
fun readTextFile {filename} =
let
open TextIO
+
val h = openIn filename
+
val contents = inputAll h
+
val () = closeIn h
in
contents
end;
-fun writeTextFile {filename,contents} =
+fun writeTextFile {contents,filename} =
let
open TextIO
val h = openOut filename
@@ -1905,11 +1378,13 @@
end;
(* ------------------------------------------------------------------------- *)
-(* Profiling *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun err x s = TextIO.output (TextIO.stdErr, x ^ ": " ^ s ^ "\n");
+(* Profiling and error reporting. *)
+(* ------------------------------------------------------------------------- *)
+
+fun chat s = TextIO.output (TextIO.stdErr, s ^ "\n");
+
+local
+ fun err x s = chat (x ^ ": " ^ s);
in
fun try f x = f x
handle e as Error _ => (err "try" (errorToString e); raise e)
@@ -1959,7 +1434,7 @@
(* ========================================================================= *)
(* SUPPORT FOR LAZY EVALUATION *)
-(* Copyright (c) 2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2007 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Lazy =
@@ -1967,6 +1442,8 @@
type 'a lazy
+val quickly : 'a -> 'a lazy
+
val delay : (unit -> 'a) -> 'a lazy
val force : 'a lazy -> 'a
@@ -1979,7 +1456,7 @@
structure Metis = struct open Metis
(* Metis-specific ML environment *)
-nonfix ++ -- RL mem;
+nonfix ++ -- RL;
val explode = String.explode;
val implode = String.implode;
val print = TextIO.print;
@@ -1988,7 +1465,7 @@
(* ========================================================================= *)
(* SUPPORT FOR LAZY EVALUATION *)
-(* Copyright (c) 2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2007 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Lazy :> Lazy =
@@ -2000,16 +1477,21 @@
datatype 'a lazy = Lazy of 'a thunk Unsynchronized.ref;
+fun quickly v = Lazy (Unsynchronized.ref (Value v));
+
fun delay f = Lazy (Unsynchronized.ref (Thunk f));
-fun force (Lazy (Unsynchronized.ref (Value v))) = v
- | force (Lazy (s as Unsynchronized.ref (Thunk f))) =
- let
- val v = f ()
- val () = s := Value v
- in
- v
- end;
+fun force (Lazy s) =
+ case !s of
+ Value v => v
+ | Thunk f =>
+ let
+ val v = f ()
+
+ val () = s := Value v
+ in
+ v
+ end;
fun memoize f =
let
@@ -2021,11 +1503,370 @@
end
end;
+(**** Original file: Stream.sig ****)
+
+(* ========================================================================= *)
+(* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML *)
+(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* ========================================================================= *)
+
+signature Stream =
+sig
+
+(* ------------------------------------------------------------------------- *)
+(* The stream type. *)
+(* ------------------------------------------------------------------------- *)
+
+datatype 'a stream = Nil | Cons of 'a * (unit -> 'a stream)
+
+(* If you're wondering how to create an infinite stream: *)
+(* val stream4 = let fun s4 () = Metis.Stream.Cons (4,s4) in s4 () end; *)
+
+(* ------------------------------------------------------------------------- *)
+(* Stream constructors. *)
+(* ------------------------------------------------------------------------- *)
+
+val repeat : 'a -> 'a stream
+
+val count : int -> int stream
+
+val funpows : ('a -> 'a) -> 'a -> 'a stream
+
+(* ------------------------------------------------------------------------- *)
+(* Stream versions of standard list operations: these should all terminate. *)
+(* ------------------------------------------------------------------------- *)
+
+val cons : 'a -> (unit -> 'a stream) -> 'a stream
+
+val null : 'a stream -> bool
+
+val hd : 'a stream -> 'a (* raises Empty *)
+
+val tl : 'a stream -> 'a stream (* raises Empty *)
+
+val hdTl : 'a stream -> 'a * 'a stream (* raises Empty *)
+
+val singleton : 'a -> 'a stream
+
+val append : 'a stream -> (unit -> 'a stream) -> 'a stream
+
+val map : ('a -> 'b) -> 'a stream -> 'b stream
+
+val maps :
+ ('a -> 's -> 'b * 's) -> ('s -> 'b stream) -> 's -> 'a stream -> 'b stream
+
+val zipwith : ('a -> 'b -> 'c) -> 'a stream -> 'b stream -> 'c stream
+
+val zip : 'a stream -> 'b stream -> ('a * 'b) stream
+
+val take : int -> 'a stream -> 'a stream (* raises Subscript *)
+
+val drop : int -> 'a stream -> 'a stream (* raises Subscript *)
+
+(* ------------------------------------------------------------------------- *)
+(* Stream versions of standard list operations: these might not terminate. *)
+(* ------------------------------------------------------------------------- *)
+
+val length : 'a stream -> int
+
+val exists : ('a -> bool) -> 'a stream -> bool
+
+val all : ('a -> bool) -> 'a stream -> bool
+
+val filter : ('a -> bool) -> 'a stream -> 'a stream
+
+val foldl : ('a * 's -> 's) -> 's -> 'a stream -> 's
+
+val concat : 'a stream stream -> 'a stream
+
+val mapPartial : ('a -> 'b option) -> 'a stream -> 'b stream
+
+val mapsPartial :
+ ('a -> 's -> 'b option * 's) -> ('s -> 'b stream) -> 's ->
+ 'a stream -> 'b stream
+
+val mapConcat : ('a -> 'b stream) -> 'a stream -> 'b stream
+
+val mapsConcat :
+ ('a -> 's -> 'b stream * 's) -> ('s -> 'b stream) -> 's ->
+ 'a stream -> 'b stream
+
+(* ------------------------------------------------------------------------- *)
+(* Stream operations. *)
+(* ------------------------------------------------------------------------- *)
+
+val memoize : 'a stream -> 'a stream
+
+val listConcat : 'a list stream -> 'a stream
+
+val concatList : 'a stream list -> 'a stream
+
+val toList : 'a stream -> 'a list
+
+val fromList : 'a list -> 'a stream
+
+val toString : char stream -> string
+
+val fromString : string -> char stream
+
+val toTextFile : {filename : string} -> string stream -> unit
+
+val fromTextFile : {filename : string} -> string stream (* line by line *)
+
+end
+
+(**** Original file: Stream.sml ****)
+
+structure Metis = struct open Metis
+(* Metis-specific ML environment *)
+nonfix ++ -- RL;
+val explode = String.explode;
+val implode = String.implode;
+val print = TextIO.print;
+val foldl = List.foldl;
+val foldr = List.foldr;
+
+(* ========================================================================= *)
+(* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML *)
+(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* ========================================================================= *)
+
+structure Stream :> Stream =
+struct
+
+val K = Useful.K;
+
+val pair = Useful.pair;
+
+val funpow = Useful.funpow;
+
+(* ------------------------------------------------------------------------- *)
+(* The stream type. *)
+(* ------------------------------------------------------------------------- *)
+
+datatype 'a stream =
+ Nil
+ | Cons of 'a * (unit -> 'a stream);
+
+(* ------------------------------------------------------------------------- *)
+(* Stream constructors. *)
+(* ------------------------------------------------------------------------- *)
+
+fun repeat x = let fun rep () = Cons (x,rep) in rep () end;
+
+fun count n = Cons (n, fn () => count (n + 1));
+
+fun funpows f x = Cons (x, fn () => funpows f (f x));
+
+(* ------------------------------------------------------------------------- *)
+(* Stream versions of standard list operations: these should all terminate. *)
+(* ------------------------------------------------------------------------- *)
+
+fun cons h t = Cons (h,t);
+
+fun null Nil = true
+ | null (Cons _) = false;
+
+fun hd Nil = raise Empty
+ | hd (Cons (h,_)) = h;
+
+fun tl Nil = raise Empty
+ | tl (Cons (_,t)) = t ();
+
+fun hdTl s = (hd s, tl s);
+
+fun singleton s = Cons (s, K Nil);
+
+fun append Nil s = s ()
+ | append (Cons (h,t)) s = Cons (h, fn () => append (t ()) s);
+
+fun map f =
+ let
+ fun m Nil = Nil
+ | m (Cons (h,t)) = Cons (f h, m o t)
+ in
+ m
+ end;
+
+fun maps f g =
+ let
+ fun mm s Nil = g s
+ | mm s (Cons (x,xs)) =
+ let
+ val (y,s') = f x s
+ in
+ Cons (y, mm s' o xs)
+ end
+ in
+ mm
+ end;
+
+fun zipwith f =
+ let
+ fun z Nil _ = Nil
+ | z _ Nil = Nil
+ | z (Cons (x,xs)) (Cons (y,ys)) =
+ Cons (f x y, fn () => z (xs ()) (ys ()))
+ in
+ z
+ end;
+
+fun zip s t = zipwith pair s t;
+
+fun take 0 _ = Nil
+ | take n Nil = raise Subscript
+ | take 1 (Cons (x,_)) = Cons (x, K Nil)
+ | take n (Cons (x,xs)) = Cons (x, fn () => take (n - 1) (xs ()));
+
+fun drop n s = funpow n tl s handle Empty => raise Subscript;
+
+(* ------------------------------------------------------------------------- *)
+(* Stream versions of standard list operations: these might not terminate. *)
+(* ------------------------------------------------------------------------- *)
+
+local
+ fun len n Nil = n
+ | len n (Cons (_,t)) = len (n + 1) (t ());
+in
+ fun length s = len 0 s;
+end;
+
+fun exists pred =
+ let
+ fun f Nil = false
+ | f (Cons (h,t)) = pred h orelse f (t ())
+ in
+ f
+ end;
+
+fun all pred = not o exists (not o pred);
+
+fun filter p Nil = Nil
+ | filter p (Cons (x,xs)) =
+ if p x then Cons (x, fn () => filter p (xs ())) else filter p (xs ());
+
+fun foldl f =
+ let
+ fun fold b Nil = b
+ | fold b (Cons (h,t)) = fold (f (h,b)) (t ())
+ in
+ fold
+ end;
+
+fun concat Nil = Nil
+ | concat (Cons (Nil, ss)) = concat (ss ())
+ | concat (Cons (Cons (x, xs), ss)) =
+ Cons (x, fn () => concat (Cons (xs (), ss)));
+
+fun mapPartial f =
+ let
+ fun mp Nil = Nil
+ | mp (Cons (h,t)) =
+ case f h of
+ NONE => mp (t ())
+ | SOME h' => Cons (h', fn () => mp (t ()))
+ in
+ mp
+ end;
+
+fun mapsPartial f g =
+ let
+ fun mp s Nil = g s
+ | mp s (Cons (h,t)) =
+ let
+ val (h,s) = f h s
+ in
+ case h of
+ NONE => mp s (t ())
+ | SOME h => Cons (h, fn () => mp s (t ()))
+ end
+ in
+ mp
+ end;
+
+fun mapConcat f =
+ let
+ fun mc Nil = Nil
+ | mc (Cons (h,t)) = append (f h) (fn () => mc (t ()))
+ in
+ mc
+ end;
+
+fun mapsConcat f g =
+ let
+ fun mc s Nil = g s
+ | mc s (Cons (h,t)) =
+ let
+ val (l,s) = f h s
+ in
+ append l (fn () => mc s (t ()))
+ end
+ in
+ mc
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Stream operations. *)
+(* ------------------------------------------------------------------------- *)
+
+fun memoize Nil = Nil
+ | memoize (Cons (h,t)) = Cons (h, Lazy.memoize (fn () => memoize (t ())));
+
+fun concatList [] = Nil
+ | concatList (h :: t) = append h (fn () => concatList t);
+
+local
+ fun toLst res Nil = rev res
+ | toLst res (Cons (x, xs)) = toLst (x :: res) (xs ());
+in
+ fun toList s = toLst [] s;
+end;
+
+fun fromList [] = Nil
+ | fromList (x :: xs) = Cons (x, fn () => fromList xs);
+
+fun listConcat s = concat (map fromList s);
+
+fun toString s = implode (toList s);
+
+fun fromString s = fromList (explode s);
+
+fun toTextFile {filename = f} s =
+ let
+ val (h,close) =
+ if f = "-" then (TextIO.stdOut, K ())
+ else (TextIO.openOut f, TextIO.closeOut)
+
+ fun toFile Nil = ()
+ | toFile (Cons (x,y)) = (TextIO.output (h,x); toFile (y ()))
+
+ val () = toFile s
+ in
+ close h
+ end;
+
+fun fromTextFile {filename = f} =
+ let
+ val (h,close) =
+ if f = "-" then (TextIO.stdIn, K ())
+ else (TextIO.openIn f, TextIO.closeIn)
+
+ fun strm () =
+ case TextIO.inputLine h of
+ NONE => (close h; Nil)
+ | SOME s => Cons (s,strm)
+ in
+ memoize (strm ())
+ end;
+
+end
+end;
+
(**** Original file: Ordered.sig ****)
(* ========================================================================= *)
(* ORDERED TYPES *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Ordered =
@@ -2061,7 +1902,7 @@
structure Metis = struct open Metis
(* Metis-specific ML environment *)
-nonfix ++ -- RL mem;
+nonfix ++ -- RL;
val explode = String.explode;
val implode = String.implode;
val print = TextIO.print;
@@ -2070,51 +1911,3377 @@
(* ========================================================================= *)
(* ORDERED TYPES *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure IntOrdered =
struct type t = int val compare = Int.compare end;
+structure IntPairOrdered =
+struct
+
+type t = int * int;
+
+fun compare ((i1,j1),(i2,j2)) =
+ case Int.compare (i1,i2) of
+ LESS => LESS
+ | EQUAL => Int.compare (j1,j2)
+ | GREATER => GREATER;
+
+end;
+
structure StringOrdered =
struct type t = string val compare = String.compare end;
end;
+(**** Original file: Map.sig ****)
+
+(* ========================================================================= *)
+(* FINITE MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
+(* ========================================================================= *)
+
+signature Map =
+sig
+
+(* ------------------------------------------------------------------------- *)
+(* A type of finite maps. *)
+(* ------------------------------------------------------------------------- *)
+
+type ('key,'a) map
+
+(* ------------------------------------------------------------------------- *)
+(* Constructors. *)
+(* ------------------------------------------------------------------------- *)
+
+val new : ('key * 'key -> order) -> ('key,'a) map
+
+val singleton : ('key * 'key -> order) -> 'key * 'a -> ('key,'a) map
+
+(* ------------------------------------------------------------------------- *)
+(* Map size. *)
+(* ------------------------------------------------------------------------- *)
+
+val null : ('key,'a) map -> bool
+
+val size : ('key,'a) map -> int
+
+(* ------------------------------------------------------------------------- *)
+(* Querying. *)
+(* ------------------------------------------------------------------------- *)
+
+val peekKey : ('key,'a) map -> 'key -> ('key * 'a) option
+
+val peek : ('key,'a) map -> 'key -> 'a option
+
+val get : ('key,'a) map -> 'key -> 'a (* raises Error *)
+
+val pick : ('key,'a) map -> 'key * 'a (* an arbitrary key/value pair *)
+
+val nth : ('key,'a) map -> int -> 'key * 'a (* in the range [0,size-1] *)
+
+val random : ('key,'a) map -> 'key * 'a
+
+(* ------------------------------------------------------------------------- *)
+(* Adding. *)
+(* ------------------------------------------------------------------------- *)
+
+val insert : ('key,'a) map -> 'key * 'a -> ('key,'a) map
+
+val insertList : ('key,'a) map -> ('key * 'a) list -> ('key,'a) map
+
+(* ------------------------------------------------------------------------- *)
+(* Removing. *)
+(* ------------------------------------------------------------------------- *)
+
+val delete : ('key,'a) map -> 'key -> ('key,'a) map (* must be present *)
+
+val remove : ('key,'a) map -> 'key -> ('key,'a) map
+
+val deletePick : ('key,'a) map -> ('key * 'a) * ('key,'a) map
+
+val deleteNth : ('key,'a) map -> int -> ('key * 'a) * ('key,'a) map
+
+val deleteRandom : ('key,'a) map -> ('key * 'a) * ('key,'a) map
+
+(* ------------------------------------------------------------------------- *)
+(* Joining (all join operations prefer keys in the second map). *)
+(* ------------------------------------------------------------------------- *)
+
+val merge :
+ {first : 'key * 'a -> 'c option,
+ second : 'key * 'b -> 'c option,
+ both : ('key * 'a) * ('key * 'b) -> 'c option} ->
+ ('key,'a) map -> ('key,'b) map -> ('key,'c) map
+
+val union :
+ (('key * 'a) * ('key * 'a) -> 'a option) ->
+ ('key,'a) map -> ('key,'a) map -> ('key,'a) map
+
+val intersect :
+ (('key * 'a) * ('key * 'b) -> 'c option) ->
+ ('key,'a) map -> ('key,'b) map -> ('key,'c) map
+
+(* ------------------------------------------------------------------------- *)
+(* Set operations on the domain. *)
+(* ------------------------------------------------------------------------- *)
+
+val inDomain : 'key -> ('key,'a) map -> bool
+
+val unionDomain : ('key,'a) map -> ('key,'a) map -> ('key,'a) map
+
+val unionListDomain : ('key,'a) map list -> ('key,'a) map
+
+val intersectDomain : ('key,'a) map -> ('key,'a) map -> ('key,'a) map
+
+val intersectListDomain : ('key,'a) map list -> ('key,'a) map
+
+val differenceDomain : ('key,'a) map -> ('key,'a) map -> ('key,'a) map
+
+val symmetricDifferenceDomain : ('key,'a) map -> ('key,'a) map -> ('key,'a) map
+
+val equalDomain : ('key,'a) map -> ('key,'a) map -> bool
+
+val subsetDomain : ('key,'a) map -> ('key,'a) map -> bool
+
+val disjointDomain : ('key,'a) map -> ('key,'a) map -> bool
+
+(* ------------------------------------------------------------------------- *)
+(* Mapping and folding. *)
+(* ------------------------------------------------------------------------- *)
+
+val mapPartial : ('key * 'a -> 'b option) -> ('key,'a) map -> ('key,'b) map
+
+val map : ('key * 'a -> 'b) -> ('key,'a) map -> ('key,'b) map
+
+val app : ('key * 'a -> unit) -> ('key,'a) map -> unit
+
+val transform : ('a -> 'b) -> ('key,'a) map -> ('key,'b) map
+
+val filter : ('key * 'a -> bool) -> ('key,'a) map -> ('key,'a) map
+
+val partition :
+ ('key * 'a -> bool) -> ('key,'a) map -> ('key,'a) map * ('key,'a) map
+
+val foldl : ('key * 'a * 's -> 's) -> 's -> ('key,'a) map -> 's
+
+val foldr : ('key * 'a * 's -> 's) -> 's -> ('key,'a) map -> 's
+
+(* ------------------------------------------------------------------------- *)
+(* Searching. *)
+(* ------------------------------------------------------------------------- *)
+
+val findl : ('key * 'a -> bool) -> ('key,'a) map -> ('key * 'a) option
+
+val findr : ('key * 'a -> bool) -> ('key,'a) map -> ('key * 'a) option
+
+val firstl : ('key * 'a -> 'b option) -> ('key,'a) map -> 'b option
+
+val firstr : ('key * 'a -> 'b option) -> ('key,'a) map -> 'b option
+
+val exists : ('key * 'a -> bool) -> ('key,'a) map -> bool
+
+val all : ('key * 'a -> bool) -> ('key,'a) map -> bool
+
+val count : ('key * 'a -> bool) -> ('key,'a) map -> int
+
+(* ------------------------------------------------------------------------- *)
+(* Comparing. *)
+(* ------------------------------------------------------------------------- *)
+
+val compare : ('a * 'a -> order) -> ('key,'a) map * ('key,'a) map -> order
+
+val equal : ('a -> 'a -> bool) -> ('key,'a) map -> ('key,'a) map -> bool
+
+(* ------------------------------------------------------------------------- *)
+(* Converting to and from lists. *)
+(* ------------------------------------------------------------------------- *)
+
+val keys : ('key,'a) map -> 'key list
+
+val values : ('key,'a) map -> 'a list
+
+val toList : ('key,'a) map -> ('key * 'a) list
+
+val fromList : ('key * 'key -> order) -> ('key * 'a) list -> ('key,'a) map
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printing. *)
+(* ------------------------------------------------------------------------- *)
+
+val toString : ('key,'a) map -> string
+
+(* ------------------------------------------------------------------------- *)
+(* Iterators over maps. *)
+(* ------------------------------------------------------------------------- *)
+
+type ('key,'a) iterator
+
+val mkIterator : ('key,'a) map -> ('key,'a) iterator option
+
+val mkRevIterator : ('key,'a) map -> ('key,'a) iterator option
+
+val readIterator : ('key,'a) iterator -> 'key * 'a
+
+val advanceIterator : ('key,'a) iterator -> ('key,'a) iterator option
+
+end
+
+(**** Original file: Map.sml ****)
+
+structure Metis = struct open Metis
+(* Metis-specific ML environment *)
+nonfix ++ -- RL;
+val explode = String.explode;
+val implode = String.implode;
+val print = TextIO.print;
+val foldl = List.foldl;
+val foldr = List.foldr;
+
+(* ========================================================================= *)
+(* FINITE MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
+(* ========================================================================= *)
+
+structure Map :> Map =
+struct
+
+(* ------------------------------------------------------------------------- *)
+(* Importing useful functionality. *)
+(* ------------------------------------------------------------------------- *)
+
+exception Bug = Useful.Bug;
+
+exception Error = Useful.Error;
+
+val pointerEqual = Portable.pointerEqual;
+
+val K = Useful.K;
+
+val randomInt = Portable.randomInt;
+
+val randomWord = Portable.randomWord;
+
+(* ------------------------------------------------------------------------- *)
+(* Converting a comparison function to an equality function. *)
+(* ------------------------------------------------------------------------- *)
+
+fun equalKey compareKey key1 key2 = compareKey (key1,key2) = EQUAL;
+
+(* ------------------------------------------------------------------------- *)
+(* Priorities. *)
+(* ------------------------------------------------------------------------- *)
+
+type priority = Word.word;
+
+val randomPriority = randomWord;
+
+val comparePriority = Word.compare;
+
+(* ------------------------------------------------------------------------- *)
+(* Priority search trees. *)
+(* ------------------------------------------------------------------------- *)
+
+datatype ('key,'value) tree =
+ E
+ | T of ('key,'value) node
+
+and ('key,'value) node =
+ Node of
+ {size : int,
+ priority : priority,
+ left : ('key,'value) tree,
+ key : 'key,
+ value : 'value,
+ right : ('key,'value) tree};
+
+fun lowerPriorityNode node1 node2 =
+ let
+ val Node {priority = p1, ...} = node1
+ and Node {priority = p2, ...} = node2
+ in
+ comparePriority (p1,p2) = LESS
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Tree debugging functions. *)
+(* ------------------------------------------------------------------------- *)
+
+(*BasicDebug
+local
+ fun checkSizes tree =
+ case tree of
+ E => 0
+ | T (Node {size,left,right,...}) =>
+ let
+ val l = checkSizes left
+ and r = checkSizes right
+
+ val () = if l + 1 + r = size then () else raise Bug "wrong size"
+ in
+ size
+ end;
+
+ fun checkSorted compareKey x tree =
+ case tree of
+ E => x
+ | T (Node {left,key,right,...}) =>
+ let
+ val x = checkSorted compareKey x left
+
+ val () =
+ case x of
+ NONE => ()
+ | SOME k =>
+ case compareKey (k,key) of
+ LESS => ()
+ | EQUAL => raise Bug "duplicate keys"
+ | GREATER => raise Bug "unsorted"
+
+ val x = SOME key
+ in
+ checkSorted compareKey x right
+ end;
+
+ fun checkPriorities compareKey tree =
+ case tree of
+ E => NONE
+ | T node =>
+ let
+ val Node {left,right,...} = node
+
+ val () =
+ case checkPriorities compareKey left of
+ NONE => ()
+ | SOME lnode =>
+ if not (lowerPriorityNode node lnode) then ()
+ else raise Bug "left child has greater priority"
+
+ val () =
+ case checkPriorities compareKey right of
+ NONE => ()
+ | SOME rnode =>
+ if not (lowerPriorityNode node rnode) then ()
+ else raise Bug "right child has greater priority"
+ in
+ SOME node
+ end;
+in
+ fun treeCheckInvariants compareKey tree =
+ let
+ val _ = checkSizes tree
+
+ val _ = checkSorted compareKey NONE tree
+
+ val _ = checkPriorities compareKey tree
+ in
+ tree
+ end
+ handle Error err => raise Bug err;
+end;
+*)
+
+(* ------------------------------------------------------------------------- *)
+(* Tree operations. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeNew () = E;
+
+fun nodeSize (Node {size = x, ...}) = x;
+
+fun treeSize tree =
+ case tree of
+ E => 0
+ | T x => nodeSize x;
+
+fun mkNode priority left key value right =
+ let
+ val size = treeSize left + 1 + treeSize right
+ in
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ end;
+
+fun mkTree priority left key value right =
+ let
+ val node = mkNode priority left key value right
+ in
+ T node
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Extracting the left and right spines of a tree. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeLeftSpine acc tree =
+ case tree of
+ E => acc
+ | T node => nodeLeftSpine acc node
+
+and nodeLeftSpine acc node =
+ let
+ val Node {left,...} = node
+ in
+ treeLeftSpine (node :: acc) left
+ end;
+
+fun treeRightSpine acc tree =
+ case tree of
+ E => acc
+ | T node => nodeRightSpine acc node
+
+and nodeRightSpine acc node =
+ let
+ val Node {right,...} = node
+ in
+ treeRightSpine (node :: acc) right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Singleton trees. *)
+(* ------------------------------------------------------------------------- *)
+
+fun mkNodeSingleton priority key value =
+ let
+ val size = 1
+ and left = E
+ and right = E
+ in
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ end;
+
+fun nodeSingleton (key,value) =
+ let
+ val priority = randomPriority ()
+ in
+ mkNodeSingleton priority key value
+ end;
+
+fun treeSingleton key_value =
+ let
+ val node = nodeSingleton key_value
+ in
+ T node
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Appending two trees, where every element of the first tree is less than *)
+(* every element of the second tree. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeAppend tree1 tree2 =
+ case tree1 of
+ E => tree2
+ | T node1 =>
+ case tree2 of
+ E => tree1
+ | T node2 =>
+ if lowerPriorityNode node1 node2 then
+ let
+ val Node {priority,left,key,value,right,...} = node2
+
+ val left = treeAppend tree1 left
+ in
+ mkTree priority left key value right
+ end
+ else
+ let
+ val Node {priority,left,key,value,right,...} = node1
+
+ val right = treeAppend right tree2
+ in
+ mkTree priority left key value right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Appending two trees and a node, where every element of the first tree is *)
+(* less than the node, which in turn is less than every element of the *)
+(* second tree. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeCombine left node right =
+ let
+ val left_node = treeAppend left (T node)
+ in
+ treeAppend left_node right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Searching a tree for a value. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treePeek compareKey pkey tree =
+ case tree of
+ E => NONE
+ | T node => nodePeek compareKey pkey node
+
+and nodePeek compareKey pkey node =
+ let
+ val Node {left,key,value,right,...} = node
+ in
+ case compareKey (pkey,key) of
+ LESS => treePeek compareKey pkey left
+ | EQUAL => SOME value
+ | GREATER => treePeek compareKey pkey right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Tree paths. *)
+(* ------------------------------------------------------------------------- *)
+
+(* Generating a path by searching a tree for a key/value pair *)
+
+fun treePeekPath compareKey pkey path tree =
+ case tree of
+ E => (path,NONE)
+ | T node => nodePeekPath compareKey pkey path node
+
+and nodePeekPath compareKey pkey path node =
+ let
+ val Node {left,key,right,...} = node
+ in
+ case compareKey (pkey,key) of
+ LESS => treePeekPath compareKey pkey ((true,node) :: path) left
+ | EQUAL => (path, SOME node)
+ | GREATER => treePeekPath compareKey pkey ((false,node) :: path) right
+ end;
+
+(* A path splits a tree into left/right components *)
+
+fun addSidePath ((wentLeft,node),(leftTree,rightTree)) =
+ let
+ val Node {priority,left,key,value,right,...} = node
+ in
+ if wentLeft then (leftTree, mkTree priority rightTree key value right)
+ else (mkTree priority left key value leftTree, rightTree)
+ end;
+
+fun addSidesPath left_right = List.foldl addSidePath left_right;
+
+fun mkSidesPath path = addSidesPath (E,E) path;
+
+(* Updating the subtree at a path *)
+
+local
+ fun updateTree ((wentLeft,node),tree) =
+ let
+ val Node {priority,left,key,value,right,...} = node
+ in
+ if wentLeft then mkTree priority tree key value right
+ else mkTree priority left key value tree
+ end;
+in
+ fun updateTreePath tree = List.foldl updateTree tree;
+end;
+
+(* Inserting a new node at a path position *)
+
+fun insertNodePath node =
+ let
+ fun insert left_right path =
+ case path of
+ [] =>
+ let
+ val (left,right) = left_right
+ in
+ treeCombine left node right
+ end
+ | (step as (_,snode)) :: rest =>
+ if lowerPriorityNode snode node then
+ let
+ val left_right = addSidePath (step,left_right)
+ in
+ insert left_right rest
+ end
+ else
+ let
+ val (left,right) = left_right
+
+ val tree = treeCombine left node right
+ in
+ updateTreePath tree path
+ end
+ in
+ insert (E,E)
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Using a key to split a node into three components: the keys comparing *)
+(* less than the supplied key, an optional equal key, and the keys comparing *)
+(* greater. *)
+(* ------------------------------------------------------------------------- *)
+
+fun nodePartition compareKey pkey node =
+ let
+ val (path,pnode) = nodePeekPath compareKey pkey [] node
+ in
+ case pnode of
+ NONE =>
+ let
+ val (left,right) = mkSidesPath path
+ in
+ (left,NONE,right)
+ end
+ | SOME node =>
+ let
+ val Node {left,key,value,right,...} = node
+
+ val (left,right) = addSidesPath (left,right) path
+ in
+ (left, SOME (key,value), right)
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Searching a tree for a key/value pair. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treePeekKey compareKey pkey tree =
+ case tree of
+ E => NONE
+ | T node => nodePeekKey compareKey pkey node
+
+and nodePeekKey compareKey pkey node =
+ let
+ val Node {left,key,value,right,...} = node
+ in
+ case compareKey (pkey,key) of
+ LESS => treePeekKey compareKey pkey left
+ | EQUAL => SOME (key,value)
+ | GREATER => treePeekKey compareKey pkey right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Inserting new key/values into the tree. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeInsert compareKey key_value tree =
+ let
+ val (key,value) = key_value
+
+ val (path,inode) = treePeekPath compareKey key [] tree
+ in
+ case inode of
+ NONE =>
+ let
+ val node = nodeSingleton (key,value)
+ in
+ insertNodePath node path
+ end
+ | SOME node =>
+ let
+ val Node {size,priority,left,right,...} = node
+
+ val node =
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ in
+ updateTreePath (T node) path
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Deleting key/value pairs: it raises an exception if the supplied key is *)
+(* not present. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeDelete compareKey dkey tree =
+ case tree of
+ E => raise Bug "Map.delete: element not found"
+ | T node => nodeDelete compareKey dkey node
+
+and nodeDelete compareKey dkey node =
+ let
+ val Node {size,priority,left,key,value,right} = node
+ in
+ case compareKey (dkey,key) of
+ LESS =>
+ let
+ val size = size - 1
+ and left = treeDelete compareKey dkey left
+
+ val node =
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ in
+ T node
+ end
+ | EQUAL => treeAppend left right
+ | GREATER =>
+ let
+ val size = size - 1
+ and right = treeDelete compareKey dkey right
+
+ val node =
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ in
+ T node
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Partial map is the basic operation for preserving tree structure. *)
+(* It applies its argument function to the elements *in order*. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeMapPartial f tree =
+ case tree of
+ E => E
+ | T node => nodeMapPartial f node
+
+and nodeMapPartial f (Node {priority,left,key,value,right,...}) =
+ let
+ val left = treeMapPartial f left
+ and vo = f (key,value)
+ and right = treeMapPartial f right
+ in
+ case vo of
+ NONE => treeAppend left right
+ | SOME value => mkTree priority left key value right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Mapping tree values. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeMap f tree =
+ case tree of
+ E => E
+ | T node => T (nodeMap f node)
+
+and nodeMap f node =
+ let
+ val Node {size,priority,left,key,value,right} = node
+
+ val left = treeMap f left
+ and value = f (key,value)
+ and right = treeMap f right
+ in
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Merge is the basic operation for joining two trees. Note that the merged *)
+(* key is always the one from the second map. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeMerge compareKey f1 f2 fb tree1 tree2 =
+ case tree1 of
+ E => treeMapPartial f2 tree2
+ | T node1 =>
+ case tree2 of
+ E => treeMapPartial f1 tree1
+ | T node2 => nodeMerge compareKey f1 f2 fb node1 node2
+
+and nodeMerge compareKey f1 f2 fb node1 node2 =
+ let
+ val Node {priority,left,key,value,right,...} = node2
+
+ val (l,kvo,r) = nodePartition compareKey key node1
+
+ val left = treeMerge compareKey f1 f2 fb l left
+ and right = treeMerge compareKey f1 f2 fb r right
+
+ val vo =
+ case kvo of
+ NONE => f2 (key,value)
+ | SOME kv => fb (kv,(key,value))
+ in
+ case vo of
+ NONE => treeAppend left right
+ | SOME value =>
+ let
+ val node = mkNodeSingleton priority key value
+ in
+ treeCombine left node right
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* A union operation on trees. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeUnion compareKey f f2 tree1 tree2 =
+ case tree1 of
+ E => tree2
+ | T node1 =>
+ case tree2 of
+ E => tree1
+ | T node2 => nodeUnion compareKey f f2 node1 node2
+
+and nodeUnion compareKey f f2 node1 node2 =
+ if pointerEqual (node1,node2) then nodeMapPartial f2 node1
+ else
+ let
+ val Node {priority,left,key,value,right,...} = node2
+
+ val (l,kvo,r) = nodePartition compareKey key node1
+
+ val left = treeUnion compareKey f f2 l left
+ and right = treeUnion compareKey f f2 r right
+
+ val vo =
+ case kvo of
+ NONE => SOME value
+ | SOME kv => f (kv,(key,value))
+ in
+ case vo of
+ NONE => treeAppend left right
+ | SOME value =>
+ let
+ val node = mkNodeSingleton priority key value
+ in
+ treeCombine left node right
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* An intersect operation on trees. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeIntersect compareKey f t1 t2 =
+ case t1 of
+ E => E
+ | T n1 =>
+ case t2 of
+ E => E
+ | T n2 => nodeIntersect compareKey f n1 n2
+
+and nodeIntersect compareKey f n1 n2 =
+ let
+ val Node {priority,left,key,value,right,...} = n2
+
+ val (l,kvo,r) = nodePartition compareKey key n1
+
+ val left = treeIntersect compareKey f l left
+ and right = treeIntersect compareKey f r right
+
+ val vo =
+ case kvo of
+ NONE => NONE
+ | SOME kv => f (kv,(key,value))
+ in
+ case vo of
+ NONE => treeAppend left right
+ | SOME value => mkTree priority left key value right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* A union operation on trees which simply chooses the second value. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeUnionDomain compareKey tree1 tree2 =
+ case tree1 of
+ E => tree2
+ | T node1 =>
+ case tree2 of
+ E => tree1
+ | T node2 =>
+ if pointerEqual (node1,node2) then tree2
+ else nodeUnionDomain compareKey node1 node2
+
+and nodeUnionDomain compareKey node1 node2 =
+ let
+ val Node {priority,left,key,value,right,...} = node2
+
+ val (l,_,r) = nodePartition compareKey key node1
+
+ val left = treeUnionDomain compareKey l left
+ and right = treeUnionDomain compareKey r right
+
+ val node = mkNodeSingleton priority key value
+ in
+ treeCombine left node right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* An intersect operation on trees which simply chooses the second value. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeIntersectDomain compareKey tree1 tree2 =
+ case tree1 of
+ E => E
+ | T node1 =>
+ case tree2 of
+ E => E
+ | T node2 =>
+ if pointerEqual (node1,node2) then tree2
+ else nodeIntersectDomain compareKey node1 node2
+
+and nodeIntersectDomain compareKey node1 node2 =
+ let
+ val Node {priority,left,key,value,right,...} = node2
+
+ val (l,kvo,r) = nodePartition compareKey key node1
+
+ val left = treeIntersectDomain compareKey l left
+ and right = treeIntersectDomain compareKey r right
+ in
+ if Option.isSome kvo then mkTree priority left key value right
+ else treeAppend left right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* A difference operation on trees. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeDifferenceDomain compareKey t1 t2 =
+ case t1 of
+ E => E
+ | T n1 =>
+ case t2 of
+ E => t1
+ | T n2 => nodeDifferenceDomain compareKey n1 n2
+
+and nodeDifferenceDomain compareKey n1 n2 =
+ if pointerEqual (n1,n2) then E
+ else
+ let
+ val Node {priority,left,key,value,right,...} = n1
+
+ val (l,kvo,r) = nodePartition compareKey key n2
+
+ val left = treeDifferenceDomain compareKey left l
+ and right = treeDifferenceDomain compareKey right r
+ in
+ if Option.isSome kvo then treeAppend left right
+ else mkTree priority left key value right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* A subset operation on trees. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeSubsetDomain compareKey tree1 tree2 =
+ case tree1 of
+ E => true
+ | T node1 =>
+ case tree2 of
+ E => false
+ | T node2 => nodeSubsetDomain compareKey node1 node2
+
+and nodeSubsetDomain compareKey node1 node2 =
+ pointerEqual (node1,node2) orelse
+ let
+ val Node {size,left,key,right,...} = node1
+ in
+ size <= nodeSize node2 andalso
+ let
+ val (l,kvo,r) = nodePartition compareKey key node2
+ in
+ Option.isSome kvo andalso
+ treeSubsetDomain compareKey left l andalso
+ treeSubsetDomain compareKey right r
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Picking an arbitrary key/value pair from a tree. *)
+(* ------------------------------------------------------------------------- *)
+
+fun nodePick node =
+ let
+ val Node {key,value,...} = node
+ in
+ (key,value)
+ end;
+
+fun treePick tree =
+ case tree of
+ E => raise Bug "Map.treePick"
+ | T node => nodePick node;
+
+(* ------------------------------------------------------------------------- *)
+(* Removing an arbitrary key/value pair from a tree. *)
+(* ------------------------------------------------------------------------- *)
+
+fun nodeDeletePick node =
+ let
+ val Node {left,key,value,right,...} = node
+ in
+ ((key,value), treeAppend left right)
+ end;
+
+fun treeDeletePick tree =
+ case tree of
+ E => raise Bug "Map.treeDeletePick"
+ | T node => nodeDeletePick node;
+
+(* ------------------------------------------------------------------------- *)
+(* Finding the nth smallest key/value (counting from 0). *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeNth n tree =
+ case tree of
+ E => raise Bug "Map.treeNth"
+ | T node => nodeNth n node
+
+and nodeNth n node =
+ let
+ val Node {left,key,value,right,...} = node
+
+ val k = treeSize left
+ in
+ if n = k then (key,value)
+ else if n < k then treeNth n left
+ else treeNth (n - (k + 1)) right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Removing the nth smallest key/value (counting from 0). *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeDeleteNth n tree =
+ case tree of
+ E => raise Bug "Map.treeDeleteNth"
+ | T node => nodeDeleteNth n node
+
+and nodeDeleteNth n node =
+ let
+ val Node {size,priority,left,key,value,right} = node
+
+ val k = treeSize left
+ in
+ if n = k then ((key,value), treeAppend left right)
+ else if n < k then
+ let
+ val (key_value,left) = treeDeleteNth n left
+
+ val size = size - 1
+
+ val node =
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ in
+ (key_value, T node)
+ end
+ else
+ let
+ val n = n - (k + 1)
+
+ val (key_value,right) = treeDeleteNth n right
+
+ val size = size - 1
+
+ val node =
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ in
+ (key_value, T node)
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Iterators. *)
+(* ------------------------------------------------------------------------- *)
+
+datatype ('key,'value) iterator =
+ LR of ('key * 'value) * ('key,'value) tree * ('key,'value) node list
+ | RL of ('key * 'value) * ('key,'value) tree * ('key,'value) node list;
+
+fun fromSpineLR nodes =
+ case nodes of
+ [] => NONE
+ | Node {key,value,right,...} :: nodes =>
+ SOME (LR ((key,value),right,nodes));
+
+fun fromSpineRL nodes =
+ case nodes of
+ [] => NONE
+ | Node {key,value,left,...} :: nodes =>
+ SOME (RL ((key,value),left,nodes));
+
+fun addLR nodes tree = fromSpineLR (treeLeftSpine nodes tree);
+
+fun addRL nodes tree = fromSpineRL (treeRightSpine nodes tree);
+
+fun treeMkIterator tree = addLR [] tree;
+
+fun treeMkRevIterator tree = addRL [] tree;
+
+fun readIterator iter =
+ case iter of
+ LR (key_value,_,_) => key_value
+ | RL (key_value,_,_) => key_value;
+
+fun advanceIterator iter =
+ case iter of
+ LR (_,tree,nodes) => addLR nodes tree
+ | RL (_,tree,nodes) => addRL nodes tree;
+
+fun foldIterator f acc io =
+ case io of
+ NONE => acc
+ | SOME iter =>
+ let
+ val (key,value) = readIterator iter
+ in
+ foldIterator f (f (key,value,acc)) (advanceIterator iter)
+ end;
+
+fun findIterator pred io =
+ case io of
+ NONE => NONE
+ | SOME iter =>
+ let
+ val key_value = readIterator iter
+ in
+ if pred key_value then SOME key_value
+ else findIterator pred (advanceIterator iter)
+ end;
+
+fun firstIterator f io =
+ case io of
+ NONE => NONE
+ | SOME iter =>
+ let
+ val key_value = readIterator iter
+ in
+ case f key_value of
+ NONE => firstIterator f (advanceIterator iter)
+ | s => s
+ end;
+
+fun compareIterator compareKey compareValue io1 io2 =
+ case (io1,io2) of
+ (NONE,NONE) => EQUAL
+ | (NONE, SOME _) => LESS
+ | (SOME _, NONE) => GREATER
+ | (SOME i1, SOME i2) =>
+ let
+ val (k1,v1) = readIterator i1
+ and (k2,v2) = readIterator i2
+ in
+ case compareKey (k1,k2) of
+ LESS => LESS
+ | EQUAL =>
+ (case compareValue (v1,v2) of
+ LESS => LESS
+ | EQUAL =>
+ let
+ val io1 = advanceIterator i1
+ and io2 = advanceIterator i2
+ in
+ compareIterator compareKey compareValue io1 io2
+ end
+ | GREATER => GREATER)
+ | GREATER => GREATER
+ end;
+
+fun equalIterator equalKey equalValue io1 io2 =
+ case (io1,io2) of
+ (NONE,NONE) => true
+ | (NONE, SOME _) => false
+ | (SOME _, NONE) => false
+ | (SOME i1, SOME i2) =>
+ let
+ val (k1,v1) = readIterator i1
+ and (k2,v2) = readIterator i2
+ in
+ equalKey k1 k2 andalso
+ equalValue v1 v2 andalso
+ let
+ val io1 = advanceIterator i1
+ and io2 = advanceIterator i2
+ in
+ equalIterator equalKey equalValue io1 io2
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* A type of finite maps. *)
+(* ------------------------------------------------------------------------- *)
+
+datatype ('key,'value) map =
+ Map of ('key * 'key -> order) * ('key,'value) tree;
+
+(* ------------------------------------------------------------------------- *)
+(* Map debugging functions. *)
+(* ------------------------------------------------------------------------- *)
+
+(*BasicDebug
+fun checkInvariants s m =
+ let
+ val Map (compareKey,tree) = m
+
+ val _ = treeCheckInvariants compareKey tree
+ in
+ m
+ end
+ handle Bug bug => raise Bug (s ^ "\n" ^ "Map.checkInvariants: " ^ bug);
+*)
+
+(* ------------------------------------------------------------------------- *)
+(* Constructors. *)
+(* ------------------------------------------------------------------------- *)
+
+fun new compareKey =
+ let
+ val tree = treeNew ()
+ in
+ Map (compareKey,tree)
+ end;
+
+fun singleton compareKey key_value =
+ let
+ val tree = treeSingleton key_value
+ in
+ Map (compareKey,tree)
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Map size. *)
+(* ------------------------------------------------------------------------- *)
+
+fun size (Map (_,tree)) = treeSize tree;
+
+fun null m = size m = 0;
+
+(* ------------------------------------------------------------------------- *)
+(* Querying. *)
+(* ------------------------------------------------------------------------- *)
+
+fun peekKey (Map (compareKey,tree)) key = treePeekKey compareKey key tree;
+
+fun peek (Map (compareKey,tree)) key = treePeek compareKey key tree;
+
+fun inDomain key m = Option.isSome (peek m key);
+
+fun get m key =
+ case peek m key of
+ NONE => raise Error "Map.get: element not found"
+ | SOME value => value;
+
+fun pick (Map (_,tree)) = treePick tree;
+
+fun nth (Map (_,tree)) n = treeNth n tree;
+
+fun random m =
+ let
+ val n = size m
+ in
+ if n = 0 then raise Bug "Map.random: empty"
+ else nth m (randomInt n)
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Adding. *)
+(* ------------------------------------------------------------------------- *)
+
+fun insert (Map (compareKey,tree)) key_value =
+ let
+ val tree = treeInsert compareKey key_value tree
+ in
+ Map (compareKey,tree)
+ end;
+
+(*BasicDebug
+val insert = fn m => fn kv =>
+ checkInvariants "Map.insert: result"
+ (insert (checkInvariants "Map.insert: input" m) kv);
+*)
+
+fun insertList m =
+ let
+ fun ins (key_value,acc) = insert acc key_value
+ in
+ List.foldl ins m
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Removing. *)
+(* ------------------------------------------------------------------------- *)
+
+fun delete (Map (compareKey,tree)) dkey =
+ let
+ val tree = treeDelete compareKey dkey tree
+ in
+ Map (compareKey,tree)
+ end;
+
+(*BasicDebug
+val delete = fn m => fn k =>
+ checkInvariants "Map.delete: result"
+ (delete (checkInvariants "Map.delete: input" m) k);
+*)
+
+fun remove m key = if inDomain key m then delete m key else m;
+
+fun deletePick (Map (compareKey,tree)) =
+ let
+ val (key_value,tree) = treeDeletePick tree
+ in
+ (key_value, Map (compareKey,tree))
+ end;
+
+(*BasicDebug
+val deletePick = fn m =>
+ let
+ val (kv,m) = deletePick (checkInvariants "Map.deletePick: input" m)
+ in
+ (kv, checkInvariants "Map.deletePick: result" m)
+ end;
+*)
+
+fun deleteNth (Map (compareKey,tree)) n =
+ let
+ val (key_value,tree) = treeDeleteNth n tree
+ in
+ (key_value, Map (compareKey,tree))
+ end;
+
+(*BasicDebug
+val deleteNth = fn m => fn n =>
+ let
+ val (kv,m) = deleteNth (checkInvariants "Map.deleteNth: input" m) n
+ in
+ (kv, checkInvariants "Map.deleteNth: result" m)
+ end;
+*)
+
+fun deleteRandom m =
+ let
+ val n = size m
+ in
+ if n = 0 then raise Bug "Map.deleteRandom: empty"
+ else deleteNth m (randomInt n)
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Joining (all join operations prefer keys in the second map). *)
+(* ------------------------------------------------------------------------- *)
+
+fun merge {first,second,both} (Map (compareKey,tree1)) (Map (_,tree2)) =
+ let
+ val tree = treeMerge compareKey first second both tree1 tree2
+ in
+ Map (compareKey,tree)
+ end;
+
+(*BasicDebug
+val merge = fn f => fn m1 => fn m2 =>
+ checkInvariants "Map.merge: result"
+ (merge f
+ (checkInvariants "Map.merge: input 1" m1)
+ (checkInvariants "Map.merge: input 2" m2));
+*)
+
+fun union f (Map (compareKey,tree1)) (Map (_,tree2)) =
+ let
+ fun f2 kv = f (kv,kv)
+
+ val tree = treeUnion compareKey f f2 tree1 tree2
+ in
+ Map (compareKey,tree)
+ end;
+
+(*BasicDebug
+val union = fn f => fn m1 => fn m2 =>
+ checkInvariants "Map.union: result"
+ (union f
+ (checkInvariants "Map.union: input 1" m1)
+ (checkInvariants "Map.union: input 2" m2));
+*)
+
+fun intersect f (Map (compareKey,tree1)) (Map (_,tree2)) =
+ let
+ val tree = treeIntersect compareKey f tree1 tree2
+ in
+ Map (compareKey,tree)
+ end;
+
+(*BasicDebug
+val intersect = fn f => fn m1 => fn m2 =>
+ checkInvariants "Map.intersect: result"
+ (intersect f
+ (checkInvariants "Map.intersect: input 1" m1)
+ (checkInvariants "Map.intersect: input 2" m2));
+*)
+
+(* ------------------------------------------------------------------------- *)
+(* Iterators over maps. *)
+(* ------------------------------------------------------------------------- *)
+
+fun mkIterator (Map (_,tree)) = treeMkIterator tree;
+
+fun mkRevIterator (Map (_,tree)) = treeMkRevIterator tree;
+
+(* ------------------------------------------------------------------------- *)
+(* Mapping and folding. *)
+(* ------------------------------------------------------------------------- *)
+
+fun mapPartial f (Map (compareKey,tree)) =
+ let
+ val tree = treeMapPartial f tree
+ in
+ Map (compareKey,tree)
+ end;
+
+(*BasicDebug
+val mapPartial = fn f => fn m =>
+ checkInvariants "Map.mapPartial: result"
+ (mapPartial f (checkInvariants "Map.mapPartial: input" m));
+*)
+
+fun map f (Map (compareKey,tree)) =
+ let
+ val tree = treeMap f tree
+ in
+ Map (compareKey,tree)
+ end;
+
+(*BasicDebug
+val map = fn f => fn m =>
+ checkInvariants "Map.map: result"
+ (map f (checkInvariants "Map.map: input" m));
+*)
+
+fun transform f = map (fn (_,value) => f value);
+
+fun filter pred =
+ let
+ fun f (key_value as (_,value)) =
+ if pred key_value then SOME value else NONE
+ in
+ mapPartial f
+ end;
+
+fun partition p =
+ let
+ fun np x = not (p x)
+ in
+ fn m => (filter p m, filter np m)
+ end;
+
+fun foldl f b m = foldIterator f b (mkIterator m);
+
+fun foldr f b m = foldIterator f b (mkRevIterator m);
+
+fun app f m = foldl (fn (key,value,()) => f (key,value)) () m;
+
+(* ------------------------------------------------------------------------- *)
+(* Searching. *)
+(* ------------------------------------------------------------------------- *)
+
+fun findl p m = findIterator p (mkIterator m);
+
+fun findr p m = findIterator p (mkRevIterator m);
+
+fun firstl f m = firstIterator f (mkIterator m);
+
+fun firstr f m = firstIterator f (mkRevIterator m);
+
+fun exists p m = Option.isSome (findl p m);
+
+fun all p =
+ let
+ fun np x = not (p x)
+ in
+ fn m => not (exists np m)
+ end;
+
+fun count pred =
+ let
+ fun f (k,v,acc) = if pred (k,v) then acc + 1 else acc
+ in
+ foldl f 0
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Comparing. *)
+(* ------------------------------------------------------------------------- *)
+
+fun compare compareValue (m1,m2) =
+ if pointerEqual (m1,m2) then EQUAL
+ else
+ case Int.compare (size m1, size m2) of
+ LESS => LESS
+ | EQUAL =>
+ let
+ val Map (compareKey,_) = m1
+
+ val io1 = mkIterator m1
+ and io2 = mkIterator m2
+ in
+ compareIterator compareKey compareValue io1 io2
+ end
+ | GREATER => GREATER;
+
+fun equal equalValue m1 m2 =
+ pointerEqual (m1,m2) orelse
+ (size m1 = size m2 andalso
+ let
+ val Map (compareKey,_) = m1
+
+ val io1 = mkIterator m1
+ and io2 = mkIterator m2
+ in
+ equalIterator (equalKey compareKey) equalValue io1 io2
+ end);
+
+(* ------------------------------------------------------------------------- *)
+(* Set operations on the domain. *)
+(* ------------------------------------------------------------------------- *)
+
+fun unionDomain (Map (compareKey,tree1)) (Map (_,tree2)) =
+ let
+ val tree = treeUnionDomain compareKey tree1 tree2
+ in
+ Map (compareKey,tree)
+ end;
+
+(*BasicDebug
+val unionDomain = fn m1 => fn m2 =>
+ checkInvariants "Map.unionDomain: result"
+ (unionDomain
+ (checkInvariants "Map.unionDomain: input 1" m1)
+ (checkInvariants "Map.unionDomain: input 2" m2));
+*)
+
+local
+ fun uncurriedUnionDomain (m,acc) = unionDomain acc m;
+in
+ fun unionListDomain ms =
+ case ms of
+ [] => raise Bug "Map.unionListDomain: no sets"
+ | m :: ms => List.foldl uncurriedUnionDomain m ms;
+end;
+
+fun intersectDomain (Map (compareKey,tree1)) (Map (_,tree2)) =
+ let
+ val tree = treeIntersectDomain compareKey tree1 tree2
+ in
+ Map (compareKey,tree)
+ end;
+
+(*BasicDebug
+val intersectDomain = fn m1 => fn m2 =>
+ checkInvariants "Map.intersectDomain: result"
+ (intersectDomain
+ (checkInvariants "Map.intersectDomain: input 1" m1)
+ (checkInvariants "Map.intersectDomain: input 2" m2));
+*)
+
+local
+ fun uncurriedIntersectDomain (m,acc) = intersectDomain acc m;
+in
+ fun intersectListDomain ms =
+ case ms of
+ [] => raise Bug "Map.intersectListDomain: no sets"
+ | m :: ms => List.foldl uncurriedIntersectDomain m ms;
+end;
+
+fun differenceDomain (Map (compareKey,tree1)) (Map (_,tree2)) =
+ let
+ val tree = treeDifferenceDomain compareKey tree1 tree2
+ in
+ Map (compareKey,tree)
+ end;
+
+(*BasicDebug
+val differenceDomain = fn m1 => fn m2 =>
+ checkInvariants "Map.differenceDomain: result"
+ (differenceDomain
+ (checkInvariants "Map.differenceDomain: input 1" m1)
+ (checkInvariants "Map.differenceDomain: input 2" m2));
+*)
+
+fun symmetricDifferenceDomain m1 m2 =
+ unionDomain (differenceDomain m1 m2) (differenceDomain m2 m1);
+
+fun equalDomain m1 m2 = equal (K (K true)) m1 m2;
+
+fun subsetDomain (Map (compareKey,tree1)) (Map (_,tree2)) =
+ treeSubsetDomain compareKey tree1 tree2;
+
+fun disjointDomain m1 m2 = null (intersectDomain m1 m2);
+
+(* ------------------------------------------------------------------------- *)
+(* Converting to and from lists. *)
+(* ------------------------------------------------------------------------- *)
+
+fun keys m = foldr (fn (key,_,l) => key :: l) [] m;
+
+fun values m = foldr (fn (_,value,l) => value :: l) [] m;
+
+fun toList m = foldr (fn (key,value,l) => (key,value) :: l) [] m;
+
+fun fromList compareKey l =
+ let
+ val m = new compareKey
+ in
+ insertList m l
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printing. *)
+(* ------------------------------------------------------------------------- *)
+
+fun toString m = "<" ^ (if null m then "" else Int.toString (size m)) ^ ">";
+
+end
+end;
+
+(**** Original file: KeyMap.sig ****)
+
+(* ========================================================================= *)
+(* FINITE MAPS WITH A FIXED KEY TYPE *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
+(* ========================================================================= *)
+
+signature KeyMap =
+sig
+
+(* ------------------------------------------------------------------------- *)
+(* A type of map keys. *)
+(* ------------------------------------------------------------------------- *)
+
+type key
+
+(* ------------------------------------------------------------------------- *)
+(* A type of finite maps. *)
+(* ------------------------------------------------------------------------- *)
+
+type 'a map
+
+(* ------------------------------------------------------------------------- *)
+(* Constructors. *)
+(* ------------------------------------------------------------------------- *)
+
+val new : unit -> 'a map
+
+val singleton : key * 'a -> 'a map
+
+(* ------------------------------------------------------------------------- *)
+(* Map size. *)
+(* ------------------------------------------------------------------------- *)
+
+val null : 'a map -> bool
+
+val size : 'a map -> int
+
+(* ------------------------------------------------------------------------- *)
+(* Querying. *)
+(* ------------------------------------------------------------------------- *)
+
+val peekKey : 'a map -> key -> (key * 'a) option
+
+val peek : 'a map -> key -> 'a option
+
+val get : 'a map -> key -> 'a (* raises Error *)
+
+val pick : 'a map -> key * 'a (* an arbitrary key/value pair *)
+
+val nth : 'a map -> int -> key * 'a (* in the range [0,size-1] *)
+
+val random : 'a map -> key * 'a
+
+(* ------------------------------------------------------------------------- *)
+(* Adding. *)
+(* ------------------------------------------------------------------------- *)
+
+val insert : 'a map -> key * 'a -> 'a map
+
+val insertList : 'a map -> (key * 'a) list -> 'a map
+
+(* ------------------------------------------------------------------------- *)
+(* Removing. *)
+(* ------------------------------------------------------------------------- *)
+
+val delete : 'a map -> key -> 'a map (* must be present *)
+
+val remove : 'a map -> key -> 'a map
+
+val deletePick : 'a map -> (key * 'a) * 'a map
+
+val deleteNth : 'a map -> int -> (key * 'a) * 'a map
+
+val deleteRandom : 'a map -> (key * 'a) * 'a map
+
+(* ------------------------------------------------------------------------- *)
+(* Joining (all join operations prefer keys in the second map). *)
+(* ------------------------------------------------------------------------- *)
+
+val merge :
+ {first : key * 'a -> 'c option,
+ second : key * 'b -> 'c option,
+ both : (key * 'a) * (key * 'b) -> 'c option} ->
+ 'a map -> 'b map -> 'c map
+
+val union :
+ ((key * 'a) * (key * 'a) -> 'a option) ->
+ 'a map -> 'a map -> 'a map
+
+val intersect :
+ ((key * 'a) * (key * 'b) -> 'c option) ->
+ 'a map -> 'b map -> 'c map
+
+(* ------------------------------------------------------------------------- *)
+(* Set operations on the domain. *)
+(* ------------------------------------------------------------------------- *)
+
+val inDomain : key -> 'a map -> bool
+
+val unionDomain : 'a map -> 'a map -> 'a map
+
+val unionListDomain : 'a map list -> 'a map
+
+val intersectDomain : 'a map -> 'a map -> 'a map
+
+val intersectListDomain : 'a map list -> 'a map
+
+val differenceDomain : 'a map -> 'a map -> 'a map
+
+val symmetricDifferenceDomain : 'a map -> 'a map -> 'a map
+
+val equalDomain : 'a map -> 'a map -> bool
+
+val subsetDomain : 'a map -> 'a map -> bool
+
+val disjointDomain : 'a map -> 'a map -> bool
+
+(* ------------------------------------------------------------------------- *)
+(* Mapping and folding. *)
+(* ------------------------------------------------------------------------- *)
+
+val mapPartial : (key * 'a -> 'b option) -> 'a map -> 'b map
+
+val map : (key * 'a -> 'b) -> 'a map -> 'b map
+
+val app : (key * 'a -> unit) -> 'a map -> unit
+
+val transform : ('a -> 'b) -> 'a map -> 'b map
+
+val filter : (key * 'a -> bool) -> 'a map -> 'a map
+
+val partition :
+ (key * 'a -> bool) -> 'a map -> 'a map * 'a map
+
+val foldl : (key * 'a * 's -> 's) -> 's -> 'a map -> 's
+
+val foldr : (key * 'a * 's -> 's) -> 's -> 'a map -> 's
+
+(* ------------------------------------------------------------------------- *)
+(* Searching. *)
+(* ------------------------------------------------------------------------- *)
+
+val findl : (key * 'a -> bool) -> 'a map -> (key * 'a) option
+
+val findr : (key * 'a -> bool) -> 'a map -> (key * 'a) option
+
+val firstl : (key * 'a -> 'b option) -> 'a map -> 'b option
+
+val firstr : (key * 'a -> 'b option) -> 'a map -> 'b option
+
+val exists : (key * 'a -> bool) -> 'a map -> bool
+
+val all : (key * 'a -> bool) -> 'a map -> bool
+
+val count : (key * 'a -> bool) -> 'a map -> int
+
+(* ------------------------------------------------------------------------- *)
+(* Comparing. *)
+(* ------------------------------------------------------------------------- *)
+
+val compare : ('a * 'a -> order) -> 'a map * 'a map -> order
+
+val equal : ('a -> 'a -> bool) -> 'a map -> 'a map -> bool
+
+(* ------------------------------------------------------------------------- *)
+(* Converting to and from lists. *)
+(* ------------------------------------------------------------------------- *)
+
+val keys : 'a map -> key list
+
+val values : 'a map -> 'a list
+
+val toList : 'a map -> (key * 'a) list
+
+val fromList : (key * 'a) list -> 'a map
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printing. *)
+(* ------------------------------------------------------------------------- *)
+
+val toString : 'a map -> string
+
+(* ------------------------------------------------------------------------- *)
+(* Iterators over maps. *)
+(* ------------------------------------------------------------------------- *)
+
+type 'a iterator
+
+val mkIterator : 'a map -> 'a iterator option
+
+val mkRevIterator : 'a map -> 'a iterator option
+
+val readIterator : 'a iterator -> key * 'a
+
+val advanceIterator : 'a iterator -> 'a iterator option
+
+end
+
+(**** Original file: KeyMap.sml ****)
+
+(* ========================================================================= *)
+(* FINITE MAPS WITH A FIXED KEY TYPE *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
+(* ========================================================================= *)
+
+functor KeyMap (Key : Ordered) :> KeyMap where type key = Key.t =
+struct
+
+(* ------------------------------------------------------------------------- *)
+(* Importing from the input signature. *)
+(* ------------------------------------------------------------------------- *)
+
+open Metis; (* MODIFIED by Jasmin Blanchette *)
+
+type key = Key.t;
+
+val compareKey = Key.compare;
+
+(* ------------------------------------------------------------------------- *)
+(* Importing useful functionality. *)
+(* ------------------------------------------------------------------------- *)
+
+exception Bug = Useful.Bug;
+
+exception Error = Useful.Error;
+
+val pointerEqual = Portable.pointerEqual;
+
+val K = Useful.K;
+
+val randomInt = Portable.randomInt;
+
+val randomWord = Portable.randomWord;
+
+(* ------------------------------------------------------------------------- *)
+(* Converting a comparison function to an equality function. *)
+(* ------------------------------------------------------------------------- *)
+
+fun equalKey key1 key2 = compareKey (key1,key2) = EQUAL;
+
+(* ------------------------------------------------------------------------- *)
+(* Priorities. *)
+(* ------------------------------------------------------------------------- *)
+
+type priority = Word.word;
+
+val randomPriority = randomWord;
+
+val comparePriority = Word.compare;
+
+(* ------------------------------------------------------------------------- *)
+(* Priority search trees. *)
+(* ------------------------------------------------------------------------- *)
+
+datatype 'value tree =
+ E
+ | T of 'value node
+
+and 'value node =
+ Node of
+ {size : int,
+ priority : priority,
+ left : 'value tree,
+ key : key,
+ value : 'value,
+ right : 'value tree};
+
+fun lowerPriorityNode node1 node2 =
+ let
+ val Node {priority = p1, ...} = node1
+ and Node {priority = p2, ...} = node2
+ in
+ comparePriority (p1,p2) = LESS
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Tree debugging functions. *)
+(* ------------------------------------------------------------------------- *)
+
+(*BasicDebug
+local
+ fun checkSizes tree =
+ case tree of
+ E => 0
+ | T (Node {size,left,right,...}) =>
+ let
+ val l = checkSizes left
+ and r = checkSizes right
+
+ val () = if l + 1 + r = size then () else raise Bug "wrong size"
+ in
+ size
+ end;
+
+ fun checkSorted x tree =
+ case tree of
+ E => x
+ | T (Node {left,key,right,...}) =>
+ let
+ val x = checkSorted x left
+
+ val () =
+ case x of
+ NONE => ()
+ | SOME k =>
+ case compareKey (k,key) of
+ LESS => ()
+ | EQUAL => raise Bug "duplicate keys"
+ | GREATER => raise Bug "unsorted"
+
+ val x = SOME key
+ in
+ checkSorted x right
+ end;
+
+ fun checkPriorities tree =
+ case tree of
+ E => NONE
+ | T node =>
+ let
+ val Node {left,right,...} = node
+
+ val () =
+ case checkPriorities left of
+ NONE => ()
+ | SOME lnode =>
+ if not (lowerPriorityNode node lnode) then ()
+ else raise Bug "left child has greater priority"
+
+ val () =
+ case checkPriorities right of
+ NONE => ()
+ | SOME rnode =>
+ if not (lowerPriorityNode node rnode) then ()
+ else raise Bug "right child has greater priority"
+ in
+ SOME node
+ end;
+in
+ fun treeCheckInvariants tree =
+ let
+ val _ = checkSizes tree
+
+ val _ = checkSorted NONE tree
+
+ val _ = checkPriorities tree
+ in
+ tree
+ end
+ handle Error err => raise Bug err;
+end;
+*)
+
+(* ------------------------------------------------------------------------- *)
+(* Tree operations. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeNew () = E;
+
+fun nodeSize (Node {size = x, ...}) = x;
+
+fun treeSize tree =
+ case tree of
+ E => 0
+ | T x => nodeSize x;
+
+fun mkNode priority left key value right =
+ let
+ val size = treeSize left + 1 + treeSize right
+ in
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ end;
+
+fun mkTree priority left key value right =
+ let
+ val node = mkNode priority left key value right
+ in
+ T node
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Extracting the left and right spines of a tree. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeLeftSpine acc tree =
+ case tree of
+ E => acc
+ | T node => nodeLeftSpine acc node
+
+and nodeLeftSpine acc node =
+ let
+ val Node {left,...} = node
+ in
+ treeLeftSpine (node :: acc) left
+ end;
+
+fun treeRightSpine acc tree =
+ case tree of
+ E => acc
+ | T node => nodeRightSpine acc node
+
+and nodeRightSpine acc node =
+ let
+ val Node {right,...} = node
+ in
+ treeRightSpine (node :: acc) right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Singleton trees. *)
+(* ------------------------------------------------------------------------- *)
+
+fun mkNodeSingleton priority key value =
+ let
+ val size = 1
+ and left = E
+ and right = E
+ in
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ end;
+
+fun nodeSingleton (key,value) =
+ let
+ val priority = randomPriority ()
+ in
+ mkNodeSingleton priority key value
+ end;
+
+fun treeSingleton key_value =
+ let
+ val node = nodeSingleton key_value
+ in
+ T node
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Appending two trees, where every element of the first tree is less than *)
+(* every element of the second tree. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeAppend tree1 tree2 =
+ case tree1 of
+ E => tree2
+ | T node1 =>
+ case tree2 of
+ E => tree1
+ | T node2 =>
+ if lowerPriorityNode node1 node2 then
+ let
+ val Node {priority,left,key,value,right,...} = node2
+
+ val left = treeAppend tree1 left
+ in
+ mkTree priority left key value right
+ end
+ else
+ let
+ val Node {priority,left,key,value,right,...} = node1
+
+ val right = treeAppend right tree2
+ in
+ mkTree priority left key value right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Appending two trees and a node, where every element of the first tree is *)
+(* less than the node, which in turn is less than every element of the *)
+(* second tree. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeCombine left node right =
+ let
+ val left_node = treeAppend left (T node)
+ in
+ treeAppend left_node right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Searching a tree for a value. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treePeek pkey tree =
+ case tree of
+ E => NONE
+ | T node => nodePeek pkey node
+
+and nodePeek pkey node =
+ let
+ val Node {left,key,value,right,...} = node
+ in
+ case compareKey (pkey,key) of
+ LESS => treePeek pkey left
+ | EQUAL => SOME value
+ | GREATER => treePeek pkey right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Tree paths. *)
+(* ------------------------------------------------------------------------- *)
+
+(* Generating a path by searching a tree for a key/value pair *)
+
+fun treePeekPath pkey path tree =
+ case tree of
+ E => (path,NONE)
+ | T node => nodePeekPath pkey path node
+
+and nodePeekPath pkey path node =
+ let
+ val Node {left,key,right,...} = node
+ in
+ case compareKey (pkey,key) of
+ LESS => treePeekPath pkey ((true,node) :: path) left
+ | EQUAL => (path, SOME node)
+ | GREATER => treePeekPath pkey ((false,node) :: path) right
+ end;
+
+(* A path splits a tree into left/right components *)
+
+fun addSidePath ((wentLeft,node),(leftTree,rightTree)) =
+ let
+ val Node {priority,left,key,value,right,...} = node
+ in
+ if wentLeft then (leftTree, mkTree priority rightTree key value right)
+ else (mkTree priority left key value leftTree, rightTree)
+ end;
+
+fun addSidesPath left_right = List.foldl addSidePath left_right;
+
+fun mkSidesPath path = addSidesPath (E,E) path;
+
+(* Updating the subtree at a path *)
+
+local
+ fun updateTree ((wentLeft,node),tree) =
+ let
+ val Node {priority,left,key,value,right,...} = node
+ in
+ if wentLeft then mkTree priority tree key value right
+ else mkTree priority left key value tree
+ end;
+in
+ fun updateTreePath tree = List.foldl updateTree tree;
+end;
+
+(* Inserting a new node at a path position *)
+
+fun insertNodePath node =
+ let
+ fun insert left_right path =
+ case path of
+ [] =>
+ let
+ val (left,right) = left_right
+ in
+ treeCombine left node right
+ end
+ | (step as (_,snode)) :: rest =>
+ if lowerPriorityNode snode node then
+ let
+ val left_right = addSidePath (step,left_right)
+ in
+ insert left_right rest
+ end
+ else
+ let
+ val (left,right) = left_right
+
+ val tree = treeCombine left node right
+ in
+ updateTreePath tree path
+ end
+ in
+ insert (E,E)
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Using a key to split a node into three components: the keys comparing *)
+(* less than the supplied key, an optional equal key, and the keys comparing *)
+(* greater. *)
+(* ------------------------------------------------------------------------- *)
+
+fun nodePartition pkey node =
+ let
+ val (path,pnode) = nodePeekPath pkey [] node
+ in
+ case pnode of
+ NONE =>
+ let
+ val (left,right) = mkSidesPath path
+ in
+ (left,NONE,right)
+ end
+ | SOME node =>
+ let
+ val Node {left,key,value,right,...} = node
+
+ val (left,right) = addSidesPath (left,right) path
+ in
+ (left, SOME (key,value), right)
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Searching a tree for a key/value pair. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treePeekKey pkey tree =
+ case tree of
+ E => NONE
+ | T node => nodePeekKey pkey node
+
+and nodePeekKey pkey node =
+ let
+ val Node {left,key,value,right,...} = node
+ in
+ case compareKey (pkey,key) of
+ LESS => treePeekKey pkey left
+ | EQUAL => SOME (key,value)
+ | GREATER => treePeekKey pkey right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Inserting new key/values into the tree. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeInsert key_value tree =
+ let
+ val (key,value) = key_value
+
+ val (path,inode) = treePeekPath key [] tree
+ in
+ case inode of
+ NONE =>
+ let
+ val node = nodeSingleton (key,value)
+ in
+ insertNodePath node path
+ end
+ | SOME node =>
+ let
+ val Node {size,priority,left,right,...} = node
+
+ val node =
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ in
+ updateTreePath (T node) path
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Deleting key/value pairs: it raises an exception if the supplied key is *)
+(* not present. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeDelete dkey tree =
+ case tree of
+ E => raise Bug "KeyMap.delete: element not found"
+ | T node => nodeDelete dkey node
+
+and nodeDelete dkey node =
+ let
+ val Node {size,priority,left,key,value,right} = node
+ in
+ case compareKey (dkey,key) of
+ LESS =>
+ let
+ val size = size - 1
+ and left = treeDelete dkey left
+
+ val node =
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ in
+ T node
+ end
+ | EQUAL => treeAppend left right
+ | GREATER =>
+ let
+ val size = size - 1
+ and right = treeDelete dkey right
+
+ val node =
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ in
+ T node
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Partial map is the basic operation for preserving tree structure. *)
+(* It applies its argument function to the elements *in order*. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeMapPartial f tree =
+ case tree of
+ E => E
+ | T node => nodeMapPartial f node
+
+and nodeMapPartial f (Node {priority,left,key,value,right,...}) =
+ let
+ val left = treeMapPartial f left
+ and vo = f (key,value)
+ and right = treeMapPartial f right
+ in
+ case vo of
+ NONE => treeAppend left right
+ | SOME value => mkTree priority left key value right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Mapping tree values. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeMap f tree =
+ case tree of
+ E => E
+ | T node => T (nodeMap f node)
+
+and nodeMap f node =
+ let
+ val Node {size,priority,left,key,value,right} = node
+
+ val left = treeMap f left
+ and value = f (key,value)
+ and right = treeMap f right
+ in
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Merge is the basic operation for joining two trees. Note that the merged *)
+(* key is always the one from the second map. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeMerge f1 f2 fb tree1 tree2 =
+ case tree1 of
+ E => treeMapPartial f2 tree2
+ | T node1 =>
+ case tree2 of
+ E => treeMapPartial f1 tree1
+ | T node2 => nodeMerge f1 f2 fb node1 node2
+
+and nodeMerge f1 f2 fb node1 node2 =
+ let
+ val Node {priority,left,key,value,right,...} = node2
+
+ val (l,kvo,r) = nodePartition key node1
+
+ val left = treeMerge f1 f2 fb l left
+ and right = treeMerge f1 f2 fb r right
+
+ val vo =
+ case kvo of
+ NONE => f2 (key,value)
+ | SOME kv => fb (kv,(key,value))
+ in
+ case vo of
+ NONE => treeAppend left right
+ | SOME value =>
+ let
+ val node = mkNodeSingleton priority key value
+ in
+ treeCombine left node right
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* A op union operation on trees. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeUnion f f2 tree1 tree2 =
+ case tree1 of
+ E => tree2
+ | T node1 =>
+ case tree2 of
+ E => tree1
+ | T node2 => nodeUnion f f2 node1 node2
+
+and nodeUnion f f2 node1 node2 =
+ if pointerEqual (node1,node2) then nodeMapPartial f2 node1
+ else
+ let
+ val Node {priority,left,key,value,right,...} = node2
+
+ val (l,kvo,r) = nodePartition key node1
+
+ val left = treeUnion f f2 l left
+ and right = treeUnion f f2 r right
+
+ val vo =
+ case kvo of
+ NONE => SOME value
+ | SOME kv => f (kv,(key,value))
+ in
+ case vo of
+ NONE => treeAppend left right
+ | SOME value =>
+ let
+ val node = mkNodeSingleton priority key value
+ in
+ treeCombine left node right
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* An intersect operation on trees. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeIntersect f t1 t2 =
+ case t1 of
+ E => E
+ | T n1 =>
+ case t2 of
+ E => E
+ | T n2 => nodeIntersect f n1 n2
+
+and nodeIntersect f n1 n2 =
+ let
+ val Node {priority,left,key,value,right,...} = n2
+
+ val (l,kvo,r) = nodePartition key n1
+
+ val left = treeIntersect f l left
+ and right = treeIntersect f r right
+
+ val vo =
+ case kvo of
+ NONE => NONE
+ | SOME kv => f (kv,(key,value))
+ in
+ case vo of
+ NONE => treeAppend left right
+ | SOME value => mkTree priority left key value right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* A op union operation on trees which simply chooses the second value. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeUnionDomain tree1 tree2 =
+ case tree1 of
+ E => tree2
+ | T node1 =>
+ case tree2 of
+ E => tree1
+ | T node2 =>
+ if pointerEqual (node1,node2) then tree2
+ else nodeUnionDomain node1 node2
+
+and nodeUnionDomain node1 node2 =
+ let
+ val Node {priority,left,key,value,right,...} = node2
+
+ val (l,_,r) = nodePartition key node1
+
+ val left = treeUnionDomain l left
+ and right = treeUnionDomain r right
+
+ val node = mkNodeSingleton priority key value
+ in
+ treeCombine left node right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* An intersect operation on trees which simply chooses the second value. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeIntersectDomain tree1 tree2 =
+ case tree1 of
+ E => E
+ | T node1 =>
+ case tree2 of
+ E => E
+ | T node2 =>
+ if pointerEqual (node1,node2) then tree2
+ else nodeIntersectDomain node1 node2
+
+and nodeIntersectDomain node1 node2 =
+ let
+ val Node {priority,left,key,value,right,...} = node2
+
+ val (l,kvo,r) = nodePartition key node1
+
+ val left = treeIntersectDomain l left
+ and right = treeIntersectDomain r right
+ in
+ if Option.isSome kvo then mkTree priority left key value right
+ else treeAppend left right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* A difference operation on trees. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeDifferenceDomain t1 t2 =
+ case t1 of
+ E => E
+ | T n1 =>
+ case t2 of
+ E => t1
+ | T n2 => nodeDifferenceDomain n1 n2
+
+and nodeDifferenceDomain n1 n2 =
+ if pointerEqual (n1,n2) then E
+ else
+ let
+ val Node {priority,left,key,value,right,...} = n1
+
+ val (l,kvo,r) = nodePartition key n2
+
+ val left = treeDifferenceDomain left l
+ and right = treeDifferenceDomain right r
+ in
+ if Option.isSome kvo then treeAppend left right
+ else mkTree priority left key value right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* A op subset operation on trees. *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeSubsetDomain tree1 tree2 =
+ case tree1 of
+ E => true
+ | T node1 =>
+ case tree2 of
+ E => false
+ | T node2 => nodeSubsetDomain node1 node2
+
+and nodeSubsetDomain node1 node2 =
+ pointerEqual (node1,node2) orelse
+ let
+ val Node {size,left,key,right,...} = node1
+ in
+ size <= nodeSize node2 andalso
+ let
+ val (l,kvo,r) = nodePartition key node2
+ in
+ Option.isSome kvo andalso
+ treeSubsetDomain left l andalso
+ treeSubsetDomain right r
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Picking an arbitrary key/value pair from a tree. *)
+(* ------------------------------------------------------------------------- *)
+
+fun nodePick node =
+ let
+ val Node {key,value,...} = node
+ in
+ (key,value)
+ end;
+
+fun treePick tree =
+ case tree of
+ E => raise Bug "KeyMap.treePick"
+ | T node => nodePick node;
+
+(* ------------------------------------------------------------------------- *)
+(* Removing an arbitrary key/value pair from a tree. *)
+(* ------------------------------------------------------------------------- *)
+
+fun nodeDeletePick node =
+ let
+ val Node {left,key,value,right,...} = node
+ in
+ ((key,value), treeAppend left right)
+ end;
+
+fun treeDeletePick tree =
+ case tree of
+ E => raise Bug "KeyMap.treeDeletePick"
+ | T node => nodeDeletePick node;
+
+(* ------------------------------------------------------------------------- *)
+(* Finding the nth smallest key/value (counting from 0). *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeNth n tree =
+ case tree of
+ E => raise Bug "KeyMap.treeNth"
+ | T node => nodeNth n node
+
+and nodeNth n node =
+ let
+ val Node {left,key,value,right,...} = node
+
+ val k = treeSize left
+ in
+ if n = k then (key,value)
+ else if n < k then treeNth n left
+ else treeNth (n - (k + 1)) right
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Removing the nth smallest key/value (counting from 0). *)
+(* ------------------------------------------------------------------------- *)
+
+fun treeDeleteNth n tree =
+ case tree of
+ E => raise Bug "KeyMap.treeDeleteNth"
+ | T node => nodeDeleteNth n node
+
+and nodeDeleteNth n node =
+ let
+ val Node {size,priority,left,key,value,right} = node
+
+ val k = treeSize left
+ in
+ if n = k then ((key,value), treeAppend left right)
+ else if n < k then
+ let
+ val (key_value,left) = treeDeleteNth n left
+
+ val size = size - 1
+
+ val node =
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ in
+ (key_value, T node)
+ end
+ else
+ let
+ val n = n - (k + 1)
+
+ val (key_value,right) = treeDeleteNth n right
+
+ val size = size - 1
+
+ val node =
+ Node
+ {size = size,
+ priority = priority,
+ left = left,
+ key = key,
+ value = value,
+ right = right}
+ in
+ (key_value, T node)
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Iterators. *)
+(* ------------------------------------------------------------------------- *)
+
+datatype 'value iterator =
+ LR of (key * 'value) * 'value tree * 'value node list
+ | RL of (key * 'value) * 'value tree * 'value node list;
+
+fun fromSpineLR nodes =
+ case nodes of
+ [] => NONE
+ | Node {key,value,right,...} :: nodes =>
+ SOME (LR ((key,value),right,nodes));
+
+fun fromSpineRL nodes =
+ case nodes of
+ [] => NONE
+ | Node {key,value,left,...} :: nodes =>
+ SOME (RL ((key,value),left,nodes));
+
+fun addLR nodes tree = fromSpineLR (treeLeftSpine nodes tree);
+
+fun addRL nodes tree = fromSpineRL (treeRightSpine nodes tree);
+
+fun treeMkIterator tree = addLR [] tree;
+
+fun treeMkRevIterator tree = addRL [] tree;
+
+fun readIterator iter =
+ case iter of
+ LR (key_value,_,_) => key_value
+ | RL (key_value,_,_) => key_value;
+
+fun advanceIterator iter =
+ case iter of
+ LR (_,tree,nodes) => addLR nodes tree
+ | RL (_,tree,nodes) => addRL nodes tree;
+
+fun foldIterator f acc io =
+ case io of
+ NONE => acc
+ | SOME iter =>
+ let
+ val (key,value) = readIterator iter
+ in
+ foldIterator f (f (key,value,acc)) (advanceIterator iter)
+ end;
+
+fun findIterator pred io =
+ case io of
+ NONE => NONE
+ | SOME iter =>
+ let
+ val key_value = readIterator iter
+ in
+ if pred key_value then SOME key_value
+ else findIterator pred (advanceIterator iter)
+ end;
+
+fun firstIterator f io =
+ case io of
+ NONE => NONE
+ | SOME iter =>
+ let
+ val key_value = readIterator iter
+ in
+ case f key_value of
+ NONE => firstIterator f (advanceIterator iter)
+ | s => s
+ end;
+
+fun compareIterator compareValue io1 io2 =
+ case (io1,io2) of
+ (NONE,NONE) => EQUAL
+ | (NONE, SOME _) => LESS
+ | (SOME _, NONE) => GREATER
+ | (SOME i1, SOME i2) =>
+ let
+ val (k1,v1) = readIterator i1
+ and (k2,v2) = readIterator i2
+ in
+ case compareKey (k1,k2) of
+ LESS => LESS
+ | EQUAL =>
+ (case compareValue (v1,v2) of
+ LESS => LESS
+ | EQUAL =>
+ let
+ val io1 = advanceIterator i1
+ and io2 = advanceIterator i2
+ in
+ compareIterator compareValue io1 io2
+ end
+ | GREATER => GREATER)
+ | GREATER => GREATER
+ end;
+
+fun equalIterator equalValue io1 io2 =
+ case (io1,io2) of
+ (NONE,NONE) => true
+ | (NONE, SOME _) => false
+ | (SOME _, NONE) => false
+ | (SOME i1, SOME i2) =>
+ let
+ val (k1,v1) = readIterator i1
+ and (k2,v2) = readIterator i2
+ in
+ equalKey k1 k2 andalso
+ equalValue v1 v2 andalso
+ let
+ val io1 = advanceIterator i1
+ and io2 = advanceIterator i2
+ in
+ equalIterator equalValue io1 io2
+ end
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* A type of finite maps. *)
+(* ------------------------------------------------------------------------- *)
+
+datatype 'value map =
+ Map of 'value tree;
+
+(* ------------------------------------------------------------------------- *)
+(* Map debugging functions. *)
+(* ------------------------------------------------------------------------- *)
+
+(*BasicDebug
+fun checkInvariants s m =
+ let
+ val Map tree = m
+
+ val _ = treeCheckInvariants tree
+ in
+ m
+ end
+ handle Bug bug => raise Bug (s ^ "\n" ^ "KeyMap.checkInvariants: " ^ bug);
+*)
+
+(* ------------------------------------------------------------------------- *)
+(* Constructors. *)
+(* ------------------------------------------------------------------------- *)
+
+fun new () =
+ let
+ val tree = treeNew ()
+ in
+ Map tree
+ end;
+
+fun singleton key_value =
+ let
+ val tree = treeSingleton key_value
+ in
+ Map tree
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Map size. *)
+(* ------------------------------------------------------------------------- *)
+
+fun size (Map tree) = treeSize tree;
+
+fun null m = size m = 0;
+
+(* ------------------------------------------------------------------------- *)
+(* Querying. *)
+(* ------------------------------------------------------------------------- *)
+
+fun peekKey (Map tree) key = treePeekKey key tree;
+
+fun peek (Map tree) key = treePeek key tree;
+
+fun inDomain key m = Option.isSome (peek m key);
+
+fun get m key =
+ case peek m key of
+ NONE => raise Error "KeyMap.get: element not found"
+ | SOME value => value;
+
+fun pick (Map tree) = treePick tree;
+
+fun nth (Map tree) n = treeNth n tree;
+
+fun random m =
+ let
+ val n = size m
+ in
+ if n = 0 then raise Bug "KeyMap.random: empty"
+ else nth m (randomInt n)
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Adding. *)
+(* ------------------------------------------------------------------------- *)
+
+fun insert (Map tree) key_value =
+ let
+ val tree = treeInsert key_value tree
+ in
+ Map tree
+ end;
+
+(*BasicDebug
+val insert = fn m => fn kv =>
+ checkInvariants "KeyMap.insert: result"
+ (insert (checkInvariants "KeyMap.insert: input" m) kv);
+*)
+
+fun insertList m =
+ let
+ fun ins (key_value,acc) = insert acc key_value
+ in
+ List.foldl ins m
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Removing. *)
+(* ------------------------------------------------------------------------- *)
+
+fun delete (Map tree) dkey =
+ let
+ val tree = treeDelete dkey tree
+ in
+ Map tree
+ end;
+
+(*BasicDebug
+val delete = fn m => fn k =>
+ checkInvariants "KeyMap.delete: result"
+ (delete (checkInvariants "KeyMap.delete: input" m) k);
+*)
+
+fun remove m key = if inDomain key m then delete m key else m;
+
+fun deletePick (Map tree) =
+ let
+ val (key_value,tree) = treeDeletePick tree
+ in
+ (key_value, Map tree)
+ end;
+
+(*BasicDebug
+val deletePick = fn m =>
+ let
+ val (kv,m) = deletePick (checkInvariants "KeyMap.deletePick: input" m)
+ in
+ (kv, checkInvariants "KeyMap.deletePick: result" m)
+ end;
+*)
+
+fun deleteNth (Map tree) n =
+ let
+ val (key_value,tree) = treeDeleteNth n tree
+ in
+ (key_value, Map tree)
+ end;
+
+(*BasicDebug
+val deleteNth = fn m => fn n =>
+ let
+ val (kv,m) = deleteNth (checkInvariants "KeyMap.deleteNth: input" m) n
+ in
+ (kv, checkInvariants "KeyMap.deleteNth: result" m)
+ end;
+*)
+
+fun deleteRandom m =
+ let
+ val n = size m
+ in
+ if n = 0 then raise Bug "KeyMap.deleteRandom: empty"
+ else deleteNth m (randomInt n)
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Joining (all join operations prefer keys in the second map). *)
+(* ------------------------------------------------------------------------- *)
+
+fun merge {first,second,both} (Map tree1) (Map tree2) =
+ let
+ val tree = treeMerge first second both tree1 tree2
+ in
+ Map tree
+ end;
+
+(*BasicDebug
+val merge = fn f => fn m1 => fn m2 =>
+ checkInvariants "KeyMap.merge: result"
+ (merge f
+ (checkInvariants "KeyMap.merge: input 1" m1)
+ (checkInvariants "KeyMap.merge: input 2" m2));
+*)
+
+fun op union f (Map tree1) (Map tree2) =
+ let
+ fun f2 kv = f (kv,kv)
+
+ val tree = treeUnion f f2 tree1 tree2
+ in
+ Map tree
+ end;
+
+(*BasicDebug
+val op union = fn f => fn m1 => fn m2 =>
+ checkInvariants "KeyMap.union: result"
+ (union f
+ (checkInvariants "KeyMap.union: input 1" m1)
+ (checkInvariants "KeyMap.union: input 2" m2));
+*)
+
+fun intersect f (Map tree1) (Map tree2) =
+ let
+ val tree = treeIntersect f tree1 tree2
+ in
+ Map tree
+ end;
+
+(*BasicDebug
+val intersect = fn f => fn m1 => fn m2 =>
+ checkInvariants "KeyMap.intersect: result"
+ (intersect f
+ (checkInvariants "KeyMap.intersect: input 1" m1)
+ (checkInvariants "KeyMap.intersect: input 2" m2));
+*)
+
+(* ------------------------------------------------------------------------- *)
+(* Iterators over maps. *)
+(* ------------------------------------------------------------------------- *)
+
+fun mkIterator (Map tree) = treeMkIterator tree;
+
+fun mkRevIterator (Map tree) = treeMkRevIterator tree;
+
+(* ------------------------------------------------------------------------- *)
+(* Mapping and folding. *)
+(* ------------------------------------------------------------------------- *)
+
+fun mapPartial f (Map tree) =
+ let
+ val tree = treeMapPartial f tree
+ in
+ Map tree
+ end;
+
+(*BasicDebug
+val mapPartial = fn f => fn m =>
+ checkInvariants "KeyMap.mapPartial: result"
+ (mapPartial f (checkInvariants "KeyMap.mapPartial: input" m));
+*)
+
+fun map f (Map tree) =
+ let
+ val tree = treeMap f tree
+ in
+ Map tree
+ end;
+
+(*BasicDebug
+val map = fn f => fn m =>
+ checkInvariants "KeyMap.map: result"
+ (map f (checkInvariants "KeyMap.map: input" m));
+*)
+
+fun transform f = map (fn (_,value) => f value);
+
+fun filter pred =
+ let
+ fun f (key_value as (_,value)) =
+ if pred key_value then SOME value else NONE
+ in
+ mapPartial f
+ end;
+
+fun partition p =
+ let
+ fun np x = not (p x)
+ in
+ fn m => (filter p m, filter np m)
+ end;
+
+fun foldl f b m = foldIterator f b (mkIterator m);
+
+fun foldr f b m = foldIterator f b (mkRevIterator m);
+
+fun app f m = foldl (fn (key,value,()) => f (key,value)) () m;
+
+(* ------------------------------------------------------------------------- *)
+(* Searching. *)
+(* ------------------------------------------------------------------------- *)
+
+fun findl p m = findIterator p (mkIterator m);
+
+fun findr p m = findIterator p (mkRevIterator m);
+
+fun firstl f m = firstIterator f (mkIterator m);
+
+fun firstr f m = firstIterator f (mkRevIterator m);
+
+fun exists p m = Option.isSome (findl p m);
+
+fun all p =
+ let
+ fun np x = not (p x)
+ in
+ fn m => not (exists np m)
+ end;
+
+fun count pred =
+ let
+ fun f (k,v,acc) = if pred (k,v) then acc + 1 else acc
+ in
+ foldl f 0
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Comparing. *)
+(* ------------------------------------------------------------------------- *)
+
+fun compare compareValue (m1,m2) =
+ if pointerEqual (m1,m2) then EQUAL
+ else
+ case Int.compare (size m1, size m2) of
+ LESS => LESS
+ | EQUAL =>
+ let
+ val Map _ = m1
+
+ val io1 = mkIterator m1
+ and io2 = mkIterator m2
+ in
+ compareIterator compareValue io1 io2
+ end
+ | GREATER => GREATER;
+
+fun equal equalValue m1 m2 =
+ pointerEqual (m1,m2) orelse
+ (size m1 = size m2 andalso
+ let
+ val Map _ = m1
+
+ val io1 = mkIterator m1
+ and io2 = mkIterator m2
+ in
+ equalIterator equalValue io1 io2
+ end);
+
+(* ------------------------------------------------------------------------- *)
+(* Set operations on the domain. *)
+(* ------------------------------------------------------------------------- *)
+
+fun unionDomain (Map tree1) (Map tree2) =
+ let
+ val tree = treeUnionDomain tree1 tree2
+ in
+ Map tree
+ end;
+
+(*BasicDebug
+val unionDomain = fn m1 => fn m2 =>
+ checkInvariants "KeyMap.unionDomain: result"
+ (unionDomain
+ (checkInvariants "KeyMap.unionDomain: input 1" m1)
+ (checkInvariants "KeyMap.unionDomain: input 2" m2));
+*)
+
+local
+ fun uncurriedUnionDomain (m,acc) = unionDomain acc m;
+in
+ fun unionListDomain ms =
+ case ms of
+ [] => raise Bug "KeyMap.unionListDomain: no sets"
+ | m :: ms => List.foldl uncurriedUnionDomain m ms;
+end;
+
+fun intersectDomain (Map tree1) (Map tree2) =
+ let
+ val tree = treeIntersectDomain tree1 tree2
+ in
+ Map tree
+ end;
+
+(*BasicDebug
+val intersectDomain = fn m1 => fn m2 =>
+ checkInvariants "KeyMap.intersectDomain: result"
+ (intersectDomain
+ (checkInvariants "KeyMap.intersectDomain: input 1" m1)
+ (checkInvariants "KeyMap.intersectDomain: input 2" m2));
+*)
+
+local
+ fun uncurriedIntersectDomain (m,acc) = intersectDomain acc m;
+in
+ fun intersectListDomain ms =
+ case ms of
+ [] => raise Bug "KeyMap.intersectListDomain: no sets"
+ | m :: ms => List.foldl uncurriedIntersectDomain m ms;
+end;
+
+fun differenceDomain (Map tree1) (Map tree2) =
+ let
+ val tree = treeDifferenceDomain tree1 tree2
+ in
+ Map tree
+ end;
+
+(*BasicDebug
+val differenceDomain = fn m1 => fn m2 =>
+ checkInvariants "KeyMap.differenceDomain: result"
+ (differenceDomain
+ (checkInvariants "KeyMap.differenceDomain: input 1" m1)
+ (checkInvariants "KeyMap.differenceDomain: input 2" m2));
+*)
+
+fun symmetricDifferenceDomain m1 m2 =
+ unionDomain (differenceDomain m1 m2) (differenceDomain m2 m1);
+
+fun equalDomain m1 m2 = equal (K (K true)) m1 m2;
+
+fun subsetDomain (Map tree1) (Map tree2) =
+ treeSubsetDomain tree1 tree2;
+
+fun disjointDomain m1 m2 = null (intersectDomain m1 m2);
+
+(* ------------------------------------------------------------------------- *)
+(* Converting to and from lists. *)
+(* ------------------------------------------------------------------------- *)
+
+fun keys m = foldr (fn (key,_,l) => key :: l) [] m;
+
+fun values m = foldr (fn (_,value,l) => value :: l) [] m;
+
+fun toList m = foldr (fn (key,value,l) => (key,value) :: l) [] m;
+
+fun fromList l =
+ let
+ val m = new ()
+ in
+ insertList m l
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printing. *)
+(* ------------------------------------------------------------------------- *)
+
+fun toString m = "<" ^ (if null m then "" else Int.toString (size m)) ^ ">";
+
+end
+
+structure IntMap =
+KeyMap (Metis.IntOrdered); (* MODIFIED by Jasmin Blanchette *)
+
+structure IntPairMap =
+KeyMap (Metis.IntPairOrdered); (* MODIFIED by Jasmin Blanchette *)
+
+structure StringMap =
+KeyMap (Metis.StringOrdered); (* MODIFIED by Jasmin Blanchette *)
+
(**** Original file: Set.sig ****)
(* ========================================================================= *)
-(* FINITE SETS *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
+(* FINITE SETS IMPLEMENTED WITH RANDOMLY BALANCED TREES *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Set =
sig
(* ------------------------------------------------------------------------- *)
-(* Finite sets *)
+(* A type of finite sets. *)
(* ------------------------------------------------------------------------- *)
type 'elt set
-val comparison : 'elt set -> ('elt * 'elt -> order)
+(* ------------------------------------------------------------------------- *)
+(* Constructors. *)
+(* ------------------------------------------------------------------------- *)
val empty : ('elt * 'elt -> order) -> 'elt set
val singleton : ('elt * 'elt -> order) -> 'elt -> 'elt set
+(* ------------------------------------------------------------------------- *)
+(* Set size. *)
+(* ------------------------------------------------------------------------- *)
+
val null : 'elt set -> bool
val size : 'elt set -> int
+(* ------------------------------------------------------------------------- *)
+(* Querying. *)
+(* ------------------------------------------------------------------------- *)
+
+val peek : 'elt set -> 'elt -> 'elt option
+
val member : 'elt -> 'elt set -> bool
+val pick : 'elt set -> 'elt (* an arbitrary element *)
+
+val nth : 'elt set -> int -> 'elt (* in the range [0,size-1] *)
+
+val random : 'elt set -> 'elt
+
+(* ------------------------------------------------------------------------- *)
+(* Adding. *)
+(* ------------------------------------------------------------------------- *)
+
val add : 'elt set -> 'elt -> 'elt set
val addList : 'elt set -> 'elt list -> 'elt set
-val delete : 'elt set -> 'elt -> 'elt set (* raises Error *)
-
-(* Union and intersect prefer elements in the second set *)
+(* ------------------------------------------------------------------------- *)
+(* Removing. *)
+(* ------------------------------------------------------------------------- *)
+
+val delete : 'elt set -> 'elt -> 'elt set (* must be present *)
+
+val remove : 'elt set -> 'elt -> 'elt set
+
+val deletePick : 'elt set -> 'elt * 'elt set
+
+val deleteNth : 'elt set -> int -> 'elt * 'elt set
+
+val deleteRandom : 'elt set -> 'elt * 'elt set
+
+(* ------------------------------------------------------------------------- *)
+(* Joining. *)
+(* ------------------------------------------------------------------------- *)
val union : 'elt set -> 'elt set -> 'elt set
@@ -2128,22 +5295,24 @@
val symmetricDifference : 'elt set -> 'elt set -> 'elt set
-val disjoint : 'elt set -> 'elt set -> bool
-
-val subset : 'elt set -> 'elt set -> bool
-
-val equal : 'elt set -> 'elt set -> bool
+(* ------------------------------------------------------------------------- *)
+(* Mapping and folding. *)
+(* ------------------------------------------------------------------------- *)
val filter : ('elt -> bool) -> 'elt set -> 'elt set
val partition : ('elt -> bool) -> 'elt set -> 'elt set * 'elt set
-val count : ('elt -> bool) -> 'elt set -> int
+val app : ('elt -> unit) -> 'elt set -> unit
val foldl : ('elt * 's -> 's) -> 's -> 'elt set -> 's
val foldr : ('elt * 's -> 's) -> 's -> 'elt set -> 's
+(* ------------------------------------------------------------------------- *)
+(* Searching. *)
+(* ------------------------------------------------------------------------- *)
+
val findl : ('elt -> bool) -> 'elt set -> 'elt option
val findr : ('elt -> bool) -> 'elt set -> 'elt option
@@ -2156,27 +5325,45 @@
val all : ('elt -> bool) -> 'elt set -> bool
-val map : ('elt -> 'a) -> 'elt set -> ('elt * 'a) list
+val count : ('elt -> bool) -> 'elt set -> int
+
+(* ------------------------------------------------------------------------- *)
+(* Comparing. *)
+(* ------------------------------------------------------------------------- *)
+
+val compare : 'elt set * 'elt set -> order
+
+val equal : 'elt set -> 'elt set -> bool
+
+val subset : 'elt set -> 'elt set -> bool
+
+val disjoint : 'elt set -> 'elt set -> bool
+
+(* ------------------------------------------------------------------------- *)
+(* Converting to and from lists. *)
+(* ------------------------------------------------------------------------- *)
val transform : ('elt -> 'a) -> 'elt set -> 'a list
-val app : ('elt -> unit) -> 'elt set -> unit
-
val toList : 'elt set -> 'elt list
val fromList : ('elt * 'elt -> order) -> 'elt list -> 'elt set
-val pick : 'elt set -> 'elt (* raises Empty *)
-
-val random : 'elt set -> 'elt (* raises Empty *)
-
-val deletePick : 'elt set -> 'elt * 'elt set (* raises Empty *)
-
-val deleteRandom : 'elt set -> 'elt * 'elt set (* raises Empty *)
-
-val compare : 'elt set * 'elt set -> order
-
-val close : ('elt set -> 'elt set) -> 'elt set -> 'elt set
+(* ------------------------------------------------------------------------- *)
+(* Converting to and from maps. *)
+(* ------------------------------------------------------------------------- *)
+
+type ('elt,'a) map = ('elt,'a) Metis.Map.map
+
+val mapPartial : ('elt -> 'a option) -> 'elt set -> ('elt,'a) map
+
+val map : ('elt -> 'a) -> 'elt set -> ('elt,'a) map
+
+val domain : ('elt,'a) map -> 'elt set
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printing. *)
+(* ------------------------------------------------------------------------- *)
val toString : 'elt set -> string
@@ -2196,11 +5383,11 @@
end
-(**** Original file: RandomSet.sml ****)
+(**** Original file: Set.sml ****)
structure Metis = struct open Metis
(* Metis-specific ML environment *)
-nonfix ++ -- RL mem;
+nonfix ++ -- RL;
val explode = String.explode;
val implode = String.implode;
val print = TextIO.print;
@@ -2209,667 +5396,414 @@
(* ========================================================================= *)
(* FINITE SETS IMPLEMENTED WITH RANDOMLY BALANCED TREES *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure RandomSet :> Set =
-struct
-
-exception Bug = Useful.Bug;
-
-exception Error = Useful.Error;
-
-val pointerEqual = Portable.pointerEqual;
-
-val K = Useful.K;
-
-val snd = Useful.snd;
-
-val randomInt = Portable.randomInt;
-
-val randomWord = Portable.randomWord;
-
-(* ------------------------------------------------------------------------- *)
-(* Random search trees. *)
-(* ------------------------------------------------------------------------- *)
-
-type priority = Word.word;
-
-datatype 'a tree =
- E
- | T of
- {size : int,
- priority : priority,
- left : 'a tree,
- key : 'a,
- right : 'a tree};
-
-type 'a node =
- {size : int,
- priority : priority,
- left : 'a tree,
- key : 'a,
- right : 'a tree};
-
-datatype 'a set = Set of ('a * 'a -> order) * 'a tree;
-
-(* ------------------------------------------------------------------------- *)
-(* Random priorities. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- val randomPriority = randomWord;
-
- val priorityOrder = Word.compare;
-in
- fun treeSingleton key =
- T {size = 1, priority = randomPriority (),
- left = E, key = key, right = E};
-
- fun nodePriorityOrder cmp (x1 : 'a node, x2 : 'a node) =
- let
- val {priority = p1, key = k1, ...} = x1
- and {priority = p2, key = k2, ...} = x2
- in
- case priorityOrder (p1,p2) of
- LESS => LESS
- | EQUAL => cmp (k1,k2)
- | GREATER => GREATER
- end;
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Debugging functions. *)
-(* ------------------------------------------------------------------------- *)
-
-local
- fun checkSizes E = 0
- | checkSizes (T {size,left,right,...}) =
- let
- val l = checkSizes left
- and r = checkSizes right
- val () = if l + 1 + r = size then () else raise Error "wrong size"
- in
- size
- end
-
- fun checkSorted _ x E = x
- | checkSorted cmp x (T {left,key,right,...}) =
- let
- val x = checkSorted cmp x left
- val () =
- case x of
- NONE => ()
- | SOME k =>
- case cmp (k,key) of
- LESS => ()
- | EQUAL => raise Error "duplicate keys"
- | GREATER => raise Error "unsorted"
- in
- checkSorted cmp (SOME key) right
- end;
-
- fun checkPriorities _ E = NONE
- | checkPriorities cmp (T (x as {left,right,...})) =
- let
- val () =
- case checkPriorities cmp left of
- NONE => ()
- | SOME l =>
- case nodePriorityOrder cmp (l,x) of
- LESS => ()
- | EQUAL => raise Error "left child has equal key"
- | GREATER => raise Error "left child has greater priority"
- val () =
- case checkPriorities cmp right of
- NONE => ()
- | SOME r =>
- case nodePriorityOrder cmp (r,x) of
- LESS => ()
- | EQUAL => raise Error "right child has equal key"
- | GREATER => raise Error "right child has greater priority"
- in
- SOME x
- end;
-in
- fun checkWellformed s (set as Set (cmp,tree)) =
- (let
- val _ = checkSizes tree
- val _ = checkSorted cmp NONE tree
- val _ = checkPriorities cmp tree
- in
- set
- end
- handle Error err => raise Bug err)
- handle Bug bug => raise Bug (s ^ "\nRandomSet.checkWellformed: " ^ bug);
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Basic operations. *)
-(* ------------------------------------------------------------------------- *)
-
-fun comparison (Set (cmp,_)) = cmp;
-
-fun empty cmp = Set (cmp,E);
-
-fun treeSize E = 0
- | treeSize (T {size = s, ...}) = s;
-
-fun size (Set (_,tree)) = treeSize tree;
-
-fun mkT p l k r =
- T {size = treeSize l + 1 + treeSize r, priority = p,
- left = l, key = k, right = r};
-
-fun singleton cmp key = Set (cmp, treeSingleton key);
-
-local
- fun treePeek cmp E pkey = NONE
- | treePeek cmp (T {left,key,right,...}) pkey =
- case cmp (pkey,key) of
- LESS => treePeek cmp left pkey
- | EQUAL => SOME key
- | GREATER => treePeek cmp right pkey
-in
- fun peek (Set (cmp,tree)) key = treePeek cmp tree key;
-end;
-
-(* treeAppend assumes that every element of the first tree is less than *)
-(* every element of the second tree. *)
-
-fun treeAppend _ t1 E = t1
- | treeAppend _ E t2 = t2
- | treeAppend cmp (t1 as T x1) (t2 as T x2) =
- case nodePriorityOrder cmp (x1,x2) of
- LESS =>
- let
- val {priority = p2, left = l2, key = k2, right = r2, ...} = x2
- in
- mkT p2 (treeAppend cmp t1 l2) k2 r2
- end
- | EQUAL => raise Bug "RandomSet.treeAppend: equal keys"
- | GREATER =>
- let
- val {priority = p1, left = l1, key = k1, right = r1, ...} = x1
- in
- mkT p1 l1 k1 (treeAppend cmp r1 t2)
- end;
-
-(* nodePartition splits the node into three parts: the keys comparing less *)
-(* than the supplied key, an optional equal key, and the keys comparing *)
-(* greater. *)
-
-local
- fun mkLeft [] t = t
- | mkLeft (({priority,left,key,...} : 'a node) :: xs) t =
- mkLeft xs (mkT priority left key t);
-
- fun mkRight [] t = t
- | mkRight (({priority,key,right,...} : 'a node) :: xs) t =
- mkRight xs (mkT priority t key right);
-
- fun treePart _ _ lefts rights E = (mkLeft lefts E, NONE, mkRight rights E)
- | treePart cmp pkey lefts rights (T x) = nodePart cmp pkey lefts rights x
- and nodePart cmp pkey lefts rights (x as {left,key,right,...}) =
- case cmp (pkey,key) of
- LESS => treePart cmp pkey lefts (x :: rights) left
- | EQUAL => (mkLeft lefts left, SOME key, mkRight rights right)
- | GREATER => treePart cmp pkey (x :: lefts) rights right;
-in
- fun nodePartition cmp x pkey = nodePart cmp pkey [] [] x;
-end;
-
-(* union first calls treeCombineRemove, to combine the values *)
-(* for equal keys into the first map and remove them from the second map. *)
-(* Note that the combined key is always the one from the second map. *)
-
-local
- fun treeCombineRemove _ t1 E = (t1,E)
- | treeCombineRemove _ E t2 = (E,t2)
- | treeCombineRemove cmp (t1 as T x1) (t2 as T x2) =
- let
- val {priority = p1, left = l1, key = k1, right = r1, ...} = x1
- val (l2,k2,r2) = nodePartition cmp x2 k1
- val (l1,l2) = treeCombineRemove cmp l1 l2
- and (r1,r2) = treeCombineRemove cmp r1 r2
- in
- case k2 of
- NONE => if treeSize l2 + treeSize r2 = #size x2 then (t1,t2)
- else (mkT p1 l1 k1 r1, treeAppend cmp l2 r2)
- | SOME k2 => (mkT p1 l1 k2 r1, treeAppend cmp l2 r2)
- end;
-
- fun treeUnionDisjoint _ t1 E = t1
- | treeUnionDisjoint _ E t2 = t2
- | treeUnionDisjoint cmp (T x1) (T x2) =
- case nodePriorityOrder cmp (x1,x2) of
- LESS => nodeUnionDisjoint cmp x2 x1
- | EQUAL => raise Bug "RandomSet.unionDisjoint: equal keys"
- | GREATER => nodeUnionDisjoint cmp x1 x2
-
- and nodeUnionDisjoint cmp x1 x2 =
- let
- val {priority = p1, left = l1, key = k1, right = r1, ...} = x1
- val (l2,_,r2) = nodePartition cmp x2 k1
- val l = treeUnionDisjoint cmp l1 l2
- and r = treeUnionDisjoint cmp r1 r2
- in
- mkT p1 l k1 r
- end;
-in
- fun union (s1 as Set (cmp,t1)) (Set (_,t2)) =
- if pointerEqual (t1,t2) then s1
- else
- let
- val (t1,t2) = treeCombineRemove cmp t1 t2
- in
- Set (cmp, treeUnionDisjoint cmp t1 t2)
- end;
-end;
-
-(*DEBUG
-val union = fn t1 => fn t2 =>
- checkWellformed "RandomSet.union: result"
- (union (checkWellformed "RandomSet.union: input 1" t1)
- (checkWellformed "RandomSet.union: input 2" t2));
-*)
-
-(* intersect is a simple case of the union algorithm. *)
-
-local
- fun treeIntersect _ _ E = E
- | treeIntersect _ E _ = E
- | treeIntersect cmp (t1 as T x1) (t2 as T x2) =
- let
- val {priority = p1, left = l1, key = k1, right = r1, ...} = x1
- val (l2,k2,r2) = nodePartition cmp x2 k1
- val l = treeIntersect cmp l1 l2
- and r = treeIntersect cmp r1 r2
- in
- case k2 of
- NONE => treeAppend cmp l r
- | SOME k2 => mkT p1 l k2 r
- end;
-in
- fun intersect (s1 as Set (cmp,t1)) (Set (_,t2)) =
- if pointerEqual (t1,t2) then s1
- else Set (cmp, treeIntersect cmp t1 t2);
-end;
-
-(*DEBUG
-val intersect = fn t1 => fn t2 =>
- checkWellformed "RandomSet.intersect: result"
- (intersect (checkWellformed "RandomSet.intersect: input 1" t1)
- (checkWellformed "RandomSet.intersect: input 2" t2));
-*)
-
-(* delete raises an exception if the supplied key is not found, which *)
-(* makes it simpler to maximize sharing. *)
-
-local
- fun treeDelete _ E _ = raise Error "RandomSet.delete: element not found"
- | treeDelete cmp (T {priority,left,key,right,...}) dkey =
- case cmp (dkey,key) of
- LESS => mkT priority (treeDelete cmp left dkey) key right
- | EQUAL => treeAppend cmp left right
- | GREATER => mkT priority left key (treeDelete cmp right dkey);
-in
- fun delete (Set (cmp,tree)) key = Set (cmp, treeDelete cmp tree key);
-end;
-
-(*DEBUG
-val delete = fn t => fn x =>
- checkWellformed "RandomSet.delete: result"
- (delete (checkWellformed "RandomSet.delete: input" t) x);
-*)
-
-(* Set difference *)
-
-local
- fun treeDifference _ t1 E = t1
- | treeDifference _ E _ = E
- | treeDifference cmp (t1 as T x1) (T x2) =
- let
- val {size = s1, priority = p1, left = l1, key = k1, right = r1} = x1
- val (l2,k2,r2) = nodePartition cmp x2 k1
- val l = treeDifference cmp l1 l2
- and r = treeDifference cmp r1 r2
- in
- if Option.isSome k2 then treeAppend cmp l r
- else if treeSize l + treeSize r + 1 = s1 then t1
- else mkT p1 l k1 r
- end;
-in
- fun difference (Set (cmp,tree1)) (Set (_,tree2)) =
- if pointerEqual (tree1,tree2) then Set (cmp,E)
- else Set (cmp, treeDifference cmp tree1 tree2);
-end;
-
-(*DEBUG
-val difference = fn t1 => fn t2 =>
- checkWellformed "RandomSet.difference: result"
- (difference (checkWellformed "RandomSet.difference: input 1" t1)
- (checkWellformed "RandomSet.difference: input 2" t2));
-*)
-
-(* Subsets *)
-
-local
- fun treeSubset _ E _ = true
- | treeSubset _ _ E = false
- | treeSubset cmp (t1 as T x1) (T x2) =
- let
- val {size = s1, left = l1, key = k1, right = r1, ...} = x1
- and {size = s2, ...} = x2
- in
- s1 <= s2 andalso
- let
- val (l2,k2,r2) = nodePartition cmp x2 k1
- in
- Option.isSome k2 andalso
- treeSubset cmp l1 l2 andalso
- treeSubset cmp r1 r2
- end
- end;
-in
- fun subset (Set (cmp,tree1)) (Set (_,tree2)) =
- pointerEqual (tree1,tree2) orelse
- treeSubset cmp tree1 tree2;
-end;
-
-(* Set equality *)
-
-local
- fun treeEqual _ E E = true
- | treeEqual _ E _ = false
- | treeEqual _ _ E = false
- | treeEqual cmp (t1 as T x1) (T x2) =
- let
- val {size = s1, left = l1, key = k1, right = r1, ...} = x1
- and {size = s2, ...} = x2
- in
- s1 = s2 andalso
- let
- val (l2,k2,r2) = nodePartition cmp x2 k1
- in
- Option.isSome k2 andalso
- treeEqual cmp l1 l2 andalso
- treeEqual cmp r1 r2
- end
- end;
-in
- fun equal (Set (cmp,tree1)) (Set (_,tree2)) =
- pointerEqual (tree1,tree2) orelse
- treeEqual cmp tree1 tree2;
-end;
-
-(* filter is the basic function for preserving the tree structure. *)
-
-local
- fun treeFilter _ _ E = E
- | treeFilter cmp pred (T {priority,left,key,right,...}) =
- let
- val left = treeFilter cmp pred left
- and right = treeFilter cmp pred right
- in
- if pred key then mkT priority left key right
- else treeAppend cmp left right
- end;
-in
- fun filter pred (Set (cmp,tree)) = Set (cmp, treeFilter cmp pred tree);
-end;
-
-(* nth picks the nth smallest key (counting from 0). *)
-
-local
- fun treeNth E _ = raise Subscript
- | treeNth (T {left,key,right,...}) n =
- let
- val k = treeSize left
- in
- if n = k then key
- else if n < k then treeNth left n
- else treeNth right (n - (k + 1))
- end;
-in
- fun nth (Set (_,tree)) n = treeNth tree n;
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Iterators. *)
-(* ------------------------------------------------------------------------- *)
-
-fun leftSpine E acc = acc
- | leftSpine (t as T {left,...}) acc = leftSpine left (t :: acc);
-
-fun rightSpine E acc = acc
- | rightSpine (t as T {right,...}) acc = rightSpine right (t :: acc);
-
-datatype 'a iterator =
- LR of 'a * 'a tree * 'a tree list
- | RL of 'a * 'a tree * 'a tree list;
-
-fun mkLR [] = NONE
- | mkLR (T {key,right,...} :: l) = SOME (LR (key,right,l))
- | mkLR (E :: _) = raise Bug "RandomSet.mkLR";
-
-fun mkRL [] = NONE
- | mkRL (T {key,left,...} :: l) = SOME (RL (key,left,l))
- | mkRL (E :: _) = raise Bug "RandomSet.mkRL";
-
-fun mkIterator (Set (_,tree)) = mkLR (leftSpine tree []);
-
-fun mkRevIterator (Set (_,tree)) = mkRL (rightSpine tree []);
-
-fun readIterator (LR (key,_,_)) = key
- | readIterator (RL (key,_,_)) = key;
-
-fun advanceIterator (LR (_,next,l)) = mkLR (leftSpine next l)
- | advanceIterator (RL (_,next,l)) = mkRL (rightSpine next l);
-
-(* ------------------------------------------------------------------------- *)
-(* Derived operations. *)
-(* ------------------------------------------------------------------------- *)
-
-fun null s = size s = 0;
-
-fun member x s = Option.isSome (peek s x);
-
-fun add s x = union s (singleton (comparison s) x);
-
-(*DEBUG
-val add = fn s => fn x =>
- checkWellformed "RandomSet.add: result"
- (add (checkWellformed "RandomSet.add: input" s) x);
-*)
-
-local
- fun unionPairs ys [] = rev ys
- | unionPairs ys (xs as [_]) = List.revAppend (ys,xs)
- | unionPairs ys (x1 :: x2 :: xs) = unionPairs (union x1 x2 :: ys) xs;
-in
- fun unionList [] = raise Error "RandomSet.unionList: no sets"
- | unionList [s] = s
- | unionList l = unionList (unionPairs [] l);
-end;
-
-local
- fun intersectPairs ys [] = rev ys
- | intersectPairs ys (xs as [_]) = List.revAppend (ys,xs)
- | intersectPairs ys (x1 :: x2 :: xs) =
- intersectPairs (intersect x1 x2 :: ys) xs;
-in
- fun intersectList [] = raise Error "RandomSet.intersectList: no sets"
- | intersectList [s] = s
- | intersectList l = intersectList (intersectPairs [] l);
-end;
-
-fun symmetricDifference s1 s2 = union (difference s1 s2) (difference s2 s1);
-
-fun disjoint s1 s2 = null (intersect s1 s2);
-
-fun partition pred set = (filter pred set, filter (not o pred) set);
-
-local
- fun fold _ NONE acc = acc
- | fold f (SOME iter) acc =
- let
- val key = readIterator iter
- in
- fold f (advanceIterator iter) (f (key,acc))
- end;
-in
- fun foldl f b m = fold f (mkIterator m) b;
-
- fun foldr f b m = fold f (mkRevIterator m) b;
-end;
-
-local
- fun find _ NONE = NONE
- | find pred (SOME iter) =
- let
- val key = readIterator iter
- in
- if pred key then SOME key
- else find pred (advanceIterator iter)
- end;
-in
- fun findl p m = find p (mkIterator m);
-
- fun findr p m = find p (mkRevIterator m);
-end;
-
-local
- fun first _ NONE = NONE
- | first f (SOME iter) =
- let
- val key = readIterator iter
- in
- case f key of
- NONE => first f (advanceIterator iter)
- | s => s
- end;
-in
- fun firstl f m = first f (mkIterator m);
-
- fun firstr f m = first f (mkRevIterator m);
-end;
-
-fun count p = foldl (fn (x,n) => if p x then n + 1 else n) 0;
-
-fun fromList cmp l = List.foldl (fn (k,s) => add s k) (empty cmp) l;
-
-fun addList s l = union s (fromList (comparison s) l);
-
-fun toList s = foldr op:: [] s;
-
-fun map f s = rev (foldl (fn (x,l) => (x, f x) :: l) [] s);
-
-fun transform f s = rev (foldl (fn (x,l) => f x :: l) [] s);
-
-fun app f s = foldl (fn (x,()) => f x) () s;
-
-fun exists p s = Option.isSome (findl p s);
-
-fun all p s = not (exists (not o p) s);
-
-local
- fun iterCompare _ NONE NONE = EQUAL
- | iterCompare _ NONE (SOME _) = LESS
- | iterCompare _ (SOME _) NONE = GREATER
- | iterCompare cmp (SOME i1) (SOME i2) =
- keyIterCompare cmp (readIterator i1) (readIterator i2) i1 i2
-
- and keyIterCompare cmp k1 k2 i1 i2 =
- case cmp (k1,k2) of
- LESS => LESS
- | EQUAL => iterCompare cmp (advanceIterator i1) (advanceIterator i2)
- | GREATER => GREATER;
-in
- fun compare (s1,s2) =
- if pointerEqual (s1,s2) then EQUAL
- else
- case Int.compare (size s1, size s2) of
- LESS => LESS
- | EQUAL => iterCompare (comparison s1) (mkIterator s1) (mkIterator s2)
- | GREATER => GREATER;
-end;
-
-fun pick s =
- case findl (K true) s of
- SOME p => p
- | NONE => raise Error "RandomSet.pick: empty";
-
-fun random s =
- case size s of
- 0 => raise Error "RandomSet.random: empty"
- | n => nth s (randomInt n);
-
-fun deletePick s = let val x = pick s in (x, delete s x) end;
-
-fun deleteRandom s = let val x = random s in (x, delete s x) end;
-
-fun close f s = let val s' = f s in if equal s s' then s else close f s' end;
-
-fun toString s = "{" ^ (if null s then "" else Int.toString (size s)) ^ "}";
-
-end
-end;
-
-(**** Original file: Set.sml ****)
-
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL mem;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
-(* ========================================================================= *)
-(* FINITE SETS *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-structure Set = RandomSet;
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
+(* ========================================================================= *)
+
+structure Set :> Set =
+struct
+
+(* ------------------------------------------------------------------------- *)
+(* A type of finite sets. *)
+(* ------------------------------------------------------------------------- *)
+
+type ('elt,'a) map = ('elt,'a) Map.map;
+
+datatype 'elt set = Set of ('elt,unit) map;
+
+(* ------------------------------------------------------------------------- *)
+(* Converting to and from maps. *)
+(* ------------------------------------------------------------------------- *)
+
+fun dest (Set m) = m;
+
+fun mapPartial f =
+ let
+ fun mf (elt,()) = f elt
+ in
+ fn Set m => Map.mapPartial mf m
+ end;
+
+fun map f =
+ let
+ fun mf (elt,()) = f elt
+ in
+ fn Set m => Map.map mf m
+ end;
+
+fun domain m = Set (Map.transform (fn _ => ()) m);
+
+(* ------------------------------------------------------------------------- *)
+(* Constructors. *)
+(* ------------------------------------------------------------------------- *)
+
+fun empty cmp = Set (Map.new cmp);
+
+fun singleton cmp elt = Set (Map.singleton cmp (elt,()));
+
+(* ------------------------------------------------------------------------- *)
+(* Set size. *)
+(* ------------------------------------------------------------------------- *)
+
+fun null (Set m) = Map.null m;
+
+fun size (Set m) = Map.size m;
+
+(* ------------------------------------------------------------------------- *)
+(* Querying. *)
+(* ------------------------------------------------------------------------- *)
+
+fun peek (Set m) elt =
+ case Map.peekKey m elt of
+ SOME (elt,()) => SOME elt
+ | NONE => NONE;
+
+fun member elt (Set m) = Map.inDomain elt m;
+
+fun pick (Set m) =
+ let
+ val (elt,_) = Map.pick m
+ in
+ elt
+ end;
+
+fun nth (Set m) n =
+ let
+ val (elt,_) = Map.nth m n
+ in
+ elt
+ end;
+
+fun random (Set m) =
+ let
+ val (elt,_) = Map.random m
+ in
+ elt
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Adding. *)
+(* ------------------------------------------------------------------------- *)
+
+fun add (Set m) elt =
+ let
+ val m = Map.insert m (elt,())
+ in
+ Set m
+ end;
+
+local
+ fun uncurriedAdd (elt,set) = add set elt;
+in
+ fun addList set = List.foldl uncurriedAdd set;
+end;
+
+(* ------------------------------------------------------------------------- *)
+(* Removing. *)
+(* ------------------------------------------------------------------------- *)
+
+fun delete (Set m) elt =
+ let
+ val m = Map.delete m elt
+ in
+ Set m
+ end;
+
+fun remove (Set m) elt =
+ let
+ val m = Map.remove m elt
+ in
+ Set m
+ end;
+
+fun deletePick (Set m) =
+ let
+ val ((elt,()),m) = Map.deletePick m
+ in
+ (elt, Set m)
+ end;
+
+fun deleteNth (Set m) n =
+ let
+ val ((elt,()),m) = Map.deleteNth m n
+ in
+ (elt, Set m)
+ end;
+
+fun deleteRandom (Set m) =
+ let
+ val ((elt,()),m) = Map.deleteRandom m
+ in
+ (elt, Set m)
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Joining. *)
+(* ------------------------------------------------------------------------- *)
+
+fun union (Set m1) (Set m2) = Set (Map.unionDomain m1 m2);
+
+fun unionList sets =
+ let
+ val ms = List.map dest sets
+ in
+ Set (Map.unionListDomain ms)
+ end;
+
+fun intersect (Set m1) (Set m2) = Set (Map.intersectDomain m1 m2);
+
+fun intersectList sets =
+ let
+ val ms = List.map dest sets
+ in
+ Set (Map.intersectListDomain ms)
+ end;
+
+fun difference (Set m1) (Set m2) =
+ Set (Map.differenceDomain m1 m2);
+
+fun symmetricDifference (Set m1) (Set m2) =
+ Set (Map.symmetricDifferenceDomain m1 m2);
+
+(* ------------------------------------------------------------------------- *)
+(* Mapping and folding. *)
+(* ------------------------------------------------------------------------- *)
+
+fun filter pred =
+ let
+ fun mpred (elt,()) = pred elt
+ in
+ fn Set m => Set (Map.filter mpred m)
+ end;
+
+fun partition pred =
+ let
+ fun mpred (elt,()) = pred elt
+ in
+ fn Set m =>
+ let
+ val (m1,m2) = Map.partition mpred m
+ in
+ (Set m1, Set m2)
+ end
+ end;
+
+fun app f =
+ let
+ fun mf (elt,()) = f elt
+ in
+ fn Set m => Map.app mf m
+ end;
+
+fun foldl f =
+ let
+ fun mf (elt,(),acc) = f (elt,acc)
+ in
+ fn acc => fn Set m => Map.foldl mf acc m
+ end;
+
+fun foldr f =
+ let
+ fun mf (elt,(),acc) = f (elt,acc)
+ in
+ fn acc => fn Set m => Map.foldr mf acc m
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Searching. *)
+(* ------------------------------------------------------------------------- *)
+
+fun findl p =
+ let
+ fun mp (elt,()) = p elt
+ in
+ fn Set m =>
+ case Map.findl mp m of
+ SOME (elt,()) => SOME elt
+ | NONE => NONE
+ end;
+
+fun findr p =
+ let
+ fun mp (elt,()) = p elt
+ in
+ fn Set m =>
+ case Map.findr mp m of
+ SOME (elt,()) => SOME elt
+ | NONE => NONE
+ end;
+
+fun firstl f =
+ let
+ fun mf (elt,()) = f elt
+ in
+ fn Set m => Map.firstl mf m
+ end;
+
+fun firstr f =
+ let
+ fun mf (elt,()) = f elt
+ in
+ fn Set m => Map.firstr mf m
+ end;
+
+fun exists p =
+ let
+ fun mp (elt,()) = p elt
+ in
+ fn Set m => Map.exists mp m
+ end;
+
+fun all p =
+ let
+ fun mp (elt,()) = p elt
+ in
+ fn Set m => Map.all mp m
+ end;
+
+fun count p =
+ let
+ fun mp (elt,()) = p elt
+ in
+ fn Set m => Map.count mp m
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Comparing. *)
+(* ------------------------------------------------------------------------- *)
+
+fun compareValue ((),()) = EQUAL;
+
+fun equalValue () () = true;
+
+fun compare (Set m1, Set m2) = Map.compare compareValue (m1,m2);
+
+fun equal (Set m1) (Set m2) = Map.equal equalValue m1 m2;
+
+fun subset (Set m1) (Set m2) = Map.subsetDomain m1 m2;
+
+fun disjoint (Set m1) (Set m2) = Map.disjointDomain m1 m2;
+
+(* ------------------------------------------------------------------------- *)
+(* Converting to and from lists. *)
+(* ------------------------------------------------------------------------- *)
+
+fun transform f =
+ let
+ fun inc (x,l) = f x :: l
+ in
+ foldr inc []
+ end;
+
+fun toList (Set m) = Map.keys m;
+
+fun fromList cmp elts = addList (empty cmp) elts;
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printing. *)
+(* ------------------------------------------------------------------------- *)
+
+fun toString set =
+ "{" ^ (if null set then "" else Int.toString (size set)) ^ "}";
+
+(* ------------------------------------------------------------------------- *)
+(* Iterators over sets *)
+(* ------------------------------------------------------------------------- *)
+
+type 'elt iterator = ('elt,unit) Map.iterator;
+
+fun mkIterator (Set m) = Map.mkIterator m;
+
+fun mkRevIterator (Set m) = Map.mkRevIterator m;
+
+fun readIterator iter =
+ let
+ val (elt,()) = Map.readIterator iter
+ in
+ elt
+ end;
+
+fun advanceIterator iter = Map.advanceIterator iter;
+
+end
end;
(**** Original file: ElementSet.sig ****)
(* ========================================================================= *)
(* FINITE SETS WITH A FIXED ELEMENT TYPE *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature ElementSet =
sig
+(* ------------------------------------------------------------------------- *)
+(* A type of set elements. *)
+(* ------------------------------------------------------------------------- *)
+
type element
(* ------------------------------------------------------------------------- *)
-(* Finite sets *)
+(* A type of finite sets. *)
(* ------------------------------------------------------------------------- *)
type set
+(* ------------------------------------------------------------------------- *)
+(* Constructors. *)
+(* ------------------------------------------------------------------------- *)
+
val empty : set
val singleton : element -> set
+(* ------------------------------------------------------------------------- *)
+(* Set size. *)
+(* ------------------------------------------------------------------------- *)
+
val null : set -> bool
val size : set -> int
+(* ------------------------------------------------------------------------- *)
+(* Querying. *)
+(* ------------------------------------------------------------------------- *)
+
+val peek : set -> element -> element option
+
val member : element -> set -> bool
+val pick : set -> element (* an arbitrary element *)
+
+val nth : set -> int -> element (* in the range [0,size-1] *)
+
+val random : set -> element
+
+(* ------------------------------------------------------------------------- *)
+(* Adding. *)
+(* ------------------------------------------------------------------------- *)
+
val add : set -> element -> set
val addList : set -> element list -> set
-val delete : set -> element -> set (* raises Error *)
-
-(* Union and intersect prefer elements in the second set *)
+(* ------------------------------------------------------------------------- *)
+(* Removing. *)
+(* ------------------------------------------------------------------------- *)
+
+val delete : set -> element -> set (* must be present *)
+
+val remove : set -> element -> set
+
+val deletePick : set -> element * set
+
+val deleteNth : set -> int -> element * set
+
+val deleteRandom : set -> element * set
+
+(* ------------------------------------------------------------------------- *)
+(* Joining. *)
+(* ------------------------------------------------------------------------- *)
val union : set -> set -> set
@@ -2883,22 +5817,24 @@
val symmetricDifference : set -> set -> set
-val disjoint : set -> set -> bool
-
-val subset : set -> set -> bool
-
-val equal : set -> set -> bool
+(* ------------------------------------------------------------------------- *)
+(* Mapping and folding. *)
+(* ------------------------------------------------------------------------- *)
val filter : (element -> bool) -> set -> set
val partition : (element -> bool) -> set -> set * set
-val count : (element -> bool) -> set -> int
+val app : (element -> unit) -> set -> unit
val foldl : (element * 's -> 's) -> 's -> set -> 's
val foldr : (element * 's -> 's) -> 's -> set -> 's
+(* ------------------------------------------------------------------------- *)
+(* Searching. *)
+(* ------------------------------------------------------------------------- *)
+
val findl : (element -> bool) -> set -> element option
val findr : (element -> bool) -> set -> element option
@@ -2911,27 +5847,45 @@
val all : (element -> bool) -> set -> bool
-val map : (element -> 'a) -> set -> (element * 'a) list
+val count : (element -> bool) -> set -> int
+
+(* ------------------------------------------------------------------------- *)
+(* Comparing. *)
+(* ------------------------------------------------------------------------- *)
+
+val compare : set * set -> order
+
+val equal : set -> set -> bool
+
+val subset : set -> set -> bool
+
+val disjoint : set -> set -> bool
+
+(* ------------------------------------------------------------------------- *)
+(* Converting to and from lists. *)
+(* ------------------------------------------------------------------------- *)
val transform : (element -> 'a) -> set -> 'a list
-val app : (element -> unit) -> set -> unit
-
val toList : set -> element list
val fromList : element list -> set
-val pick : set -> element (* raises Empty *)
-
-val random : set -> element (* raises Empty *)
-
-val deletePick : set -> element * set (* raises Empty *)
-
-val deleteRandom : set -> element * set (* raises Empty *)
-
-val compare : set * set -> order
-
-val close : (set -> set) -> set -> set
+(* ------------------------------------------------------------------------- *)
+(* Converting to and from maps. *)
+(* ------------------------------------------------------------------------- *)
+
+type 'a map
+
+val mapPartial : (element -> 'a option) -> set -> 'a map
+
+val map : (element -> 'a) -> set -> 'a map
+
+val domain : 'a map -> set
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printing. *)
+(* ------------------------------------------------------------------------- *)
val toString : set -> string
@@ -2955,1128 +5909,369 @@
(* ========================================================================= *)
(* FINITE SETS WITH A FIXED ELEMENT TYPE *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-functor ElementSet (Key : Ordered) :> ElementSet where type element = Key.t =
-struct
-
- open Metis;
-
-type element = Key.t;
-
-(* ------------------------------------------------------------------------- *)
-(* Finite sets *)
-(* ------------------------------------------------------------------------- *)
-
-type set = Key.t Set.set;
-
-val empty = Set.empty Key.compare;
-
-fun singleton key = Set.singleton Key.compare key;
-
-val null = Set.null;
-
-val size = Set.size;
-
-val member = Set.member;
-
-val add = Set.add;
-
-val addList = Set.addList;
-
-val delete = Set.delete;
-
-val op union = Set.union;
-
-val unionList = Set.unionList;
-
-val intersect = Set.intersect;
-
-val intersectList = Set.intersectList;
-
-val difference = Set.difference;
-
-val symmetricDifference = Set.symmetricDifference;
-
-val disjoint = Set.disjoint;
-
-val op subset = Set.subset;
-
-val equal = Set.equal;
-
-val filter = Set.filter;
-
-val partition = Set.partition;
-
-val count = Set.count;
-
-val foldl = Set.foldl;
-
-val foldr = Set.foldr;
-
-val findl = Set.findl;
-
-val findr = Set.findr;
-
-val firstl = Set.firstl;
-
-val firstr = Set.firstr;
-
-val exists = Set.exists;
-
-val all = Set.all;
-
-val map = Set.map;
-
-val transform = Set.transform;
-
-val app = Set.app;
-
-val toList = Set.toList;
-
-fun fromList l = Set.fromList Key.compare l;
-
-val pick = Set.pick;
-
-val random = Set.random;
-
-val deletePick = Set.deletePick;
-
-val deleteRandom = Set.deleteRandom;
-
-val compare = Set.compare;
-
-val close = Set.close;
-
-val toString = Set.toString;
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
+(* ========================================================================= *)
+
+functor ElementSet (KM : KeyMap) :> ElementSet
+where type element = KM.key and type 'a map = 'a KM.map =
+struct
+
+(* ------------------------------------------------------------------------- *)
+(* A type of set elements. *)
+(* ------------------------------------------------------------------------- *)
+
+type element = KM.key;
+
+(* ------------------------------------------------------------------------- *)
+(* A type of finite sets. *)
+(* ------------------------------------------------------------------------- *)
+
+type 'a map = 'a KM.map;
+
+datatype set = Set of unit map;
+
+(* ------------------------------------------------------------------------- *)
+(* Converting to and from maps. *)
+(* ------------------------------------------------------------------------- *)
+
+fun dest (Set m) = m;
+
+fun mapPartial f =
+ let
+ fun mf (elt,()) = f elt
+ in
+ fn Set m => KM.mapPartial mf m
+ end;
+
+fun map f =
+ let
+ fun mf (elt,()) = f elt
+ in
+ fn Set m => KM.map mf m
+ end;
+
+fun domain m = Set (KM.transform (fn _ => ()) m);
+
+(* ------------------------------------------------------------------------- *)
+(* Constructors. *)
+(* ------------------------------------------------------------------------- *)
+
+val empty = Set (KM.new ());
+
+fun singleton elt = Set (KM.singleton (elt,()));
+
+(* ------------------------------------------------------------------------- *)
+(* Set size. *)
+(* ------------------------------------------------------------------------- *)
+
+fun null (Set m) = KM.null m;
+
+fun size (Set m) = KM.size m;
+
+(* ------------------------------------------------------------------------- *)
+(* Querying. *)
+(* ------------------------------------------------------------------------- *)
+
+fun peek (Set m) elt =
+ case KM.peekKey m elt of
+ SOME (elt,()) => SOME elt
+ | NONE => NONE;
+
+fun member elt (Set m) = KM.inDomain elt m;
+
+fun pick (Set m) =
+ let
+ val (elt,_) = KM.pick m
+ in
+ elt
+ end;
+
+fun nth (Set m) n =
+ let
+ val (elt,_) = KM.nth m n
+ in
+ elt
+ end;
+
+fun random (Set m) =
+ let
+ val (elt,_) = KM.random m
+ in
+ elt
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Adding. *)
+(* ------------------------------------------------------------------------- *)
+
+fun add (Set m) elt =
+ let
+ val m = KM.insert m (elt,())
+ in
+ Set m
+ end;
+
+local
+ fun uncurriedAdd (elt,set) = add set elt;
+in
+ fun addList set = List.foldl uncurriedAdd set;
+end;
+
+(* ------------------------------------------------------------------------- *)
+(* Removing. *)
+(* ------------------------------------------------------------------------- *)
+
+fun delete (Set m) elt =
+ let
+ val m = KM.delete m elt
+ in
+ Set m
+ end;
+
+fun remove (Set m) elt =
+ let
+ val m = KM.remove m elt
+ in
+ Set m
+ end;
+
+fun deletePick (Set m) =
+ let
+ val ((elt,()),m) = KM.deletePick m
+ in
+ (elt, Set m)
+ end;
+
+fun deleteNth (Set m) n =
+ let
+ val ((elt,()),m) = KM.deleteNth m n
+ in
+ (elt, Set m)
+ end;
+
+fun deleteRandom (Set m) =
+ let
+ val ((elt,()),m) = KM.deleteRandom m
+ in
+ (elt, Set m)
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Joining. *)
+(* ------------------------------------------------------------------------- *)
+
+fun op union (Set m1) (Set m2) = Set (KM.unionDomain m1 m2);
+
+fun unionList sets =
+ let
+ val ms = List.map dest sets
+ in
+ Set (KM.unionListDomain ms)
+ end;
+
+fun intersect (Set m1) (Set m2) = Set (KM.intersectDomain m1 m2);
+
+fun intersectList sets =
+ let
+ val ms = List.map dest sets
+ in
+ Set (KM.intersectListDomain ms)
+ end;
+
+fun difference (Set m1) (Set m2) =
+ Set (KM.differenceDomain m1 m2);
+
+fun symmetricDifference (Set m1) (Set m2) =
+ Set (KM.symmetricDifferenceDomain m1 m2);
+
+(* ------------------------------------------------------------------------- *)
+(* Mapping and folding. *)
+(* ------------------------------------------------------------------------- *)
+
+fun filter pred =
+ let
+ fun mpred (elt,()) = pred elt
+ in
+ fn Set m => Set (KM.filter mpred m)
+ end;
+
+fun partition pred =
+ let
+ fun mpred (elt,()) = pred elt
+ in
+ fn Set m =>
+ let
+ val (m1,m2) = KM.partition mpred m
+ in
+ (Set m1, Set m2)
+ end
+ end;
+
+fun app f =
+ let
+ fun mf (elt,()) = f elt
+ in
+ fn Set m => KM.app mf m
+ end;
+
+fun foldl f =
+ let
+ fun mf (elt,(),acc) = f (elt,acc)
+ in
+ fn acc => fn Set m => KM.foldl mf acc m
+ end;
+
+fun foldr f =
+ let
+ fun mf (elt,(),acc) = f (elt,acc)
+ in
+ fn acc => fn Set m => KM.foldr mf acc m
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Searching. *)
+(* ------------------------------------------------------------------------- *)
+
+fun findl p =
+ let
+ fun mp (elt,()) = p elt
+ in
+ fn Set m =>
+ case KM.findl mp m of
+ SOME (elt,()) => SOME elt
+ | NONE => NONE
+ end;
+
+fun findr p =
+ let
+ fun mp (elt,()) = p elt
+ in
+ fn Set m =>
+ case KM.findr mp m of
+ SOME (elt,()) => SOME elt
+ | NONE => NONE
+ end;
+
+fun firstl f =
+ let
+ fun mf (elt,()) = f elt
+ in
+ fn Set m => KM.firstl mf m
+ end;
+
+fun firstr f =
+ let
+ fun mf (elt,()) = f elt
+ in
+ fn Set m => KM.firstr mf m
+ end;
+
+fun exists p =
+ let
+ fun mp (elt,()) = p elt
+ in
+ fn Set m => KM.exists mp m
+ end;
+
+fun all p =
+ let
+ fun mp (elt,()) = p elt
+ in
+ fn Set m => KM.all mp m
+ end;
+
+fun count p =
+ let
+ fun mp (elt,()) = p elt
+ in
+ fn Set m => KM.count mp m
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* Comparing. *)
+(* ------------------------------------------------------------------------- *)
+
+fun compareValue ((),()) = EQUAL;
+
+fun equalValue () () = true;
+
+fun compare (Set m1, Set m2) = KM.compare compareValue (m1,m2);
+
+fun equal (Set m1) (Set m2) = KM.equal equalValue m1 m2;
+
+fun op subset (Set m1) (Set m2) = KM.subsetDomain m1 m2;
+
+fun disjoint (Set m1) (Set m2) = KM.disjointDomain m1 m2;
+
+(* ------------------------------------------------------------------------- *)
+(* Converting to and from lists. *)
+(* ------------------------------------------------------------------------- *)
+
+fun transform f =
+ let
+ fun inc (x,l) = f x :: l
+ in
+ foldr inc []
+ end;
+
+fun toList (Set m) = KM.keys m;
+
+fun fromList elts = addList empty elts;
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printing. *)
+(* ------------------------------------------------------------------------- *)
+
+fun toString set =
+ "{" ^ (if null set then "" else Int.toString (size set)) ^ "}";
(* ------------------------------------------------------------------------- *)
(* Iterators over sets *)
(* ------------------------------------------------------------------------- *)
-type iterator = Key.t Set.iterator;
-
-val mkIterator = Set.mkIterator;
-
-val mkRevIterator = Set.mkRevIterator;
-
-val readIterator = Set.readIterator;
-
-val advanceIterator = Set.advanceIterator;
-
-end
-
- structure Metis = struct open Metis;
+type iterator = unit KM.iterator;
+
+fun mkIterator (Set m) = KM.mkIterator m;
+
+fun mkRevIterator (Set m) = KM.mkRevIterator m;
+
+fun readIterator iter =
+ let
+ val (elt,()) = KM.readIterator iter
+ in
+ elt
+ end;
+
+fun advanceIterator iter = KM.advanceIterator iter;
+
+end
structure IntSet =
-ElementSet (IntOrdered);
+ElementSet (IntMap);
+
+structure IntPairSet =
+ElementSet (IntPairMap);
structure StringSet =
-ElementSet (StringOrdered);
-
- end;
-
-(**** Original file: Map.sig ****)
-
-(* ========================================================================= *)
-(* FINITE MAPS *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
-
-signature Map =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* Finite maps *)
-(* ------------------------------------------------------------------------- *)
-
-type ('key,'a) map
-
-val new : ('key * 'key -> order) -> ('key,'a) map
-
-val null : ('key,'a) map -> bool
-
-val size : ('key,'a) map -> int
-
-val singleton : ('key * 'key -> order) -> 'key * 'a -> ('key,'a) map
-
-val inDomain : 'key -> ('key,'a) map -> bool
-
-val peek : ('key,'a) map -> 'key -> 'a option
-
-val insert : ('key,'a) map -> 'key * 'a -> ('key,'a) map
-
-val insertList : ('key,'a) map -> ('key * 'a) list -> ('key,'a) map
-
-val get : ('key,'a) map -> 'key -> 'a (* raises Error *)
-
-(* Union and intersect prefer keys in the second map *)
-
-val union :
- ('a * 'a -> 'a option) -> ('key,'a) map -> ('key,'a) map -> ('key,'a) map
-
-val intersect :
- ('a * 'a -> 'a option) -> ('key,'a) map -> ('key,'a) map -> ('key,'a) map
-
-val delete : ('key,'a) map -> 'key -> ('key,'a) map (* raises Error *)
-
-val difference : ('key,'a) map -> ('key,'b) map -> ('key,'a) map
-
-val subsetDomain : ('key,'a) map -> ('key,'a) map -> bool
-
-val equalDomain : ('key,'a) map -> ('key,'a) map -> bool
-
-val mapPartial : ('key * 'a -> 'b option) -> ('key,'a) map -> ('key,'b) map
-
-val filter : ('key * 'a -> bool) -> ('key,'a) map -> ('key,'a) map
-
-val map : ('key * 'a -> 'b) -> ('key,'a) map -> ('key,'b) map
-
-val app : ('key * 'a -> unit) -> ('key,'a) map -> unit
-
-val transform : ('a -> 'b) -> ('key,'a) map -> ('key,'b) map
-
-val foldl : ('key * 'a * 's -> 's) -> 's -> ('key,'a) map -> 's
-
-val foldr : ('key * 'a * 's -> 's) -> 's -> ('key,'a) map -> 's
-
-val findl : ('key * 'a -> bool) -> ('key,'a) map -> ('key * 'a) option
-
-val findr : ('key * 'a -> bool) -> ('key,'a) map -> ('key * 'a) option
-
-val firstl : ('key * 'a -> 'b option) -> ('key,'a) map -> 'b option
-
-val firstr : ('key * 'a -> 'b option) -> ('key,'a) map -> 'b option
-
-val exists : ('key * 'a -> bool) -> ('key,'a) map -> bool
-
-val all : ('key * 'a -> bool) -> ('key,'a) map -> bool
-
-val domain : ('key,'a) map -> 'key list
-
-val toList : ('key,'a) map -> ('key * 'a) list
-
-val fromList : ('key * 'key -> order) -> ('key * 'a) list -> ('key,'a) map
-
-val random : ('key,'a) map -> 'key * 'a (* raises Empty *)
-
-val compare : ('a * 'a -> order) -> ('key,'a) map * ('key,'a) map -> order
-
-val equal : ('a -> 'a -> bool) -> ('key,'a) map -> ('key,'a) map -> bool
-
-val toString : ('key,'a) map -> string
-
-(* ------------------------------------------------------------------------- *)
-(* Iterators over maps *)
-(* ------------------------------------------------------------------------- *)
-
-type ('key,'a) iterator
-
-val mkIterator : ('key,'a) map -> ('key,'a) iterator option
-
-val mkRevIterator : ('key,'a) map -> ('key,'a) iterator option
-
-val readIterator : ('key,'a) iterator -> 'key * 'a
-
-val advanceIterator : ('key,'a) iterator -> ('key,'a) iterator option
-
-end
-
-(**** Original file: RandomMap.sml ****)
-
-structure Metis = stru