# HG changeset patch # User blanchet # Date 1284405850 -7200 # Node ID 7f11d833d65b0e5976eeda434dc9345e4b684bd6 # Parent 41ce0b56d85808b6363c7e7a0fc88bf797d0dadc# Parent 7d850b17a7a6b1c1ad0cb09db3d7dc4f83a4e70d merged diff -r 41ce0b56d858 -r 7f11d833d65b doc-src/Nitpick/nitpick.tex --- 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 diff -r 41ce0b56d858 -r 7f11d833d65b doc-src/Sledgehammer/sledgehammer.tex --- 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'' diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/HOL.thy --- 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 *} diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/IsaMakefile --- 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 \ diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.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; diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Mutabelle/MutabelleExtra.thy --- 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Mutabelle/mutabelle_extra.ML --- 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) diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Sledgehammer.thy --- 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/ATP/atp_systems.ML --- 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Nitpick/kodkod.ML --- 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))) | _ => ()); diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Nitpick/nitpick.ML --- 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Nitpick/nitpick_hol.ML --- 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, diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Nitpick/nitpick_isar.ML --- 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; diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Nitpick/nitpick_kodkod.ML --- 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Nitpick/nitpick_model.ML --- 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Nitpick/nitpick_mono.ML --- 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 () => " \ \ " ^ 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) diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Nitpick/nitpick_nut.ML --- 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 = diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Nitpick/nitpick_preproc.ML --- 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 => diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Nitpick/nitpick_rep.ML --- 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) diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Sledgehammer/metis_clauses.ML --- 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), diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Sledgehammer/metis_tactics.ML --- 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); diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Sledgehammer/sledgehammer.ML --- 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML --- 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; diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML --- 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\" \ diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML --- 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/Sledgehammer/sledgehammer_util.ML --- 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/meson.ML --- 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/HOL/Tools/try.ML --- 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; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Code_Generator.thy --- 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/make-metis --- 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/metis.ML --- 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 = 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 MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES *) -(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) -(* ========================================================================= *) - -structure RandomMap :> Map = -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,'b) tree = - E - | T of - {size : int, - priority : priority, - left : ('a,'b) tree, - key : 'a, - value : 'b, - right : ('a,'b) tree}; - -type ('a,'b) node = - {size : int, - priority : priority, - left : ('a,'b) tree, - key : 'a, - value : 'b, - right : ('a,'b) tree}; - -datatype ('a,'b) map = Map of ('a * 'a -> order) * ('a,'b) tree; - -(* ------------------------------------------------------------------------- *) -(* Random priorities. *) -(* ------------------------------------------------------------------------- *) - -local - val randomPriority = randomWord; - - val priorityOrder = Word.compare; -in - fun treeSingleton (key,value) = - T {size = 1, priority = randomPriority (), - left = E, key = key, value = value, right = E}; - - fun nodePriorityOrder cmp (x1 : ('a,'b) node, x2 : ('a,'b) 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 (m as Map (cmp,tree)) = - (let - val _ = checkSizes tree - val _ = checkSorted cmp NONE tree - val _ = checkPriorities cmp tree - in - m - end - handle Error err => raise Bug err) - handle Bug bug => raise Bug (s ^ "\nRandomMap.checkWellformed: " ^ bug); -end; - -(* ------------------------------------------------------------------------- *) -(* Basic operations. *) -(* ------------------------------------------------------------------------- *) - -fun comparison (Map (cmp,_)) = cmp; - -fun new cmp = Map (cmp,E); - -fun treeSize E = 0 - | treeSize (T {size = s, ...}) = s; - -fun size (Map (_,tree)) = treeSize tree; - -fun mkT p l k v r = - T {size = treeSize l + 1 + treeSize r, priority = p, - left = l, key = k, value = v, right = r}; - -fun singleton cmp key_value = Map (cmp, treeSingleton key_value); - -local - fun treePeek cmp E pkey = NONE - | treePeek cmp (T {left,key,value,right,...}) pkey = - case cmp (pkey,key) of - LESS => treePeek cmp left pkey - | EQUAL => SOME value - | GREATER => treePeek cmp right pkey -in - fun peek (Map (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, value = v2, right = r2, ...} = x2 - in - mkT p2 (treeAppend cmp t1 l2) k2 v2 r2 - end - | EQUAL => raise Bug "RandomSet.treeAppend: equal keys" - | GREATER => - let - val {priority = p1, - left = l1, key = k1, value = v1, right = r1, ...} = x1 - in - mkT p1 l1 k1 v1 (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,value,...} : ('a,'b) node) :: xs) t = - mkLeft xs (mkT priority left key value t); - - fun mkRight [] t = t - | mkRight (({priority,key,value,right,...} : ('a,'b) node) :: xs) t = - mkRight xs (mkT priority t key value 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,value,right,...}) = - case cmp (pkey,key) of - LESS => treePart cmp pkey lefts (x :: rights) left - | EQUAL => (mkLeft lefts left, SOME (key,value), 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 f (t1 as T x1) (t2 as T x2) = - let - val {priority = p1, - left = l1, key = k1, value = v1, right = r1, ...} = x1 - val (l2,k2_v2,r2) = nodePartition cmp x2 k1 - val (l1,l2) = treeCombineRemove cmp f l1 l2 - and (r1,r2) = treeCombineRemove cmp f r1 r2 - in - case k2_v2 of - NONE => - if treeSize l2 + treeSize r2 = #size x2 then (t1,t2) - else (mkT p1 l1 k1 v1 r1, treeAppend cmp l2 r2) - | SOME (k2,v2) => - case f (v1,v2) of - NONE => (treeAppend cmp l1 r1, treeAppend cmp l2 r2) - | SOME v => (mkT p1 l1 k2 v 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, value = v1, 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 v1 r - end; -in - fun union f (m1 as Map (cmp,t1)) (Map (_,t2)) = - if pointerEqual (t1,t2) then m1 - else - let - val (t1,t2) = treeCombineRemove cmp f t1 t2 - in - Map (cmp, treeUnionDisjoint cmp t1 t2) - end; -end; - -(*DEBUG -val union = fn f => fn t1 => fn t2 => - checkWellformed "RandomMap.union: result" - (union f (checkWellformed "RandomMap.union: input 1" t1) - (checkWellformed "RandomMap.union: input 2" t2)); -*) - -(* intersect is a simple case of the union algorithm. *) - -local - fun treeIntersect _ _ _ E = E - | treeIntersect _ _ E _ = E - | treeIntersect cmp f (t1 as T x1) (t2 as T x2) = - let - val {priority = p1, - left = l1, key = k1, value = v1, right = r1, ...} = x1 - val (l2,k2_v2,r2) = nodePartition cmp x2 k1 - val l = treeIntersect cmp f l1 l2 - and r = treeIntersect cmp f r1 r2 - in - case k2_v2 of - NONE => treeAppend cmp l r - | SOME (k2,v2) => - case f (v1,v2) of - NONE => treeAppend cmp l r - | SOME v => mkT p1 l k2 v r - end; -in - fun intersect f (m1 as Map (cmp,t1)) (Map (_,t2)) = - if pointerEqual (t1,t2) then m1 - else Map (cmp, treeIntersect cmp f t1 t2); -end; - -(*DEBUG -val intersect = fn f => fn t1 => fn t2 => - checkWellformed "RandomMap.intersect: result" - (intersect f (checkWellformed "RandomMap.intersect: input 1" t1) - (checkWellformed "RandomMap.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 "RandomMap.delete: element not found" - | treeDelete cmp (T {priority,left,key,value,right,...}) dkey = - case cmp (dkey,key) of - LESS => mkT priority (treeDelete cmp left dkey) key value right - | EQUAL => treeAppend cmp left right - | GREATER => mkT priority left key value (treeDelete cmp right dkey); -in - fun delete (Map (cmp,tree)) key = Map (cmp, treeDelete cmp tree key); -end; - -(*DEBUG -val delete = fn t => fn x => - checkWellformed "RandomMap.delete: result" - (delete (checkWellformed "RandomMap.delete: input" t) x); -*) - -(* Set difference on domains *) - -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, value = v1, right = r1} = x1 - val (l2,k2_v2,r2) = nodePartition cmp x2 k1 - val l = treeDifference cmp l1 l2 - and r = treeDifference cmp r1 r2 - in - if Option.isSome k2_v2 then treeAppend cmp l r - else if treeSize l + treeSize r + 1 = s1 then t1 - else mkT p1 l k1 v1 r - end; -in - fun difference (Map (cmp,tree1)) (Map (_,tree2)) = - Map (cmp, treeDifference cmp tree1 tree2); -end; - -(*DEBUG -val difference = fn t1 => fn t2 => - checkWellformed "RandomMap.difference: result" - (difference (checkWellformed "RandomMap.difference: input 1" t1) - (checkWellformed "RandomMap.difference: input 2" t2)); -*) - -(* subsetDomain is mainly used when using maps as sets. *) - -local - fun treeSubsetDomain _ E _ = true - | treeSubsetDomain _ _ E = false - | treeSubsetDomain 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_v2,r2) = nodePartition cmp x2 k1 - in - Option.isSome k2_v2 andalso - treeSubsetDomain cmp l1 l2 andalso - treeSubsetDomain cmp r1 r2 - end - end; -in - fun subsetDomain (Map (cmp,tree1)) (Map (_,tree2)) = - pointerEqual (tree1,tree2) orelse - treeSubsetDomain cmp tree1 tree2; -end; - -(* Map equality *) - -local - fun treeEqual _ _ E E = true - | treeEqual _ _ E _ = false - | treeEqual _ _ _ E = false - | treeEqual cmp veq (t1 as T x1) (T x2) = - let - val {size = s1, left = l1, key = k1, value = v1, right = r1, ...} = x1 - and {size = s2, ...} = x2 - in - s1 = s2 andalso - let - val (l2,k2_v2,r2) = nodePartition cmp x2 k1 - in - (case k2_v2 of NONE => false | SOME (_,v2) => veq v1 v2) andalso - treeEqual cmp veq l1 l2 andalso - treeEqual cmp veq r1 r2 - end - end; -in - fun equal veq (Map (cmp,tree1)) (Map (_,tree2)) = - pointerEqual (tree1,tree2) orelse - treeEqual cmp veq tree1 tree2; -end; - -(* mapPartial is the basic function for preserving the tree structure. *) -(* It applies the argument function to the elements *in order*. *) - -local - fun treeMapPartial cmp _ E = E - | treeMapPartial cmp f (T {priority,left,key,value,right,...}) = - let - val left = treeMapPartial cmp f left - and value' = f (key,value) - and right = treeMapPartial cmp f right - in - case value' of - NONE => treeAppend cmp left right - | SOME value => mkT priority left key value right - end; -in - fun mapPartial f (Map (cmp,tree)) = Map (cmp, treeMapPartial cmp f tree); -end; - -(* map is a primitive function for efficiency reasons. *) -(* It also applies the argument function to the elements *in order*. *) - -local - fun treeMap _ E = E - | treeMap f (T {size,priority,left,key,value,right}) = - let - val left = treeMap f left - and value = f (key,value) - and right = treeMap f right - in - T {size = size, priority = priority, left = left, - key = key, value = value, right = right} - end; -in - fun map f (Map (cmp,tree)) = Map (cmp, treeMap f tree); -end; - -(* nth picks the nth smallest key/value (counting from 0). *) - -local - fun treeNth E _ = raise Subscript - | treeNth (T {left,key,value,right,...}) n = - let - val k = treeSize left - in - if n = k then (key,value) - else if n < k then treeNth left n - else treeNth right (n - (k + 1)) - end; -in - fun nth (Map (_,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 ('key,'a) iterator = - LR of ('key * 'a) * ('key,'a) tree * ('key,'a) tree list - | RL of ('key * 'a) * ('key,'a) tree * ('key,'a) tree list; - -fun mkLR [] = NONE - | mkLR (T {key,value,right,...} :: l) = SOME (LR ((key,value),right,l)) - | mkLR (E :: _) = raise Bug "RandomMap.mkLR"; - -fun mkRL [] = NONE - | mkRL (T {key,value,left,...} :: l) = SOME (RL ((key,value),left,l)) - | mkRL (E :: _) = raise Bug "RandomMap.mkRL"; - -fun mkIterator (Map (_,tree)) = mkLR (leftSpine tree []); - -fun mkRevIterator (Map (_,tree)) = mkRL (rightSpine tree []); - -fun readIterator (LR (key_value,_,_)) = key_value - | readIterator (RL (key_value,_,_)) = key_value; - -fun advanceIterator (LR (_,next,l)) = mkLR (leftSpine next l) - | advanceIterator (RL (_,next,l)) = mkRL (rightSpine next l); - -(* ------------------------------------------------------------------------- *) -(* Derived operations. *) -(* ------------------------------------------------------------------------- *) - -fun null m = size m = 0; - -fun get m key = - case peek m key of - NONE => raise Error "RandomMap.get: element not found" - | SOME value => value; - -fun inDomain key m = Option.isSome (peek m key); - -fun insert m key_value = - union (SOME o snd) m (singleton (comparison m) key_value); - -(*DEBUG -val insert = fn m => fn x => - checkWellformed "RandomMap.insert: result" - (insert (checkWellformed "RandomMap.insert: input" m) x); -*) - -local - fun fold _ NONE acc = acc - | fold f (SOME iter) acc = - let - val (key,value) = readIterator iter - in - fold f (advanceIterator iter) (f (key,value,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_value = readIterator iter - in - if pred key_value then SOME key_value - 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_value = readIterator iter - in - case f key_value 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 fromList cmp l = List.foldl (fn (k_v,m) => insert m k_v) (new cmp) l; - -fun insertList m l = union (SOME o snd) m (fromList (comparison m) l); - -fun filter p = - let - fun f (key_value as (_,value)) = - if p key_value then SOME value else NONE - in - mapPartial f - end; - -fun app f m = foldl (fn (key,value,()) => f (key,value)) () m; - -fun transform f = map (fn (_,value) => f value); - -fun toList m = foldr (fn (key,value,l) => (key,value) :: l) [] m; - -fun domain m = foldr (fn (key,_,l) => key :: l) [] m; - -fun exists p m = Option.isSome (findl p m); - -fun all p m = not (exists (not o p) m); - -fun random m = - case size m of - 0 => raise Error "RandomMap.random: empty" - | n => nth m (randomInt n); - -local - fun iterCompare _ _ NONE NONE = EQUAL - | iterCompare _ _ NONE (SOME _) = LESS - | iterCompare _ _ (SOME _) NONE = GREATER - | iterCompare kcmp vcmp (SOME i1) (SOME i2) = - keyIterCompare kcmp vcmp (readIterator i1) (readIterator i2) i1 i2 - - and keyIterCompare kcmp vcmp (k1,v1) (k2,v2) i1 i2 = - case kcmp (k1,k2) of - LESS => LESS - | EQUAL => - (case vcmp (v1,v2) of - LESS => LESS - | EQUAL => - iterCompare kcmp vcmp (advanceIterator i1) (advanceIterator i2) - | GREATER => GREATER) - | GREATER => GREATER; -in - fun compare vcmp (m1,m2) = - if pointerEqual (m1,m2) then EQUAL - else - case Int.compare (size m1, size m2) of - LESS => LESS - | EQUAL => - iterCompare (comparison m1) vcmp (mkIterator m1) (mkIterator m2) - | GREATER => GREATER; -end; - -fun equalDomain m1 m2 = equal (K (K true)) m1 m2; - -fun toString m = "<" ^ (if null m then "" else Int.toString (size m)) ^ ">"; - -end -end; - -(**** Original file: Map.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 MAPS *) -(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) -(* ========================================================================= *) - -structure Map = RandomMap; -end; - -(**** Original file: KeyMap.sig ****) - -(* ========================================================================= *) -(* FINITE MAPS WITH A FIXED KEY TYPE *) -(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) -(* ========================================================================= *) - -signature KeyMap = -sig - -type key - -(* ------------------------------------------------------------------------- *) -(* Finite maps *) -(* ------------------------------------------------------------------------- *) - -type 'a map - -val new : unit -> 'a map - -val null : 'a map -> bool - -val size : 'a map -> int - -val singleton : key * 'a -> 'a map - -val inDomain : key -> 'a map -> bool - -val peek : 'a map -> key -> 'a option - -val insert : 'a map -> key * 'a -> 'a map - -val insertList : 'a map -> (key * 'a) list -> 'a map - -val get : 'a map -> key -> 'a (* raises Error *) - -(* Union and intersect prefer keys in the second map *) - -val union : ('a * 'a -> 'a option) -> 'a map -> 'a map -> 'a map - -val intersect : ('a * 'a -> 'a option) -> 'a map -> 'a map -> 'a map - -val delete : 'a map -> key -> 'a map (* raises Error *) - -val difference : 'a map -> 'a map -> 'a map - -val subsetDomain : 'a map -> 'a map -> bool - -val equalDomain : 'a map -> 'a map -> bool - -val mapPartial : (key * 'a -> 'b option) -> 'a map -> 'b map - -val filter : (key * 'a -> bool) -> 'a map -> 'a 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 foldl : (key * 'a * 's -> 's) -> 's -> 'a map -> 's - -val foldr : (key * 'a * 's -> 's) -> 's -> 'a map -> 's - -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 domain : 'a map -> key list - -val toList : 'a map -> (key * 'a) list - -val fromList : (key * 'a) list -> 'a map - -val compare : ('a * 'a -> order) -> 'a map * 'a map -> order - -val equal : ('a -> 'a -> bool) -> 'a map -> 'a map -> bool - -val random : 'a map -> key * 'a (* raises Empty *) - -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-2006 Joe Hurd, distributed under the BSD License *) -(* ========================================================================= *) - -functor KeyMap (Key : Ordered) :> KeyMap where type key = Key.t = -struct - - open Metis; - -type key = Key.t; - -(* ------------------------------------------------------------------------- *) -(* Finite maps *) -(* ------------------------------------------------------------------------- *) - -type 'a map = (Key.t,'a) Map.map; - -fun new () = Map.new Key.compare; - -val null = Map.null; - -val size = Map.size; - -fun singleton key_value = Map.singleton Key.compare key_value; - -val inDomain = Map.inDomain; - -val peek = Map.peek; - -val insert = Map.insert; - -val insertList = Map.insertList; - -val get = Map.get; - -(* Both op union and intersect prefer keys in the second map *) - -val op union = Map.union; - -val intersect = Map.intersect; - -val delete = Map.delete; - -val difference = Map.difference; - -val subsetDomain = Map.subsetDomain; - -val equalDomain = Map.equalDomain; - -val mapPartial = Map.mapPartial; - -val filter = Map.filter; - -val map = Map.map; - -val app = Map.app; - -val transform = Map.transform; - -val foldl = Map.foldl; - -val foldr = Map.foldr; - -val findl = Map.findl; - -val findr = Map.findr; - -val firstl = Map.firstl; - -val firstr = Map.firstr; - -val exists = Map.exists; - -val all = Map.all; - -val domain = Map.domain; - -val toList = Map.toList; - -fun fromList l = Map.fromList Key.compare l; - -val compare = Map.compare; - -val equal = Map.equal; - -val random = Map.random; - -val toString = Map.toString; - -(* ------------------------------------------------------------------------- *) -(* Iterators over maps *) -(* ------------------------------------------------------------------------- *) - -type 'a iterator = (Key.t,'a) Map.iterator; - -val mkIterator = Map.mkIterator; - -val mkRevIterator = Map.mkRevIterator; - -val readIterator = Map.readIterator; - -val advanceIterator = Map.advanceIterator; - -end - - structure Metis = struct open Metis - -structure IntMap = -KeyMap (IntOrdered); - -structure StringMap = -KeyMap (StringOrdered); - - end; +ElementSet (StringMap); (**** Original file: Sharing.sig ****) (* ========================================================================= *) (* PRESERVING SHARING OF ML VALUES *) -(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Sharing = sig (* ------------------------------------------------------------------------- *) -(* Pointer equality. *) -(* ------------------------------------------------------------------------- *) - -val pointerEqual : 'a * 'a -> bool +(* Option operations. *) +(* ------------------------------------------------------------------------- *) + +val mapOption : ('a -> 'a) -> 'a option -> 'a option + +val mapsOption : ('a -> 's -> 'a * 's) -> 'a option -> 's -> 'a option * 's (* ------------------------------------------------------------------------- *) (* List operations. *) @@ -4084,6 +6279,12 @@ val map : ('a -> 'a) -> 'a list -> 'a list +val revMap : ('a -> 'a) -> 'a list -> 'a list + +val maps : ('a -> 's -> 'a * 's) -> 'a list -> 's -> 'a list * 's + +val revMaps : ('a -> 's -> 'a * 's) -> 'a list -> 's -> 'a list * 's + val updateNth : int * 'a -> 'a list -> 'a list val setify : ''a list -> ''a list @@ -4106,7 +6307,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; @@ -4115,7 +6316,7 @@ (* ========================================================================= *) (* PRESERVING SHARING OF ML VALUES *) -(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Sharing :> Sharing = @@ -4123,13 +6324,31 @@ infix == -(* ------------------------------------------------------------------------- *) -(* Pointer equality. *) -(* ------------------------------------------------------------------------- *) - -val pointerEqual = Portable.pointerEqual; - -val op== = pointerEqual; +val op== = Portable.pointerEqual; + +(* ------------------------------------------------------------------------- *) +(* Option operations. *) +(* ------------------------------------------------------------------------- *) + +fun mapOption f xo = + case xo of + SOME x => + let + val y = f x + in + if x == y then xo else SOME y + end + | NONE => xo; + +fun mapsOption f xo acc = + case xo of + SOME x => + let + val (y,acc) = f x acc + in + if x == y then (xo,acc) else (SOME y, acc) + end + | NONE => (xo,acc); (* ------------------------------------------------------------------------- *) (* List operations. *) @@ -4137,17 +6356,78 @@ fun map f = let - fun m _ a_b [] = List.revAppend a_b - | m ys a_b (x :: xs) = - let - val y = f x - val ys = y :: ys - in - m ys (if x == y then a_b else (ys,xs)) xs - end - in - fn l => m [] ([],l) l - end; + fun m ys ys_xs xs = + case xs of + [] => List.revAppend ys_xs + | x :: xs => + let + val y = f x + val ys = y :: ys + val ys_xs = if x == y then ys_xs else (ys,xs) + in + m ys ys_xs xs + end + in + fn xs => m [] ([],xs) xs + end; + +fun maps f = + let + fun m acc ys ys_xs xs = + case xs of + [] => (List.revAppend ys_xs, acc) + | x :: xs => + let + val (y,acc) = f x acc + val ys = y :: ys + val ys_xs = if x == y then ys_xs else (ys,xs) + in + m acc ys ys_xs xs + end + in + fn xs => fn acc => m acc [] ([],xs) xs + end; + +local + fun revTails acc xs = + case xs of + [] => acc + | x :: xs' => revTails ((x,xs) :: acc) xs'; +in + fun revMap f = + let + fun m ys same xxss = + case xxss of + [] => ys + | (x,xs) :: xxss => + let + val y = f x + val same = same andalso x == y + val ys = if same then xs else y :: ys + in + m ys same xxss + end + in + fn xs => m [] true (revTails [] xs) + end; + + fun revMaps f = + let + fun m acc ys same xxss = + case xxss of + [] => (ys,acc) + | (x,xs) :: xxss => + let + val (y,acc) = f x acc + val same = same andalso x == y + val ys = if same then xs else y :: ys + in + m acc ys same xxss + end + in + fn xs => fn acc => m acc [] true (revTails [] xs) + end; +end; fun updateNth (n,x) l = let @@ -4194,329 +6474,11 @@ 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 -> '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 -> 'a stream -> 'b stream - -(* ------------------------------------------------------------------------- *) -(* Stream operations *) -(* ------------------------------------------------------------------------- *) - -val memoize : 'a stream -> '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 mem; -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 datatype declaration encapsulates all the primitive operations *) -(* ------------------------------------------------------------------------- *) - -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, fn () => m (t ())) - in - m - end; - -fun maps f = - let - fun mm _ NIL = NIL - | mm s (CONS (x, xs)) = - let - val (y, s') = f x s - in - CONS (y, fn () => mm s' (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 = - let - fun mm _ NIL = NIL - | mm s (CONS (x, xs)) = - let - val (yo, s') = f x s - val t = mm s' o xs - in - case yo of NONE => t () | SOME y => CONS (y, t) - end - in - mm - end; - -(* ------------------------------------------------------------------------- *) -(* Stream operations *) -(* ------------------------------------------------------------------------- *) - -fun memoize NIL = NIL - | memoize (CONS (h,t)) = CONS (h, Lazy.memoize (fn () => memoize (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 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: Heap.sig ****) (* ========================================================================= *) (* A HEAP DATATYPE FOR ML *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Heap = @@ -4550,7 +6512,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; @@ -4559,7 +6521,7 @@ (* ========================================================================= *) (* A HEAP DATATYPE FOR ML *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Heap :> Heap = @@ -4622,12 +6584,12 @@ end; fun toStream h = - if null h then Stream.NIL + if null h then Stream.Nil else let val (x,h) = remove h in - Stream.CONS (x, fn () => toStream h) + Stream.Cons (x, fn () => toStream h) end; fun toString h = @@ -4636,89 +6598,1345 @@ end end; -(**** Original file: Parser.sig ****) - -(* ========================================================================= *) -(* PARSING AND PRETTY PRINTING *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) -(* ========================================================================= *) - -signature Parser = +(**** Original file: Print.sig ****) + +(* ========================================================================= *) +(* PRETTY-PRINTING *) +(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *) +(* ========================================================================= *) + +signature Print = sig (* ------------------------------------------------------------------------- *) -(* Pretty printing for built-in types *) -(* ------------------------------------------------------------------------- *) - -type ppstream = Metis.PP.ppstream +(* A type of pretty-printers. *) +(* ------------------------------------------------------------------------- *) datatype breakStyle = Consistent | Inconsistent -type 'a pp = ppstream -> 'a -> unit - -val lineLength : int Unsynchronized.ref - -val beginBlock : ppstream -> breakStyle -> int -> unit - -val endBlock : ppstream -> unit - -val addString : ppstream -> string -> unit - -val addBreak : ppstream -> int * int -> unit - -val addNewline : ppstream -> unit +datatype ppStep = + BeginBlock of breakStyle * int + | EndBlock + | AddString of string + | AddBreak of int + | AddNewline + +type ppstream = ppStep Metis.Stream.stream + +type 'a pp = 'a -> ppstream + +(* ------------------------------------------------------------------------- *) +(* Pretty-printer primitives. *) +(* ------------------------------------------------------------------------- *) + +val beginBlock : breakStyle -> int -> ppstream + +val endBlock : ppstream + +val addString : string -> ppstream + +val addBreak : int -> ppstream + +val addNewline : ppstream + +val skip : ppstream + +val sequence : ppstream -> ppstream -> ppstream + +val duplicate : int -> ppstream -> ppstream + +val program : ppstream list -> ppstream + +val stream : ppstream Metis.Stream.stream -> ppstream + +val block : breakStyle -> int -> ppstream -> ppstream + +val blockProgram : breakStyle -> int -> ppstream list -> ppstream + +val bracket : string -> string -> ppstream -> ppstream + +val field : string -> ppstream -> ppstream + +val record : (string * ppstream) list -> ppstream + +(* ------------------------------------------------------------------------- *) +(* Pretty-printer combinators. *) +(* ------------------------------------------------------------------------- *) val ppMap : ('a -> 'b) -> 'b pp -> 'a pp val ppBracket : string -> string -> 'a pp -> 'a pp -val ppSequence : string -> 'a pp -> 'a list pp - -val ppBinop : string -> 'a pp -> 'b pp -> ('a * 'b) pp +val ppOp : string -> ppstream + +val ppOp2 : string -> 'a pp -> 'b pp -> ('a * 'b) pp + +val ppOp3 : string -> string -> 'a pp -> 'b pp -> 'c pp -> ('a * 'b * 'c) pp + +val ppOpList : string -> 'a pp -> 'a list pp + +val ppOpStream : string -> 'a pp -> 'a Metis.Stream.stream pp + +(* ------------------------------------------------------------------------- *) +(* Pretty-printers for common types. *) +(* ------------------------------------------------------------------------- *) val ppChar : char pp val ppString : string pp +val ppEscapeString : {escape : char list} -> string pp + val ppUnit : unit pp val ppBool : bool pp val ppInt : int pp +val ppPrettyInt : int pp + val ppReal : real pp +val ppPercent : real pp + val ppOrder : order pp val ppList : 'a pp -> 'a list pp +val ppStream : 'a pp -> 'a Metis.Stream.stream pp + val ppOption : 'a pp -> 'a option pp val ppPair : 'a pp -> 'b pp -> ('a * 'b) pp val ppTriple : 'a pp -> 'b pp -> 'c pp -> ('a * 'b * 'c) pp -val toString : 'a pp -> 'a -> string (* Uses !lineLength *) - -val fromString : ('a -> string) -> 'a pp - -val ppTrace : 'a pp -> string -> 'a -> unit - -(* ------------------------------------------------------------------------- *) -(* Recursive descent parsing combinators *) -(* ------------------------------------------------------------------------- *) - -(* Generic parsers - -Recommended fixities: +val ppBreakStyle : breakStyle pp + +val ppPpStep : ppStep pp + +val ppPpstream : ppstream pp + +(* ------------------------------------------------------------------------- *) +(* Pretty-printing infix operators. *) +(* ------------------------------------------------------------------------- *) + +datatype infixes = + Infixes of + {token : string, + precedence : int, + leftAssoc : bool} list + +val tokensInfixes : infixes -> StringSet.set (* MODIFIED by Jasmin Blanchette *) + +val layerInfixes : + infixes -> + {tokens : {leftSpaces : int, token : string, rightSpaces : int} list, + leftAssoc : bool} list + +val ppInfixes : + infixes -> ('a -> (string * 'a * 'a) option) -> ('a * bool) pp -> + ('a * bool) pp + +(* ------------------------------------------------------------------------- *) +(* Executing pretty-printers to generate lines. *) +(* ------------------------------------------------------------------------- *) + +val execute : {lineLength : int} -> ppstream -> string Metis.Stream.stream + +(* ------------------------------------------------------------------------- *) +(* Executing pretty-printers with a global line length. *) +(* ------------------------------------------------------------------------- *) + +val lineLength : int Unsynchronized.ref + +val toString : 'a pp -> 'a -> string + +val toStream : 'a pp -> 'a -> string Metis.Stream.stream + +val trace : 'a pp -> string -> 'a -> unit + +end + +(**** Original file: Print.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; + +(* ========================================================================= *) +(* PRETTY-PRINTING *) +(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *) +(* ========================================================================= *) + +structure Print :> Print = +struct + +open Useful; + +(* ------------------------------------------------------------------------- *) +(* Constants. *) +(* ------------------------------------------------------------------------- *) + +val initialLineLength = 75; + +(* ------------------------------------------------------------------------- *) +(* Helper functions. *) +(* ------------------------------------------------------------------------- *) + +fun revAppend xs s = + case xs of + [] => s () + | x :: xs => revAppend xs (K (Stream.Cons (x,s))); + +fun revConcat strm = + case strm of + Stream.Nil => Stream.Nil + | Stream.Cons (h,t) => revAppend h (revConcat o t); + +local + fun join current prev = (prev ^ "\n", current); +in + fun joinNewline strm = + case strm of + Stream.Nil => Stream.Nil + | Stream.Cons (h,t) => Stream.maps join Stream.singleton h (t ()); +end; + +local + fun calcSpaces n = nChars #" " n; + + val cachedSpaces = Vector.tabulate (initialLineLength,calcSpaces); +in + fun nSpaces n = + if n < initialLineLength then Vector.sub (cachedSpaces,n) + else calcSpaces n; +end; + +(* ------------------------------------------------------------------------- *) +(* A type of pretty-printers. *) +(* ------------------------------------------------------------------------- *) + +datatype breakStyle = Consistent | Inconsistent; + +datatype ppStep = + BeginBlock of breakStyle * int + | EndBlock + | AddString of string + | AddBreak of int + | AddNewline; + +type ppstream = ppStep Stream.stream; + +type 'a pp = 'a -> ppstream; + +fun breakStyleToString style = + case style of + Consistent => "Consistent" + | Inconsistent => "Inconsistent"; + +fun ppStepToString step = + case step of + BeginBlock _ => "BeginBlock" + | EndBlock => "EndBlock" + | AddString _ => "AddString" + | AddBreak _ => "AddBreak" + | AddNewline => "AddNewline"; + +(* ------------------------------------------------------------------------- *) +(* Pretty-printer primitives. *) +(* ------------------------------------------------------------------------- *) + +fun beginBlock style indent = Stream.singleton (BeginBlock (style,indent)); + +val endBlock = Stream.singleton EndBlock; + +fun addString s = Stream.singleton (AddString s); + +fun addBreak spaces = Stream.singleton (AddBreak spaces); + +val addNewline = Stream.singleton AddNewline; + +val skip : ppstream = Stream.Nil; + +fun sequence pp1 pp2 : ppstream = Stream.append pp1 (K pp2); + +local + fun dup pp n () = if n = 1 then pp else Stream.append pp (dup pp (n - 1)); +in + fun duplicate n pp = if n = 0 then skip else dup pp n (); +end; + +val program : ppstream list -> ppstream = Stream.concatList; + +val stream : ppstream Stream.stream -> ppstream = Stream.concat; + +fun block style indent pp = program [beginBlock style indent, pp, endBlock]; + +fun blockProgram style indent pps = block style indent (program pps); + +fun bracket l r pp = + blockProgram Inconsistent (size l) + [addString l, + pp, + addString r]; + +fun field f pp = + blockProgram Inconsistent 2 + [addString (f ^ " ="), + addBreak 1, + pp]; + +val record = + let + val sep = sequence (addString ",") (addBreak 1) + + fun recordField (f,pp) = field f pp + + fun sepField f = sequence sep (recordField f) + + fun fields [] = [] + | fields (f :: fs) = recordField f :: map sepField fs + in + bracket "{" "}" o blockProgram Consistent 0 o fields + end; + +(* ------------------------------------------------------------------------- *) +(* Pretty-printer combinators. *) +(* ------------------------------------------------------------------------- *) + +fun ppMap f ppB a : ppstream = ppB (f a); + +fun ppBracket l r ppA a = bracket l r (ppA a); + +fun ppOp s = sequence (if s = "" then skip else addString s) (addBreak 1); + +fun ppOp2 ab ppA ppB (a,b) = + blockProgram Inconsistent 0 + [ppA a, + ppOp ab, + ppB b]; + +fun ppOp3 ab bc ppA ppB ppC (a,b,c) = + blockProgram Inconsistent 0 + [ppA a, + ppOp ab, + ppB b, + ppOp bc, + ppC c]; + +fun ppOpList s ppA = + let + fun ppOpA a = sequence (ppOp s) (ppA a) + in + fn [] => skip + | h :: t => blockProgram Inconsistent 0 (ppA h :: map ppOpA t) + end; + +fun ppOpStream s ppA = + let + fun ppOpA a = sequence (ppOp s) (ppA a) + in + fn Stream.Nil => skip + | Stream.Cons (h,t) => + blockProgram Inconsistent 0 + [ppA h, + Stream.concat (Stream.map ppOpA (t ()))] + end; + +(* ------------------------------------------------------------------------- *) +(* Pretty-printers for common types. *) +(* ------------------------------------------------------------------------- *) + +fun ppChar c = addString (str c); + +val ppString = addString; + +fun ppEscapeString {escape} = + let + val escapeMap = map (fn c => (c, "\\" ^ str c)) escape + + fun escapeChar c = + case c of + #"\\" => "\\\\" + | #"\n" => "\\n" + | #"\t" => "\\t" + | _ => + case List.find (equal c o fst) escapeMap of + SOME (_,s) => s + | NONE => str c + in + fn s => addString (String.translate escapeChar s) + end; + +val ppUnit : unit pp = K (addString "()"); + +fun ppBool b = addString (if b then "true" else "false"); + +fun ppInt i = addString (Int.toString i); + +local + val ppNeg = addString "~" + and ppSep = addString "," + and ppZero = addString "0" + and ppZeroZero = addString "00"; + + fun ppIntBlock i = + if i < 10 then sequence ppZeroZero (ppInt i) + else if i < 100 then sequence ppZero (ppInt i) + else ppInt i; + + fun ppIntBlocks i = + if i < 1000 then ppInt i + else sequence (ppIntBlocks (i div 1000)) + (sequence ppSep (ppIntBlock (i mod 1000))); +in + fun ppPrettyInt i = + if i < 0 then sequence ppNeg (ppIntBlocks (~i)) + else ppIntBlocks i; +end; + +fun ppReal r = addString (Real.toString r); + +fun ppPercent p = addString (percentToString p); + +fun ppOrder x = + addString + (case x of + LESS => "Less" + | EQUAL => "Equal" + | GREATER => "Greater"); + +fun ppList ppA = ppBracket "[" "]" (ppOpList "," ppA); + +fun ppStream ppA = ppBracket "[" "]" (ppOpStream "," ppA); + +fun ppOption ppA ao = + case ao of + SOME a => ppA a + | NONE => addString "-"; + +fun ppPair ppA ppB = ppBracket "(" ")" (ppOp2 "," ppA ppB); + +fun ppTriple ppA ppB ppC = ppBracket "(" ")" (ppOp3 "," "," ppA ppB ppC); + +fun ppBreakStyle style = addString (breakStyleToString style); + +fun ppPpStep step = + let + val cmd = ppStepToString step + in + blockProgram Inconsistent 2 + (addString cmd :: + (case step of + BeginBlock style_indent => + [addBreak 1, + ppPair ppBreakStyle ppInt style_indent] + | EndBlock => [] + | AddString s => + [addBreak 1, + addString ("\"" ^ s ^ "\"")] + | AddBreak n => + [addBreak 1, + ppInt n] + | AddNewline => [])) + end; + +val ppPpstream = ppStream ppPpStep; + +(* ------------------------------------------------------------------------- *) +(* Pretty-printing infix operators. *) +(* ------------------------------------------------------------------------- *) + +datatype infixes = + Infixes of + {token : string, + precedence : int, + leftAssoc : bool} list; + +local + fun chop l = + case l of + #" " :: l => let val (n,l) = chop l in (n + 1, l) end + | _ => (0,l); +in + fun opSpaces tok = + let + val tok = explode tok + val (r,tok) = chop (rev tok) + val (l,tok) = chop (rev tok) + val tok = implode tok + in + {leftSpaces = l, token = tok, rightSpaces = r} + end; +end; + +fun ppOpSpaces {leftSpaces,token,rightSpaces} = + let + val leftSpacesToken = + if leftSpaces = 0 then token else nSpaces leftSpaces ^ token + in + sequence + (addString leftSpacesToken) + (addBreak rightSpaces) + end; + +local + fun new t l acc = {tokens = [opSpaces t], leftAssoc = l} :: acc; + + fun add t l acc = + case acc of + [] => raise Bug "Print.layerInfixOps.layer" + | {tokens = ts, leftAssoc = l'} :: acc => + if l = l' then {tokens = opSpaces t :: ts, leftAssoc = l} :: acc + else raise Bug "Print.layerInfixOps: mixed assocs"; + + fun layer ({token = t, precedence = p, leftAssoc = l}, (acc,p')) = + let + val acc = if p = p' then add t l acc else new t l acc + in + (acc,p) + end; +in + fun layerInfixes (Infixes i) = + case sortMap #precedence Int.compare i of + [] => [] + | {token = t, precedence = p, leftAssoc = l} :: i => + let + val acc = new t l [] + + val (acc,_) = List.foldl layer (acc,p) i + in + acc + end; +end; + +val tokensLayeredInfixes = + let + fun addToken ({leftSpaces = _, token = t, rightSpaces = _}, s) = + StringSet.add s t + + fun addTokens ({tokens = t, leftAssoc = _}, s) = + List.foldl addToken s t + in + List.foldl addTokens StringSet.empty + end; + +val tokensInfixes = tokensLayeredInfixes o layerInfixes; + +local + val mkTokenMap = + let + fun add (token,m) = + let + val {leftSpaces = _, token = t, rightSpaces = _} = token + in + StringMap.insert m (t, ppOpSpaces token) + end + in + List.foldl add (StringMap.new ()) + end; +in + fun ppGenInfix {tokens,leftAssoc} = + let + val tokenMap = mkTokenMap tokens + in + fn dest => fn ppSub => + let + fun dest' tm = + case dest tm of + NONE => NONE + | SOME (t,a,b) => + case StringMap.peek tokenMap t of + NONE => NONE + | SOME p => SOME (p,a,b) + + fun ppGo (tmr as (tm,r)) = + case dest' tm of + NONE => ppSub tmr + | SOME (p,a,b) => + program + [(if leftAssoc then ppGo else ppSub) (a,true), + p, + (if leftAssoc then ppSub else ppGo) (b,r)] + in + fn tmr as (tm,_) => + if Option.isSome (dest' tm) then + block Inconsistent 0 (ppGo tmr) + else + ppSub tmr + end + end; +end + +fun ppInfixes ops = + let + val layeredOps = layerInfixes ops + + val toks = tokensLayeredInfixes layeredOps + + val iprinters = List.map ppGenInfix layeredOps + in + fn dest => fn ppSub => + let + fun printer sub = foldl (fn (ip,p) => ip dest p) sub iprinters + + fun isOp t = + case dest t of + SOME (x,_,_) => StringSet.member x toks + | NONE => false + + fun subpr (tmr as (tm,_)) = + if isOp tm then + blockProgram Inconsistent 1 + [addString "(", + printer subpr (tm,false), + addString ")"] + else + ppSub tmr + in + fn tmr => block Inconsistent 0 (printer subpr tmr) + end + end; + +(* ------------------------------------------------------------------------- *) +(* Executing pretty-printers to generate lines. *) +(* ------------------------------------------------------------------------- *) + +datatype blockBreakStyle = + InconsistentBlock + | ConsistentBlock + | BreakingBlock; + +datatype block = + Block of + {indent : int, + style : blockBreakStyle, + size : int, + chunks : chunk list} + +and chunk = + StringChunk of {size : int, string : string} + | BreakChunk of int + | BlockChunk of block; + +datatype state = + State of + {blocks : block list, + lineIndent : int, + lineSize : int}; + +val initialIndent = 0; + +val initialStyle = Inconsistent; + +fun liftStyle style = + case style of + Inconsistent => InconsistentBlock + | Consistent => ConsistentBlock; + +fun breakStyle style = + case style of + ConsistentBlock => BreakingBlock + | _ => style; + +fun sizeBlock (Block {size,...}) = size; + +fun sizeChunk chunk = + case chunk of + StringChunk {size,...} => size + | BreakChunk spaces => spaces + | BlockChunk block => sizeBlock block; + +val splitChunks = + let + fun split _ [] = NONE + | split acc (chunk :: chunks) = + case chunk of + BreakChunk _ => SOME (rev acc, chunks) + | _ => split (chunk :: acc) chunks + in + split [] + end; + +val sizeChunks = List.foldl (fn (c,z) => sizeChunk c + z) 0; + +local + fun render acc [] = acc + | render acc (chunk :: chunks) = + case chunk of + StringChunk {string = s, ...} => render (acc ^ s) chunks + | BreakChunk n => render (acc ^ nSpaces n) chunks + | BlockChunk (Block {chunks = l, ...}) => + render acc (List.revAppend (l,chunks)); +in + fun renderChunks indent chunks = render (nSpaces indent) (rev chunks); + + fun renderChunk indent chunk = renderChunks indent [chunk]; +end; + +fun isEmptyBlock block = + let + val Block {indent = _, style = _, size, chunks} = block + + val empty = null chunks + +(*BasicDebug + val _ = not empty orelse size = 0 orelse + raise Bug "Print.isEmptyBlock: bad size" +*) + in + empty + end; + +fun checkBlock ind block = + let + val Block {indent, style = _, size, chunks} = block + val _ = indent >= ind orelse raise Bug "Print.checkBlock: bad indents" + val size' = checkChunks indent chunks + val _ = size = size' orelse raise Bug "Print.checkBlock: wrong size" + in + size + end + +and checkChunks ind chunks = + case chunks of + [] => 0 + | chunk :: chunks => checkChunk ind chunk + checkChunks ind chunks + +and checkChunk ind chunk = + case chunk of + StringChunk {size,...} => size + | BreakChunk spaces => spaces + | BlockChunk block => checkBlock ind block; + +val checkBlocks = + let + fun check ind blocks = + case blocks of + [] => 0 + | block :: blocks => + let + val Block {indent,...} = block + in + checkBlock ind block + check indent blocks + end + in + check initialIndent o rev + end; + +val initialBlock = + let + val indent = initialIndent + val style = liftStyle initialStyle + val size = 0 + val chunks = [] + in + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + end; + +val initialState = + let + val blocks = [initialBlock] + val lineIndent = initialIndent + val lineSize = 0 + in + State + {blocks = blocks, + lineIndent = lineIndent, + lineSize = lineSize} + end; + +(*BasicDebug +fun checkState state = + (let + val State {blocks, lineIndent = _, lineSize} = state + val lineSize' = checkBlocks blocks + val _ = lineSize = lineSize' orelse + raise Error "wrong lineSize" + in + () + end + handle Error err => raise Bug err) + handle Bug bug => raise Bug ("Print.checkState: " ^ bug); +*) + +fun isFinalState state = + let + val State {blocks,lineIndent,lineSize} = state + in + case blocks of + [] => raise Bug "Print.isFinalState: no block" + | [block] => isEmptyBlock block + | _ :: _ :: _ => false + end; + +local + fun renderBreak lineIndent (chunks,lines) = + let + val line = renderChunks lineIndent chunks + + val lines = line :: lines + in + lines + end; + + fun renderBreaks lineIndent lineIndent' breaks lines = + case rev breaks of + [] => raise Bug "Print.renderBreaks" + | c :: cs => + let + val lines = renderBreak lineIndent (c,lines) + in + List.foldl (renderBreak lineIndent') lines cs + end; + + fun splitAllChunks cumulatingChunks = + let + fun split chunks = + case splitChunks chunks of + SOME (prefix,chunks) => prefix :: split chunks + | NONE => [List.concat (chunks :: cumulatingChunks)] + in + split + end; + + fun mkBreak style cumulatingChunks chunks = + case splitChunks chunks of + NONE => NONE + | SOME (chunks,broken) => + let + val breaks = + case style of + InconsistentBlock => + [List.concat (broken :: cumulatingChunks)] + | _ => splitAllChunks cumulatingChunks broken + in + SOME (breaks,chunks) + end; + + fun naturalBreak blocks = + case blocks of + [] => Right ([],[]) + | block :: blocks => + case naturalBreak blocks of + Left (breaks,blocks,lineIndent,lineSize) => + let + val Block {size,...} = block + + val blocks = block :: blocks + + val lineSize = lineSize + size + in + Left (breaks,blocks,lineIndent,lineSize) + end + | Right (cumulatingChunks,blocks) => + let + val Block {indent,style,size,chunks} = block + + val style = breakStyle style + in + case mkBreak style cumulatingChunks chunks of + SOME (breaks,chunks) => + let + val size = sizeChunks chunks + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + + val blocks = block :: blocks + + val lineIndent = indent + + val lineSize = size + in + Left (breaks,blocks,lineIndent,lineSize) + end + | NONE => + let + val cumulatingChunks = chunks :: cumulatingChunks + + val size = 0 + + val chunks = [] + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + + val blocks = block :: blocks + in + Right (cumulatingChunks,blocks) + end + end; + + fun forceBreakBlock cumulatingChunks block = + let + val Block {indent, style, size = _, chunks} = block + + val style = breakStyle style + + val break = + case mkBreak style cumulatingChunks chunks of + SOME (breaks,chunks) => + let + val lineSize = sizeChunks chunks + val lineIndent = indent + in + SOME (breaks,chunks,lineIndent,lineSize) + end + | NONE => forceBreakChunks cumulatingChunks chunks + in + case break of + SOME (breaks,chunks,lineIndent,lineSize) => + let + val size = lineSize + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + in + SOME (breaks,block,lineIndent,lineSize) + end + | NONE => NONE + end + + and forceBreakChunks cumulatingChunks chunks = + case chunks of + [] => NONE + | chunk :: chunks => + case forceBreakChunk (chunks :: cumulatingChunks) chunk of + SOME (breaks,chunk,lineIndent,lineSize) => + let + val chunks = [chunk] + in + SOME (breaks,chunks,lineIndent,lineSize) + end + | NONE => + case forceBreakChunks cumulatingChunks chunks of + SOME (breaks,chunks,lineIndent,lineSize) => + let + val chunks = chunk :: chunks + + val lineSize = lineSize + sizeChunk chunk + in + SOME (breaks,chunks,lineIndent,lineSize) + end + | NONE => NONE + + and forceBreakChunk cumulatingChunks chunk = + case chunk of + StringChunk _ => NONE + | BreakChunk _ => raise Bug "Print.forceBreakChunk: BreakChunk" + | BlockChunk block => + case forceBreakBlock cumulatingChunks block of + SOME (breaks,block,lineIndent,lineSize) => + let + val chunk = BlockChunk block + in + SOME (breaks,chunk,lineIndent,lineSize) + end + | NONE => NONE; + + fun forceBreak cumulatingChunks blocks' blocks = + case blocks of + [] => NONE + | block :: blocks => + let + val cumulatingChunks = + case cumulatingChunks of + [] => raise Bug "Print.forceBreak: null cumulatingChunks" + | _ :: cumulatingChunks => cumulatingChunks + + val blocks' = + case blocks' of + [] => raise Bug "Print.forceBreak: null blocks'" + | _ :: blocks' => blocks' + in + case forceBreakBlock cumulatingChunks block of + SOME (breaks,block,lineIndent,lineSize) => + let + val blocks = block :: blocks' + in + SOME (breaks,blocks,lineIndent,lineSize) + end + | NONE => + case forceBreak cumulatingChunks blocks' blocks of + SOME (breaks,blocks,lineIndent,lineSize) => + let + val blocks = block :: blocks + + val Block {size,...} = block + + val lineSize = lineSize + size + in + SOME (breaks,blocks,lineIndent,lineSize) + end + | NONE => NONE + end; + + fun normalize lineLength lines state = + let + val State {blocks,lineIndent,lineSize} = state + in + if lineIndent + lineSize <= lineLength then (lines,state) + else + let + val break = + case naturalBreak blocks of + Left break => SOME break + | Right (c,b) => forceBreak c b blocks + in + case break of + SOME (breaks,blocks,lineIndent',lineSize) => + let + val lines = renderBreaks lineIndent lineIndent' breaks lines + + val state = + State + {blocks = blocks, + lineIndent = lineIndent', + lineSize = lineSize} + in + normalize lineLength lines state + end + | NONE => (lines,state) + end + end; + +(*BasicDebug + val normalize = fn lineLength => fn lines => fn state => + let + val () = checkState state + in + normalize lineLength lines state + end + handle Bug bug => + raise Bug ("Print.normalize: before normalize:\n" ^ bug) +*) + + fun executeBeginBlock (style,ind) lines state = + let + val State {blocks,lineIndent,lineSize} = state + + val Block {indent,...} = + case blocks of + [] => raise Bug "Print.executeBeginBlock: no block" + | block :: _ => block + + val indent = indent + ind + + val style = liftStyle style + + val size = 0 + + val chunks = [] + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + + val blocks = block :: blocks + + val state = + State + {blocks = blocks, + lineIndent = lineIndent, + lineSize = lineSize} + in + (lines,state) + end; + + fun executeEndBlock lines state = + let + val State {blocks,lineIndent,lineSize} = state + + val (lineSize,blocks) = + case blocks of + [] => raise Bug "Print.executeEndBlock: no block" + | topBlock :: blocks => + let + val Block + {indent = topIndent, + style = topStyle, + size = topSize, + chunks = topChunks} = topBlock + in + case topChunks of + [] => (lineSize,blocks) + | headTopChunks :: tailTopChunks => + let + val (lineSize,topSize,topChunks) = + case headTopChunks of + BreakChunk spaces => + let + val lineSize = lineSize - spaces + and topSize = topSize - spaces + and topChunks = tailTopChunks + in + (lineSize,topSize,topChunks) + end + | _ => (lineSize,topSize,topChunks) + + val topBlock = + Block + {indent = topIndent, + style = topStyle, + size = topSize, + chunks = topChunks} + in + case blocks of + [] => raise Error "Print.executeEndBlock: no block" + | block :: blocks => + let + val Block {indent,style,size,chunks} = block + + val size = size + topSize + + val chunks = BlockChunk topBlock :: chunks + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + + val blocks = block :: blocks + in + (lineSize,blocks) + end + end + end + + val state = + State + {blocks = blocks, + lineIndent = lineIndent, + lineSize = lineSize} + in + (lines,state) + end; + + fun executeAddString lineLength s lines state = + let + val State {blocks,lineIndent,lineSize} = state + + val n = size s + + val blocks = + case blocks of + [] => raise Bug "Print.executeAddString: no block" + | Block {indent,style,size,chunks} :: blocks => + let + val size = size + n + + val chunk = StringChunk {size = n, string = s} + + val chunks = chunk :: chunks + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + + val blocks = block :: blocks + in + blocks + end + + val lineSize = lineSize + n + + val state = + State + {blocks = blocks, + lineIndent = lineIndent, + lineSize = lineSize} + in + normalize lineLength lines state + end; + + fun executeAddBreak lineLength spaces lines state = + let + val State {blocks,lineIndent,lineSize} = state + + val (blocks,lineSize) = + case blocks of + [] => raise Bug "Print.executeAddBreak: no block" + | Block {indent,style,size,chunks} :: blocks' => + case chunks of + [] => (blocks,lineSize) + | chunk :: chunks' => + let + val spaces = + case style of + BreakingBlock => lineLength + 1 + | _ => spaces + + val size = size + spaces + + val chunks = + case chunk of + BreakChunk k => BreakChunk (k + spaces) :: chunks' + | _ => BreakChunk spaces :: chunks + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + + val blocks = block :: blocks' + + val lineSize = lineSize + spaces + in + (blocks,lineSize) + end + + val state = + State + {blocks = blocks, + lineIndent = lineIndent, + lineSize = lineSize} + in + normalize lineLength lines state + end; + + fun executeBigBreak lineLength lines state = + executeAddBreak lineLength (lineLength + 1) lines state; + + fun executeAddNewline lineLength lines state = + let + val (lines,state) = executeAddString lineLength "" lines state + val (lines,state) = executeBigBreak lineLength lines state + in + executeAddString lineLength "" lines state + end; + + fun final lineLength lines state = + let + val lines = + if isFinalState state then lines + else + let + val (lines,state) = executeBigBreak lineLength lines state + +(*BasicDebug + val _ = isFinalState state orelse raise Bug "Print.final" +*) + in + lines + end + in + if null lines then Stream.Nil else Stream.singleton lines + end; +in + fun execute {lineLength} = + let + fun advance step state = + let + val lines = [] + in + case step of + BeginBlock style_ind => executeBeginBlock style_ind lines state + | EndBlock => executeEndBlock lines state + | AddString s => executeAddString lineLength s lines state + | AddBreak spaces => executeAddBreak lineLength spaces lines state + | AddNewline => executeAddNewline lineLength lines state + end + +(*BasicDebug + val advance = fn step => fn state => + let + val (lines,state) = advance step state + val () = checkState state + in + (lines,state) + end + handle Bug bug => + raise Bug ("Print.advance: after " ^ ppStepToString step ^ + " command:\n" ^ bug) +*) + in + revConcat o Stream.maps advance (final lineLength []) initialState + end; +end; + +(* ------------------------------------------------------------------------- *) +(* Executing pretty-printers with a global line length. *) +(* ------------------------------------------------------------------------- *) + +val lineLength = Unsynchronized.ref initialLineLength; + +fun toStream ppA a = + Stream.map (fn s => s ^ "\n") + (execute {lineLength = !lineLength} (ppA a)); + +fun toString ppA a = + case execute {lineLength = !lineLength} (ppA a) of + Stream.Nil => "" + | Stream.Cons (h,t) => Stream.foldl (fn (s,z) => z ^ "\n" ^ s) h (t ()); + +fun trace ppX nameX x = + Useful.trace (toString (ppOp2 " =" ppString ppX) (nameX,x) ^ "\n"); + +end +end; + +(**** Original file: Parse.sig ****) + +(* ========================================================================= *) +(* PARSING *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) +(* ========================================================================= *) + +signature Parse = +sig + +(* ------------------------------------------------------------------------- *) +(* A "cannot parse" exception. *) +(* ------------------------------------------------------------------------- *) + +exception NoParse + +(* ------------------------------------------------------------------------- *) +(* Recursive descent parsing combinators. *) +(* ------------------------------------------------------------------------- *) + +(* + Recommended fixities: + infixr 9 >>++ infixr 8 ++ infixr 7 >> infixr 6 || *) -exception NoParse - val error : 'a -> 'b * 'a val ++ : ('a -> 'b * 'a) * ('a -> 'c * 'a) -> 'a -> ('b * 'c) * 'a @@ -4741,12 +7959,12 @@ val optional : ('a -> 'b * 'a) -> 'a -> 'b option * 'a -(* Stream based parsers *) +(* ------------------------------------------------------------------------- *) +(* Stream-based parsers. *) +(* ------------------------------------------------------------------------- *) type ('a,'b) parser = 'a Metis.Stream.stream -> 'b * 'a Metis.Stream.stream -val everything : ('a, 'b list) parser -> 'a Metis.Stream.stream -> 'b Metis.Stream.stream - val maybe : ('a -> 'b option) -> ('a,'b) parser val finished : ('a,unit) parser @@ -4755,26 +7973,49 @@ val any : ('a,'a) parser -val exact : ''a -> (''a,''a) parser - -(* ------------------------------------------------------------------------- *) -(* Infix operators *) -(* ------------------------------------------------------------------------- *) - -type infixities = {token : string, precedence : int, leftAssoc : bool} list - -val infixTokens : infixities -> string list +(* ------------------------------------------------------------------------- *) +(* Parsing whole streams. *) +(* ------------------------------------------------------------------------- *) + +val fromStream : ('a,'b) parser -> 'a Metis.Stream.stream -> 'b + +val fromList : ('a,'b) parser -> 'a list -> 'b + +val everything : ('a, 'b list) parser -> 'a Metis.Stream.stream -> 'b Metis.Stream.stream + +(* ------------------------------------------------------------------------- *) +(* Parsing lines of text. *) +(* ------------------------------------------------------------------------- *) + +val initialize : + {lines : string Metis.Stream.stream} -> + {chars : char list Metis.Stream.stream, + parseErrorLocation : unit -> string} + +val exactChar : char -> (char,unit) parser + +val exactCharList : char list -> (char,unit) parser + +val exactString : string -> (char,unit) parser + +val escapeString : {escape : char list} -> (char,string) parser + +val manySpace : (char,unit) parser + +val atLeastOneSpace : (char,unit) parser + +val fromString : (char,'a) parser -> string -> 'a + +(* ------------------------------------------------------------------------- *) +(* Infix operators. *) +(* ------------------------------------------------------------------------- *) val parseInfixes : - infixities -> (string * 'a * 'a -> 'a) -> (string,'a) parser -> + Metis.Print.infixes -> (string * 'a * 'a -> 'a) -> (string,'a) parser -> (string,'a) parser -val ppInfixes : - infixities -> ('a -> (string * 'a * 'a) option) -> ('a * bool) pp -> - ('a * bool) pp - -(* ------------------------------------------------------------------------- *) -(* Quotations *) +(* ------------------------------------------------------------------------- *) +(* Quotations. *) (* ------------------------------------------------------------------------- *) type 'a quotation = 'a Metis.frag list @@ -4783,11 +8024,11 @@ end -(**** Original file: Parser.sml ****) +(**** Original file: Parse.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; @@ -4795,12 +8036,14 @@ val foldr = List.foldr; (* ========================================================================= *) -(* PARSER COMBINATORS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) -(* ========================================================================= *) - -structure Parser :> Parser = -struct +(* PARSING *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) +(* ========================================================================= *) + +structure Parse :> Parse = +struct + +open Useful; infixr 9 >>++ infixr 8 ++ @@ -4808,137 +8051,15 @@ infixr 6 || (* ------------------------------------------------------------------------- *) -(* Helper functions. *) -(* ------------------------------------------------------------------------- *) - -exception Bug = Useful.Bug; - -val trace = Useful.trace -and equal = Useful.equal -and I = Useful.I -and K = Useful.K -and C = Useful.C -and fst = Useful.fst -and snd = Useful.snd -and pair = Useful.pair -and curry = Useful.curry -and funpow = Useful.funpow -and mem = Useful.mem -and sortMap = Useful.sortMap; - -(* ------------------------------------------------------------------------- *) -(* Pretty printing for built-in types *) -(* ------------------------------------------------------------------------- *) - -type ppstream = PP.ppstream - -datatype breakStyle = Consistent | Inconsistent - -type 'a pp = PP.ppstream -> 'a -> unit; - -val lineLength = Unsynchronized.ref 75; - -fun beginBlock pp Consistent = PP.begin_block pp PP.CONSISTENT - | beginBlock pp Inconsistent = PP.begin_block pp PP.INCONSISTENT; - -val endBlock = PP.end_block -and addString = PP.add_string -and addBreak = PP.add_break -and addNewline = PP.add_newline; - -fun ppMap f ppA (ppstrm : PP.ppstream) x : unit = ppA ppstrm (f x); - -fun ppBracket l r ppA pp a = - let - val ln = size l - in - beginBlock pp Inconsistent ln; - if ln = 0 then () else addString pp l; - ppA pp a; - if r = "" then () else addString pp r; - endBlock pp - end; - -fun ppSequence sep ppA pp = - let - fun ppX x = (addString pp sep; addBreak pp (1,0); ppA pp x) - in - fn [] => () - | h :: t => - (beginBlock pp Inconsistent 0; - ppA pp h; - app ppX t; - endBlock pp) - end; - -fun ppBinop s ppA ppB pp (a,b) = - (beginBlock pp Inconsistent 0; - ppA pp a; - if s = "" then () else addString pp s; - addBreak pp (1,0); - ppB pp b; - endBlock pp); - -fun ppTrinop ab bc ppA ppB ppC pp (a,b,c) = - (beginBlock pp Inconsistent 0; - ppA pp a; - if ab = "" then () else addString pp ab; - addBreak pp (1,0); - ppB pp b; - if bc = "" then () else addString pp bc; - addBreak pp (1,0); - ppC pp c; - endBlock pp); - -(* Pretty-printers for common types *) - -fun ppString pp s = - (beginBlock pp Inconsistent 0; - addString pp s; - endBlock pp); - -val ppUnit = ppMap (fn () => "()") ppString; - -val ppChar = ppMap str ppString; - -val ppBool = ppMap (fn true => "true" | false => "false") ppString; - -val ppInt = ppMap Int.toString ppString; - -val ppReal = ppMap Real.toString ppString; - -val ppOrder = - let - fun f LESS = "Less" - | f EQUAL = "Equal" - | f GREATER = "Greater" - in - ppMap f ppString - end; - -fun ppList ppA = ppBracket "[" "]" (ppSequence "," ppA); - -fun ppOption _ pp NONE = ppString pp "-" - | ppOption ppA pp (SOME a) = ppA pp a; - -fun ppPair ppA ppB = ppBracket "(" ")" (ppBinop "," ppA ppB); - -fun ppTriple ppA ppB ppC = ppBracket "(" ")" (ppTrinop "," "," ppA ppB ppC); - -(* Keep eta expanded to evaluate lineLength when supplied with a *) -fun toString ppA a = PP.pp_to_string (!lineLength) ppA a; - -fun fromString toS = ppMap toS ppString; - -fun ppTrace ppX nameX x = - trace (toString (ppBinop " =" ppString ppX) (nameX,x) ^ "\n"); - -(* ------------------------------------------------------------------------- *) -(* Generic. *) +(* A "cannot parse" exception. *) (* ------------------------------------------------------------------------- *) exception NoParse; +(* ------------------------------------------------------------------------- *) +(* Recursive descent parsing combinators. *) +(* ------------------------------------------------------------------------- *) + val error : 'a -> 'b * 'a = fn _ => raise NoParse; fun op ++ (parser1,parser2) input = @@ -4981,7 +8102,7 @@ let fun sparser l = parser >> (fn x => x :: l) in - mmany sparser [] >> rev + mmany sparser [] >> rev end; fun atLeastOne p = (p ++ many p) >> op::; @@ -4991,191 +8112,179 @@ fun optional p = (p >> SOME) || (nothing >> K NONE); (* ------------------------------------------------------------------------- *) -(* Stream-based. *) +(* Stream-based parsers. *) (* ------------------------------------------------------------------------- *) type ('a,'b) parser = 'a Stream.stream -> 'b * 'a Stream.stream -fun everything parser = - let - fun f input = - let - val (result,input) = parser input - in - Stream.append (Stream.fromList result) (fn () => f input) - end - handle NoParse => - if Stream.null input then Stream.NIL else raise NoParse - in - f - end; - -fun maybe p Stream.NIL = raise NoParse - | maybe p (Stream.CONS (h,t)) = +fun maybe p Stream.Nil = raise NoParse + | maybe p (Stream.Cons (h,t)) = case p h of SOME r => (r, t ()) | NONE => raise NoParse; -fun finished Stream.NIL = ((), Stream.NIL) - | finished (Stream.CONS _) = raise NoParse; +fun finished Stream.Nil = ((), Stream.Nil) + | finished (Stream.Cons _) = raise NoParse; fun some p = maybe (fn x => if p x then SOME x else NONE); fun any input = some (K true) input; -fun exact tok = some (fn item => item = tok); - -(* ------------------------------------------------------------------------- *) -(* Parsing and pretty-printing for infix operators. *) -(* ------------------------------------------------------------------------- *) - -type infixities = {token : string, precedence : int, leftAssoc : bool} list; - -local - fun unflatten ({token,precedence,leftAssoc}, ([],_)) = - ([(leftAssoc, [token])], precedence) - | unflatten ({token,precedence,leftAssoc}, ((a,l) :: dealt, p)) = - if p = precedence then - let - val _ = leftAssoc = a orelse - raise Bug "infix parser/printer: mixed assocs" - in - ((a, token :: l) :: dealt, p) - end - else - ((leftAssoc,[token]) :: (a,l) :: dealt, precedence); -in - fun layerOps infixes = - let - val infixes = sortMap #precedence Int.compare infixes - val (parsers,_) = foldl unflatten ([],0) infixes - in - parsers - end; -end; - -local - fun chop (#" " :: chs) = let val (n,l) = chop chs in (n + 1, l) end - | chop chs = (0,chs); - - fun nspaces n = funpow n (curry op^ " ") ""; - - fun spacify tok = - let - val chs = explode tok - val (r,chs) = chop (rev chs) - val (l,chs) = chop (rev chs) - in - ((l,r), implode chs) - end; - - fun lrspaces (l,r) = - (if l = 0 then K () else C addString (nspaces l), - if r = 0 then K () else C addBreak (r, 0)); -in - fun opSpaces s = let val (l_r,s) = spacify s in (lrspaces l_r, s) end; - - val opClean = snd o spacify; -end; - -val infixTokens : infixities -> string list = - List.map (fn {token,...} => opClean token); +(* ------------------------------------------------------------------------- *) +(* Parsing whole streams. *) +(* ------------------------------------------------------------------------- *) + +fun fromStream parser input = + let + val (res,_) = (parser ++ finished >> fst) input + in + res + end; + +fun fromList parser l = fromStream parser (Stream.fromList l); + +fun everything parser = + let + fun parserOption input = + SOME (parser input) + handle e as NoParse => if Stream.null input then NONE else raise e + + fun parserList input = + case parserOption input of + NONE => Stream.Nil + | SOME (result,input) => + Stream.append (Stream.fromList result) (fn () => parserList input) + in + parserList + end; + +(* ------------------------------------------------------------------------- *) +(* Parsing lines of text. *) +(* ------------------------------------------------------------------------- *) + +fun initialize {lines} = + let + val lastLine = Unsynchronized.ref (~1,"","","") + + val chars = + let + fun saveLast line = + let + val Unsynchronized.ref (n,_,l2,l3) = lastLine + val () = lastLine := (n + 1, l2, l3, line) + in + explode line + end + in + Stream.memoize (Stream.map saveLast lines) + end + + fun parseErrorLocation () = + let + val Unsynchronized.ref (n,l1,l2,l3) = lastLine + in + (if n <= 0 then "at start" + else "around line " ^ Int.toString n) ^ + chomp (":\n" ^ l1 ^ l2 ^ l3) + end + in + {chars = chars, + parseErrorLocation = parseErrorLocation} + end; + +fun exactChar (c : char) = some (equal c) >> K (); + +fun exactCharList cs = + case cs of + [] => nothing + | c :: cs => (exactChar c ++ exactCharList cs) >> snd; + +fun exactString s = exactCharList (explode s); + +fun escapeString {escape} = + let + fun isEscape c = mem c escape + + fun isNormal c = + case c of + #"\\" => false + | #"\n" => false + | #"\t" => false + | _ => not (isEscape c) + + val escapeParser = + (exactChar #"\\" >> K #"\\") || + (exactChar #"n" >> K #"\n") || + (exactChar #"t" >> K #"\t") || + some isEscape + + val charParser = + ((exactChar #"\\" ++ escapeParser) >> snd) || + some isNormal + in + many charParser >> implode + end; + +local + val isSpace = Char.isSpace; + + val space = some isSpace; +in + val manySpace = many space >> K (); + + val atLeastOneSpace = atLeastOne space >> K (); +end; + +fun fromString parser s = fromList parser (explode s); + +(* ------------------------------------------------------------------------- *) +(* Infix operators. *) +(* ------------------------------------------------------------------------- *) fun parseGenInfix update sof toks parse inp = let - val (e, rest) = parse inp - + val (e,rest) = parse inp + val continue = case rest of - Stream.NIL => NONE - | Stream.CONS (h, t) => if mem h toks then SOME (h, t) else NONE + Stream.Nil => NONE + | Stream.Cons (h_t as (h,_)) => + if StringSet.member h toks then SOME h_t else NONE in case continue of NONE => (sof e, rest) | SOME (h,t) => parseGenInfix update (update sof h e) toks parse (t ()) end; -fun parseLeftInfix toks con = - parseGenInfix (fn f => fn t => fn a => fn b => con (t, f a, b)) I toks; - -fun parseRightInfix toks con = - parseGenInfix (fn f => fn t => fn a => fn b => f (con (t, a, b))) I toks; +local + fun add ({leftSpaces = _, token = t, rightSpaces = _}, s) = StringSet.add s t; + + fun parse leftAssoc toks con = + let + val update = + if leftAssoc then (fn f => fn t => fn a => fn b => con (t, f a, b)) + else (fn f => fn t => fn a => fn b => f (con (t, a, b))) + in + parseGenInfix update I toks + end; +in + fun parseLayeredInfixes {tokens,leftAssoc} = + let + val toks = List.foldl add StringSet.empty tokens + in + parse leftAssoc toks + end; +end; fun parseInfixes ops = let - fun layeredOp (x,y) = (x, List.map opClean y) - - val layeredOps = List.map layeredOp (layerOps ops) - - fun iparser (a,t) = (if a then parseLeftInfix else parseRightInfix) t - - val iparsers = List.map iparser layeredOps + val layeredOps = Print.layerInfixes ops + + val iparsers = List.map parseLayeredInfixes layeredOps in fn con => fn subparser => foldl (fn (p,sp) => p con sp) subparser iparsers end; -fun ppGenInfix left toks = - let - val spc = List.map opSpaces toks - in - fn dest => fn ppSub => - let - fun dest' tm = - case dest tm of - NONE => NONE - | SOME (t, a, b) => - Option.map (pair (a,b)) (List.find (equal t o snd) spc) - - open PP - - fun ppGo pp (tmr as (tm,r)) = - case dest' tm of - NONE => ppSub pp tmr - | SOME ((a,b),((lspc,rspc),tok)) => - ((if left then ppGo else ppSub) pp (a,true); - lspc pp; addString pp tok; rspc pp; - (if left then ppSub else ppGo) pp (b,r)) - in - fn pp => fn tmr as (tm,_) => - case dest' tm of - NONE => ppSub pp tmr - | SOME _ => (beginBlock pp Inconsistent 0; ppGo pp tmr; endBlock pp) - end - end; - -fun ppLeftInfix toks = ppGenInfix true toks; - -fun ppRightInfix toks = ppGenInfix false toks; - -fun ppInfixes ops = - let - val layeredOps = layerOps ops - - val toks = List.concat (List.map (List.map opClean o snd) layeredOps) - - fun iprinter (a,t) = (if a then ppLeftInfix else ppRightInfix) t - - val iprinters = List.map iprinter layeredOps - in - fn dest => fn ppSub => - let - fun printer sub = foldl (fn (ip,p) => ip dest p) sub iprinters - - fun isOp t = case dest t of SOME (x,_,_) => mem x toks | _ => false - - open PP - - fun subpr pp (tmr as (tm,_)) = - if isOp tm then - (beginBlock pp Inconsistent 1; addString pp "("; - printer subpr pp (tm, false); addString pp ")"; endBlock pp) - else ppSub pp tmr - in - fn pp => fn tmr => - (beginBlock pp Inconsistent 0; printer subpr pp tmr; endBlock pp) - end - end; - -(* ------------------------------------------------------------------------- *) -(* Quotations *) +(* ------------------------------------------------------------------------- *) +(* Quotations. *) (* ------------------------------------------------------------------------- *) type 'a quotation = 'a frag list; @@ -5197,7 +8306,7 @@ (* ========================================================================= *) (* PROCESSING COMMAND LINE OPTIONS *) -(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Options = @@ -5293,7 +8402,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; @@ -5302,7 +8411,7 @@ (* ========================================================================= *) (* PROCESSING COMMAND LINE OPTIONS *) -(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Options :> Options = @@ -5418,9 +8527,10 @@ description = "no more options", processor = fn _ => raise Fail "basicOptions: --"}, {switches = ["-?","-h","--help"], arguments = [], - description = "display all options and exit", + description = "display option information and exit", processor = fn _ => raise OptionExit - {message = SOME "displaying all options", usage = true, success = true}}, + {message = SOME "displaying option information", + usage = true, success = true}}, {switches = ["-v", "--version"], arguments = [], description = "display version information", processor = fn _ => raise Fail "basicOptions: -v, --version"}]; @@ -5429,8 +8539,9 @@ (* All the command line options of a program *) (* ------------------------------------------------------------------------- *) -type allOptions = {name : string, version : string, header : string, - footer : string, options : opt list}; +type allOptions = + {name : string, version : string, header : string, + footer : string, options : opt list}; (* ------------------------------------------------------------------------- *) (* Usage information *) @@ -5547,17 +8658,47 @@ (* ========================================================================= *) (* NAMES *) -(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Name = sig -type name = string +(* ------------------------------------------------------------------------- *) +(* A type of names. *) +(* ------------------------------------------------------------------------- *) + +type name = string (* MODIFIED by Jasmin Blanchette *) + +(* ------------------------------------------------------------------------- *) +(* A total ordering. *) +(* ------------------------------------------------------------------------- *) val compare : name * name -> order -val pp : name Metis.Parser.pp +val equal : name -> name -> bool + +(* ------------------------------------------------------------------------- *) +(* Fresh names. *) +(* ------------------------------------------------------------------------- *) + +val newName : unit -> name + +val newNames : int -> name list + +val variantPrime : (name -> bool) -> name -> name + +val variantNum : (name -> bool) -> name -> name + +(* ------------------------------------------------------------------------- *) +(* Parsing and pretty printing. *) +(* ------------------------------------------------------------------------- *) + +val pp : name Metis.Print.pp + +val toString : name -> string + +val fromString : string -> name end @@ -5565,7 +8706,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; @@ -5574,43 +8715,176 @@ (* ========================================================================= *) (* NAMES *) -(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Name :> Name = struct +open Useful; + +(* ------------------------------------------------------------------------- *) +(* A type of names. *) +(* ------------------------------------------------------------------------- *) + type name = string; +(* ------------------------------------------------------------------------- *) +(* A total ordering. *) +(* ------------------------------------------------------------------------- *) + val compare = String.compare; -val pp = Parser.ppString; +fun equal n1 n2 = n1 = n2; + +(* ------------------------------------------------------------------------- *) +(* Fresh variables. *) +(* ------------------------------------------------------------------------- *) + +local + val prefix = "_"; + + fun numName i = mkPrefix prefix (Int.toString i); +in + fun newName () = numName (newInt ()); + + fun newNames n = map numName (newInts n); +end; + +fun variantPrime acceptable = + let + fun variant n = if acceptable n then n else variant (n ^ "'") + in + variant + end; + +local + fun isDigitOrPrime #"'" = true + | isDigitOrPrime c = Char.isDigit c; +in + fun variantNum acceptable n = + if acceptable n then n + else + let + val n = stripSuffix isDigitOrPrime n + + fun variant i = + let + val n_i = n ^ Int.toString i + in + if acceptable n_i then n_i else variant (i + 1) + end + in + variant 0 + end; +end; + +(* ------------------------------------------------------------------------- *) +(* Parsing and pretty printing. *) +(* ------------------------------------------------------------------------- *) + +val pp = Print.ppString; + +fun toString s : string = s; + +fun fromString s : name = s; end structure NameOrdered = struct type t = Name.name val compare = Name.compare end +structure NameMap = KeyMap (NameOrdered); + structure NameSet = struct local - structure S = ElementSet (NameOrdered); + structure S = ElementSet (NameMap); in open S; end; val pp = - Parser.ppMap + Print.ppMap toList - (Parser.ppBracket "{" "}" (Parser.ppSequence "," Name.pp)); - -end - -structure NameMap = KeyMap (NameOrdered); - -structure NameArity = -struct + (Print.ppBracket "{" "}" (Print.ppOpList "," Name.pp)); + +end +end; + +(**** Original file: NameArity.sig ****) + +(* ========================================================================= *) +(* NAME/ARITY PAIRS *) +(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) +(* ========================================================================= *) + +signature NameArity = +sig + +(* ------------------------------------------------------------------------- *) +(* A type of name/arity pairs. *) +(* ------------------------------------------------------------------------- *) + +type nameArity = Metis.Name.name * int + +val name : nameArity -> Metis.Name.name + +val arity : nameArity -> int + +(* ------------------------------------------------------------------------- *) +(* Testing for different arities. *) +(* ------------------------------------------------------------------------- *) + +val nary : int -> nameArity -> bool + +val nullary : nameArity -> bool + +val unary : nameArity -> bool + +val binary : nameArity -> bool + +val ternary : nameArity -> bool + +(* ------------------------------------------------------------------------- *) +(* A total ordering. *) +(* ------------------------------------------------------------------------- *) + +val compare : nameArity * nameArity -> order + +val equal : nameArity -> nameArity -> bool + +(* ------------------------------------------------------------------------- *) +(* Parsing and pretty printing. *) +(* ------------------------------------------------------------------------- *) + +val pp : nameArity Metis.Print.pp + +end + +(**** Original file: NameArity.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; + +(* ========================================================================= *) +(* NAME/ARITY PAIRS *) +(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) +(* ========================================================================= *) + +structure NameArity :> NameArity = +struct + +(* ------------------------------------------------------------------------- *) +(* A type of name/arity pairs. *) +(* ------------------------------------------------------------------------- *) type nameArity = Name.name * int; @@ -5618,6 +8892,10 @@ fun arity ((_,i) : nameArity) = i; +(* ------------------------------------------------------------------------- *) +(* Testing for different arities. *) +(* ------------------------------------------------------------------------- *) + fun nary i n_i = arity n_i = i; val nullary = nary 0 @@ -5625,24 +8903,56 @@ and binary = nary 2 and ternary = nary 3; +(* ------------------------------------------------------------------------- *) +(* A total ordering. *) +(* ------------------------------------------------------------------------- *) + fun compare ((n1,i1),(n2,i2)) = case Name.compare (n1,n2) of LESS => LESS | EQUAL => Int.compare (i1,i2) | GREATER => GREATER; -val pp = Parser.ppMap (fn (n,i) => n ^ "/" ^ Int.toString i) Parser.ppString; +fun equal (n1,i1) (n2,i2) = i1 = i2 andalso Name.equal n1 n2; + +(* ------------------------------------------------------------------------- *) +(* Parsing and pretty printing. *) +(* ------------------------------------------------------------------------- *) + +fun pp (n,i) = + Print.blockProgram Print.Inconsistent 0 + [Name.pp n, + Print.addString "/", + Print.ppInt i]; end structure NameArityOrdered = struct type t = NameArity.nameArity val compare = NameArity.compare end +structure NameArityMap = +struct + + local + structure S = KeyMap (NameArityOrdered); + in + open S; + end; + + fun compose m1 m2 = + let + fun pk ((_,a),n) = peek m2 (n,a) + in + mapPartial pk m1 + end; + +end + structure NameAritySet = struct local - structure S = ElementSet (NameArityOrdered); + structure S = ElementSet (NameArityMap); in open S; end; @@ -5650,20 +8960,18 @@ val allNullary = all NameArity.nullary; val pp = - Parser.ppMap + Print.ppMap toList - (Parser.ppBracket "{" "}" (Parser.ppSequence "," NameArity.pp)); - -end - -structure NameArityMap = KeyMap (NameArityOrdered); + (Print.ppBracket "{" "}" (Print.ppOpList "," NameArity.pp)); + +end end; (**** Original file: Term.sig ****) (* ========================================================================= *) (* FIRST ORDER LOGIC TERMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Term = @@ -5743,6 +9051,8 @@ val compare : term * term -> order +val equal : term -> term -> bool + (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) @@ -5757,7 +9067,7 @@ val find : (term -> bool) -> term -> path option -val ppPath : path Metis.Parser.pp +val ppPath : path Metis.Print.pp val pathToString : path -> string @@ -5769,6 +9079,8 @@ val freeVars : term -> Metis.NameSet.set +val freeVarsList : term list -> Metis.NameSet.set + (* ------------------------------------------------------------------------- *) (* Fresh variables. *) (* ------------------------------------------------------------------------- *) @@ -5785,6 +9097,10 @@ (* Special support for terms with type annotations. *) (* ------------------------------------------------------------------------- *) +val hasTypeFunctionName : functionName + +val hasTypeFunction : function + val isTypedVar : term -> bool val typedSymbols : term -> int @@ -5795,15 +9111,17 @@ (* Special support for terms with an explicit function application operator. *) (* ------------------------------------------------------------------------- *) -val mkComb : term * term -> term - -val destComb : term -> term * term - -val isComb : term -> bool - -val listMkComb : term * term list -> term - -val stripComb : term -> term * term list +val appName : Metis.Name.name + +val mkApp : term * term -> term + +val destApp : term -> term * term + +val isApp : term -> bool + +val listMkApp : term * term list -> term + +val stripApp : term -> term * term list (* ------------------------------------------------------------------------- *) (* Parsing and pretty printing. *) @@ -5811,23 +9129,23 @@ (* Infix symbols *) -val infixes : Metis.Parser.infixities Unsynchronized.ref +val infixes : Metis.Print.infixes Unsynchronized.ref (* The negation symbol *) -val negation : Metis.Name.name Unsynchronized.ref +val negation : string Unsynchronized.ref (* Binder symbols *) -val binders : Metis.Name.name list Unsynchronized.ref +val binders : string list Unsynchronized.ref (* Bracket symbols *) -val brackets : (Metis.Name.name * Metis.Name.name) list Unsynchronized.ref +val brackets : (string * string) list Unsynchronized.ref (* Pretty printing *) -val pp : term Metis.Parser.pp +val pp : term Metis.Print.pp val toString : term -> string @@ -5835,7 +9153,7 @@ val fromString : string -> term -val parse : term Metis.Parser.quotation -> term +val parse : term Metis.Parse.quotation -> term end @@ -5843,7 +9161,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; @@ -5852,7 +9170,7 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC TERMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Term :> Term = @@ -5861,24 +9179,6 @@ open Useful; (* ------------------------------------------------------------------------- *) -(* Helper functions. *) -(* ------------------------------------------------------------------------- *) - -fun stripSuffix pred s = - let - fun f 0 = "" - | f n = - let - val n' = n - 1 - in - if pred (String.sub (s,n')) then f n' - else String.substring (s,0,n) - end - in - f (size s) - end; - -(* ------------------------------------------------------------------------- *) (* A type of first order logic terms. *) (* ------------------------------------------------------------------------- *) @@ -5905,7 +9205,7 @@ val isVar = can destVar; -fun equalVar v (Var v') = v = v' +fun equalVar v (Var v') = Name.equal v v' | equalVar _ _ = false; (* Functions *) @@ -5954,7 +9254,7 @@ fun mkBinop f (a,b) = Fn (f,[a,b]); fun destBinop f (Fn (x,[a,b])) = - if x = f then (a,b) else raise Error "Term.destBinop: wrong binop" + if Name.equal x f then (a,b) else raise Error "Term.destBinop: wrong binop" | destBinop _ _ = raise Error "Term.destBinop: not a binop"; fun isBinop f = can (destBinop f); @@ -5981,27 +9281,37 @@ local fun cmp [] [] = EQUAL - | cmp (Var _ :: _) (Fn _ :: _) = LESS - | cmp (Fn _ :: _) (Var _ :: _) = GREATER - | cmp (Var v1 :: tms1) (Var v2 :: tms2) = - (case Name.compare (v1,v2) of - LESS => LESS - | EQUAL => cmp tms1 tms2 - | GREATER => GREATER) - | cmp (Fn (f1,a1) :: tms1) (Fn (f2,a2) :: tms2) = - (case Name.compare (f1,f2) of - LESS => LESS - | EQUAL => - (case Int.compare (length a1, length a2) of - LESS => LESS - | EQUAL => cmp (a1 @ tms1) (a2 @ tms2) - | GREATER => GREATER) - | GREATER => GREATER) + | cmp (tm1 :: tms1) (tm2 :: tms2) = + let + val tm1_tm2 = (tm1,tm2) + in + if Portable.pointerEqual tm1_tm2 then cmp tms1 tms2 + else + case tm1_tm2 of + (Var v1, Var v2) => + (case Name.compare (v1,v2) of + LESS => LESS + | EQUAL => cmp tms1 tms2 + | GREATER => GREATER) + | (Var _, Fn _) => LESS + | (Fn _, Var _) => GREATER + | (Fn (f1,a1), Fn (f2,a2)) => + (case Name.compare (f1,f2) of + LESS => LESS + | EQUAL => + (case Int.compare (length a1, length a2) of + LESS => LESS + | EQUAL => cmp (a1 @ tms1) (a2 @ tms2) + | GREATER => GREATER) + | GREATER => GREATER) + end | cmp _ _ = raise Bug "Term.compare"; in fun compare (tm1,tm2) = cmp [tm1] [tm2]; end; +fun equal tm1 tm2 = compare (tm1,tm2) = EQUAL; + (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) @@ -6030,7 +9340,7 @@ fun subterms tm = subtms [([],tm)] []; end; -fun replace tm ([],res) = if res = tm then tm else res +fun replace tm ([],res) = if equal res tm then tm else res | replace tm (h :: t, res) = case tm of Var _ => raise Error "Term.replace: Var" @@ -6041,7 +9351,7 @@ val arg = List.nth (tms,h) val arg' = replace arg (t,res) in - if Sharing.pointerEqual (arg',arg) then tm + if Portable.pointerEqual (arg',arg) then tm else Fn (func, updateNth (h,arg') tms) end; @@ -6063,9 +9373,9 @@ fn tm => search [([],tm)] end; -val ppPath = Parser.ppList Parser.ppInt; - -val pathToString = Parser.toString ppPath; +val ppPath = Print.ppList Print.ppInt; + +val pathToString = Print.toString ppPath; (* ------------------------------------------------------------------------- *) (* Free variables. *) @@ -6073,7 +9383,7 @@ local fun free _ [] = false - | free v (Var w :: tms) = v = w orelse free v tms + | free v (Var w :: tms) = Name.equal v w orelse free v tms | free v (Fn (_,args) :: tms) = free v (args @ tms); in fun freeIn v tm = free v [tm]; @@ -6084,77 +9394,100 @@ | free vs (Var v :: tms) = free (NameSet.add vs v) tms | free vs (Fn (_,args) :: tms) = free vs (args @ tms); in - fun freeVars tm = free NameSet.empty [tm]; + val freeVarsList = free NameSet.empty; + + fun freeVars tm = freeVarsList [tm]; end; (* ------------------------------------------------------------------------- *) (* Fresh variables. *) (* ------------------------------------------------------------------------- *) -local - val prefix = "_"; - - fun numVar i = Var (mkPrefix prefix (Int.toString i)); -in - fun newVar () = numVar (newInt ()); - - fun newVars n = map numVar (newInts n); -end; - -fun variantPrime avoid = - let - fun f v = if NameSet.member v avoid then f (v ^ "'") else v - in - f - end; - -fun variantNum avoid v = - if not (NameSet.member v avoid) then v - else - let - val v = stripSuffix Char.isDigit v - - fun f n = - let - val v_n = v ^ Int.toString n - in - if NameSet.member v_n avoid then f (n + 1) else v_n - end - in - f 0 - end; +fun newVar () = Var (Name.newName ()); + +fun newVars n = map Var (Name.newNames n); + +local + fun avoidAcceptable avoid n = not (NameSet.member n avoid); +in + fun variantPrime avoid = Name.variantPrime (avoidAcceptable avoid); + + fun variantNum avoid = Name.variantNum (avoidAcceptable avoid); +end; (* ------------------------------------------------------------------------- *) (* Special support for terms with type annotations. *) (* ------------------------------------------------------------------------- *) -fun isTypedVar (Var _) = true - | isTypedVar (Fn (":", [Var _, _])) = true - | isTypedVar (Fn _) = false; +val hasTypeFunctionName = Name.fromString ":"; + +val hasTypeFunction = (hasTypeFunctionName,2); + +fun destFnHasType ((f,a) : functionName * term list) = + if not (Name.equal f hasTypeFunctionName) then + raise Error "Term.destFnHasType" + else + case a of + [tm,ty] => (tm,ty) + | _ => raise Error "Term.destFnHasType"; + +val isFnHasType = can destFnHasType; + +fun isTypedVar tm = + case tm of + Var _ => true + | Fn func => + case total destFnHasType func of + SOME (Var _, _) => true + | _ => false; local fun sz n [] = n - | sz n (Var _ :: tms) = sz (n + 1) tms - | sz n (Fn (":",[tm,_]) :: tms) = sz n (tm :: tms) - | sz n (Fn (_,args) :: tms) = sz (n + 1) (args @ tms); + | sz n (tm :: tms) = + case tm of + Var _ => sz (n + 1) tms + | Fn func => + case total destFnHasType func of + SOME (tm,_) => sz n (tm :: tms) + | NONE => + let + val (_,a) = func + in + sz (n + 1) (a @ tms) + end; in fun typedSymbols tm = sz 0 [tm]; end; local fun subtms [] acc = acc - | subtms ((_, Var _) :: rest) acc = subtms rest acc - | subtms ((_, Fn (":", [Var _, _])) :: rest) acc = subtms rest acc - | subtms ((path, tm as Fn func) :: rest) acc = - let - fun f (n,arg) = (n :: path, arg) - - val acc = (rev path, tm) :: acc - in - case func of - (":",[arg,_]) => subtms ((0 :: path, arg) :: rest) acc - | (_,args) => subtms (map f (enumerate args) @ rest) acc - end; + | subtms ((path,tm) :: rest) acc = + case tm of + Var _ => subtms rest acc + | Fn func => + case total destFnHasType func of + SOME (t,_) => + (case t of + Var _ => subtms rest acc + | Fn _ => + let + val acc = (rev path, tm) :: acc + val rest = (0 :: path, t) :: rest + in + subtms rest acc + end) + | NONE => + let + fun f (n,arg) = (n :: path, arg) + + val (_,args) = func + + val acc = (rev path, tm) :: acc + + val rest = map f (enumerate args) @ rest + in + subtms rest acc + end; in fun nonVarTypedSubterms tm = subtms [([],tm)] []; end; @@ -6163,20 +9496,37 @@ (* Special support for terms with an explicit function application operator. *) (* ------------------------------------------------------------------------- *) -fun mkComb (f,a) = Fn (".",[f,a]); - -fun destComb (Fn (".",[f,a])) = (f,a) - | destComb _ = raise Error "destComb"; - -val isComb = can destComb; - -fun listMkComb (f,l) = foldl mkComb f l; - -local - fun strip tms (Fn (".",[f,a])) = strip (a :: tms) f - | strip tms tm = (tm,tms); -in - fun stripComb tm = strip [] tm; +val appName = Name.fromString "."; + +fun mkFnApp (fTm,aTm) = (appName, [fTm,aTm]); + +fun mkApp f_a = Fn (mkFnApp f_a); + +fun destFnApp ((f,a) : Name.name * term list) = + if not (Name.equal f appName) then raise Error "Term.destFnApp" + else + case a of + [fTm,aTm] => (fTm,aTm) + | _ => raise Error "Term.destFnApp"; + +val isFnApp = can destFnApp; + +fun destApp tm = + case tm of + Var _ => raise Error "Term.destApp" + | Fn func => destFnApp func; + +val isApp = can destApp; + +fun listMkApp (f,l) = foldl mkApp f l; + +local + fun strip tms tm = + case total destApp tm of + SOME (f,a) => strip (a :: tms) f + | NONE => (tm,tms); +in + fun stripApp tm = strip [] tm; end; (* ------------------------------------------------------------------------- *) @@ -6185,185 +9535,204 @@ (* Operators parsed and printed infix *) -val infixes : Parser.infixities Unsynchronized.ref = Unsynchronized.ref - [(* ML symbols *) - {token = " / ", precedence = 7, leftAssoc = true}, - {token = " div ", precedence = 7, leftAssoc = true}, - {token = " mod ", precedence = 7, leftAssoc = true}, - {token = " * ", precedence = 7, leftAssoc = true}, - {token = " + ", precedence = 6, leftAssoc = true}, - {token = " - ", precedence = 6, leftAssoc = true}, - {token = " ^ ", precedence = 6, leftAssoc = true}, - {token = " @ ", precedence = 5, leftAssoc = false}, - {token = " :: ", precedence = 5, leftAssoc = false}, - {token = " = ", precedence = 4, leftAssoc = true}, - {token = " <> ", precedence = 4, leftAssoc = true}, - {token = " <= ", precedence = 4, leftAssoc = true}, - {token = " < ", precedence = 4, leftAssoc = true}, - {token = " >= ", precedence = 4, leftAssoc = true}, - {token = " > ", precedence = 4, leftAssoc = true}, - {token = " o ", precedence = 3, leftAssoc = true}, - {token = " -> ", precedence = 2, leftAssoc = false}, (* inferred prec *) - {token = " : ", precedence = 1, leftAssoc = false}, (* inferred prec *) - {token = ", ", precedence = 0, leftAssoc = false}, (* inferred prec *) - - (* Logical connectives *) - {token = " /\\ ", precedence = ~1, leftAssoc = false}, - {token = " \\/ ", precedence = ~2, leftAssoc = false}, - {token = " ==> ", precedence = ~3, leftAssoc = false}, - {token = " <=> ", precedence = ~4, leftAssoc = false}, - - (* Other symbols *) - {token = " . ", precedence = 9, leftAssoc = true}, (* function app *) - {token = " ** ", precedence = 8, leftAssoc = true}, - {token = " ++ ", precedence = 6, leftAssoc = true}, - {token = " -- ", precedence = 6, leftAssoc = true}, - {token = " == ", precedence = 4, leftAssoc = true}]; +val infixes = + (Unsynchronized.ref o Print.Infixes) + [(* ML symbols *) + {token = " / ", precedence = 7, leftAssoc = true}, + {token = " div ", precedence = 7, leftAssoc = true}, + {token = " mod ", precedence = 7, leftAssoc = true}, + {token = " * ", precedence = 7, leftAssoc = true}, + {token = " + ", precedence = 6, leftAssoc = true}, + {token = " - ", precedence = 6, leftAssoc = true}, + {token = " ^ ", precedence = 6, leftAssoc = true}, + {token = " @ ", precedence = 5, leftAssoc = false}, + {token = " :: ", precedence = 5, leftAssoc = false}, + {token = " = ", precedence = 4, leftAssoc = true}, + {token = " <> ", precedence = 4, leftAssoc = true}, + {token = " <= ", precedence = 4, leftAssoc = true}, + {token = " < ", precedence = 4, leftAssoc = true}, + {token = " >= ", precedence = 4, leftAssoc = true}, + {token = " > ", precedence = 4, leftAssoc = true}, + {token = " o ", precedence = 3, leftAssoc = true}, + {token = " -> ", precedence = 2, leftAssoc = false}, (* inferred prec *) + {token = " : ", precedence = 1, leftAssoc = false}, (* inferred prec *) + {token = ", ", precedence = 0, leftAssoc = false}, (* inferred prec *) + + (* Logical connectives *) + {token = " /\\ ", precedence = ~1, leftAssoc = false}, + {token = " \\/ ", precedence = ~2, leftAssoc = false}, + {token = " ==> ", precedence = ~3, leftAssoc = false}, + {token = " <=> ", precedence = ~4, leftAssoc = false}, + + (* Other symbols *) + {token = " . ", precedence = 9, leftAssoc = true}, (* function app *) + {token = " ** ", precedence = 8, leftAssoc = true}, + {token = " ++ ", precedence = 6, leftAssoc = true}, + {token = " -- ", precedence = 6, leftAssoc = true}, + {token = " == ", precedence = 4, leftAssoc = true}]; (* The negation symbol *) -val negation : Name.name Unsynchronized.ref = Unsynchronized.ref "~"; +val negation : string Unsynchronized.ref = Unsynchronized.ref "~"; (* Binder symbols *) -val binders : Name.name list Unsynchronized.ref = Unsynchronized.ref ["\\","!","?","?!"]; +val binders : string list Unsynchronized.ref = Unsynchronized.ref ["\\","!","?","?!"]; (* Bracket symbols *) -val brackets : (Name.name * Name.name) list Unsynchronized.ref = Unsynchronized.ref [("[","]"),("{","}")]; +val brackets : (string * string) list Unsynchronized.ref = Unsynchronized.ref [("[","]"),("{","}")]; (* Pretty printing *) -local - open Parser; -in - fun pp inputPpstrm inputTerm = - let - val quants = !binders - and iOps = !infixes - and neg = !negation - and bracks = !brackets - - val bracks = map (fn (b1,b2) => (b1 ^ b2, b1, b2)) bracks - - val bTokens = map #2 bracks @ map #3 bracks - - val iTokens = infixTokens iOps - - fun destI (Fn (f,[a,b])) = - if mem f iTokens then SOME (f,a,b) else NONE - | destI _ = NONE - - val iPrinter = ppInfixes iOps destI - - val specialTokens = neg :: quants @ ["$","(",")"] @ bTokens @ iTokens - - fun vName bv s = NameSet.member s bv - - fun checkVarName bv s = if vName bv s then s else "$" ^ s - - fun varName bv = ppMap (checkVarName bv) ppString - - fun checkFunctionName bv s = - if mem s specialTokens orelse vName bv s then "(" ^ s ^ ")" else s - - fun functionName bv = ppMap (checkFunctionName bv) ppString - - fun isI tm = Option.isSome (destI tm) - - fun stripNeg (tm as Fn (f,[a])) = - if f <> neg then (0,tm) +fun pp inputTerm = + let + val quants = !binders + and iOps = !infixes + and neg = !negation + and bracks = !brackets + + val bracks = map (fn (b1,b2) => (b1 ^ b2, b1, b2)) bracks + + val bTokens = map #2 bracks @ map #3 bracks + + val iTokens = Print.tokensInfixes iOps + + fun destI tm = + case tm of + Fn (f,[a,b]) => + let + val f = Name.toString f + in + if StringSet.member f iTokens then SOME (f,a,b) else NONE + end + | _ => NONE + + val iPrinter = Print.ppInfixes iOps destI + + val specialTokens = + StringSet.addList iTokens (neg :: quants @ ["$","(",")"] @ bTokens) + + fun vName bv s = StringSet.member s bv + + fun checkVarName bv n = + let + val s = Name.toString n + in + if vName bv s then s else "$" ^ s + end + + fun varName bv = Print.ppMap (checkVarName bv) Print.ppString + + fun checkFunctionName bv n = + let + val s = Name.toString n + in + if StringSet.member s specialTokens orelse vName bv s then + "(" ^ s ^ ")" + else + s + end + + fun functionName bv = Print.ppMap (checkFunctionName bv) Print.ppString + + fun isI tm = Option.isSome (destI tm) + + fun stripNeg tm = + case tm of + Fn (f,[a]) => + if Name.toString f <> neg then (0,tm) else let val (n,tm) = stripNeg a in (n + 1, tm) end - | stripNeg tm = (0,tm) - - val destQuant = - let - fun dest q (Fn (q', [Var v, body])) = - if q <> q' then NONE - else - (case dest q body of - NONE => SOME (q,v,[],body) - | SOME (_,v',vs,body) => SOME (q, v, v' :: vs, body)) - | dest _ _ = NONE - in - fn tm => Useful.first (fn q => dest q tm) quants - end - - fun isQuant tm = Option.isSome (destQuant tm) - - fun destBrack (Fn (b,[tm])) = - (case List.find (fn (n,_,_) => n = b) bracks of - NONE => NONE - | SOME (_,b1,b2) => SOME (b1,tm,b2)) - | destBrack _ = NONE - - fun isBrack tm = Option.isSome (destBrack tm) - - fun functionArgument bv ppstrm tm = - (addBreak ppstrm (1,0); - if isBrack tm then customBracket bv ppstrm tm - else if isVar tm orelse isConst tm then basic bv ppstrm tm - else bracket bv ppstrm tm) - - and basic bv ppstrm (Var v) = varName bv ppstrm v - | basic bv ppstrm (Fn (f,args)) = - (beginBlock ppstrm Inconsistent 2; - functionName bv ppstrm f; - app (functionArgument bv ppstrm) args; - endBlock ppstrm) - - and customBracket bv ppstrm tm = - case destBrack tm of - SOME (b1,tm,b2) => ppBracket b1 b2 (term bv) ppstrm tm - | NONE => basic bv ppstrm tm - - and innerQuant bv ppstrm tm = - case destQuant tm of - NONE => term bv ppstrm tm - | SOME (q,v,vs,tm) => - let - val bv = NameSet.addList (NameSet.add bv v) vs - in - addString ppstrm q; - varName bv ppstrm v; - app (fn v => (addBreak ppstrm (1,0); varName bv ppstrm v)) vs; - addString ppstrm "."; - addBreak ppstrm (1,0); - innerQuant bv ppstrm tm - end - - and quantifier bv ppstrm tm = - if not (isQuant tm) then customBracket bv ppstrm tm - else - (beginBlock ppstrm Inconsistent 2; - innerQuant bv ppstrm tm; - endBlock ppstrm) - - and molecule bv ppstrm (tm,r) = - let - val (n,tm) = stripNeg tm - in - beginBlock ppstrm Inconsistent n; - funpow n (fn () => addString ppstrm neg) (); - if isI tm orelse (r andalso isQuant tm) then bracket bv ppstrm tm - else quantifier bv ppstrm tm; - endBlock ppstrm - end - - and term bv ppstrm tm = iPrinter (molecule bv) ppstrm (tm,false) - - and bracket bv ppstrm tm = ppBracket "(" ")" (term bv) ppstrm tm - in - term NameSet.empty - end inputPpstrm inputTerm; -end; - -fun toString tm = Parser.toString pp tm; + | _ => (0,tm) + + val destQuant = + let + fun dest q (Fn (q', [Var v, body])) = + if Name.toString q' <> q then NONE + else + (case dest q body of + NONE => SOME (q,v,[],body) + | SOME (_,v',vs,body) => SOME (q, v, v' :: vs, body)) + | dest _ _ = NONE + in + fn tm => Useful.first (fn q => dest q tm) quants + end + + fun isQuant tm = Option.isSome (destQuant tm) + + fun destBrack (Fn (b,[tm])) = + let + val s = Name.toString b + in + case List.find (fn (n,_,_) => n = s) bracks of + NONE => NONE + | SOME (_,b1,b2) => SOME (b1,tm,b2) + end + | destBrack _ = NONE + + fun isBrack tm = Option.isSome (destBrack tm) + + fun functionArgument bv tm = + Print.sequence + (Print.addBreak 1) + (if isBrack tm then customBracket bv tm + else if isVar tm orelse isConst tm then basic bv tm + else bracket bv tm) + + and basic bv (Var v) = varName bv v + | basic bv (Fn (f,args)) = + Print.blockProgram Print.Inconsistent 2 + (functionName bv f :: map (functionArgument bv) args) + + and customBracket bv tm = + case destBrack tm of + SOME (b1,tm,b2) => Print.ppBracket b1 b2 (term bv) tm + | NONE => basic bv tm + + and innerQuant bv tm = + case destQuant tm of + NONE => term bv tm + | SOME (q,v,vs,tm) => + let + val bv = StringSet.addList bv (map Name.toString (v :: vs)) + in + Print.program + [Print.addString q, + varName bv v, + Print.program + (map (Print.sequence (Print.addBreak 1) o varName bv) vs), + Print.addString ".", + Print.addBreak 1, + innerQuant bv tm] + end + + and quantifier bv tm = + if not (isQuant tm) then customBracket bv tm + else Print.block Print.Inconsistent 2 (innerQuant bv tm) + + and molecule bv (tm,r) = + let + val (n,tm) = stripNeg tm + in + Print.blockProgram Print.Inconsistent n + [Print.duplicate n (Print.addString neg), + if isI tm orelse (r andalso isQuant tm) then bracket bv tm + else quantifier bv tm] + end + + and term bv tm = iPrinter (molecule bv) (tm,false) + + and bracket bv tm = Print.ppBracket "(" ")" (term bv) tm + in + term StringSet.empty + end inputTerm; + +val toString = Print.toString pp; (* Parsing *) local - open Parser; + open Parse; infixr 9 >>++ infixr 8 ++ @@ -6383,7 +9752,7 @@ val symbolToken = let fun isNeg c = str c = !negation - + val symbolChars = explode "<>=-*+/\\?@|!$%&#^:;~" fun isSymbol c = mem c symbolChars @@ -6424,42 +9793,50 @@ fun possibleVarName "" = false | possibleVarName s = isAlphaNum (String.sub (s,0)) - fun vName bv s = NameSet.member s bv - - val iTokens = infixTokens iOps - - val iParser = parseInfixes iOps (fn (f,a,b) => Fn (f,[a,b])) - - val specialTokens = neg :: quants @ ["$"] @ bTokens @ iTokens + fun vName bv s = StringSet.member s bv + + val iTokens = Print.tokensInfixes iOps + + val iParser = + parseInfixes iOps (fn (f,a,b) => Fn (Name.fromString f, [a,b])) + + val specialTokens = + StringSet.addList iTokens (neg :: quants @ ["$"] @ bTokens) fun varName bv = - Parser.some (vName bv) || - (exact "$" ++ some possibleVarName) >> (fn (_,s) => s) - - fun fName bv s = not (mem s specialTokens) andalso not (vName bv s) + some (vName bv) || + (some (Useful.equal "$") ++ some possibleVarName) >> snd + + fun fName bv s = + not (StringSet.member s specialTokens) andalso not (vName bv s) fun functionName bv = - Parser.some (fName bv) || - (exact "(" ++ any ++ exact ")") >> (fn (_,(s,_)) => s) + some (fName bv) || + (some (Useful.equal "(") ++ any ++ some (Useful.equal ")")) >> + (fn (_,(s,_)) => s) fun basic bv tokens = let - val var = varName bv >> Var - - val const = functionName bv >> (fn f => Fn (f,[])) + val var = varName bv >> (Var o Name.fromString) + + val const = + functionName bv >> (fn f => Fn (Name.fromString f, [])) fun bracket (ab,a,b) = - (exact a ++ term bv ++ exact b) >> - (fn (_,(tm,_)) => if ab = "()" then tm else Fn (ab,[tm])) + (some (Useful.equal a) ++ term bv ++ some (Useful.equal b)) >> + (fn (_,(tm,_)) => + if ab = "()" then tm else Fn (Name.fromString ab, [tm])) fun quantifier q = let - fun bind (v,t) = Fn (q, [Var v, t]) + fun bind (v,t) = + Fn (Name.fromString q, [Var (Name.fromString v), t]) in - (exact q ++ atLeastOne (some possibleVarName) ++ - exact ".") >>++ + (some (Useful.equal q) ++ + atLeastOne (some possibleVarName) ++ + some (Useful.equal ".")) >>++ (fn (_,(vs,_)) => - term (NameSet.addList bv vs) >> + term (StringSet.addList bv vs) >> (fn body => foldr bind body vs)) end in @@ -6471,18 +9848,20 @@ and molecule bv tokens = let - val negations = many (exact neg) >> length + val negations = many (some (Useful.equal neg)) >> length val function = - (functionName bv ++ many (basic bv)) >> Fn || basic bv + (functionName bv ++ many (basic bv)) >> + (fn (f,args) => Fn (Name.fromString f, args)) || + basic bv in (negations ++ function) >> - (fn (n,tm) => funpow n (fn t => Fn (neg,[t])) tm) + (fn (n,tm) => funpow n (fn t => Fn (Name.fromString neg, [t])) tm) end tokens and term bv tokens = iParser (molecule bv) tokens in - term NameSet.empty + term StringSet.empty end inputStream; in fun fromString input = @@ -6495,15 +9874,14 @@ in case Stream.toList terms of [tm] => tm - | _ => raise Error "Syntax.stringToTerm" - end; -end; - -local - val antiquotedTermToString = - Parser.toString (Parser.ppBracket "(" ")" pp); -in - val parse = Parser.parseQuotation antiquotedTermToString fromString; + | _ => raise Error "Term.fromString" + end; +end; + +local + val antiquotedTermToString = Print.toString (Print.ppBracket "(" ")" pp); +in + val parse = Parse.parseQuotation antiquotedTermToString fromString; end; end @@ -6511,16 +9889,16 @@ structure TermOrdered = struct type t = Term.term val compare = Term.compare end -structure TermSet = ElementSet (TermOrdered); - structure TermMap = KeyMap (TermOrdered); + +structure TermSet = ElementSet (TermMap); end; (**** Original file: Subst.sig ****) (* ========================================================================= *) (* FIRST ORDER LOGIC SUBSTITUTIONS *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Subst = @@ -6548,8 +9926,6 @@ val singleton : Metis.Term.var * Metis.Term.term -> subst -val union : subst -> subst -> subst - val toList : subst -> (Metis.Term.var * Metis.Term.term) list val fromList : (Metis.Term.var * Metis.Term.term) list -> subst @@ -6558,7 +9934,7 @@ val foldr : (Metis.Term.var * Metis.Term.term * 'a -> 'a) -> 'a -> subst -> 'a -val pp : subst Metis.Parser.pp +val pp : subst Metis.Print.pp val toString : subst -> string @@ -6591,7 +9967,13 @@ val compose : subst -> subst -> subst (* ------------------------------------------------------------------------- *) -(* Substitutions can be inverted iff they are renaming substitutions. *) +(* Creating the union of two compatible substitutions. *) +(* ------------------------------------------------------------------------- *) + +val union : subst -> subst -> subst (* raises Error *) + +(* ------------------------------------------------------------------------- *) +(* Substitutions can be inverted iff they are renaming substitutions. *) (* ------------------------------------------------------------------------- *) val invert : subst -> subst (* raises Error *) @@ -6605,6 +9987,22 @@ val freshVars : Metis.NameSet.set -> subst (* ------------------------------------------------------------------------- *) +(* Free variables. *) +(* ------------------------------------------------------------------------- *) + +val redexes : subst -> Metis.NameSet.set + +val residueFreeVars : subst -> Metis.NameSet.set + +val freeVars : subst -> Metis.NameSet.set + +(* ------------------------------------------------------------------------- *) +(* Functions. *) +(* ------------------------------------------------------------------------- *) + +val functions : subst -> Metis.NameAritySet.set + +(* ------------------------------------------------------------------------- *) (* Matching for first order logic terms. *) (* ------------------------------------------------------------------------- *) @@ -6622,7 +10020,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; @@ -6631,7 +10029,7 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC SUBSTITUTIONS *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Subst :> Subst = @@ -6661,16 +10059,6 @@ fun singleton v_tm = insert empty v_tm; -local - fun compatible (tm1,tm2) = - if tm1 = tm2 then SOME tm1 else raise Error "Subst.union: incompatible"; -in - fun union (s1 as Subst m1) (s2 as Subst m2) = - if NameMap.null m1 then s2 - else if NameMap.null m2 then s1 - else Subst (NameMap.union compatible m1 m2); -end; - fun toList (Subst m) = NameMap.toList m; fun fromList l = Subst (NameMap.fromList l); @@ -6679,12 +10067,12 @@ fun foldr f b (Subst m) = NameMap.foldr f b m; -fun pp ppstrm sub = - Parser.ppBracket "<[" "]>" - (Parser.ppSequence "," (Parser.ppBinop " |->" Parser.ppString Term.pp)) - ppstrm (toList sub); - -val toString = Parser.toString pp; +fun pp sub = + Print.ppBracket "<[" "]>" + (Print.ppOpList "," (Print.ppOp2 " |->" Name.pp Term.pp)) + (toList sub); + +val toString = Print.toString pp; (* ------------------------------------------------------------------------- *) (* Normalizing removes identity substitutions. *) @@ -6709,13 +10097,13 @@ let fun tmSub (tm as Term.Var v) = (case peek sub v of - SOME tm' => if Sharing.pointerEqual (tm,tm') then tm else tm' + SOME tm' => if Portable.pointerEqual (tm,tm') then tm else tm' | NONE => tm) | tmSub (tm as Term.Fn (f,args)) = let val args' = Sharing.map tmSub args in - if Sharing.pointerEqual (args,args') then tm + if Portable.pointerEqual (args,args') then tm else Term.Fn (f,args') end in @@ -6758,7 +10146,22 @@ end; (* ------------------------------------------------------------------------- *) -(* Substitutions can be inverted iff they are renaming substitutions. *) +(* Creating the union of two compatible substitutions. *) +(* ------------------------------------------------------------------------- *) + +local + fun compatible ((_,tm1),(_,tm2)) = + if Term.equal tm1 tm2 then SOME tm1 + else raise Error "Subst.union: incompatible"; +in + fun union (s1 as Subst m1) (s2 as Subst m2) = + if NameMap.null m1 then s2 + else if NameMap.null m2 then s1 + else Subst (NameMap.union compatible m1 m2); +end; + +(* ------------------------------------------------------------------------- *) +(* Substitutions can be inverted iff they are renaming substitutions. *) (* ------------------------------------------------------------------------- *) local @@ -6784,6 +10187,42 @@ end; (* ------------------------------------------------------------------------- *) +(* Free variables. *) +(* ------------------------------------------------------------------------- *) + +val redexes = + let + fun add (v,_,s) = NameSet.add s v + in + foldl add NameSet.empty + end; + +val residueFreeVars = + let + fun add (_,t,s) = NameSet.union s (Term.freeVars t) + in + foldl add NameSet.empty + end; + +val freeVars = + let + fun add (v,t,s) = NameSet.union (NameSet.add s v) (Term.freeVars t) + in + foldl add NameSet.empty + end; + +(* ------------------------------------------------------------------------- *) +(* Functions. *) +(* ------------------------------------------------------------------------- *) + +val functions = + let + fun add (_,t,s) = NameAritySet.union s (Term.functions t) + in + foldl add NameAritySet.empty + end; + +(* ------------------------------------------------------------------------- *) (* Matching for first order logic terms. *) (* ------------------------------------------------------------------------- *) @@ -6795,13 +10234,13 @@ case peek sub v of NONE => insert sub (v,tm) | SOME tm' => - if tm = tm' then sub + if Term.equal tm tm' then sub else raise Error "Subst.match: incompatible matches" in matchList sub rest end | matchList sub ((Term.Fn (f1,args1), Term.Fn (f2,args2)) :: rest) = - if f1 = f2 andalso length args1 = length args2 then + if Name.equal f1 f2 andalso length args1 = length args2 then matchList sub (zip args1 args2 @ rest) else raise Error "Subst.match: different structure" | matchList _ _ = raise Error "Subst.match: functions can't match vars"; @@ -6828,7 +10267,7 @@ | SOME tm' => solve' sub tm' tm rest) | solve' sub tm1 (tm2 as Term.Var _) rest = solve' sub tm2 tm1 rest | solve' sub (Term.Fn (f1,args1)) (Term.Fn (f2,args2)) rest = - if f1 = f2 andalso length args1 = length args2 then + if Name.equal f1 f2 andalso length args1 = length args2 then solve sub (zip args1 args2 @ rest) else raise Error "Subst.unify: different structure"; @@ -6843,7 +10282,7 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC ATOMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Atom = @@ -6895,6 +10334,8 @@ val compare : atom * atom -> order +val equal : atom -> atom -> bool + (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) @@ -6937,6 +10378,8 @@ (* The equality relation. *) (* ------------------------------------------------------------------------- *) +val eqRelationName : relationName + val eqRelation : relation val mkEq : Metis.Term.term * Metis.Term.term -> atom @@ -6969,13 +10412,13 @@ (* Parsing and pretty printing. *) (* ------------------------------------------------------------------------- *) -val pp : atom Metis.Parser.pp +val pp : atom Metis.Print.pp val toString : atom -> string val fromString : string -> atom -val parse : Metis.Term.term Metis.Parser.quotation -> atom +val parse : Metis.Term.term Metis.Parse.quotation -> atom end @@ -6983,7 +10426,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; @@ -6992,7 +10435,7 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC ATOMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Atom :> Atom = @@ -7041,7 +10484,7 @@ fun mkBinop p (a,b) : atom = (p,[a,b]); fun destBinop p (x,[a,b]) = - if x = p then (a,b) else raise Error "Atom.destBinop: wrong binop" + if Name.equal x p then (a,b) else raise Error "Atom.destBinop: wrong binop" | destBinop _ _ = raise Error "Atom.destBinop: not a binop"; fun isBinop p = can (destBinop p); @@ -7062,6 +10505,8 @@ | EQUAL => lexCompare Term.compare (tms1,tms2) | GREATER => GREATER; +fun equal atm1 atm2 = compare (atm1,atm2) = EQUAL; + (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) @@ -7086,7 +10531,7 @@ val tm = List.nth (tms,h) val tm' = Term.replace tm (t,res) in - if Sharing.pointerEqual (tm,tm') then atm + if Portable.pointerEqual (tm,tm') then atm else (rel, updateNth (h,tm') tms) end; @@ -7121,7 +10566,7 @@ let val tms' = Sharing.map (Subst.subst sub) tms in - if Sharing.pointerEqual (tms',tms) then atm else (p,tms') + if Portable.pointerEqual (tms',tms) then atm else (p,tms') end; (* ------------------------------------------------------------------------- *) @@ -7133,7 +10578,7 @@ in fun match sub (p1,tms1) (p2,tms2) = let - val _ = (p1 = p2 andalso length tms1 = length tms2) orelse + val _ = (Name.equal p1 p2 andalso length tms1 = length tms2) orelse raise Error "Atom.match" in foldl matchArg sub (zip tms1 tms2) @@ -7149,7 +10594,7 @@ in fun unify sub (p1,tms1) (p2,tms2) = let - val _ = (p1 = p2 andalso length tms1 = length tms2) orelse + val _ = (Name.equal p1 p2 andalso length tms1 = length tms2) orelse raise Error "Atom.unify" in foldl unifyArg sub (zip tms1 tms2) @@ -7160,24 +10605,24 @@ (* The equality relation. *) (* ------------------------------------------------------------------------- *) -val eqName = "="; - -val eqArity = 2; - -val eqRelation = (eqName,eqArity); - -val mkEq = mkBinop eqName; - -fun destEq x = destBinop eqName x; - -fun isEq x = isBinop eqName x; +val eqRelationName = Name.fromString "="; + +val eqRelationArity = 2; + +val eqRelation = (eqRelationName,eqRelationArity); + +val mkEq = mkBinop eqRelationName; + +fun destEq x = destBinop eqRelationName x; + +fun isEq x = isBinop eqRelationName x; fun mkRefl tm = mkEq (tm,tm); fun destRefl atm = let val (l,r) = destEq atm - val _ = l = r orelse raise Error "Atom.destRefl" + val _ = Term.equal l r orelse raise Error "Atom.destRefl" in l end; @@ -7187,7 +10632,7 @@ fun sym atm = let val (l,r) = destEq atm - val _ = l <> r orelse raise Error "Atom.sym: refl" + val _ = not (Term.equal l r) orelse raise Error "Atom.sym: refl" in mkEq (r,l) end; @@ -7219,29 +10664,29 @@ (* Parsing and pretty printing. *) (* ------------------------------------------------------------------------- *) -val pp = Parser.ppMap Term.Fn Term.pp; - -val toString = Parser.toString pp; +val pp = Print.ppMap Term.Fn Term.pp; + +val toString = Print.toString pp; fun fromString s = Term.destFn (Term.fromString s); -val parse = Parser.parseQuotation Term.toString fromString; +val parse = Parse.parseQuotation Term.toString fromString; end structure AtomOrdered = struct type t = Atom.atom val compare = Atom.compare end -structure AtomSet = ElementSet (AtomOrdered); - structure AtomMap = KeyMap (AtomOrdered); + +structure AtomSet = ElementSet (AtomMap); end; (**** Original file: Formula.sig ****) (* ========================================================================= *) (* FIRST ORDER LOGIC FORMULAS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Formula = @@ -7275,6 +10720,10 @@ val isBoolean : formula -> bool +val isTrue : formula -> bool + +val isFalse : formula -> bool + (* Functions *) val functions : formula -> Metis.NameAritySet.set @@ -7333,6 +10782,8 @@ val listMkForall : Metis.Term.var list * formula -> formula +val setMkForall : Metis.NameSet.set * formula -> formula + val stripForall : formula -> Metis.Term.var list * formula (* Existential quantification *) @@ -7343,6 +10794,8 @@ val listMkExists : Metis.Term.var list * formula -> formula +val setMkExists : Metis.NameSet.set * formula -> formula + val stripExists : formula -> Metis.Term.var list * formula (* ------------------------------------------------------------------------- *) @@ -7357,6 +10810,8 @@ val compare : formula * formula -> order +val equal : formula -> formula -> bool + (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) @@ -7365,6 +10820,8 @@ val freeVars : formula -> Metis.NameSet.set +val freeVarsList : formula list -> Metis.NameSet.set + val specialize : formula -> formula val generalize : formula -> formula @@ -7404,12 +10861,18 @@ val rhs : formula -> Metis.Term.term (* ------------------------------------------------------------------------- *) +(* Splitting goals. *) +(* ------------------------------------------------------------------------- *) + +val splitGoal : formula -> formula list + +(* ------------------------------------------------------------------------- *) (* Parsing and pretty-printing. *) (* ------------------------------------------------------------------------- *) -type quotation = formula Metis.Parser.quotation - -val pp : formula Metis.Parser.pp +type quotation = formula Metis.Parse.quotation + +val pp : formula Metis.Print.pp val toString : formula -> string @@ -7423,7 +10886,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; @@ -7432,7 +10895,7 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC FORMULAS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Formula :> Formula = @@ -7471,6 +10934,16 @@ val isBoolean = can destBoolean; +fun isTrue fm = + case fm of + True => true + | _ => false; + +fun isFalse fm = + case fm of + False => true + | _ => false; + (* Functions *) local @@ -7643,6 +11116,8 @@ fun listMkForall ([],body) = body | listMkForall (v :: vs, body) = Forall (v, listMkForall (vs,body)); +fun setMkForall (vs,body) = NameSet.foldr Forall body vs; + local fun strip vs (Forall (v,b)) = strip (v :: vs) b | strip vs tm = (rev vs, tm); @@ -7660,6 +11135,8 @@ fun listMkExists ([],body) = body | listMkExists (v :: vs, body) = Exists (v, listMkExists (vs,body)); +fun setMkExists (vs,body) = NameSet.foldr Exists body vs; + local fun strip vs (Exists (v,b)) = strip (v :: vs) b | strip vs tm = (rev vs, tm); @@ -7693,50 +11170,56 @@ local fun cmp [] = EQUAL - | cmp ((True,True) :: l) = cmp l - | cmp ((True,_) :: _) = LESS - | cmp ((_,True) :: _) = GREATER - | cmp ((False,False) :: l) = cmp l - | cmp ((False,_) :: _) = LESS - | cmp ((_,False) :: _) = GREATER - | cmp ((Atom atm1, Atom atm2) :: l) = - (case Atom.compare (atm1,atm2) of - LESS => LESS - | EQUAL => cmp l - | GREATER => GREATER) - | cmp ((Atom _, _) :: _) = LESS - | cmp ((_, Atom _) :: _) = GREATER - | cmp ((Not p1, Not p2) :: l) = cmp ((p1,p2) :: l) - | cmp ((Not _, _) :: _) = LESS - | cmp ((_, Not _) :: _) = GREATER - | cmp ((And (p1,q1), And (p2,q2)) :: l) = cmp ((p1,p2) :: (q1,q2) :: l) - | cmp ((And _, _) :: _) = LESS - | cmp ((_, And _) :: _) = GREATER - | cmp ((Or (p1,q1), Or (p2,q2)) :: l) = cmp ((p1,p2) :: (q1,q2) :: l) - | cmp ((Or _, _) :: _) = LESS - | cmp ((_, Or _) :: _) = GREATER - | cmp ((Imp (p1,q1), Imp (p2,q2)) :: l) = cmp ((p1,p2) :: (q1,q2) :: l) - | cmp ((Imp _, _) :: _) = LESS - | cmp ((_, Imp _) :: _) = GREATER - | cmp ((Iff (p1,q1), Iff (p2,q2)) :: l) = cmp ((p1,p2) :: (q1,q2) :: l) - | cmp ((Iff _, _) :: _) = LESS - | cmp ((_, Iff _) :: _) = GREATER - | cmp ((Forall (v1,p1), Forall (v2,p2)) :: l) = - (case Name.compare (v1,v2) of - LESS => LESS - | EQUAL => cmp ((p1,p2) :: l) - | GREATER => GREATER) - | cmp ((Forall _, Exists _) :: _) = LESS - | cmp ((Exists _, Forall _) :: _) = GREATER - | cmp ((Exists (v1,p1), Exists (v2,p2)) :: l) = - (case Name.compare (v1,v2) of - LESS => LESS - | EQUAL => cmp ((p1,p2) :: l) - | GREATER => GREATER); + | cmp (f1_f2 :: fs) = + if Portable.pointerEqual f1_f2 then cmp fs + else + case f1_f2 of + (True,True) => cmp fs + | (True,_) => LESS + | (_,True) => GREATER + | (False,False) => cmp fs + | (False,_) => LESS + | (_,False) => GREATER + | (Atom atm1, Atom atm2) => + (case Atom.compare (atm1,atm2) of + LESS => LESS + | EQUAL => cmp fs + | GREATER => GREATER) + | (Atom _, _) => LESS + | (_, Atom _) => GREATER + | (Not p1, Not p2) => cmp ((p1,p2) :: fs) + | (Not _, _) => LESS + | (_, Not _) => GREATER + | (And (p1,q1), And (p2,q2)) => cmp ((p1,p2) :: (q1,q2) :: fs) + | (And _, _) => LESS + | (_, And _) => GREATER + | (Or (p1,q1), Or (p2,q2)) => cmp ((p1,p2) :: (q1,q2) :: fs) + | (Or _, _) => LESS + | (_, Or _) => GREATER + | (Imp (p1,q1), Imp (p2,q2)) => cmp ((p1,p2) :: (q1,q2) :: fs) + | (Imp _, _) => LESS + | (_, Imp _) => GREATER + | (Iff (p1,q1), Iff (p2,q2)) => cmp ((p1,p2) :: (q1,q2) :: fs) + | (Iff _, _) => LESS + | (_, Iff _) => GREATER + | (Forall (v1,p1), Forall (v2,p2)) => + (case Name.compare (v1,v2) of + LESS => LESS + | EQUAL => cmp ((p1,p2) :: fs) + | GREATER => GREATER) + | (Forall _, Exists _) => LESS + | (Exists _, Forall _) => GREATER + | (Exists (v1,p1), Exists (v2,p2)) => + (case Name.compare (v1,v2) of + LESS => LESS + | EQUAL => cmp ((p1,p2) :: fs) + | GREATER => GREATER); in fun compare fm1_fm2 = cmp [fm1_fm2]; end; +fun equal fm1 fm2 = compare (fm1,fm2) = EQUAL; + (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) @@ -7752,8 +11235,10 @@ | f (Or (p,q) :: fms) = f (p :: q :: fms) | f (Imp (p,q) :: fms) = f (p :: q :: fms) | f (Iff (p,q) :: fms) = f (p :: q :: fms) - | f (Forall (w,p) :: fms) = if v = w then f fms else f (p :: fms) - | f (Exists (w,p) :: fms) = if v = w then f fms else f (p :: fms) + | f (Forall (w,p) :: fms) = + if Name.equal v w then f fms else f (p :: fms) + | f (Exists (w,p) :: fms) = + if Name.equal v w then f fms else f (p :: fms) in fn fm => f [fm] end; @@ -7771,8 +11256,12 @@ | fv vs ((bv, Iff (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms) | fv vs ((bv, Forall (v,p)) :: fms) = fv vs ((NameSet.add bv v, p) :: fms) | fv vs ((bv, Exists (v,p)) :: fms) = fv vs ((NameSet.add bv v, p) :: fms); -in - fun freeVars fm = fv NameSet.empty [(NameSet.empty,fm)]; + + fun add (fm,vs) = fv vs [(NameSet.empty,fm)]; +in + fun freeVars fm = add (fm,NameSet.empty); + + fun freeVarsList fms = List.foldl add NameSet.empty fms; end; fun specialize fm = snd (stripForall fm); @@ -7794,13 +11283,13 @@ let val tms' = Sharing.map (Subst.subst sub) tms in - if Sharing.pointerEqual (tms,tms') then fm else Atom (p,tms') + if Portable.pointerEqual (tms,tms') then fm else Atom (p,tms') end | Not p => let val p' = substFm sub p in - if Sharing.pointerEqual (p,p') then fm else Not p' + if Portable.pointerEqual (p,p') then fm else Not p' end | And (p,q) => substConn sub fm And p q | Or (p,q) => substConn sub fm Or p q @@ -7814,8 +11303,8 @@ val p' = substFm sub p and q' = substFm sub q in - if Sharing.pointerEqual (p,p') andalso - Sharing.pointerEqual (q,q') + if Portable.pointerEqual (p,p') andalso + Portable.pointerEqual (q,q') then fm else conn (p',q') end @@ -7825,12 +11314,12 @@ val v' = let fun f (w,s) = - if w = v then s + if Name.equal w v then s else case Subst.peek sub w of NONE => NameSet.add s w | SOME tm => NameSet.union s (Term.freeVars tm) - + val vars = freeVars p val vars = NameSet.foldl f NameSet.empty vars in @@ -7838,12 +11327,12 @@ end val sub = - if v = v' then Subst.remove sub (NameSet.singleton v) + if Name.equal v v' then Subst.remove sub (NameSet.singleton v) else Subst.insert sub (v, Term.Var v') val p' = substCheck sub p in - if v = v' andalso Sharing.pointerEqual (p,p') then fm + if Name.equal v v' andalso Portable.pointerEqual (p,p') then fm else quant (v',p') end; in @@ -7883,34 +11372,39 @@ (* Parsing and pretty-printing. *) (* ------------------------------------------------------------------------- *) -type quotation = formula Parser.quotation - -val truthSymbol = "T" -and falsitySymbol = "F" -and conjunctionSymbol = "/\\" -and disjunctionSymbol = "\\/" -and implicationSymbol = "==>" -and equivalenceSymbol = "<=>" -and universalSymbol = "!" -and existentialSymbol = "?"; - -local - fun demote True = Term.Fn (truthSymbol,[]) - | demote False = Term.Fn (falsitySymbol,[]) +type quotation = formula Parse.quotation; + +val truthName = Name.fromString "T" +and falsityName = Name.fromString "F" +and conjunctionName = Name.fromString "/\\" +and disjunctionName = Name.fromString "\\/" +and implicationName = Name.fromString "==>" +and equivalenceName = Name.fromString "<=>" +and universalName = Name.fromString "!" +and existentialName = Name.fromString "?"; + +local + fun demote True = Term.Fn (truthName,[]) + | demote False = Term.Fn (falsityName,[]) | demote (Atom (p,tms)) = Term.Fn (p,tms) - | demote (Not p) = Term.Fn (!Term.negation, [demote p]) - | demote (And (p,q)) = Term.Fn (conjunctionSymbol, [demote p, demote q]) - | demote (Or (p,q)) = Term.Fn (disjunctionSymbol, [demote p, demote q]) - | demote (Imp (p,q)) = Term.Fn (implicationSymbol, [demote p, demote q]) - | demote (Iff (p,q)) = Term.Fn (equivalenceSymbol, [demote p, demote q]) - | demote (Forall (v,b)) = Term.Fn (universalSymbol, [Term.Var v, demote b]) + | demote (Not p) = + let + val Unsynchronized.ref s = Term.negation + in + Term.Fn (Name.fromString s, [demote p]) + end + | demote (And (p,q)) = Term.Fn (conjunctionName, [demote p, demote q]) + | demote (Or (p,q)) = Term.Fn (disjunctionName, [demote p, demote q]) + | demote (Imp (p,q)) = Term.Fn (implicationName, [demote p, demote q]) + | demote (Iff (p,q)) = Term.Fn (equivalenceName, [demote p, demote q]) + | demote (Forall (v,b)) = Term.Fn (universalName, [Term.Var v, demote b]) | demote (Exists (v,b)) = - Term.Fn (existentialSymbol, [Term.Var v, demote b]); -in - fun pp ppstrm fm = Term.pp ppstrm (demote fm); -end; - -val toString = Parser.toString pp; + Term.Fn (existentialName, [Term.Var v, demote b]); +in + fun pp fm = Term.pp (demote fm); +end; + +val toString = Print.toString pp; local fun isQuant [Term.Var _, _] = true @@ -7918,23 +11412,23 @@ fun promote (Term.Var v) = Atom (v,[]) | promote (Term.Fn (f,tms)) = - if f = truthSymbol andalso null tms then + if Name.equal f truthName andalso null tms then True - else if f = falsitySymbol andalso null tms then + else if Name.equal f falsityName andalso null tms then False - else if f = !Term.negation andalso length tms = 1 then + else if Name.toString f = !Term.negation andalso length tms = 1 then Not (promote (hd tms)) - else if f = conjunctionSymbol andalso length tms = 2 then + else if Name.equal f conjunctionName andalso length tms = 2 then And (promote (hd tms), promote (List.nth (tms,1))) - else if f = disjunctionSymbol andalso length tms = 2 then + else if Name.equal f disjunctionName andalso length tms = 2 then Or (promote (hd tms), promote (List.nth (tms,1))) - else if f = implicationSymbol andalso length tms = 2 then + else if Name.equal f implicationName andalso length tms = 2 then Imp (promote (hd tms), promote (List.nth (tms,1))) - else if f = equivalenceSymbol andalso length tms = 2 then + else if Name.equal f equivalenceName andalso length tms = 2 then Iff (promote (hd tms), promote (List.nth (tms,1))) - else if f = universalSymbol andalso isQuant tms then + else if Name.equal f universalName andalso isQuant tms then Forall (Term.destVar (hd tms), promote (List.nth (tms,1))) - else if f = existentialSymbol andalso isQuant tms then + else if Name.equal f existentialName andalso isQuant tms then Exists (Term.destVar (hd tms), promote (List.nth (tms,1))) else Atom (f,tms); @@ -7942,16 +11436,71 @@ fun fromString s = promote (Term.fromString s); end; -val parse = Parser.parseQuotation toString fromString; - -end +val parse = Parse.parseQuotation toString fromString; + +(* ------------------------------------------------------------------------- *) +(* Splitting goals. *) +(* ------------------------------------------------------------------------- *) + +local + fun add_asms asms goal = + if null asms then goal else Imp (listMkConj (rev asms), goal); + + fun add_var_asms asms v goal = add_asms asms (Forall (v,goal)); + + fun split asms pol fm = + case (pol,fm) of + (* Positive splittables *) + (true,True) => [] + | (true, Not f) => split asms false f + | (true, And (f1,f2)) => split asms true f1 @ split (f1 :: asms) true f2 + | (true, Or (f1,f2)) => split (Not f1 :: asms) true f2 + | (true, Imp (f1,f2)) => split (f1 :: asms) true f2 + | (true, Iff (f1,f2)) => + split (f1 :: asms) true f2 @ split (f2 :: asms) true f1 + | (true, Forall (v,f)) => map (add_var_asms asms v) (split [] true f) + (* Negative splittables *) + | (false,False) => [] + | (false, Not f) => split asms true f + | (false, And (f1,f2)) => split (f1 :: asms) false f2 + | (false, Or (f1,f2)) => + split asms false f1 @ split (Not f1 :: asms) false f2 + | (false, Imp (f1,f2)) => split asms true f1 @ split (f1 :: asms) false f2 + | (false, Iff (f1,f2)) => + split (f1 :: asms) false f2 @ split (f2 :: asms) false f1 + | (false, Exists (v,f)) => map (add_var_asms asms v) (split [] false f) + (* Unsplittables *) + | _ => [add_asms asms (if pol then fm else Not fm)]; +in + fun splitGoal fm = split [] true fm; +end; + +(*MetisTrace3 +val splitGoal = fn fm => + let + val result = splitGoal fm + val () = Print.trace pp "Formula.splitGoal: fm" fm + val () = Print.trace (Print.ppList pp) "Formula.splitGoal: result" result + in + result + end; +*) + +end + +structure FormulaOrdered = +struct type t = Formula.formula val compare = Formula.compare end + +structure FormulaMap = KeyMap (FormulaOrdered); + +structure FormulaSet = ElementSet (FormulaMap); end; (**** Original file: Literal.sig ****) (* ========================================================================= *) (* FIRST ORDER LOGIC LITERALS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Literal = @@ -8017,6 +11566,8 @@ val compare : literal * literal -> order (* negative < positive *) +val equal : literal -> literal -> bool + (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) @@ -8103,13 +11654,13 @@ (* Parsing and pretty-printing. *) (* ------------------------------------------------------------------------- *) -val pp : literal Metis.Parser.pp +val pp : literal Metis.Print.pp val toString : literal -> string val fromString : string -> literal -val parse : Metis.Term.term Metis.Parser.quotation -> literal +val parse : Metis.Term.term Metis.Parse.quotation -> literal end @@ -8117,7 +11668,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; @@ -8126,7 +11677,7 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC LITERALS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Literal :> Literal = @@ -8196,11 +11747,9 @@ (* A total comparison function for literals. *) (* ------------------------------------------------------------------------- *) -fun compare ((pol1,atm1),(pol2,atm2)) = - case boolCompare (pol1,pol2) of - LESS => GREATER - | EQUAL => Atom.compare (atm1,atm2) - | GREATER => LESS; +val compare = prodCompare boolCompare Atom.compare; + +fun equal (p1,atm1) (p2,atm2) = p1 = p2 andalso Atom.equal atm1 atm2; (* ------------------------------------------------------------------------- *) (* Subterms. *) @@ -8214,7 +11763,7 @@ let val atm' = Atom.replace atm path_tm in - if Sharing.pointerEqual (atm,atm') then lit else (pol,atm') + if Portable.pointerEqual (atm,atm') then lit else (pol,atm') end; (* ------------------------------------------------------------------------- *) @@ -8233,7 +11782,7 @@ let val atm' = Atom.subst sub atm in - if Sharing.pointerEqual (atm',atm) then lit else (pol,atm') + if Portable.pointerEqual (atm',atm) then lit else (pol,atm') end; (* ------------------------------------------------------------------------- *) @@ -8308,24 +11857,26 @@ (* Parsing and pretty-printing. *) (* ------------------------------------------------------------------------- *) -val pp = Parser.ppMap toFormula Formula.pp; - -val toString = Parser.toString pp; +val pp = Print.ppMap toFormula Formula.pp; + +val toString = Print.toString pp; fun fromString s = fromFormula (Formula.fromString s); -val parse = Parser.parseQuotation Term.toString fromString; +val parse = Parse.parseQuotation Term.toString fromString; end structure LiteralOrdered = struct type t = Literal.literal val compare = Literal.compare end +structure LiteralMap = KeyMap (LiteralOrdered); + structure LiteralSet = struct local - structure S = ElementSet (LiteralOrdered); + structure S = ElementSet (LiteralMap); in open S; end; @@ -8353,6 +11904,8 @@ foldl f NameAritySet.empty end; + fun freeIn v = exists (Literal.freeIn v); + val freeVars = let fun f (lit,set) = NameSet.union set (Literal.freeVars lit) @@ -8360,6 +11913,13 @@ foldl f NameSet.empty end; + val freeVarsList = + let + fun f (lits,set) = NameSet.union set (freeVars lits) + in + List.foldl f NameSet.empty + end; + val symbols = let fun f (lit,z) = Literal.symbols lit + z @@ -8379,31 +11939,42 @@ fun substLit (lit,(eq,lits')) = let val lit' = Literal.subst sub lit - val eq = eq andalso Sharing.pointerEqual (lit,lit') + val eq = eq andalso Portable.pointerEqual (lit,lit') in (eq, add lits' lit') end - + val (eq,lits') = foldl substLit (true,empty) lits in if eq then lits else lits' end; + fun conjoin set = + Formula.listMkConj (List.map Literal.toFormula (toList set)); + + fun disjoin set = + Formula.listMkDisj (List.map Literal.toFormula (toList set)); + val pp = - Parser.ppMap + Print.ppMap toList - (Parser.ppBracket "{" "}" (Parser.ppSequence "," Literal.pp)); - -end - -structure LiteralMap = KeyMap (LiteralOrdered); + (Print.ppBracket "{" "}" (Print.ppOpList "," Literal.pp)); + +end + +structure LiteralSetOrdered = +struct type t = LiteralSet.set val compare = LiteralSet.compare end + +structure LiteralSetMap = KeyMap (LiteralSetOrdered); + +structure LiteralSetSet = ElementSet (LiteralSetMap); end; (**** Original file: Thm.sig ****) (* ========================================================================= *) -(* A LOGICAL KERNEL FOR FIRST ORDER CLAUSES *) -(* Copyright (c) 2001-2004 Joe Hurd, distributed under the BSD License *) +(* A LOGICAL KERNEL FOR FIRST ORDER CLAUSAL THEOREMS *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Thm = @@ -8413,6 +11984,12 @@ (* An abstract type of first order logic theorems. *) (* ------------------------------------------------------------------------- *) +type thm + +(* ------------------------------------------------------------------------- *) +(* Theorem destructors. *) +(* ------------------------------------------------------------------------- *) + type clause = Metis.LiteralSet.set datatype inferenceType = @@ -8424,14 +12001,8 @@ | Refl | Equality -type thm - type inference = inferenceType * thm list -(* ------------------------------------------------------------------------- *) -(* Theorem destructors. *) -(* ------------------------------------------------------------------------- *) - val clause : thm -> clause val inference : thm -> inference @@ -8482,11 +12053,11 @@ (* Pretty-printing. *) (* ------------------------------------------------------------------------- *) -val ppInferenceType : inferenceType Metis.Parser.pp +val ppInferenceType : inferenceType Metis.Print.pp val inferenceTypeToString : inferenceType -> string -val pp : thm Metis.Parser.pp +val pp : thm Metis.Print.pp val toString : thm -> string @@ -8554,7 +12125,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; @@ -8562,8 +12133,8 @@ val foldr = List.foldr; (* ========================================================================= *) -(* A LOGICAL KERNEL FOR FIRST ORDER CLAUSES *) -(* Copyright (c) 2001-2004 Joe Hurd, distributed under the BSD License *) +(* A LOGICAL KERNEL FOR FIRST ORDER CLAUSAL THEOREMS *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Thm :> Thm = @@ -8648,13 +12219,9 @@ (* Free variables. *) (* ------------------------------------------------------------------------- *) -fun freeIn v (Thm (cl,_)) = LiteralSet.exists (Literal.freeIn v) cl; - -local - fun free (lit,set) = NameSet.union (Literal.freeVars lit) set; -in - fun freeVars (Thm (cl,_)) = LiteralSet.foldl free NameSet.empty cl; -end; +fun freeIn v (Thm (cl,_)) = LiteralSet.freeIn v cl; + +fun freeVars (Thm (cl,_)) = LiteralSet.freeVars cl; (* ------------------------------------------------------------------------- *) (* Pretty-printing. *) @@ -8668,26 +12235,21 @@ | inferenceTypeToString Refl = "Refl" | inferenceTypeToString Equality = "Equality"; -fun ppInferenceType ppstrm inf = - Parser.ppString ppstrm (inferenceTypeToString inf); +fun ppInferenceType inf = + Print.ppString (inferenceTypeToString inf); local fun toFormula th = Formula.listMkDisj (map Literal.toFormula (LiteralSet.toList (clause th))); in - fun pp ppstrm th = - let - open PP - in - begin_block ppstrm INCONSISTENT 3; - add_string ppstrm "|- "; - Formula.pp ppstrm (toFormula th); - end_block ppstrm - end; -end; - -val toString = Parser.toString pp; + fun pp th = + Print.blockProgram Print.Inconsistent 3 + [Print.addString "|- ", + Formula.pp (toFormula th)]; +end; + +val toString = Print.toString pp; (* ------------------------------------------------------------------------- *) (* Primitive rules of inference. *) @@ -8720,7 +12282,7 @@ let val cl' = LiteralSet.subst sub cl in - if Sharing.pointerEqual (cl,cl') then th + if Portable.pointerEqual (cl,cl') then th else case inf of (Subst,_) => Thm (cl',inf) @@ -8744,7 +12306,7 @@ Thm (LiteralSet.union cl1' cl2', (Resolve,[th1,th2])) end; -(*DEBUG +(*MetisDebug val resolve = fn lit => fn pos => fn neg => resolve lit pos neg handle Error err => @@ -8790,7 +12352,7 @@ (* ========================================================================= *) (* PROOFS IN FIRST ORDER LOGIC *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Proof = @@ -8829,14 +12391,22 @@ val proof : Metis.Thm.thm -> proof (* ------------------------------------------------------------------------- *) +(* Free variables. *) +(* ------------------------------------------------------------------------- *) + +val freeIn : Metis.Term.var -> proof -> bool + +val freeVars : proof -> Metis.NameSet.set + +(* ------------------------------------------------------------------------- *) (* Printing. *) (* ------------------------------------------------------------------------- *) -val ppInference : inference Metis.Parser.pp +val ppInference : inference Metis.Print.pp val inferenceToString : inference -> string -val pp : proof Metis.Parser.pp +val pp : proof Metis.Print.pp val toString : proof -> string @@ -8846,7 +12416,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; @@ -8855,7 +12425,7 @@ (* ========================================================================= *) (* PROOFS IN FIRST ORDER LOGIC *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Proof :> Proof = @@ -8889,120 +12459,117 @@ | inferenceType (Equality _) = Thm.Equality; local - fun ppAssume pp atm = (Parser.addBreak pp (1,0); Atom.pp pp atm); - - fun ppSubst ppThm pp (sub,thm) = - (Parser.addBreak pp (1,0); - Parser.beginBlock pp Parser.Inconsistent 1; - Parser.addString pp "{"; - Parser.ppBinop " =" Parser.ppString Subst.pp pp ("sub",sub); - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Parser.ppBinop " =" Parser.ppString ppThm pp ("thm",thm); - Parser.addString pp "}"; - Parser.endBlock pp); - - fun ppResolve ppThm pp (res,pos,neg) = - (Parser.addBreak pp (1,0); - Parser.beginBlock pp Parser.Inconsistent 1; - Parser.addString pp "{"; - Parser.ppBinop " =" Parser.ppString Atom.pp pp ("res",res); - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Parser.ppBinop " =" Parser.ppString ppThm pp ("pos",pos); - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Parser.ppBinop " =" Parser.ppString ppThm pp ("neg",neg); - Parser.addString pp "}"; - Parser.endBlock pp); - - fun ppRefl pp tm = (Parser.addBreak pp (1,0); Term.pp pp tm); - - fun ppEquality pp (lit,path,res) = - (Parser.addBreak pp (1,0); - Parser.beginBlock pp Parser.Inconsistent 1; - Parser.addString pp "{"; - Parser.ppBinop " =" Parser.ppString Literal.pp pp ("lit",lit); - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Parser.ppBinop " =" Parser.ppString Term.ppPath pp ("path",path); - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Parser.ppBinop " =" Parser.ppString Term.pp pp ("res",res); - Parser.addString pp "}"; - Parser.endBlock pp); - - fun ppInf ppAxiom ppThm pp inf = + fun ppAssume atm = Print.sequence (Print.addBreak 1) (Atom.pp atm); + + fun ppSubst ppThm (sub,thm) = + Print.sequence (Print.addBreak 1) + (Print.blockProgram Print.Inconsistent 1 + [Print.addString "{", + Print.ppOp2 " =" Print.ppString Subst.pp ("sub",sub), + Print.addString ",", + Print.addBreak 1, + Print.ppOp2 " =" Print.ppString ppThm ("thm",thm), + Print.addString "}"]); + + fun ppResolve ppThm (res,pos,neg) = + Print.sequence (Print.addBreak 1) + (Print.blockProgram Print.Inconsistent 1 + [Print.addString "{", + Print.ppOp2 " =" Print.ppString Atom.pp ("res",res), + Print.addString ",", + Print.addBreak 1, + Print.ppOp2 " =" Print.ppString ppThm ("pos",pos), + Print.addString ",", + Print.addBreak 1, + Print.ppOp2 " =" Print.ppString ppThm ("neg",neg), + Print.addString "}"]); + + fun ppRefl tm = Print.sequence (Print.addBreak 1) (Term.pp tm); + + fun ppEquality (lit,path,res) = + Print.sequence (Print.addBreak 1) + (Print.blockProgram Print.Inconsistent 1 + [Print.addString "{", + Print.ppOp2 " =" Print.ppString Literal.pp ("lit",lit), + Print.addString ",", + Print.addBreak 1, + Print.ppOp2 " =" Print.ppString Term.ppPath ("path",path), + Print.addString ",", + Print.addBreak 1, + Print.ppOp2 " =" Print.ppString Term.pp ("res",res), + Print.addString "}"]); + + fun ppInf ppAxiom ppThm inf = let val infString = Thm.inferenceTypeToString (inferenceType inf) in - Parser.beginBlock pp Parser.Inconsistent (size infString + 1); - Parser.ppString pp infString; - case inf of - Axiom cl => ppAxiom pp cl - | Assume x => ppAssume pp x - | Subst x => ppSubst ppThm pp x - | Resolve x => ppResolve ppThm pp x - | Refl x => ppRefl pp x - | Equality x => ppEquality pp x; - Parser.endBlock pp - end; - - fun ppAxiom pp cl = - (Parser.addBreak pp (1,0); - Parser.ppMap - LiteralSet.toList - (Parser.ppBracket "{" "}" (Parser.ppSequence "," Literal.pp)) pp cl); + Print.block Print.Inconsistent 2 + (Print.sequence + (Print.addString infString) + (case inf of + Axiom cl => ppAxiom cl + | Assume x => ppAssume x + | Subst x => ppSubst ppThm x + | Resolve x => ppResolve ppThm x + | Refl x => ppRefl x + | Equality x => ppEquality x)) + end; + + fun ppAxiom cl = + Print.sequence + (Print.addBreak 1) + (Print.ppMap + LiteralSet.toList + (Print.ppBracket "{" "}" (Print.ppOpList "," Literal.pp)) cl); in val ppInference = ppInf ppAxiom Thm.pp; - fun pp p prf = + fun pp prf = let fun thmString n = "(" ^ Int.toString n ^ ")" - + val prf = enumerate prf - fun ppThm p th = + fun ppThm th = + Print.addString let val cl = Thm.clause th fun pred (_,(th',_)) = LiteralSet.equal (Thm.clause th') cl in case List.find pred prf of - NONE => Parser.addString p "(?)" - | SOME (n,_) => Parser.addString p (thmString n) + NONE => "(?)" + | SOME (n,_) => thmString n end fun ppStep (n,(th,inf)) = let val s = thmString n in - Parser.beginBlock p Parser.Consistent (1 + size s); - Parser.addString p (s ^ " "); - Thm.pp p th; - Parser.addBreak p (2,0); - Parser.ppBracket "[" "]" (ppInf (K (K ())) ppThm) p inf; - Parser.endBlock p; - Parser.addNewline p - end - in - Parser.beginBlock p Parser.Consistent 0; - Parser.addString p "START OF PROOF"; - Parser.addNewline p; - app ppStep prf; - Parser.addString p "END OF PROOF"; - Parser.addNewline p; - Parser.endBlock p - end -(*DEBUG + Print.sequence + (Print.blockProgram Print.Consistent (1 + size s) + [Print.addString (s ^ " "), + Thm.pp th, + Print.addBreak 2, + Print.ppBracket "[" "]" (ppInf (K Print.skip) ppThm) inf]) + Print.addNewline + end + in + Print.blockProgram Print.Consistent 0 + [Print.addString "START OF PROOF", + Print.addNewline, + Print.program (map ppStep prf), + Print.addString "END OF PROOF"] + end +(*MetisDebug handle Error err => raise Bug ("Proof.pp: shouldn't fail:\n" ^ err); *) end; -val inferenceToString = Parser.toString ppInference; - -val toString = Parser.toString pp; +val inferenceToString = Print.toString ppInference; + +val toString = Print.toString pp; (* ------------------------------------------------------------------------- *) (* Reconstructing single inferences. *) @@ -9027,9 +12594,9 @@ let fun recon [] = let -(*TRACE3 - val () = Parser.ppTrace LiteralSet.pp "reconstructSubst: cl" cl - val () = Parser.ppTrace LiteralSet.pp "reconstructSubst: cl'" cl' +(*MetisTrace3 + val () = Print.trace LiteralSet.pp "reconstructSubst: cl" cl + val () = Print.trace LiteralSet.pp "reconstructSubst: cl'" cl' *) in raise Bug "can't reconstruct Subst rule" @@ -9049,7 +12616,7 @@ in Subst.normalize (recon [(LiteralSet.toList cl, Subst.empty)]) end -(*DEBUG +(*MetisDebug handle Error err => raise Bug ("Proof.recontructSubst: shouldn't fail:\n" ^ err); *) @@ -9069,32 +12636,37 @@ if not (LiteralSet.null lits) then LiteralSet.pick lits else raise Bug "can't reconstruct Resolve rule" end) -(*DEBUG +(*MetisDebug handle Error err => raise Bug ("Proof.recontructResolvant: shouldn't fail:\n" ^ err); *) fun reconstructEquality cl = let -(*TRACE3 - val () = Parser.ppTrace LiteralSet.pp "Proof.reconstructEquality: cl" cl +(*MetisTrace3 + val () = Print.trace LiteralSet.pp "Proof.reconstructEquality: cl" cl *) fun sync s t path (f,a) (f',a') = - if f <> f' orelse length a <> length a' then NONE + if not (Name.equal f f' andalso length a = length a') then NONE else - case List.filter (op<> o snd) (enumerate (zip a a')) of - [(i,(tm,tm'))] => - let - val path = i :: path - in - if tm = s andalso tm' = t then SOME (rev path) - else - case (tm,tm') of - (Term.Fn f_a, Term.Fn f_a') => sync s t path f_a f_a' - | _ => NONE - end - | _ => NONE + let + val itms = enumerate (zip a a') + in + case List.filter (not o uncurry Term.equal o snd) itms of + [(i,(tm,tm'))] => + let + val path = i :: path + in + if Term.equal tm s andalso Term.equal tm' t then + SOME (rev path) + else + case (tm,tm') of + (Term.Fn f_a, Term.Fn f_a') => sync s t path f_a f_a' + | _ => NONE + end + | _ => NONE + end fun recon (neq,(pol,atm),(pol',atm')) = if pol = pol' then NONE @@ -9103,9 +12675,9 @@ val (s,t) = Literal.destNeq neq val path = - if s <> t then sync s t [] atm atm' - else if atm <> atm' then NONE - else Atom.find (equal s) atm + if not (Term.equal s t) then sync s t [] atm atm' + else if not (Atom.equal atm atm') then NONE + else Atom.find (Term.equal s) atm in case path of SOME path => SOME ((pol',atm),path,t) @@ -9119,10 +12691,10 @@ | ([l1],[l2]) => [(l1,l1,l2),(l1,l2,l1)] | _ => raise Bug "reconstructEquality: malformed" -(*TRACE3 +(*MetisTrace3 val ppCands = - Parser.ppList (Parser.ppTriple Literal.pp Literal.pp Literal.pp) - val () = Parser.ppTrace ppCands + Print.ppList (Print.ppTriple Literal.pp Literal.pp Literal.pp) + val () = Print.trace ppCands "Proof.reconstructEquality: candidates" candidates *) in @@ -9130,7 +12702,7 @@ SOME info => info | NONE => raise Bug "can't reconstruct Equality rule" end -(*DEBUG +(*MetisDebug handle Error err => raise Bug ("Proof.recontructEquality: shouldn't fail:\n" ^ err); *) @@ -9159,25 +12731,25 @@ in fun thmToInference th = let -(*TRACE3 - val () = Parser.ppTrace Thm.pp "Proof.thmToInference: th" th +(*MetisTrace3 + val () = Print.trace Thm.pp "Proof.thmToInference: th" th *) val cl = Thm.clause th val thmInf = Thm.inference th -(*TRACE3 - val ppThmInf = Parser.ppPair Thm.ppInferenceType (Parser.ppList Thm.pp) - val () = Parser.ppTrace ppThmInf "Proof.thmToInference: thmInf" thmInf +(*MetisTrace3 + val ppThmInf = Print.ppPair Thm.ppInferenceType (Print.ppList Thm.pp) + val () = Print.trace ppThmInf "Proof.thmToInference: thmInf" thmInf *) val inf = reconstruct cl thmInf -(*TRACE3 - val () = Parser.ppTrace ppInference "Proof.thmToInference: inf" inf -*) -(*DEBUG +(*MetisTrace3 + val () = Print.trace ppInference "Proof.thmToInference: inf" inf +*) +(*MetisDebug val () = let val th' = inferenceToThm inf @@ -9195,7 +12767,7 @@ in inf end -(*DEBUG +(*MetisDebug handle Error err => raise Bug ("Proof.thmToInference: shouldn't fail:\n" ^ err); *) @@ -9206,40 +12778,101 @@ (* ------------------------------------------------------------------------- *) local - fun thmCompare (th1,th2) = - LiteralSet.compare (Thm.clause th1, Thm.clause th2); - - fun buildProof (th,(m,l)) = - if Map.inDomain th m then (m,l) - else - let - val (_,deps) = Thm.inference th - val (m,l) = foldl buildProof (m,l) deps - in - if Map.inDomain th m then (m,l) - else - let - val l = (th, thmToInference th) :: l - in - (Map.insert m (th,l), l) - end - end; + val emptyThms : Thm.thm LiteralSetMap.map = LiteralSetMap.new (); + + fun addThms (th,ths) = + let + val cl = Thm.clause th + in + if LiteralSetMap.inDomain cl ths then ths + else + let + val (_,pars) = Thm.inference th + val ths = List.foldl addThms ths pars + in + if LiteralSetMap.inDomain cl ths then ths + else LiteralSetMap.insert ths (cl,th) + end + end; + + fun mkThms th = addThms (th,emptyThms); + + fun addProof (th,(ths,acc)) = + let + val cl = Thm.clause th + in + case LiteralSetMap.peek ths cl of + NONE => (ths,acc) + | SOME th => + let + val (_,pars) = Thm.inference th + val (ths,acc) = List.foldl addProof (ths,acc) pars + val ths = LiteralSetMap.delete ths cl + val acc = (th, thmToInference th) :: acc + in + (ths,acc) + end + end; + + fun mkProof ths th = + let + val (ths,acc) = addProof (th,(ths,[])) +(*MetisTrace4 + val () = Print.trace Print.ppInt "Proof.proof: unnecessary clauses" (LiteralSetMap.size ths) +*) + in + rev acc + end; in fun proof th = let -(*TRACE3 - val () = Parser.ppTrace Thm.pp "Proof.proof: th" th -*) - val (m,_) = buildProof (th, (Map.new thmCompare, [])) -(*TRACE3 - val () = Parser.ppTrace Parser.ppInt "Proof.proof: size" (Map.size m) -*) - in - case Map.peek m th of - SOME l => rev l - | NONE => raise Bug "Proof.proof" - end; -end; +(*MetisTrace3 + val () = Print.trace Thm.pp "Proof.proof: th" th +*) + val ths = mkThms th + val infs = mkProof ths th +(*MetisTrace3 + val () = Print.trace Print.ppInt "Proof.proof: size" (length infs) +*) + in + infs + end; +end; + +(* ------------------------------------------------------------------------- *) +(* Free variables. *) +(* ------------------------------------------------------------------------- *) + +fun freeIn v = + let + fun free th_inf = + case th_inf of + (_, Axiom lits) => LiteralSet.freeIn v lits + | (_, Assume atm) => Atom.freeIn v atm + | (th, Subst _) => Thm.freeIn v th + | (_, Resolve _) => false + | (_, Refl tm) => Term.freeIn v tm + | (_, Equality (lit,_,tm)) => + Literal.freeIn v lit orelse Term.freeIn v tm + in + List.exists free + end; + +val freeVars = + let + fun inc (th_inf,set) = + NameSet.union set + (case th_inf of + (_, Axiom lits) => LiteralSet.freeVars lits + | (_, Assume atm) => Atom.freeVars atm + | (th, Subst _) => Thm.freeVars th + | (_, Resolve _) => NameSet.empty + | (_, Refl tm) => Term.freeVars tm + | (_, Equality (lit,_,tm)) => + NameSet.union (Literal.freeVars lit) (Term.freeVars tm)) + in + List.foldl inc NameSet.empty + end; end end; @@ -9248,7 +12881,7 @@ (* ========================================================================= *) (* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Rule = @@ -9261,7 +12894,7 @@ type equation = (Metis.Term.term * Metis.Term.term) * Metis.Thm.thm -val ppEquation : equation Metis.Parser.pp +val ppEquation : equation Metis.Print.pp val equationToString : equation -> string @@ -9391,6 +13024,8 @@ (* x = x *) (* ------------------------------------------------------------------------- *) +val reflexivityRule : Metis.Term.term -> Metis.Thm.thm + val reflexivity : Metis.Thm.thm (* ------------------------------------------------------------------------- *) @@ -9399,6 +13034,8 @@ (* ~(x = y) \/ y = x *) (* ------------------------------------------------------------------------- *) +val symmetryRule : Metis.Term.term -> Metis.Term.term -> Metis.Thm.thm + val symmetry : Metis.Thm.thm (* ------------------------------------------------------------------------- *) @@ -9521,7 +13158,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; @@ -9530,7 +13167,7 @@ (* ========================================================================= *) (* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Rule :> Rule = @@ -9539,12 +13176,33 @@ open Useful; (* ------------------------------------------------------------------------- *) +(* Variable names. *) +(* ------------------------------------------------------------------------- *) + +val xVarName = Name.fromString "x"; +val xVar = Term.Var xVarName; + +val yVarName = Name.fromString "y"; +val yVar = Term.Var yVarName; + +val zVarName = Name.fromString "z"; +val zVar = Term.Var zVarName; + +fun xIVarName i = Name.fromString ("x" ^ Int.toString i); +fun xIVar i = Term.Var (xIVarName i); + +fun yIVarName i = Name.fromString ("y" ^ Int.toString i); +fun yIVar i = Term.Var (yIVarName i); + +(* ------------------------------------------------------------------------- *) (* *) (* --------- reflexivity *) (* x = x *) (* ------------------------------------------------------------------------- *) -val reflexivity = Thm.refl (Term.Var "x"); +fun reflexivityRule x = Thm.refl x; + +val reflexivity = reflexivityRule xVar; (* ------------------------------------------------------------------------- *) (* *) @@ -9552,17 +13210,17 @@ (* ~(x = y) \/ y = x *) (* ------------------------------------------------------------------------- *) -val symmetry = - let - val x = Term.Var "x" - and y = Term.Var "y" - val reflTh = reflexivity +fun symmetryRule x y = + let + val reflTh = reflexivityRule x val reflLit = Thm.destUnit reflTh val eqTh = Thm.equality reflLit [0] y in Thm.resolve reflLit reflTh eqTh end; +val symmetry = symmetryRule xVar yVar; + (* ------------------------------------------------------------------------- *) (* *) (* --------------------------------- transitivity *) @@ -9571,12 +13229,9 @@ val transitivity = let - val x = Term.Var "x" - and y = Term.Var "y" - and z = Term.Var "z" - val eqTh = Thm.equality (Literal.mkEq (y,z)) [0] x - in - Thm.resolve (Literal.mkEq (y,x)) symmetry eqTh + val eqTh = Thm.equality (Literal.mkEq (yVar,zVar)) [0] xVar + in + Thm.resolve (Literal.mkEq (yVar,xVar)) symmetry eqTh end; (* ------------------------------------------------------------------------- *) @@ -9589,10 +13244,10 @@ let val (x,y) = Literal.destEq lit in - if x = y then th + if Term.equal x y then th else let - val sub = Subst.fromList [("x",x),("y",y)] + val sub = Subst.fromList [(xVarName,x),(yVarName,y)] val symTh = Thm.subst sub symmetry in Thm.resolve lit th symTh @@ -9606,9 +13261,9 @@ type equation = (Term.term * Term.term) * Thm.thm; -fun ppEquation pp (eqn as (_,th)) = Thm.pp pp th; - -fun equationToString x = Parser.toString ppEquation x; +fun ppEquation (_,th) = Thm.pp th; + +val equationToString = Print.toString ppEquation; fun equationLiteral (t_u,th) = let @@ -9620,7 +13275,7 @@ fun reflEqn t = ((t,t), Thm.refl t); fun symEqn (eqn as ((t,u), th)) = - if t = u then eqn + if Term.equal t u then eqn else ((u,t), case equationLiteral eqn of @@ -9628,9 +13283,9 @@ | NONE => th); fun transEqn (eqn1 as ((x,y), th1)) (eqn2 as ((_,z), th2)) = - if x = y then eqn2 - else if y = z then eqn1 - else if x = z then reflEqn x + if Term.equal x y then eqn2 + else if Term.equal y z then eqn1 + else if Term.equal x z then reflEqn x else ((x,z), case equationLiteral eqn1 of @@ -9640,7 +13295,7 @@ NONE => th2 | SOME y_z => let - val sub = Subst.fromList [("x",x),("y",y),("z",z)] + val sub = Subst.fromList [(xVarName,x),(yVarName,y),(zVarName,z)] val th = Thm.subst sub transitivity val th = Thm.resolve x_y th1 th val th = Thm.resolve y_z th2 th @@ -9648,7 +13303,7 @@ th end); -(*DEBUG +(*MetisDebug val transEqn = fn eqn1 => fn eqn2 => transEqn eqn1 eqn2 handle Error err => @@ -9679,7 +13334,7 @@ handle Error err => (print (s ^ ": " ^ Term.toString tm ^ " --> Error: " ^ err ^ "\n"); raise Error (s ^ ": " ^ err)); - + fun thenConvTrans tm (tm',th1) (tm'',th2) = let val eqn1 = ((tm,tm'),th1) @@ -9719,7 +13374,7 @@ | everyConv (conv :: convs) tm = thenConv conv (everyConv convs) tm; fun rewrConv (eqn as ((x,y), eqTh)) path tm = - if x = y then allConv tm + if Term.equal x y then allConv tm else if null path then (y,eqTh) else let @@ -9736,7 +13391,7 @@ (tm',th) end; -(*DEBUG +(*MetisDebug val rewrConv = fn eqn as ((x,y),eqTh) => fn path => fn tm => rewrConv eqn path tm handle Error err => @@ -9780,7 +13435,7 @@ f end; -(*DEBUG +(*MetisDebug val repeatTopDownConv = fn conv => fn tm => repeatTopDownConv conv tm handle Error err => raise Error ("repeatTopDownConv: " ^ err); @@ -9803,9 +13458,9 @@ val res1 as (lit',th1) = literule1 lit val res2 as (lit'',th2) = literule2 lit' in - if lit = lit' then res2 - else if lit' = lit'' then res1 - else if lit = lit'' then allLiterule lit + if Literal.equal lit lit' then res2 + else if Literal.equal lit' lit'' then res1 + else if Literal.equal lit lit'' then allLiterule lit else (lit'', if not (Thm.member lit' th1) then th1 @@ -9839,7 +13494,7 @@ thenLiterule literule (everyLiterule literules) lit; fun rewrLiterule (eqn as ((x,y),eqTh)) path lit = - if x = y then allLiterule lit + if Term.equal x y then allLiterule lit else let val th = Thm.equality lit path y @@ -9852,7 +13507,7 @@ (lit',th) end; -(*DEBUG +(*MetisDebug val rewrLiterule = fn eqn => fn path => fn lit => rewrLiterule eqn path lit handle Error err => @@ -9914,12 +13569,12 @@ let val (lit',litTh) = literule lit in - if lit = lit' then th + if Literal.equal lit lit' then th else if not (Thm.negateMember lit litTh) then litTh else Thm.resolve lit th litTh end; -(*DEBUG +(*MetisDebug val literalRule = fn literule => fn lit => fn th => literalRule literule lit th handle Error err => @@ -9942,7 +13597,7 @@ fun allLiteralsRule literule th = literalsRule literule (Thm.clause th) th; fun convRule conv = allLiteralsRule (allArgumentsLiterule conv); - + (* ------------------------------------------------------------------------- *) (* *) (* ---------------------------------------------- functionCongruence (f,n) *) @@ -9952,8 +13607,8 @@ fun functionCongruence (f,n) = let - val xs = List.tabulate (n, fn i => Term.Var ("x" ^ Int.toString i)) - and ys = List.tabulate (n, fn i => Term.Var ("y" ^ Int.toString i)) + val xs = List.tabulate (n,xIVar) + and ys = List.tabulate (n,yIVar) fun cong ((i,yi),(th,lit)) = let @@ -9979,8 +13634,8 @@ fun relationCongruence (R,n) = let - val xs = List.tabulate (n, fn i => Term.Var ("x" ^ Int.toString i)) - and ys = List.tabulate (n, fn i => Term.Var ("y" ^ Int.toString i)) + val xs = List.tabulate (n,xIVar) + and ys = List.tabulate (n,yIVar) fun cong ((i,yi),(th,lit)) = let @@ -10007,10 +13662,10 @@ let val (x,y) = Literal.destNeq lit in - if x = y then th + if Term.equal x y then th else let - val sub = Subst.fromList [("x",y),("y",x)] + val sub = Subst.fromList [(xVarName,y),(yVarName,x)] val symTh = Thm.subst sub symmetry in Thm.resolve lit th symTh @@ -10075,10 +13730,12 @@ fun expand lit = let val (x,y) = Literal.destNeq lit - in - if (Term.isTypedVar x orelse Term.isTypedVar y) andalso x <> y then - Subst.unify Subst.empty x y - else raise Error "expand" + val _ = Term.isTypedVar x orelse Term.isTypedVar y orelse + raise Error "Rule.expandAbbrevs: no vars" + val _ = not (Term.equal x y) orelse + raise Error "Rule.expandAbbrevs: equal vars" + in + Subst.unify Subst.empty x y end; in fun expandAbbrevs th = @@ -10133,8 +13790,8 @@ FactorEdge of Atom.atom * Atom.atom | ReflEdge of Term.term * Term.term; - fun ppEdge p (FactorEdge atm_atm') = Parser.ppPair Atom.pp Atom.pp p atm_atm' - | ppEdge p (ReflEdge tm_tm') = Parser.ppPair Term.pp Term.pp p tm_tm'; + fun ppEdge (FactorEdge atm_atm') = Print.ppPair Atom.pp Atom.pp atm_atm' + | ppEdge (ReflEdge tm_tm') = Print.ppPair Term.pp Term.pp tm_tm'; datatype joinStatus = Joined @@ -10209,7 +13866,7 @@ end | init_edges acc apart ((sub,edge) :: sub_edges) = let -(*DEBUG +(*MetisDebug val () = if not (Subst.null sub) then () else raise Bug "Rule.factor.init_edges: empty subst" *) @@ -10262,21 +13919,21 @@ in fun factor' cl = let -(*TRACE6 - val () = Parser.ppTrace LiteralSet.pp "Rule.factor': cl" cl +(*MetisTrace6 + val () = Print.trace LiteralSet.pp "Rule.factor': cl" cl *) val edges = mk_edges [] [] (LiteralSet.toList cl) -(*TRACE6 - val ppEdgesSize = Parser.ppMap length Parser.ppInt - val ppEdgel = Parser.ppList ppEdge - val ppEdges = Parser.ppList (Parser.ppTriple ppEdgel Subst.pp ppEdgel) - val () = Parser.ppTrace ppEdgesSize "Rule.factor': |edges|" edges - val () = Parser.ppTrace ppEdges "Rule.factor': edges" edges +(*MetisTrace6 + val ppEdgesSize = Print.ppMap length Print.ppInt + val ppEdgel = Print.ppList ppEdge + val ppEdges = Print.ppList (Print.ppTriple ppEdgel Subst.pp ppEdgel) + val () = Print.trace ppEdgesSize "Rule.factor': |edges|" edges + val () = Print.trace ppEdges "Rule.factor': edges" edges *) val result = fact [] edges -(*TRACE6 - val ppResult = Parser.ppList Subst.pp - val () = Parser.ppTrace ppResult "Rule.factor': result" result +(*MetisTrace6 + val ppResult = Print.ppList Subst.pp + val () = Print.trace ppResult "Rule.factor': result" result *) in result @@ -10297,7 +13954,7 @@ (* ========================================================================= *) (* NORMALIZING FORMULAS *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2009 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Normalize = @@ -10310,10 +13967,44 @@ val nnf : Metis.Formula.formula -> Metis.Formula.formula (* ------------------------------------------------------------------------- *) +(* Conjunctive normal form derivations. *) +(* ------------------------------------------------------------------------- *) + +type thm + +datatype inference = + Axiom of Metis.Formula.formula + | Definition of string * Metis.Formula.formula + | Simplify of thm * thm list + | Conjunct of thm + | Specialize of thm + | Skolemize of thm + | Clausify of thm + +val mkAxiom : Metis.Formula.formula -> thm + +val destThm : thm -> Metis.Formula.formula * inference + +val proveThms : + thm list -> (Metis.Formula.formula * inference * Metis.Formula.formula list) list + +val toStringInference : inference -> string + +val ppInference : inference Metis.Print.pp + +(* ------------------------------------------------------------------------- *) (* Conjunctive normal form. *) (* ------------------------------------------------------------------------- *) -val cnf : Metis.Formula.formula -> Metis.Formula.formula list +type cnf + +val initialCnf : cnf + +val addCnf : thm -> cnf -> (Metis.Thm.clause * thm) list * cnf + +val proveCnf : thm list -> (Metis.Thm.clause * thm) list + +val cnf : Metis.Formula.formula -> Metis.Thm.clause list end @@ -10321,7 +14012,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; @@ -10330,7 +14021,7 @@ (* ========================================================================= *) (* NORMALIZING FORMULAS *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Normalize :> Normalize = @@ -10339,38 +14030,102 @@ open Useful; (* ------------------------------------------------------------------------- *) +(* Constants. *) +(* ------------------------------------------------------------------------- *) + +val prefix = "FOFtoCNF"; + +val skolemPrefix = "skolem" ^ prefix; + +val definitionPrefix = "definition" ^ prefix; + +(* ------------------------------------------------------------------------- *) +(* Storing huge real numbers as their log. *) +(* ------------------------------------------------------------------------- *) + +datatype logReal = LogReal of real; + +fun compareLogReal (LogReal logX, LogReal logY) = + Real.compare (logX,logY); + +val zeroLogReal = LogReal ~1.0; + +val oneLogReal = LogReal 0.0; + +local + fun isZero logX = logX < 0.0; + + (* Assume logX >= logY >= 0.0 *) + fun add logX logY = logX + Math.ln (1.0 + Math.exp (logY - logX)); +in + fun isZeroLogReal (LogReal logX) = isZero logX; + + fun multiplyLogReal (LogReal logX) (LogReal logY) = + if isZero logX orelse isZero logY then zeroLogReal + else LogReal (logX + logY); + + fun addLogReal (lx as LogReal logX) (ly as LogReal logY) = + if isZero logX then ly + else if isZero logY then lx + else if logX < logY then LogReal (add logY logX) + else LogReal (add logX logY); + + fun withinRelativeLogReal logDelta (LogReal logX) (LogReal logY) = + isZero logX orelse + (not (isZero logY) andalso logX < logY + logDelta); +end; + +fun toStringLogReal (LogReal logX) = Real.toString logX; + +(* ------------------------------------------------------------------------- *) (* Counting the clauses that would be generated by conjunctive normal form. *) (* ------------------------------------------------------------------------- *) -datatype count = Count of {positive : real, negative : real}; - -fun positive (Count {positive = p, ...}) = p; - -fun negative (Count {negative = n, ...}) = n; +val countLogDelta = 0.01; + +datatype count = Count of {positive : logReal, negative : logReal}; + +fun countCompare (count1,count2) = + let + val Count {positive = p1, negative = _} = count1 + and Count {positive = p2, negative = _} = count2 + in + compareLogReal (p1,p2) + end; fun countNegate (Count {positive = p, negative = n}) = Count {positive = n, negative = p}; +fun countLeqish count1 count2 = + let + val Count {positive = p1, negative = _} = count1 + and Count {positive = p2, negative = _} = count2 + in + withinRelativeLogReal countLogDelta p1 p2 + end; + +(*MetisDebug fun countEqualish count1 count2 = - let - val Count {positive = p1, negative = n1} = count1 - and Count {positive = p2, negative = n2} = count2 - in - Real.abs (p1 - p2) < 0.5 andalso Real.abs (n1 - n2) < 0.5 - end; - -val countTrue = Count {positive = 0.0, negative = 1.0}; - -val countFalse = Count {positive = 1.0, negative = 0.0}; - -val countLiteral = Count {positive = 1.0, negative = 1.0}; + countLeqish count1 count2 andalso + countLeqish count2 count1; + +fun countEquivish count1 count2 = + countEqualish count1 count2 andalso + countEqualish (countNegate count1) (countNegate count2); +*) + +val countTrue = Count {positive = zeroLogReal, negative = oneLogReal}; + +val countFalse = Count {positive = oneLogReal, negative = zeroLogReal}; + +val countLiteral = Count {positive = oneLogReal, negative = oneLogReal}; fun countAnd2 (count1,count2) = let val Count {positive = p1, negative = n1} = count1 and Count {positive = p2, negative = n2} = count2 - val p = p1 + p2 - and n = n1 * n2 + val p = addLogReal p1 p2 + and n = multiplyLogReal n1 n2 in Count {positive = p, negative = n} end; @@ -10379,25 +14134,36 @@ let val Count {positive = p1, negative = n1} = count1 and Count {positive = p2, negative = n2} = count2 - val p = p1 * p2 - and n = n1 + n2 + val p = multiplyLogReal p1 p2 + and n = addLogReal n1 n2 in Count {positive = p, negative = n} end; -(*** Is this associative? ***) +(* Whether countXor2 is associative or not is an open question. *) + fun countXor2 (count1,count2) = let val Count {positive = p1, negative = n1} = count1 and Count {positive = p2, negative = n2} = count2 - val p = p1 * p2 + n1 * n2 - and n = p1 * n2 + n1 * p2 + val p = addLogReal (multiplyLogReal p1 p2) (multiplyLogReal n1 n2) + and n = addLogReal (multiplyLogReal p1 n2) (multiplyLogReal n1 p2) in Count {positive = p, negative = n} end; fun countDefinition body_count = countXor2 (countLiteral,body_count); +val countToString = + let + val rToS = toStringLogReal + in + fn Count {positive = p, negative = n} => + "(+" ^ rToS p ^ ",-" ^ rToS n ^ ")" + end; + +val ppCount = Print.ppMap countToString Print.ppString; + (* ------------------------------------------------------------------------- *) (* A type of normalized formula. *) (* ------------------------------------------------------------------------- *) @@ -10413,41 +14179,43 @@ | Forall of NameSet.set * count * NameSet.set * formula; fun compare f1_f2 = - case f1_f2 of - (True,True) => EQUAL - | (True,_) => LESS - | (_,True) => GREATER - | (False,False) => EQUAL - | (False,_) => LESS - | (_,False) => GREATER - | (Literal (_,l1), Literal (_,l2)) => Literal.compare (l1,l2) - | (Literal _, _) => LESS - | (_, Literal _) => GREATER - | (And (_,_,s1), And (_,_,s2)) => Set.compare (s1,s2) - | (And _, _) => LESS - | (_, And _) => GREATER - | (Or (_,_,s1), Or (_,_,s2)) => Set.compare (s1,s2) - | (Or _, _) => LESS - | (_, Or _) => GREATER - | (Xor (_,_,p1,s1), Xor (_,_,p2,s2)) => - (case boolCompare (p1,p2) of - LESS => LESS - | EQUAL => Set.compare (s1,s2) - | GREATER => GREATER) - | (Xor _, _) => LESS - | (_, Xor _) => GREATER - | (Exists (_,_,n1,f1), Exists (_,_,n2,f2)) => - (case NameSet.compare (n1,n2) of - LESS => LESS - | EQUAL => compare (f1,f2) - | GREATER => GREATER) - | (Exists _, _) => LESS - | (_, Exists _) => GREATER - | (Forall (_,_,n1,f1), Forall (_,_,n2,f2)) => - (case NameSet.compare (n1,n2) of - LESS => LESS - | EQUAL => compare (f1,f2) - | GREATER => GREATER); + if Portable.pointerEqual f1_f2 then EQUAL + else + case f1_f2 of + (True,True) => EQUAL + | (True,_) => LESS + | (_,True) => GREATER + | (False,False) => EQUAL + | (False,_) => LESS + | (_,False) => GREATER + | (Literal (_,l1), Literal (_,l2)) => Literal.compare (l1,l2) + | (Literal _, _) => LESS + | (_, Literal _) => GREATER + | (And (_,_,s1), And (_,_,s2)) => Set.compare (s1,s2) + | (And _, _) => LESS + | (_, And _) => GREATER + | (Or (_,_,s1), Or (_,_,s2)) => Set.compare (s1,s2) + | (Or _, _) => LESS + | (_, Or _) => GREATER + | (Xor (_,_,p1,s1), Xor (_,_,p2,s2)) => + (case boolCompare (p1,p2) of + LESS => LESS + | EQUAL => Set.compare (s1,s2) + | GREATER => GREATER) + | (Xor _, _) => LESS + | (_, Xor _) => GREATER + | (Exists (_,_,n1,f1), Exists (_,_,n2,f2)) => + (case NameSet.compare (n1,n2) of + LESS => LESS + | EQUAL => compare (f1,f2) + | GREATER => GREATER) + | (Exists _, _) => LESS + | (_, Exists _) => GREATER + | (Forall (_,_,n1,f1), Forall (_,_,n2,f2)) => + (case NameSet.compare (n1,n2) of + LESS => LESS + | EQUAL => compare (f1,f2) + | GREATER => GREATER); val empty = Set.empty compare; @@ -10491,7 +14259,7 @@ | polarity (Exists _) = true | polarity (Forall _) = false; -(*DEBUG +(*MetisDebug val polarity = fn f => let val res1 = compare (f, negate f) = LESS @@ -10585,7 +14353,7 @@ end end; -val AndList = foldl And2 True; +val AndList = List.foldl And2 True; val AndSet = Set.foldl And2 True; @@ -10621,7 +14389,7 @@ end end; -val OrList = foldl Or2 False; +val OrList = List.foldl Or2 False; val OrSet = Set.foldl Or2 False; @@ -10631,12 +14399,13 @@ and s2 = case f2 of And (_,_,s) => s | _ => singleton f2 fun g x1 (x2,acc) = And2 (Or2 (x1,x2), acc) + fun f (x1,acc) = Set.foldl (g x1) acc s2 in Set.foldl f True s1 end; -val pushOrList = foldl pushOr2 False; +val pushOrList = List.foldl pushOr2 False; local fun normalize fm = @@ -10674,7 +14443,7 @@ end; end; -val XorList = foldl Xor2 False; +val XorList = List.foldl Xor2 False; val XorSet = Set.foldl Xor2 False; @@ -10736,7 +14505,7 @@ exists init_fm end; -fun ExistsList (vs,f) = foldl Exists1 f vs; +fun ExistsList (vs,f) = List.foldl Exists1 f vs; fun ExistsSet (n,f) = NameSet.foldl Exists1 f n; @@ -10774,10 +14543,12 @@ forall init_fm end; -fun ForallList (vs,f) = foldl Forall1 f vs; +fun ForallList (vs,f) = List.foldl Forall1 f vs; fun ForallSet (n,f) = NameSet.foldl Forall1 f n; +fun generalize f = ForallSet (freeVars f, f); + local fun subst_fv fvSub = let @@ -10910,9 +14681,20 @@ val toFormula = form; end; -val pp = Parser.ppMap toFormula Formula.pp; - -val toString = Parser.toString pp; +fun toLiteral (Literal (_,lit)) = lit + | toLiteral _ = raise Error "Normalize.toLiteral"; + +local + fun addLiteral (l,s) = LiteralSet.add s (toLiteral l); +in + fun toClause False = LiteralSet.empty + | toClause (Or (_,_,s)) = Set.foldl addLiteral LiteralSet.empty s + | toClause l = LiteralSet.singleton (toLiteral l); +end; + +val pp = Print.ppMap toFormula Formula.pp; + +val toString = Print.toString pp; (* ------------------------------------------------------------------------- *) (* Negation normal form. *) @@ -10921,224 +14703,30 @@ fun nnf fm = toFormula (fromFormula fm); (* ------------------------------------------------------------------------- *) -(* Simplifying with definitions. *) -(* ------------------------------------------------------------------------- *) - -datatype simplify = - Simplify of - {formula : (formula,formula) Map.map, - andSet : (formula Set.set * formula) list, - orSet : (formula Set.set * formula) list, - xorSet : (formula Set.set * formula) list}; - -val simplifyEmpty = - Simplify {formula = Map.new compare, andSet = [], orSet = [], xorSet = []}; - -local - fun simpler fm s = - Set.size s <> 1 orelse - case Set.pick s of - True => false - | False => false - | Literal _ => false - | _ => true; - - fun addSet set_defs body_def = - let - fun def_body_size (body,_) = Set.size body - - val body_size = def_body_size body_def - - val (body,_) = body_def - - fun add acc [] = List.revAppend (acc,[body_def]) - | add acc (l as (bd as (b,_)) :: bds) = - case Int.compare (def_body_size bd, body_size) of - LESS => List.revAppend (acc, body_def :: l) - | EQUAL => if Set.equal b body then List.revAppend (acc,l) - else add (bd :: acc) bds - | GREATER => add (bd :: acc) bds - in - add [] set_defs - end; - - fun add simp (body,False) = add simp (negate body, True) - | add simp (True,_) = simp - | add (Simplify {formula,andSet,orSet,xorSet}) (And (_,_,s), def) = - let - val andSet = addSet andSet (s,def) - and orSet = addSet orSet (negateSet s, negate def) - in - Simplify - {formula = formula, - andSet = andSet, orSet = orSet, xorSet = xorSet} - end - | add (Simplify {formula,andSet,orSet,xorSet}) (Or (_,_,s), def) = - let - val orSet = addSet orSet (s,def) - and andSet = addSet andSet (negateSet s, negate def) - in - Simplify - {formula = formula, - andSet = andSet, orSet = orSet, xorSet = xorSet} - end - | add simp (Xor (_,_,p,s), def) = - let - val simp = addXorSet simp (s, applyPolarity p def) - in - case def of - True => - let - fun addXorLiteral (fm as Literal _, simp) = - let - val s = Set.delete s fm - in - if not (simpler fm s) then simp - else addXorSet simp (s, applyPolarity (not p) fm) - end - | addXorLiteral (_,simp) = simp - in - Set.foldl addXorLiteral simp s - end - | _ => simp - end - | add (simp as Simplify {formula,andSet,orSet,xorSet}) (body,def) = - if Map.inDomain body formula then simp - else - let - val formula = Map.insert formula (body,def) - val formula = Map.insert formula (negate body, negate def) - in - Simplify - {formula = formula, - andSet = andSet, orSet = orSet, xorSet = xorSet} - end - - and addXorSet (simp as Simplify {formula,andSet,orSet,xorSet}) (s,def) = - if Set.size s = 1 then add simp (Set.pick s, def) - else - let - val xorSet = addSet xorSet (s,def) - in - Simplify - {formula = formula, - andSet = andSet, orSet = orSet, xorSet = xorSet} - end; -in - fun simplifyAdd simp fm = add simp (fm,True); -end; - -local - fun simplifySet set_defs set = - let - fun pred (s,_) = Set.subset s set - in - case List.find pred set_defs of - NONE => NONE - | SOME (s,f) => SOME (Set.add (Set.difference set s) f) - end; -in - fun simplify (Simplify {formula,andSet,orSet,xorSet}) = - let - fun simp fm = simp_top (simp_sub fm) - - and simp_top (changed_fm as (_, And (_,_,s))) = - (case simplifySet andSet s of - NONE => changed_fm - | SOME s => simp_top (true, AndSet s)) - | simp_top (changed_fm as (_, Or (_,_,s))) = - (case simplifySet orSet s of - NONE => changed_fm - | SOME s => simp_top (true, OrSet s)) - | simp_top (changed_fm as (_, Xor (_,_,p,s))) = - (case simplifySet xorSet s of - NONE => changed_fm - | SOME s => simp_top (true, XorPolaritySet (p,s))) - | simp_top (changed_fm as (_,fm)) = - (case Map.peek formula fm of - NONE => changed_fm - | SOME fm => simp_top (true,fm)) - - and simp_sub fm = - case fm of - And (_,_,s) => - let - val l = Set.transform simp s - val changed = List.exists fst l - val fm = if changed then AndList (map snd l) else fm - in - (changed,fm) - end - | Or (_,_,s) => - let - val l = Set.transform simp s - val changed = List.exists fst l - val fm = if changed then OrList (map snd l) else fm - in - (changed,fm) - end - | Xor (_,_,p,s) => - let - val l = Set.transform simp s - val changed = List.exists fst l - val fm = if changed then XorPolarityList (p, map snd l) else fm - in - (changed,fm) - end - | Exists (_,_,n,f) => - let - val (changed,f) = simp f - val fm = if changed then ExistsSet (n,f) else fm - in - (changed,fm) - end - | Forall (_,_,n,f) => - let - val (changed,f) = simp f - val fm = if changed then ForallSet (n,f) else fm - in - (changed,fm) - end - | _ => (false,fm); - in - fn fm => snd (simp fm) - end; -end; - -(*TRACE2 -val simplify = fn simp => fn fm => - let - val fm' = simplify simp fm - val () = if compare (fm,fm') = EQUAL then () - else (Parser.ppTrace pp "Normalize.simplify: fm" fm; - Parser.ppTrace pp "Normalize.simplify: fm'" fm') - in - fm' - end; -*) - -(* ------------------------------------------------------------------------- *) (* Basic conjunctive normal form. *) (* ------------------------------------------------------------------------- *) val newSkolemFunction = let - val counter : int NameMap.map Unsynchronized.ref = Unsynchronized.ref (NameMap.new ()) - in - fn n => CRITICAL (fn () => - let - val Unsynchronized.ref m = counter - val i = Option.getOpt (NameMap.peek m n, 0) - val () = counter := NameMap.insert m (n, i + 1) - in - "skolem_" ^ n ^ (if i = 0 then "" else "_" ^ Int.toString i) - end) + val counter : int StringMap.map Unsynchronized.ref = Unsynchronized.ref (StringMap.new ()) + in + fn n => + let + val Unsynchronized.ref m = counter + val s = Name.toString n + val i = Option.getOpt (StringMap.peek m s, 0) + val () = counter := StringMap.insert m (s, i + 1) + val i = if i = 0 then "" else "_" ^ Int.toString i + val s = skolemPrefix ^ "_" ^ s ^ i + in + Name.fromString s + end end; fun skolemize fv bv fm = let val fv = NameSet.transform Term.Var fv - + fun mk (v,s) = Subst.insert s (v, Term.Fn (newSkolemFunction v, fv)) in subst (NameSet.foldl mk Subst.empty bv) fm @@ -11158,7 +14746,7 @@ in (NameSet.add a v', Subst.insert s (v, Term.Var v')) end - + val avoid = NameSet.union (NameSet.union avoid fv) bv val (_,sub) = NameSet.foldl ren (avoid,Subst.empty) captured @@ -11167,35 +14755,43 @@ end end; - fun cnf avoid fm = -(*TRACE5 - let - val fm' = cnf' avoid fm - val () = Parser.ppTrace pp "Normalize.cnf: fm" fm - val () = Parser.ppTrace pp "Normalize.cnf: fm'" fm' + fun cnfFm avoid fm = +(*MetisTrace5 + let + val fm' = cnfFm' avoid fm + val () = Print.trace pp "Normalize.cnfFm: fm" fm + val () = Print.trace pp "Normalize.cnfFm: fm'" fm' in fm' end - and cnf' avoid fm = + and cnfFm' avoid fm = *) case fm of True => True | False => False | Literal _ => fm - | And (_,_,s) => AndList (Set.transform (cnf avoid) s) - | Or (_,_,s) => pushOrList (snd (Set.foldl cnfOr (avoid,[]) s)) - | Xor _ => cnf avoid (pushXor fm) - | Exists (fv,_,n,f) => cnf avoid (skolemize fv n f) - | Forall (fv,_,n,f) => cnf avoid (rename avoid fv n f) - - and cnfOr (fm,(avoid,acc)) = - let - val fm = cnf avoid fm - in - (NameSet.union (freeVars fm) avoid, fm :: acc) - end; -in - val basicCnf = cnf NameSet.empty; + | And (_,_,s) => AndList (Set.transform (cnfFm avoid) s) + | Or (fv,_,s) => + let + val avoid = NameSet.union avoid fv + val (fms,_) = Set.foldl cnfOr ([],avoid) s + in + pushOrList fms + end + | Xor _ => cnfFm avoid (pushXor fm) + | Exists (fv,_,n,f) => cnfFm avoid (skolemize fv n f) + | Forall (fv,_,n,f) => cnfFm avoid (rename avoid fv n f) + + and cnfOr (fm,(fms,avoid)) = + let + val fm = cnfFm avoid fm + val fms = fm :: fms + val avoid = NameSet.union avoid (freeVars fm) + in + (fms,avoid) + end; +in + val basicCnf = cnfFm NameSet.empty; end; (* ------------------------------------------------------------------------- *) @@ -11203,7 +14799,7 @@ (* ------------------------------------------------------------------------- *) local - type best = real * formula option; + type best = count * formula option; fun minBreak countClauses fm best = case fm of @@ -11218,7 +14814,7 @@ minBreakSet countClauses countXor2 countFalse XorSet s best | Exists (_,_,_,f) => minBreak countClauses f best | Forall (_,_,_,f) => minBreak countClauses f best - + and minBreakSet countClauses count2 count0 mkSet fmSet best = let fun cumulatives fms = @@ -11242,7 +14838,11 @@ val (c1,_,fms) = foldl fwd (count0,empty,[]) fms val (c2,_,fms) = foldl bwd (count0,empty,[]) fms - val _ = countEqualish c1 c2 orelse raise Bug "cumulativeCounts" +(*MetisDebug + val _ = countEquivish c1 c2 orelse + raise Bug ("cumulativeCounts: c1 = " ^ countToString c1 ^ + ", c2 = " ^ countToString c2) +*) in fms end @@ -11250,6 +14850,7 @@ fun breakSing ((c1,_,fm,c2,_),best) = let val cFms = count2 (c1,c2) + fun countCls cFm = countClauses (count2 (cFms,cFm)) in minBreak countCls fm best @@ -11261,13 +14862,13 @@ if Set.null s1 then best else let - val cDef = countDefinition (count2 (c1, count fm)) + val cDef = countDefinition (countXor2 (c1, count fm)) val cFm = count2 (countLiteral,c2) - val cl = positive cDef + countClauses cFm - val better = cl < bcl - 0.5 + val cl = countAnd2 (cDef, countClauses cFm) + val noBetter = countLeqish bcl cl in - if better then (cl, SOME (mkSet (Set.add s1 fm))) - else best + if noBetter then best + else (cl, SOME (mkSet (Set.add s1 fm))) end in fn ((c1,s1,fm,c2,s2),best) => @@ -11278,14 +14879,14 @@ fun breakSet measure best = let - val fms = sortMap (measure o count) Real.compare fms + val fms = sortMap (measure o count) countCompare fms in foldl breakSet1 best (cumulatives fms) end val best = foldl breakSing best (cumulatives fms) - val best = breakSet positive best - val best = breakSet negative best + val best = breakSet I best + val best = breakSet countNegate best val best = breakSet countClauses best in best @@ -11293,21 +14894,20 @@ in fun minimumDefinition fm = let - val countClauses = positive - val cl = countClauses (count fm) - in - if cl < 1.5 then NONE + val cl = count fm + in + if countLeqish cl countLiteral then NONE else let - val (cl',def) = minBreak countClauses fm (cl,NONE) -(*TRACE1 + val (cl',def) = minBreak I fm (cl,NONE) +(*MetisTrace1 val () = case def of NONE => () | SOME d => - Parser.ppTrace pp ("defCNF: before = " ^ Real.toString cl ^ - ", after = " ^ Real.toString cl' ^ - ", definition") d + Print.trace pp ("defCNF: before = " ^ countToString cl ^ + ", after = " ^ countToString cl' ^ + ", definition") d *) in def @@ -11316,78 +14916,493 @@ end; (* ------------------------------------------------------------------------- *) -(* Conjunctive normal form. *) +(* Conjunctive normal form derivations. *) +(* ------------------------------------------------------------------------- *) + +datatype thm = Thm of formula * inference + +and inference = + Axiom of Formula.formula + | Definition of string * Formula.formula + | Simplify of thm * thm list + | Conjunct of thm + | Specialize of thm + | Skolemize of thm + | Clausify of thm; + +fun parentsInference inf = + case inf of + Axiom _ => [] + | Definition _ => [] + | Simplify (th,ths) => th :: ths + | Conjunct th => [th] + | Specialize th => [th] + | Skolemize th => [th] + | Clausify th => [th]; + +fun compareThm (Thm (fm1,_), Thm (fm2,_)) = compare (fm1,fm2); + +fun parentsThm (Thm (_,inf)) = parentsInference inf; + +fun mkAxiom fm = Thm (fromFormula fm, Axiom fm); + +fun destThm (Thm (fm,inf)) = (toFormula fm, inf); + +local + val emptyProved : (thm,Formula.formula) Map.map = Map.new compareThm; + + fun isProved proved th = Map.inDomain th proved; + + fun isUnproved proved th = not (isProved proved th); + + fun lookupProved proved th = + case Map.peek proved th of + SOME fm => fm + | NONE => raise Bug "Normalize.lookupProved"; + + fun prove acc proved ths = + case ths of + [] => rev acc + | th :: ths' => + if isProved proved th then prove acc proved ths' + else + let + val pars = parentsThm th + + val deps = List.filter (isUnproved proved) pars + in + if null deps then + let + val (fm,inf) = destThm th + + val fms = map (lookupProved proved) pars + + val acc = (fm,inf,fms) :: acc + + val proved = Map.insert proved (th,fm) + in + prove acc proved ths' + end + else + let + val ths = deps @ ths + in + prove acc proved ths + end + end; +in + val proveThms = prove [] emptyProved; +end; + +fun toStringInference inf = + case inf of + Axiom _ => "Axiom" + | Definition _ => "Definition" + | Simplify _ => "Simplify" + | Conjunct _ => "Conjunct" + | Specialize _ => "Specialize" + | Skolemize _ => "Skolemize" + | Clausify _ => "Clausify"; + +val ppInference = Print.ppMap toStringInference Print.ppString; + +(* ------------------------------------------------------------------------- *) +(* Simplifying with definitions. *) +(* ------------------------------------------------------------------------- *) + +datatype simplify = + Simp of + {formula : (formula, formula * thm) Map.map, + andSet : (formula Set.set * formula * thm) list, + orSet : (formula Set.set * formula * thm) list, + xorSet : (formula Set.set * formula * thm) list}; + +val simplifyEmpty = + Simp + {formula = Map.new compare, + andSet = [], + orSet = [], + xorSet = []}; + +local + fun simpler fm s = + Set.size s <> 1 orelse + case Set.pick s of + True => false + | False => false + | Literal _ => false + | _ => true; + + fun addSet set_defs body_def = + let + fun def_body_size (body,_,_) = Set.size body + + val body_size = def_body_size body_def + + val (body,_,_) = body_def + + fun add acc [] = List.revAppend (acc,[body_def]) + | add acc (l as (bd as (b,_,_)) :: bds) = + case Int.compare (def_body_size bd, body_size) of + LESS => List.revAppend (acc, body_def :: l) + | EQUAL => + if Set.equal b body then List.revAppend (acc,l) + else add (bd :: acc) bds + | GREATER => add (bd :: acc) bds + in + add [] set_defs + end; + + fun add simp (body,False,th) = add simp (negate body, True, th) + | add simp (True,_,_) = simp + | add (Simp {formula,andSet,orSet,xorSet}) (And (_,_,s), def, th) = + let + val andSet = addSet andSet (s,def,th) + and orSet = addSet orSet (negateSet s, negate def, th) + in + Simp + {formula = formula, + andSet = andSet, + orSet = orSet, + xorSet = xorSet} + end + | add (Simp {formula,andSet,orSet,xorSet}) (Or (_,_,s), def, th) = + let + val orSet = addSet orSet (s,def,th) + and andSet = addSet andSet (negateSet s, negate def, th) + in + Simp + {formula = formula, + andSet = andSet, + orSet = orSet, + xorSet = xorSet} + end + | add simp (Xor (_,_,p,s), def, th) = + let + val simp = addXorSet simp (s, applyPolarity p def, th) + in + case def of + True => + let + fun addXorLiteral (fm as Literal _, simp) = + let + val s = Set.delete s fm + in + if not (simpler fm s) then simp + else addXorSet simp (s, applyPolarity (not p) fm, th) + end + | addXorLiteral (_,simp) = simp + in + Set.foldl addXorLiteral simp s + end + | _ => simp + end + | add (simp as Simp {formula,andSet,orSet,xorSet}) (body,def,th) = + if Map.inDomain body formula then simp + else + let + val formula = Map.insert formula (body,(def,th)) + val formula = Map.insert formula (negate body, (negate def, th)) + in + Simp + {formula = formula, + andSet = andSet, + orSet = orSet, + xorSet = xorSet} + end + + and addXorSet (simp as Simp {formula,andSet,orSet,xorSet}) (s,def,th) = + if Set.size s = 1 then add simp (Set.pick s, def, th) + else + let + val xorSet = addSet xorSet (s,def,th) + in + Simp + {formula = formula, + andSet = andSet, + orSet = orSet, + xorSet = xorSet} + end; +in + fun simplifyAdd simp (th as Thm (fm,_)) = add simp (fm,True,th); +end; + +local + fun simplifySet set_defs set = + let + fun pred (s,_,_) = Set.subset s set + in + case List.find pred set_defs of + NONE => NONE + | SOME (s,f,th) => + let + val set = Set.add (Set.difference set s) f + in + SOME (set,th) + end + end; +in + fun simplify (Simp {formula,andSet,orSet,xorSet}) = + let + fun simp fm inf = + case simp_sub fm inf of + NONE => simp_top fm inf + | SOME (fm,inf) => try_simp_top fm inf + + and try_simp_top fm inf = + case simp_top fm inf of + NONE => SOME (fm,inf) + | x => x + + and simp_top fm inf = + case fm of + And (_,_,s) => + (case simplifySet andSet s of + NONE => NONE + | SOME (s,th) => + let + val fm = AndSet s + val inf = th :: inf + in + try_simp_top fm inf + end) + | Or (_,_,s) => + (case simplifySet orSet s of + NONE => NONE + | SOME (s,th) => + let + val fm = OrSet s + val inf = th :: inf + in + try_simp_top fm inf + end) + | Xor (_,_,p,s) => + (case simplifySet xorSet s of + NONE => NONE + | SOME (s,th) => + let + val fm = XorPolaritySet (p,s) + val inf = th :: inf + in + try_simp_top fm inf + end) + | _ => + (case Map.peek formula fm of + NONE => NONE + | SOME (fm,th) => + let + val inf = th :: inf + in + try_simp_top fm inf + end) + + and simp_sub fm inf = + case fm of + And (_,_,s) => + (case simp_set s inf of + NONE => NONE + | SOME (l,inf) => SOME (AndList l, inf)) + | Or (_,_,s) => + (case simp_set s inf of + NONE => NONE + | SOME (l,inf) => SOME (OrList l, inf)) + | Xor (_,_,p,s) => + (case simp_set s inf of + NONE => NONE + | SOME (l,inf) => SOME (XorPolarityList (p,l), inf)) + | Exists (_,_,n,f) => + (case simp f inf of + NONE => NONE + | SOME (f,inf) => SOME (ExistsSet (n,f), inf)) + | Forall (_,_,n,f) => + (case simp f inf of + NONE => NONE + | SOME (f,inf) => SOME (ForallSet (n,f), inf)) + | _ => NONE + + and simp_set s inf = + let + val (changed,l,inf) = Set.foldr simp_set_elt (false,[],inf) s + in + if changed then SOME (l,inf) else NONE + end + + and simp_set_elt (fm,(changed,l,inf)) = + case simp fm inf of + NONE => (changed, fm :: l, inf) + | SOME (fm,inf) => (true, fm :: l, inf) + in + fn th as Thm (fm,_) => + case simp fm [] of + SOME (fm,ths) => + let + val inf = Simplify (th,ths) + in + Thm (fm,inf) + end + | NONE => th + end; +end; + +(*MetisTrace2 +val simplify = fn simp => fn th as Thm (fm,_) => + let + val th' as Thm (fm',_) = simplify simp th + val () = if compare (fm,fm') = EQUAL then () + else (Print.trace pp "Normalize.simplify: fm" fm; + Print.trace pp "Normalize.simplify: fm'" fm') + in + th' + end; +*) + +(* ------------------------------------------------------------------------- *) +(* Definitions. *) (* ------------------------------------------------------------------------- *) val newDefinitionRelation = let val counter : int Unsynchronized.ref = Unsynchronized.ref 0 in - fn () => CRITICAL (fn () => - let - val Unsynchronized.ref i = counter - val () = counter := i + 1 - in - "defCNF_" ^ Int.toString i - end) + fn () => + let + val Unsynchronized.ref i = counter + val () = counter := i + 1 + in + definitionPrefix ^ "_" ^ Int.toString i + end end; fun newDefinition def = let val fv = freeVars def - val atm = (newDefinitionRelation (), NameSet.transform Term.Var fv) - val lit = Literal (fv,(true,atm)) - in - Xor2 (lit,def) - end; - -local - fun def_cnf acc [] = acc - | def_cnf acc ((prob,simp,fms) :: probs) = - def_cnf_problem acc prob simp fms probs - - and def_cnf_problem acc prob _ [] probs = def_cnf (prob :: acc) probs - | def_cnf_problem acc prob simp (fm :: fms) probs = - def_cnf_formula acc prob simp (simplify simp fm) fms probs - - and def_cnf_formula acc prob simp fm fms probs = + val rel = newDefinitionRelation () + val atm = (Name.fromString rel, NameSet.transform Term.Var fv) + val fm = Formula.Iff (Formula.Atom atm, toFormula def) + val fm = Formula.setMkForall (fv,fm) + val inf = Definition (rel,fm) + val lit = Literal (fv,(false,atm)) + val fm = Xor2 (lit,def) + in + Thm (fm,inf) + end; + +(* ------------------------------------------------------------------------- *) +(* Definitional conjunctive normal form. *) +(* ------------------------------------------------------------------------- *) + +datatype cnf = + ConsistentCnf of simplify + | InconsistentCnf; + +val initialCnf = ConsistentCnf simplifyEmpty; + +local + fun def_cnf_inconsistent th = + let + val cls = [(LiteralSet.empty,th)] + in + (cls,InconsistentCnf) + end; + + fun def_cnf_clause inf (fm,acc) = + let + val cl = toClause fm + val th = Thm (fm,inf) + in + (cl,th) :: acc + end +(*MetisDebug + handle Error err => + (Print.trace pp "Normalize.addCnf.def_cnf_clause: fm" fm; + raise Bug ("Normalize.addCnf.def_cnf_clause: " ^ err)); +*) + + fun def_cnf cls simp ths = + case ths of + [] => (cls, ConsistentCnf simp) + | th :: ths => def_cnf_formula cls simp (simplify simp th) ths + + and def_cnf_formula cls simp (th as Thm (fm,_)) ths = case fm of - True => def_cnf_problem acc prob simp fms probs - | False => def_cnf acc probs - | And (_,_,s) => def_cnf_problem acc prob simp (Set.toList s @ fms) probs + True => def_cnf cls simp ths + | False => def_cnf_inconsistent th + | And (_,_,s) => + let + fun add (f,z) = Thm (f, Conjunct th) :: z + in + def_cnf cls simp (Set.foldr add ths s) + end | Exists (fv,_,n,f) => - def_cnf_formula acc prob simp (skolemize fv n f) fms probs - | Forall (_,_,_,f) => def_cnf_formula acc prob simp f fms probs + let + val th = Thm (skolemize fv n f, Skolemize th) + in + def_cnf_formula cls simp th ths + end + | Forall (_,_,_,f) => + let + val th = Thm (f, Specialize th) + in + def_cnf_formula cls simp th ths + end | _ => case minimumDefinition fm of - NONE => - let - val prob = fm :: prob - and simp = simplifyAdd simp fm - in - def_cnf_problem acc prob simp fms probs - end - | SOME def => - let - val def_fm = newDefinition def - and fms = fm :: fms - in - def_cnf_formula acc prob simp def_fm fms probs + SOME def => + let + val ths = th :: ths + val th = newDefinition def + in + def_cnf_formula cls simp th ths + end + | NONE => + let + val simp = simplifyAdd simp th + + val fm = basicCnf fm + + val inf = Clausify th + in + case fm of + True => def_cnf cls simp ths + | False => def_cnf_inconsistent (Thm (fm,inf)) + | And (_,_,s) => + let + val inf = Conjunct (Thm (fm,inf)) + val cls = Set.foldl (def_cnf_clause inf) cls s + in + def_cnf cls simp ths + end + | fm => def_cnf (def_cnf_clause inf (fm,cls)) simp ths end; - - fun cnf_prob prob = toFormula (AndList (map basicCnf prob)); -in - fun cnf fm = - let - val fm = fromFormula fm -(*TRACE2 - val () = Parser.ppTrace pp "Normalize.cnf: fm" fm -*) - val probs = def_cnf [] [([],simplifyEmpty,[fm])] - in - map cnf_prob probs - end; -end; +in + fun addCnf th cnf = + case cnf of + ConsistentCnf simp => def_cnf [] simp [th] + | InconsistentCnf => ([],cnf); +end; + +local + fun add (th,(cls,cnf)) = + let + val (cls',cnf) = addCnf th cnf + in + (cls' @ cls, cnf) + end; +in + fun proveCnf ths = + let + val (cls,_) = List.foldl add ([],initialCnf) ths + in + rev cls + end; +end; + +fun cnf fm = + let + val cls = proveCnf [mkAxiom fm] + in + map fst cls + end; end end; @@ -11396,40 +15411,198 @@ (* ========================================================================= *) (* RANDOM FINITE MODELS *) -(* Copyright (c) 2003-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Model = sig (* ------------------------------------------------------------------------- *) +(* Model size. *) +(* ------------------------------------------------------------------------- *) + +type size = {size : int} + +(* ------------------------------------------------------------------------- *) +(* A model of size N has integer elements 0...N-1. *) +(* ------------------------------------------------------------------------- *) + +type element = int + +val zeroElement : element + +val incrementElement : size -> element -> element option + +(* ------------------------------------------------------------------------- *) (* The parts of the model that are fixed. *) -(* Note: a model of size N has integer elements 0...N-1. *) -(* ------------------------------------------------------------------------- *) - -type fixed = - {size : int} -> - {functions : (Metis.Term.functionName * int list) -> int option, - relations : (Metis.Atom.relationName * int list) -> bool option} - -val fixedMerge : fixed -> fixed -> fixed (* Prefers the second fixed *) - -val fixedMergeList : fixed list -> fixed - -val fixedPure : fixed (* : = *) - -val fixedBasic : fixed (* id fst snd #1 #2 #3 <> *) - -val fixedModulo : fixed (* suc pre ~ + - * exp div mod *) - (* is_0 divides even odd *) - -val fixedOverflowNum : fixed (* suc pre + - * exp div mod *) - (* is_0 <= < >= > divides even odd *) - -val fixedOverflowInt : fixed (* suc pre + - * exp div mod *) - (* is_0 <= < >= > divides even odd *) - -val fixedSet : fixed (* empty univ union intersect compl card in subset *) +(* ------------------------------------------------------------------------- *) + +type fixedFunction = size -> element list -> element option + +type fixedRelation = size -> element list -> bool option + +datatype fixed = + Fixed of + {functions : fixedFunction Metis.NameArityMap.map, + relations : fixedRelation Metis.NameArityMap.map} + +val emptyFixed : fixed + +val unionFixed : fixed -> fixed -> fixed + +val getFunctionFixed : fixed -> Metis.NameArity.nameArity -> fixedFunction + +val getRelationFixed : fixed -> Metis.NameArity.nameArity -> fixedRelation + +val insertFunctionFixed : fixed -> Metis.NameArity.nameArity * fixedFunction -> fixed + +val insertRelationFixed : fixed -> Metis.NameArity.nameArity * fixedRelation -> fixed + +val unionListFixed : fixed list -> fixed + +val basicFixed : fixed (* interprets equality and hasType *) + +(* ------------------------------------------------------------------------- *) +(* Renaming fixed model parts. *) +(* ------------------------------------------------------------------------- *) + +type fixedMap = + {functionMap : Metis.Name.name Metis.NameArityMap.map, + relationMap : Metis.Name.name Metis.NameArityMap.map} + +val mapFixed : fixedMap -> fixed -> fixed + +val ppFixedMap : fixedMap Metis.Print.pp + +(* ------------------------------------------------------------------------- *) +(* Standard fixed model parts. *) +(* ------------------------------------------------------------------------- *) + +(* Projections *) + +val projectionMin : int + +val projectionMax : int + +val projectionName : int -> Metis.Name.name + +val projectionFixed : fixed + +(* Arithmetic *) + +val numeralMin : int + +val numeralMax : int + +val numeralName : int -> Metis.Name.name + +val addName : Metis.Name.name + +val divName : Metis.Name.name + +val dividesName : Metis.Name.name + +val evenName : Metis.Name.name + +val expName : Metis.Name.name + +val geName : Metis.Name.name + +val gtName : Metis.Name.name + +val isZeroName : Metis.Name.name + +val leName : Metis.Name.name + +val ltName : Metis.Name.name + +val modName : Metis.Name.name + +val multName : Metis.Name.name + +val negName : Metis.Name.name + +val oddName : Metis.Name.name + +val preName : Metis.Name.name + +val subName : Metis.Name.name + +val sucName : Metis.Name.name + +val modularFixed : fixed + +val overflowFixed : fixed + +(* Sets *) + +val cardName : Metis.Name.name + +val complementName : Metis.Name.name + +val differenceName : Metis.Name.name + +val emptyName : Metis.Name.name + +val memberName : Metis.Name.name + +val insertName : Metis.Name.name + +val intersectName : Metis.Name.name + +val singletonName : Metis.Name.name + +val subsetName : Metis.Name.name + +val symmetricDifferenceName : Metis.Name.name + +val unionName : Metis.Name.name + +val universeName : Metis.Name.name + +val setFixed : fixed + +(* Lists *) + +val appendName : Metis.Name.name + +val consName : Metis.Name.name + +val lengthName : Metis.Name.name + +val nilName : Metis.Name.name + +val nullName : Metis.Name.name + +val tailName : Metis.Name.name + +val listFixed : fixed + +(* ------------------------------------------------------------------------- *) +(* Valuations. *) +(* ------------------------------------------------------------------------- *) + +type valuation + +val emptyValuation : valuation + +val zeroValuation : Metis.NameSet.set -> valuation + +val constantValuation : element -> Metis.NameSet.set -> valuation + +val peekValuation : valuation -> Metis.Name.name -> element option + +val getValuation : valuation -> Metis.Name.name -> element + +val insertValuation : valuation -> Metis.Name.name * element -> valuation + +val randomValuation : {size : int} -> Metis.NameSet.set -> valuation + +val incrementValuation : + {size : int} -> Metis.NameSet.set -> valuation -> valuation option + +val foldValuation : + {size : int} -> Metis.NameSet.set -> (valuation * 'a -> 'a) -> 'a -> 'a (* ------------------------------------------------------------------------- *) (* A type of random finite models. *) @@ -11439,28 +15612,21 @@ type model +val default : parameters + val new : parameters -> model val size : model -> int (* ------------------------------------------------------------------------- *) -(* Valuations. *) -(* ------------------------------------------------------------------------- *) - -type valuation = int Metis.NameMap.map - -val valuationEmpty : valuation - -val valuationRandom : {size : int} -> Metis.NameSet.set -> valuation - -val valuationFold : - {size : int} -> Metis.NameSet.set -> (valuation * 'a -> 'a) -> 'a -> 'a - -(* ------------------------------------------------------------------------- *) (* Interpreting terms and formulas in the model. *) (* ------------------------------------------------------------------------- *) -val interpretTerm : model -> valuation -> Metis.Term.term -> int +val interpretFunction : model -> Metis.Term.functionName * element list -> element + +val interpretRelation : model -> Metis.Atom.relationName * element list -> bool + +val interpretTerm : model -> valuation -> Metis.Term.term -> element val interpretAtom : model -> valuation -> Metis.Atom.atom -> bool @@ -11475,17 +15641,49 @@ (* Note: if it's cheaper, a systematic check will be performed instead. *) (* ------------------------------------------------------------------------- *) +val check : + (model -> valuation -> 'a -> bool) -> {maxChecks : int option} -> model -> + Metis.NameSet.set -> 'a -> {T : int, F : int} + val checkAtom : - {maxChecks : int} -> model -> Metis.Atom.atom -> {T : int, F : int} + {maxChecks : int option} -> model -> Metis.Atom.atom -> {T : int, F : int} val checkFormula : - {maxChecks : int} -> model -> Metis.Formula.formula -> {T : int, F : int} + {maxChecks : int option} -> model -> Metis.Formula.formula -> {T : int, F : int} val checkLiteral : - {maxChecks : int} -> model -> Metis.Literal.literal -> {T : int, F : int} + {maxChecks : int option} -> model -> Metis.Literal.literal -> {T : int, F : int} val checkClause : - {maxChecks : int} -> model -> Metis.Thm.clause -> {T : int, F : int} + {maxChecks : int option} -> model -> Metis.Thm.clause -> {T : int, F : int} + +(* ------------------------------------------------------------------------- *) +(* Updating the model. *) +(* ------------------------------------------------------------------------- *) + +val updateFunction : + model -> (Metis.Term.functionName * element list) * element -> unit + +val updateRelation : + model -> (Metis.Atom.relationName * element list) * bool -> unit + +(* ------------------------------------------------------------------------- *) +(* Choosing a random perturbation to make a formula true in the model. *) +(* ------------------------------------------------------------------------- *) + +val perturbTerm : model -> valuation -> Metis.Term.term * element list -> unit + +val perturbAtom : model -> valuation -> Metis.Atom.atom * bool -> unit + +val perturbLiteral : model -> valuation -> Metis.Literal.literal -> unit + +val perturbClause : model -> valuation -> Metis.Thm.clause -> unit + +(* ------------------------------------------------------------------------- *) +(* Pretty printing. *) +(* ------------------------------------------------------------------------- *) + +val pp : model Metis.Print.pp end @@ -11493,7 +15691,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; @@ -11502,7 +15700,7 @@ (* ========================================================================= *) (* RANDOM FINITE MODELS *) -(* Copyright (c) 2003-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Model :> Model = @@ -11511,378 +15709,912 @@ open Useful; (* ------------------------------------------------------------------------- *) +(* Constants. *) +(* ------------------------------------------------------------------------- *) + +val maxSpace = 1000; + +(* ------------------------------------------------------------------------- *) (* Helper functions. *) (* ------------------------------------------------------------------------- *) -fun intExp x y = exp op* x y 1; - -fun natFromString "" = NONE - | natFromString "0" = SOME 0 - | natFromString s = - case charToInt (String.sub (s,0)) of +val multInt = + case Int.maxInt of + NONE => (fn x => fn y => SOME (x * y)) + | SOME m => + let + val m = Real.floor (Math.sqrt (Real.fromInt m)) + in + fn x => fn y => if x <= m andalso y <= m then SOME (x * y) else NONE + end; + +local + fun iexp x y acc = + if y mod 2 = 0 then iexp' x y acc + else + case multInt acc x of + SOME acc => iexp' x y acc + | NONE => NONE + + and iexp' x y acc = + if y = 1 then SOME acc + else + let + val y = y div 2 + in + case multInt x x of + SOME x => iexp x y acc + | NONE => NONE + end; +in + fun expInt x y = + if y <= 1 then + if y = 0 then SOME 1 + else if y = 1 then SOME x + else raise Bug "expInt: negative exponent" + else if x <= 1 then + if 0 <= x then SOME x + else raise Bug "expInt: negative exponand" + else iexp x y 1; +end; + +fun boolToInt true = 1 + | boolToInt false = 0; + +fun intToBool 1 = true + | intToBool 0 = false + | intToBool _ = raise Bug "Model.intToBool"; + +fun minMaxInterval i j = interval i (1 + j - i); + +(* ------------------------------------------------------------------------- *) +(* Model size. *) +(* ------------------------------------------------------------------------- *) + +type size = {size : int}; + +(* ------------------------------------------------------------------------- *) +(* A model of size N has integer elements 0...N-1. *) +(* ------------------------------------------------------------------------- *) + +type element = int; + +val zeroElement = 0; + +fun incrementElement {size = N} i = + let + val i = i + 1 + in + if i = N then NONE else SOME i + end; + +fun elementListSpace {size = N} arity = + case expInt N arity of NONE => NONE - | SOME 0 => NONE - | SOME d => - let - fun parse 0 _ acc = SOME acc - | parse n i acc = - case charToInt (String.sub (s,i)) of - NONE => NONE - | SOME d => parse (n - 1) (i + 1) (10 * acc + d) - in - parse (size s - 1) 1 d - end; - -fun projection (_,[]) = NONE - | projection ("#1", x :: _) = SOME x - | projection ("#2", _ :: x :: _) = SOME x - | projection ("#3", _ :: _ :: x :: _) = SOME x - | projection (func,args) = - let - val f = size func - and n = length args - - val p = - if f < 2 orelse n <= 3 orelse String.sub (func,0) <> #"#" then NONE - else if f = 2 then - (case charToInt (String.sub (func,1)) of - NONE => NONE - | p as SOME d => if d <= 3 then NONE else p) - else if (n < intExp 10 (f - 2) handle Overflow => true) then NONE - else - (natFromString (String.extract (func,1,NONE)) - handle Overflow => NONE) - in - case p of - NONE => NONE - | SOME k => if k > n then NONE else SOME (List.nth (args, k - 1)) + | s as SOME m => if m <= maxSpace then s else NONE; + +fun elementListIndex {size = N} = + let + fun f acc elts = + case elts of + [] => acc + | elt :: elts => f (N * acc + elt) elts + in + f 0 end; (* ------------------------------------------------------------------------- *) (* The parts of the model that are fixed. *) -(* Note: a model of size N has integer elements 0...N-1. *) -(* ------------------------------------------------------------------------- *) - -type fixedModel = - {functions : (Term.functionName * int list) -> int option, - relations : (Atom.relationName * int list) -> bool option}; - -type fixed = {size : int} -> fixedModel - -fun fixedMerge fixed1 fixed2 parm = - let - val {functions = f1, relations = r1} = fixed1 parm - and {functions = f2, relations = r2} = fixed2 parm - - fun functions x = case f2 x of NONE => f1 x | s => s - - fun relations x = case r2 x of NONE => r1 x | s => s - in - {functions = functions, relations = relations} - end; - -fun fixedMergeList [] = raise Bug "fixedMergeList: empty" - | fixedMergeList (f :: l) = foldl (uncurry fixedMerge) f l; - -fun fixedPure {size = _} = - let - fun functions (":",[x,_]) = SOME x - | functions _ = NONE - - fun relations (rel,[x,y]) = - if (rel,2) = Atom.eqRelation then SOME (x = y) else NONE - | relations _ = NONE - in - {functions = functions, relations = relations} - end; - -fun fixedBasic {size = _} = - let - fun functions ("id",[x]) = SOME x - | functions ("fst",[x,_]) = SOME x - | functions ("snd",[_,x]) = SOME x - | functions func_args = projection func_args - - fun relations ("<>",[x,y]) = SOME (x <> y) - | relations _ = NONE - in - {functions = functions, relations = relations} - end; - -fun fixedModulo {size = N} = - let - fun mod_N k = k mod N - - val one = mod_N 1 - - fun mult (x,y) = mod_N (x * y) - - fun divides_N 0 = false - | divides_N x = N mod x = 0 - - val even_N = divides_N 2 - - fun functions (numeral,[]) = - Option.map mod_N (natFromString numeral handle Overflow => NONE) - | functions ("suc",[x]) = SOME (if x = N - 1 then 0 else x + 1) - | functions ("pre",[x]) = SOME (if x = 0 then N - 1 else x - 1) - | functions ("~",[x]) = SOME (if x = 0 then 0 else N - x) - | functions ("+",[x,y]) = SOME (mod_N (x + y)) - | functions ("-",[x,y]) = SOME (if x < y then N + x - y else x - y) - | functions ("*",[x,y]) = SOME (mult (x,y)) - | functions ("exp",[x,y]) = SOME (exp mult x y one) - | functions ("div",[x,y]) = if divides_N y then SOME (x div y) else NONE - | functions ("mod",[x,y]) = if divides_N y then SOME (x mod y) else NONE - | functions _ = NONE - - fun relations ("is_0",[x]) = SOME (x = 0) - | relations ("divides",[x,y]) = - if x = 0 then SOME (y = 0) - else if divides_N x then SOME (y mod x = 0) else NONE - | relations ("even",[x]) = if even_N then SOME (x mod 2 = 0) else NONE - | relations ("odd",[x]) = if even_N then SOME (x mod 2 = 1) else NONE - | relations _ = NONE - in - {functions = functions, relations = relations} - end; - -local - datatype onum = ONeg | ONum of int | OInf; - - val zero = ONum 0 - and one = ONum 1 - and two = ONum 2; - - fun suc (ONum x) = ONum (x + 1) - | suc v = v; - - fun pre (ONum 0) = ONeg - | pre (ONum x) = ONum (x - 1) - | pre v = v; - - fun neg ONeg = NONE - | neg (n as ONum 0) = SOME n - | neg _ = SOME ONeg; - - fun add ONeg ONeg = SOME ONeg - | add ONeg (ONum y) = if y = 0 then SOME ONeg else NONE - | add ONeg OInf = NONE - | add (ONum x) ONeg = if x = 0 then SOME ONeg else NONE - | add (ONum x) (ONum y) = SOME (ONum (x + y)) - | add (ONum _) OInf = SOME OInf - | add OInf ONeg = NONE - | add OInf (ONum _) = SOME OInf - | add OInf OInf = SOME OInf; - - fun sub ONeg ONeg = NONE - | sub ONeg (ONum _) = SOME ONeg - | sub ONeg OInf = SOME ONeg - | sub (ONum _) ONeg = NONE - | sub (ONum x) (ONum y) = SOME (if x < y then ONeg else ONum (x - y)) - | sub (ONum _) OInf = SOME ONeg - | sub OInf ONeg = SOME OInf - | sub OInf (ONum y) = if y = 0 then SOME OInf else NONE - | sub OInf OInf = NONE; - - fun mult ONeg ONeg = NONE - | mult ONeg (ONum y) = SOME (if y = 0 then zero else ONeg) - | mult ONeg OInf = SOME ONeg - | mult (ONum x) ONeg = SOME (if x = 0 then zero else ONeg) - | mult (ONum x) (ONum y) = SOME (ONum (x * y)) - | mult (ONum x) OInf = SOME (if x = 0 then zero else OInf) - | mult OInf ONeg = SOME ONeg - | mult OInf (ONum y) = SOME (if y = 0 then zero else OInf) - | mult OInf OInf = SOME OInf; - - fun exp ONeg ONeg = NONE - | exp ONeg (ONum y) = - if y = 0 then SOME one else if y mod 2 = 0 then NONE else SOME ONeg - | exp ONeg OInf = NONE - | exp (ONum x) ONeg = NONE - | exp (ONum x) (ONum y) = SOME (ONum (intExp x y) handle Overflow => OInf) - | exp (ONum x) OInf = - SOME (if x = 0 then zero else if x = 1 then one else OInf) - | exp OInf ONeg = NONE - | exp OInf (ONum y) = SOME (if y = 0 then one else OInf) - | exp OInf OInf = SOME OInf; - - fun odiv ONeg ONeg = NONE - | odiv ONeg (ONum y) = if y = 1 then SOME ONeg else NONE - | odiv ONeg OInf = NONE - | odiv (ONum _) ONeg = NONE - | odiv (ONum x) (ONum y) = if y = 0 then NONE else SOME (ONum (x div y)) - | odiv (ONum _) OInf = SOME zero - | odiv OInf ONeg = NONE - | odiv OInf (ONum y) = if y = 1 then SOME OInf else NONE - | odiv OInf OInf = NONE; - - fun omod ONeg ONeg = NONE - | omod ONeg (ONum y) = if y = 1 then SOME zero else NONE - | omod ONeg OInf = NONE - | omod (ONum _) ONeg = NONE - | omod (ONum x) (ONum y) = if y = 0 then NONE else SOME (ONum (x mod y)) - | omod (x as ONum _) OInf = SOME x - | omod OInf ONeg = NONE - | omod OInf (ONum y) = if y = 1 then SOME OInf else NONE - | omod OInf OInf = NONE; - - fun le ONeg ONeg = NONE - | le ONeg (ONum y) = SOME true - | le ONeg OInf = SOME true - | le (ONum _) ONeg = SOME false - | le (ONum x) (ONum y) = SOME (x <= y) - | le (ONum _) OInf = SOME true - | le OInf ONeg = SOME false - | le OInf (ONum _) = SOME false - | le OInf OInf = NONE; - - fun lt x y = Option.map not (le y x); - - fun ge x y = le y x; - - fun gt x y = lt y x; - - fun divides ONeg ONeg = NONE - | divides ONeg (ONum y) = if y = 0 then SOME true else NONE - | divides ONeg OInf = NONE - | divides (ONum x) ONeg = - if x = 0 then SOME false else if x = 1 then SOME true else NONE - | divides (ONum x) (ONum y) = SOME (Useful.divides x y) - | divides (ONum x) OInf = - if x = 0 then SOME false else if x = 1 then SOME true else NONE - | divides OInf ONeg = NONE - | divides OInf (ONum y) = SOME (y = 0) - | divides OInf OInf = NONE; - - fun even n = divides two n; - - fun odd n = Option.map not (even n); - - fun fixedOverflow mk_onum dest_onum = - let - fun partial_dest_onum NONE = NONE - | partial_dest_onum (SOME n) = dest_onum n - - fun functions (numeral,[]) = - (case (natFromString numeral handle Overflow => NONE) of - NONE => NONE - | SOME n => dest_onum (ONum n)) - | functions ("suc",[x]) = dest_onum (suc (mk_onum x)) - | functions ("pre",[x]) = dest_onum (pre (mk_onum x)) - | functions ("~",[x]) = partial_dest_onum (neg (mk_onum x)) - | functions ("+",[x,y]) = - partial_dest_onum (add (mk_onum x) (mk_onum y)) - | functions ("-",[x,y]) = - partial_dest_onum (sub (mk_onum x) (mk_onum y)) - | functions ("*",[x,y]) = - partial_dest_onum (mult (mk_onum x) (mk_onum y)) - | functions ("exp",[x,y]) = - partial_dest_onum (exp (mk_onum x) (mk_onum y)) - | functions ("div",[x,y]) = - partial_dest_onum (odiv (mk_onum x) (mk_onum y)) - | functions ("mod",[x,y]) = - partial_dest_onum (omod (mk_onum x) (mk_onum y)) - | functions _ = NONE - - fun relations ("is_0",[x]) = SOME (mk_onum x = zero) - | relations ("<=",[x,y]) = le (mk_onum x) (mk_onum y) - | relations ("<",[x,y]) = lt (mk_onum x) (mk_onum y) - | relations (">=",[x,y]) = ge (mk_onum x) (mk_onum y) - | relations (">",[x,y]) = gt (mk_onum x) (mk_onum y) - | relations ("divides",[x,y]) = divides (mk_onum x) (mk_onum y) - | relations ("even",[x]) = even (mk_onum x) - | relations ("odd",[x]) = odd (mk_onum x) - | relations _ = NONE - in - {functions = functions, relations = relations} - end; -in - fun fixedOverflowNum {size = N} = - let - val oinf = N - 1 - - fun mk_onum x = if x = oinf then OInf else ONum x - - fun dest_onum ONeg = NONE - | dest_onum (ONum x) = SOME (if x < oinf then x else oinf) - | dest_onum OInf = SOME oinf - in - fixedOverflow mk_onum dest_onum - end; - - fun fixedOverflowInt {size = N} = - let - val oinf = N - 2 - val oneg = N - 1 - - fun mk_onum x = - if x = oneg then ONeg else if x = oinf then OInf else ONum x - - fun dest_onum ONeg = SOME oneg - | dest_onum (ONum x) = SOME (if x < oinf then x else oinf) - | dest_onum OInf = SOME oinf - in - fixedOverflow mk_onum dest_onum - end; -end; - -fun fixedSet {size = N} = - let - val M = - let - fun f 0 acc = acc - | f x acc = f (x div 2) (acc + 1) - in - f N 0 - end - - val univ = IntSet.fromList (interval 0 M) - - val mk_set = - let - fun f _ s 0 = s - | f k s x = - let - val s = if x mod 2 = 0 then s else IntSet.add s k - in - f (k + 1) s (x div 2) - end - in - f 0 IntSet.empty - end - - fun dest_set s = - let - fun f 0 x = x - | f k x = - let - val k = k - 1 - in - f k (if IntSet.member k s then 2 * x + 1 else 2 * x) - end - - val x = case IntSet.findr (K true) s of NONE => 0 | SOME k => f k 1 - in - if x < N then SOME x else NONE - end - - fun functions ("empty",[]) = dest_set IntSet.empty - | functions ("univ",[]) = dest_set univ - | functions ("union",[x,y]) = - dest_set (IntSet.union (mk_set x) (mk_set y)) - | functions ("intersect",[x,y]) = - dest_set (IntSet.intersect (mk_set x) (mk_set y)) - | functions ("compl",[x]) = - dest_set (IntSet.difference univ (mk_set x)) - | functions ("card",[x]) = SOME (IntSet.size (mk_set x)) - | functions _ = NONE - - fun relations ("in",[x,y]) = SOME (IntSet.member (x mod M) (mk_set y)) - | relations ("subset",[x,y]) = - SOME (IntSet.subset (mk_set x) (mk_set y)) - | relations _ = NONE - in - {functions = functions, relations = relations} +(* ------------------------------------------------------------------------- *) + +type fixedFunction = size -> element list -> element option; + +type fixedRelation = size -> element list -> bool option; + +datatype fixed = + Fixed of + {functions : fixedFunction NameArityMap.map, + relations : fixedRelation NameArityMap.map}; + +val uselessFixedFunction : fixedFunction = K (K NONE); + +val uselessFixedRelation : fixedRelation = K (K NONE); + +val emptyFunctions : fixedFunction NameArityMap.map = NameArityMap.new (); + +val emptyRelations : fixedRelation NameArityMap.map = NameArityMap.new (); + +fun fixed0 f sz elts = + case elts of + [] => f sz + | _ => raise Bug "Model.fixed0: wrong arity"; + +fun fixed1 f sz elts = + case elts of + [x] => f sz x + | _ => raise Bug "Model.fixed1: wrong arity"; + +fun fixed2 f sz elts = + case elts of + [x,y] => f sz x y + | _ => raise Bug "Model.fixed2: wrong arity"; + +val emptyFixed = + let + val fns = emptyFunctions + and rels = emptyRelations + in + Fixed + {functions = fns, + relations = rels} + end; + +fun peekFunctionFixed fix name_arity = + let + val Fixed {functions = fns, ...} = fix + in + NameArityMap.peek fns name_arity + end; + +fun peekRelationFixed fix name_arity = + let + val Fixed {relations = rels, ...} = fix + in + NameArityMap.peek rels name_arity + end; + +fun getFunctionFixed fix name_arity = + case peekFunctionFixed fix name_arity of + SOME f => f + | NONE => uselessFixedFunction; + +fun getRelationFixed fix name_arity = + case peekRelationFixed fix name_arity of + SOME rel => rel + | NONE => uselessFixedRelation; + +fun insertFunctionFixed fix name_arity_fn = + let + val Fixed {functions = fns, relations = rels} = fix + + val fns = NameArityMap.insert fns name_arity_fn + in + Fixed + {functions = fns, + relations = rels} + end; + +fun insertRelationFixed fix name_arity_rel = + let + val Fixed {functions = fns, relations = rels} = fix + + val rels = NameArityMap.insert rels name_arity_rel + in + Fixed + {functions = fns, + relations = rels} + end; + +local + fun union _ = raise Bug "Model.unionFixed: nameArity clash"; +in + fun unionFixed fix1 fix2 = + let + val Fixed {functions = fns1, relations = rels1} = fix1 + and Fixed {functions = fns2, relations = rels2} = fix2 + + val fns = NameArityMap.union union fns1 fns2 + + val rels = NameArityMap.union union rels1 rels2 + in + Fixed + {functions = fns, + relations = rels} + end; +end; + +val unionListFixed = + let + fun union (fix,acc) = unionFixed acc fix + in + List.foldl union emptyFixed + end; + +local + fun hasTypeFn _ elts = + case elts of + [x,_] => SOME x + | _ => raise Bug "Model.hasTypeFn: wrong arity"; + + fun eqRel _ elts = + case elts of + [x,y] => SOME (x = y) + | _ => raise Bug "Model.eqRel: wrong arity"; +in + val basicFixed = + let + val fns = NameArityMap.singleton (Term.hasTypeFunction,hasTypeFn) + + val rels = NameArityMap.singleton (Atom.eqRelation,eqRel) + in + Fixed + {functions = fns, + relations = rels} + end; +end; + +(* ------------------------------------------------------------------------- *) +(* Renaming fixed model parts. *) +(* ------------------------------------------------------------------------- *) + +type fixedMap = + {functionMap : Name.name NameArityMap.map, + relationMap : Name.name NameArityMap.map}; + +fun mapFixed fixMap fix = + let + val {functionMap = fnMap, relationMap = relMap} = fixMap + and Fixed {functions = fns, relations = rels} = fix + + val fns = NameArityMap.compose fnMap fns + + val rels = NameArityMap.compose relMap rels + in + Fixed + {functions = fns, + relations = rels} + end; + +local + fun mkEntry tag (na,n) = (tag,na,n); + + fun mkList tag m = map (mkEntry tag) (NameArityMap.toList m); + + fun ppEntry (tag,source_arity,target) = + Print.blockProgram Print.Inconsistent 2 + [Print.addString tag, + Print.addBreak 1, + NameArity.pp source_arity, + Print.addString " ->", + Print.addBreak 1, + Name.pp target]; +in + fun ppFixedMap fixMap = + let + val {functionMap = fnMap, relationMap = relMap} = fixMap + in + case mkList "function" fnMap @ mkList "relation" relMap of + [] => Print.skip + | entry :: entries => + Print.blockProgram Print.Consistent 0 + (ppEntry entry :: + map (Print.sequence Print.addNewline o ppEntry) entries) + end; +end; + +(* ------------------------------------------------------------------------- *) +(* Standard fixed model parts. *) +(* ------------------------------------------------------------------------- *) + +(* Projections *) + +val projectionMin = 1 +and projectionMax = 9; + +val projectionList = minMaxInterval projectionMin projectionMax; + +fun projectionName i = + let + val _ = projectionMin <= i orelse + raise Bug "Model.projectionName: less than projectionMin" + + val _ = i <= projectionMax orelse + raise Bug "Model.projectionName: greater than projectionMax" + in + Name.fromString ("project" ^ Int.toString i) + end; + +fun projectionFn i _ elts = SOME (List.nth (elts, i - 1)); + +fun arityProjectionFixed arity = + let + fun mkProj i = ((projectionName i, arity), projectionFn i) + + fun addProj i acc = + if i > arity then acc + else addProj (i + 1) (NameArityMap.insert acc (mkProj i)) + + val fns = addProj projectionMin emptyFunctions + + val rels = emptyRelations + in + Fixed + {functions = fns, + relations = rels} + end; + +val projectionFixed = + unionListFixed (map arityProjectionFixed projectionList); + +(* Arithmetic *) + +val numeralMin = ~100 +and numeralMax = 100; + +val numeralList = minMaxInterval numeralMin numeralMax; + +fun numeralName i = + let + val _ = numeralMin <= i orelse + raise Bug "Model.numeralName: less than numeralMin" + + val _ = i <= numeralMax orelse + raise Bug "Model.numeralName: greater than numeralMax" + + val s = if i < 0 then "negative" ^ Int.toString (~i) else Int.toString i + in + Name.fromString s + end; + +val addName = Name.fromString "+" +and divName = Name.fromString "div" +and dividesName = Name.fromString "divides" +and evenName = Name.fromString "even" +and expName = Name.fromString "exp" +and geName = Name.fromString ">=" +and gtName = Name.fromString ">" +and isZeroName = Name.fromString "isZero" +and leName = Name.fromString "<=" +and ltName = Name.fromString "<" +and modName = Name.fromString "mod" +and multName = Name.fromString "*" +and negName = Name.fromString "~" +and oddName = Name.fromString "odd" +and preName = Name.fromString "pre" +and subName = Name.fromString "-" +and sucName = Name.fromString "suc"; + +local + (* Support *) + + fun modN {size = N} x = x mod N; + + fun oneN sz = modN sz 1; + + fun multN sz (x,y) = modN sz (x * y); + + (* Functions *) + + fun numeralFn i sz = SOME (modN sz i); + + fun addFn sz x y = SOME (modN sz (x + y)); + + fun divFn {size = N} x y = + let + val y = if y = 0 then N else y + in + SOME (x div y) + end; + + fun expFn sz x y = SOME (exp (multN sz) x y (oneN sz)); + + fun modFn {size = N} x y = + let + val y = if y = 0 then N else y + in + SOME (x mod y) + end; + + fun multFn sz x y = SOME (multN sz (x,y)); + + fun negFn {size = N} x = SOME (if x = 0 then 0 else N - x); + + fun preFn {size = N} x = SOME (if x = 0 then N - 1 else x - 1); + + fun subFn {size = N} x y = SOME (if x < y then N + x - y else x - y); + + fun sucFn {size = N} x = SOME (if x = N - 1 then 0 else x + 1); + + (* Relations *) + + fun dividesRel _ x y = SOME (divides x y); + + fun evenRel _ x = SOME (x mod 2 = 0); + + fun geRel _ x y = SOME (x >= y); + + fun gtRel _ x y = SOME (x > y); + + fun isZeroRel _ x = SOME (x = 0); + + fun leRel _ x y = SOME (x <= y); + + fun ltRel _ x y = SOME (x < y); + + fun oddRel _ x = SOME (x mod 2 = 1); +in + val modularFixed = + let + val fns = + NameArityMap.fromList + (map (fn i => ((numeralName i,0), fixed0 (numeralFn i))) + numeralList @ + [((addName,2), fixed2 addFn), + ((divName,2), fixed2 divFn), + ((expName,2), fixed2 expFn), + ((modName,2), fixed2 modFn), + ((multName,2), fixed2 multFn), + ((negName,1), fixed1 negFn), + ((preName,1), fixed1 preFn), + ((subName,2), fixed2 subFn), + ((sucName,1), fixed1 sucFn)]) + + val rels = + NameArityMap.fromList + [((dividesName,2), fixed2 dividesRel), + ((evenName,1), fixed1 evenRel), + ((geName,2), fixed2 geRel), + ((gtName,2), fixed2 gtRel), + ((isZeroName,1), fixed1 isZeroRel), + ((leName,2), fixed2 leRel), + ((ltName,2), fixed2 ltRel), + ((oddName,1), fixed1 oddRel)] + in + Fixed + {functions = fns, + relations = rels} + end; +end; + +local + (* Support *) + + fun cutN {size = N} x = if x >= N then N - 1 else x; + + fun oneN sz = cutN sz 1; + + fun multN sz (x,y) = cutN sz (x * y); + + (* Functions *) + + fun numeralFn i sz = if i < 0 then NONE else SOME (cutN sz i); + + fun addFn sz x y = SOME (cutN sz (x + y)); + + fun divFn _ x y = if y = 0 then NONE else SOME (x div y); + + fun expFn sz x y = SOME (exp (multN sz) x y (oneN sz)); + + fun modFn {size = N} x y = + if y = 0 orelse x = N - 1 then NONE else SOME (x mod y); + + fun multFn sz x y = SOME (multN sz (x,y)); + + fun negFn _ x = if x = 0 then SOME 0 else NONE; + + fun preFn _ x = if x = 0 then NONE else SOME (x - 1); + + fun subFn {size = N} x y = + if y = 0 then SOME x + else if x = N - 1 orelse x < y then NONE + else SOME (x - y); + + fun sucFn sz x = SOME (cutN sz (x + 1)); + + (* Relations *) + + fun dividesRel {size = N} x y = + if x = 1 orelse y = 0 then SOME true + else if x = 0 then SOME false + else if y = N - 1 then NONE + else SOME (divides x y); + + fun evenRel {size = N} x = + if x = N - 1 then NONE else SOME (x mod 2 = 0); + + fun geRel {size = N} y x = + if x = N - 1 then if y = N - 1 then NONE else SOME false + else if y = N - 1 then SOME true else SOME (x <= y); + + fun gtRel {size = N} y x = + if x = N - 1 then if y = N - 1 then NONE else SOME false + else if y = N - 1 then SOME true else SOME (x < y); + + fun isZeroRel _ x = SOME (x = 0); + + fun leRel {size = N} x y = + if x = N - 1 then if y = N - 1 then NONE else SOME false + else if y = N - 1 then SOME true else SOME (x <= y); + + fun ltRel {size = N} x y = + if x = N - 1 then if y = N - 1 then NONE else SOME false + else if y = N - 1 then SOME true else SOME (x < y); + + fun oddRel {size = N} x = + if x = N - 1 then NONE else SOME (x mod 2 = 1); +in + val overflowFixed = + let + val fns = + NameArityMap.fromList + (map (fn i => ((numeralName i,0), fixed0 (numeralFn i))) + numeralList @ + [((addName,2), fixed2 addFn), + ((divName,2), fixed2 divFn), + ((expName,2), fixed2 expFn), + ((modName,2), fixed2 modFn), + ((multName,2), fixed2 multFn), + ((negName,1), fixed1 negFn), + ((preName,1), fixed1 preFn), + ((subName,2), fixed2 subFn), + ((sucName,1), fixed1 sucFn)]) + + val rels = + NameArityMap.fromList + [((dividesName,2), fixed2 dividesRel), + ((evenName,1), fixed1 evenRel), + ((geName,2), fixed2 geRel), + ((gtName,2), fixed2 gtRel), + ((isZeroName,1), fixed1 isZeroRel), + ((leName,2), fixed2 leRel), + ((ltName,2), fixed2 ltRel), + ((oddName,1), fixed1 oddRel)] + in + Fixed + {functions = fns, + relations = rels} + end; +end; + +(* Sets *) + +val cardName = Name.fromString "card" +and complementName = Name.fromString "complement" +and differenceName = Name.fromString "difference" +and emptyName = Name.fromString "empty" +and memberName = Name.fromString "member" +and insertName = Name.fromString "insert" +and intersectName = Name.fromString "intersect" +and singletonName = Name.fromString "singleton" +and subsetName = Name.fromString "subset" +and symmetricDifferenceName = Name.fromString "symmetricDifference" +and unionName = Name.fromString "union" +and universeName = Name.fromString "universe"; + +local + (* Support *) + + fun eltN {size = N} = + let + fun f 0 acc = acc + | f x acc = f (x div 2) (acc + 1) + in + f N ~1 + end; + + fun posN i = Word.<< (0w1, Word.fromInt i); + + fun univN sz = Word.- (posN (eltN sz), 0w1); + + fun setN sz x = Word.andb (Word.fromInt x, univN sz); + + (* Functions *) + + fun cardFn sz x = + let + fun f 0w0 acc = acc + | f s acc = + let + val acc = if Word.andb (s,0w1) = 0w0 then acc else acc + 1 + in + f (Word.>> (s,0w1)) acc + end + in + SOME (f (setN sz x) 0) + end; + + fun complementFn sz x = SOME (Word.toInt (Word.xorb (univN sz, setN sz x))); + + fun differenceFn sz x y = + let + val x = setN sz x + and y = setN sz y + in + SOME (Word.toInt (Word.andb (x, Word.notb y))) + end; + + fun emptyFn _ = SOME 0; + + fun insertFn sz x y = + let + val x = x mod eltN sz + and y = setN sz y + in + SOME (Word.toInt (Word.orb (posN x, y))) + end; + + fun intersectFn sz x y = + SOME (Word.toInt (Word.andb (setN sz x, setN sz y))); + + fun singletonFn sz x = + let + val x = x mod eltN sz + in + SOME (Word.toInt (posN x)) + end; + + fun symmetricDifferenceFn sz x y = + let + val x = setN sz x + and y = setN sz y + in + SOME (Word.toInt (Word.xorb (x,y))) + end; + + fun unionFn sz x y = + SOME (Word.toInt (Word.orb (setN sz x, setN sz y))); + + fun universeFn sz = SOME (Word.toInt (univN sz)); + + (* Relations *) + + fun memberRel sz x y = + let + val x = x mod eltN sz + and y = setN sz y + in + SOME (Word.andb (posN x, y) <> 0w0) + end; + + fun subsetRel sz x y = + let + val x = setN sz x + and y = setN sz y + in + SOME (Word.andb (x, Word.notb y) = 0w0) + end; +in + val setFixed = + let + val fns = + NameArityMap.fromList + [((cardName,1), fixed1 cardFn), + ((complementName,1), fixed1 complementFn), + ((differenceName,2), fixed2 differenceFn), + ((emptyName,0), fixed0 emptyFn), + ((insertName,2), fixed2 insertFn), + ((intersectName,2), fixed2 intersectFn), + ((singletonName,1), fixed1 singletonFn), + ((symmetricDifferenceName,2), fixed2 symmetricDifferenceFn), + ((unionName,2), fixed2 unionFn), + ((universeName,0), fixed0 universeFn)] + + val rels = + NameArityMap.fromList + [((memberName,2), fixed2 memberRel), + ((subsetName,2), fixed2 subsetRel)] + in + Fixed + {functions = fns, + relations = rels} + end; +end; + +(* Lists *) + +val appendName = Name.fromString "@" +and consName = Name.fromString "::" +and lengthName = Name.fromString "length" +and nilName = Name.fromString "nil" +and nullName = Name.fromString "null" +and tailName = Name.fromString "tail"; + +local + val baseFix = + let + val fix = unionFixed projectionFixed overflowFixed + + val sucFn = getFunctionFixed fix (sucName,1) + + fun suc2Fn sz _ x = sucFn sz [x] + in + insertFunctionFixed fix ((sucName,2), fixed2 suc2Fn) + end; + + val fixMap = + {functionMap = NameArityMap.fromList + [((appendName,2),addName), + ((consName,2),sucName), + ((lengthName,1), projectionName 1), + ((nilName,0), numeralName 0), + ((tailName,1),preName)], + relationMap = NameArityMap.fromList + [((nullName,1),isZeroName)]}; + +in + val listFixed = mapFixed fixMap baseFix; +end; + +(* ------------------------------------------------------------------------- *) +(* Valuations. *) +(* ------------------------------------------------------------------------- *) + +datatype valuation = Valuation of element NameMap.map; + +val emptyValuation = Valuation (NameMap.new ()); + +fun insertValuation (Valuation m) v_i = Valuation (NameMap.insert m v_i); + +fun peekValuation (Valuation m) v = NameMap.peek m v; + +fun constantValuation i = + let + fun add (v,V) = insertValuation V (v,i) + in + NameSet.foldl add emptyValuation + end; + +val zeroValuation = constantValuation zeroElement; + +fun getValuation V v = + case peekValuation V v of + SOME i => i + | NONE => raise Error "Model.getValuation: incomplete valuation"; + +fun randomValuation {size = N} vs = + let + fun f (v,V) = insertValuation V (v, Portable.randomInt N) + in + NameSet.foldl f emptyValuation vs + end; + +fun incrementValuation N vars = + let + fun inc vs V = + case vs of + [] => NONE + | v :: vs => + let + val (carry,i) = + case incrementElement N (getValuation V v) of + SOME i => (false,i) + | NONE => (true,zeroElement) + + val V = insertValuation V (v,i) + in + if carry then inc vs V else SOME V + end + in + inc (NameSet.toList vars) + end; + +fun foldValuation N vars f = + let + val inc = incrementValuation N vars + + fun fold V acc = + let + val acc = f (V,acc) + in + case inc V of + NONE => acc + | SOME V => fold V acc + end + + val zero = zeroValuation vars + in + fold zero + end; + +(* ------------------------------------------------------------------------- *) +(* A type of random finite mapping Z^n -> Z. *) +(* ------------------------------------------------------------------------- *) + +val UNKNOWN = ~1; + +datatype table = + ForgetfulTable + | ArrayTable of int Array.array; + +fun newTable N arity = + case elementListSpace {size = N} arity of + NONE => ForgetfulTable + | SOME space => ArrayTable (Array.array (space,UNKNOWN)); + +local + fun randomResult R = Portable.randomInt R; +in + fun lookupTable N R table elts = + case table of + ForgetfulTable => randomResult R + | ArrayTable a => + let + val i = elementListIndex {size = N} elts + + val r = Array.sub (a,i) + in + if r <> UNKNOWN then r + else + let + val r = randomResult R + + val () = Array.update (a,i,r) + in + r + end + end; +end; + +fun updateTable N table (elts,r) = + case table of + ForgetfulTable => () + | ArrayTable a => + let + val i = elementListIndex {size = N} elts + + val () = Array.update (a,i,r) + in + () + end; + +(* ------------------------------------------------------------------------- *) +(* A type of random finite mappings name * arity -> Z^arity -> Z. *) +(* ------------------------------------------------------------------------- *) + +datatype tables = + Tables of + {domainSize : int, + rangeSize : int, + tableMap : table NameArityMap.map Unsynchronized.ref}; + +fun newTables N R = + Tables + {domainSize = N, + rangeSize = R, + tableMap = Unsynchronized.ref (NameArityMap.new ())}; + +fun getTables tables n_a = + let + val Tables {domainSize = N, rangeSize = _, tableMap = tm} = tables + + val Unsynchronized.ref m = tm + in + case NameArityMap.peek m n_a of + SOME t => t + | NONE => + let + val (_,a) = n_a + + val t = newTable N a + + val m = NameArityMap.insert m (n_a,t) + + val () = tm := m + in + t + end + end; + +fun lookupTables tables (n,elts) = + let + val Tables {domainSize = N, rangeSize = R, ...} = tables + + val a = length elts + + val table = getTables tables (n,a) + in + lookupTable N R table elts + end; + +fun updateTables tables ((n,elts),r) = + let + val Tables {domainSize = N, ...} = tables + + val a = length elts + + val table = getTables tables (n,a) + in + updateTable N table (elts,r) end; (* ------------------------------------------------------------------------- *) @@ -11894,166 +16626,153 @@ datatype model = Model of {size : int, - fixed : fixedModel, - functions : (Term.functionName * int list, int) Map.map Unsynchronized.ref, - relations : (Atom.relationName * int list, bool) Map.map Unsynchronized.ref}; - -local - fun cmp ((n1,l1),(n2,l2)) = - case String.compare (n1,n2) of - LESS => LESS - | EQUAL => lexCompare Int.compare (l1,l2) - | GREATER => GREATER; -in - fun new {size = N, fixed} = + fixedFunctions : (element list -> element option) NameArityMap.map, + fixedRelations : (element list -> bool option) NameArityMap.map, + randomFunctions : tables, + randomRelations : tables}; + +fun new {size = N, fixed} = + let + val Fixed {functions = fns, relations = rels} = fixed + + val fixFns = NameArityMap.transform (fn f => f {size = N}) fns + and fixRels = NameArityMap.transform (fn r => r {size = N}) rels + + val rndFns = newTables N N + and rndRels = newTables N 2 + in Model {size = N, - fixed = fixed {size = N}, - functions = Unsynchronized.ref (Map.new cmp), - relations = Unsynchronized.ref (Map.new cmp)}; -end; - -fun size (Model {size = s, ...}) = s; - -(* ------------------------------------------------------------------------- *) -(* Valuations. *) -(* ------------------------------------------------------------------------- *) - -type valuation = int NameMap.map; - -val valuationEmpty : valuation = NameMap.new (); - -fun valuationRandom {size = N} vs = - let - fun f (v,V) = NameMap.insert V (v, Portable.randomInt N) - in - NameSet.foldl f valuationEmpty vs - end; - -fun valuationFold {size = N} vs f = - let - val vs = NameSet.toList vs - - fun inc [] _ = NONE - | inc (v :: l) V = - case NameMap.peek V v of - NONE => raise Bug "Model.valuationFold" - | SOME k => - let - val k = if k = N - 1 then 0 else k + 1 - val V = NameMap.insert V (v,k) - in - if k = 0 then inc l V else SOME V - end - - val zero = foldl (fn (v,V) => NameMap.insert V (v,0)) valuationEmpty vs - - fun fold V acc = - let - val acc = f (V,acc) - in - case inc vs V of NONE => acc | SOME V => fold V acc - end - in - fold zero - end; + fixedFunctions = fixFns, + fixedRelations = fixRels, + randomFunctions = rndFns, + randomRelations = rndRels} + end; + +fun size (Model {size = N, ...}) = N; + +fun peekFixedFunction M (n,elts) = + let + val Model {fixedFunctions = fixFns, ...} = M + in + case NameArityMap.peek fixFns (n, length elts) of + NONE => NONE + | SOME fixFn => fixFn elts + end; + +fun isFixedFunction M n_elts = Option.isSome (peekFixedFunction M n_elts); + +fun peekFixedRelation M (n,elts) = + let + val Model {fixedRelations = fixRels, ...} = M + in + case NameArityMap.peek fixRels (n, length elts) of + NONE => NONE + | SOME fixRel => fixRel elts + end; + +fun isFixedRelation M n_elts = Option.isSome (peekFixedRelation M n_elts); + +(* A default model *) + +val defaultSize = 8; + +val defaultFixed = + unionListFixed + [basicFixed, + projectionFixed, + modularFixed, + setFixed, + listFixed]; + +val default = {size = defaultSize, fixed = defaultFixed}; + +(* ------------------------------------------------------------------------- *) +(* Taking apart terms to interpret them. *) +(* ------------------------------------------------------------------------- *) + +fun destTerm tm = + case tm of + Term.Var _ => tm + | Term.Fn f_tms => + case Term.stripApp tm of + (_,[]) => tm + | (v as Term.Var _, tms) => Term.Fn (Term.appName, v :: tms) + | (Term.Fn (f,tms), tms') => Term.Fn (f, tms @ tms'); (* ------------------------------------------------------------------------- *) (* Interpreting terms and formulas in the model. *) (* ------------------------------------------------------------------------- *) +fun interpretFunction M n_elts = + case peekFixedFunction M n_elts of + SOME r => r + | NONE => + let + val Model {randomFunctions = rndFns, ...} = M + in + lookupTables rndFns n_elts + end; + +fun interpretRelation M n_elts = + case peekFixedRelation M n_elts of + SOME r => r + | NONE => + let + val Model {randomRelations = rndRels, ...} = M + in + intToBool (lookupTables rndRels n_elts) + end; + fun interpretTerm M V = let - val Model {size = N, fixed, functions, ...} = M - val {functions = fixed_functions, ...} = fixed - - fun interpret (Term.Var v) = - (case NameMap.peek V v of - NONE => raise Error "Model.interpretTerm: incomplete valuation" - | SOME i => i) - | interpret (tm as Term.Fn f_tms) = - let - val (f,tms) = - case Term.stripComb tm of - (_,[]) => f_tms - | (v as Term.Var _, tms) => (".", v :: tms) - | (Term.Fn (f,tms), tms') => (f, tms @ tms') - val elts = map interpret tms - val f_elts = (f,elts) - val Unsynchronized.ref funcs = functions - in - case Map.peek funcs f_elts of - SOME k => k - | NONE => - let - val k = - case fixed_functions f_elts of - SOME k => k - | NONE => Portable.randomInt N - - val () = functions := Map.insert funcs (f_elts,k) - in - k - end - end; + fun interpret tm = + case destTerm tm of + Term.Var v => getValuation V v + | Term.Fn (f,tms) => interpretFunction M (f, map interpret tms) in interpret end; fun interpretAtom M V (r,tms) = - let - val Model {fixed,relations,...} = M - val {relations = fixed_relations, ...} = fixed - - val elts = map (interpretTerm M V) tms - val r_elts = (r,elts) - val Unsynchronized.ref rels = relations - in - case Map.peek rels r_elts of - SOME b => b - | NONE => - let - val b = - case fixed_relations r_elts of - SOME b => b - | NONE => Portable.randomBool () - - val () = relations := Map.insert rels (r_elts,b) - in - b - end - end; + interpretRelation M (r, map (interpretTerm M V) tms); fun interpretFormula M = let - val Model {size = N, ...} = M - - fun interpret _ Formula.True = true - | interpret _ Formula.False = false - | interpret V (Formula.Atom atm) = interpretAtom M V atm - | interpret V (Formula.Not p) = not (interpret V p) - | interpret V (Formula.Or (p,q)) = interpret V p orelse interpret V q - | interpret V (Formula.And (p,q)) = interpret V p andalso interpret V q - | interpret V (Formula.Imp (p,q)) = - interpret V (Formula.Or (Formula.Not p, q)) - | interpret V (Formula.Iff (p,q)) = interpret V p = interpret V q - | interpret V (Formula.Forall (v,p)) = interpret' V v p N - | interpret V (Formula.Exists (v,p)) = - interpret V (Formula.Not (Formula.Forall (v, Formula.Not p))) - and interpret' _ _ _ 0 = true - | interpret' V v p i = + val N = size M + + fun interpret V fm = + case fm of + Formula.True => true + | Formula.False => false + | Formula.Atom atm => interpretAtom M V atm + | Formula.Not p => not (interpret V p) + | Formula.Or (p,q) => interpret V p orelse interpret V q + | Formula.And (p,q) => interpret V p andalso interpret V q + | Formula.Imp (p,q) => interpret V (Formula.Or (Formula.Not p, q)) + | Formula.Iff (p,q) => interpret V p = interpret V q + | Formula.Forall (v,p) => interpret' V p v N + | Formula.Exists (v,p) => + interpret V (Formula.Not (Formula.Forall (v, Formula.Not p))) + + and interpret' V fm v i = + i = 0 orelse let val i = i - 1 - val V' = NameMap.insert V (v,i) - in - interpret V' p andalso interpret' V v p i + val V' = insertValuation V (v,i) + in + interpret V' fm andalso interpret' V fm v i end in interpret end; -fun interpretLiteral M V (true,atm) = interpretAtom M V atm - | interpretLiteral M V (false,atm) = not (interpretAtom M V atm); +fun interpretLiteral M V (pol,atm) = + let + val b = interpretAtom M V atm + in + if pol then b else not b + end; fun interpretClause M V cl = LiteralSet.exists (interpretLiteral M V) cl; @@ -12062,33 +16781,199 @@ (* Note: if it's cheaper, a systematic check will be performed instead. *) (* ------------------------------------------------------------------------- *) -local - fun checkGen freeVars interpret {maxChecks} M x = - let - val Model {size = N, ...} = M - - fun score (V,{T,F}) = - if interpret M V x then {T = T + 1, F = F} else {T = T, F = F + 1} - - val vs = freeVars x - - fun randomCheck acc = score (valuationRandom {size = N} vs, acc) - - val small = - intExp N (NameSet.size vs) <= maxChecks handle Overflow => false - in - if small then valuationFold {size = N} vs score {T = 0, F = 0} - else funpow maxChecks randomCheck {T = 0, F = 0} - end; -in - val checkAtom = checkGen Atom.freeVars interpretAtom; - - val checkFormula = checkGen Formula.freeVars interpretFormula; - - val checkLiteral = checkGen Literal.freeVars interpretLiteral; - - val checkClause = checkGen LiteralSet.freeVars interpretClause; -end; +fun check interpret {maxChecks} M fv x = + let + val N = size M + + fun score (V,{T,F}) = + if interpret M V x then {T = T + 1, F = F} else {T = T, F = F + 1} + + fun randomCheck acc = score (randomValuation {size = N} fv, acc) + + val maxChecks = + case maxChecks of + NONE => maxChecks + | SOME m => + case expInt N (NameSet.size fv) of + SOME n => if n <= m then NONE else maxChecks + | NONE => maxChecks + in + case maxChecks of + SOME m => funpow m randomCheck {T = 0, F = 0} + | NONE => foldValuation {size = N} fv score {T = 0, F = 0} + end; + +fun checkAtom maxChecks M atm = + check interpretAtom maxChecks M (Atom.freeVars atm) atm; + +fun checkFormula maxChecks M fm = + check interpretFormula maxChecks M (Formula.freeVars fm) fm; + +fun checkLiteral maxChecks M lit = + check interpretLiteral maxChecks M (Literal.freeVars lit) lit; + +fun checkClause maxChecks M cl = + check interpretClause maxChecks M (LiteralSet.freeVars cl) cl; + +(* ------------------------------------------------------------------------- *) +(* Updating the model. *) +(* ------------------------------------------------------------------------- *) + +fun updateFunction M func_elts_elt = + let + val Model {randomFunctions = rndFns, ...} = M + + val () = updateTables rndFns func_elts_elt + in + () + end; + +fun updateRelation M (rel_elts,pol) = + let + val Model {randomRelations = rndRels, ...} = M + + val () = updateTables rndRels (rel_elts, boolToInt pol) + in + () + end; + +(* ------------------------------------------------------------------------- *) +(* A type of terms with interpretations embedded in the subterms. *) +(* ------------------------------------------------------------------------- *) + +datatype modelTerm = + ModelVar + | ModelFn of Term.functionName * modelTerm list * int list; + +fun modelTerm M V = + let + fun modelTm tm = + case destTerm tm of + Term.Var v => (ModelVar, getValuation V v) + | Term.Fn (f,tms) => + let + val (tms,xs) = unzip (map modelTm tms) + in + (ModelFn (f,tms,xs), interpretFunction M (f,xs)) + end + in + modelTm + end; + +(* ------------------------------------------------------------------------- *) +(* Perturbing the model. *) +(* ------------------------------------------------------------------------- *) + +datatype perturbation = + FunctionPerturbation of (Term.functionName * element list) * element + | RelationPerturbation of (Atom.relationName * element list) * bool; + +fun perturb M pert = + case pert of + FunctionPerturbation func_elts_elt => updateFunction M func_elts_elt + | RelationPerturbation rel_elts_pol => updateRelation M rel_elts_pol; + +local + fun pertTerm _ [] _ acc = acc + | pertTerm M target tm acc = + case tm of + ModelVar => acc + | ModelFn (func,tms,xs) => + let + fun onTarget ys = mem (interpretFunction M (func,ys)) target + + val func_xs = (func,xs) + + val acc = + if isFixedFunction M func_xs then acc + else + let + fun add (y,acc) = FunctionPerturbation (func_xs,y) :: acc + in + foldl add acc target + end + in + pertTerms M onTarget tms xs acc + end + + and pertTerms M onTarget = + let + val N = size M + + fun filterElements pred = + let + fun filt 0 acc = acc + | filt i acc = + let + val i = i - 1 + val acc = if pred i then i :: acc else acc + in + filt i acc + end + in + filt N [] + end + + fun pert _ [] [] acc = acc + | pert ys (tm :: tms) (x :: xs) acc = + let + fun pred y = + y <> x andalso onTarget (List.revAppend (ys, y :: xs)) + + val target = filterElements pred + + val acc = pertTerm M target tm acc + in + pert (x :: ys) tms xs acc + end + | pert _ _ _ _ = raise Bug "Model.pertTerms.pert" + in + pert [] + end; + + fun pertAtom M V target (rel,tms) acc = + let + fun onTarget ys = interpretRelation M (rel,ys) = target + + val (tms,xs) = unzip (map (modelTerm M V) tms) + + val rel_xs = (rel,xs) + + val acc = + if isFixedRelation M rel_xs then acc + else RelationPerturbation (rel_xs,target) :: acc + in + pertTerms M onTarget tms xs acc + end; + + fun pertLiteral M V ((pol,atm),acc) = pertAtom M V pol atm acc; + + fun pertClause M V cl acc = LiteralSet.foldl (pertLiteral M V) acc cl; + + fun pickPerturb M perts = + if null perts then () + else perturb M (List.nth (perts, Portable.randomInt (length perts))); +in + fun perturbTerm M V (tm,target) = + pickPerturb M (pertTerm M target (fst (modelTerm M V tm)) []); + + fun perturbAtom M V (atm,target) = + pickPerturb M (pertAtom M V target atm []); + + fun perturbLiteral M V lit = pickPerturb M (pertLiteral M V (lit,[])); + + fun perturbClause M V cl = pickPerturb M (pertClause M V cl []); +end; + +(* ------------------------------------------------------------------------- *) +(* Pretty printing. *) +(* ------------------------------------------------------------------------- *) + +fun pp M = + Print.program + [Print.addString "Model{", + Print.ppInt (size M), + Print.addString "}"]; end end; @@ -12096,8 +16981,8 @@ (**** Original file: Problem.sig ****) (* ========================================================================= *) -(* SOME SAMPLE PROBLEMS TO TEST PROOF PROCEDURES *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* CNF PROBLEMS *) +(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Problem = @@ -12107,16 +16992,22 @@ (* Problems. *) (* ------------------------------------------------------------------------- *) -type problem = Metis.Thm.clause list +type problem = + {axioms : Metis.Thm.clause list, + conjecture : Metis.Thm.clause list} val size : problem -> {clauses : int, literals : int, symbols : int, typedSymbols : int} -val fromGoal : Metis.Formula.formula -> problem list - -val toClauses : problem -> Metis.Formula.formula +val freeVars : problem -> Metis.NameSet.set + +val toClauses : problem -> Metis.Thm.clause list + +val toFormula : problem -> Metis.Formula.formula + +val toGoal : problem -> Metis.Formula.formula val toString : problem -> string @@ -12157,7 +17048,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; @@ -12165,8 +17056,8 @@ val foldr = List.foldr; (* ========================================================================= *) -(* SOME SAMPLE PROBLEMS TO TEST PROOF PROCEDURES *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* CNF PROBLEMS *) +(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Problem :> Problem = @@ -12178,83 +17069,56 @@ (* Problems. *) (* ------------------------------------------------------------------------- *) -type problem = Thm.clause list; - -fun size cls = - {clauses = length cls, - literals = foldl (fn (cl,n) => n + LiteralSet.size cl) 0 cls, - symbols = foldl (fn (cl,n) => n + LiteralSet.symbols cl) 0 cls, - typedSymbols = foldl (fn (cl,n) => n + LiteralSet.typedSymbols cl) 0 cls}; - -fun checkFormula {models,checks} exp fm = - let - fun check 0 = true - | check n = - let - val N = 3 + Portable.randomInt 3 - val M = Model.new {size = N, fixed = Model.fixedPure} - val {T,F} = Model.checkFormula {maxChecks = checks} M fm - in - (if exp then F = 0 else T = 0) andalso check (n - 1) - end - in - check models - end; - -val checkGoal = checkFormula {models = 10, checks = 100} true; - -val checkClauses = checkFormula {models = 10, checks = 100} false; - -fun fromGoal goal = - let - fun fromLiterals fms = LiteralSet.fromList (map Literal.fromFormula fms) - - fun fromClause fm = fromLiterals (Formula.stripDisj fm) - - fun fromCnf fm = map fromClause (Formula.stripConj fm) - -(*DEBUG - val () = if checkGoal goal then () - else raise Error "goal failed the finite model test" -*) - - val refute = Formula.Not (Formula.generalize goal) - -(*TRACE2 - val () = Parser.ppTrace Formula.pp "Problem.fromGoal: refute" refute -*) - - val cnfs = Normalize.cnf refute - -(* - val () = - let - fun check fm = - let -(*TRACE2 - val () = Parser.ppTrace Formula.pp "Problem.fromGoal: cnf" fm -*) - in - if checkClauses fm then () - else raise Error "cnf failed the finite model test" - end - in - app check cnfs - end -*) - in - map fromCnf cnfs - end; - -fun toClauses cls = - let - fun formulize cl = - Formula.listMkDisj (LiteralSet.transform Literal.toFormula cl) - in - Formula.listMkConj (map formulize cls) - end; - -fun toString cls = Formula.toString (toClauses cls); +type problem = + {axioms : Thm.clause list, + conjecture : Thm.clause list}; + +fun toClauses {axioms,conjecture} = axioms @ conjecture; + +fun size prob = + let + fun lits (cl,n) = n + LiteralSet.size cl + + fun syms (cl,n) = n + LiteralSet.symbols cl + + fun typedSyms (cl,n) = n + LiteralSet.typedSymbols cl + + val cls = toClauses prob + in + {clauses = length cls, + literals = foldl lits 0 cls, + symbols = foldl syms 0 cls, + typedSymbols = foldl typedSyms 0 cls} + end; + +fun freeVars {axioms,conjecture} = + NameSet.union + (LiteralSet.freeVarsList axioms) + (LiteralSet.freeVarsList conjecture); + +local + fun clauseToFormula cl = + Formula.listMkDisj (LiteralSet.transform Literal.toFormula cl); +in + fun toFormula prob = + Formula.listMkConj (map clauseToFormula (toClauses prob)); + + fun toGoal {axioms,conjecture} = + let + val clToFm = Formula.generalize o clauseToFormula + val clsToFm = Formula.listMkConj o map clToFm + + val fm = Formula.False + val fm = + if null conjecture then fm + else Formula.Imp (clsToFm conjecture, fm) + val fm = Formula.Imp (clsToFm axioms, fm) + in + fm + end; +end; + +fun toString prob = Formula.toString (toFormula prob); (* ------------------------------------------------------------------------- *) (* Categorizing problems. *) @@ -12283,8 +17147,10 @@ equality : equality, horn : horn}; -fun categorize cls = - let +fun categorize prob = + let + val cls = toClauses prob + val rels = let fun f (cl,set) = NameAritySet.union set (LiteralSet.relations cl) @@ -12355,7 +17221,7 @@ (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature TermNet = @@ -12387,7 +17253,7 @@ val toString : 'a termNet -> string -val pp : 'a Metis.Parser.pp -> 'a termNet Metis.Parser.pp +val pp : 'a Metis.Print.pp -> 'a termNet Metis.Print.pp (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) @@ -12408,7 +17274,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; @@ -12417,7 +17283,7 @@ (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure TermNet :> TermNet = @@ -12426,29 +17292,65 @@ open Useful; (* ------------------------------------------------------------------------- *) -(* Quotient terms *) -(* ------------------------------------------------------------------------- *) - -datatype qterm = VAR | FN of NameArity.nameArity * qterm list; - -fun termToQterm (Term.Var _) = VAR - | termToQterm (Term.Fn (f,l)) = FN ((f, length l), map termToQterm l); +(* Anonymous variables. *) +(* ------------------------------------------------------------------------- *) + +val anonymousName = Name.fromString "_"; +val anonymousVar = Term.Var anonymousName; + +(* ------------------------------------------------------------------------- *) +(* Quotient terms. *) +(* ------------------------------------------------------------------------- *) + +datatype qterm = + Var + | Fn of NameArity.nameArity * qterm list; + +local + fun cmp [] = EQUAL + | cmp (q1_q2 :: qs) = + if Portable.pointerEqual q1_q2 then cmp qs + else + case q1_q2 of + (Var,Var) => EQUAL + | (Var, Fn _) => LESS + | (Fn _, Var) => GREATER + | (Fn f1, Fn f2) => fnCmp f1 f2 qs + + and fnCmp (n1,q1) (n2,q2) qs = + case NameArity.compare (n1,n2) of + LESS => LESS + | EQUAL => cmp (zip q1 q2 @ qs) + | GREATER => GREATER; +in + fun compareQterm q1_q2 = cmp [q1_q2]; + + fun compareFnQterm (f1,f2) = fnCmp f1 f2 []; +end; + +fun equalQterm q1 q2 = compareQterm (q1,q2) = EQUAL; + +fun equalFnQterm f1 f2 = compareFnQterm (f1,f2) = EQUAL; + +fun termToQterm (Term.Var _) = Var + | termToQterm (Term.Fn (f,l)) = Fn ((f, length l), map termToQterm l); local fun qm [] = true - | qm ((VAR,_) :: rest) = qm rest - | qm ((FN _, VAR) :: _) = false - | qm ((FN (f,a), FN (g,b)) :: rest) = f = g andalso qm (zip a b @ rest); + | qm ((Var,_) :: rest) = qm rest + | qm ((Fn _, Var) :: _) = false + | qm ((Fn (f,a), Fn (g,b)) :: rest) = + NameArity.equal f g andalso qm (zip a b @ rest); in fun matchQtermQterm qtm qtm' = qm [(qtm,qtm')]; end; local fun qm [] = true - | qm ((VAR,_) :: rest) = qm rest - | qm ((FN _, Term.Var _) :: _) = false - | qm ((FN ((f,n),a), Term.Fn (g,b)) :: rest) = - f = g andalso n = length b andalso qm (zip a b @ rest); + | qm ((Var,_) :: rest) = qm rest + | qm ((Fn _, Term.Var _) :: _) = false + | qm ((Fn ((f,n),a), Term.Fn (g,b)) :: rest) = + Name.equal f g andalso n = length b andalso qm (zip a b @ rest); in fun matchQtermTerm qtm tm = qm [(qtm,tm)]; end; @@ -12458,26 +17360,27 @@ | qn qsub ((Term.Var v, qtm) :: rest) = (case NameMap.peek qsub v of NONE => qn (NameMap.insert qsub (v,qtm)) rest - | SOME qtm' => if qtm = qtm' then qn qsub rest else NONE) - | qn _ ((Term.Fn _, VAR) :: _) = NONE - | qn qsub ((Term.Fn (f,a), FN ((g,n),b)) :: rest) = - if f = g andalso length a = n then qn qsub (zip a b @ rest) else NONE; + | SOME qtm' => if equalQterm qtm qtm' then qn qsub rest else NONE) + | qn _ ((Term.Fn _, Var) :: _) = NONE + | qn qsub ((Term.Fn (f,a), Fn ((g,n),b)) :: rest) = + if Name.equal f g andalso length a = n then qn qsub (zip a b @ rest) + else NONE; in fun matchTermQterm qsub tm qtm = qn qsub [(tm,qtm)]; end; local - fun qv VAR x = x - | qv x VAR = x - | qv (FN (f,a)) (FN (g,b)) = - let - val _ = f = g orelse raise Error "TermNet.qv" - in - FN (f, zipwith qv a b) + fun qv Var x = x + | qv x Var = x + | qv (Fn (f,a)) (Fn (g,b)) = + let + val _ = NameArity.equal f g orelse raise Error "TermNet.qv" + in + Fn (f, zipWith qv a b) end; fun qu qsub [] = qsub - | qu qsub ((VAR, _) :: rest) = qu qsub rest + | qu qsub ((Var, _) :: rest) = qu qsub rest | qu qsub ((qtm, Term.Var v) :: rest) = let val qtm = @@ -12485,8 +17388,8 @@ in qu (NameMap.insert qsub (v,qtm)) rest end - | qu qsub ((FN ((f,n),a), Term.Fn (g,b)) :: rest) = - if f = g andalso n = length b then qu qsub (zip a b @ rest) + | qu qsub ((Fn ((f,n),a), Term.Fn (g,b)) :: rest) = + if Name.equal f g andalso n = length b then qu qsub (zip a b @ rest) else raise Error "TermNet.qu"; in fun unifyQtermQterm qtm qtm' = total (qv qtm) qtm'; @@ -12495,10 +17398,10 @@ end; local - fun qtermToTerm VAR = Term.Var "" - | qtermToTerm (FN ((f,_),l)) = Term.Fn (f, map qtermToTerm l); -in - val ppQterm = Parser.ppMap qtermToTerm Term.pp; + fun qtermToTerm Var = anonymousVar + | qtermToTerm (Fn ((f,_),l)) = Term.Fn (f, map qtermToTerm l); +in + val ppQterm = Print.ppMap qtermToTerm Term.pp; end; (* ------------------------------------------------------------------------- *) @@ -12508,22 +17411,22 @@ type parameters = {fifo : bool}; datatype 'a net = - RESULT of 'a list - | SINGLE of qterm * 'a net - | MULTIPLE of 'a net option * 'a net NameArityMap.map; - -datatype 'a termNet = NET of parameters * int * (int * (int * 'a) net) option; + Result of 'a list + | Single of qterm * 'a net + | Multiple of 'a net option * 'a net NameArityMap.map; + +datatype 'a termNet = Net of parameters * int * (int * (int * 'a) net) option; (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) -fun new parm = NET (parm,0,NONE); - -local - fun computeSize (RESULT l) = length l - | computeSize (SINGLE (_,n)) = computeSize n - | computeSize (MULTIPLE (vs,fs)) = +fun new parm = Net (parm,0,NONE); + +local + fun computeSize (Result l) = length l + | computeSize (Single (_,n)) = computeSize n + | computeSize (Multiple (vs,fs)) = NameArityMap.foldl (fn (_,n,acc) => acc + computeSize n) (case vs of SOME n => computeSize n | NONE => 0) @@ -12533,38 +17436,38 @@ | netSize (SOME n) = SOME (computeSize n, n); end; -fun size (NET (_,_,NONE)) = 0 - | size (NET (_, _, SOME (i,_))) = i; +fun size (Net (_,_,NONE)) = 0 + | size (Net (_, _, SOME (i,_))) = i; fun null net = size net = 0; -fun singles qtms a = foldr SINGLE a qtms; +fun singles qtms a = foldr Single a qtms; local fun pre NONE = (0,NONE) | pre (SOME (i,n)) = (i, SOME n); - fun add (RESULT l) [] (RESULT l') = RESULT (l @ l') - | add a (input1 as qtm :: qtms) (SINGLE (qtm',n)) = - if qtm = qtm' then SINGLE (qtm, add a qtms n) - else add a input1 (add n [qtm'] (MULTIPLE (NONE, NameArityMap.new ()))) - | add a (VAR :: qtms) (MULTIPLE (vs,fs)) = - MULTIPLE (SOME (oadd a qtms vs), fs) - | add a (FN (f,l) :: qtms) (MULTIPLE (vs,fs)) = + fun add (Result l) [] (Result l') = Result (l @ l') + | add a (input1 as qtm :: qtms) (Single (qtm',n)) = + if equalQterm qtm qtm' then Single (qtm, add a qtms n) + else add a input1 (add n [qtm'] (Multiple (NONE, NameArityMap.new ()))) + | add a (Var :: qtms) (Multiple (vs,fs)) = + Multiple (SOME (oadd a qtms vs), fs) + | add a (Fn (f,l) :: qtms) (Multiple (vs,fs)) = let val n = NameArityMap.peek fs f in - MULTIPLE (vs, NameArityMap.insert fs (f, oadd a (l @ qtms) n)) + Multiple (vs, NameArityMap.insert fs (f, oadd a (l @ qtms) n)) end | add _ _ _ = raise Bug "TermNet.insert: Match" and oadd a qtms NONE = singles qtms a | oadd a qtms (SOME n) = add a qtms n; - fun ins a qtm (i,n) = SOME (i + 1, oadd (RESULT [a]) [qtm] n); -in - fun insert (NET (p,k,n)) (tm,a) = - NET (p, k + 1, ins (k,a) (termToQterm tm) (pre n)) + fun ins a qtm (i,n) = SOME (i + 1, oadd (Result [a]) [qtm] n); +in + fun insert (Net (p,k,n)) (tm,a) = + Net (p, k + 1, ins (k,a) (termToQterm tm) (pre n)) handle Error _ => raise Bug "TermNet.insert: should never fail"; end; @@ -12572,26 +17475,26 @@ fun filter pred = let - fun filt (RESULT l) = + fun filt (Result l) = (case List.filter (fn (_,a) => pred a) l of [] => NONE - | l => SOME (RESULT l)) - | filt (SINGLE (qtm,n)) = + | l => SOME (Result l)) + | filt (Single (qtm,n)) = (case filt n of NONE => NONE - | SOME n => SOME (SINGLE (qtm,n))) - | filt (MULTIPLE (vs,fs)) = + | SOME n => SOME (Single (qtm,n))) + | filt (Multiple (vs,fs)) = let val vs = Option.mapPartial filt vs val fs = NameArityMap.mapPartial (fn (_,n) => filt n) fs in if not (Option.isSome vs) andalso NameArityMap.null fs then NONE - else SOME (MULTIPLE (vs,fs)) - end - in - fn net as NET (_,_,NONE) => net - | NET (p, k, SOME (_,n)) => NET (p, k, netSize (filt n)) + else SOME (Multiple (vs,fs)) + end + in + fn net as Net (_,_,NONE) => net + | Net (p, k, SOME (_,n)) => Net (p, k, netSize (filt n)) end handle Error _ => raise Bug "TermNet.filter: should never fail"; @@ -12606,7 +17509,7 @@ let val (a,qtms) = revDivide qtms n in - addQterm (FN (f,a)) (ks,fs,qtms) + addQterm (Fn (f,a)) (ks,fs,qtms) end | norm stack = stack @@ -12620,7 +17523,7 @@ and addFn (f as (_,n)) (ks,fs,qtms) = norm (n :: ks, f :: fs, qtms); in val stackEmpty = ([],[],[]); - + val stackAddQterm = addQterm; val stackAddFn = addFn; @@ -12633,16 +17536,16 @@ fun fold _ acc [] = acc | fold inc acc ((0,stack,net) :: rest) = fold inc (inc (stackValue stack, net, acc)) rest - | fold inc acc ((n, stack, SINGLE (qtm,net)) :: rest) = + | fold inc acc ((n, stack, Single (qtm,net)) :: rest) = fold inc acc ((n - 1, stackAddQterm qtm stack, net) :: rest) - | fold inc acc ((n, stack, MULTIPLE (v,fns)) :: rest) = + | fold inc acc ((n, stack, Multiple (v,fns)) :: rest) = let val n = n - 1 val rest = case v of NONE => rest - | SOME net => (n, stackAddQterm VAR stack, net) :: rest + | SOME net => (n, stackAddQterm Var stack, net) :: rest fun getFns (f as (_,k), net, x) = (k + n, stackAddFn f stack, net) :: x @@ -12657,11 +17560,11 @@ fun foldEqualTerms pat inc acc = let fun fold ([],net) = inc (pat,net,acc) - | fold (pat :: pats, SINGLE (qtm,net)) = - if pat = qtm then fold (pats,net) else acc - | fold (VAR :: pats, MULTIPLE (v,_)) = + | fold (pat :: pats, Single (qtm,net)) = + if equalQterm pat qtm then fold (pats,net) else acc + | fold (Var :: pats, Multiple (v,_)) = (case v of NONE => acc | SOME net => fold (pats,net)) - | fold (FN (f,a) :: pats, MULTIPLE (_,fns)) = + | fold (Fn (f,a) :: pats, Multiple (_,fns)) = (case NameArityMap.peek fns f of NONE => acc | SOME net => fold (a @ pats, net)) @@ -12674,20 +17577,20 @@ fun fold _ acc [] = acc | fold inc acc (([],stack,net) :: rest) = fold inc (inc (stackValue stack, net, acc)) rest - | fold inc acc ((VAR :: pats, stack, net) :: rest) = + | fold inc acc ((Var :: pats, stack, net) :: rest) = let fun harvest (qtm,n,l) = (pats, stackAddQterm qtm stack, n) :: l in fold inc acc (foldTerms harvest rest net) end - | fold inc acc ((pat :: pats, stack, SINGLE (qtm,net)) :: rest) = + | fold inc acc ((pat :: pats, stack, Single (qtm,net)) :: rest) = (case unifyQtermQterm pat qtm of NONE => fold inc acc rest | SOME qtm => fold inc acc ((pats, stackAddQterm qtm stack, net) :: rest)) | fold inc acc - (((pat as FN (f,a)) :: pats, stack, MULTIPLE (v,fns)) :: rest) = + (((pat as Fn (f,a)) :: pats, stack, Multiple (v,fns)) :: rest) = let val rest = case v of @@ -12724,10 +17627,10 @@ local fun mat acc [] = acc - | mat acc ((RESULT l, []) :: rest) = mat (l @ acc) rest - | mat acc ((SINGLE (qtm,n), tm :: tms) :: rest) = + | mat acc ((Result l, []) :: rest) = mat (l @ acc) rest + | mat acc ((Single (qtm,n), tm :: tms) :: rest) = mat acc (if matchQtermTerm qtm tm then (n,tms) :: rest else rest) - | mat acc ((MULTIPLE (vs,fs), tm :: tms) :: rest) = + | mat acc ((Multiple (vs,fs), tm :: tms) :: rest) = let val rest = case vs of NONE => rest | SOME n => (n,tms) :: rest @@ -12743,8 +17646,8 @@ end | mat _ _ = raise Bug "TermNet.match: Match"; in - fun match (NET (_,_,NONE)) _ = [] - | match (NET (p, _, SOME (_,n))) tm = + fun match (Net (_,_,NONE)) _ = [] + | match (Net (p, _, SOME (_,n))) tm = finally p (mat [] [(n,[tm])]) handle Error _ => raise Bug "TermNet.match: should never fail"; end; @@ -12756,16 +17659,16 @@ fun seenInc qsub tms (_,net,rest) = (qsub,net,tms) :: rest; fun mat acc [] = acc - | mat acc ((_, RESULT l, []) :: rest) = mat (l @ acc) rest - | mat acc ((qsub, SINGLE (qtm,net), tm :: tms) :: rest) = + | mat acc ((_, Result l, []) :: rest) = mat (l @ acc) rest + | mat acc ((qsub, Single (qtm,net), tm :: tms) :: rest) = (case matchTermQterm qsub tm qtm of NONE => mat acc rest | SOME qsub => mat acc ((qsub,net,tms) :: rest)) - | mat acc ((qsub, net as MULTIPLE _, Term.Var v :: tms) :: rest) = + | mat acc ((qsub, net as Multiple _, Term.Var v :: tms) :: rest) = (case NameMap.peek qsub v of NONE => mat acc (foldTerms (unseenInc qsub v tms) rest net) | SOME qtm => mat acc (foldEqualTerms qtm (seenInc qsub tms) rest net)) - | mat acc ((qsub, MULTIPLE (_,fns), Term.Fn (f,a) :: tms) :: rest) = + | mat acc ((qsub, Multiple (_,fns), Term.Fn (f,a) :: tms) :: rest) = let val rest = case NameArityMap.peek fns (f, length a) of @@ -12776,8 +17679,8 @@ end | mat _ _ = raise Bug "TermNet.matched.mat"; in - fun matched (NET (_,_,NONE)) _ = [] - | matched (NET (parm, _, SOME (_,net))) tm = + fun matched (Net (_,_,NONE)) _ = [] + | matched (Net (parm, _, SOME (_,net))) tm = finally parm (mat [] [(NameMap.new (), net, [tm])]) handle Error _ => raise Bug "TermNet.matched: should never fail"; end; @@ -12787,16 +17690,16 @@ (NameMap.insert qsub (v,qtm), net, tms) :: rest; fun mat acc [] = acc - | mat acc ((_, RESULT l, []) :: rest) = mat (l @ acc) rest - | mat acc ((qsub, SINGLE (qtm,net), tm :: tms) :: rest) = + | mat acc ((_, Result l, []) :: rest) = mat (l @ acc) rest + | mat acc ((qsub, Single (qtm,net), tm :: tms) :: rest) = (case unifyQtermTerm qsub qtm tm of NONE => mat acc rest | SOME qsub => mat acc ((qsub,net,tms) :: rest)) - | mat acc ((qsub, net as MULTIPLE _, Term.Var v :: tms) :: rest) = + | mat acc ((qsub, net as Multiple _, Term.Var v :: tms) :: rest) = (case NameMap.peek qsub v of NONE => mat acc (foldTerms (inc qsub v tms) rest net) | SOME qtm => mat acc (foldUnifiableTerms qtm (inc qsub v tms) rest net)) - | mat acc ((qsub, MULTIPLE (v,fns), Term.Fn (f,a) :: tms) :: rest) = + | mat acc ((qsub, Multiple (v,fns), Term.Fn (f,a) :: tms) :: rest) = let val rest = case v of NONE => rest | SOME net => (qsub,net,tms) :: rest @@ -12809,8 +17712,8 @@ end | mat _ _ = raise Bug "TermNet.unify.mat"; in - fun unify (NET (_,_,NONE)) _ = [] - | unify (NET (parm, _, SOME (_,net))) tm = + fun unify (Net (_,_,NONE)) _ = [] + | unify (Net (parm, _, SOME (_,net))) tm = finally parm (mat [] [(NameMap.new (), net, [tm])]) handle Error _ => raise Bug "TermNet.unify: should never fail"; end; @@ -12820,16 +17723,16 @@ (* ------------------------------------------------------------------------- *) local - fun inc (qtm, RESULT l, acc) = + fun inc (qtm, Result l, acc) = foldl (fn ((n,a),acc) => (n,(qtm,a)) :: acc) acc l | inc _ = raise Bug "TermNet.pp.inc"; - - fun toList (NET (_,_,NONE)) = [] - | toList (NET (parm, _, SOME (_,net))) = + + fun toList (Net (_,_,NONE)) = [] + | toList (Net (parm, _, SOME (_,net))) = finally parm (foldTerms inc [] net); in fun pp ppA = - Parser.ppMap toList (Parser.ppList (Parser.ppBinop " |->" ppQterm ppA)); + Print.ppMap toList (Print.ppList (Print.ppOp2 " |->" ppQterm ppA)); end; end @@ -12839,7 +17742,7 @@ (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature AtomNet = @@ -12869,7 +17772,7 @@ val toString : 'a atomNet -> string -val pp : 'a Metis.Parser.pp -> 'a atomNet Metis.Parser.pp +val pp : 'a Metis.Print.pp -> 'a atomNet Metis.Print.pp (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) @@ -12890,7 +17793,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; @@ -12899,7 +17802,7 @@ (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure AtomNet :> AtomNet = @@ -12962,7 +17865,7 @@ (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature LiteralNet = @@ -12994,7 +17897,7 @@ val toString : 'a literalNet -> string -val pp : 'a Metis.Parser.pp -> 'a literalNet Metis.Parser.pp +val pp : 'a Metis.Print.pp -> 'a literalNet Metis.Print.pp (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) @@ -13015,7 +17918,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; @@ -13024,7 +17927,7 @@ (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure LiteralNet :> LiteralNet = @@ -13072,9 +17975,9 @@ fun toString net = "LiteralNet[" ^ Int.toString (size net) ^ "]"; fun pp ppA = - Parser.ppMap + Print.ppMap (fn {positive,negative} => (positive,negative)) - (Parser.ppBinop " + NEGATIVE" (AtomNet.pp ppA) (AtomNet.pp ppA)); + (Print.ppOp2 " + NEGATIVE" (AtomNet.pp ppA) (AtomNet.pp ppA)); (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) @@ -13102,7 +18005,7 @@ (* ========================================================================= *) (* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Subsume = @@ -13122,7 +18025,7 @@ val filter : ('a -> bool) -> 'a subsume -> 'a subsume -val pp : 'a subsume Metis.Parser.pp +val pp : 'a subsume Metis.Print.pp val toString : 'a subsume -> string @@ -13156,7 +18059,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; @@ -13165,7 +18068,7 @@ (* ========================================================================= *) (* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Subsume :> Subsume = @@ -13324,7 +18227,7 @@ fun toString subsume = "Subsume{" ^ Int.toString (size subsume) ^ "}"; -fun pp p = Parser.ppMap toString Parser.ppString p; +fun pp subsume = Print.ppMap toString Print.ppString subsume; (* ------------------------------------------------------------------------- *) (* Subsumption checking. *) @@ -13439,19 +18342,19 @@ genSubsumes pred subsume (SOME (LiteralSet.size cl)) cl; end; -(*TRACE4 +(*MetisTrace4 val subsumes = fn pred => fn subsume => fn cl => let val ppCl = LiteralSet.pp val ppSub = Subst.pp - val () = Parser.ppTrace ppCl "Subsume.subsumes: cl" cl + val () = Print.trace ppCl "Subsume.subsumes: cl" cl val result = subsumes pred subsume cl val () = case result of NONE => trace "Subsume.subsumes: not subsumed\n" | SOME (cl,sub,_) => - (Parser.ppTrace ppCl "Subsume.subsumes: subsuming cl" cl; - Parser.ppTrace ppSub "Subsume.subsumes: subsuming sub" sub) + (Print.trace ppCl "Subsume.subsumes: subsuming cl" cl; + Print.trace ppSub "Subsume.subsumes: subsuming sub" sub) in result end; @@ -13460,14 +18363,14 @@ let val ppCl = LiteralSet.pp val ppSub = Subst.pp - val () = Parser.ppTrace ppCl "Subsume.strictlySubsumes: cl" cl + val () = Print.trace ppCl "Subsume.strictlySubsumes: cl" cl val result = strictlySubsumes pred subsume cl val () = case result of NONE => trace "Subsume.subsumes: not subsumed\n" | SOME (cl,sub,_) => - (Parser.ppTrace ppCl "Subsume.subsumes: subsuming cl" cl; - Parser.ppTrace ppSub "Subsume.subsumes: subsuming sub" sub) + (Print.trace ppCl "Subsume.subsumes: subsuming cl" cl; + Print.trace ppSub "Subsume.subsumes: subsuming sub" sub) in result end; @@ -13503,7 +18406,7 @@ (* ========================================================================= *) (* THE KNUTH-BENDIX TERM ORDERING *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature KnuthBendixOrder = @@ -13528,7 +18431,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; @@ -13537,7 +18440,7 @@ (* ========================================================================= *) (* KNUTH-BENDIX TERM ORDERING CONSTRAINTS *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure KnuthBendixOrder :> KnuthBendixOrder = @@ -13549,10 +18452,12 @@ (* Helper functions. *) (* ------------------------------------------------------------------------- *) -fun firstNotEqual f l = - case List.find op<> l of +fun notEqualTerm (x,y) = not (Term.equal x y); + +fun firstNotEqualTerm f l = + case List.find notEqualTerm l of SOME (x,y) => f x y - | NONE => raise Bug "firstNotEqual"; + | NONE => raise Bug "firstNotEqualTerm"; (* ------------------------------------------------------------------------- *) (* The weight of all constants must be at least 1, and there must be at most *) @@ -13573,7 +18478,7 @@ fn ((f1,n1),(f2,n2)) => case Int.compare (n1,n2) of LESS => LESS - | EQUAL => String.compare (f1,f2) + | EQUAL => Name.compare (f1,f2) | GREATER => GREATER; (* The default order *) @@ -13599,7 +18504,7 @@ fun weightNeg (Weight (m,c)) = Weight (NameMap.transform ~ m, ~c); local - fun add (n1,n2) = + fun add ((_,n1),(_,n2)) = let val n = n1 + n2 in @@ -13612,15 +18517,6 @@ fun weightSubtract w1 w2 = weightAdd w1 (weightNeg w2); -fun weightMult 0 _ = weightZero - | weightMult 1 w = w - | weightMult k (Weight (m,c)) = - let - fun mult n = k * n - in - Weight (NameMap.transform mult m, k * c) - end; - fun weightTerm weight = let fun wt m c [] = Weight (m,c) @@ -13636,80 +18532,41 @@ fn tm => wt weightEmpty ~1 [tm] end; -fun weightIsVar v (Weight (m,c)) = - c = 0 andalso NameMap.size m = 1 andalso NameMap.peek m v = SOME 1; - -fun weightConst (Weight (_,c)) = c; - -fun weightVars (Weight (m,_)) = - NameMap.foldl (fn (v,_,s) => NameSet.add s v) NameSet.empty m; - -val weightsVars = - List.foldl (fn (w,s) => NameSet.union (weightVars w) s) NameSet.empty; - -fun weightVarList w = NameSet.toList (weightVars w); - -fun weightNumVars (Weight (m,_)) = NameMap.size m; - -fun weightNumVarsWithPositiveCoeff (Weight (m,_)) = - NameMap.foldl (fn (_,n,z) => if n > 0 then z + 1 else z) 0 m; - -fun weightCoeff var (Weight (m,_)) = Option.getOpt (NameMap.peek m var, 0); - -fun weightCoeffs varList w = map (fn var => weightCoeff var w) varList; - -fun weightCoeffSum (Weight (m,_)) = NameMap.foldl (fn (_,n,z) => n + z) 0 m; - fun weightLowerBound (w as Weight (m,c)) = if NameMap.exists (fn (_,n) => n < 0) m then NONE else SOME c; -fun weightNoLowerBound w = not (Option.isSome (weightLowerBound w)); - -fun weightAlwaysPositive w = - case weightLowerBound w of NONE => false | SOME n => n > 0; - -fun weightUpperBound (w as Weight (m,c)) = - if NameMap.exists (fn (_,n) => n > 0) m then NONE else SOME c; - -fun weightNoUpperBound w = not (Option.isSome (weightUpperBound w)); - -fun weightAlwaysNegative w = - case weightUpperBound w of NONE => false | SOME n => n < 0; - -fun weightGcd (w as Weight (m,c)) = - NameMap.foldl (fn (_,i,k) => gcd i k) (Int.abs c) m; - -fun ppWeightList pp = - let - fun coeffToString n = - if n < 0 then "~" ^ coeffToString (~n) - else if n = 1 then "" - else Int.toString n - - fun pp_tm pp ("",n) = Parser.ppInt pp n - | pp_tm pp (v,n) = Parser.ppString pp (coeffToString n ^ v) - in - fn [] => Parser.ppInt pp 0 - | tms => Parser.ppSequence " +" pp_tm pp tms - end; - -fun ppWeight pp (Weight (m,c)) = +(*MetisDebug +val ppWeightList = + let + fun ppCoeff n = + if n < 0 then Print.sequence (Print.addString "~") (ppCoeff (~n)) + else if n = 1 then Print.skip + else Print.ppInt n + + fun pp_tm (NONE,n) = Print.ppInt n + | pp_tm (SOME v, n) = Print.sequence (ppCoeff n) (Name.pp v) + in + fn [] => Print.ppInt 0 + | tms => Print.ppOpList " +" pp_tm tms + end; + +fun ppWeight (Weight (m,c)) = let val l = NameMap.toList m - val l = if c = 0 then l else l @ [("",c)] - in - ppWeightList pp l - end; - -val weightToString = Parser.toString ppWeight; + val l = map (fn (v,n) => (SOME v, n)) l + val l = if c = 0 then l else l @ [(NONE,c)] + in + ppWeightList l + end; + +val weightToString = Print.toString ppWeight; +*) (* ------------------------------------------------------------------------- *) (* The Knuth-Bendix term order. *) (* ------------------------------------------------------------------------- *) -datatype kboOrder = Less | Equal | Greater | SignOf of weight; - -fun kboOrder {weight,precedence} = +fun compare {weight,precedence} = let fun weightDifference tm1 tm2 = let @@ -13736,7 +18593,7 @@ and precedenceLess (Term.Fn (f1,a1)) (Term.Fn (f2,a2)) = (case precedence ((f1, length a1), (f2, length a2)) of LESS => true - | EQUAL => firstNotEqual weightLess (zip a1 a2) + | EQUAL => firstNotEqualTerm weightLess (zip a1 a2) | GREATER => false) | precedenceLess _ _ = false @@ -13747,39 +18604,33 @@ val w = weightDifference tm1 tm2 in if weightIsZero w then precedenceCmp tm1 tm2 - else if weightDiffLess w tm1 tm2 then Less - else if weightDiffGreater w tm1 tm2 then Greater - else SignOf w + else if weightDiffLess w tm1 tm2 then SOME LESS + else if weightDiffGreater w tm1 tm2 then SOME GREATER + else NONE end and precedenceCmp (Term.Fn (f1,a1)) (Term.Fn (f2,a2)) = (case precedence ((f1, length a1), (f2, length a2)) of - LESS => Less - | EQUAL => firstNotEqual weightCmp (zip a1 a2) - | GREATER => Greater) + LESS => SOME LESS + | EQUAL => firstNotEqualTerm weightCmp (zip a1 a2) + | GREATER => SOME GREATER) | precedenceCmp _ _ = raise Bug "kboOrder.precendenceCmp" in - fn (tm1,tm2) => if tm1 = tm2 then Equal else weightCmp tm1 tm2 - end; - -fun compare kbo tm1_tm2 = - case kboOrder kbo tm1_tm2 of - Less => SOME LESS - | Equal => SOME EQUAL - | Greater => SOME GREATER - | SignOf _ => NONE; - -(*TRACE7 + fn (tm1,tm2) => + if Term.equal tm1 tm2 then SOME EQUAL else weightCmp tm1 tm2 + end; + +(*MetisTrace7 val compare = fn kbo => fn (tm1,tm2) => let - val () = Parser.ppTrace Term.pp "KnuthBendixOrder.compare: tm1" tm1 - val () = Parser.ppTrace Term.pp "KnuthBendixOrder.compare: tm2" tm2 + val () = Print.trace Term.pp "KnuthBendixOrder.compare: tm1" tm1 + val () = Print.trace Term.pp "KnuthBendixOrder.compare: tm2" tm2 val result = compare kbo (tm1,tm2) val () = case result of NONE => trace "KnuthBendixOrder.compare: result = Incomparable\n" | SOME x => - Parser.ppTrace Parser.ppOrder "KnuthBendixOrder.compare: result" x + Print.trace Print.ppOrder "KnuthBendixOrder.compare: result" x in result end; @@ -13792,18 +18643,30 @@ (* ========================================================================= *) (* ORDERED REWRITING FOR FIRST ORDER TERMS *) -(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Rewrite = sig (* ------------------------------------------------------------------------- *) -(* A type of rewrite systems. *) +(* Orientations of equations. *) (* ------------------------------------------------------------------------- *) datatype orient = LeftToRight | RightToLeft +val toStringOrient : orient -> string + +val ppOrient : orient Metis.Print.pp + +val toStringOrientOption : orient option -> string + +val ppOrientOption : orient option Metis.Print.pp + +(* ------------------------------------------------------------------------- *) +(* A type of rewrite systems. *) +(* ------------------------------------------------------------------------- *) + type reductionOrder = Metis.Term.term * Metis.Term.term -> order option type equationId = int @@ -13826,7 +18689,7 @@ val toString : rewrite -> string -val pp : rewrite Metis.Parser.pp +val pp : rewrite Metis.Print.pp (* ------------------------------------------------------------------------- *) (* Add equations into the system. *) @@ -13882,7 +18745,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; @@ -13891,7 +18754,7 @@ (* ========================================================================= *) (* ORDERED REWRITING FOR FIRST ORDER TERMS *) -(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Rewrite :> Rewrite = @@ -13900,11 +18763,29 @@ open Useful; (* ------------------------------------------------------------------------- *) -(* A type of rewrite systems. *) +(* Orientations of equations. *) (* ------------------------------------------------------------------------- *) datatype orient = LeftToRight | RightToLeft; +fun toStringOrient ort = + case ort of + LeftToRight => "-->" + | RightToLeft => "<--"; + +val ppOrient = Print.ppMap toStringOrient Print.ppString; + +fun toStringOrientOption orto = + case orto of + SOME ort => toStringOrient ort + | NONE => "<->"; + +val ppOrientOption = Print.ppMap toStringOrientOption Print.ppString; + +(* ------------------------------------------------------------------------- *) +(* A type of rewrite systems. *) +(* ------------------------------------------------------------------------- *) + type reductionOrder = Term.term * Term.term -> order option; type equationId = int; @@ -13950,72 +18831,63 @@ fun equations (Rewrite {known,...}) = IntMap.foldr (fn (_,(eqn,_),eqns) => eqn :: eqns) [] known; -val pp = Parser.ppMap equations (Parser.ppList Rule.ppEquation); - -(*DEBUG -local - fun orientOptionToString ort = - case ort of - SOME LeftToRight => "-->" - | SOME RightToLeft => "<--" - | NONE => "<->"; - - open Parser; - - fun ppEq p ((x_y,_),ort) = - ppBinop (" " ^ orientOptionToString ort) Term.pp Term.pp p x_y; - - fun ppField f ppA p a = - (beginBlock p Inconsistent 2; - addString p (f ^ " ="); - addBreak p (1,0); - ppA p a; - endBlock p); +val pp = Print.ppMap equations (Print.ppList Rule.ppEquation); + +(*MetisTrace1 +local + fun ppEq ((x_y,_),ort) = + Print.ppOp2 (" " ^ toStringOrientOption ort) Term.pp Term.pp x_y; + + fun ppField f ppA a = + Print.blockProgram Print.Inconsistent 2 + [Print.addString (f ^ " ="), + Print.addBreak 1, + ppA a]; val ppKnown = - ppField "known" (ppMap IntMap.toList (ppList (ppPair ppInt ppEq))); + ppField "known" + (Print.ppMap IntMap.toList + (Print.ppList (Print.ppPair Print.ppInt ppEq))); val ppRedexes = - ppField - "redexes" - (TermNet.pp - (ppPair ppInt (ppMap (orientOptionToString o SOME) ppString))); + ppField "redexes" + (TermNet.pp (Print.ppPair Print.ppInt ppOrient)); val ppSubterms = - ppField - "subterms" + ppField "subterms" (TermNet.pp - (ppMap + (Print.ppMap (fn (i,l,p) => (i, (if l then 0 else 1) :: p)) - (ppPair ppInt Term.ppPath))); - - val ppWaiting = ppField "waiting" (ppMap (IntSet.toList) (ppList ppInt)); -in - fun pp p (Rewrite {known,redexes,subterms,waiting,...}) = - (beginBlock p Inconsistent 2; - addString p "Rewrite"; - addBreak p (1,0); - beginBlock p Inconsistent 1; - addString p "{"; - ppKnown p known; -(*TRACE5 - addString p ","; - addBreak p (1,0); - ppRedexes p redexes; - addString p ","; - addBreak p (1,0); - ppSubterms p subterms; - addString p ","; - addBreak p (1,0); - ppWaiting p waiting; -*) - endBlock p; - addString p "}"; - endBlock p); -end; -*) - -val toString = Parser.toString pp; + (Print.ppPair Print.ppInt Term.ppPath))); + + val ppWaiting = + ppField "waiting" + (Print.ppMap (IntSet.toList) (Print.ppList Print.ppInt)); +in + fun pp (Rewrite {known,redexes,subterms,waiting,...}) = + Print.blockProgram Print.Inconsistent 2 + [Print.addString "Rewrite", + Print.addBreak 1, + Print.blockProgram Print.Inconsistent 1 + [Print.addString "{", + ppKnown known, +(*MetisTrace5 + Print.addString ",", + Print.addBreak 1, + ppRedexes redexes, + Print.addString ",", + Print.addBreak 1, + ppSubterms subterms, + Print.addString ",", + Print.addBreak 1, + ppWaiting waiting, +*) + Print.skip], + Print.addString "}"] +end; +*) + +val toString = Print.toString pp; (* ------------------------------------------------------------------------- *) (* Debug functions. *) @@ -14028,7 +18900,7 @@ NONE => false | SOME sub => order (tm, Subst.subst (Subst.normalize sub) r) = SOME GREATER - + fun knownRed tm (eqnId,(eqn,ort)) = eqnId <> id andalso ((ort <> SOME RightToLeft andalso eqnRed eqn tm) orelse @@ -14082,8 +18954,8 @@ Rewrite {order = order, known = known, redexes = redexes, subterms = subterms, waiting = waiting} -(*TRACE5 - val () = Parser.ppTrace pp "Rewrite.add: result" rw +(*MetisTrace5 + val () = Print.trace pp "Rewrite.add: result" rw *) in rw @@ -14147,17 +19019,18 @@ NONE => raise Error "incomparable" | SOME LESS => let - val sub = Subst.fromList [("x",l),("y",r)] - val th = Thm.subst sub Rule.symmetry - in - fn tm => if tm = r then (l,th) else raise Error "mkNeqConv: RL" + val th = Rule.symmetryRule l r + in + fn tm => + if Term.equal tm r then (l,th) else raise Error "mkNeqConv: RL" end | SOME EQUAL => raise Error "irreflexive" | SOME GREATER => let val th = Thm.assume lit in - fn tm => if tm = l then (r,th) else raise Error "mkNeqConv: LR" + fn tm => + if Term.equal tm l then (r,th) else raise Error "mkNeqConv: LR" end end; @@ -14212,14 +19085,14 @@ | SOME lit => (false,lit) val (lit',litTh) = literule lit in - if lit = lit' then eqn + if Literal.equal lit lit' then eqn else (Literal.destEq lit', if strongEqn then th else if not (Thm.negateMember lit litTh) then litTh else Thm.resolve lit th litTh) end -(*DEBUG +(*MetisDebug handle Error err => raise Error ("Rewrite.rewriteIdEqn':\n" ^ err); *) @@ -14232,7 +19105,7 @@ val neq = neqConvsDelete neq lit val (lit',litTh) = mk_literule neq lit in - if lit = lit' then acc + if Literal.equal lit lit' then acc else let val th = Thm.resolve lit th litTh @@ -14268,15 +19141,15 @@ fun rewriteIdRule' order known redexes id th = rewriteIdLiteralsRule' order known redexes id (Thm.clause th) th; -(*DEBUG +(*MetisDebug val rewriteIdRule' = fn order => fn known => fn redexes => fn id => fn th => let -(*TRACE6 - val () = Parser.ppTrace Thm.pp "Rewrite.rewriteIdRule': th" th +(*MetisTrace6 + val () = Print.trace Thm.pp "Rewrite.rewriteIdRule': th" th *) val result = rewriteIdRule' order known redexes id th -(*TRACE6 - val () = Parser.ppTrace Thm.pp "Rewrite.rewriteIdRule': result" result +(*MetisTrace6 + val () = Print.trace Thm.pp "Rewrite.rewriteIdRule': result" result *) val _ = not (thmReducible order known id result) orelse raise Bug "rewriteIdRule: should be normalized" @@ -14322,8 +19195,8 @@ end; fun sameRedexes NONE _ _ = false - | sameRedexes (SOME LeftToRight) (l0,_) (l,_) = l0 = l - | sameRedexes (SOME RightToLeft) (_,r0) (_,r) = r0 = r; + | sameRedexes (SOME LeftToRight) (l0,_) (l,_) = Term.equal l0 l + | sameRedexes (SOME RightToLeft) (_,r0) (_,r) = Term.equal r0 r; fun redexResidues NONE (l,r) = [(l,r,false),(r,l,false)] | redexResidues (SOME LeftToRight) (l,r) = [(l,r,true)] @@ -14346,13 +19219,13 @@ else raise Error "order" end end - + fun addRed lr ((id',left,path),todo) = if id <> id' andalso not (IntSet.member id' todo) andalso can (checkValidRewr lr id' left) path then IntSet.add todo id' else todo - + fun findRed (lr as (l,_,_), todo) = List.foldl (addRed lr) todo (TermNet.matched subterms l) in @@ -14364,7 +19237,13 @@ val (eq0,_) = eqn0 val Rewrite {order,known,redexes,subterms,waiting} = rw val eqn as (eq,_) = rewriteIdEqn' order known redexes id eqn0 - val identical = eq = eq0 + val identical = + let + val (l0,r0) = eq0 + and (l,r) = eq + in + Term.equal l l0 andalso Term.equal r r0 + end val same_redexes = identical orelse sameRedexes ort0 eq0 eq val rpl = if same_redexes then rpl else IntSet.add rpl id val spl = if new orelse identical then spl else IntSet.add spl id @@ -14434,7 +19313,7 @@ case IntMap.peek known id of NONE => reds | SOME eqn_ort => addRedexes id eqn_ort reds - + val redexes = TermNet.filter filt redexes val redexes = IntSet.foldl addReds redexes rpl in @@ -14451,7 +19330,7 @@ case IntMap.peek known id of NONE => subtms | SOME (eqn,_) => addSubterms id eqn subtms - + val subterms = TermNet.filter filt subterms val subterms = IntSet.foldl addSubtms subterms spl in @@ -14460,18 +19339,21 @@ in fun rebuild rpl spl rw = let -(*TRACE5 - val ppPl = Parser.ppMap IntSet.toList (Parser.ppList Parser.ppInt) - val () = Parser.ppTrace ppPl "Rewrite.rebuild: rpl" rpl - val () = Parser.ppTrace ppPl "Rewrite.rebuild: spl" spl +(*MetisTrace5 + val ppPl = Print.ppMap IntSet.toList (Print.ppList Print.ppInt) + val () = Print.trace ppPl "Rewrite.rebuild: rpl" rpl + val () = Print.trace ppPl "Rewrite.rebuild: spl" spl *) val Rewrite {order,known,redexes,subterms,waiting} = rw val redexes = cleanRedexes known redexes rpl val subterms = cleanSubterms known subterms spl in Rewrite - {order = order, known = known, redexes = redexes, - subterms = subterms, waiting = waiting} + {order = order, + known = known, + redexes = redexes, + subterms = subterms, + waiting = waiting} end; end; @@ -14499,17 +19381,17 @@ if isReduced rw then (rw,[]) else reduceAcc (IntSet.empty,IntSet.empty,IntSet.empty,rw,IntSet.empty); -(*DEBUG +(*MetisDebug val reduce' = fn rw => let -(*TRACE4 - val () = Parser.ppTrace pp "Rewrite.reduce': rw" rw +(*MetisTrace4 + val () = Print.trace pp "Rewrite.reduce': rw" rw *) val Rewrite {known,order,...} = rw val result as (Rewrite {known = known', ...}, _) = reduce' rw -(*TRACE4 - val ppResult = Parser.ppPair pp (Parser.ppList Parser.ppInt) - val () = Parser.ppTrace ppResult "Rewrite.reduce': result" result +(*MetisTrace4 + val ppResult = Print.ppPair pp (Print.ppList Print.ppInt) + val () = Print.trace ppResult "Rewrite.reduce': result" result *) val ths = map (fn (id,((_,th),_)) => (id,th)) (IntMap.toList known') val _ = @@ -14547,7 +19429,7 @@ (* ========================================================================= *) (* A STORE FOR UNIT THEOREMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Units = @@ -14571,7 +19453,7 @@ val toString : units -> string -val pp : units Metis.Parser.pp +val pp : units Metis.Print.pp (* ------------------------------------------------------------------------- *) (* Add units into the store. *) @@ -14599,7 +19481,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; @@ -14608,7 +19490,7 @@ (* ========================================================================= *) (* A STORE FOR UNIT THEOREMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Units :> Units = @@ -14634,7 +19516,7 @@ fun toString units = "U{" ^ Int.toString (size units) ^ "}"; -val pp = Parser.ppMap toString Parser.ppString; +val pp = Print.ppMap toString Print.ppString; (* ------------------------------------------------------------------------- *) (* Add units into the store. *) @@ -14718,7 +19600,7 @@ (* ========================================================================= *) (* CLAUSE = ID + THEOREM *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Clause = @@ -14818,7 +19700,9 @@ val showId : bool Unsynchronized.ref -val pp : clause Metis.Parser.pp +val pp : clause Metis.Print.pp + +val toString : clause -> string end @@ -14826,7 +19710,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; @@ -14835,7 +19719,7 @@ (* ========================================================================= *) (* CLAUSE = ID + THEOREM *) -(* Copyright (c) 2002-2004 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Clause :> Clause = @@ -14851,8 +19735,7 @@ let val r = Unsynchronized.ref 0 in - fn () => CRITICAL (fn () => - case r of Unsynchronized.ref n => let val () = r := n + 1 in n end) + fn () => case r of Unsynchronized.ref n => let val () = r := n + 1 in n end end; (* ------------------------------------------------------------------------- *) @@ -14882,19 +19765,21 @@ val showId = Unsynchronized.ref false; local - val ppIdThm = Parser.ppPair Parser.ppInt Thm.pp; -in - fun pp p (Clause {id,thm,...}) = - if !showId then ppIdThm p (id,thm) else Thm.pp p thm; -end; + val ppIdThm = Print.ppPair Print.ppInt Thm.pp; +in + fun pp (Clause {id,thm,...}) = + if !showId then ppIdThm (id,thm) else Thm.pp thm; +end; + +fun toString cl = Print.toString pp cl; (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) -val default : parameters = +val default : parameters = {ordering = KnuthBendixOrder.default, - orderLiterals = UnsignedLiteralOrder, (*LCP: changed from PositiveLiteralOrder*) + orderLiterals = UnsignedLiteralOrder (* PositiveLiteralOrder *) (* MODIFIED by Jasmin Blanchette *), orderTerms = true}; fun mk info = Clause info @@ -14977,13 +19862,13 @@ LiteralSet.foldr addLit LiteralSet.empty litSet end; -(*TRACE6 +(*MetisTrace6 val largestLiterals = fn cl => let val ppResult = LiteralSet.pp - val () = Parser.ppTrace pp "Clause.largestLiterals: cl" cl + val () = Print.trace pp "Clause.largestLiterals: cl" cl val result = largestLiterals cl - val () = Parser.ppTrace ppResult "Clause.largestLiterals: result" result + val () = Print.trace ppResult "Clause.largestLiterals: result" result in result end; @@ -15053,10 +19938,10 @@ Rewrite.rewriteIdRule rewr cmp id th end -(*TRACE4 - val () = Parser.ppTrace Rewrite.pp "Clause.rewrite: rewr" rewr - val () = Parser.ppTrace Parser.ppInt "Clause.rewrite: id" id - val () = Parser.ppTrace pp "Clause.rewrite: cl" cl +(*MetisTrace4 + val () = Print.trace Rewrite.pp "Clause.rewrite: rewr" rewr + val () = Print.trace Print.ppInt "Clause.rewrite: id" id + val () = Print.trace pp "Clause.rewrite: cl" cl *) val thm = @@ -15066,13 +19951,13 @@ val result = Clause {parameters = parameters, id = id, thm = thm} -(*TRACE4 - val () = Parser.ppTrace pp "Clause.rewrite: result" result +(*MetisTrace4 + val () = Print.trace pp "Clause.rewrite: result" result *) in result end -(*DEBUG +(*MetisDebug handle Error err => raise Error ("Clause.rewrite:\n" ^ err); *) @@ -15089,12 +19974,12 @@ map apply (Rule.factor' lits) end; -(*TRACE5 +(*MetisTrace5 val factor = fn cl => let - val () = Parser.ppTrace pp "Clause.factor: cl" cl + val () = Print.trace pp "Clause.factor: cl" cl val result = factor cl - val () = Parser.ppTrace (Parser.ppList pp) "Clause.factor: result" result + val () = Print.trace (Print.ppList pp) "Clause.factor: result" result in result end; @@ -15102,51 +19987,55 @@ fun resolve (cl1,lit1) (cl2,lit2) = let -(*TRACE5 - val () = Parser.ppTrace pp "Clause.resolve: cl1" cl1 - val () = Parser.ppTrace Literal.pp "Clause.resolve: lit1" lit1 - val () = Parser.ppTrace pp "Clause.resolve: cl2" cl2 - val () = Parser.ppTrace Literal.pp "Clause.resolve: lit2" lit2 +(*MetisTrace5 + val () = Print.trace pp "Clause.resolve: cl1" cl1 + val () = Print.trace Literal.pp "Clause.resolve: lit1" lit1 + val () = Print.trace pp "Clause.resolve: cl2" cl2 + val () = Print.trace Literal.pp "Clause.resolve: lit2" lit2 *) val Clause {parameters, thm = th1, ...} = cl1 and Clause {thm = th2, ...} = cl2 val sub = Literal.unify Subst.empty lit1 (Literal.negate lit2) -(*TRACE5 - val () = Parser.ppTrace Subst.pp "Clause.resolve: sub" sub +(*MetisTrace5 + val () = Print.trace Subst.pp "Clause.resolve: sub" sub *) val lit1 = Literal.subst sub lit1 val lit2 = Literal.negate lit1 val th1 = Thm.subst sub th1 and th2 = Thm.subst sub th2 val _ = isLargerLiteral parameters (Thm.clause th1) lit1 orelse -(*TRACE5 +(*MetisTrace5 (trace "Clause.resolve: th1 violates ordering\n"; false) orelse *) raise Error "resolve: clause1: ordering constraints" val _ = isLargerLiteral parameters (Thm.clause th2) lit2 orelse -(*TRACE5 +(*MetisTrace5 (trace "Clause.resolve: th2 violates ordering\n"; false) orelse *) raise Error "resolve: clause2: ordering constraints" val th = Thm.resolve lit1 th1 th2 -(*TRACE5 - val () = Parser.ppTrace Thm.pp "Clause.resolve: th" th +(*MetisTrace5 + val () = Print.trace Thm.pp "Clause.resolve: th" th *) val cl = Clause {parameters = parameters, id = newId (), thm = th} -(*TRACE5 - val () = Parser.ppTrace pp "Clause.resolve: cl" cl +(*MetisTrace5 + val () = Print.trace pp "Clause.resolve: cl" cl *) in cl end; -fun paramodulate (cl1,lit1,ort,tm1) (cl2,lit2,path,tm2) = - let -(*TRACE5 - val () = Parser.ppTrace pp "Clause.paramodulate: cl1" cl1 - val () = Parser.ppTrace Literal.pp "Clause.paramodulate: lit1" lit1 - val () = Parser.ppTrace pp "Clause.paramodulate: cl2" cl2 - val () = Parser.ppTrace Literal.pp "Clause.paramodulate: lit2" lit2 +fun paramodulate (cl1,lit1,ort1,tm1) (cl2,lit2,path2,tm2) = + let +(*MetisTrace5 + val () = Print.trace pp "Clause.paramodulate: cl1" cl1 + val () = Print.trace Literal.pp "Clause.paramodulate: lit1" lit1 + val () = Print.trace Rewrite.ppOrient "Clause.paramodulate: ort1" ort1 + val () = Print.trace Term.pp "Clause.paramodulate: tm1" tm1 + val () = Print.trace pp "Clause.paramodulate: cl2" cl2 + val () = Print.trace Literal.pp "Clause.paramodulate: lit2" lit2 + val () = Print.trace Term.ppPath "Clause.paramodulate: path2" path2 + val () = Print.trace Term.pp "Clause.paramodulate: tm2" tm2 *) val Clause {parameters, thm = th1, ...} = cl1 and Clause {thm = th2, ...} = cl2 @@ -15155,33 +20044,37 @@ and lit2 = Literal.subst sub lit2 and th1 = Thm.subst sub th1 and th2 = Thm.subst sub th2 + val _ = isLargerLiteral parameters (Thm.clause th1) lit1 orelse -(*TRACE5 - (trace "Clause.paramodulate: cl1 ordering\n"; false) orelse -*) - raise Error "paramodulate: with clause: ordering constraints" + raise Error "Clause.paramodulate: with clause: ordering" val _ = isLargerLiteral parameters (Thm.clause th2) lit2 orelse -(*TRACE5 - (trace "Clause.paramodulate: cl2 ordering\n"; false) orelse -*) - raise Error "paramodulate: into clause: ordering constraints" + raise Error "Clause.paramodulate: into clause: ordering" + val eqn = (Literal.destEq lit1, th1) val eqn as (l_r,_) = - case ort of + case ort1 of Rewrite.LeftToRight => eqn | Rewrite.RightToLeft => Rule.symEqn eqn +(*MetisTrace6 + val () = Print.trace Rule.ppEquation "Clause.paramodulate: eqn" eqn +*) val _ = isLargerTerm parameters l_r orelse -(*TRACE5 - (trace "Clause.paramodulate: eqn ordering\n"; false) orelse -*) - raise Error "paramodulate: equation: ordering constraints" - val th = Rule.rewrRule eqn lit2 path th2 -(*TRACE5 - val () = Parser.ppTrace Thm.pp "Clause.paramodulate: th" th + raise Error "Clause.paramodulate: equation: ordering constraints" + val th = Rule.rewrRule eqn lit2 path2 th2 +(*MetisTrace5 + val () = Print.trace Thm.pp "Clause.paramodulate: th" th *) in Clause {parameters = parameters, id = newId (), thm = th} - end; + end +(*MetisTrace5 + handle Error err => + let + val () = trace ("Clause.paramodulate: failed: " ^ err ^ "\n") + in + raise Error err + end; +*) end end; @@ -15190,7 +20083,7 @@ (* ========================================================================= *) (* THE ACTIVE SET OF CLAUSES *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Active = @@ -15200,7 +20093,10 @@ (* A type of active clause sets. *) (* ------------------------------------------------------------------------- *) -type simplify = {subsume : bool, reduce : bool, rewrite : bool} +type simplify = + {subsume : bool, + reduce : bool, + rewrite : bool} type parameters = {clause : Metis.Clause.parameters, @@ -15217,13 +20113,15 @@ val size : active -> int -val saturated : active -> Metis.Clause.clause list +val saturation : active -> Metis.Clause.clause list (* ------------------------------------------------------------------------- *) (* Create a new active clause set and initialize clauses. *) (* ------------------------------------------------------------------------- *) -val new : parameters -> Metis.Thm.thm list -> active * Metis.Clause.clause list +val new : + parameters -> {axioms : Metis.Thm.thm list, conjecture : Metis.Thm.thm list} -> + active * {axioms : Metis.Clause.clause list, conjecture : Metis.Clause.clause list} (* ------------------------------------------------------------------------- *) (* Add a clause into the active set and deduce all consequences. *) @@ -15235,7 +20133,7 @@ (* Pretty printing. *) (* ------------------------------------------------------------------------- *) -val pp : active Metis.Parser.pp +val pp : active Metis.Print.pp end @@ -15243,7 +20141,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; @@ -15252,7 +20150,7 @@ (* ========================================================================= *) (* THE ACTIVE SET OF CLAUSES *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Active :> Active = @@ -15264,10 +20162,32 @@ (* Helper functions. *) (* ------------------------------------------------------------------------- *) -local +(*MetisDebug +local + fun mkRewrite ordering = + let + fun add (cl,rw) = + let + val {id, thm = th, ...} = Clause.dest cl + in + case total Thm.destUnitEq th of + SOME l_r => Rewrite.add rw (id,(l_r,th)) + | NONE => rw + end + in + foldl add (Rewrite.new (KnuthBendixOrder.compare ordering)) + end; + fun allFactors red = let - fun allClause cl = List.all red (cl :: Clause.factor cl) + fun allClause cl = + List.all red (cl :: Clause.factor cl) orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.allFactors: cl" cl + in + false + end in List.all allClause end; @@ -15282,6 +20202,12 @@ | SOME cl => allFactors red [cl] in LiteralSet.all allLiteral2 (Clause.literals cl) + end orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.allResolutions: cl2" cl + in + false end fun allClause1 allCls cl = @@ -15291,7 +20217,14 @@ fun allLiteral1 lit = List.all (allClause2 (cl,lit)) allCls in LiteralSet.all allLiteral1 (Clause.literals cl) - end + end orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.allResolutions: cl1" cl + in + false + end + in fn [] => true | allCls as cl :: cls => @@ -15312,9 +20245,24 @@ | SOME cl => allFactors red [cl] in List.all allSubterms (Literal.nonVarTypedSubterms lit) + end orelse + let + val () = Print.trace Literal.pp + "Active.isSaturated.allParamodulations: lit2" lit + in + false end in LiteralSet.all allLiteral2 (Clause.literals cl) + end orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.allParamodulations: cl2" cl + val (_,_,ort,_) = cl_lit_ort_tm + val () = Print.trace Rewrite.ppOrient + "Active.isSaturated.allParamodulations: ort1" ort + in + false end fun allClause1 cl = @@ -15330,9 +20278,21 @@ | SOME (l,r) => allCl2 (cl,lit,Rewrite.LeftToRight,l) andalso allCl2 (cl,lit,Rewrite.RightToLeft,r) + end orelse + let + val () = Print.trace Literal.pp + "Active.isSaturated.allParamodulations: lit1" lit + in + false end in LiteralSet.all allLiteral1 (Clause.literals cl) + end orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.allParamodulations: cl1" cl + in + false end in List.all allClause1 cls @@ -15350,30 +20310,49 @@ val cl' = Clause.reduce reduce cl' val cl' = Clause.rewrite rewrite cl' in - not (Clause.equalThms cl cl') andalso simp cl' + not (Clause.equalThms cl cl') andalso + (simp cl' orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.redundant: cl'" cl' + in + false + end) end in - simp + fn cl => + simp cl orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.redundant: cl" cl + in + false + end end; in fun isSaturated ordering subs cls = let -(*TRACE2 - val ppCls = Parser.ppList Clause.pp - val () = Parser.ppTrace ppCls "Active.checkSaturated: clauses" cls -*) - val red = Units.empty - val rw = Rewrite.new (KnuthBendixOrder.compare ordering) - val red = redundant {subsume = subs, reduce = red, rewrite = rw} - in - allFactors red cls andalso - allResolutions red cls andalso - allParamodulations red cls - end; - - fun checkSaturated ordering subs cls = - if isSaturated ordering subs cls then () else raise Bug "unsaturated"; -end; + val rd = Units.empty + val rw = mkRewrite ordering cls + val red = redundant {subsume = subs, reduce = rd, rewrite = rw} + in + (allFactors red cls andalso + allResolutions red cls andalso + allParamodulations red cls) orelse + let + val () = Print.trace Rewrite.pp "Active.isSaturated: rw" rw + val () = Print.trace (Print.ppList Clause.pp) + "Active.isSaturated: clauses" cls + in + false + end + end; +end; + +fun checkSaturated ordering subs cls = + if isSaturated ordering subs cls then () + else raise Bug "Active.checkSaturated"; +*) (* ------------------------------------------------------------------------- *) (* A type of active clause sets. *) @@ -15453,7 +20432,7 @@ IntMap.foldr add [] cls end; -fun saturated active = +fun saturation active = let fun remove (cl,(cls,subs)) = let @@ -15467,7 +20446,7 @@ val (cls,_) = foldl remove ([], Subsume.new ()) cls val (cls,subs) = foldl remove ([], Subsume.new ()) cls -(*DEBUG +(*MetisDebug val Active {parameters,...} = active val {clause,...} = parameters val {ordering,...} = clause @@ -15485,57 +20464,53 @@ let fun toStr active = "Active{" ^ Int.toString (size active) ^ "}" in - Parser.ppMap toStr Parser.ppString - end; - -(*DEBUG -local - open Parser; - - fun ppField f ppA p a = - (beginBlock p Inconsistent 2; - addString p (f ^ " ="); - addBreak p (1,0); - ppA p a; - endBlock p); + Print.ppMap toStr Print.ppString + end; + +(*MetisDebug +local + fun ppField f ppA a = + Print.blockProgram Print.Inconsistent 2 + [Print.addString (f ^ " ="), + Print.addBreak 1, + ppA a]; val ppClauses = ppField "clauses" - (Parser.ppMap IntMap.toList - (Parser.ppList (Parser.ppPair Parser.ppInt Clause.pp))); + (Print.ppMap IntMap.toList + (Print.ppList (Print.ppPair Print.ppInt Clause.pp))); val ppRewrite = ppField "rewrite" Rewrite.pp; val ppSubterms = ppField "subterms" (TermNet.pp - (Parser.ppMap (fn (c,l,p,t) => ((Clause.id c, l, p), t)) - (Parser.ppPair - (Parser.ppTriple Parser.ppInt Literal.pp Term.ppPath) + (Print.ppMap (fn (c,l,p,t) => ((Clause.id c, l, p), t)) + (Print.ppPair + (Print.ppTriple Print.ppInt Literal.pp Term.ppPath) Term.pp))); in - fun pp p (Active {clauses,rewrite,subterms,...}) = - (beginBlock p Inconsistent 2; - addString p "Active"; - addBreak p (1,0); - beginBlock p Inconsistent 1; - addString p "{"; - ppClauses p clauses; - addString p ","; - addBreak p (1,0); - ppRewrite p rewrite; -(*TRACE5 - addString p ","; - addBreak p (1,0); - ppSubterms p subterms; -*) - endBlock p; - addString p "}"; - endBlock p); -end; -*) - -val toString = Parser.toString pp; + fun pp (Active {clauses,rewrite,subterms,...}) = + Print.blockProgram Print.Inconsistent 2 + [Print.addString "Active", + Print.addBreak 1, + Print.blockProgram Print.Inconsistent 1 + [Print.addString "{", + ppClauses clauses, + Print.addString ",", + Print.addBreak 1, + ppRewrite rewrite, +(*MetisTrace5 + Print.addString ",", + Print.addBreak 1, + ppSubterms subterms, +*) + Print.skip], + Print.addString "}"]; +end; +*) + +val toString = Print.toString pp; (* ------------------------------------------------------------------------- *) (* Simplify clauses. *) @@ -15566,17 +20541,17 @@ end end; -(*DEBUG +(*MetisDebug val simplify = fn simp => fn units => fn rewr => fn subs => fn cl => let - fun traceCl s = Parser.ppTrace Clause.pp ("Active.simplify: " ^ s) -(*TRACE4 - val ppClOpt = Parser.ppOption Clause.pp + fun traceCl s = Print.trace Clause.pp ("Active.simplify: " ^ s) +(*MetisTrace4 + val ppClOpt = Print.ppOption Clause.pp val () = traceCl "cl" cl *) val cl' = simplify simp units rewr subs cl -(*TRACE4 - val () = Parser.ppTrace ppClOpt "Active.simplify: cl'" cl' +(*MetisTrace4 + val () = Print.trace ppClOpt "Active.simplify: cl'" cl' *) val () = case cl' of @@ -15711,8 +20686,8 @@ case total (Clause.resolve cl_lit) (cl,lit) of SOME cl' => cl' :: acc | NONE => acc -(*TRACE4 - val () = Parser.ppTrace Literal.pp "Active.deduceResolution: lit" lit +(*MetisTrace4 + val () = Print.trace Literal.pp "Active.deduceResolution: lit" lit *) in if Atom.isEq atm then acc @@ -15747,13 +20722,30 @@ val eqns = Clause.largestEquations cl val subtms = if TermNet.null equations then [] else Clause.largestSubterms cl +(*MetisTrace5 + val () = Print.trace LiteralSet.pp "Active.deduce: lits" lits + val () = Print.trace + (Print.ppList + (Print.ppMap (fn (lit,ort,_) => (lit,ort)) + (Print.ppPair Literal.pp Rewrite.ppOrient))) + "Active.deduce: eqns" eqns + val () = Print.trace + (Print.ppList + (Print.ppTriple Literal.pp Term.ppPath Term.pp)) + "Active.deduce: subtms" subtms +*) val acc = [] val acc = LiteralSet.foldl (deduceResolution literals cl) acc lits val acc = foldl (deduceParamodulationWith subterms cl) acc eqns val acc = foldl (deduceParamodulationInto equations cl) acc subtms - in - rev acc + val acc = rev acc + +(*MetisTrace5 + val () = Print.trace (Print.ppList Clause.pp) "Active.deduce: acc" acc +*) + in + acc end; (* ------------------------------------------------------------------------- *) @@ -15807,12 +20799,12 @@ in order (tm,tm') = SOME GREATER end - + fun addRed ((cl,tm),acc) = let -(*TRACE5 - val () = Parser.ppTrace Clause.pp "Active.addRed: cl" cl - val () = Parser.ppTrace Term.pp "Active.addRed: tm" tm +(*MetisTrace5 + val () = Print.trace Clause.pp "Active.addRed: cl" cl + val () = Print.trace Term.pp "Active.addRed: tm" tm *) val id = Clause.id cl in @@ -15821,15 +20813,15 @@ else IntSet.add acc id end -(*TRACE5 - val () = Parser.ppTrace Term.pp "Active.addReduce: l" l - val () = Parser.ppTrace Term.pp "Active.addReduce: r" r - val () = Parser.ppTrace Parser.ppBool "Active.addReduce: ord" ord +(*MetisTrace5 + val () = Print.trace Term.pp "Active.addReduce: l" l + val () = Print.trace Term.pp "Active.addReduce: r" r + val () = Print.trace Print.ppBool "Active.addReduce: ord" ord *) in List.foldl addRed acc (TermNet.matched allSubterms l) end - + fun addEquation redexResidues (id,acc) = case Rewrite.peek rewrite id of NONE => acc @@ -15853,7 +20845,7 @@ if choose_clause_rewritables active ids then clause_rewritables active else rewrite_rewritables active ids; -(*DEBUG +(*MetisDebug val rewritables = fn active => fn ids => let val clause_ids = clause_rewritables active @@ -15863,13 +20855,13 @@ if IntSet.equal rewrite_ids clause_ids then () else let - val ppIdl = Parser.ppList Parser.ppInt - val ppIds = Parser.ppMap IntSet.toList ppIdl - val () = Parser.ppTrace pp "Active.rewritables: active" active - val () = Parser.ppTrace ppIdl "Active.rewritables: ids" ids - val () = Parser.ppTrace ppIds + val ppIdl = Print.ppList Print.ppInt + val ppIds = Print.ppMap IntSet.toList ppIdl + val () = Print.trace pp "Active.rewritables: active" active + val () = Print.trace ppIdl "Active.rewritables: ids" ids + val () = Print.trace ppIds "Active.rewritables: clause_ids" clause_ids - val () = Parser.ppTrace ppIds + val () = Print.trace ppIds "Active.rewritables: rewrite_ids" rewrite_ids in raise Bug "Active.rewritables: ~(rewrite_ids SUBSET clause_ids)" @@ -15884,12 +20876,19 @@ else let fun idPred id = not (IntSet.member id ids) - + fun clausePred cl = idPred (Clause.id cl) - + val Active - {parameters,clauses,units,rewrite,subsume,literals, - equations,subterms,allSubterms} = active + {parameters, + clauses, + units, + rewrite, + subsume, + literals, + equations, + subterms, + allSubterms} = active val clauses = IntMap.filter (idPred o fst) clauses and subsume = Subsume.filter clausePred subsume @@ -15899,9 +20898,14 @@ and allSubterms = TermNet.filter (clausePred o fst) allSubterms in Active - {parameters = parameters, clauses = clauses, units = units, - rewrite = rewrite, subsume = subsume, literals = literals, - equations = equations, subterms = subterms, + {parameters = parameters, + clauses = clauses, + units = units, + rewrite = rewrite, + subsume = subsume, + literals = literals, + equations = equations, + subterms = subterms, allSubterms = allSubterms} end; in @@ -15909,21 +20913,21 @@ if Rewrite.isReduced rewrite then (active,[]) else let -(*TRACE3 +(*MetisTrace3 val () = trace "Active.extract_rewritables: inter-reducing\n" *) val (rewrite,ids) = Rewrite.reduce' rewrite val active = setRewrite active rewrite val ids = rewritables active ids val cls = IntSet.transform (IntMap.get clauses) ids -(*TRACE3 - val ppCls = Parser.ppList Clause.pp - val () = Parser.ppTrace ppCls "Active.extract_rewritables: cls" cls +(*MetisTrace3 + val ppCls = Print.ppList Clause.pp + val () = Print.trace ppCls "Active.extract_rewritables: cls" cls *) in (delete active ids, cls) end -(*DEBUG +(*MetisDebug handle Error err => raise Bug ("Active.extract_rewritables: shouldn't fail\n" ^ err); *) @@ -15997,13 +21001,13 @@ fun factor active cls = factor' active [] cls; end; -(*TRACE4 +(*MetisTrace4 val factor = fn active => fn cls => let - val ppCls = Parser.ppList Clause.pp - val () = Parser.ppTrace ppCls "Active.factor: cls" cls + val ppCls = Print.ppList Clause.pp + val () = Print.trace ppCls "Active.factor: cls" cls val (active,cls') = factor active cls - val () = Parser.ppTrace ppCls "Active.factor: cls'" cls' + val () = Print.trace ppCls "Active.factor: cls'" cls' in (active,cls') end; @@ -16013,16 +21017,18 @@ (* Create a new active clause set and initialize clauses. *) (* ------------------------------------------------------------------------- *) -fun new parameters ths = +fun new parameters {axioms,conjecture} = let val {clause,...} = parameters fun mk_clause th = Clause.mk {parameters = clause, id = Clause.newId (), thm = th} - val cls = map mk_clause ths - in - factor (empty parameters) cls + val active = empty parameters + val (active,axioms) = factor active (map mk_clause axioms) + val (active,conjecture) = factor active (map mk_clause conjecture) + in + (active, {axioms = axioms, conjecture = conjecture}) end; (* ------------------------------------------------------------------------- *) @@ -16037,16 +21043,16 @@ else if not (Clause.equalThms cl cl') then factor active [cl'] else let -(*TRACE3 - val () = Parser.ppTrace Clause.pp "Active.add: cl" cl +(*MetisTrace2 + val () = Print.trace Clause.pp "Active.add: cl" cl *) val active = addClause active cl val cl = Clause.freshVars cl val cls = deduce active cl val (active,cls) = factor active cls -(*TRACE2 - val ppCls = Parser.ppList Clause.pp - val () = Parser.ppTrace ppCls "Active.add: cls" cls +(*MetisTrace2 + val ppCls = Print.ppList Clause.pp + val () = Print.trace ppCls "Active.add: cls" cls *) in (active,cls) @@ -16059,22 +21065,48 @@ (* ========================================================================= *) (* THE WAITING SET OF CLAUSES *) -(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Waiting = sig (* ------------------------------------------------------------------------- *) -(* A type of waiting sets of clauses. *) -(* ------------------------------------------------------------------------- *) +(* The parameters control the order that clauses are removed from the *) +(* waiting set: clauses are assigned a weight and removed in strict weight *) +(* order, with smaller weights being removed before larger weights. *) +(* *) +(* The weight of a clause is defined to be *) +(* *) +(* d * s^symbolsWeight * v^variablesWeight * l^literalsWeight * m *) +(* *) +(* where *) +(* *) +(* d = the derivation distance of the clause from the axioms *) +(* s = the number of symbols in the clause *) +(* v = the number of distinct variables in the clause *) +(* l = the number of literals in the clause *) +(* m = the truth of the clause wrt the models *) +(* ------------------------------------------------------------------------- *) + +type weight = real + +type modelParameters = + {model : Metis.Model.parameters, + initialPerturbations : int, + maxChecks : int option, + perturbations : int, + weight : weight} type parameters = - {symbolsWeight : real, - literalsWeight : real, - modelsWeight : real, - modelChecks : int, - models : Metis.Model.parameters list} + {symbolsWeight : weight, + variablesWeight : weight, + literalsWeight : weight, + models : modelParameters list} + +(* ------------------------------------------------------------------------- *) +(* A type of waiting sets of clauses. *) +(* ------------------------------------------------------------------------- *) type waiting @@ -16086,11 +21118,14 @@ val default : parameters -val new : parameters -> Metis.Clause.clause list -> waiting +val new : + parameters -> + {axioms : Metis.Clause.clause list, + conjecture : Metis.Clause.clause list} -> waiting val size : waiting -> int -val pp : waiting Metis.Parser.pp +val pp : waiting Metis.Print.pp (* ------------------------------------------------------------------------- *) (* Adding new clauses. *) @@ -16110,7 +21145,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; @@ -16119,7 +21154,7 @@ (* ========================================================================= *) (* THE WAITING SET OF CLAUSES *) -(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Waiting :> Waiting = @@ -16131,35 +21166,23 @@ (* A type of waiting sets of clauses. *) (* ------------------------------------------------------------------------- *) -(* The parameter type controls the heuristics for clause selection. *) -(* Increasing any of the *Weight parameters will favour clauses with low *) -(* values of that field. *) - -(* Note that there is an extra parameter of inference distance from the *) -(* starting axioms (a.k.a. time) which has a fixed weight of 1, so all *) -(* the other parameters should be set relative to this baseline. *) - -(* The first two parameters, symbolsWeight and literalsWeight, control the *) -(* time:weight ratio, i.e., whether to favour clauses derived in a few *) -(* steps from the axioms (time) or whether to favour small clauses (weight). *) -(* Small can be a combination of low number of symbols (the symbolWeight *) -(* parameter) or literals (the literalsWeight parameter). *) - -(* modelsWeight controls the semantic guidance. Increasing this parameter *) -(* favours clauses that return false more often when interpreted *) -(* modelChecks times over the given list of models. *) +type weight = real; + +type modelParameters = + {model : Model.parameters, + initialPerturbations : int, + maxChecks : int option, + perturbations : int, + weight : weight} type parameters = - {symbolsWeight : real, - literalsWeight : real, - modelsWeight : real, - modelChecks : int, - models : Model.parameters list}; + {symbolsWeight : weight, + variablesWeight : weight, + literalsWeight : weight, + models : modelParameters list}; type distance = real; -type weight = real; - datatype waiting = Waiting of {parameters : parameters, @@ -16170,27 +21193,105 @@ (* Basic operations. *) (* ------------------------------------------------------------------------- *) +val defaultModels : modelParameters list = + [(* MODIFIED by Jasmin Blanchette + {model = Model.default, + initialPerturbations = 100, + maxChecks = SOME 20, + perturbations = 0, + weight = 1.0} *)]; + val default : parameters = {symbolsWeight = 1.0, - literalsWeight = 0.0, - modelsWeight = 0.0, - modelChecks = 20, - models = []}; + literalsWeight = (* 1.0 *) 0.0, (* MODIFIED by Jasmin Blanchette *) + variablesWeight = (* 1.0 *) 0.0, (* MODIFIED by Jasmin Blanchette *) + models = defaultModels}; fun size (Waiting {clauses,...}) = Heap.size clauses; val pp = - Parser.ppMap + Print.ppMap (fn w => "Waiting{" ^ Int.toString (size w) ^ "}") - Parser.ppString; - -(*DEBUG + Print.ppString; + +(*MetisDebug val pp = - Parser.ppMap + Print.ppMap (fn Waiting {clauses,...} => map (fn (w,(_,cl)) => (w, Clause.id cl, cl)) (Heap.toList clauses)) - (Parser.ppList (Parser.ppTriple Parser.ppReal Parser.ppInt Clause.pp)); -*) + (Print.ppList (Print.ppTriple Print.ppReal Print.ppInt Clause.pp)); +*) + +(* ------------------------------------------------------------------------- *) +(* Perturbing the models. *) +(* ------------------------------------------------------------------------- *) + +type modelClause = NameSet.set * Thm.clause; + +fun mkModelClause cl = + let + val lits = Clause.literals cl + val fvs = LiteralSet.freeVars lits + in + (fvs,lits) + end; + +val mkModelClauses = map mkModelClause; + +fun perturbModel M cls = + if null cls then K () + else + let + val N = {size = Model.size M} + + fun perturbClause (fv,cl) = + let + val V = Model.randomValuation N fv + in + if Model.interpretClause M V cl then () + else Model.perturbClause M V cl + end + + fun perturbClauses () = app perturbClause cls + in + fn n => funpow n perturbClauses () + end; + +fun initialModel axioms conjecture parm = + let + val {model,initialPerturbations,...} : modelParameters = parm + val m = Model.new model + val () = perturbModel m conjecture initialPerturbations + val () = perturbModel m axioms initialPerturbations + in + m + end; + +fun checkModels parms models (fv,cl) = + let + fun check ((parm,model),z) = + let + val {maxChecks,weight,...} : modelParameters = parm + val n = {maxChecks = maxChecks} + val {T,F} = Model.check Model.interpretClause n model fv cl + in + Math.pow (1.0 + Real.fromInt T / Real.fromInt (T + F), weight) * z + end + in + List.foldl check 1.0 (zip parms models) + end; + +fun perturbModels parms models cls = + let + fun perturb (parm,model) = + let + val {perturbations,...} : modelParameters = parm + in + perturbModel model cls perturbations + end + in + app perturb (zip parms models) + end; (* ------------------------------------------------------------------------- *) (* Clause weights. *) @@ -16199,40 +21300,39 @@ local fun clauseSymbols cl = Real.fromInt (LiteralSet.typedSymbols cl); + fun clauseVariables cl = + Real.fromInt (NameSet.size (LiteralSet.freeVars cl) + 1); + fun clauseLiterals cl = Real.fromInt (LiteralSet.size cl); - fun clauseSat modelChecks models cl = - let - fun g {T,F} = (Real.fromInt T / Real.fromInt (T + F)) + 1.0 - fun f (m,z) = g (Model.checkClause {maxChecks = modelChecks} m cl) * z - in - foldl f 1.0 models - end; - - fun priority cl = 1e~12 * Real.fromInt (Clause.id cl); -in - fun clauseWeight (parm : parameters) models dist cl = - let -(*TRACE3 - val () = Parser.ppTrace Clause.pp "Waiting.clauseWeight: cl" cl -*) - val {symbolsWeight,literalsWeight,modelsWeight,modelChecks,...} = parm + fun clausePriority cl = 1e~12 * Real.fromInt (Clause.id cl); +in + fun clauseWeight (parm : parameters) mods dist mcl cl = + let +(*MetisTrace3 + val () = Print.trace Clause.pp "Waiting.clauseWeight: cl" cl +*) + val {symbolsWeight,variablesWeight,literalsWeight,models,...} = parm val lits = Clause.literals cl val symbolsW = Math.pow (clauseSymbols lits, symbolsWeight) + val variablesW = Math.pow (clauseVariables lits, variablesWeight) val literalsW = Math.pow (clauseLiterals lits, literalsWeight) - val modelsW = Math.pow (clauseSat modelChecks models lits, modelsWeight) -(*TRACE4 + val modelsW = (* checkModels models mods mcl *) 1.0 (* MODIFIED by Jasmin Blanchette *) +(*MetisTrace4 val () = trace ("Waiting.clauseWeight: dist = " ^ Real.toString dist ^ "\n") val () = trace ("Waiting.clauseWeight: symbolsW = " ^ Real.toString symbolsW ^ "\n") + val () = trace ("Waiting.clauseWeight: variablesW = " ^ + Real.toString variablesW ^ "\n") val () = trace ("Waiting.clauseWeight: literalsW = " ^ Real.toString literalsW ^ "\n") val () = trace ("Waiting.clauseWeight: modelsW = " ^ Real.toString modelsW ^ "\n") *) - val weight = dist * symbolsW * literalsW * modelsW + priority cl -(*TRACE3 + val weight = dist * symbolsW * variablesW * literalsW * modelsW + val weight = weight + clausePriority cl +(*MetisTrace3 val () = trace ("Waiting.clauseWeight: weight = " ^ Real.toString weight ^ "\n") *) @@ -16245,29 +21345,39 @@ (* Adding new clauses. *) (* ------------------------------------------------------------------------- *) -fun add waiting (_,[]) = waiting - | add waiting (dist,cls) = - let -(*TRACE3 - val () = Parser.ppTrace pp "Waiting.add: waiting" waiting - val () = Parser.ppTrace (Parser.ppList Clause.pp) "Waiting.add: cls" cls -*) - +fun add' waiting dist mcls cls = + let val Waiting {parameters,clauses,models} = waiting + val {models = modelParameters, ...} = parameters val dist = dist + Math.ln (Real.fromInt (length cls)) - val weight = clauseWeight parameters models dist - - fun f (cl,acc) = Heap.add acc (weight cl, (dist,cl)) - - val clauses = foldl f clauses cls - - val waiting = - Waiting {parameters = parameters, clauses = clauses, models = models} - -(*TRACE3 - val () = Parser.ppTrace pp "Waiting.add: waiting" waiting + fun addCl ((mcl,cl),acc) = + let + val weight = clauseWeight parameters models dist mcl cl + in + Heap.add acc (weight,(dist,cl)) + end + + val clauses = List.foldl addCl clauses (zip mcls cls) + + val () = perturbModels modelParameters models mcls + in + Waiting {parameters = parameters, clauses = clauses, models = models} + end; + +fun add waiting (_,[]) = waiting + | add waiting (dist,cls) = + let +(*MetisTrace3 + val () = Print.trace pp "Waiting.add: waiting" waiting + val () = Print.trace (Print.ppList Clause.pp) "Waiting.add: cls" cls +*) + + val waiting = add' waiting dist (mkModelClauses cls) cls + +(*MetisTrace3 + val () = Print.trace pp "Waiting.add: waiting" waiting *) in waiting @@ -16276,15 +21386,24 @@ local fun cmp ((w1,_),(w2,_)) = Real.compare (w1,w2); - fun empty parameters = - let + fun empty parameters axioms conjecture = + let + val {models = modelParameters, ...} = parameters val clauses = Heap.new cmp - and models = map Model.new (#models parameters) + and models = map (initialModel axioms conjecture) modelParameters in Waiting {parameters = parameters, clauses = clauses, models = models} end; in - fun new parameters cls = add (empty parameters) (0.0,cls); + fun new parameters {axioms,conjecture} = + let + val mAxioms = mkModelClauses axioms + and mConjecture = mkModelClauses conjecture + + val waiting = empty parameters mAxioms mConjecture + in + add' waiting 0.0 (mAxioms @ mConjecture) (axioms @ conjecture) + end; end; (* ------------------------------------------------------------------------- *) @@ -16310,7 +21429,7 @@ (* ========================================================================= *) (* THE RESOLUTION PROOF PROCEDURE *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Resolution = @@ -16332,13 +21451,15 @@ val default : parameters -val new : parameters -> Metis.Thm.thm list -> resolution +val new : + parameters -> {axioms : Metis.Thm.thm list, conjecture : Metis.Thm.thm list} -> + resolution val active : resolution -> Metis.Active.active val waiting : resolution -> Metis.Waiting.waiting -val pp : resolution Metis.Parser.pp +val pp : resolution Metis.Print.pp (* ------------------------------------------------------------------------- *) (* The main proof loop. *) @@ -16362,7 +21483,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; @@ -16371,7 +21492,7 @@ (* ========================================================================= *) (* THE RESOLUTION PROOF PROCEDURE *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Resolution :> Resolution = @@ -16380,7 +21501,7 @@ open Useful; (* ------------------------------------------------------------------------- *) -(* Parameters. *) +(* A type of resolution proof procedures. *) (* ------------------------------------------------------------------------- *) type parameters = @@ -16415,11 +21536,11 @@ fun waiting (Resolution {waiting = w, ...}) = w; val pp = - Parser.ppMap + Print.ppMap (fn Resolution {active,waiting,...} => "Resolution(" ^ Int.toString (Active.size active) ^ "<-" ^ Int.toString (Waiting.size waiting) ^ ")") - Parser.ppString; + Print.ppString; (* ------------------------------------------------------------------------- *) (* The main proof loop. *) @@ -16436,21 +21557,21 @@ fun iterate resolution = let val Resolution {parameters,active,waiting} = resolution -(*TRACE2 - val () = Parser.ppTrace Active.pp "Resolution.iterate: active" active - val () = Parser.ppTrace Waiting.pp "Resolution.iterate: waiting" waiting +(*MetisTrace2 + val () = Print.trace Active.pp "Resolution.iterate: active" active + val () = Print.trace Waiting.pp "Resolution.iterate: waiting" waiting *) in case Waiting.remove waiting of NONE => - Decided (Satisfiable (map Clause.thm (Active.saturated active))) + Decided (Satisfiable (map Clause.thm (Active.saturation active))) | SOME ((d,cl),waiting) => if Clause.isContradiction cl then Decided (Contradiction (Clause.thm cl)) else let -(*TRACE1 - val () = Parser.ppTrace Clause.pp "Resolution.iterate: cl" cl +(*MetisTrace1 + val () = Print.trace Clause.pp "Resolution.iterate: cl" cl *) val (active,cls) = Active.add active cl val waiting = Waiting.add waiting (d,cls) @@ -16472,20 +21593,75 @@ (**** Original file: Tptp.sig ****) (* ========================================================================= *) -(* THE TPTP PROBLEM FILE FORMAT (TPTP v2) *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* THE TPTP PROBLEM FILE FORMAT *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Tptp = sig (* ------------------------------------------------------------------------- *) -(* Mapping TPTP functions and relations to different names. *) -(* ------------------------------------------------------------------------- *) - -val functionMapping : {name : string, arity : int, tptp : string} list Unsynchronized.ref - -val relationMapping : {name : string, arity : int, tptp : string} list Unsynchronized.ref +(* Mapping to and from TPTP variable, function and relation names. *) +(* ------------------------------------------------------------------------- *) + +type mapping + +val defaultMapping : mapping + +val mkMapping : + {functionMapping : {name : Metis.Name.name, arity : int, tptp : string} list, + relationMapping : {name : Metis.Name.name, arity : int, tptp : string} list} -> + mapping + +val addVarSetMapping : mapping -> Metis.NameSet.set -> mapping + +(* ------------------------------------------------------------------------- *) +(* Interpreting TPTP functions and relations in a finite model. *) +(* ------------------------------------------------------------------------- *) + +val defaultFixedMap : Metis.Model.fixedMap + +val defaultModel : Metis.Model.parameters + +val ppFixedMap : mapping -> Metis.Model.fixedMap Metis.Print.pp + +(* ------------------------------------------------------------------------- *) +(* TPTP roles. *) +(* ------------------------------------------------------------------------- *) + +datatype role = + AxiomRole + | ConjectureRole + | DefinitionRole + | NegatedConjectureRole + | PlainRole + | TheoremRole + | OtherRole of string; + +val isCnfConjectureRole : role -> bool + +val isFofConjectureRole : role -> bool + +val toStringRole : role -> string + +val fromStringRole : string -> role + +val ppRole : role Metis.Print.pp + +(* ------------------------------------------------------------------------- *) +(* SZS statuses. *) +(* ------------------------------------------------------------------------- *) + +datatype status = + CounterSatisfiableStatus + | TheoremStatus + | SatisfiableStatus + | UnknownStatus + | UnsatisfiableStatus + +val toStringStatus : status -> string + +val ppStatus : status Metis.Print.pp (* ------------------------------------------------------------------------- *) (* TPTP literals. *) @@ -16495,67 +21671,153 @@ Boolean of bool | Literal of Metis.Literal.literal -val negate : literal -> literal - -val literalFunctions : literal -> Metis.NameAritySet.set - -val literalRelation : literal -> Metis.Atom.relation option - -val literalFreeVars : literal -> Metis.NameSet.set +val negateLiteral : literal -> literal + +val functionsLiteral : literal -> Metis.NameAritySet.set + +val relationLiteral : literal -> Metis.Atom.relation option + +val freeVarsLiteral : literal -> Metis.NameSet.set + +(* ------------------------------------------------------------------------- *) +(* TPTP formula names. *) +(* ------------------------------------------------------------------------- *) + +datatype formulaName = + FormulaName of string + +val ppFormulaName : formulaName Metis.Print.pp + +(* ------------------------------------------------------------------------- *) +(* TPTP formula bodies. *) +(* ------------------------------------------------------------------------- *) + +datatype formulaBody = + CnfFormulaBody of literal list + | FofFormulaBody of Metis.Formula.formula + +(* ------------------------------------------------------------------------- *) +(* TPTP formula sources. *) +(* ------------------------------------------------------------------------- *) + +datatype formulaSource = + NoFormulaSource + | StripFormulaSource of + {inference : string, + parents : formulaName list} + | NormalizeFormulaSource of + {inference : Metis.Normalize.inference, + parents : formulaName list} + | ProofFormulaSource of + {inference : Metis.Proof.inference, + parents : formulaName list} (* ------------------------------------------------------------------------- *) (* TPTP formulas. *) (* ------------------------------------------------------------------------- *) datatype formula = - CnfFormula of {name : string, role : string, clause : literal list} - | FofFormula of {name : string, role : string, formula : Metis.Formula.formula} - -val formulaFunctions : formula -> Metis.NameAritySet.set - -val formulaRelations : formula -> Metis.NameAritySet.set - -val formulaFreeVars : formula -> Metis.NameSet.set - -val formulaIsConjecture : formula -> bool - -val ppFormula : formula Metis.Parser.pp - -val parseFormula : char Metis.Stream.stream -> formula Metis.Stream.stream - -val formulaToString : formula -> string - -val formulaFromString : string -> formula + Formula of + {name : formulaName, + role : role, + body : formulaBody, + source : formulaSource} + +val nameFormula : formula -> formulaName + +val roleFormula : formula -> role + +val bodyFormula : formula -> formulaBody + +val sourceFormula : formula -> formulaSource + +val functionsFormula : formula -> Metis.NameAritySet.set + +val relationsFormula : formula -> Metis.NameAritySet.set + +val freeVarsFormula : formula -> Metis.NameSet.set + +val freeVarsListFormula : formula list -> Metis.NameSet.set + +val isCnfConjectureFormula : formula -> bool +val isFofConjectureFormula : formula -> bool +val isConjectureFormula : formula -> bool + +(* ------------------------------------------------------------------------- *) +(* Clause information. *) +(* ------------------------------------------------------------------------- *) + +datatype clauseSource = + CnfClauseSource of formulaName * literal list + | FofClauseSource of Metis.Normalize.thm + +type 'a clauseInfo = 'a Metis.LiteralSetMap.map + +type clauseNames = formulaName clauseInfo + +type clauseRoles = role clauseInfo + +type clauseSources = clauseSource clauseInfo + +val noClauseNames : clauseNames + +val noClauseRoles : clauseRoles + +val noClauseSources : clauseSources (* ------------------------------------------------------------------------- *) (* TPTP problems. *) (* ------------------------------------------------------------------------- *) -datatype goal = - Cnf of Metis.Problem.problem - | Fof of Metis.Formula.formula - -type problem = {comments : string list, formulas : formula list} - -val read : {filename : string} -> problem - -val write : {filename : string} -> problem -> unit - +type comments = string list + +type includes = string list + +datatype problem = + Problem of + {comments : comments, + includes : includes, + formulas : formula list} + +val hasCnfConjecture : problem -> bool +val hasFofConjecture : problem -> bool val hasConjecture : problem -> bool -val toGoal : problem -> goal - -val fromProblem : Metis.Problem.problem -> problem - -val prove : {filename : string} -> bool +val freeVars : problem -> Metis.NameSet.set + +val mkProblem : + {comments : comments, + includes : includes, + names : clauseNames, + roles : clauseRoles, + problem : Metis.Problem.problem} -> problem + +val normalize : + problem -> + {subgoal : Metis.Formula.formula * formulaName list, + problem : Metis.Problem.problem, + sources : clauseSources} list + +val goal : problem -> Metis.Formula.formula + +val read : {mapping : mapping, filename : string} -> problem + +val write : + {problem : problem, + mapping : mapping, + filename : string} -> unit + +val prove : {filename : string, mapping : mapping} -> bool (* ------------------------------------------------------------------------- *) (* TSTP proofs. *) (* ------------------------------------------------------------------------- *) -val ppProof : Metis.Proof.proof Metis.Parser.pp - -val proofToString : Metis.Proof.proof -> string +val fromProof : + {problem : problem, + proofs : {subgoal : Metis.Formula.formula * formulaName list, + sources : clauseSources, + refutation : Metis.Thm.thm} list} -> formula list end @@ -16563,7 +21825,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; @@ -16571,8 +21833,8 @@ val foldr = List.foldr; (* ========================================================================= *) -(* THE TPTP PROBLEM FILE FORMAT (TPTP v2) *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* THE TPTP PROBLEM FILE FORMAT *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Tptp :> Tptp = @@ -16581,11 +21843,95 @@ open Useful; (* ------------------------------------------------------------------------- *) -(* Constants. *) -(* ------------------------------------------------------------------------- *) - -val ROLE_NEGATED_CONJECTURE = "negated_conjecture" -and ROLE_CONJECTURE = "conjecture"; +(* Default TPTP function and relation name mapping. *) +(* ------------------------------------------------------------------------- *) + +val defaultFunctionMapping = + [(* Mapping TPTP functions to infix symbols *) + {name = "~", arity = 1, tptp = "negate"}, + {name = "*", arity = 2, tptp = "multiply"}, + {name = "/", arity = 2, tptp = "divide"}, + {name = "+", arity = 2, tptp = "add"}, + {name = "-", arity = 2, tptp = "subtract"}, + {name = "::", arity = 2, tptp = "cons"}, + {name = "@", arity = 2, tptp = "append"}, + {name = ",", arity = 2, tptp = "pair"}, + (* Expanding HOL symbols to TPTP alphanumerics *) + {name = ":", arity = 2, tptp = "has_type"}, + {name = ".", arity = 2, tptp = "apply"}]; + +val defaultRelationMapping = + [(* Mapping TPTP relations to infix symbols *) + {name = "=", arity = 2, tptp = "="}, (* this preserves the = symbol *) + {name = "==", arity = 2, tptp = "equalish"}, + {name = "<=", arity = 2, tptp = "less_equal"}, + {name = "<", arity = 2, tptp = "less_than"}, + {name = ">=", arity = 2, tptp = "greater_equal"}, + {name = ">", arity = 2, tptp = "greater_than"}, + (* Expanding HOL symbols to TPTP alphanumerics *) + {name = "{}", arity = 1, tptp = "bool"}]; + +(* ------------------------------------------------------------------------- *) +(* Interpreting TPTP functions and relations in a finite model. *) +(* ------------------------------------------------------------------------- *) + +val defaultFunctionModel = + [{name = "~", arity = 1, model = Model.negName}, + {name = "*", arity = 2, model = Model.multName}, + {name = "/", arity = 2, model = Model.divName}, + {name = "+", arity = 2, model = Model.addName}, + {name = "-", arity = 2, model = Model.subName}, + {name = "::", arity = 2, model = Model.consName}, + {name = "@", arity = 2, model = Model.appendName}, + {name = ":", arity = 2, model = Term.hasTypeFunctionName}, + {name = "additive_identity", arity = 0, model = Model.numeralName 0}, + {name = "app", arity = 2, model = Model.appendName}, + {name = "complement", arity = 1, model = Model.complementName}, + {name = "difference", arity = 2, model = Model.differenceName}, + {name = "divide", arity = 2, model = Model.divName}, + {name = "empty_set", arity = 0, model = Model.emptyName}, + {name = "identity", arity = 0, model = Model.numeralName 1}, + {name = "identity_map", arity = 1, model = Model.projectionName 1}, + {name = "intersection", arity = 2, model = Model.intersectName}, + {name = "minus", arity = 1, model = Model.negName}, + {name = "multiplicative_identity", arity = 0, model = Model.numeralName 1}, + {name = "n0", arity = 0, model = Model.numeralName 0}, + {name = "n1", arity = 0, model = Model.numeralName 1}, + {name = "n2", arity = 0, model = Model.numeralName 2}, + {name = "n3", arity = 0, model = Model.numeralName 3}, + {name = "n4", arity = 0, model = Model.numeralName 4}, + {name = "n5", arity = 0, model = Model.numeralName 5}, + {name = "n6", arity = 0, model = Model.numeralName 6}, + {name = "n7", arity = 0, model = Model.numeralName 7}, + {name = "n8", arity = 0, model = Model.numeralName 8}, + {name = "n9", arity = 0, model = Model.numeralName 9}, + {name = "nil", arity = 0, model = Model.nilName}, + {name = "null_class", arity = 0, model = Model.emptyName}, + {name = "singleton", arity = 1, model = Model.singletonName}, + {name = "successor", arity = 1, model = Model.sucName}, + {name = "symmetric_difference", arity = 2, + model = Model.symmetricDifferenceName}, + {name = "union", arity = 2, model = Model.unionName}, + {name = "universal_class", arity = 0, model = Model.universeName}]; + +val defaultRelationModel = + [{name = "=", arity = 2, model = Atom.eqRelationName}, + {name = "==", arity = 2, model = Atom.eqRelationName}, + {name = "<=", arity = 2, model = Model.leName}, + {name = "<", arity = 2, model = Model.ltName}, + {name = ">=", arity = 2, model = Model.geName}, + {name = ">", arity = 2, model = Model.gtName}, + {name = "divides", arity = 2, model = Model.dividesName}, + {name = "element_of_set", arity = 2, model = Model.memberName}, + {name = "equal", arity = 2, model = Atom.eqRelationName}, + {name = "equal_elements", arity = 2, model = Atom.eqRelationName}, + {name = "equal_sets", arity = 2, model = Atom.eqRelationName}, + {name = "equivalent", arity = 2, model = Atom.eqRelationName}, + {name = "less", arity = 2, model = Model.ltName}, + {name = "less_or_equal", arity = 2, model = Model.leName}, + {name = "member", arity = 2, model = Model.memberName}, + {name = "subclass", arity = 2, model = Model.subsetName}, + {name = "subset", arity = 2, model = Model.subsetName}]; (* ------------------------------------------------------------------------- *) (* Helper functions. *) @@ -16601,97 +21947,451 @@ n > 0 andalso hp (String.sub (s,0)) andalso ct (n - 1) end; -(* ------------------------------------------------------------------------- *) -(* Mapping TPTP functions and relations to different names. *) -(* ------------------------------------------------------------------------- *) - -val functionMapping = Unsynchronized.ref - [(* Mapping TPTP functions to infix symbols *) - {name = "*", arity = 2, tptp = "multiply"}, - {name = "/", arity = 2, tptp = "divide"}, - {name = "+", arity = 2, tptp = "add"}, - {name = "-", arity = 2, tptp = "subtract"}, - {name = "::", arity = 2, tptp = "cons"}, - {name = ",", arity = 2, tptp = "pair"}, - (* Expanding HOL symbols to TPTP alphanumerics *) - {name = ":", arity = 2, tptp = "has_type"}, - {name = ".", arity = 2, tptp = "apply"}, - {name = "<=", arity = 0, tptp = "less_equal"}]; - -val relationMapping = Unsynchronized.ref - [(* Mapping TPTP relations to infix symbols *) - {name = "=", arity = 2, tptp = "="}, - {name = "==", arity = 2, tptp = "equalish"}, - {name = "<=", arity = 2, tptp = "less_equal"}, - {name = "<", arity = 2, tptp = "less_than"}, - (* Expanding HOL symbols to TPTP alphanumerics *) - {name = "{}", arity = 1, tptp = "bool"}]; - -fun mappingToTptp x = - let - fun mk ({name,arity,tptp},m) = NameArityMap.insert m ((name,arity),tptp) - in - foldl mk (NameArityMap.new ()) x - end; - -fun mappingFromTptp x = - let - fun mk ({name,arity,tptp},m) = NameArityMap.insert m ((tptp,arity),name) - in - foldl mk (NameArityMap.new ()) x - end; - -fun findMapping mapping (name_arity as (n,_)) = - Option.getOpt (NameArityMap.peek mapping name_arity, n); - -fun mapTerm functionMap = - let - val mapName = findMapping functionMap - - fun mapTm (tm as Term.Var _) = tm - | mapTm (Term.Fn (f,a)) = Term.Fn (mapName (f, length a), map mapTm a) - in - mapTm - end; - -fun mapAtom {functionMap,relationMap} (p,a) = - (findMapping relationMap (p, length a), map (mapTerm functionMap) a); - -fun mapFof maps = - let - open Formula - - fun form True = True - | form False = False - | form (Atom a) = Atom (mapAtom maps a) - | form (Not p) = Not (form p) - | form (And (p,q)) = And (form p, form q) - | form (Or (p,q)) = Or (form p, form q) - | form (Imp (p,q)) = Imp (form p, form q) - | form (Iff (p,q)) = Iff (form p, form q) - | form (Forall (v,p)) = Forall (v, form p) - | form (Exists (v,p)) = Exists (v, form p) - in - form - end; - -(* ------------------------------------------------------------------------- *) -(* Comments. *) -(* ------------------------------------------------------------------------- *) - -fun mkComment "" = "%" - | mkComment line = "% " ^ line; - -fun destComment "" = "" - | destComment l = - let - val _ = String.sub (l,0) = #"%" orelse raise Error "destComment" - val n = if size l >= 2 andalso String.sub (l,1) = #" " then 2 else 1 - in - String.extract (l,n,NONE) - end; - -val isComment = can destComment; +fun stripSuffix pred s = + let + fun f 0 = "" + | f n = + let + val n' = n - 1 + in + if pred (String.sub (s,n')) then f n' + else String.substring (s,0,n) + end + in + f (size s) + end; + +fun variant avoid s = + if not (StringSet.member s avoid) then s + else + let + val s = stripSuffix Char.isDigit s + + fun var i = + let + val s_i = s ^ Int.toString i + in + if not (StringSet.member s_i avoid) then s_i else var (i + 1) + end + in + var 0 + end; + +(* ------------------------------------------------------------------------- *) +(* Mapping to legal TPTP names. *) +(* ------------------------------------------------------------------------- *) + +local + fun nonEmptyPred p l = + case l of + [] => false + | c :: cs => p (c,cs); + + fun existsPred l x = List.exists (fn p => p x) l; + + fun isTptpChar #"_" = true + | isTptpChar c = Char.isAlphaNum c; + + fun isTptpName l s = nonEmptyPred (existsPred l) (explode s); + + fun isRegular (c,cs) = + Char.isLower c andalso List.all isTptpChar cs; + + fun isNumber (c,cs) = + Char.isDigit c andalso List.all Char.isDigit cs; + + fun isDefined (c,cs) = + c = #"$" andalso nonEmptyPred isRegular cs; + + fun isSystem (c,cs) = + c = #"$" andalso nonEmptyPred isDefined cs; +in + fun mkTptpVarName s = + let + val s = + case List.filter isTptpChar (explode s) of + [] => [#"X"] + | l as c :: cs => + if Char.isUpper c then l + else if Char.isLower c then Char.toUpper c :: cs + else #"X" :: l + in + implode s + end; + + val isTptpConstName = isTptpName [isRegular,isNumber,isDefined,isSystem] + and isTptpFnName = isTptpName [isRegular,isDefined,isSystem] + and isTptpPropName = isTptpName [isRegular,isDefined,isSystem] + and isTptpRelName = isTptpName [isRegular,isDefined,isSystem]; + + val isTptpFormulaName = isTptpName [isRegular,isNumber]; +end; + +(* ------------------------------------------------------------------------- *) +(* Mapping to legal TPTP variable names. *) +(* ------------------------------------------------------------------------- *) + +datatype varToTptp = VarToTptp of StringSet.set * string NameMap.map; + +val emptyVarToTptp = VarToTptp (StringSet.empty, NameMap.new ()); + +fun addVarToTptp vm v = + let + val VarToTptp (avoid,mapping) = vm + in + if NameMap.inDomain v mapping then vm + else + let + val s = variant avoid (mkTptpVarName (Name.toString v)) + + val avoid = StringSet.add avoid s + and mapping = NameMap.insert mapping (v,s) + in + VarToTptp (avoid,mapping) + end + end; + +local + fun add (v,vm) = addVarToTptp vm v; +in + val addListVarToTptp = List.foldl add; + + val addSetVarToTptp = NameSet.foldl add; +end; + +val fromListVarToTptp = addListVarToTptp emptyVarToTptp; + +val fromSetVarToTptp = addSetVarToTptp emptyVarToTptp; + +fun getVarToTptp vm v = + let + val VarToTptp (_,mapping) = vm + in + case NameMap.peek mapping v of + SOME s => s + | NONE => raise Bug "Tptp.getVarToTptp: unknown var" + end; + +(* ------------------------------------------------------------------------- *) +(* Mapping from TPTP variable names. *) +(* ------------------------------------------------------------------------- *) + +fun getVarFromTptp s = Name.fromString s; + +(* ------------------------------------------------------------------------- *) +(* Mapping to TPTP function and relation names. *) +(* ------------------------------------------------------------------------- *) + +datatype nameToTptp = NameToTptp of string NameArityMap.map; + +local + val emptyNames : string NameArityMap.map = NameArityMap.new (); + + fun addNames ({name,arity,tptp},mapping) = + NameArityMap.insert mapping ((name,arity),tptp); + + val fromListNames = List.foldl addNames emptyNames; +in + fun mkNameToTptp mapping = NameToTptp (fromListNames mapping); +end; + +local + fun escapeChar c = + case c of + #"\\" => "\\\\" + | #"'" => "\\'" + | #"\n" => "\\n" + | #"\t" => "\\t" + | _ => str c; + + val escapeString = String.translate escapeChar; +in + fun singleQuote s = "'" ^ escapeString s ^ "'"; +end; + +fun getNameToTptp isTptp s = if isTptp s then s else singleQuote s; + +fun getNameArityToTptp isZeroTptp isPlusTptp (NameToTptp mapping) na = + case NameArityMap.peek mapping na of + SOME s => s + | NONE => + let + val (n,a) = na + val isTptp = if a = 0 then isZeroTptp else isPlusTptp + val s = Name.toString n + in + getNameToTptp isTptp s + end; + +(* ------------------------------------------------------------------------- *) +(* Mapping from TPTP function and relation names. *) +(* ------------------------------------------------------------------------- *) + +datatype nameFromTptp = NameFromTptp of (string * int, Name.name) Map.map; + +local + val stringArityCompare = prodCompare String.compare Int.compare; + + val emptyStringArityMap = Map.new stringArityCompare; + + fun addStringArityMap ({name,arity,tptp},mapping) = + Map.insert mapping ((tptp,arity),name); + + val fromListStringArityMap = + List.foldl addStringArityMap emptyStringArityMap; +in + fun mkNameFromTptp mapping = NameFromTptp (fromListStringArityMap mapping); +end; + +fun getNameFromTptp (NameFromTptp mapping) sa = + case Map.peek mapping sa of + SOME n => n + | NONE => + let + val (s,_) = sa + in + Name.fromString s + end; + +(* ------------------------------------------------------------------------- *) +(* Mapping to and from TPTP variable, function and relation names. *) +(* ------------------------------------------------------------------------- *) + +datatype mapping = + Mapping of + {varTo : varToTptp, + fnTo : nameToTptp, + relTo : nameToTptp, + fnFrom : nameFromTptp, + relFrom : nameFromTptp}; + +fun mkMapping mapping = + let + val {functionMapping,relationMapping} = mapping + + val varTo = emptyVarToTptp + val fnTo = mkNameToTptp functionMapping + val relTo = mkNameToTptp relationMapping + + val fnFrom = mkNameFromTptp functionMapping + val relFrom = mkNameFromTptp relationMapping + in + Mapping + {varTo = varTo, + fnTo = fnTo, + relTo = relTo, + fnFrom = fnFrom, + relFrom = relFrom} + end; + +fun addVarListMapping mapping vs = + let + val Mapping + {varTo, + fnTo, + relTo, + fnFrom, + relFrom} = mapping + + val varTo = addListVarToTptp varTo vs + in + Mapping + {varTo = varTo, + fnTo = fnTo, + relTo = relTo, + fnFrom = fnFrom, + relFrom = relFrom} + end; + +fun addVarSetMapping mapping vs = + let + val Mapping + {varTo, + fnTo, + relTo, + fnFrom, + relFrom} = mapping + + val varTo = addSetVarToTptp varTo vs + in + Mapping + {varTo = varTo, + fnTo = fnTo, + relTo = relTo, + fnFrom = fnFrom, + relFrom = relFrom} + end; + +fun varToTptp mapping v = + let + val Mapping {varTo,...} = mapping + in + getVarToTptp varTo v + end; + +fun fnToTptp mapping fa = + let + val Mapping {fnTo,...} = mapping + in + getNameArityToTptp isTptpConstName isTptpFnName fnTo fa + end; + +fun relToTptp mapping ra = + let + val Mapping {relTo,...} = mapping + in + getNameArityToTptp isTptpPropName isTptpRelName relTo ra + end; + +fun varFromTptp (_ : mapping) v = getVarFromTptp v; + +fun fnFromTptp mapping fa = + let + val Mapping {fnFrom,...} = mapping + in + getNameFromTptp fnFrom fa + end; + +fun relFromTptp mapping ra = + let + val Mapping {relFrom,...} = mapping + in + getNameFromTptp relFrom ra + end; + +val defaultMapping = + let + fun lift {name,arity,tptp} = + {name = Name.fromString name, arity = arity, tptp = tptp} + + val functionMapping = map lift defaultFunctionMapping + and relationMapping = map lift defaultRelationMapping + + val mapping = + {functionMapping = functionMapping, + relationMapping = relationMapping} + in + mkMapping mapping + end; + +(* ------------------------------------------------------------------------- *) +(* Interpreting TPTP functions and relations in a finite model. *) +(* ------------------------------------------------------------------------- *) + +fun mkFixedMap funcModel relModel = + let + fun mkEntry {name,arity,model} = ((Name.fromString name, arity), model) + + fun mkMap l = NameArityMap.fromList (map mkEntry l) + in + {functionMap = mkMap funcModel, + relationMap = mkMap relModel} + end; + +val defaultFixedMap = mkFixedMap defaultFunctionModel defaultRelationModel; + +val defaultModel = + let + val {size = N, fixed = fix} = Model.default + + val fix = Model.mapFixed defaultFixedMap fix + in + {size = N, fixed = fix} + end; + +local + fun toTptpMap toTptp = + let + fun add ((src,arity),dest,m) = + let + val src = Name.fromString (toTptp (src,arity)) + in + NameArityMap.insert m ((src,arity),dest) + end + in + fn m => NameArityMap.foldl add (NameArityMap.new ()) m + end; + + fun toTptpFixedMap mapping fixMap = + let + val {functionMap = fnMap, relationMap = relMap} = fixMap + + val fnMap = toTptpMap (fnToTptp mapping) fnMap + and relMap = toTptpMap (relToTptp mapping) relMap + in + {functionMap = fnMap, + relationMap = relMap} + end; +in + fun ppFixedMap mapping fixMap = + Model.ppFixedMap (toTptpFixedMap mapping fixMap); +end; + +(* ------------------------------------------------------------------------- *) +(* TPTP roles. *) +(* ------------------------------------------------------------------------- *) + +datatype role = + AxiomRole + | ConjectureRole + | DefinitionRole + | NegatedConjectureRole + | PlainRole + | TheoremRole + | OtherRole of string; + +fun isCnfConjectureRole role = + case role of + NegatedConjectureRole => true + | _ => false; + +fun isFofConjectureRole role = + case role of + ConjectureRole => true + | _ => false; + +fun toStringRole role = + case role of + AxiomRole => "axiom" + | ConjectureRole => "conjecture" + | DefinitionRole => "definition" + | NegatedConjectureRole => "negated_conjecture" + | PlainRole => "plain" + | TheoremRole => "theorem" + | OtherRole s => s; + +fun fromStringRole s = + case s of + "axiom" => AxiomRole + | "conjecture" => ConjectureRole + | "definition" => DefinitionRole + | "negated_conjecture" => NegatedConjectureRole + | "plain" => PlainRole + | "theorem" => TheoremRole + | _ => OtherRole s; + +val ppRole = Print.ppMap toStringRole Print.ppString; + +(* ------------------------------------------------------------------------- *) +(* SZS statuses. *) +(* ------------------------------------------------------------------------- *) + +datatype status = + CounterSatisfiableStatus + | TheoremStatus + | SatisfiableStatus + | UnknownStatus + | UnsatisfiableStatus; + +fun toStringStatus status = + case status of + CounterSatisfiableStatus => "CounterSatisfiable" + | TheoremStatus => "Theorem" + | SatisfiableStatus => "Satisfiable" + | UnknownStatus => "Unknown" + | UnsatisfiableStatus => "Unsatisfiable"; + +val ppStatus = Print.ppMap toStringStatus Print.ppString; (* ------------------------------------------------------------------------- *) (* TPTP literals. *) @@ -16701,14 +22401,29 @@ Boolean of bool | Literal of Literal.literal; -fun negate (Boolean b) = (Boolean (not b)) - | negate (Literal l) = (Literal (Literal.negate l)); - -fun literalFunctions (Boolean _) = NameAritySet.empty - | literalFunctions (Literal lit) = Literal.functions lit; - -fun literalRelation (Boolean _) = NONE - | literalRelation (Literal lit) = SOME (Literal.relation lit); +fun destLiteral lit = + case lit of + Literal l => l + | _ => raise Error "Tptp.destLiteral"; + +fun isBooleanLiteral lit = + case lit of + Boolean _ => true + | _ => false; + +fun equalBooleanLiteral b lit = + case lit of + Boolean b' => b = b' + | _ => false; + +fun negateLiteral (Boolean b) = (Boolean (not b)) + | negateLiteral (Literal l) = (Literal (Literal.negate l)); + +fun functionsLiteral (Boolean _) = NameAritySet.empty + | functionsLiteral (Literal lit) = Literal.functions lit; + +fun relationLiteral (Boolean _) = NONE + | relationLiteral (Literal lit) = SOME (Literal.relation lit); fun literalToFormula (Boolean true) = Formula.True | literalToFormula (Boolean false) = Formula.False @@ -16718,107 +22433,174 @@ | literalFromFormula Formula.False = Boolean false | literalFromFormula fm = Literal (Literal.fromFormula fm); -fun literalFreeVars (Boolean _) = NameSet.empty - | literalFreeVars (Literal lit) = Literal.freeVars lit; +fun freeVarsLiteral (Boolean _) = NameSet.empty + | freeVarsLiteral (Literal lit) = Literal.freeVars lit; fun literalSubst sub lit = case lit of Boolean _ => lit | Literal l => Literal (Literal.subst sub l); -fun mapLiteral maps lit = - case lit of - Boolean _ => lit - | Literal (p,a) => Literal (p, mapAtom maps a); - -fun destLiteral (Literal l) = l - | destLiteral _ = raise Error "destLiteral"; - (* ------------------------------------------------------------------------- *) (* Printing formulas using TPTP syntax. *) (* ------------------------------------------------------------------------- *) -val ppVar = Parser.ppString; - -local - fun term pp (Term.Var v) = ppVar pp v - | term pp (Term.Fn (c,[])) = Parser.addString pp c - | term pp (Term.Fn (f,tms)) = - (Parser.beginBlock pp Parser.Inconsistent 2; - Parser.addString pp (f ^ "("); - Parser.ppSequence "," term pp tms; - Parser.addString pp ")"; - Parser.endBlock pp); -in - fun ppTerm pp tm = - (Parser.beginBlock pp Parser.Inconsistent 0; - term pp tm; - Parser.endBlock pp); -end; - -fun ppAtom pp atm = ppTerm pp (Term.Fn atm); - -local - open Formula; - - fun fof pp (fm as And _) = assoc_binary pp ("&", stripConj fm) - | fof pp (fm as Or _) = assoc_binary pp ("|", stripDisj fm) - | fof pp (Imp a_b) = nonassoc_binary pp ("=>",a_b) - | fof pp (Iff a_b) = nonassoc_binary pp ("<=>",a_b) - | fof pp fm = unitary pp fm - - and nonassoc_binary pp (s,a_b) = - Parser.ppBinop (" " ^ s) unitary unitary pp a_b - - and assoc_binary pp (s,l) = Parser.ppSequence (" " ^ s) unitary pp l - - and unitary pp fm = - if isForall fm then quantified pp ("!", stripForall fm) - else if isExists fm then quantified pp ("?", stripExists fm) - else if atom pp fm then () - else if isNeg fm then - let - fun pr () = (Parser.addString pp "~"; Parser.addBreak pp (1,0)) - val (n,fm) = Formula.stripNeg fm - in - Parser.beginBlock pp Parser.Inconsistent 2; - funpow n pr (); - unitary pp fm; - Parser.endBlock pp - end - else - (Parser.beginBlock pp Parser.Inconsistent 1; - Parser.addString pp "("; - fof pp fm; - Parser.addString pp ")"; - Parser.endBlock pp) - - and quantified pp (q,(vs,fm)) = - (Parser.beginBlock pp Parser.Inconsistent 2; - Parser.addString pp (q ^ " "); - Parser.beginBlock pp Parser.Inconsistent (String.size q); - Parser.addString pp "["; - Parser.ppSequence "," ppVar pp vs; - Parser.addString pp "] :"; - Parser.endBlock pp; - Parser.addBreak pp (1,0); - unitary pp fm; - Parser.endBlock pp) - - and atom pp True = (Parser.addString pp "$true"; true) - | atom pp False = (Parser.addString pp "$false"; true) - | atom pp fm = - case total destEq fm of - SOME a_b => (Parser.ppBinop " =" ppTerm ppTerm pp a_b; true) - | NONE => - case total destNeq fm of - SOME a_b => (Parser.ppBinop " !=" ppTerm ppTerm pp a_b; true) - | NONE => case fm of Atom atm => (ppAtom pp atm; true) | _ => false; -in - fun ppFof pp fm = - (Parser.beginBlock pp Parser.Inconsistent 0; - fof pp fm; - Parser.endBlock pp); +fun ppVar mapping v = + let + val s = varToTptp mapping v + in + Print.addString s + end; + +fun ppFnName mapping fa = Print.addString (fnToTptp mapping fa); + +fun ppConst mapping c = ppFnName mapping (c,0); + +fun ppTerm mapping = + let + fun term tm = + case tm of + Term.Var v => ppVar mapping v + | Term.Fn (f,tms) => + case length tms of + 0 => ppConst mapping f + | a => + Print.blockProgram Print.Inconsistent 2 + [ppFnName mapping (f,a), + Print.addString "(", + Print.ppOpList "," term tms, + Print.addString ")"] + in + Print.block Print.Inconsistent 0 o term + end; + +fun ppRelName mapping ra = Print.addString (relToTptp mapping ra); + +fun ppProp mapping p = ppRelName mapping (p,0); + +fun ppAtom mapping (r,tms) = + case length tms of + 0 => ppProp mapping r + | a => + Print.blockProgram Print.Inconsistent 2 + [ppRelName mapping (r,a), + Print.addString "(", + Print.ppOpList "," (ppTerm mapping) tms, + Print.addString ")"]; + +local + val neg = Print.sequence (Print.addString "~") (Print.addBreak 1); + + fun fof mapping fm = + case fm of + Formula.And _ => assoc_binary mapping ("&", Formula.stripConj fm) + | Formula.Or _ => assoc_binary mapping ("|", Formula.stripDisj fm) + | Formula.Imp a_b => nonassoc_binary mapping ("=>",a_b) + | Formula.Iff a_b => nonassoc_binary mapping ("<=>",a_b) + | _ => unitary mapping fm + + and nonassoc_binary mapping (s,a_b) = + Print.ppOp2 (" " ^ s) (unitary mapping) (unitary mapping) a_b + + and assoc_binary mapping (s,l) = Print.ppOpList (" " ^ s) (unitary mapping) l + + and unitary mapping fm = + case fm of + Formula.True => Print.addString "$true" + | Formula.False => Print.addString "$false" + | Formula.Forall _ => quantified mapping ("!", Formula.stripForall fm) + | Formula.Exists _ => quantified mapping ("?", Formula.stripExists fm) + | Formula.Not _ => + (case total Formula.destNeq fm of + SOME a_b => Print.ppOp2 " !=" (ppTerm mapping) (ppTerm mapping) a_b + | NONE => + let + val (n,fm) = Formula.stripNeg fm + in + Print.blockProgram Print.Inconsistent 2 + [Print.duplicate n neg, + unitary mapping fm] + end) + | Formula.Atom atm => + (case total Formula.destEq fm of + SOME a_b => Print.ppOp2 " =" (ppTerm mapping) (ppTerm mapping) a_b + | NONE => ppAtom mapping atm) + | _ => + Print.blockProgram Print.Inconsistent 1 + [Print.addString "(", + fof mapping fm, + Print.addString ")"] + + and quantified mapping (q,(vs,fm)) = + let + val mapping = addVarListMapping mapping vs + in + Print.blockProgram Print.Inconsistent 2 + [Print.addString q, + Print.addString " ", + Print.blockProgram Print.Inconsistent (String.size q) + [Print.addString "[", + Print.ppOpList "," (ppVar mapping) vs, + Print.addString "] :"], + Print.addBreak 1, + unitary mapping fm] + end; +in + fun ppFof mapping fm = Print.block Print.Inconsistent 0 (fof mapping fm); +end; + +(* ------------------------------------------------------------------------- *) +(* Lexing TPTP files. *) +(* ------------------------------------------------------------------------- *) + +datatype token = + AlphaNum of string + | Punct of char + | Quote of string; + +fun isAlphaNum #"_" = true + | isAlphaNum c = Char.isAlphaNum c; + +local + open Parse; + + infixr 9 >>++ + infixr 8 ++ + infixr 7 >> + infixr 6 || + + val alphaNumToken = atLeastOne (some isAlphaNum) >> (AlphaNum o implode); + + val punctToken = + let + val punctChars = "<>=-*+/\\?@|!$%&#^:;~()[]{}.," + in + some (Char.contains punctChars) >> Punct + end; + + val quoteToken = + let + val escapeParser = + some (equal #"'") >> singleton || + some (equal #"\\") >> singleton + + fun stopOn #"'" = true + | stopOn #"\n" = true + | stopOn _ = false + + val quotedParser = + some (equal #"\\") ++ escapeParser >> op:: || + some (not o stopOn) >> singleton + in + exactChar #"'" ++ many quotedParser ++ exactChar #"'" >> + (fn (_,(l,_)) => Quote (implode (List.concat l))) + end; + + val lexToken = alphaNumToken || punctToken || quoteToken; + + val space = many (some Char.isSpace) >> K (); +in + val lexer = (space ++ lexToken ++ space) >> (fn ((),(tok,())) => tok); end; (* ------------------------------------------------------------------------- *) @@ -16829,7 +22611,7 @@ val clauseFunctions = let - fun funcs (lit,acc) = NameAritySet.union (literalFunctions lit) acc + fun funcs (lit,acc) = NameAritySet.union (functionsLiteral lit) acc in foldl funcs NameAritySet.empty end; @@ -16837,7 +22619,7 @@ val clauseRelations = let fun rels (lit,acc) = - case literalRelation lit of + case relationLiteral lit of NONE => acc | SOME r => NameAritySet.add acc r in @@ -16846,15 +22628,13 @@ val clauseFreeVars = let - fun fvs (lit,acc) = NameSet.union (literalFreeVars lit) acc + fun fvs (lit,acc) = NameSet.union (freeVarsLiteral lit) acc in foldl fvs NameSet.empty end; fun clauseSubst sub lits = map (literalSubst sub) lits; -fun mapClause maps lits = map (mapLiteral maps) lits; - fun clauseToFormula lits = Formula.listMkDisj (map literalToFormula lits); fun clauseFromFormula fm = map literalFromFormula (Formula.stripDisj fm); @@ -16865,115 +22645,475 @@ fun clauseFromThm th = clauseFromLiteralSet (Thm.clause th); -val ppClause = Parser.ppMap clauseToFormula ppFof; +fun ppClause mapping = Print.ppMap clauseToFormula (ppFof mapping); + +(* ------------------------------------------------------------------------- *) +(* TPTP formula names. *) +(* ------------------------------------------------------------------------- *) + +datatype formulaName = FormulaName of string; + +datatype formulaNameSet = FormulaNameSet of formulaName Set.set; + +fun compareFormulaName (FormulaName s1, FormulaName s2) = + String.compare (s1,s2); + +fun toTptpFormulaName (FormulaName s) = + getNameToTptp isTptpFormulaName s; + +val ppFormulaName = Print.ppMap toTptpFormulaName Print.ppString; + +val emptyFormulaNameSet = FormulaNameSet (Set.empty compareFormulaName); + +fun memberFormulaNameSet n (FormulaNameSet s) = Set.member n s; + +fun addFormulaNameSet (FormulaNameSet s) n = FormulaNameSet (Set.add s n); + +fun addListFormulaNameSet (FormulaNameSet s) l = + FormulaNameSet (Set.addList s l); + +(* ------------------------------------------------------------------------- *) +(* TPTP formula bodies. *) +(* ------------------------------------------------------------------------- *) + +datatype formulaBody = + CnfFormulaBody of literal list + | FofFormulaBody of Formula.formula; + +fun destCnfFormulaBody body = + case body of + CnfFormulaBody x => x + | _ => raise Error "destCnfFormulaBody"; + +val isCnfFormulaBody = can destCnfFormulaBody; + +fun destFofFormulaBody body = + case body of + FofFormulaBody x => x + | _ => raise Error "destFofFormulaBody"; + +val isFofFormulaBody = can destFofFormulaBody; + +fun formulaBodyFunctions body = + case body of + CnfFormulaBody cl => clauseFunctions cl + | FofFormulaBody fm => Formula.functions fm; + +fun formulaBodyRelations body = + case body of + CnfFormulaBody cl => clauseRelations cl + | FofFormulaBody fm => Formula.relations fm; + +fun formulaBodyFreeVars body = + case body of + CnfFormulaBody cl => clauseFreeVars cl + | FofFormulaBody fm => Formula.freeVars fm; + +fun ppFormulaBody mapping body = + case body of + CnfFormulaBody cl => ppClause mapping cl + | FofFormulaBody fm => ppFof mapping (Formula.generalize fm); + +(* ------------------------------------------------------------------------- *) +(* TPTP formula sources. *) +(* ------------------------------------------------------------------------- *) + +datatype formulaSource = + NoFormulaSource + | StripFormulaSource of + {inference : string, + parents : formulaName list} + | NormalizeFormulaSource of + {inference : Normalize.inference, + parents : formulaName list} + | ProofFormulaSource of + {inference : Proof.inference, + parents : formulaName list}; + +fun isNoFormulaSource source = + case source of + NoFormulaSource => true + | _ => false; + +fun functionsFormulaSource source = + case source of + NoFormulaSource => NameAritySet.empty + | StripFormulaSource _ => NameAritySet.empty + | NormalizeFormulaSource data => + let + val {inference = inf, parents = _} = data + in + case inf of + Normalize.Axiom fm => Formula.functions fm + | Normalize.Definition (_,fm) => Formula.functions fm + | _ => NameAritySet.empty + end + | ProofFormulaSource data => + let + val {inference = inf, parents = _} = data + in + case inf of + Proof.Axiom cl => LiteralSet.functions cl + | Proof.Assume atm => Atom.functions atm + | Proof.Subst (sub,_) => Subst.functions sub + | Proof.Resolve (atm,_,_) => Atom.functions atm + | Proof.Refl tm => Term.functions tm + | Proof.Equality (lit,_,tm) => + NameAritySet.union (Literal.functions lit) (Term.functions tm) + end; + +fun relationsFormulaSource source = + case source of + NoFormulaSource => NameAritySet.empty + | StripFormulaSource _ => NameAritySet.empty + | NormalizeFormulaSource data => + let + val {inference = inf, parents = _} = data + in + case inf of + Normalize.Axiom fm => Formula.relations fm + | Normalize.Definition (_,fm) => Formula.relations fm + | _ => NameAritySet.empty + end + | ProofFormulaSource data => + let + val {inference = inf, parents = _} = data + in + case inf of + Proof.Axiom cl => LiteralSet.relations cl + | Proof.Assume atm => NameAritySet.singleton (Atom.relation atm) + | Proof.Subst _ => NameAritySet.empty + | Proof.Resolve (atm,_,_) => NameAritySet.singleton (Atom.relation atm) + | Proof.Refl tm => NameAritySet.empty + | Proof.Equality (lit,_,_) => + NameAritySet.singleton (Literal.relation lit) + end; + +fun freeVarsFormulaSource source = + case source of + NoFormulaSource => NameSet.empty + | StripFormulaSource _ => NameSet.empty + | NormalizeFormulaSource data => NameSet.empty + | ProofFormulaSource data => + let + val {inference = inf, parents = _} = data + in + case inf of + Proof.Axiom cl => LiteralSet.freeVars cl + | Proof.Assume atm => Atom.freeVars atm + | Proof.Subst (sub,_) => Subst.freeVars sub + | Proof.Resolve (atm,_,_) => Atom.freeVars atm + | Proof.Refl tm => Term.freeVars tm + | Proof.Equality (lit,_,tm) => + NameSet.union (Literal.freeVars lit) (Term.freeVars tm) + end; + +local + val GEN_INFERENCE = "inference" + and GEN_INTRODUCED = "introduced"; + + fun nameStrip inf = inf; + + fun ppStrip mapping inf = Print.skip; + + fun nameNormalize inf = + case inf of + Normalize.Axiom _ => "canonicalize" + | Normalize.Definition _ => "canonicalize" + | Normalize.Simplify _ => "simplify" + | Normalize.Conjunct _ => "conjunct" + | Normalize.Specialize _ => "specialize" + | Normalize.Skolemize _ => "skolemize" + | Normalize.Clausify _ => "clausify"; + + fun ppNormalize mapping inf = Print.skip; + + fun nameProof inf = + case inf of + Proof.Axiom _ => "canonicalize" + | Proof.Assume _ => "assume" + | Proof.Subst _ => "subst" + | Proof.Resolve _ => "resolve" + | Proof.Refl _ => "refl" + | Proof.Equality _ => "equality"; + + local + fun ppTermInf mapping = ppTerm mapping; + + fun ppAtomInf mapping atm = + case total Atom.destEq atm of + SOME (a,b) => ppAtom mapping (Name.fromString "$equal", [a,b]) + | NONE => ppAtom mapping atm; + + fun ppLiteralInf mapping (pol,atm) = + Print.sequence + (if pol then Print.skip else Print.addString "~ ") + (ppAtomInf mapping atm); + in + fun ppProofTerm mapping = + Print.ppBracket "$fot(" ")" (ppTermInf mapping); + + fun ppProofAtom mapping = + Print.ppBracket "$cnf(" ")" (ppAtomInf mapping); + + fun ppProofLiteral mapping = + Print.ppBracket "$cnf(" ")" (ppLiteralInf mapping); + end; + + val ppProofVar = ppVar; + + val ppProofPath = Term.ppPath; + + fun ppProof mapping inf = + Print.blockProgram Print.Inconsistent 1 + [Print.addString "[", + (case inf of + Proof.Axiom _ => Print.skip + | Proof.Assume atm => ppProofAtom mapping atm + | Proof.Subst _ => Print.skip + | Proof.Resolve (atm,_,_) => ppProofAtom mapping atm + | Proof.Refl tm => ppProofTerm mapping tm + | Proof.Equality (lit,path,tm) => + Print.program + [ppProofLiteral mapping lit, + Print.addString ",", + Print.addBreak 1, + ppProofPath path, + Print.addString ",", + Print.addBreak 1, + ppProofTerm mapping tm]), + Print.addString "]"]; + + val ppParent = ppFormulaName; + + fun ppProofSubst mapping = + Print.ppMap Subst.toList + (Print.ppList + (Print.ppBracket "bind(" ")" + (Print.ppOp2 "," (ppProofVar mapping) + (ppProofTerm mapping)))); + + fun ppProofParent mapping (p,s) = + if Subst.null s then ppParent p + else Print.ppOp2 " :" ppParent (ppProofSubst mapping) (p,s); +in + fun ppFormulaSource mapping source = + case source of + NoFormulaSource => Print.skip + | StripFormulaSource {inference,parents} => + let + val gen = GEN_INFERENCE + + val name = nameStrip inference + in + Print.blockProgram Print.Inconsistent (size gen + 1) + [Print.addString gen, + Print.addString "(", + Print.addString name, + Print.addString ",", + Print.addBreak 1, + Print.ppBracket "[" "]" (ppStrip mapping) inference, + Print.addString ",", + Print.addBreak 1, + Print.ppList ppParent parents, + Print.addString ")"] + end + | NormalizeFormulaSource {inference,parents} => + let + val gen = GEN_INFERENCE + + val name = nameNormalize inference + in + Print.blockProgram Print.Inconsistent (size gen + 1) + [Print.addString gen, + Print.addString "(", + Print.addString name, + Print.addString ",", + Print.addBreak 1, + Print.ppBracket "[" "]" (ppNormalize mapping) inference, + Print.addString ",", + Print.addBreak 1, + Print.ppList ppParent parents, + Print.addString ")"] + end + | ProofFormulaSource {inference,parents} => + let + val isTaut = null parents + + val gen = if isTaut then GEN_INTRODUCED else GEN_INFERENCE + + val name = nameProof inference + + val parents = + let + val sub = + case inference of + Proof.Subst (s,_) => s + | _ => Subst.empty + in + map (fn parent => (parent,sub)) parents + end + in + Print.blockProgram Print.Inconsistent (size gen + 1) + ([Print.addString gen, + Print.addString "("] @ + (if isTaut then + [Print.addString "tautology", + Print.addString ",", + Print.addBreak 1, + Print.blockProgram Print.Inconsistent 1 + [Print.addString "[", + Print.addString name, + Print.addString ",", + Print.addBreak 1, + ppProof mapping inference, + Print.addString "]"]] + else + [Print.addString name, + Print.addString ",", + Print.addBreak 1, + ppProof mapping inference, + Print.addString ",", + Print.addBreak 1, + Print.ppList (ppProofParent mapping) parents]) @ + [Print.addString ")"]) + end +end; (* ------------------------------------------------------------------------- *) (* TPTP formulas. *) (* ------------------------------------------------------------------------- *) datatype formula = - CnfFormula of {name : string, role : string, clause : clause} - | FofFormula of {name : string, role : string, formula : Formula.formula}; - -fun destCnfFormula (CnfFormula x) = x - | destCnfFormula _ = raise Error "destCnfFormula"; + Formula of + {name : formulaName, + role : role, + body : formulaBody, + source : formulaSource}; + +fun nameFormula (Formula {name,...}) = name; + +fun roleFormula (Formula {role,...}) = role; + +fun bodyFormula (Formula {body,...}) = body; + +fun sourceFormula (Formula {source,...}) = source; + +fun destCnfFormula fm = destCnfFormulaBody (bodyFormula fm); val isCnfFormula = can destCnfFormula; -fun destFofFormula (FofFormula x) = x - | destFofFormula _ = raise Error "destFofFormula"; +fun destFofFormula fm = destFofFormulaBody (bodyFormula fm); val isFofFormula = can destFofFormula; -fun formulaFunctions (CnfFormula {clause,...}) = clauseFunctions clause - | formulaFunctions (FofFormula {formula,...}) = Formula.functions formula; - -fun formulaRelations (CnfFormula {clause,...}) = clauseRelations clause - | formulaRelations (FofFormula {formula,...}) = Formula.relations formula; - -fun formulaFreeVars (CnfFormula {clause,...}) = clauseFreeVars clause - | formulaFreeVars (FofFormula {formula,...}) = Formula.freeVars formula; - -fun mapFormula maps (CnfFormula {name,role,clause}) = - CnfFormula {name = name, role = role, clause = mapClause maps clause} - | mapFormula maps (FofFormula {name,role,formula}) = - FofFormula {name = name, role = role, formula = mapFof maps formula}; +fun functionsFormula fm = + let + val bodyFns = formulaBodyFunctions (bodyFormula fm) + and sourceFns = functionsFormulaSource (sourceFormula fm) + in + NameAritySet.union bodyFns sourceFns + end; + +fun relationsFormula fm = + let + val bodyRels = formulaBodyRelations (bodyFormula fm) + and sourceRels = relationsFormulaSource (sourceFormula fm) + in + NameAritySet.union bodyRels sourceRels + end; + +fun freeVarsFormula fm = + let + val bodyFvs = formulaBodyFreeVars (bodyFormula fm) + and sourceFvs = freeVarsFormulaSource (sourceFormula fm) + in + NameSet.union bodyFvs sourceFvs + end; + +val freeVarsListFormula = + let + fun add (fm,vs) = NameSet.union vs (freeVarsFormula fm) + in + List.foldl add NameSet.empty + end; val formulasFunctions = let - fun funcs (fm,acc) = NameAritySet.union (formulaFunctions fm) acc + fun funcs (fm,acc) = NameAritySet.union (functionsFormula fm) acc in foldl funcs NameAritySet.empty end; val formulasRelations = let - fun rels (fm,acc) = NameAritySet.union (formulaRelations fm) acc + fun rels (fm,acc) = NameAritySet.union (relationsFormula fm) acc in foldl rels NameAritySet.empty end; -fun formulaIsConjecture (CnfFormula {role,...}) = role = ROLE_NEGATED_CONJECTURE - | formulaIsConjecture (FofFormula {role,...}) = role = ROLE_CONJECTURE; - -local - open Parser; - +fun isCnfConjectureFormula fm = + case fm of + Formula {role, body = CnfFormulaBody _, ...} => isCnfConjectureRole role + | _ => false; + +fun isFofConjectureFormula fm = + case fm of + Formula {role, body = FofFormulaBody _, ...} => isFofConjectureRole role + | _ => false; + +fun isConjectureFormula fm = + isCnfConjectureFormula fm orelse + isFofConjectureFormula fm; + +(* Parsing and pretty-printing *) + +fun ppFormula mapping fm = + let + val Formula {name,role,body,source} = fm + + val gen = + case body of + CnfFormulaBody _ => "cnf" + | FofFormulaBody _ => "fof" + in + Print.blockProgram Print.Inconsistent (size gen + 1) + ([Print.addString gen, + Print.addString "(", + ppFormulaName name, + Print.addString ",", + Print.addBreak 1, + ppRole role, + Print.addString ",", + Print.addBreak 1, + Print.blockProgram Print.Consistent 1 + [Print.addString "(", + ppFormulaBody mapping body, + Print.addString ")"]] @ + (if isNoFormulaSource source then [] + else + [Print.addString ",", + Print.addBreak 1, + ppFormulaSource mapping source]) @ + [Print.addString ")."]) + end; + +fun formulaToString mapping = Print.toString (ppFormula mapping); + +local + open Parse; + + infixr 9 >>++ infixr 8 ++ infixr 7 >> infixr 6 || - datatype token = - AlphaNum of string - | Punct of char - | Quote of string; - - fun isAlphaNum #"_" = true - | isAlphaNum c = Char.isAlphaNum c; - - local - val alphaNumToken = atLeastOne (some isAlphaNum) >> (AlphaNum o implode); - - val punctToken = - let - val punctChars = "<>=-*+/\\?@|!$%&#^:;~()[]{}.," - in - some (Char.contains punctChars) >> Punct - end; - - val quoteToken = - let - val escapeParser = - exact #"'" >> singleton || - exact #"\\" >> singleton - - fun stopOn #"'" = true - | stopOn #"\n" = true - | stopOn _ = false - - val quotedParser = - exact #"\\" ++ escapeParser >> op:: || - some (not o stopOn) >> singleton - in - exact #"'" ++ many quotedParser ++ exact #"'" >> - (fn (_,(l,_)) => Quote (implode (List.concat l))) - end; - - val lexToken = alphaNumToken || punctToken || quoteToken; - - val space = many (some Char.isSpace) >> K (); - in - val lexer = (space ++ lexToken ++ space) >> (fn ((),(tok,())) => tok); - end; - fun someAlphaNum p = maybe (fn AlphaNum s => if p s then SOME s else NONE | _ => NONE); fun alphaNumParser s = someAlphaNum (equal s) >> K (); - fun isLower s = Char.isLower (String.sub (s,0)); - - val lowerParser = someAlphaNum isLower; + val lowerParser = someAlphaNum (fn s => Char.isLower (String.sub (s,0))); val upperParser = someAlphaNum (fn s => Char.isUpper (String.sub (s,0))); @@ -16986,12 +23126,7 @@ fun punctParser c = somePunct (equal c) >> K (); - fun quoteParser p = - let - fun q s = if p s then s else "'" ^ s ^ "'" - in - maybe (fn Quote s => SOME (q s) | _ => NONE) - end; + val quoteParser = maybe (fn Quote s => SOME s | _ => NONE); local fun f [] = raise Bug "symbolParser" @@ -17008,16 +23143,19 @@ punctParser #"$" ++ punctParser #"$" ++ someAlphaNum (K true) >> (fn ((),((),s)) => "$$" ^ s); - val nameParser = stringParser || numberParser || quoteParser (K false); - - val roleParser = lowerParser; + val nameParser = + (stringParser || numberParser || quoteParser) >> FormulaName; + + val roleParser = lowerParser >> fromStringRole; local fun isProposition s = isHdTlString Char.isLower isAlphaNum s; in val propositionParser = someAlphaNum isProposition || - definedParser || systemParser || quoteParser isProposition; + definedParser || + systemParser || + quoteParser; end; local @@ -17025,17 +23163,20 @@ in val functionParser = someAlphaNum isFunction || - definedParser || systemParser || quoteParser isFunction; + definedParser || + systemParser || + quoteParser; end; local - fun isConstant s = - isHdTlString Char.isLower isAlphaNum s orelse - isHdTlString Char.isDigit Char.isDigit s; + fun isConstant s = isHdTlString Char.isLower isAlphaNum s; in val constantParser = someAlphaNum isConstant || - definedParser || systemParser || quoteParser isConstant; + definedParser || + numberParser || + systemParser || + quoteParser; end; val varParser = upperParser; @@ -17046,536 +23187,819 @@ punctParser #"]") >> (fn ((),(h,(t,()))) => h :: t); - fun termParser input = - ((functionArgumentsParser >> Term.Fn) || - nonFunctionArgumentsTermParser) input - - and functionArgumentsParser input = - ((functionParser ++ punctParser #"(" ++ termParser ++ - many ((punctParser #"," ++ termParser) >> snd) ++ - punctParser #")") >> - (fn (f,((),(t,(ts,())))) => (f, t :: ts))) input - - and nonFunctionArgumentsTermParser input = - ((varParser >> Term.Var) || - (constantParser >> (fn n => Term.Fn (n,[])))) input - - val binaryAtomParser = - ((punctParser #"=" ++ termParser) >> - (fn ((),r) => fn l => Literal.mkEq (l,r))) || - ((symbolParser "!=" ++ termParser) >> - (fn ((),r) => fn l => Literal.mkNeq (l,r))); - - val maybeBinaryAtomParser = - optional binaryAtomParser >> - (fn SOME f => (fn a => f (Term.Fn a)) - | NONE => (fn a => (true,a))); - - val literalAtomParser = - ((functionArgumentsParser ++ maybeBinaryAtomParser) >> - (fn (a,f) => f a)) || - ((nonFunctionArgumentsTermParser ++ binaryAtomParser) >> - (fn (a,f) => f a)) || - (propositionParser >> - (fn n => (true,(n,[])))); - - val atomParser = - literalAtomParser >> - (fn (pol,("$true",[])) => Boolean pol - | (pol,("$false",[])) => Boolean (not pol) - | (pol,("$equal",[a,b])) => Literal (pol, Atom.mkEq (a,b)) - | lit => Literal lit); - - val literalParser = - ((punctParser #"~" ++ atomParser) >> (negate o snd)) || - atomParser; - - val disjunctionParser = - (literalParser ++ many ((punctParser #"|" ++ literalParser) >> snd)) >> - (fn (h,t) => h :: t); - - val clauseParser = - ((punctParser #"(" ++ disjunctionParser ++ punctParser #")") >> - (fn ((),(c,())) => c)) || - disjunctionParser; - -(* - An exact transcription of the fof_formula syntax from - - TPTP-v3.2.0/Documents/SyntaxBNF, - - fun fofFormulaParser input = - (binaryFormulaParser || unitaryFormulaParser) input - - and binaryFormulaParser input = - (nonAssocBinaryFormulaParser || assocBinaryFormulaParser) input - - and nonAssocBinaryFormulaParser input = - ((unitaryFormulaParser ++ binaryConnectiveParser ++ - unitaryFormulaParser) >> - (fn (f,(c,g)) => c (f,g))) input - - and binaryConnectiveParser input = - ((symbolParser "<=>" >> K Formula.Iff) || - (symbolParser "=>" >> K Formula.Imp) || - (symbolParser "<=" >> K (fn (f,g) => Formula.Imp (g,f))) || - (symbolParser "<~>" >> K (Formula.Not o Formula.Iff)) || - (symbolParser "~|" >> K (Formula.Not o Formula.Or)) || - (symbolParser "~&" >> K (Formula.Not o Formula.And))) input - - and assocBinaryFormulaParser input = - (orFormulaParser || andFormulaParser) input - - and orFormulaParser input = - ((unitaryFormulaParser ++ - atLeastOne ((punctParser #"|" ++ unitaryFormulaParser) >> snd)) >> - (fn (f,fs) => Formula.listMkDisj (f :: fs))) input - - and andFormulaParser input = - ((unitaryFormulaParser ++ - atLeastOne ((punctParser #"&" ++ unitaryFormulaParser) >> snd)) >> - (fn (f,fs) => Formula.listMkConj (f :: fs))) input - - and unitaryFormulaParser input = - (quantifiedFormulaParser || - unaryFormulaParser || - ((punctParser #"(" ++ fofFormulaParser ++ punctParser #")") >> - (fn ((),(f,())) => f)) || - (atomParser >> - (fn Boolean b => Formula.mkBoolean b - | Literal l => Literal.toFormula l))) input - - and quantifiedFormulaParser input = - ((quantifierParser ++ varListParser ++ punctParser #":" ++ - unitaryFormulaParser) >> - (fn (q,(v,((),f))) => q (v,f))) input - - and quantifierParser input = - ((punctParser #"!" >> K Formula.listMkForall) || - (punctParser #"?" >> K Formula.listMkExists)) input - - and unaryFormulaParser input = - ((unaryConnectiveParser ++ unitaryFormulaParser) >> - (fn (c,f) => c f)) input - - and unaryConnectiveParser input = - (punctParser #"~" >> K Formula.Not) input; -*) - -(* - This version is supposed to be equivalent to the spec version above, - but uses closures to avoid reparsing prefixes. -*) - - fun fofFormulaParser input = - ((unitaryFormulaParser ++ optional binaryFormulaParser) >> - (fn (f,NONE) => f | (f, SOME t) => t f)) input - - and binaryFormulaParser input = - (nonAssocBinaryFormulaParser || assocBinaryFormulaParser) input - - and nonAssocBinaryFormulaParser input = - ((binaryConnectiveParser ++ unitaryFormulaParser) >> - (fn (c,g) => fn f => c (f,g))) input - - and binaryConnectiveParser input = - ((symbolParser "<=>" >> K Formula.Iff) || - (symbolParser "=>" >> K Formula.Imp) || - (symbolParser "<=" >> K (fn (f,g) => Formula.Imp (g,f))) || - (symbolParser "<~>" >> K (Formula.Not o Formula.Iff)) || - (symbolParser "~|" >> K (Formula.Not o Formula.Or)) || - (symbolParser "~&" >> K (Formula.Not o Formula.And))) input - - and assocBinaryFormulaParser input = - (orFormulaParser || andFormulaParser) input - - and orFormulaParser input = - (atLeastOne ((punctParser #"|" ++ unitaryFormulaParser) >> snd) >> - (fn fs => fn f => Formula.listMkDisj (f :: fs))) input - - and andFormulaParser input = - (atLeastOne ((punctParser #"&" ++ unitaryFormulaParser) >> snd) >> - (fn fs => fn f => Formula.listMkConj (f :: fs))) input - - and unitaryFormulaParser input = - (quantifiedFormulaParser || - unaryFormulaParser || - ((punctParser #"(" ++ fofFormulaParser ++ punctParser #")") >> - (fn ((),(f,())) => f)) || - (atomParser >> - (fn Boolean b => Formula.mkBoolean b - | Literal l => Literal.toFormula l))) input - - and quantifiedFormulaParser input = - ((quantifierParser ++ varListParser ++ punctParser #":" ++ - unitaryFormulaParser) >> - (fn (q,(v,((),f))) => q (v,f))) input - - and quantifierParser input = - ((punctParser #"!" >> K Formula.listMkForall) || - (punctParser #"?" >> K Formula.listMkExists)) input - - and unaryFormulaParser input = - ((unaryConnectiveParser ++ unitaryFormulaParser) >> - (fn (c,f) => c f)) input + fun mkVarName mapping v = varFromTptp mapping v; + + fun mkVar mapping v = + let + val v = mkVarName mapping v + in + Term.Var v + end + + fun mkFn mapping (f,tms) = + let + val f = fnFromTptp mapping (f, length tms) + in + Term.Fn (f,tms) + end; + + fun mkConst mapping c = mkFn mapping (c,[]); + + fun mkAtom mapping (r,tms) = + let + val r = relFromTptp mapping (r, length tms) + in + (r,tms) + end; + + fun termParser mapping input = + let + val fnP = functionArgumentsParser mapping >> mkFn mapping + val nonFnP = nonFunctionArgumentsTermParser mapping + in + fnP || nonFnP + end input + + and functionArgumentsParser mapping input = + let + val commaTmP = (punctParser #"," ++ termParser mapping) >> snd + in + (functionParser ++ punctParser #"(" ++ termParser mapping ++ + many commaTmP ++ punctParser #")") >> + (fn (f,((),(t,(ts,())))) => (f, t :: ts)) + end input + + and nonFunctionArgumentsTermParser mapping input = + let + val varP = varParser >> mkVar mapping + val constP = constantParser >> mkConst mapping + in + varP || constP + end input; + + fun binaryAtomParser mapping tm input = + let + val eqP = + (punctParser #"=" ++ termParser mapping) >> + (fn ((),r) => (true,("$equal",[tm,r]))) + + val neqP = + (symbolParser "!=" ++ termParser mapping) >> + (fn ((),r) => (false,("$equal",[tm,r]))) + in + eqP || neqP + end input; + + fun maybeBinaryAtomParser mapping (s,tms) input = + let + val tm = mkFn mapping (s,tms) + in + optional (binaryAtomParser mapping tm) >> + (fn SOME lit => lit + | NONE => (true,(s,tms))) + end input; + + fun literalAtomParser mapping input = + let + val fnP = + functionArgumentsParser mapping >>++ + maybeBinaryAtomParser mapping + + val nonFnP = + nonFunctionArgumentsTermParser mapping >>++ + binaryAtomParser mapping + + val propP = propositionParser >> (fn s => (true,(s,[]))) + in + fnP || nonFnP || propP + end input; + + fun atomParser mapping input = + let + fun mk (pol,rel) = + case rel of + ("$true",[]) => Boolean pol + | ("$false",[]) => Boolean (not pol) + | ("$equal",[l,r]) => Literal (pol, Atom.mkEq (l,r)) + | (r,tms) => Literal (pol, mkAtom mapping (r,tms)) + in + literalAtomParser mapping >> mk + end input; + + fun literalParser mapping input = + let + val negP = + (punctParser #"~" ++ atomParser mapping) >> + (negateLiteral o snd) + + val posP = atomParser mapping + in + negP || posP + end input; + + fun disjunctionParser mapping input = + let + val orLitP = (punctParser #"|" ++ literalParser mapping) >> snd + in + (literalParser mapping ++ many orLitP) >> (fn (h,t) => h :: t) + end input; + + fun clauseParser mapping input = + let + val disjP = disjunctionParser mapping + + val bracketDisjP = + (punctParser #"(" ++ disjP ++ punctParser #")") >> + (fn ((),(c,())) => c) + in + bracketDisjP || disjP + end input; + + val binaryConnectiveParser = + (symbolParser "<=>" >> K Formula.Iff) || + (symbolParser "=>" >> K Formula.Imp) || + (symbolParser "<=" >> K (fn (f,g) => Formula.Imp (g,f))) || + (symbolParser "<~>" >> K (Formula.Not o Formula.Iff)) || + (symbolParser "~|" >> K (Formula.Not o Formula.Or)) || + (symbolParser "~&" >> K (Formula.Not o Formula.And)); + + val quantifierParser = + (punctParser #"!" >> K Formula.listMkForall) || + (punctParser #"?" >> K Formula.listMkExists); + + fun fofFormulaParser mapping input = + let + fun mk (f,NONE) = f + | mk (f, SOME t) = t f + in + (unitaryFormulaParser mapping ++ + optional (binaryFormulaParser mapping)) >> mk + end input + + and binaryFormulaParser mapping input = + let + val nonAssocP = nonAssocBinaryFormulaParser mapping + + val assocP = assocBinaryFormulaParser mapping + in + nonAssocP || assocP + end input + + and nonAssocBinaryFormulaParser mapping input = + let + fun mk (c,g) f = c (f,g) + in + (binaryConnectiveParser ++ unitaryFormulaParser mapping) >> mk + end input + + and assocBinaryFormulaParser mapping input = + let + val orP = orFormulaParser mapping + + val andP = andFormulaParser mapping + in + orP || andP + end input + + and orFormulaParser mapping input = + let + val orFmP = (punctParser #"|" ++ unitaryFormulaParser mapping) >> snd + in + atLeastOne orFmP >> + (fn fs => fn f => Formula.listMkDisj (f :: fs)) + end input + + and andFormulaParser mapping input = + let + val andFmP = (punctParser #"&" ++ unitaryFormulaParser mapping) >> snd + in + atLeastOne andFmP >> + (fn fs => fn f => Formula.listMkConj (f :: fs)) + end input + + and unitaryFormulaParser mapping input = + let + val quantP = quantifiedFormulaParser mapping + + val unaryP = unaryFormulaParser mapping + + val brackP = + (punctParser #"(" ++ fofFormulaParser mapping ++ + punctParser #")") >> + (fn ((),(f,())) => f) + + val atomP = + atomParser mapping >> + (fn Boolean b => Formula.mkBoolean b + | Literal l => Literal.toFormula l) + in + quantP || + unaryP || + brackP || + atomP + end input + + and quantifiedFormulaParser mapping input = + let + fun mk (q,(vs,((),f))) = q (map (mkVarName mapping) vs, f) + in + (quantifierParser ++ varListParser ++ punctParser #":" ++ + unitaryFormulaParser mapping) >> mk + end input + + and unaryFormulaParser mapping input = + let + fun mk (c,f) = c f + in + (unaryConnectiveParser ++ unitaryFormulaParser mapping) >> mk + end input and unaryConnectiveParser input = (punctParser #"~" >> K Formula.Not) input; - val cnfParser = - (alphaNumParser "cnf" ++ punctParser #"(" ++ - nameParser ++ punctParser #"," ++ - roleParser ++ punctParser #"," ++ - clauseParser ++ punctParser #")" ++ - punctParser #".") >> - (fn ((),((),(n,((),(r,((),(c,((),())))))))) => - CnfFormula {name = n, role = r, clause = c}); - - val fofParser = - (alphaNumParser "fof" ++ punctParser #"(" ++ - nameParser ++ punctParser #"," ++ - roleParser ++ punctParser #"," ++ - fofFormulaParser ++ punctParser #")" ++ - punctParser #".") >> - (fn ((),((),(n,((),(r,((),(f,((),())))))))) => - FofFormula {name = n, role = r, formula = f}); - - val formulaParser = cnfParser || fofParser; + fun cnfParser mapping input = + let + fun mk ((),((),(name,((),(role,((),(cl,((),())))))))) = + let + val body = CnfFormulaBody cl + val source = NoFormulaSource + in + Formula + {name = name, + role = role, + body = body, + source = source} + end + in + (alphaNumParser "cnf" ++ punctParser #"(" ++ + nameParser ++ punctParser #"," ++ + roleParser ++ punctParser #"," ++ + clauseParser mapping ++ punctParser #")" ++ + punctParser #".") >> mk + end input; + + fun fofParser mapping input = + let + fun mk ((),((),(name,((),(role,((),(fm,((),())))))))) = + let + val body = FofFormulaBody fm + val source = NoFormulaSource + in + Formula + {name = name, + role = role, + body = body, + source = source} + end + in + (alphaNumParser "fof" ++ punctParser #"(" ++ + nameParser ++ punctParser #"," ++ + roleParser ++ punctParser #"," ++ + fofFormulaParser mapping ++ punctParser #")" ++ + punctParser #".") >> mk + end input; +in + fun formulaParser mapping input = + let + val cnfP = cnfParser mapping + + val fofP = fofParser mapping + in + cnfP || fofP + end input; +end; + +(* ------------------------------------------------------------------------- *) +(* Include declarations. *) +(* ------------------------------------------------------------------------- *) + +fun ppInclude i = + Print.blockProgram Print.Inconsistent 2 + [Print.addString "include('", + Print.addString i, + Print.addString "')."]; + +val includeToString = Print.toString ppInclude; + +local + open Parse; + + infixr 9 >>++ + infixr 8 ++ + infixr 7 >> + infixr 6 || + + val filenameParser = maybe (fn Quote s => SOME s | _ => NONE); +in + val includeParser = + (some (equal (AlphaNum "include")) ++ + some (equal (Punct #"(")) ++ + filenameParser ++ + some (equal (Punct #")")) ++ + some (equal (Punct #"."))) >> + (fn (_,(_,(f,(_,_)))) => f); +end; + +(* ------------------------------------------------------------------------- *) +(* Parsing TPTP files. *) +(* ------------------------------------------------------------------------- *) + +datatype declaration = + IncludeDeclaration of string + | FormulaDeclaration of formula; + +val partitionDeclarations = + let + fun part (d,(il,fl)) = + case d of + IncludeDeclaration i => (i :: il, fl) + | FormulaDeclaration f => (il, f :: fl) + in + fn l => List.foldl part ([],[]) (rev l) + end; + +local + open Parse; + + infixr 9 >>++ + infixr 8 ++ + infixr 7 >> + infixr 6 || + + fun declarationParser mapping = + (includeParser >> IncludeDeclaration) || + (formulaParser mapping >> FormulaDeclaration); fun parseChars parser chars = let - val tokens = Parser.everything (lexer >> singleton) chars - in - Parser.everything (parser >> singleton) tokens - end; - - fun canParseString parser s = - let - val chars = Stream.fromString s - in - case Stream.toList (parseChars parser chars) of - [_] => true - | _ => false - end - handle NoParse => false; -in - val parseFormula = parseChars formulaParser; - - val isTptpRelation = canParseString functionParser - and isTptpProposition = canParseString propositionParser - and isTptpFunction = canParseString functionParser - and isTptpConstant = canParseString constantParser; -end; - -fun formulaFromString s = - case Stream.toList (parseFormula (Stream.fromList (explode s))) of - [fm] => fm - | _ => raise Parser.NoParse; - -local - local - fun explodeAlpha s = List.filter Char.isAlpha (explode s); - in - fun normTptpName s n = - case explodeAlpha n of - [] => s - | c :: cs => implode (Char.toLower c :: cs); - - fun normTptpVar s n = - case explodeAlpha n of - [] => s - | c :: cs => implode (Char.toUpper c :: cs); - end; - - fun normTptpFunc (n,0) = if isTptpConstant n then n else normTptpName "c" n - | normTptpFunc (n,_) = if isTptpFunction n then n else normTptpName "f" n; - - fun normTptpRel (n,0) = if isTptpProposition n then n else normTptpName "p" n - | normTptpRel (n,_) = if isTptpRelation n then n else normTptpName "r" n; - - fun mkMap set norm mapping = - let - val mapping = mappingToTptp mapping - - fun mk (n_r,(a,m)) = - case NameArityMap.peek mapping n_r of - SOME t => (a, NameArityMap.insert m (n_r,t)) - | NONE => - let - val t = norm n_r - val (n,_) = n_r - val t = if t = n then n else Term.variantNum a t - in - (NameSet.add a t, NameArityMap.insert m (n_r,t)) - end - - val avoid = - let - fun mk ((n,r),s) = - let - val n = Option.getOpt (NameArityMap.peek mapping (n,r), n) - in - NameSet.add s n - end - in - NameAritySet.foldl mk NameSet.empty set - end - in - snd (NameAritySet.foldl mk (avoid, NameArityMap.new ()) set) - end; - - fun mkTptpVar a v = Term.variantNum a (normTptpVar "V" v); - - fun isTptpVar v = mkTptpVar NameSet.empty v = v; - - fun alphaFormula fm = - let - fun addVar v a s = - let - val v' = mkTptpVar a v - val a = NameSet.add a v' - and s = if v = v' then s else Subst.insert s (v, Term.Var v') - in - (v',(a,s)) - end - - fun initVar (v,(a,s)) = snd (addVar v a s) - - open Formula - - fun alpha _ _ True = True - | alpha _ _ False = False - | alpha _ s (Atom atm) = Atom (Atom.subst s atm) - | alpha a s (Not p) = Not (alpha a s p) - | alpha a s (And (p,q)) = And (alpha a s p, alpha a s q) - | alpha a s (Or (p,q)) = Or (alpha a s p, alpha a s q) - | alpha a s (Imp (p,q)) = Imp (alpha a s p, alpha a s q) - | alpha a s (Iff (p,q)) = Iff (alpha a s p, alpha a s q) - | alpha a s (Forall (v,p)) = - let val (v,(a,s)) = addVar v a s in Forall (v, alpha a s p) end - | alpha a s (Exists (v,p)) = - let val (v,(a,s)) = addVar v a s in Exists (v, alpha a s p) end - - val fvs = formulaFreeVars fm - val (avoid,fvs) = NameSet.partition isTptpVar fvs - val (avoid,sub) = NameSet.foldl initVar (avoid,Subst.empty) fvs -(*TRACE5 - val () = Parser.ppTrace Subst.pp "Tptp.alpha: sub" sub -*) - in - case fm of - CnfFormula {name,role,clause} => - CnfFormula {name = name, role = role, clause = clauseSubst sub clause} - | FofFormula {name,role,formula} => - FofFormula {name = name, role = role, formula = alpha avoid sub formula} - end; - - fun formulaToTptp maps fm = alphaFormula (mapFormula maps fm); -in - fun formulasToTptp formulas = - let - val funcs = formulasFunctions formulas - and rels = formulasRelations formulas - - val functionMap = mkMap funcs normTptpFunc (!functionMapping) - and relationMap = mkMap rels normTptpRel (!relationMapping) - - val maps = {functionMap = functionMap, relationMap = relationMap} - in - map (formulaToTptp maps) formulas - end; -end; - -fun formulasFromTptp formulas = - let - val functionMap = mappingFromTptp (!functionMapping) - and relationMap = mappingFromTptp (!relationMapping) - - val maps = {functionMap = functionMap, relationMap = relationMap} - in - map (mapFormula maps) formulas - end; - -local - fun ppGen ppX pp (gen,name,role,x) = - (Parser.beginBlock pp Parser.Inconsistent (size gen + 1); - Parser.addString pp (gen ^ "(" ^ name ^ ","); - Parser.addBreak pp (1,0); - Parser.addString pp (role ^ ","); - Parser.addBreak pp (1,0); - Parser.beginBlock pp Parser.Consistent 1; - Parser.addString pp "("; - ppX pp x; - Parser.addString pp ")"; - Parser.endBlock pp; - Parser.addString pp ")."; - Parser.endBlock pp); -in - fun ppFormula pp (CnfFormula {name,role,clause}) = - ppGen ppClause pp ("cnf",name,role,clause) - | ppFormula pp (FofFormula {name,role,formula}) = - ppGen ppFof pp ("fof",name,role,formula); -end; - -val formulaToString = Parser.toString ppFormula; + val tokens = Parse.everything (lexer >> singleton) chars + in + Parse.everything (parser >> singleton) tokens + end; +in + fun parseDeclaration mapping = parseChars (declarationParser mapping); +end; + +(* ------------------------------------------------------------------------- *) +(* Clause information. *) +(* ------------------------------------------------------------------------- *) + +datatype clauseSource = + CnfClauseSource of formulaName * literal list + | FofClauseSource of Normalize.thm; + +type 'a clauseInfo = 'a LiteralSetMap.map; + +type clauseNames = formulaName clauseInfo; + +type clauseRoles = role clauseInfo; + +type clauseSources = clauseSource clauseInfo; + +val noClauseNames : clauseNames = LiteralSetMap.new (); + +val allClauseNames : clauseNames -> formulaNameSet = + let + fun add (_,n,s) = addFormulaNameSet s n + in + LiteralSetMap.foldl add emptyFormulaNameSet + end; + +val noClauseRoles : clauseRoles = LiteralSetMap.new (); + +val noClauseSources : clauseSources = LiteralSetMap.new (); + +(* ------------------------------------------------------------------------- *) +(* Comments. *) +(* ------------------------------------------------------------------------- *) + +fun mkLineComment "" = "%" + | mkLineComment line = "% " ^ line; + +fun destLineComment cs = + case cs of + [] => "" + | #"%" :: #" " :: rest => implode rest + | #"%" :: rest => implode rest + | _ => raise Error "Tptp.destLineComment"; + +val isLineComment = can destLineComment; (* ------------------------------------------------------------------------- *) (* TPTP problems. *) (* ------------------------------------------------------------------------- *) -datatype goal = - Cnf of Problem.problem - | Fof of Formula.formula; - -type problem = {comments : string list, formulas : formula list}; - -local - fun stripComments acc strm = +type comments = string list; + +type includes = string list; + +datatype problem = + Problem of + {comments : comments, + includes : includes, + formulas : formula list}; + +fun hasCnfConjecture (Problem {formulas,...}) = + List.exists isCnfConjectureFormula formulas; + +fun hasFofConjecture (Problem {formulas,...}) = + List.exists isFofConjectureFormula formulas; + +fun hasConjecture (Problem {formulas,...}) = + List.exists isConjectureFormula formulas; + +fun freeVars (Problem {formulas,...}) = freeVarsListFormula formulas; + +local + fun bump n avoid = + let + val s = FormulaName (Int.toString n) + in + if memberFormulaNameSet s avoid then bump (n + 1) avoid + else (s, n, addFormulaNameSet avoid s) + end; + + fun fromClause defaultRole names roles cl (n,avoid) = + let + val (name,n,avoid) = + case LiteralSetMap.peek names cl of + SOME name => (name,n,avoid) + | NONE => bump n avoid + + val role = Option.getOpt (LiteralSetMap.peek roles cl, defaultRole) + + val body = CnfFormulaBody (clauseFromLiteralSet cl) + + val source = NoFormulaSource + + val formula = + Formula + {name = name, + role = role, + body = body, + source = source} + in + (formula,(n,avoid)) + end; +in + fun mkProblem {comments,includes,names,roles,problem} = + let + fun fromCl defaultRole = fromClause defaultRole names roles + + val {axioms,conjecture} = problem + + val n_avoid = (0, allClauseNames names) + + val (axiomFormulas,n_avoid) = maps (fromCl AxiomRole) axioms n_avoid + + val (conjectureFormulas,_) = + maps (fromCl NegatedConjectureRole) conjecture n_avoid + + val formulas = axiomFormulas @ conjectureFormulas + in + Problem + {comments = comments, + includes = includes, + formulas = formulas} + end; +end; + +type normalization = + {problem : Problem.problem, + sources : clauseSources}; + +val initialNormalization : normalization = + {problem = {axioms = [], conjecture = []}, + sources = noClauseSources}; + +datatype problemGoal = + NoGoal + | CnfGoal of (formulaName * clause) list + | FofGoal of (formulaName * Formula.formula) list; + +local + fun partitionFormula (formula,(cnfAxioms,fofAxioms,cnfGoals,fofGoals)) = + let + val Formula {name,role,body,...} = formula + in + case body of + CnfFormulaBody cl => + if isCnfConjectureRole role then + let + val cnfGoals = (name,cl) :: cnfGoals + in + (cnfAxioms,fofAxioms,cnfGoals,fofGoals) + end + else + let + val cnfAxioms = (name,cl) :: cnfAxioms + in + (cnfAxioms,fofAxioms,cnfGoals,fofGoals) + end + | FofFormulaBody fm => + if isFofConjectureRole role then + let + val fofGoals = (name,fm) :: fofGoals + in + (cnfAxioms,fofAxioms,cnfGoals,fofGoals) + end + else + let + val fofAxioms = (name,fm) :: fofAxioms + in + (cnfAxioms,fofAxioms,cnfGoals,fofGoals) + end + end; + + fun partitionFormulas fms = + let + val (cnfAxioms,fofAxioms,cnfGoals,fofGoals) = + List.foldl partitionFormula ([],[],[],[]) fms + + val goal = + case (rev cnfGoals, rev fofGoals) of + ([],[]) => NoGoal + | (cnfGoals,[]) => CnfGoal cnfGoals + | ([],fofGoals) => FofGoal fofGoals + | (_ :: _, _ :: _) => + raise Error "TPTP problem has both cnf and fof conjecture formulas" + in + {cnfAxioms = rev cnfAxioms, + fofAxioms = rev fofAxioms, + goal = goal} + end; + + fun addClauses role clauses acc : normalization = + let + fun addClause (cl_src,sources) = + LiteralSetMap.insert sources cl_src + + val {problem,sources} : normalization = acc + val {axioms,conjecture} = problem + + val cls = map fst clauses + val (axioms,conjecture) = + if isCnfConjectureRole role then (axioms, cls @ conjecture) + else (cls @ axioms, conjecture) + + val problem = {axioms = axioms, conjecture = conjecture} + and sources = List.foldl addClause sources clauses + in + {problem = problem, + sources = sources} + end; + + fun addCnf role ((name,clause),(norm,cnf)) = + if List.exists (equalBooleanLiteral true) clause then (norm,cnf) + else + let + val cl = List.mapPartial (total destLiteral) clause + val cl = LiteralSet.fromList cl + + val src = CnfClauseSource (name,clause) + + val norm = addClauses role [(cl,src)] norm + in + (norm,cnf) + end; + + val addCnfAxiom = addCnf AxiomRole; + + val addCnfGoal = addCnf NegatedConjectureRole; + + fun addFof role (th,(norm,cnf)) = + let + fun sourcify (cl,inf) = (cl, FofClauseSource inf) + + val (clauses,cnf) = Normalize.addCnf th cnf + val clauses = map sourcify clauses + val norm = addClauses role clauses norm + in + (norm,cnf) + end; + + fun addFofAxiom ((_,fm),acc) = + addFof AxiomRole (Normalize.mkAxiom fm, acc); + + fun normProblem subgoal (norm,_) = + let + val {problem,sources} = norm + val {axioms,conjecture} = problem + val problem = {axioms = rev axioms, conjecture = rev conjecture} + in + {subgoal = subgoal, + problem = problem, + sources = sources} + end; + + val normProblemFalse = normProblem (Formula.False,[]); + + fun splitProblem acc = + let + fun mk parents subgoal = + let + val subgoal = Formula.generalize subgoal + + val th = Normalize.mkAxiom (Formula.Not subgoal) + + val acc = addFof NegatedConjectureRole (th,acc) + in + normProblem (subgoal,parents) acc + end + + fun split (name,goal) = + let + val subgoals = Formula.splitGoal goal + val subgoals = + if null subgoals then [Formula.True] else subgoals + + val parents = [name] + in + map (mk parents) subgoals + end + in + fn goals => List.concat (map split goals) + end; + + fun clausesToGoal cls = + let + val cls = map (Formula.generalize o clauseToFormula o snd) cls + in + Formula.listMkConj cls + end; + + fun formulasToGoal fms = + let + val fms = map (Formula.generalize o snd) fms + in + Formula.listMkConj fms + end; +in + fun goal (Problem {formulas,...}) = + let + val {cnfAxioms,fofAxioms,goal} = partitionFormulas formulas + + val fm = + case goal of + NoGoal => Formula.False + | CnfGoal cls => Formula.Imp (clausesToGoal cls, Formula.False) + | FofGoal goals => formulasToGoal goals + + val fm = + if null fofAxioms then fm + else Formula.Imp (formulasToGoal fofAxioms, fm) + + val fm = + if null cnfAxioms then fm + else Formula.Imp (clausesToGoal cnfAxioms, fm) + in + fm + end; + + fun normalize (Problem {formulas,...}) = + let + val {cnfAxioms,fofAxioms,goal} = partitionFormulas formulas + + val acc = (initialNormalization, Normalize.initialCnf) + val acc = List.foldl addCnfAxiom acc cnfAxioms + val acc = List.foldl addFofAxiom acc fofAxioms + in + case goal of + NoGoal => [normProblemFalse acc] + | CnfGoal cls => [normProblemFalse (List.foldl addCnfGoal acc cls)] + | FofGoal goals => splitProblem acc goals + end; +end; + +local + datatype blockComment = + OutsideBlockComment + | EnteringBlockComment + | InsideBlockComment + | LeavingBlockComment; + + fun stripLineComments acc strm = case strm of - Stream.NIL => (rev acc, Stream.NIL) - | Stream.CONS (line,rest) => - case total destComment line of - SOME s => stripComments (s :: acc) (rest ()) - | NONE => (rev acc, Stream.filter (not o isComment) strm); -in - fun read {filename} = - let + Stream.Nil => (rev acc, Stream.Nil) + | Stream.Cons (line,rest) => + case total destLineComment line of + SOME s => stripLineComments (s :: acc) (rest ()) + | NONE => (rev acc, Stream.filter (not o isLineComment) strm); + + fun advanceBlockComment c state = + case state of + OutsideBlockComment => + if c = #"/" then (Stream.Nil, EnteringBlockComment) + else (Stream.singleton c, OutsideBlockComment) + | EnteringBlockComment => + if c = #"*" then (Stream.Nil, InsideBlockComment) + else if c = #"/" then (Stream.singleton #"/", EnteringBlockComment) + else (Stream.fromList [#"/",c], OutsideBlockComment) + | InsideBlockComment => + if c = #"*" then (Stream.Nil, LeavingBlockComment) + else (Stream.Nil, InsideBlockComment) + | LeavingBlockComment => + if c = #"/" then (Stream.Nil, OutsideBlockComment) + else if c = #"*" then (Stream.Nil, LeavingBlockComment) + else (Stream.Nil, InsideBlockComment); + + fun eofBlockComment state = + case state of + OutsideBlockComment => Stream.Nil + | EnteringBlockComment => Stream.singleton #"/" + | _ => raise Error "EOF inside a block comment"; + + val stripBlockComments = + Stream.mapsConcat advanceBlockComment eofBlockComment + OutsideBlockComment; +in + fun read {mapping,filename} = + let + (* Estimating parse error line numbers *) + val lines = Stream.fromTextFile {filename = filename} - val lines = Stream.map chomp lines - - val (comments,lines) = stripComments [] lines - - val chars = Stream.concat (Stream.map Stream.fromString lines) - - val formulas = Stream.toList (parseFormula chars) - - val formulas = formulasFromTptp formulas - in - {comments = comments, formulas = formulas} - end; -end; - -(* Quick testing -installPP Term.pp; -installPP Formula.pp; -val () = Term.isVarName := (fn s => Char.isUpper (String.sub (s,0))); -val TPTP_DIR = "/Users/Joe/ptr/tptp/tptp/"; -val num1 = read {filename = TPTP_DIR ^ "NUM/NUM001-1.tptp"}; -val lcl9 = read {filename = TPTP_DIR ^ "LCL/LCL009-1.tptp"}; -val set11 = read {filename = TPTP_DIR ^ "SET/SET011+3.tptp"}; -val swc128 = read {filename = TPTP_DIR ^ "SWC/SWC128-1.tptp"}; -*) - -local - fun mkCommentLine comment = mkComment comment ^ "\n"; - - fun formulaStream [] () = Stream.NIL - | formulaStream (h :: t) () = - Stream.CONS ("\n" ^ formulaToString h, formulaStream t); -in - fun write {filename} {comments,formulas} = - Stream.toTextFile - {filename = filename} - (Stream.append - (Stream.map mkCommentLine (Stream.fromList comments)) - (formulaStream (formulasToTptp formulas))); -end; - -(* ------------------------------------------------------------------------- *) -(* Converting TPTP problems to goal formulas. *) -(* ------------------------------------------------------------------------- *) - -fun isCnfProblem ({formulas,...} : problem) = - let - val cnf = List.exists isCnfFormula formulas - and fof = List.exists isFofFormula formulas - in - case (cnf,fof) of - (false,false) => raise Error "TPTP problem has no formulas" - | (true,true) => raise Error "TPTP problem has both cnf and fof formulas" - | (cnf,_) => cnf - end; - -fun hasConjecture ({formulas,...} : problem) = - List.exists formulaIsConjecture formulas; - -local - fun cnfFormulaToClause (CnfFormula {clause,...}) = - if mem (Boolean true) clause then NONE - else - let - val lits = List.mapPartial (total destLiteral) clause - in - SOME (LiteralSet.fromList lits) - end - | cnfFormulaToClause (FofFormula _) = raise Bug "cnfFormulaToClause"; - - fun fofFormulaToGoal (FofFormula {formula,role,...}, {axioms,goals}) = - let - val fm = Formula.generalize formula - in - if role = ROLE_CONJECTURE then - {axioms = axioms, goals = fm :: goals} - else - {axioms = fm :: axioms, goals = goals} - end - | fofFormulaToGoal (CnfFormula _, _) = raise Bug "fofFormulaToGoal"; -in - fun toGoal (prob as {formulas,...}) = - if isCnfProblem prob then - Cnf (List.mapPartial cnfFormulaToClause formulas) - else - Fof - let - val axioms_goals = {axioms = [], goals = []} - val axioms_goals = List.foldl fofFormulaToGoal axioms_goals formulas - in - case axioms_goals of - {axioms, goals = []} => - Formula.Imp (Formula.listMkConj axioms, Formula.False) - | {axioms = [], goals} => Formula.listMkConj goals - | {axioms,goals} => - Formula.Imp (Formula.listMkConj axioms, Formula.listMkConj goals) - end; -end; - -local - fun fromClause cl n = - let - val name = "clause_" ^ Int.toString n - val role = ROLE_NEGATED_CONJECTURE - val clause = clauseFromLiteralSet cl - in - (CnfFormula {name = name, role = role, clause = clause}, n + 1) - end; -in - fun fromProblem prob = - let - val comments = [] - and (formulas,_) = maps fromClause prob 0 - in - {comments = comments, formulas = formulas} - end; -end; - -local - fun refute cls = - let - val res = Resolution.new Resolution.default (map Thm.axiom cls) - in - case Resolution.loop res of + val {chars,parseErrorLocation} = Parse.initialize {lines = lines} + in + (let + (* The character stream *) + + val (comments,chars) = stripLineComments [] chars + + val chars = Parse.everything Parse.any chars + + val chars = stripBlockComments chars + + (* The declaration stream *) + + val declarations = Stream.toList (parseDeclaration mapping chars) + + val (includes,formulas) = partitionDeclarations declarations + in + Problem + {comments = comments, + includes = includes, + formulas = formulas} + end + handle Parse.NoParse => raise Error "parse error") + handle Error err => + raise Error ("error in TPTP file \"" ^ filename ^ "\" " ^ + parseErrorLocation () ^ "\n" ^ err) + end; +end; + +local + val newline = Stream.singleton "\n"; + + fun spacer top = if top then Stream.Nil else newline; + + fun mkComment comment = mkLineComment comment ^ "\n"; + + fun mkInclude inc = includeToString inc ^ "\n"; + + fun formulaStream _ _ [] = Stream.Nil + | formulaStream mapping top (h :: t) = + Stream.append + (Stream.concatList + [spacer top, + Stream.singleton (formulaToString mapping h), + newline]) + (fn () => formulaStream mapping false t); +in + fun write {problem,mapping,filename} = + let + val Problem {comments,includes,formulas} = problem + + val includesTop = null comments + val formulasTop = includesTop andalso null includes + in + Stream.toTextFile + {filename = filename} + (Stream.concatList + [Stream.map mkComment (Stream.fromList comments), + spacer includesTop, + Stream.map mkInclude (Stream.fromList includes), + formulaStream mapping formulasTop formulas]) + end; +end; + +local + fun refute {axioms,conjecture} = + let + val axioms = map Thm.axiom axioms + and conjecture = map Thm.axiom conjecture + val problem = {axioms = axioms, conjecture = conjecture} + val resolution = Resolution.new Resolution.default problem + in + case Resolution.loop resolution of Resolution.Contradiction _ => true | Resolution.Satisfiable _ => false end; in fun prove filename = let - val tptp = read filename - val problems = - case toGoal tptp of - Cnf prob => [prob] - | Fof goal => Problem.fromGoal goal + val problem = read filename + val problems = map #problem (normalize problem) in List.all refute problems end; @@ -17586,133 +24010,385 @@ (* ------------------------------------------------------------------------- *) local - fun ppAtomInfo pp atm = - case total Atom.destEq atm of - SOME (a,b) => ppAtom pp ("$equal",[a,b]) - | NONE => ppAtom pp atm; - - fun ppLiteralInfo pp (pol,atm) = - if pol then ppAtomInfo pp atm - else - (Parser.beginBlock pp Parser.Inconsistent 2; - Parser.addString pp "~"; - Parser.addBreak pp (1,0); - ppAtomInfo pp atm; - Parser.endBlock pp); - - val ppAssumeInfo = Parser.ppBracket "(" ")" ppAtomInfo; - - val ppSubstInfo = - Parser.ppMap - Subst.toList - (Parser.ppSequence "," - (Parser.ppBracket "[" "]" - (Parser.ppBinop "," ppVar (Parser.ppBracket "(" ")" ppTerm)))); - - val ppResolveInfo = Parser.ppBracket "(" ")" ppAtomInfo; - - val ppReflInfo = Parser.ppBracket "(" ")" ppTerm; - - fun ppEqualityInfo pp (lit,path,res) = - (Parser.ppBracket "(" ")" ppLiteralInfo pp lit; - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Term.ppPath pp path; - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Parser.ppBracket "(" ")" ppTerm pp res); - - fun ppInfInfo pp inf = + fun newName avoid prefix = + let + fun bump i = + let + val name = FormulaName (prefix ^ Int.toString i) + val i = i + 1 + in + if memberFormulaNameSet name avoid then bump i else (name,i) + end + in + bump + end; + + fun lookupClauseSource sources cl = + case LiteralSetMap.peek sources cl of + SOME src => src + | NONE => raise Bug "Tptp.lookupClauseSource"; + + fun lookupFormulaName fmNames fm = + case FormulaMap.peek fmNames fm of + SOME name => name + | NONE => raise Bug "Tptp.lookupFormulaName"; + + fun lookupClauseName clNames cl = + case LiteralSetMap.peek clNames cl of + SOME name => name + | NONE => raise Bug "Tptp.lookupClauseName"; + + fun lookupClauseSourceName sources fmNames cl = + case lookupClauseSource sources cl of + CnfClauseSource (name,_) => name + | FofClauseSource th => + let + val (fm,_) = Normalize.destThm th + in + lookupFormulaName fmNames fm + end; + + fun collectProofDeps sources ((_,inf),names_ths) = + case inf of + Proof.Axiom cl => + let + val (names,ths) = names_ths + in + case lookupClauseSource sources cl of + CnfClauseSource (name,_) => + let + val names = addFormulaNameSet names name + in + (names,ths) + end + | FofClauseSource th => + let + val ths = th :: ths + in + (names,ths) + end + end + | _ => names_ths; + + fun collectNormalizeDeps ((_,inf,_),fofs_defs) = case inf of - Proof.Axiom _ => raise Bug "ppInfInfo" - | Proof.Assume atm => ppAssumeInfo pp atm - | Proof.Subst (sub,_) => ppSubstInfo pp sub - | Proof.Resolve (res,_,_) => ppResolveInfo pp res - | Proof.Refl tm => ppReflInfo pp tm - | Proof.Equality x => ppEqualityInfo pp x; -in - fun ppProof p prf = - let - fun thmString n = Int.toString n - - val prf = enumerate prf - - fun ppThm p th = - let - val cl = Thm.clause th - - fun pred (_,(th',_)) = LiteralSet.equal (Thm.clause th') cl - in - case List.find pred prf of - NONE => Parser.addString p "(?)" - | SOME (n,_) => Parser.addString p (thmString n) - end - - fun ppInf p inf = - let - val name = Thm.inferenceTypeToString (Proof.inferenceType inf) - val name = String.map Char.toLower name - in - Parser.addString p (name ^ ","); - Parser.addBreak p (1,0); - Parser.ppBracket "[" "]" ppInfInfo p inf; - case Proof.parents inf of - [] => () - | ths => - (Parser.addString p ","; - Parser.addBreak p (1,0); - Parser.ppList ppThm p ths) - end - - fun ppTaut p inf = - (Parser.addString p "tautology,"; - Parser.addBreak p (1,0); - Parser.ppBracket "[" "]" ppInf p inf) - - fun ppStepInfo p (n,(th,inf)) = - let - val is_axiom = case inf of Proof.Axiom _ => true | _ => false - val name = thmString n - val role = - if is_axiom then "axiom" - else if Thm.isContradiction th then "theorem" - else "plain" - val cl = clauseFromThm th - in - Parser.addString p (name ^ ","); - Parser.addBreak p (1,0); - Parser.addString p (role ^ ","); - Parser.addBreak p (1,0); - Parser.ppBracket "(" ")" ppClause p cl; - if is_axiom then () - else - let - val is_tautology = null (Proof.parents inf) - in - Parser.addString p ","; - Parser.addBreak p (1,0); - if is_tautology then - Parser.ppBracket "introduced(" ")" ppTaut p inf - else - Parser.ppBracket "inference(" ")" ppInf p inf - end - end - - fun ppStep p step = - (Parser.ppBracket "cnf(" ")" ppStepInfo p step; - Parser.addString p "."; - Parser.addNewline p) - in - Parser.beginBlock p Parser.Consistent 0; - app (ppStep p) prf; - Parser.endBlock p - end -(*DEBUG - handle Error err => raise Bug ("Tptp.ppProof: shouldn't fail:\n" ^ err); -*) -end; - -val proofToString = Parser.toString ppProof; + Normalize.Axiom fm => + let + val (fofs,defs) = fofs_defs + val fofs = FormulaSet.add fofs fm + in + (fofs,defs) + end + | Normalize.Definition n_d => + let + val (fofs,defs) = fofs_defs + val defs = StringMap.insert defs n_d + in + (fofs,defs) + end + | _ => fofs_defs; + + fun collectSubgoalProofDeps subgoalProof (names,fofs,defs) = + let + val {subgoal,sources,refutation} = subgoalProof + + val names = addListFormulaNameSet names (snd subgoal) + + val proof = Proof.proof refutation + + val (names,ths) = + List.foldl (collectProofDeps sources) (names,[]) proof + + val normalization = Normalize.proveThms (rev ths) + + val (fofs,defs) = + List.foldl collectNormalizeDeps (fofs,defs) normalization + + val subgoalProof = + {subgoal = subgoal, + normalization = normalization, + sources = sources, + proof = proof} + in + (subgoalProof,(names,fofs,defs)) + end; + + fun addProblemFormula names fofs (formula,(avoid,formulas,fmNames)) = + let + val name = nameFormula formula + + val avoid = addFormulaNameSet avoid name + + val (formulas,fmNames) = + if memberFormulaNameSet name names then + (formula :: formulas, fmNames) + else + case bodyFormula formula of + CnfFormulaBody _ => (formulas,fmNames) + | FofFormulaBody fm => + if not (FormulaSet.member fm fofs) then (formulas,fmNames) + else (formula :: formulas, FormulaMap.insert fmNames (fm,name)) + in + (avoid,formulas,fmNames) + end; + + fun addDefinitionFormula avoid (_,def,(formulas,i,fmNames)) = + let + val (name,i) = newName avoid "definition_" i + + val role = DefinitionRole + + val body = FofFormulaBody def + + val source = NoFormulaSource + + val formula = + Formula + {name = name, + role = role, + body = body, + source = source} + + val formulas = formula :: formulas + + val fmNames = FormulaMap.insert fmNames (def,name) + in + (formulas,i,fmNames) + end; + + fun addSubgoalFormula avoid subgoalProof (formulas,i) = + let + val {subgoal,normalization,sources,proof} = subgoalProof + + val (fm,pars) = subgoal + + val (name,i) = newName avoid "subgoal_" i + + val number = i - 1 + + val (subgoal,formulas) = + if null pars then (NONE,formulas) + else + let + val role = PlainRole + + val body = FofFormulaBody fm + + val source = + StripFormulaSource + {inference = "strip", + parents = pars} + + val formula = + Formula + {name = name, + role = role, + body = body, + source = source} + in + (SOME (name,fm), formula :: formulas) + end + + val subgoalProof = + {number = number, + subgoal = subgoal, + normalization = normalization, + sources = sources, + proof = proof} + in + (subgoalProof,(formulas,i)) + end; + + fun mkNormalizeFormulaSource fmNames inference fms = + let + val fms = + case inference of + Normalize.Axiom fm => fm :: fms + | Normalize.Definition (_,fm) => fm :: fms + | _ => fms + + val parents = map (lookupFormulaName fmNames) fms + in + NormalizeFormulaSource + {inference = inference, + parents = parents} + end; + + fun mkProofFormulaSource sources fmNames clNames inference = + let + val parents = + case inference of + Proof.Axiom cl => [lookupClauseSourceName sources fmNames cl] + | _ => + let + val cls = map Thm.clause (Proof.parents inference) + in + map (lookupClauseName clNames) cls + end + in + ProofFormulaSource + {inference = inference, + parents = parents} + end; + + fun addNormalizeFormula avoid prefix ((fm,inf,fms),acc) = + let + val (formulas,i,fmNames) = acc + + val (name,i) = newName avoid prefix i + + val role = PlainRole + + val body = FofFormulaBody fm + + val source = mkNormalizeFormulaSource fmNames inf fms + + val formula = + Formula + {name = name, + role = role, + body = body, + source = source} + + val formulas = formula :: formulas + + val fmNames = FormulaMap.insert fmNames (fm,name) + in + (formulas,i,fmNames) + end; + + fun isSameClause sources formulas inf = + case inf of + Proof.Axiom cl => + (case lookupClauseSource sources cl of + CnfClauseSource (name,lits) => + if List.exists isBooleanLiteral lits then NONE + else SOME name + | _ => NONE) + | _ => NONE; + + fun addProofFormula avoid sources fmNames prefix ((th,inf),acc) = + let + val (formulas,i,clNames) = acc + + val cl = Thm.clause th + in + case isSameClause sources formulas inf of + SOME name => + let + val clNames = LiteralSetMap.insert clNames (cl,name) + in + (formulas,i,clNames) + end + | NONE => + let + val (name,i) = newName avoid prefix i + + val role = PlainRole + + val body = CnfFormulaBody (clauseFromLiteralSet cl) + + val source = mkProofFormulaSource sources fmNames clNames inf + + val formula = + Formula + {name = name, + role = role, + body = body, + source = source} + + val formulas = formula :: formulas + + val clNames = LiteralSetMap.insert clNames (cl,name) + in + (formulas,i,clNames) + end + end; + + fun addSubgoalProofFormulas avoid fmNames (subgoalProof,formulas) = + let + val {number,subgoal,normalization,sources,proof} = subgoalProof + + val (formulas,fmNames) = + case subgoal of + NONE => (formulas,fmNames) + | SOME (name,fm) => + let + val source = + StripFormulaSource + {inference = "negate", + parents = [name]} + + val prefix = "negate_" ^ Int.toString number ^ "_" + + val (name,_) = newName avoid prefix 0 + + val role = PlainRole + + val fm = Formula.Not fm + + val body = FofFormulaBody fm + + val formula = + Formula + {name = name, + role = role, + body = body, + source = source} + + val formulas = formula :: formulas + + val fmNames = FormulaMap.insert fmNames (fm,name) + in + (formulas,fmNames) + end + + val prefix = "normalize_" ^ Int.toString number ^ "_" + val (formulas,_,fmNames) = + List.foldl (addNormalizeFormula avoid prefix) + (formulas,0,fmNames) normalization + + val prefix = "refute_" ^ Int.toString number ^ "_" + val clNames : formulaName LiteralSetMap.map = LiteralSetMap.new () + val (formulas,_,_) = + List.foldl (addProofFormula avoid sources fmNames prefix) + (formulas,0,clNames) proof + in + formulas + end; +in + fun fromProof {problem,proofs} = + let + val names = emptyFormulaNameSet + and fofs = FormulaSet.empty + and defs : Formula.formula StringMap.map = StringMap.new () + + val (proofs,(names,fofs,defs)) = + maps collectSubgoalProofDeps proofs (names,fofs,defs) + + val Problem {formulas,...} = problem + + val fmNames : formulaName FormulaMap.map = FormulaMap.new () + val (avoid,formulas,fmNames) = + List.foldl (addProblemFormula names fofs) + (emptyFormulaNameSet,[],fmNames) formulas + + val (formulas,_,fmNames) = + StringMap.foldl (addDefinitionFormula avoid) + (formulas,0,fmNames) defs + + val (proofs,(formulas,_)) = + maps (addSubgoalFormula avoid) proofs (formulas,0) + + val formulas = + List.foldl (addSubgoalProofFormulas avoid fmNames) formulas proofs + in + rev formulas + end +(*MetisDebug + handle Error err => raise Bug ("Tptp.fromProof: shouldn't fail:\n" ^ err); +*) +end; end end; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/scripts/mlpp --- a/src/Tools/Metis/scripts/mlpp Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/scripts/mlpp Mon Sep 13 21:24:10 2010 +0200 @@ -18,8 +18,8 @@ } if (!$opt_c) { die "mlpp: you must specify the SML compiler\n"; } -if ($opt_c ne "mosml" && $opt_c ne "mlton" && $opt_c ne "isabelle") { - die "mlpp: the SML compiler must be one of {mosml,mlton,isabelle}.\n"; +if ($opt_c ne "mosml" && $opt_c ne "mlton" && $opt_c ne "polyml") { + die "mlpp: the SML compiler must be one of {mosml,mlton,polyml}.\n"; } # Autoflush STDIN @@ -72,8 +72,17 @@ my $text = shift @_; if ($opt_c eq "mosml") { + $text =~ s/Array\.copy/Array_copy/g; + $text =~ s/Array\.foldli/Array_foldli/g; + $text =~ s/Array\.foldri/Array_foldri/g; + $text =~ s/Array\.modifyi/Array_modifyi/g; + $text =~ s/OS\.Process\.isSuccess/OS_Process_isSuccess/g; $text =~ s/PP\.ppstream/ppstream/g; + $text =~ s/String\.isSuffix/String_isSuffix/g; + $text =~ s/Substring\.full/Substring_full/g; $text =~ s/TextIO\.inputLine/TextIO_inputLine/g; + $text =~ s/Vector\.foldli/Vector_foldli/g; + $text =~ s/Vector\.mapi/Vector_mapi/g; } print STDOUT $text; @@ -88,7 +97,8 @@ print STDOUT "(*#line 0.0 \"$filename\"*)\n"; } - open my $INPUT, "$filename"; + open my $INPUT, "$filename" or + die "mlpp: couldn't open $filename: $!\n"; my $state = "normal"; my $comment = 0; @@ -173,7 +183,7 @@ my $pat = $2; $line = $3; print_normal $leadup; - + if ($pat eq "`") { $state = "quote"; } @@ -197,7 +207,7 @@ } elsif ($pat eq "*)") { if ($revealed_comment == 0) { - die "Too many comment closers.\n" + die "mlpp: too many comment closers.\n" } --$revealed_comment; } @@ -223,13 +233,13 @@ } if ($state eq "quote") { - die "EOF inside \` quote\n"; + die "mlpp: EOF inside \` quote\n"; } elsif ($state eq "dquote") { - die "EOF inside \" quote\n"; + die "mlpp: EOF inside \" quote\n"; } elsif ($state eq "comment") { - die "EOF inside comment\n"; + die "mlpp: EOF inside comment\n"; } else { ($state eq "normal") or die; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Active.sig --- a/src/Tools/Metis/src/Active.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Active.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* THE ACTIVE SET OF CLAUSES *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Active = @@ -10,7 +10,10 @@ (* A type of active clause sets. *) (* ------------------------------------------------------------------------- *) -type simplify = {subsume : bool, reduce : bool, rewrite : bool} +type simplify = + {subsume : bool, + reduce : bool, + rewrite : bool} type parameters = {clause : Clause.parameters, @@ -27,13 +30,15 @@ val size : active -> int -val saturated : active -> Clause.clause list +val saturation : active -> Clause.clause list (* ------------------------------------------------------------------------- *) (* Create a new active clause set and initialize clauses. *) (* ------------------------------------------------------------------------- *) -val new : parameters -> Thm.thm list -> active * Clause.clause list +val new : + parameters -> {axioms : Thm.thm list, conjecture : Thm.thm list} -> + active * {axioms : Clause.clause list, conjecture : Clause.clause list} (* ------------------------------------------------------------------------- *) (* Add a clause into the active set and deduce all consequences. *) @@ -45,6 +50,6 @@ (* Pretty printing. *) (* ------------------------------------------------------------------------- *) -val pp : active Parser.pp +val pp : active Print.pp end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Active.sml --- a/src/Tools/Metis/src/Active.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Active.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* THE ACTIVE SET OF CLAUSES *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Active :> Active = @@ -12,10 +12,32 @@ (* Helper functions. *) (* ------------------------------------------------------------------------- *) +(*MetisDebug local + fun mkRewrite ordering = + let + fun add (cl,rw) = + let + val {id, thm = th, ...} = Clause.dest cl + in + case total Thm.destUnitEq th of + SOME l_r => Rewrite.add rw (id,(l_r,th)) + | NONE => rw + end + in + foldl add (Rewrite.new (KnuthBendixOrder.compare ordering)) + end; + fun allFactors red = let - fun allClause cl = List.all red (cl :: Clause.factor cl) + fun allClause cl = + List.all red (cl :: Clause.factor cl) orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.allFactors: cl" cl + in + false + end in List.all allClause end; @@ -30,6 +52,12 @@ | SOME cl => allFactors red [cl] in LiteralSet.all allLiteral2 (Clause.literals cl) + end orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.allResolutions: cl2" cl + in + false end fun allClause1 allCls cl = @@ -39,7 +67,14 @@ fun allLiteral1 lit = List.all (allClause2 (cl,lit)) allCls in LiteralSet.all allLiteral1 (Clause.literals cl) + end orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.allResolutions: cl1" cl + in + false end + in fn [] => true | allCls as cl :: cls => @@ -60,9 +95,24 @@ | SOME cl => allFactors red [cl] in List.all allSubterms (Literal.nonVarTypedSubterms lit) + end orelse + let + val () = Print.trace Literal.pp + "Active.isSaturated.allParamodulations: lit2" lit + in + false end in LiteralSet.all allLiteral2 (Clause.literals cl) + end orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.allParamodulations: cl2" cl + val (_,_,ort,_) = cl_lit_ort_tm + val () = Print.trace Rewrite.ppOrient + "Active.isSaturated.allParamodulations: ort1" ort + in + false end fun allClause1 cl = @@ -78,9 +128,21 @@ | SOME (l,r) => allCl2 (cl,lit,Rewrite.LeftToRight,l) andalso allCl2 (cl,lit,Rewrite.RightToLeft,r) + end orelse + let + val () = Print.trace Literal.pp + "Active.isSaturated.allParamodulations: lit1" lit + in + false end in LiteralSet.all allLiteral1 (Clause.literals cl) + end orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.allParamodulations: cl1" cl + in + false end in List.all allClause1 cls @@ -98,30 +160,49 @@ val cl' = Clause.reduce reduce cl' val cl' = Clause.rewrite rewrite cl' in - not (Clause.equalThms cl cl') andalso simp cl' + not (Clause.equalThms cl cl') andalso + (simp cl' orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.redundant: cl'" cl' + in + false + end) end in - simp + fn cl => + simp cl orelse + let + val () = Print.trace Clause.pp + "Active.isSaturated.redundant: cl" cl + in + false + end end; in fun isSaturated ordering subs cls = let -(*TRACE2 - val ppCls = Parser.ppList Clause.pp - val () = Parser.ppTrace ppCls "Active.checkSaturated: clauses" cls -*) - val red = Units.empty - val rw = Rewrite.new (KnuthBendixOrder.compare ordering) - val red = redundant {subsume = subs, reduce = red, rewrite = rw} + val rd = Units.empty + val rw = mkRewrite ordering cls + val red = redundant {subsume = subs, reduce = rd, rewrite = rw} in - allFactors red cls andalso - allResolutions red cls andalso - allParamodulations red cls + (allFactors red cls andalso + allResolutions red cls andalso + allParamodulations red cls) orelse + let + val () = Print.trace Rewrite.pp "Active.isSaturated: rw" rw + val () = Print.trace (Print.ppList Clause.pp) + "Active.isSaturated: clauses" cls + in + false + end end; +end; - fun checkSaturated ordering subs cls = - if isSaturated ordering subs cls then () else raise Bug "unsaturated"; -end; +fun checkSaturated ordering subs cls = + if isSaturated ordering subs cls then () + else raise Bug "Active.checkSaturated"; +*) (* ------------------------------------------------------------------------- *) (* A type of active clause sets. *) @@ -201,7 +282,7 @@ IntMap.foldr add [] cls end; -fun saturated active = +fun saturation active = let fun remove (cl,(cls,subs)) = let @@ -215,7 +296,7 @@ val (cls,_) = foldl remove ([], Subsume.new ()) cls val (cls,subs) = foldl remove ([], Subsume.new ()) cls -(*DEBUG +(*MetisDebug val Active {parameters,...} = active val {clause,...} = parameters val {ordering,...} = clause @@ -233,57 +314,53 @@ let fun toStr active = "Active{" ^ Int.toString (size active) ^ "}" in - Parser.ppMap toStr Parser.ppString + Print.ppMap toStr Print.ppString end; -(*DEBUG +(*MetisDebug local - open Parser; - - fun ppField f ppA p a = - (beginBlock p Inconsistent 2; - addString p (f ^ " ="); - addBreak p (1,0); - ppA p a; - endBlock p); + fun ppField f ppA a = + Print.blockProgram Print.Inconsistent 2 + [Print.addString (f ^ " ="), + Print.addBreak 1, + ppA a]; val ppClauses = ppField "clauses" - (Parser.ppMap IntMap.toList - (Parser.ppList (Parser.ppPair Parser.ppInt Clause.pp))); + (Print.ppMap IntMap.toList + (Print.ppList (Print.ppPair Print.ppInt Clause.pp))); val ppRewrite = ppField "rewrite" Rewrite.pp; val ppSubterms = ppField "subterms" (TermNet.pp - (Parser.ppMap (fn (c,l,p,t) => ((Clause.id c, l, p), t)) - (Parser.ppPair - (Parser.ppTriple Parser.ppInt Literal.pp Term.ppPath) + (Print.ppMap (fn (c,l,p,t) => ((Clause.id c, l, p), t)) + (Print.ppPair + (Print.ppTriple Print.ppInt Literal.pp Term.ppPath) Term.pp))); in - fun pp p (Active {clauses,rewrite,subterms,...}) = - (beginBlock p Inconsistent 2; - addString p "Active"; - addBreak p (1,0); - beginBlock p Inconsistent 1; - addString p "{"; - ppClauses p clauses; - addString p ","; - addBreak p (1,0); - ppRewrite p rewrite; -(*TRACE5 - addString p ","; - addBreak p (1,0); - ppSubterms p subterms; + fun pp (Active {clauses,rewrite,subterms,...}) = + Print.blockProgram Print.Inconsistent 2 + [Print.addString "Active", + Print.addBreak 1, + Print.blockProgram Print.Inconsistent 1 + [Print.addString "{", + ppClauses clauses, + Print.addString ",", + Print.addBreak 1, + ppRewrite rewrite, +(*MetisTrace5 + Print.addString ",", + Print.addBreak 1, + ppSubterms subterms, *) - endBlock p; - addString p "}"; - endBlock p); + Print.skip], + Print.addString "}"]; end; *) -val toString = Parser.toString pp; +val toString = Print.toString pp; (* ------------------------------------------------------------------------- *) (* Simplify clauses. *) @@ -314,17 +391,17 @@ end end; -(*DEBUG +(*MetisDebug val simplify = fn simp => fn units => fn rewr => fn subs => fn cl => let - fun traceCl s = Parser.ppTrace Clause.pp ("Active.simplify: " ^ s) -(*TRACE4 - val ppClOpt = Parser.ppOption Clause.pp + fun traceCl s = Print.trace Clause.pp ("Active.simplify: " ^ s) +(*MetisTrace4 + val ppClOpt = Print.ppOption Clause.pp val () = traceCl "cl" cl *) val cl' = simplify simp units rewr subs cl -(*TRACE4 - val () = Parser.ppTrace ppClOpt "Active.simplify: cl'" cl' +(*MetisTrace4 + val () = Print.trace ppClOpt "Active.simplify: cl'" cl' *) val () = case cl' of @@ -459,8 +536,8 @@ case total (Clause.resolve cl_lit) (cl,lit) of SOME cl' => cl' :: acc | NONE => acc -(*TRACE4 - val () = Parser.ppTrace Literal.pp "Active.deduceResolution: lit" lit +(*MetisTrace4 + val () = Print.trace Literal.pp "Active.deduceResolution: lit" lit *) in if Atom.isEq atm then acc @@ -495,13 +572,30 @@ val eqns = Clause.largestEquations cl val subtms = if TermNet.null equations then [] else Clause.largestSubterms cl +(*MetisTrace5 + val () = Print.trace LiteralSet.pp "Active.deduce: lits" lits + val () = Print.trace + (Print.ppList + (Print.ppMap (fn (lit,ort,_) => (lit,ort)) + (Print.ppPair Literal.pp Rewrite.ppOrient))) + "Active.deduce: eqns" eqns + val () = Print.trace + (Print.ppList + (Print.ppTriple Literal.pp Term.ppPath Term.pp)) + "Active.deduce: subtms" subtms +*) val acc = [] val acc = LiteralSet.foldl (deduceResolution literals cl) acc lits val acc = foldl (deduceParamodulationWith subterms cl) acc eqns val acc = foldl (deduceParamodulationInto equations cl) acc subtms + val acc = rev acc + +(*MetisTrace5 + val () = Print.trace (Print.ppList Clause.pp) "Active.deduce: acc" acc +*) in - rev acc + acc end; (* ------------------------------------------------------------------------- *) @@ -555,12 +649,12 @@ in order (tm,tm') = SOME GREATER end - + fun addRed ((cl,tm),acc) = let -(*TRACE5 - val () = Parser.ppTrace Clause.pp "Active.addRed: cl" cl - val () = Parser.ppTrace Term.pp "Active.addRed: tm" tm +(*MetisTrace5 + val () = Print.trace Clause.pp "Active.addRed: cl" cl + val () = Print.trace Term.pp "Active.addRed: tm" tm *) val id = Clause.id cl in @@ -569,15 +663,15 @@ else IntSet.add acc id end -(*TRACE5 - val () = Parser.ppTrace Term.pp "Active.addReduce: l" l - val () = Parser.ppTrace Term.pp "Active.addReduce: r" r - val () = Parser.ppTrace Parser.ppBool "Active.addReduce: ord" ord +(*MetisTrace5 + val () = Print.trace Term.pp "Active.addReduce: l" l + val () = Print.trace Term.pp "Active.addReduce: r" r + val () = Print.trace Print.ppBool "Active.addReduce: ord" ord *) in List.foldl addRed acc (TermNet.matched allSubterms l) end - + fun addEquation redexResidues (id,acc) = case Rewrite.peek rewrite id of NONE => acc @@ -601,7 +695,7 @@ if choose_clause_rewritables active ids then clause_rewritables active else rewrite_rewritables active ids; -(*DEBUG +(*MetisDebug val rewritables = fn active => fn ids => let val clause_ids = clause_rewritables active @@ -611,13 +705,13 @@ if IntSet.equal rewrite_ids clause_ids then () else let - val ppIdl = Parser.ppList Parser.ppInt - val ppIds = Parser.ppMap IntSet.toList ppIdl - val () = Parser.ppTrace pp "Active.rewritables: active" active - val () = Parser.ppTrace ppIdl "Active.rewritables: ids" ids - val () = Parser.ppTrace ppIds + val ppIdl = Print.ppList Print.ppInt + val ppIds = Print.ppMap IntSet.toList ppIdl + val () = Print.trace pp "Active.rewritables: active" active + val () = Print.trace ppIdl "Active.rewritables: ids" ids + val () = Print.trace ppIds "Active.rewritables: clause_ids" clause_ids - val () = Parser.ppTrace ppIds + val () = Print.trace ppIds "Active.rewritables: rewrite_ids" rewrite_ids in raise Bug "Active.rewritables: ~(rewrite_ids SUBSET clause_ids)" @@ -632,12 +726,19 @@ else let fun idPred id = not (IntSet.member id ids) - + fun clausePred cl = idPred (Clause.id cl) - + val Active - {parameters,clauses,units,rewrite,subsume,literals, - equations,subterms,allSubterms} = active + {parameters, + clauses, + units, + rewrite, + subsume, + literals, + equations, + subterms, + allSubterms} = active val clauses = IntMap.filter (idPred o fst) clauses and subsume = Subsume.filter clausePred subsume @@ -647,9 +748,14 @@ and allSubterms = TermNet.filter (clausePred o fst) allSubterms in Active - {parameters = parameters, clauses = clauses, units = units, - rewrite = rewrite, subsume = subsume, literals = literals, - equations = equations, subterms = subterms, + {parameters = parameters, + clauses = clauses, + units = units, + rewrite = rewrite, + subsume = subsume, + literals = literals, + equations = equations, + subterms = subterms, allSubterms = allSubterms} end; in @@ -657,21 +763,21 @@ if Rewrite.isReduced rewrite then (active,[]) else let -(*TRACE3 +(*MetisTrace3 val () = trace "Active.extract_rewritables: inter-reducing\n" *) val (rewrite,ids) = Rewrite.reduce' rewrite val active = setRewrite active rewrite val ids = rewritables active ids val cls = IntSet.transform (IntMap.get clauses) ids -(*TRACE3 - val ppCls = Parser.ppList Clause.pp - val () = Parser.ppTrace ppCls "Active.extract_rewritables: cls" cls +(*MetisTrace3 + val ppCls = Print.ppList Clause.pp + val () = Print.trace ppCls "Active.extract_rewritables: cls" cls *) in (delete active ids, cls) end -(*DEBUG +(*MetisDebug handle Error err => raise Bug ("Active.extract_rewritables: shouldn't fail\n" ^ err); *) @@ -745,13 +851,13 @@ fun factor active cls = factor' active [] cls; end; -(*TRACE4 +(*MetisTrace4 val factor = fn active => fn cls => let - val ppCls = Parser.ppList Clause.pp - val () = Parser.ppTrace ppCls "Active.factor: cls" cls + val ppCls = Print.ppList Clause.pp + val () = Print.trace ppCls "Active.factor: cls" cls val (active,cls') = factor active cls - val () = Parser.ppTrace ppCls "Active.factor: cls'" cls' + val () = Print.trace ppCls "Active.factor: cls'" cls' in (active,cls') end; @@ -761,16 +867,18 @@ (* Create a new active clause set and initialize clauses. *) (* ------------------------------------------------------------------------- *) -fun new parameters ths = +fun new parameters {axioms,conjecture} = let val {clause,...} = parameters fun mk_clause th = Clause.mk {parameters = clause, id = Clause.newId (), thm = th} - val cls = map mk_clause ths + val active = empty parameters + val (active,axioms) = factor active (map mk_clause axioms) + val (active,conjecture) = factor active (map mk_clause conjecture) in - factor (empty parameters) cls + (active, {axioms = axioms, conjecture = conjecture}) end; (* ------------------------------------------------------------------------- *) @@ -785,16 +893,16 @@ else if not (Clause.equalThms cl cl') then factor active [cl'] else let -(*TRACE3 - val () = Parser.ppTrace Clause.pp "Active.add: cl" cl +(*MetisTrace2 + val () = Print.trace Clause.pp "Active.add: cl" cl *) val active = addClause active cl val cl = Clause.freshVars cl val cls = deduce active cl val (active,cls) = factor active cls -(*TRACE2 - val ppCls = Parser.ppList Clause.pp - val () = Parser.ppTrace ppCls "Active.add: cls" cls +(*MetisTrace2 + val ppCls = Print.ppList Clause.pp + val () = Print.trace ppCls "Active.add: cls" cls *) in (active,cls) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Atom.sig --- a/src/Tools/Metis/src/Atom.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Atom.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC ATOMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Atom = @@ -52,6 +52,8 @@ val compare : atom * atom -> order +val equal : atom -> atom -> bool + (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) @@ -94,6 +96,8 @@ (* The equality relation. *) (* ------------------------------------------------------------------------- *) +val eqRelationName : relationName + val eqRelation : relation val mkEq : Term.term * Term.term -> atom @@ -126,12 +130,12 @@ (* Parsing and pretty printing. *) (* ------------------------------------------------------------------------- *) -val pp : atom Parser.pp +val pp : atom Print.pp val toString : atom -> string val fromString : string -> atom -val parse : Term.term Parser.quotation -> atom +val parse : Term.term Parse.quotation -> atom end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Atom.sml --- a/src/Tools/Metis/src/Atom.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Atom.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC ATOMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Atom :> Atom = @@ -49,7 +49,7 @@ fun mkBinop p (a,b) : atom = (p,[a,b]); fun destBinop p (x,[a,b]) = - if x = p then (a,b) else raise Error "Atom.destBinop: wrong binop" + if Name.equal x p then (a,b) else raise Error "Atom.destBinop: wrong binop" | destBinop _ _ = raise Error "Atom.destBinop: not a binop"; fun isBinop p = can (destBinop p); @@ -70,6 +70,8 @@ | EQUAL => lexCompare Term.compare (tms1,tms2) | GREATER => GREATER; +fun equal atm1 atm2 = compare (atm1,atm2) = EQUAL; + (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) @@ -94,7 +96,7 @@ val tm = List.nth (tms,h) val tm' = Term.replace tm (t,res) in - if Sharing.pointerEqual (tm,tm') then atm + if Portable.pointerEqual (tm,tm') then atm else (rel, updateNth (h,tm') tms) end; @@ -129,7 +131,7 @@ let val tms' = Sharing.map (Subst.subst sub) tms in - if Sharing.pointerEqual (tms',tms) then atm else (p,tms') + if Portable.pointerEqual (tms',tms) then atm else (p,tms') end; (* ------------------------------------------------------------------------- *) @@ -141,7 +143,7 @@ in fun match sub (p1,tms1) (p2,tms2) = let - val _ = (p1 = p2 andalso length tms1 = length tms2) orelse + val _ = (Name.equal p1 p2 andalso length tms1 = length tms2) orelse raise Error "Atom.match" in foldl matchArg sub (zip tms1 tms2) @@ -157,7 +159,7 @@ in fun unify sub (p1,tms1) (p2,tms2) = let - val _ = (p1 = p2 andalso length tms1 = length tms2) orelse + val _ = (Name.equal p1 p2 andalso length tms1 = length tms2) orelse raise Error "Atom.unify" in foldl unifyArg sub (zip tms1 tms2) @@ -168,24 +170,24 @@ (* The equality relation. *) (* ------------------------------------------------------------------------- *) -val eqName = "="; +val eqRelationName = Name.fromString "="; -val eqArity = 2; +val eqRelationArity = 2; -val eqRelation = (eqName,eqArity); +val eqRelation = (eqRelationName,eqRelationArity); -val mkEq = mkBinop eqName; +val mkEq = mkBinop eqRelationName; -fun destEq x = destBinop eqName x; +fun destEq x = destBinop eqRelationName x; -fun isEq x = isBinop eqName x; +fun isEq x = isBinop eqRelationName x; fun mkRefl tm = mkEq (tm,tm); fun destRefl atm = let val (l,r) = destEq atm - val _ = l = r orelse raise Error "Atom.destRefl" + val _ = Term.equal l r orelse raise Error "Atom.destRefl" in l end; @@ -195,7 +197,7 @@ fun sym atm = let val (l,r) = destEq atm - val _ = l <> r orelse raise Error "Atom.sym: refl" + val _ = not (Term.equal l r) orelse raise Error "Atom.sym: refl" in mkEq (r,l) end; @@ -227,19 +229,19 @@ (* Parsing and pretty printing. *) (* ------------------------------------------------------------------------- *) -val pp = Parser.ppMap Term.Fn Term.pp; +val pp = Print.ppMap Term.Fn Term.pp; -val toString = Parser.toString pp; +val toString = Print.toString pp; fun fromString s = Term.destFn (Term.fromString s); -val parse = Parser.parseQuotation Term.toString fromString; +val parse = Parse.parseQuotation Term.toString fromString; end structure AtomOrdered = struct type t = Atom.atom val compare = Atom.compare end -structure AtomSet = ElementSet (AtomOrdered); +structure AtomMap = KeyMap (AtomOrdered); -structure AtomMap = KeyMap (AtomOrdered); +structure AtomSet = ElementSet (AtomMap); diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/AtomNet.sig --- a/src/Tools/Metis/src/AtomNet.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/AtomNet.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature AtomNet = @@ -30,7 +30,7 @@ val toString : 'a atomNet -> string -val pp : 'a Parser.pp -> 'a atomNet Parser.pp +val pp : 'a Print.pp -> 'a atomNet Print.pp (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/AtomNet.sml --- a/src/Tools/Metis/src/AtomNet.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/AtomNet.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure AtomNet :> AtomNet = diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Clause.sig --- a/src/Tools/Metis/src/Clause.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Clause.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* CLAUSE = ID + THEOREM *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Clause = @@ -100,6 +100,8 @@ val showId : bool ref -val pp : clause Parser.pp +val pp : clause Print.pp + +val toString : clause -> string end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Clause.sml --- a/src/Tools/Metis/src/Clause.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Clause.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* CLAUSE = ID + THEOREM *) -(* Copyright (c) 2002-2004 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Clause :> Clause = @@ -16,8 +16,7 @@ let val r = ref 0 in - fn () => CRITICAL (fn () => - case r of ref n => let val () = r := n + 1 in n end) + fn () => case r of ref n => let val () = r := n + 1 in n end end; (* ------------------------------------------------------------------------- *) @@ -47,19 +46,21 @@ val showId = ref false; local - val ppIdThm = Parser.ppPair Parser.ppInt Thm.pp; + val ppIdThm = Print.ppPair Print.ppInt Thm.pp; in - fun pp p (Clause {id,thm,...}) = - if !showId then ppIdThm p (id,thm) else Thm.pp p thm; + fun pp (Clause {id,thm,...}) = + if !showId then ppIdThm (id,thm) else Thm.pp thm; end; +fun toString cl = Print.toString pp cl; + (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) -val default : parameters = +val default : parameters = {ordering = KnuthBendixOrder.default, - orderLiterals = UnsignedLiteralOrder, (*LCP: changed from PositiveLiteralOrder*) + orderLiterals = PositiveLiteralOrder, orderTerms = true}; fun mk info = Clause info @@ -142,13 +143,13 @@ LiteralSet.foldr addLit LiteralSet.empty litSet end; -(*TRACE6 +(*MetisTrace6 val largestLiterals = fn cl => let val ppResult = LiteralSet.pp - val () = Parser.ppTrace pp "Clause.largestLiterals: cl" cl + val () = Print.trace pp "Clause.largestLiterals: cl" cl val result = largestLiterals cl - val () = Parser.ppTrace ppResult "Clause.largestLiterals: result" result + val () = Print.trace ppResult "Clause.largestLiterals: result" result in result end; @@ -218,10 +219,10 @@ Rewrite.rewriteIdRule rewr cmp id th end -(*TRACE4 - val () = Parser.ppTrace Rewrite.pp "Clause.rewrite: rewr" rewr - val () = Parser.ppTrace Parser.ppInt "Clause.rewrite: id" id - val () = Parser.ppTrace pp "Clause.rewrite: cl" cl +(*MetisTrace4 + val () = Print.trace Rewrite.pp "Clause.rewrite: rewr" rewr + val () = Print.trace Print.ppInt "Clause.rewrite: id" id + val () = Print.trace pp "Clause.rewrite: cl" cl *) val thm = @@ -231,13 +232,13 @@ val result = Clause {parameters = parameters, id = id, thm = thm} -(*TRACE4 - val () = Parser.ppTrace pp "Clause.rewrite: result" result +(*MetisTrace4 + val () = Print.trace pp "Clause.rewrite: result" result *) in result end -(*DEBUG +(*MetisDebug handle Error err => raise Error ("Clause.rewrite:\n" ^ err); *) @@ -254,12 +255,12 @@ map apply (Rule.factor' lits) end; -(*TRACE5 +(*MetisTrace5 val factor = fn cl => let - val () = Parser.ppTrace pp "Clause.factor: cl" cl + val () = Print.trace pp "Clause.factor: cl" cl val result = factor cl - val () = Parser.ppTrace (Parser.ppList pp) "Clause.factor: result" result + val () = Print.trace (Print.ppList pp) "Clause.factor: result" result in result end; @@ -267,51 +268,55 @@ fun resolve (cl1,lit1) (cl2,lit2) = let -(*TRACE5 - val () = Parser.ppTrace pp "Clause.resolve: cl1" cl1 - val () = Parser.ppTrace Literal.pp "Clause.resolve: lit1" lit1 - val () = Parser.ppTrace pp "Clause.resolve: cl2" cl2 - val () = Parser.ppTrace Literal.pp "Clause.resolve: lit2" lit2 +(*MetisTrace5 + val () = Print.trace pp "Clause.resolve: cl1" cl1 + val () = Print.trace Literal.pp "Clause.resolve: lit1" lit1 + val () = Print.trace pp "Clause.resolve: cl2" cl2 + val () = Print.trace Literal.pp "Clause.resolve: lit2" lit2 *) val Clause {parameters, thm = th1, ...} = cl1 and Clause {thm = th2, ...} = cl2 val sub = Literal.unify Subst.empty lit1 (Literal.negate lit2) -(*TRACE5 - val () = Parser.ppTrace Subst.pp "Clause.resolve: sub" sub +(*MetisTrace5 + val () = Print.trace Subst.pp "Clause.resolve: sub" sub *) val lit1 = Literal.subst sub lit1 val lit2 = Literal.negate lit1 val th1 = Thm.subst sub th1 and th2 = Thm.subst sub th2 val _ = isLargerLiteral parameters (Thm.clause th1) lit1 orelse -(*TRACE5 +(*MetisTrace5 (trace "Clause.resolve: th1 violates ordering\n"; false) orelse *) raise Error "resolve: clause1: ordering constraints" val _ = isLargerLiteral parameters (Thm.clause th2) lit2 orelse -(*TRACE5 +(*MetisTrace5 (trace "Clause.resolve: th2 violates ordering\n"; false) orelse *) raise Error "resolve: clause2: ordering constraints" val th = Thm.resolve lit1 th1 th2 -(*TRACE5 - val () = Parser.ppTrace Thm.pp "Clause.resolve: th" th +(*MetisTrace5 + val () = Print.trace Thm.pp "Clause.resolve: th" th *) val cl = Clause {parameters = parameters, id = newId (), thm = th} -(*TRACE5 - val () = Parser.ppTrace pp "Clause.resolve: cl" cl +(*MetisTrace5 + val () = Print.trace pp "Clause.resolve: cl" cl *) in cl end; -fun paramodulate (cl1,lit1,ort,tm1) (cl2,lit2,path,tm2) = +fun paramodulate (cl1,lit1,ort1,tm1) (cl2,lit2,path2,tm2) = let -(*TRACE5 - val () = Parser.ppTrace pp "Clause.paramodulate: cl1" cl1 - val () = Parser.ppTrace Literal.pp "Clause.paramodulate: lit1" lit1 - val () = Parser.ppTrace pp "Clause.paramodulate: cl2" cl2 - val () = Parser.ppTrace Literal.pp "Clause.paramodulate: lit2" lit2 +(*MetisTrace5 + val () = Print.trace pp "Clause.paramodulate: cl1" cl1 + val () = Print.trace Literal.pp "Clause.paramodulate: lit1" lit1 + val () = Print.trace Rewrite.ppOrient "Clause.paramodulate: ort1" ort1 + val () = Print.trace Term.pp "Clause.paramodulate: tm1" tm1 + val () = Print.trace pp "Clause.paramodulate: cl2" cl2 + val () = Print.trace Literal.pp "Clause.paramodulate: lit2" lit2 + val () = Print.trace Term.ppPath "Clause.paramodulate: path2" path2 + val () = Print.trace Term.pp "Clause.paramodulate: tm2" tm2 *) val Clause {parameters, thm = th1, ...} = cl1 and Clause {thm = th2, ...} = cl2 @@ -320,32 +325,36 @@ and lit2 = Literal.subst sub lit2 and th1 = Thm.subst sub th1 and th2 = Thm.subst sub th2 + val _ = isLargerLiteral parameters (Thm.clause th1) lit1 orelse -(*TRACE5 - (trace "Clause.paramodulate: cl1 ordering\n"; false) orelse -*) - raise Error "paramodulate: with clause: ordering constraints" + raise Error "Clause.paramodulate: with clause: ordering" val _ = isLargerLiteral parameters (Thm.clause th2) lit2 orelse -(*TRACE5 - (trace "Clause.paramodulate: cl2 ordering\n"; false) orelse -*) - raise Error "paramodulate: into clause: ordering constraints" + raise Error "Clause.paramodulate: into clause: ordering" + val eqn = (Literal.destEq lit1, th1) val eqn as (l_r,_) = - case ort of + case ort1 of Rewrite.LeftToRight => eqn | Rewrite.RightToLeft => Rule.symEqn eqn - val _ = isLargerTerm parameters l_r orelse -(*TRACE5 - (trace "Clause.paramodulate: eqn ordering\n"; false) orelse +(*MetisTrace6 + val () = Print.trace Rule.ppEquation "Clause.paramodulate: eqn" eqn *) - raise Error "paramodulate: equation: ordering constraints" - val th = Rule.rewrRule eqn lit2 path th2 -(*TRACE5 - val () = Parser.ppTrace Thm.pp "Clause.paramodulate: th" th + val _ = isLargerTerm parameters l_r orelse + raise Error "Clause.paramodulate: equation: ordering constraints" + val th = Rule.rewrRule eqn lit2 path2 th2 +(*MetisTrace5 + val () = Print.trace Thm.pp "Clause.paramodulate: th" th *) in Clause {parameters = parameters, id = newId (), thm = th} - end; + end +(*MetisTrace5 + handle Error err => + let + val () = trace ("Clause.paramodulate: failed: " ^ err ^ "\n") + in + raise Error err + end; +*) end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/ElementSet.sig --- a/src/Tools/Metis/src/ElementSet.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/ElementSet.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,36 +1,78 @@ (* ========================================================================= *) (* 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 *) +(* ------------------------------------------------------------------------- *) +(* Removing. *) +(* ------------------------------------------------------------------------- *) + +val delete : set -> element -> set (* must be present *) + +val remove : set -> element -> set -(* Union and intersect prefer elements in the second set *) +val deletePick : set -> element * set + +val deleteNth : set -> int -> element * set + +val deleteRandom : set -> element * set + +(* ------------------------------------------------------------------------- *) +(* Joining. *) +(* ------------------------------------------------------------------------- *) val union : set -> set -> set @@ -44,22 +86,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 @@ -72,27 +116,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 *) +(* ------------------------------------------------------------------------- *) +(* Converting to and from maps. *) +(* ------------------------------------------------------------------------- *) -val random : set -> element (* raises Empty *) +type 'a map -val deletePick : set -> element * set (* raises Empty *) +val mapPartial : (element -> 'a option) -> set -> 'a map -val deleteRandom : set -> element * set (* raises Empty *) +val map : (element -> 'a) -> set -> 'a map + +val domain : 'a map -> set -val compare : set * set -> order - -val close : (set -> set) -> set -> set +(* ------------------------------------------------------------------------- *) +(* Pretty-printing. *) +(* ------------------------------------------------------------------------- *) val toString : set -> string diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/ElementSet.sml --- a/src/Tools/Metis/src/ElementSet.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/ElementSet.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,123 +1,347 @@ (* ========================================================================= *) (* 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 *) (* ========================================================================= *) -functor ElementSet (Key : Ordered) :> ElementSet where type element = Key.t = +functor ElementSet (KM : KeyMap) :> ElementSet +where type element = KM.key and type 'a map = 'a KM.map = struct -(*isabelle open Metis;*) +(* ------------------------------------------------------------------------- *) +(* 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; -type element = Key.t; +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; (* ------------------------------------------------------------------------- *) -(* Finite sets *) +(* 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. *) (* ------------------------------------------------------------------------- *) -type set = Key.t Set.set; - -val empty = Set.empty Key.compare; +fun union (Set m1) (Set m2) = Set (KM.unionDomain m1 m2); -fun singleton key = Set.singleton Key.compare key; +fun unionList sets = + let + val ms = List.map dest sets + in + Set (KM.unionListDomain ms) + end; -val null = Set.null; - -val size = Set.size; +fun intersect (Set m1) (Set m2) = Set (KM.intersectDomain m1 m2); -val member = Set.member; +fun intersectList sets = + let + val ms = List.map dest sets + in + Set (KM.intersectListDomain ms) + end; -val add = Set.add; +fun difference (Set m1) (Set m2) = + Set (KM.differenceDomain m1 m2); -val addList = Set.addList; +fun symmetricDifference (Set m1) (Set m2) = + Set (KM.symmetricDifferenceDomain m1 m2); + +(* ------------------------------------------------------------------------- *) +(* Mapping and folding. *) +(* ------------------------------------------------------------------------- *) -val delete = Set.delete; - -val union = Set.union; +fun filter pred = + let + fun mpred (elt,()) = pred elt + in + fn Set m => Set (KM.filter mpred m) + end; -val unionList = Set.unionList; - -val intersect = Set.intersect; +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; -val intersectList = Set.intersectList; - -val difference = Set.difference; - -val symmetricDifference = Set.symmetricDifference; +fun app f = + let + fun mf (elt,()) = f elt + in + fn Set m => KM.app mf m + end; -val disjoint = Set.disjoint; +fun foldl f = + let + fun mf (elt,(),acc) = f (elt,acc) + in + fn acc => fn Set m => KM.foldl mf acc m + end; -val subset = Set.subset; +fun foldr f = + let + fun mf (elt,(),acc) = f (elt,acc) + in + fn acc => fn Set m => KM.foldr mf acc m + end; -val equal = Set.equal; +(* ------------------------------------------------------------------------- *) +(* Searching. *) +(* ------------------------------------------------------------------------- *) -val filter = Set.filter; - -val partition = Set.partition; +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; -val count = Set.count; - -val foldl = Set.foldl; +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; -val foldr = Set.foldr; - -val findl = Set.findl; - -val findr = Set.findr; +fun firstl f = + let + fun mf (elt,()) = f elt + in + fn Set m => KM.firstl mf m + end; -val firstl = Set.firstl; - -val firstr = Set.firstr; +fun firstr f = + let + fun mf (elt,()) = f elt + in + fn Set m => KM.firstr mf m + end; -val exists = Set.exists; +fun exists p = + let + fun mp (elt,()) = p elt + in + fn Set m => KM.exists mp m + end; -val all = Set.all; - -val map = Set.map; +fun all p = + let + fun mp (elt,()) = p elt + in + fn Set m => KM.all mp m + end; -val transform = Set.transform; - -val app = Set.app; +fun count p = + let + fun mp (elt,()) = p elt + in + fn Set m => KM.count mp m + end; -val toList = Set.toList; +(* ------------------------------------------------------------------------- *) +(* Comparing. *) +(* ------------------------------------------------------------------------- *) + +fun compareValue ((),()) = EQUAL; -fun fromList l = Set.fromList Key.compare l; +fun equalValue () () = true; + +fun compare (Set m1, Set m2) = KM.compare compareValue (m1,m2); -val pick = Set.pick; +fun equal (Set m1) (Set m2) = KM.equal equalValue m1 m2; + +fun subset (Set m1) (Set m2) = KM.subsetDomain m1 m2; -val random = Set.random; +fun disjoint (Set m1) (Set m2) = KM.disjointDomain m1 m2; -val deletePick = Set.deletePick; - -val deleteRandom = Set.deleteRandom; +(* ------------------------------------------------------------------------- *) +(* Converting to and from lists. *) +(* ------------------------------------------------------------------------- *) -val compare = Set.compare; +fun transform f = + let + fun inc (x,l) = f x :: l + in + foldr inc [] + end; + +fun toList (Set m) = KM.keys m; -val close = Set.close; +fun fromList elts = addList empty elts; -val toString = Set.toString; +(* ------------------------------------------------------------------------- *) +(* Pretty-printing. *) +(* ------------------------------------------------------------------------- *) + +fun toString set = + "{" ^ (if null set then "" else Int.toString (size set)) ^ "}"; (* ------------------------------------------------------------------------- *) (* Iterators over sets *) (* ------------------------------------------------------------------------- *) -type iterator = Key.t Set.iterator; +type iterator = unit KM.iterator; -val mkIterator = Set.mkIterator; +fun mkIterator (Set m) = KM.mkIterator m; + +fun mkRevIterator (Set m) = KM.mkRevIterator m; -val mkRevIterator = Set.mkRevIterator; +fun readIterator iter = + let + val (elt,()) = KM.readIterator iter + in + elt + end; -val readIterator = Set.readIterator; - -val advanceIterator = Set.advanceIterator; +fun advanceIterator iter = KM.advanceIterator iter; end -(*isabelle structure Metis = struct open Metis;*) +structure IntSet = +ElementSet (IntMap); -structure IntSet = -ElementSet (IntOrdered); +structure IntPairSet = +ElementSet (IntPairMap); structure StringSet = -ElementSet (StringOrdered); - -(*isabelle end;*) +ElementSet (StringMap); diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/FILES --- a/src/Tools/Metis/src/FILES Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/FILES Mon Sep 13 21:24:10 2010 +0200 @@ -1,18 +1,20 @@ -Portable.sig PortableIsabelle.sml -PP.sig PP.sml +Random.sig Random.sml +Portable.sig PortablePolyml.sml Useful.sig Useful.sml Lazy.sig Lazy.sml +Stream.sig Stream.sml Ordered.sig Ordered.sml -Set.sig RandomSet.sml Set.sml -ElementSet.sig ElementSet.sml -Map.sig RandomMap.sml Map.sml +Map.sig Map.sml KeyMap.sig KeyMap.sml +Set.sig Set.sml +ElementSet.sig ElementSet.sml Sharing.sig Sharing.sml -Stream.sig Stream.sml Heap.sig Heap.sml -Parser.sig Parser.sml +Print.sig Print.sml +Parse.sig Parse.sml Options.sig Options.sml Name.sig Name.sml +NameArity.sig NameArity.sml Term.sig Term.sml Subst.sig Subst.sml Atom.sig Atom.sml diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Formula.sig --- a/src/Tools/Metis/src/Formula.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Formula.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC FORMULAS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Formula = @@ -34,6 +34,10 @@ val isBoolean : formula -> bool +val isTrue : formula -> bool + +val isFalse : formula -> bool + (* Functions *) val functions : formula -> NameAritySet.set @@ -92,6 +96,8 @@ val listMkForall : Term.var list * formula -> formula +val setMkForall : NameSet.set * formula -> formula + val stripForall : formula -> Term.var list * formula (* Existential quantification *) @@ -102,6 +108,8 @@ val listMkExists : Term.var list * formula -> formula +val setMkExists : NameSet.set * formula -> formula + val stripExists : formula -> Term.var list * formula (* ------------------------------------------------------------------------- *) @@ -116,6 +124,8 @@ val compare : formula * formula -> order +val equal : formula -> formula -> bool + (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) @@ -124,6 +134,8 @@ val freeVars : formula -> NameSet.set +val freeVarsList : formula list -> NameSet.set + val specialize : formula -> formula val generalize : formula -> formula @@ -163,12 +175,18 @@ val rhs : formula -> Term.term (* ------------------------------------------------------------------------- *) +(* Splitting goals. *) +(* ------------------------------------------------------------------------- *) + +val splitGoal : formula -> formula list + +(* ------------------------------------------------------------------------- *) (* Parsing and pretty-printing. *) (* ------------------------------------------------------------------------- *) -type quotation = formula Parser.quotation +type quotation = formula Parse.quotation -val pp : formula Parser.pp +val pp : formula Print.pp val toString : formula -> string diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Formula.sml --- a/src/Tools/Metis/src/Formula.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Formula.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC FORMULAS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Formula :> Formula = @@ -39,6 +39,16 @@ val isBoolean = can destBoolean; +fun isTrue fm = + case fm of + True => true + | _ => false; + +fun isFalse fm = + case fm of + False => true + | _ => false; + (* Functions *) local @@ -211,6 +221,8 @@ fun listMkForall ([],body) = body | listMkForall (v :: vs, body) = Forall (v, listMkForall (vs,body)); +fun setMkForall (vs,body) = NameSet.foldr Forall body vs; + local fun strip vs (Forall (v,b)) = strip (v :: vs) b | strip vs tm = (rev vs, tm); @@ -228,6 +240,8 @@ fun listMkExists ([],body) = body | listMkExists (v :: vs, body) = Exists (v, listMkExists (vs,body)); +fun setMkExists (vs,body) = NameSet.foldr Exists body vs; + local fun strip vs (Exists (v,b)) = strip (v :: vs) b | strip vs tm = (rev vs, tm); @@ -261,50 +275,56 @@ local fun cmp [] = EQUAL - | cmp ((True,True) :: l) = cmp l - | cmp ((True,_) :: _) = LESS - | cmp ((_,True) :: _) = GREATER - | cmp ((False,False) :: l) = cmp l - | cmp ((False,_) :: _) = LESS - | cmp ((_,False) :: _) = GREATER - | cmp ((Atom atm1, Atom atm2) :: l) = - (case Atom.compare (atm1,atm2) of - LESS => LESS - | EQUAL => cmp l - | GREATER => GREATER) - | cmp ((Atom _, _) :: _) = LESS - | cmp ((_, Atom _) :: _) = GREATER - | cmp ((Not p1, Not p2) :: l) = cmp ((p1,p2) :: l) - | cmp ((Not _, _) :: _) = LESS - | cmp ((_, Not _) :: _) = GREATER - | cmp ((And (p1,q1), And (p2,q2)) :: l) = cmp ((p1,p2) :: (q1,q2) :: l) - | cmp ((And _, _) :: _) = LESS - | cmp ((_, And _) :: _) = GREATER - | cmp ((Or (p1,q1), Or (p2,q2)) :: l) = cmp ((p1,p2) :: (q1,q2) :: l) - | cmp ((Or _, _) :: _) = LESS - | cmp ((_, Or _) :: _) = GREATER - | cmp ((Imp (p1,q1), Imp (p2,q2)) :: l) = cmp ((p1,p2) :: (q1,q2) :: l) - | cmp ((Imp _, _) :: _) = LESS - | cmp ((_, Imp _) :: _) = GREATER - | cmp ((Iff (p1,q1), Iff (p2,q2)) :: l) = cmp ((p1,p2) :: (q1,q2) :: l) - | cmp ((Iff _, _) :: _) = LESS - | cmp ((_, Iff _) :: _) = GREATER - | cmp ((Forall (v1,p1), Forall (v2,p2)) :: l) = - (case Name.compare (v1,v2) of - LESS => LESS - | EQUAL => cmp ((p1,p2) :: l) - | GREATER => GREATER) - | cmp ((Forall _, Exists _) :: _) = LESS - | cmp ((Exists _, Forall _) :: _) = GREATER - | cmp ((Exists (v1,p1), Exists (v2,p2)) :: l) = - (case Name.compare (v1,v2) of - LESS => LESS - | EQUAL => cmp ((p1,p2) :: l) - | GREATER => GREATER); + | cmp (f1_f2 :: fs) = + if Portable.pointerEqual f1_f2 then cmp fs + else + case f1_f2 of + (True,True) => cmp fs + | (True,_) => LESS + | (_,True) => GREATER + | (False,False) => cmp fs + | (False,_) => LESS + | (_,False) => GREATER + | (Atom atm1, Atom atm2) => + (case Atom.compare (atm1,atm2) of + LESS => LESS + | EQUAL => cmp fs + | GREATER => GREATER) + | (Atom _, _) => LESS + | (_, Atom _) => GREATER + | (Not p1, Not p2) => cmp ((p1,p2) :: fs) + | (Not _, _) => LESS + | (_, Not _) => GREATER + | (And (p1,q1), And (p2,q2)) => cmp ((p1,p2) :: (q1,q2) :: fs) + | (And _, _) => LESS + | (_, And _) => GREATER + | (Or (p1,q1), Or (p2,q2)) => cmp ((p1,p2) :: (q1,q2) :: fs) + | (Or _, _) => LESS + | (_, Or _) => GREATER + | (Imp (p1,q1), Imp (p2,q2)) => cmp ((p1,p2) :: (q1,q2) :: fs) + | (Imp _, _) => LESS + | (_, Imp _) => GREATER + | (Iff (p1,q1), Iff (p2,q2)) => cmp ((p1,p2) :: (q1,q2) :: fs) + | (Iff _, _) => LESS + | (_, Iff _) => GREATER + | (Forall (v1,p1), Forall (v2,p2)) => + (case Name.compare (v1,v2) of + LESS => LESS + | EQUAL => cmp ((p1,p2) :: fs) + | GREATER => GREATER) + | (Forall _, Exists _) => LESS + | (Exists _, Forall _) => GREATER + | (Exists (v1,p1), Exists (v2,p2)) => + (case Name.compare (v1,v2) of + LESS => LESS + | EQUAL => cmp ((p1,p2) :: fs) + | GREATER => GREATER); in fun compare fm1_fm2 = cmp [fm1_fm2]; end; +fun equal fm1 fm2 = compare (fm1,fm2) = EQUAL; + (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) @@ -320,8 +340,10 @@ | f (Or (p,q) :: fms) = f (p :: q :: fms) | f (Imp (p,q) :: fms) = f (p :: q :: fms) | f (Iff (p,q) :: fms) = f (p :: q :: fms) - | f (Forall (w,p) :: fms) = if v = w then f fms else f (p :: fms) - | f (Exists (w,p) :: fms) = if v = w then f fms else f (p :: fms) + | f (Forall (w,p) :: fms) = + if Name.equal v w then f fms else f (p :: fms) + | f (Exists (w,p) :: fms) = + if Name.equal v w then f fms else f (p :: fms) in fn fm => f [fm] end; @@ -339,8 +361,12 @@ | fv vs ((bv, Iff (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms) | fv vs ((bv, Forall (v,p)) :: fms) = fv vs ((NameSet.add bv v, p) :: fms) | fv vs ((bv, Exists (v,p)) :: fms) = fv vs ((NameSet.add bv v, p) :: fms); + + fun add (fm,vs) = fv vs [(NameSet.empty,fm)]; in - fun freeVars fm = fv NameSet.empty [(NameSet.empty,fm)]; + fun freeVars fm = add (fm,NameSet.empty); + + fun freeVarsList fms = List.foldl add NameSet.empty fms; end; fun specialize fm = snd (stripForall fm); @@ -362,13 +388,13 @@ let val tms' = Sharing.map (Subst.subst sub) tms in - if Sharing.pointerEqual (tms,tms') then fm else Atom (p,tms') + if Portable.pointerEqual (tms,tms') then fm else Atom (p,tms') end | Not p => let val p' = substFm sub p in - if Sharing.pointerEqual (p,p') then fm else Not p' + if Portable.pointerEqual (p,p') then fm else Not p' end | And (p,q) => substConn sub fm And p q | Or (p,q) => substConn sub fm Or p q @@ -382,8 +408,8 @@ val p' = substFm sub p and q' = substFm sub q in - if Sharing.pointerEqual (p,p') andalso - Sharing.pointerEqual (q,q') + if Portable.pointerEqual (p,p') andalso + Portable.pointerEqual (q,q') then fm else conn (p',q') end @@ -393,12 +419,12 @@ val v' = let fun f (w,s) = - if w = v then s + if Name.equal w v then s else case Subst.peek sub w of NONE => NameSet.add s w | SOME tm => NameSet.union s (Term.freeVars tm) - + val vars = freeVars p val vars = NameSet.foldl f NameSet.empty vars in @@ -406,12 +432,12 @@ end val sub = - if v = v' then Subst.remove sub (NameSet.singleton v) + if Name.equal v v' then Subst.remove sub (NameSet.singleton v) else Subst.insert sub (v, Term.Var v') val p' = substCheck sub p in - if v = v' andalso Sharing.pointerEqual (p,p') then fm + if Name.equal v v' andalso Portable.pointerEqual (p,p') then fm else quant (v',p') end; in @@ -451,34 +477,39 @@ (* Parsing and pretty-printing. *) (* ------------------------------------------------------------------------- *) -type quotation = formula Parser.quotation +type quotation = formula Parse.quotation; -val truthSymbol = "T" -and falsitySymbol = "F" -and conjunctionSymbol = "/\\" -and disjunctionSymbol = "\\/" -and implicationSymbol = "==>" -and equivalenceSymbol = "<=>" -and universalSymbol = "!" -and existentialSymbol = "?"; +val truthName = Name.fromString "T" +and falsityName = Name.fromString "F" +and conjunctionName = Name.fromString "/\\" +and disjunctionName = Name.fromString "\\/" +and implicationName = Name.fromString "==>" +and equivalenceName = Name.fromString "<=>" +and universalName = Name.fromString "!" +and existentialName = Name.fromString "?"; local - fun demote True = Term.Fn (truthSymbol,[]) - | demote False = Term.Fn (falsitySymbol,[]) + fun demote True = Term.Fn (truthName,[]) + | demote False = Term.Fn (falsityName,[]) | demote (Atom (p,tms)) = Term.Fn (p,tms) - | demote (Not p) = Term.Fn (!Term.negation, [demote p]) - | demote (And (p,q)) = Term.Fn (conjunctionSymbol, [demote p, demote q]) - | demote (Or (p,q)) = Term.Fn (disjunctionSymbol, [demote p, demote q]) - | demote (Imp (p,q)) = Term.Fn (implicationSymbol, [demote p, demote q]) - | demote (Iff (p,q)) = Term.Fn (equivalenceSymbol, [demote p, demote q]) - | demote (Forall (v,b)) = Term.Fn (universalSymbol, [Term.Var v, demote b]) + | demote (Not p) = + let + val ref s = Term.negation + in + Term.Fn (Name.fromString s, [demote p]) + end + | demote (And (p,q)) = Term.Fn (conjunctionName, [demote p, demote q]) + | demote (Or (p,q)) = Term.Fn (disjunctionName, [demote p, demote q]) + | demote (Imp (p,q)) = Term.Fn (implicationName, [demote p, demote q]) + | demote (Iff (p,q)) = Term.Fn (equivalenceName, [demote p, demote q]) + | demote (Forall (v,b)) = Term.Fn (universalName, [Term.Var v, demote b]) | demote (Exists (v,b)) = - Term.Fn (existentialSymbol, [Term.Var v, demote b]); + Term.Fn (existentialName, [Term.Var v, demote b]); in - fun pp ppstrm fm = Term.pp ppstrm (demote fm); + fun pp fm = Term.pp (demote fm); end; -val toString = Parser.toString pp; +val toString = Print.toString pp; local fun isQuant [Term.Var _, _] = true @@ -486,23 +517,23 @@ fun promote (Term.Var v) = Atom (v,[]) | promote (Term.Fn (f,tms)) = - if f = truthSymbol andalso null tms then + if Name.equal f truthName andalso null tms then True - else if f = falsitySymbol andalso null tms then + else if Name.equal f falsityName andalso null tms then False - else if f = !Term.negation andalso length tms = 1 then + else if Name.toString f = !Term.negation andalso length tms = 1 then Not (promote (hd tms)) - else if f = conjunctionSymbol andalso length tms = 2 then + else if Name.equal f conjunctionName andalso length tms = 2 then And (promote (hd tms), promote (List.nth (tms,1))) - else if f = disjunctionSymbol andalso length tms = 2 then + else if Name.equal f disjunctionName andalso length tms = 2 then Or (promote (hd tms), promote (List.nth (tms,1))) - else if f = implicationSymbol andalso length tms = 2 then + else if Name.equal f implicationName andalso length tms = 2 then Imp (promote (hd tms), promote (List.nth (tms,1))) - else if f = equivalenceSymbol andalso length tms = 2 then + else if Name.equal f equivalenceName andalso length tms = 2 then Iff (promote (hd tms), promote (List.nth (tms,1))) - else if f = universalSymbol andalso isQuant tms then + else if Name.equal f universalName andalso isQuant tms then Forall (Term.destVar (hd tms), promote (List.nth (tms,1))) - else if f = existentialSymbol andalso isQuant tms then + else if Name.equal f existentialName andalso isQuant tms then Exists (Term.destVar (hd tms), promote (List.nth (tms,1))) else Atom (f,tms); @@ -510,6 +541,61 @@ fun fromString s = promote (Term.fromString s); end; -val parse = Parser.parseQuotation toString fromString; +val parse = Parse.parseQuotation toString fromString; + +(* ------------------------------------------------------------------------- *) +(* Splitting goals. *) +(* ------------------------------------------------------------------------- *) + +local + fun add_asms asms goal = + if null asms then goal else Imp (listMkConj (rev asms), goal); + + fun add_var_asms asms v goal = add_asms asms (Forall (v,goal)); + + fun split asms pol fm = + case (pol,fm) of + (* Positive splittables *) + (true,True) => [] + | (true, Not f) => split asms false f + | (true, And (f1,f2)) => split asms true f1 @ split (f1 :: asms) true f2 + | (true, Or (f1,f2)) => split (Not f1 :: asms) true f2 + | (true, Imp (f1,f2)) => split (f1 :: asms) true f2 + | (true, Iff (f1,f2)) => + split (f1 :: asms) true f2 @ split (f2 :: asms) true f1 + | (true, Forall (v,f)) => map (add_var_asms asms v) (split [] true f) + (* Negative splittables *) + | (false,False) => [] + | (false, Not f) => split asms true f + | (false, And (f1,f2)) => split (f1 :: asms) false f2 + | (false, Or (f1,f2)) => + split asms false f1 @ split (Not f1 :: asms) false f2 + | (false, Imp (f1,f2)) => split asms true f1 @ split (f1 :: asms) false f2 + | (false, Iff (f1,f2)) => + split (f1 :: asms) false f2 @ split (f2 :: asms) false f1 + | (false, Exists (v,f)) => map (add_var_asms asms v) (split [] false f) + (* Unsplittables *) + | _ => [add_asms asms (if pol then fm else Not fm)]; +in + fun splitGoal fm = split [] true fm; +end; + +(*MetisTrace3 +val splitGoal = fn fm => + let + val result = splitGoal fm + val () = Print.trace pp "Formula.splitGoal: fm" fm + val () = Print.trace (Print.ppList pp) "Formula.splitGoal: result" result + in + result + end; +*) end + +structure FormulaOrdered = +struct type t = Formula.formula val compare = Formula.compare end + +structure FormulaMap = KeyMap (FormulaOrdered); + +structure FormulaSet = ElementSet (FormulaMap); diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Heap.sig --- a/src/Tools/Metis/src/Heap.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Heap.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* A HEAP DATATYPE FOR ML *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Heap = diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Heap.sml --- a/src/Tools/Metis/src/Heap.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Heap.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* A HEAP DATATYPE FOR ML *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Heap :> Heap = @@ -63,12 +63,12 @@ end; fun toStream h = - if null h then Stream.NIL + if null h then Stream.Nil else let val (x,h) = remove h in - Stream.CONS (x, fn () => toStream h) + Stream.Cons (x, fn () => toStream h) end; fun toString h = diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/KeyMap.sig --- a/src/Tools/Metis/src/KeyMap.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/KeyMap.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,65 +1,144 @@ (* ========================================================================= *) (* FINITE MAPS WITH A FIXED KEY TYPE *) -(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature KeyMap = sig +(* ------------------------------------------------------------------------- *) +(* A type of map keys. *) +(* ------------------------------------------------------------------------- *) + type key (* ------------------------------------------------------------------------- *) -(* Finite maps *) +(* 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 -val singleton : key * 'a -> 'a map +(* ------------------------------------------------------------------------- *) +(* Querying. *) +(* ------------------------------------------------------------------------- *) -val inDomain : key -> 'a map -> bool +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 -val get : 'a map -> key -> 'a (* raises Error *) +(* ------------------------------------------------------------------------- *) +(* 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 -(* Union and intersect prefer keys in the second map *) +val deleteRandom : 'a map -> (key * 'a) * 'a map + +(* ------------------------------------------------------------------------- *) +(* Joining (all join operations prefer keys in the second map). *) +(* ------------------------------------------------------------------------- *) -val union : ('a * 'a -> 'a option) -> 'a map -> 'a map -> 'a 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 intersect : ('a * 'a -> 'a option) -> 'a map -> 'a map -> 'a 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 delete : 'a map -> key -> 'a map (* raises Error *) +val unionDomain : 'a map -> 'a map -> 'a map + +val unionListDomain : 'a map list -> 'a map + +val intersectDomain : 'a map -> 'a map -> 'a map -val difference : '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 equalDomain : '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 filter : (key * 'a -> bool) -> 'a map -> 'a 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 @@ -72,22 +151,36 @@ val all : (key * 'a -> bool) -> 'a map -> bool -val domain : 'a map -> key list +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 -val compare : ('a * 'a -> order) -> 'a map * 'a map -> order - -val equal : ('a -> 'a -> bool) -> 'a map -> 'a map -> bool - -val random : 'a map -> key * 'a (* raises Empty *) +(* ------------------------------------------------------------------------- *) +(* Pretty-printing. *) +(* ------------------------------------------------------------------------- *) val toString : 'a map -> string (* ------------------------------------------------------------------------- *) -(* Iterators over maps *) +(* Iterators over maps. *) (* ------------------------------------------------------------------------- *) type 'a iterator diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/KeyMap.sml --- a/src/Tools/Metis/src/KeyMap.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/KeyMap.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,115 +1,1442 @@ (* ========================================================================= *) (* FINITE MAPS WITH A FIXED KEY TYPE *) -(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) functor KeyMap (Key : Ordered) :> KeyMap where type key = Key.t = struct -(*isabelle open Metis;*) +(* ------------------------------------------------------------------------- *) +(* Importing from the input signature. *) +(* ------------------------------------------------------------------------- *) type key = Key.t; +val compareKey = Key.compare; + (* ------------------------------------------------------------------------- *) -(* Finite maps *) +(* 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 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 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. *) (* ------------------------------------------------------------------------- *) -type 'a map = (Key.t,'a) Map.map; +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 -fun new () = Map.new Key.compare; +and nodeIntersectDomain node1 node2 = + let + val Node {priority,left,key,value,right,...} = node2 + + val (l,kvo,r) = nodePartition key node1 -val null = Map.null; + 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; -val size = Map.size; +(* ------------------------------------------------------------------------- *) +(* 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 -fun singleton key_value = Map.singleton Key.compare key_value; +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 inDomain = Map.inDomain; + 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 subset operation on trees. *) +(* ------------------------------------------------------------------------- *) -val peek = Map.peek; +fun treeSubsetDomain tree1 tree2 = + case tree1 of + E => true + | T node1 => + case tree2 of + E => false + | T node2 => nodeSubsetDomain node1 node2 -val insert = Map.insert; +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; -val insertList = Map.insertList; +(* ------------------------------------------------------------------------- *) +(* Picking an arbitrary key/value pair from a tree. *) +(* ------------------------------------------------------------------------- *) -val get = Map.get; +fun nodePick node = + let + val Node {key,value,...} = node + in + (key,value) + end; -(* Both union and intersect prefer keys in the second map *) +fun treePick tree = + case tree of + E => raise Bug "KeyMap.treePick" + | T node => nodePick node; + +(* ------------------------------------------------------------------------- *) +(* Removing an arbitrary key/value pair from a tree. *) +(* ------------------------------------------------------------------------- *) -val union = Map.union; +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; -val intersect = Map.intersect; +(* ------------------------------------------------------------------------- *) +(* 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 delete = Map.delete; + 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; -val difference = Map.difference; +(* ------------------------------------------------------------------------- *) +(* 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 subsetDomain = Map.subsetDomain; + 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 equalDomain = Map.equalDomain; + 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 mapPartial = Map.mapPartial; + val (key_value,right) = treeDeleteNth n right + + val size = size - 1 -val filter = Map.filter; + val node = + Node + {size = size, + priority = priority, + left = left, + key = key, + value = value, + right = right} + in + (key_value, T node) + end + end; -val map = Map.map; +(* ------------------------------------------------------------------------- *) +(* Iterators. *) +(* ------------------------------------------------------------------------- *) -val app = Map.app; +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)); -val transform = Map.transform; +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; -val foldl = Map.foldl; +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; -val foldr = Map.foldr; +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; -val findl = Map.findl; +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; -val findr = Map.findr; - -val firstl = Map.firstl; +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; -val firstr = Map.firstr; - -val exists = Map.exists; +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; -val all = Map.all; - -val domain = Map.domain; - -val toList = Map.toList; +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; -fun fromList l = Map.fromList Key.compare l; +(* ------------------------------------------------------------------------- *) +(* A type of finite maps. *) +(* ------------------------------------------------------------------------- *) + +datatype 'value map = + Map of 'value tree; + +(* ------------------------------------------------------------------------- *) +(* Map debugging functions. *) +(* ------------------------------------------------------------------------- *) -val compare = Map.compare; +(*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); +*) -val equal = Map.equal; +(* ------------------------------------------------------------------------- *) +(* Constructors. *) +(* ------------------------------------------------------------------------- *) + +fun new () = + let + val tree = treeNew () + in + Map tree + end; -val random = Map.random; +fun singleton key_value = + let + val tree = treeSingleton key_value + in + Map tree + end; -val toString = Map.toString; +(* ------------------------------------------------------------------------- *) +(* Map size. *) +(* ------------------------------------------------------------------------- *) + +fun size (Map tree) = treeSize tree; + +fun null m = size m = 0; (* ------------------------------------------------------------------------- *) -(* Iterators over maps *) +(* 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 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 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. *) (* ------------------------------------------------------------------------- *) -type 'a iterator = (Key.t,'a) Map.iterator; +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; -val mkIterator = Map.mkIterator; +(* ------------------------------------------------------------------------- *) +(* 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 mkRevIterator = Map.mkRevIterator; + 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)); +*) -val readIterator = Map.readIterator; +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); -val advanceIterator = Map.advanceIterator; +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 -(*isabelle structure Metis = struct open Metis*) - structure IntMap = KeyMap (IntOrdered); +structure IntPairMap = +KeyMap (IntPairOrdered); + structure StringMap = KeyMap (StringOrdered); - -(*isabelle end;*) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/KnuthBendixOrder.sig --- a/src/Tools/Metis/src/KnuthBendixOrder.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/KnuthBendixOrder.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* THE KNUTH-BENDIX TERM ORDERING *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature KnuthBendixOrder = diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/KnuthBendixOrder.sml --- a/src/Tools/Metis/src/KnuthBendixOrder.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/KnuthBendixOrder.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* KNUTH-BENDIX TERM ORDERING CONSTRAINTS *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure KnuthBendixOrder :> KnuthBendixOrder = @@ -12,10 +12,12 @@ (* Helper functions. *) (* ------------------------------------------------------------------------- *) -fun firstNotEqual f l = - case List.find op<> l of +fun notEqualTerm (x,y) = not (Term.equal x y); + +fun firstNotEqualTerm f l = + case List.find notEqualTerm l of SOME (x,y) => f x y - | NONE => raise Bug "firstNotEqual"; + | NONE => raise Bug "firstNotEqualTerm"; (* ------------------------------------------------------------------------- *) (* The weight of all constants must be at least 1, and there must be at most *) @@ -36,7 +38,7 @@ fn ((f1,n1),(f2,n2)) => case Int.compare (n1,n2) of LESS => LESS - | EQUAL => String.compare (f1,f2) + | EQUAL => Name.compare (f1,f2) | GREATER => GREATER; (* The default order *) @@ -62,7 +64,7 @@ fun weightNeg (Weight (m,c)) = Weight (NameMap.transform ~ m, ~c); local - fun add (n1,n2) = + fun add ((_,n1),(_,n2)) = let val n = n1 + n2 in @@ -75,15 +77,6 @@ fun weightSubtract w1 w2 = weightAdd w1 (weightNeg w2); -fun weightMult 0 _ = weightZero - | weightMult 1 w = w - | weightMult k (Weight (m,c)) = - let - fun mult n = k * n - in - Weight (NameMap.transform mult m, k * c) - end; - fun weightTerm weight = let fun wt m c [] = Weight (m,c) @@ -99,80 +92,41 @@ fn tm => wt weightEmpty ~1 [tm] end; -fun weightIsVar v (Weight (m,c)) = - c = 0 andalso NameMap.size m = 1 andalso NameMap.peek m v = SOME 1; - -fun weightConst (Weight (_,c)) = c; - -fun weightVars (Weight (m,_)) = - NameMap.foldl (fn (v,_,s) => NameSet.add s v) NameSet.empty m; - -val weightsVars = - List.foldl (fn (w,s) => NameSet.union (weightVars w) s) NameSet.empty; - -fun weightVarList w = NameSet.toList (weightVars w); - -fun weightNumVars (Weight (m,_)) = NameMap.size m; - -fun weightNumVarsWithPositiveCoeff (Weight (m,_)) = - NameMap.foldl (fn (_,n,z) => if n > 0 then z + 1 else z) 0 m; - -fun weightCoeff var (Weight (m,_)) = Option.getOpt (NameMap.peek m var, 0); - -fun weightCoeffs varList w = map (fn var => weightCoeff var w) varList; - -fun weightCoeffSum (Weight (m,_)) = NameMap.foldl (fn (_,n,z) => n + z) 0 m; - fun weightLowerBound (w as Weight (m,c)) = if NameMap.exists (fn (_,n) => n < 0) m then NONE else SOME c; -fun weightNoLowerBound w = not (Option.isSome (weightLowerBound w)); - -fun weightAlwaysPositive w = - case weightLowerBound w of NONE => false | SOME n => n > 0; - -fun weightUpperBound (w as Weight (m,c)) = - if NameMap.exists (fn (_,n) => n > 0) m then NONE else SOME c; - -fun weightNoUpperBound w = not (Option.isSome (weightUpperBound w)); - -fun weightAlwaysNegative w = - case weightUpperBound w of NONE => false | SOME n => n < 0; +(*MetisDebug +val ppWeightList = + let + fun ppCoeff n = + if n < 0 then Print.sequence (Print.addString "~") (ppCoeff (~n)) + else if n = 1 then Print.skip + else Print.ppInt n -fun weightGcd (w as Weight (m,c)) = - NameMap.foldl (fn (_,i,k) => gcd i k) (Int.abs c) m; - -fun ppWeightList pp = - let - fun coeffToString n = - if n < 0 then "~" ^ coeffToString (~n) - else if n = 1 then "" - else Int.toString n - - fun pp_tm pp ("",n) = Parser.ppInt pp n - | pp_tm pp (v,n) = Parser.ppString pp (coeffToString n ^ v) + fun pp_tm (NONE,n) = Print.ppInt n + | pp_tm (SOME v, n) = Print.sequence (ppCoeff n) (Name.pp v) in - fn [] => Parser.ppInt pp 0 - | tms => Parser.ppSequence " +" pp_tm pp tms + fn [] => Print.ppInt 0 + | tms => Print.ppOpList " +" pp_tm tms end; -fun ppWeight pp (Weight (m,c)) = +fun ppWeight (Weight (m,c)) = let val l = NameMap.toList m - val l = if c = 0 then l else l @ [("",c)] + val l = map (fn (v,n) => (SOME v, n)) l + val l = if c = 0 then l else l @ [(NONE,c)] in - ppWeightList pp l + ppWeightList l end; -val weightToString = Parser.toString ppWeight; +val weightToString = Print.toString ppWeight; +*) (* ------------------------------------------------------------------------- *) (* The Knuth-Bendix term order. *) (* ------------------------------------------------------------------------- *) -datatype kboOrder = Less | Equal | Greater | SignOf of weight; - -fun kboOrder {weight,precedence} = +fun compare {weight,precedence} = let fun weightDifference tm1 tm2 = let @@ -199,7 +153,7 @@ and precedenceLess (Term.Fn (f1,a1)) (Term.Fn (f2,a2)) = (case precedence ((f1, length a1), (f2, length a2)) of LESS => true - | EQUAL => firstNotEqual weightLess (zip a1 a2) + | EQUAL => firstNotEqualTerm weightLess (zip a1 a2) | GREATER => false) | precedenceLess _ _ = false @@ -210,39 +164,33 @@ val w = weightDifference tm1 tm2 in if weightIsZero w then precedenceCmp tm1 tm2 - else if weightDiffLess w tm1 tm2 then Less - else if weightDiffGreater w tm1 tm2 then Greater - else SignOf w + else if weightDiffLess w tm1 tm2 then SOME LESS + else if weightDiffGreater w tm1 tm2 then SOME GREATER + else NONE end and precedenceCmp (Term.Fn (f1,a1)) (Term.Fn (f2,a2)) = (case precedence ((f1, length a1), (f2, length a2)) of - LESS => Less - | EQUAL => firstNotEqual weightCmp (zip a1 a2) - | GREATER => Greater) + LESS => SOME LESS + | EQUAL => firstNotEqualTerm weightCmp (zip a1 a2) + | GREATER => SOME GREATER) | precedenceCmp _ _ = raise Bug "kboOrder.precendenceCmp" in - fn (tm1,tm2) => if tm1 = tm2 then Equal else weightCmp tm1 tm2 + fn (tm1,tm2) => + if Term.equal tm1 tm2 then SOME EQUAL else weightCmp tm1 tm2 end; -fun compare kbo tm1_tm2 = - case kboOrder kbo tm1_tm2 of - Less => SOME LESS - | Equal => SOME EQUAL - | Greater => SOME GREATER - | SignOf _ => NONE; - -(*TRACE7 +(*MetisTrace7 val compare = fn kbo => fn (tm1,tm2) => let - val () = Parser.ppTrace Term.pp "KnuthBendixOrder.compare: tm1" tm1 - val () = Parser.ppTrace Term.pp "KnuthBendixOrder.compare: tm2" tm2 + val () = Print.trace Term.pp "KnuthBendixOrder.compare: tm1" tm1 + val () = Print.trace Term.pp "KnuthBendixOrder.compare: tm2" tm2 val result = compare kbo (tm1,tm2) val () = case result of NONE => trace "KnuthBendixOrder.compare: result = Incomparable\n" | SOME x => - Parser.ppTrace Parser.ppOrder "KnuthBendixOrder.compare: result" x + Print.trace Print.ppOrder "KnuthBendixOrder.compare: result" x in result end; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Lazy.sig --- a/src/Tools/Metis/src/Lazy.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Lazy.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* 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 = @@ -8,6 +8,8 @@ type 'a lazy +val quickly : 'a -> 'a lazy + val delay : (unit -> 'a) -> 'a lazy val force : 'a lazy -> 'a diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Lazy.sml --- a/src/Tools/Metis/src/Lazy.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Lazy.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* 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 = @@ -12,16 +12,21 @@ datatype 'a lazy = Lazy of 'a thunk ref; +fun quickly v = Lazy (ref (Value v)); + fun delay f = Lazy (ref (Thunk f)); -fun force (Lazy (ref (Value v))) = v - | force (Lazy (s as 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Literal.sig --- a/src/Tools/Metis/src/Literal.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Literal.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC LITERALS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Literal = @@ -66,6 +66,8 @@ val compare : literal * literal -> order (* negative < positive *) +val equal : literal -> literal -> bool + (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) @@ -152,12 +154,12 @@ (* Parsing and pretty-printing. *) (* ------------------------------------------------------------------------- *) -val pp : literal Parser.pp +val pp : literal Print.pp val toString : literal -> string val fromString : string -> literal -val parse : Term.term Parser.quotation -> literal +val parse : Term.term Parse.quotation -> literal end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Literal.sml --- a/src/Tools/Metis/src/Literal.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Literal.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC LITERALS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Literal :> Literal = @@ -70,11 +70,9 @@ (* A total comparison function for literals. *) (* ------------------------------------------------------------------------- *) -fun compare ((pol1,atm1),(pol2,atm2)) = - case boolCompare (pol1,pol2) of - LESS => GREATER - | EQUAL => Atom.compare (atm1,atm2) - | GREATER => LESS; +val compare = prodCompare boolCompare Atom.compare; + +fun equal (p1,atm1) (p2,atm2) = p1 = p2 andalso Atom.equal atm1 atm2; (* ------------------------------------------------------------------------- *) (* Subterms. *) @@ -88,7 +86,7 @@ let val atm' = Atom.replace atm path_tm in - if Sharing.pointerEqual (atm,atm') then lit else (pol,atm') + if Portable.pointerEqual (atm,atm') then lit else (pol,atm') end; (* ------------------------------------------------------------------------- *) @@ -107,7 +105,7 @@ let val atm' = Atom.subst sub atm in - if Sharing.pointerEqual (atm',atm) then lit else (pol,atm') + if Portable.pointerEqual (atm',atm) then lit else (pol,atm') end; (* ------------------------------------------------------------------------- *) @@ -182,24 +180,26 @@ (* Parsing and pretty-printing. *) (* ------------------------------------------------------------------------- *) -val pp = Parser.ppMap toFormula Formula.pp; +val pp = Print.ppMap toFormula Formula.pp; -val toString = Parser.toString pp; +val toString = Print.toString pp; fun fromString s = fromFormula (Formula.fromString s); -val parse = Parser.parseQuotation Term.toString fromString; +val parse = Parse.parseQuotation Term.toString fromString; end structure LiteralOrdered = struct type t = Literal.literal val compare = Literal.compare end +structure LiteralMap = KeyMap (LiteralOrdered); + structure LiteralSet = struct local - structure S = ElementSet (LiteralOrdered); + structure S = ElementSet (LiteralMap); in open S; end; @@ -227,6 +227,8 @@ foldl f NameAritySet.empty end; + fun freeIn v = exists (Literal.freeIn v); + val freeVars = let fun f (lit,set) = NameSet.union set (Literal.freeVars lit) @@ -234,6 +236,13 @@ foldl f NameSet.empty end; + val freeVarsList = + let + fun f (lits,set) = NameSet.union set (freeVars lits) + in + List.foldl f NameSet.empty + end; + val symbols = let fun f (lit,z) = Literal.symbols lit + z @@ -253,21 +262,32 @@ fun substLit (lit,(eq,lits')) = let val lit' = Literal.subst sub lit - val eq = eq andalso Sharing.pointerEqual (lit,lit') + val eq = eq andalso Portable.pointerEqual (lit,lit') in (eq, add lits' lit') end - + val (eq,lits') = foldl substLit (true,empty) lits in if eq then lits else lits' end; + fun conjoin set = + Formula.listMkConj (List.map Literal.toFormula (toList set)); + + fun disjoin set = + Formula.listMkDisj (List.map Literal.toFormula (toList set)); + val pp = - Parser.ppMap + Print.ppMap toList - (Parser.ppBracket "{" "}" (Parser.ppSequence "," Literal.pp)); + (Print.ppBracket "{" "}" (Print.ppOpList "," Literal.pp)); end -structure LiteralMap = KeyMap (LiteralOrdered); +structure LiteralSetOrdered = +struct type t = LiteralSet.set val compare = LiteralSet.compare end + +structure LiteralSetMap = KeyMap (LiteralSetOrdered); + +structure LiteralSetSet = ElementSet (LiteralSetMap); diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/LiteralNet.sig --- a/src/Tools/Metis/src/LiteralNet.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/LiteralNet.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature LiteralNet = @@ -32,7 +32,7 @@ val toString : 'a literalNet -> string -val pp : 'a Parser.pp -> 'a literalNet Parser.pp +val pp : 'a Print.pp -> 'a literalNet Print.pp (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/LiteralNet.sml --- a/src/Tools/Metis/src/LiteralNet.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/LiteralNet.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure LiteralNet :> LiteralNet = @@ -48,9 +48,9 @@ fun toString net = "LiteralNet[" ^ Int.toString (size net) ^ "]"; fun pp ppA = - Parser.ppMap + Print.ppMap (fn {positive,negative} => (positive,negative)) - (Parser.ppBinop " + NEGATIVE" (AtomNet.pp ppA) (AtomNet.pp ppA)); + (Print.ppOp2 " + NEGATIVE" (AtomNet.pp ppA) (AtomNet.pp ppA)); (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Map.sig --- a/src/Tools/Metis/src/Map.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Map.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,65 +1,138 @@ (* ========================================================================= *) -(* FINITE MAPS *) -(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) +(* FINITE MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES *) +(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Map = sig (* ------------------------------------------------------------------------- *) -(* Finite maps *) +(* 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 -val singleton : ('key * 'key -> order) -> 'key * 'a -> ('key,'a) map +(* ------------------------------------------------------------------------- *) +(* Querying. *) +(* ------------------------------------------------------------------------- *) -val inDomain : 'key -> ('key,'a) map -> bool +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 -val get : ('key,'a) map -> 'key -> 'a (* raises Error *) +(* ------------------------------------------------------------------------- *) +(* 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 -(* Union and intersect prefer keys in the second 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 : - ('a * 'a -> 'a option) -> ('key,'a) map -> ('key,'a) map -> ('key,'a) map + (('key * 'a) * ('key * '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 + (('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 delete : ('key,'a) map -> 'key -> ('key,'a) map (* raises Error *) +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 difference : ('key,'a) map -> ('key,'b) map -> ('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 equalDomain : ('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 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 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 @@ -72,22 +145,36 @@ val all : ('key * 'a -> bool) -> ('key,'a) map -> bool -val domain : ('key,'a) map -> 'key list +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 -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 +(* ------------------------------------------------------------------------- *) +(* Pretty-printing. *) +(* ------------------------------------------------------------------------- *) val toString : ('key,'a) map -> string (* ------------------------------------------------------------------------- *) -(* Iterators over maps *) +(* Iterators over maps. *) (* ------------------------------------------------------------------------- *) type ('key,'a) iterator diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Map.sml --- a/src/Tools/Metis/src/Map.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Map.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,1425 @@ (* ========================================================================= *) -(* FINITE MAPS *) -(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) +(* FINITE MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES *) +(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) -structure Map = RandomMap; +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 diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Model.sig --- a/src/Tools/Metis/src/Model.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Model.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,39 +1,197 @@ (* ========================================================================= *) (* RANDOM FINITE MODELS *) -(* Copyright (c) 2003-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Model = sig (* ------------------------------------------------------------------------- *) +(* Model size. *) +(* ------------------------------------------------------------------------- *) + +type size = {size : int} + +(* ------------------------------------------------------------------------- *) +(* A model of size N has integer elements 0...N-1. *) +(* ------------------------------------------------------------------------- *) + +type element = int + +val zeroElement : element + +val incrementElement : size -> element -> element option + +(* ------------------------------------------------------------------------- *) (* The parts of the model that are fixed. *) -(* Note: a model of size N has integer elements 0...N-1. *) +(* ------------------------------------------------------------------------- *) + +type fixedFunction = size -> element list -> element option + +type fixedRelation = size -> element list -> bool option + +datatype fixed = + Fixed of + {functions : fixedFunction NameArityMap.map, + relations : fixedRelation NameArityMap.map} + +val emptyFixed : fixed + +val unionFixed : fixed -> fixed -> fixed + +val getFunctionFixed : fixed -> NameArity.nameArity -> fixedFunction + +val getRelationFixed : fixed -> NameArity.nameArity -> fixedRelation + +val insertFunctionFixed : fixed -> NameArity.nameArity * fixedFunction -> fixed + +val insertRelationFixed : fixed -> NameArity.nameArity * fixedRelation -> fixed + +val unionListFixed : fixed list -> fixed + +val basicFixed : fixed (* interprets equality and hasType *) + +(* ------------------------------------------------------------------------- *) +(* Renaming fixed model parts. *) +(* ------------------------------------------------------------------------- *) + +type fixedMap = + {functionMap : Name.name NameArityMap.map, + relationMap : Name.name NameArityMap.map} + +val mapFixed : fixedMap -> fixed -> fixed + +val ppFixedMap : fixedMap Print.pp + +(* ------------------------------------------------------------------------- *) +(* Standard fixed model parts. *) (* ------------------------------------------------------------------------- *) -type fixed = - {size : int} -> - {functions : (Term.functionName * int list) -> int option, - relations : (Atom.relationName * int list) -> bool option} +(* Projections *) + +val projectionMin : int + +val projectionMax : int + +val projectionName : int -> Name.name + +val projectionFixed : fixed + +(* Arithmetic *) + +val numeralMin : int + +val numeralMax : int + +val numeralName : int -> Name.name + +val addName : Name.name + +val divName : Name.name + +val dividesName : Name.name + +val evenName : Name.name + +val expName : Name.name + +val geName : Name.name -val fixedMerge : fixed -> fixed -> fixed (* Prefers the second fixed *) +val gtName : Name.name + +val isZeroName : Name.name + +val leName : Name.name + +val ltName : Name.name + +val modName : Name.name + +val multName : Name.name + +val negName : Name.name + +val oddName : Name.name -val fixedMergeList : fixed list -> fixed +val preName : Name.name + +val subName : Name.name + +val sucName : Name.name + +val modularFixed : fixed -val fixedPure : fixed (* : = *) +val overflowFixed : fixed + +(* Sets *) + +val cardName : Name.name + +val complementName : Name.name -val fixedBasic : fixed (* id fst snd #1 #2 #3 <> *) +val differenceName : Name.name + +val emptyName : Name.name + +val memberName : Name.name + +val insertName : Name.name + +val intersectName : Name.name + +val singletonName : Name.name + +val subsetName : Name.name + +val symmetricDifferenceName : Name.name -val fixedModulo : fixed (* suc pre ~ + - * exp div mod *) - (* is_0 divides even odd *) +val unionName : Name.name + +val universeName : Name.name + +val setFixed : fixed + +(* Lists *) + +val appendName : Name.name + +val consName : Name.name + +val lengthName : Name.name + +val nilName : Name.name -val fixedOverflowNum : fixed (* suc pre + - * exp div mod *) - (* is_0 <= < >= > divides even odd *) +val nullName : Name.name + +val tailName : Name.name + +val listFixed : fixed + +(* ------------------------------------------------------------------------- *) +(* Valuations. *) +(* ------------------------------------------------------------------------- *) + +type valuation + +val emptyValuation : valuation + +val zeroValuation : NameSet.set -> valuation -val fixedOverflowInt : fixed (* suc pre + - * exp div mod *) - (* is_0 <= < >= > divides even odd *) +val constantValuation : element -> NameSet.set -> valuation + +val peekValuation : valuation -> Name.name -> element option + +val getValuation : valuation -> Name.name -> element + +val insertValuation : valuation -> Name.name * element -> valuation -val fixedSet : fixed (* empty univ union intersect compl card in subset *) +val randomValuation : {size : int} -> NameSet.set -> valuation + +val incrementValuation : + {size : int} -> NameSet.set -> valuation -> valuation option + +val foldValuation : + {size : int} -> NameSet.set -> (valuation * 'a -> 'a) -> 'a -> 'a (* ------------------------------------------------------------------------- *) (* A type of random finite models. *) @@ -43,28 +201,21 @@ type model +val default : parameters + val new : parameters -> model val size : model -> int (* ------------------------------------------------------------------------- *) -(* Valuations. *) -(* ------------------------------------------------------------------------- *) - -type valuation = int NameMap.map - -val valuationEmpty : valuation - -val valuationRandom : {size : int} -> NameSet.set -> valuation - -val valuationFold : - {size : int} -> NameSet.set -> (valuation * 'a -> 'a) -> 'a -> 'a - -(* ------------------------------------------------------------------------- *) (* Interpreting terms and formulas in the model. *) (* ------------------------------------------------------------------------- *) -val interpretTerm : model -> valuation -> Term.term -> int +val interpretFunction : model -> Term.functionName * element list -> element + +val interpretRelation : model -> Atom.relationName * element list -> bool + +val interpretTerm : model -> valuation -> Term.term -> element val interpretAtom : model -> valuation -> Atom.atom -> bool @@ -79,16 +230,48 @@ (* Note: if it's cheaper, a systematic check will be performed instead. *) (* ------------------------------------------------------------------------- *) +val check : + (model -> valuation -> 'a -> bool) -> {maxChecks : int option} -> model -> + NameSet.set -> 'a -> {T : int, F : int} + val checkAtom : - {maxChecks : int} -> model -> Atom.atom -> {T : int, F : int} + {maxChecks : int option} -> model -> Atom.atom -> {T : int, F : int} val checkFormula : - {maxChecks : int} -> model -> Formula.formula -> {T : int, F : int} + {maxChecks : int option} -> model -> Formula.formula -> {T : int, F : int} val checkLiteral : - {maxChecks : int} -> model -> Literal.literal -> {T : int, F : int} + {maxChecks : int option} -> model -> Literal.literal -> {T : int, F : int} val checkClause : - {maxChecks : int} -> model -> Thm.clause -> {T : int, F : int} + {maxChecks : int option} -> model -> Thm.clause -> {T : int, F : int} + +(* ------------------------------------------------------------------------- *) +(* Updating the model. *) +(* ------------------------------------------------------------------------- *) + +val updateFunction : + model -> (Term.functionName * element list) * element -> unit + +val updateRelation : + model -> (Atom.relationName * element list) * bool -> unit + +(* ------------------------------------------------------------------------- *) +(* Choosing a random perturbation to make a formula true in the model. *) +(* ------------------------------------------------------------------------- *) + +val perturbTerm : model -> valuation -> Term.term * element list -> unit + +val perturbAtom : model -> valuation -> Atom.atom * bool -> unit + +val perturbLiteral : model -> valuation -> Literal.literal -> unit + +val perturbClause : model -> valuation -> Thm.clause -> unit + +(* ------------------------------------------------------------------------- *) +(* Pretty printing. *) +(* ------------------------------------------------------------------------- *) + +val pp : model Print.pp end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Model.sml --- a/src/Tools/Metis/src/Model.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Model.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* RANDOM FINITE MODELS *) -(* Copyright (c) 2003-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Model :> Model = @@ -9,378 +9,912 @@ open Useful; (* ------------------------------------------------------------------------- *) +(* Constants. *) +(* ------------------------------------------------------------------------- *) + +val maxSpace = 1000; + +(* ------------------------------------------------------------------------- *) (* Helper functions. *) (* ------------------------------------------------------------------------- *) -fun intExp x y = exp op* x y 1; - -fun natFromString "" = NONE - | natFromString "0" = SOME 0 - | natFromString s = - case charToInt (String.sub (s,0)) of - NONE => NONE - | SOME 0 => NONE - | SOME d => +val multInt = + case Int.maxInt of + NONE => (fn x => fn y => SOME (x * y)) + | SOME m => let - fun parse 0 _ acc = SOME acc - | parse n i acc = - case charToInt (String.sub (s,i)) of - NONE => NONE - | SOME d => parse (n - 1) (i + 1) (10 * acc + d) + val m = Real.floor (Math.sqrt (Real.fromInt m)) in - parse (size s - 1) 1 d + fn x => fn y => if x <= m andalso y <= m then SOME (x * y) else NONE end; -fun projection (_,[]) = NONE - | projection ("#1", x :: _) = SOME x - | projection ("#2", _ :: x :: _) = SOME x - | projection ("#3", _ :: _ :: x :: _) = SOME x - | projection (func,args) = - let - val f = size func - and n = length args +local + fun iexp x y acc = + if y mod 2 = 0 then iexp' x y acc + else + case multInt acc x of + SOME acc => iexp' x y acc + | NONE => NONE + + and iexp' x y acc = + if y = 1 then SOME acc + else + let + val y = y div 2 + in + case multInt x x of + SOME x => iexp x y acc + | NONE => NONE + end; +in + fun expInt x y = + if y <= 1 then + if y = 0 then SOME 1 + else if y = 1 then SOME x + else raise Bug "expInt: negative exponent" + else if x <= 1 then + if 0 <= x then SOME x + else raise Bug "expInt: negative exponand" + else iexp x y 1; +end; + +fun boolToInt true = 1 + | boolToInt false = 0; + +fun intToBool 1 = true + | intToBool 0 = false + | intToBool _ = raise Bug "Model.intToBool"; - val p = - if f < 2 orelse n <= 3 orelse String.sub (func,0) <> #"#" then NONE - else if f = 2 then - (case charToInt (String.sub (func,1)) of - NONE => NONE - | p as SOME d => if d <= 3 then NONE else p) - else if (n < intExp 10 (f - 2) handle Overflow => true) then NONE - else - (natFromString (String.extract (func,1,NONE)) - handle Overflow => NONE) +fun minMaxInterval i j = interval i (1 + j - i); + +(* ------------------------------------------------------------------------- *) +(* Model size. *) +(* ------------------------------------------------------------------------- *) + +type size = {size : int}; + +(* ------------------------------------------------------------------------- *) +(* A model of size N has integer elements 0...N-1. *) +(* ------------------------------------------------------------------------- *) + +type element = int; + +val zeroElement = 0; + +fun incrementElement {size = N} i = + let + val i = i + 1 in - case p of - NONE => NONE - | SOME k => if k > n then NONE else SOME (List.nth (args, k - 1)) + if i = N then NONE else SOME i + end; + +fun elementListSpace {size = N} arity = + case expInt N arity of + NONE => NONE + | s as SOME m => if m <= maxSpace then s else NONE; + +fun elementListIndex {size = N} = + let + fun f acc elts = + case elts of + [] => acc + | elt :: elts => f (N * acc + elt) elts + in + f 0 end; (* ------------------------------------------------------------------------- *) (* The parts of the model that are fixed. *) -(* Note: a model of size N has integer elements 0...N-1. *) (* ------------------------------------------------------------------------- *) -type fixedModel = - {functions : (Term.functionName * int list) -> int option, - relations : (Atom.relationName * int list) -> bool option}; +type fixedFunction = size -> element list -> element option; + +type fixedRelation = size -> element list -> bool option; + +datatype fixed = + Fixed of + {functions : fixedFunction NameArityMap.map, + relations : fixedRelation NameArityMap.map}; + +val uselessFixedFunction : fixedFunction = K (K NONE); + +val uselessFixedRelation : fixedRelation = K (K NONE); + +val emptyFunctions : fixedFunction NameArityMap.map = NameArityMap.new (); + +val emptyRelations : fixedRelation NameArityMap.map = NameArityMap.new (); -type fixed = {size : int} -> fixedModel +fun fixed0 f sz elts = + case elts of + [] => f sz + | _ => raise Bug "Model.fixed0: wrong arity"; -fun fixedMerge fixed1 fixed2 parm = +fun fixed1 f sz elts = + case elts of + [x] => f sz x + | _ => raise Bug "Model.fixed1: wrong arity"; + +fun fixed2 f sz elts = + case elts of + [x,y] => f sz x y + | _ => raise Bug "Model.fixed2: wrong arity"; + +val emptyFixed = let - val {functions = f1, relations = r1} = fixed1 parm - and {functions = f2, relations = r2} = fixed2 parm - - fun functions x = case f2 x of NONE => f1 x | s => s - - fun relations x = case r2 x of NONE => r1 x | s => s + val fns = emptyFunctions + and rels = emptyRelations in - {functions = functions, relations = relations} + Fixed + {functions = fns, + relations = rels} end; -fun fixedMergeList [] = raise Bug "fixedMergeList: empty" - | fixedMergeList (f :: l) = foldl (uncurry fixedMerge) f l; +fun peekFunctionFixed fix name_arity = + let + val Fixed {functions = fns, ...} = fix + in + NameArityMap.peek fns name_arity + end; -fun fixedPure {size = _} = +fun peekRelationFixed fix name_arity = let - fun functions (":",[x,_]) = SOME x - | functions _ = NONE + val Fixed {relations = rels, ...} = fix + in + NameArityMap.peek rels name_arity + end; + +fun getFunctionFixed fix name_arity = + case peekFunctionFixed fix name_arity of + SOME f => f + | NONE => uselessFixedFunction; - fun relations (rel,[x,y]) = - if (rel,2) = Atom.eqRelation then SOME (x = y) else NONE - | relations _ = NONE +fun getRelationFixed fix name_arity = + case peekRelationFixed fix name_arity of + SOME rel => rel + | NONE => uselessFixedRelation; + +fun insertFunctionFixed fix name_arity_fn = + let + val Fixed {functions = fns, relations = rels} = fix + + val fns = NameArityMap.insert fns name_arity_fn in - {functions = functions, relations = relations} + Fixed + {functions = fns, + relations = rels} end; -fun fixedBasic {size = _} = +fun insertRelationFixed fix name_arity_rel = let - fun functions ("id",[x]) = SOME x - | functions ("fst",[x,_]) = SOME x - | functions ("snd",[_,x]) = SOME x - | functions func_args = projection func_args + val Fixed {functions = fns, relations = rels} = fix + + val rels = NameArityMap.insert rels name_arity_rel + in + Fixed + {functions = fns, + relations = rels} + end; - fun relations ("<>",[x,y]) = SOME (x <> y) - | relations _ = NONE +local + fun union _ = raise Bug "Model.unionFixed: nameArity clash"; +in + fun unionFixed fix1 fix2 = + let + val Fixed {functions = fns1, relations = rels1} = fix1 + and Fixed {functions = fns2, relations = rels2} = fix2 + + val fns = NameArityMap.union union fns1 fns2 + + val rels = NameArityMap.union union rels1 rels2 + in + Fixed + {functions = fns, + relations = rels} + end; +end; + +val unionListFixed = + let + fun union (fix,acc) = unionFixed acc fix in - {functions = functions, relations = relations} + List.foldl union emptyFixed end; -fun fixedModulo {size = N} = - let - fun mod_N k = k mod N +local + fun hasTypeFn _ elts = + case elts of + [x,_] => SOME x + | _ => raise Bug "Model.hasTypeFn: wrong arity"; - val one = mod_N 1 - - fun mult (x,y) = mod_N (x * y) - - fun divides_N 0 = false - | divides_N x = N mod x = 0 - - val even_N = divides_N 2 + fun eqRel _ elts = + case elts of + [x,y] => SOME (x = y) + | _ => raise Bug "Model.eqRel: wrong arity"; +in + val basicFixed = + let + val fns = NameArityMap.singleton (Term.hasTypeFunction,hasTypeFn) - fun functions (numeral,[]) = - Option.map mod_N (natFromString numeral handle Overflow => NONE) - | functions ("suc",[x]) = SOME (if x = N - 1 then 0 else x + 1) - | functions ("pre",[x]) = SOME (if x = 0 then N - 1 else x - 1) - | functions ("~",[x]) = SOME (if x = 0 then 0 else N - x) - | functions ("+",[x,y]) = SOME (mod_N (x + y)) - | functions ("-",[x,y]) = SOME (if x < y then N + x - y else x - y) - | functions ("*",[x,y]) = SOME (mult (x,y)) - | functions ("exp",[x,y]) = SOME (exp mult x y one) - | functions ("div",[x,y]) = if divides_N y then SOME (x div y) else NONE - | functions ("mod",[x,y]) = if divides_N y then SOME (x mod y) else NONE - | functions _ = NONE + val rels = NameArityMap.singleton (Atom.eqRelation,eqRel) + in + Fixed + {functions = fns, + relations = rels} + end; +end; + +(* ------------------------------------------------------------------------- *) +(* Renaming fixed model parts. *) +(* ------------------------------------------------------------------------- *) - fun relations ("is_0",[x]) = SOME (x = 0) - | relations ("divides",[x,y]) = - if x = 0 then SOME (y = 0) - else if divides_N x then SOME (y mod x = 0) else NONE - | relations ("even",[x]) = if even_N then SOME (x mod 2 = 0) else NONE - | relations ("odd",[x]) = if even_N then SOME (x mod 2 = 1) else NONE - | relations _ = NONE +type fixedMap = + {functionMap : Name.name NameArityMap.map, + relationMap : Name.name NameArityMap.map}; + +fun mapFixed fixMap fix = + let + val {functionMap = fnMap, relationMap = relMap} = fixMap + and Fixed {functions = fns, relations = rels} = fix + + val fns = NameArityMap.compose fnMap fns + + val rels = NameArityMap.compose relMap rels in - {functions = functions, relations = relations} + Fixed + {functions = fns, + relations = rels} end; local - datatype onum = ONeg | ONum of int | OInf; - - val zero = ONum 0 - and one = ONum 1 - and two = ONum 2; + fun mkEntry tag (na,n) = (tag,na,n); - fun suc (ONum x) = ONum (x + 1) - | suc v = v; - - fun pre (ONum 0) = ONeg - | pre (ONum x) = ONum (x - 1) - | pre v = v; - - fun neg ONeg = NONE - | neg (n as ONum 0) = SOME n - | neg _ = SOME ONeg; + fun mkList tag m = map (mkEntry tag) (NameArityMap.toList m); - fun add ONeg ONeg = SOME ONeg - | add ONeg (ONum y) = if y = 0 then SOME ONeg else NONE - | add ONeg OInf = NONE - | add (ONum x) ONeg = if x = 0 then SOME ONeg else NONE - | add (ONum x) (ONum y) = SOME (ONum (x + y)) - | add (ONum _) OInf = SOME OInf - | add OInf ONeg = NONE - | add OInf (ONum _) = SOME OInf - | add OInf OInf = SOME OInf; + fun ppEntry (tag,source_arity,target) = + Print.blockProgram Print.Inconsistent 2 + [Print.addString tag, + Print.addBreak 1, + NameArity.pp source_arity, + Print.addString " ->", + Print.addBreak 1, + Name.pp target]; +in + fun ppFixedMap fixMap = + let + val {functionMap = fnMap, relationMap = relMap} = fixMap + in + case mkList "function" fnMap @ mkList "relation" relMap of + [] => Print.skip + | entry :: entries => + Print.blockProgram Print.Consistent 0 + (ppEntry entry :: + map (Print.sequence Print.addNewline o ppEntry) entries) + end; +end; - fun sub ONeg ONeg = NONE - | sub ONeg (ONum _) = SOME ONeg - | sub ONeg OInf = SOME ONeg - | sub (ONum _) ONeg = NONE - | sub (ONum x) (ONum y) = SOME (if x < y then ONeg else ONum (x - y)) - | sub (ONum _) OInf = SOME ONeg - | sub OInf ONeg = SOME OInf - | sub OInf (ONum y) = if y = 0 then SOME OInf else NONE - | sub OInf OInf = NONE; +(* ------------------------------------------------------------------------- *) +(* Standard fixed model parts. *) +(* ------------------------------------------------------------------------- *) + +(* Projections *) - fun mult ONeg ONeg = NONE - | mult ONeg (ONum y) = SOME (if y = 0 then zero else ONeg) - | mult ONeg OInf = SOME ONeg - | mult (ONum x) ONeg = SOME (if x = 0 then zero else ONeg) - | mult (ONum x) (ONum y) = SOME (ONum (x * y)) - | mult (ONum x) OInf = SOME (if x = 0 then zero else OInf) - | mult OInf ONeg = SOME ONeg - | mult OInf (ONum y) = SOME (if y = 0 then zero else OInf) - | mult OInf OInf = SOME OInf; +val projectionMin = 1 +and projectionMax = 9; + +val projectionList = minMaxInterval projectionMin projectionMax; + +fun projectionName i = + let + val _ = projectionMin <= i orelse + raise Bug "Model.projectionName: less than projectionMin" + + val _ = i <= projectionMax orelse + raise Bug "Model.projectionName: greater than projectionMax" + in + Name.fromString ("project" ^ Int.toString i) + end; - fun exp ONeg ONeg = NONE - | exp ONeg (ONum y) = - if y = 0 then SOME one else if y mod 2 = 0 then NONE else SOME ONeg - | exp ONeg OInf = NONE - | exp (ONum x) ONeg = NONE - | exp (ONum x) (ONum y) = SOME (ONum (intExp x y) handle Overflow => OInf) - | exp (ONum x) OInf = - SOME (if x = 0 then zero else if x = 1 then one else OInf) - | exp OInf ONeg = NONE - | exp OInf (ONum y) = SOME (if y = 0 then one else OInf) - | exp OInf OInf = SOME OInf; +fun projectionFn i _ elts = SOME (List.nth (elts, i - 1)); + +fun arityProjectionFixed arity = + let + fun mkProj i = ((projectionName i, arity), projectionFn i) - fun odiv ONeg ONeg = NONE - | odiv ONeg (ONum y) = if y = 1 then SOME ONeg else NONE - | odiv ONeg OInf = NONE - | odiv (ONum _) ONeg = NONE - | odiv (ONum x) (ONum y) = if y = 0 then NONE else SOME (ONum (x div y)) - | odiv (ONum _) OInf = SOME zero - | odiv OInf ONeg = NONE - | odiv OInf (ONum y) = if y = 1 then SOME OInf else NONE - | odiv OInf OInf = NONE; + fun addProj i acc = + if i > arity then acc + else addProj (i + 1) (NameArityMap.insert acc (mkProj i)) + + val fns = addProj projectionMin emptyFunctions - fun omod ONeg ONeg = NONE - | omod ONeg (ONum y) = if y = 1 then SOME zero else NONE - | omod ONeg OInf = NONE - | omod (ONum _) ONeg = NONE - | omod (ONum x) (ONum y) = if y = 0 then NONE else SOME (ONum (x mod y)) - | omod (x as ONum _) OInf = SOME x - | omod OInf ONeg = NONE - | omod OInf (ONum y) = if y = 1 then SOME OInf else NONE - | omod OInf OInf = NONE; + val rels = emptyRelations + in + Fixed + {functions = fns, + relations = rels} + end; - fun le ONeg ONeg = NONE - | le ONeg (ONum y) = SOME true - | le ONeg OInf = SOME true - | le (ONum _) ONeg = SOME false - | le (ONum x) (ONum y) = SOME (x <= y) - | le (ONum _) OInf = SOME true - | le OInf ONeg = SOME false - | le OInf (ONum _) = SOME false - | le OInf OInf = NONE; +val projectionFixed = + unionListFixed (map arityProjectionFixed projectionList); + +(* Arithmetic *) + +val numeralMin = ~100 +and numeralMax = 100; + +val numeralList = minMaxInterval numeralMin numeralMax; - fun lt x y = Option.map not (le y x); - - fun ge x y = le y x; - - fun gt x y = lt y x; - - fun divides ONeg ONeg = NONE - | divides ONeg (ONum y) = if y = 0 then SOME true else NONE - | divides ONeg OInf = NONE - | divides (ONum x) ONeg = - if x = 0 then SOME false else if x = 1 then SOME true else NONE - | divides (ONum x) (ONum y) = SOME (Useful.divides x y) - | divides (ONum x) OInf = - if x = 0 then SOME false else if x = 1 then SOME true else NONE - | divides OInf ONeg = NONE - | divides OInf (ONum y) = SOME (y = 0) - | divides OInf OInf = NONE; - - fun even n = divides two n; - - fun odd n = Option.map not (even n); +fun numeralName i = + let + val _ = numeralMin <= i orelse + raise Bug "Model.numeralName: less than numeralMin" - fun fixedOverflow mk_onum dest_onum = - let - fun partial_dest_onum NONE = NONE - | partial_dest_onum (SOME n) = dest_onum n + val _ = i <= numeralMax orelse + raise Bug "Model.numeralName: greater than numeralMax" + + val s = if i < 0 then "negative" ^ Int.toString (~i) else Int.toString i + in + Name.fromString s + end; - fun functions (numeral,[]) = - (case (natFromString numeral handle Overflow => NONE) of - NONE => NONE - | SOME n => dest_onum (ONum n)) - | functions ("suc",[x]) = dest_onum (suc (mk_onum x)) - | functions ("pre",[x]) = dest_onum (pre (mk_onum x)) - | functions ("~",[x]) = partial_dest_onum (neg (mk_onum x)) - | functions ("+",[x,y]) = - partial_dest_onum (add (mk_onum x) (mk_onum y)) - | functions ("-",[x,y]) = - partial_dest_onum (sub (mk_onum x) (mk_onum y)) - | functions ("*",[x,y]) = - partial_dest_onum (mult (mk_onum x) (mk_onum y)) - | functions ("exp",[x,y]) = - partial_dest_onum (exp (mk_onum x) (mk_onum y)) - | functions ("div",[x,y]) = - partial_dest_onum (odiv (mk_onum x) (mk_onum y)) - | functions ("mod",[x,y]) = - partial_dest_onum (omod (mk_onum x) (mk_onum y)) - | functions _ = NONE - - fun relations ("is_0",[x]) = SOME (mk_onum x = zero) - | relations ("<=",[x,y]) = le (mk_onum x) (mk_onum y) - | relations ("<",[x,y]) = lt (mk_onum x) (mk_onum y) - | relations (">=",[x,y]) = ge (mk_onum x) (mk_onum y) - | relations (">",[x,y]) = gt (mk_onum x) (mk_onum y) - | relations ("divides",[x,y]) = divides (mk_onum x) (mk_onum y) - | relations ("even",[x]) = even (mk_onum x) - | relations ("odd",[x]) = odd (mk_onum x) - | relations _ = NONE +val addName = Name.fromString "+" +and divName = Name.fromString "div" +and dividesName = Name.fromString "divides" +and evenName = Name.fromString "even" +and expName = Name.fromString "exp" +and geName = Name.fromString ">=" +and gtName = Name.fromString ">" +and isZeroName = Name.fromString "isZero" +and leName = Name.fromString "<=" +and ltName = Name.fromString "<" +and modName = Name.fromString "mod" +and multName = Name.fromString "*" +and negName = Name.fromString "~" +and oddName = Name.fromString "odd" +and preName = Name.fromString "pre" +and subName = Name.fromString "-" +and sucName = Name.fromString "suc"; + +local + (* Support *) + + fun modN {size = N} x = x mod N; + + fun oneN sz = modN sz 1; + + fun multN sz (x,y) = modN sz (x * y); + + (* Functions *) + + fun numeralFn i sz = SOME (modN sz i); + + fun addFn sz x y = SOME (modN sz (x + y)); + + fun divFn {size = N} x y = + let + val y = if y = 0 then N else y in - {functions = functions, relations = relations} + SOME (x div y) end; -in - fun fixedOverflowNum {size = N} = + + fun expFn sz x y = SOME (exp (multN sz) x y (oneN sz)); + + fun modFn {size = N} x y = let - val oinf = N - 1 - - fun mk_onum x = if x = oinf then OInf else ONum x - - fun dest_onum ONeg = NONE - | dest_onum (ONum x) = SOME (if x < oinf then x else oinf) - | dest_onum OInf = SOME oinf + val y = if y = 0 then N else y in - fixedOverflow mk_onum dest_onum + SOME (x mod y) end; - fun fixedOverflowInt {size = N} = + fun multFn sz x y = SOME (multN sz (x,y)); + + fun negFn {size = N} x = SOME (if x = 0 then 0 else N - x); + + fun preFn {size = N} x = SOME (if x = 0 then N - 1 else x - 1); + + fun subFn {size = N} x y = SOME (if x < y then N + x - y else x - y); + + fun sucFn {size = N} x = SOME (if x = N - 1 then 0 else x + 1); + + (* Relations *) + + fun dividesRel _ x y = SOME (divides x y); + + fun evenRel _ x = SOME (x mod 2 = 0); + + fun geRel _ x y = SOME (x >= y); + + fun gtRel _ x y = SOME (x > y); + + fun isZeroRel _ x = SOME (x = 0); + + fun leRel _ x y = SOME (x <= y); + + fun ltRel _ x y = SOME (x < y); + + fun oddRel _ x = SOME (x mod 2 = 1); +in + val modularFixed = let - val oinf = N - 2 - val oneg = N - 1 - - fun mk_onum x = - if x = oneg then ONeg else if x = oinf then OInf else ONum x - - fun dest_onum ONeg = SOME oneg - | dest_onum (ONum x) = SOME (if x < oinf then x else oinf) - | dest_onum OInf = SOME oinf + val fns = + NameArityMap.fromList + (map (fn i => ((numeralName i,0), fixed0 (numeralFn i))) + numeralList @ + [((addName,2), fixed2 addFn), + ((divName,2), fixed2 divFn), + ((expName,2), fixed2 expFn), + ((modName,2), fixed2 modFn), + ((multName,2), fixed2 multFn), + ((negName,1), fixed1 negFn), + ((preName,1), fixed1 preFn), + ((subName,2), fixed2 subFn), + ((sucName,1), fixed1 sucFn)]) + + val rels = + NameArityMap.fromList + [((dividesName,2), fixed2 dividesRel), + ((evenName,1), fixed1 evenRel), + ((geName,2), fixed2 geRel), + ((gtName,2), fixed2 gtRel), + ((isZeroName,1), fixed1 isZeroRel), + ((leName,2), fixed2 leRel), + ((ltName,2), fixed2 ltRel), + ((oddName,1), fixed1 oddRel)] in - fixedOverflow mk_onum dest_onum + Fixed + {functions = fns, + relations = rels} + end; +end; + +local + (* Support *) + + fun cutN {size = N} x = if x >= N then N - 1 else x; + + fun oneN sz = cutN sz 1; + + fun multN sz (x,y) = cutN sz (x * y); + + (* Functions *) + + fun numeralFn i sz = if i < 0 then NONE else SOME (cutN sz i); + + fun addFn sz x y = SOME (cutN sz (x + y)); + + fun divFn _ x y = if y = 0 then NONE else SOME (x div y); + + fun expFn sz x y = SOME (exp (multN sz) x y (oneN sz)); + + fun modFn {size = N} x y = + if y = 0 orelse x = N - 1 then NONE else SOME (x mod y); + + fun multFn sz x y = SOME (multN sz (x,y)); + + fun negFn _ x = if x = 0 then SOME 0 else NONE; + + fun preFn _ x = if x = 0 then NONE else SOME (x - 1); + + fun subFn {size = N} x y = + if y = 0 then SOME x + else if x = N - 1 orelse x < y then NONE + else SOME (x - y); + + fun sucFn sz x = SOME (cutN sz (x + 1)); + + (* Relations *) + + fun dividesRel {size = N} x y = + if x = 1 orelse y = 0 then SOME true + else if x = 0 then SOME false + else if y = N - 1 then NONE + else SOME (divides x y); + + fun evenRel {size = N} x = + if x = N - 1 then NONE else SOME (x mod 2 = 0); + + fun geRel {size = N} y x = + if x = N - 1 then if y = N - 1 then NONE else SOME false + else if y = N - 1 then SOME true else SOME (x <= y); + + fun gtRel {size = N} y x = + if x = N - 1 then if y = N - 1 then NONE else SOME false + else if y = N - 1 then SOME true else SOME (x < y); + + fun isZeroRel _ x = SOME (x = 0); + + fun leRel {size = N} x y = + if x = N - 1 then if y = N - 1 then NONE else SOME false + else if y = N - 1 then SOME true else SOME (x <= y); + + fun ltRel {size = N} x y = + if x = N - 1 then if y = N - 1 then NONE else SOME false + else if y = N - 1 then SOME true else SOME (x < y); + + fun oddRel {size = N} x = + if x = N - 1 then NONE else SOME (x mod 2 = 1); +in + val overflowFixed = + let + val fns = + NameArityMap.fromList + (map (fn i => ((numeralName i,0), fixed0 (numeralFn i))) + numeralList @ + [((addName,2), fixed2 addFn), + ((divName,2), fixed2 divFn), + ((expName,2), fixed2 expFn), + ((modName,2), fixed2 modFn), + ((multName,2), fixed2 multFn), + ((negName,1), fixed1 negFn), + ((preName,1), fixed1 preFn), + ((subName,2), fixed2 subFn), + ((sucName,1), fixed1 sucFn)]) + + val rels = + NameArityMap.fromList + [((dividesName,2), fixed2 dividesRel), + ((evenName,1), fixed1 evenRel), + ((geName,2), fixed2 geRel), + ((gtName,2), fixed2 gtRel), + ((isZeroName,1), fixed1 isZeroRel), + ((leName,2), fixed2 leRel), + ((ltName,2), fixed2 ltRel), + ((oddName,1), fixed1 oddRel)] + in + Fixed + {functions = fns, + relations = rels} end; end; -fun fixedSet {size = N} = - let - val M = - let - fun f 0 acc = acc - | f x acc = f (x div 2) (acc + 1) - in - f N 0 - end +(* Sets *) + +val cardName = Name.fromString "card" +and complementName = Name.fromString "complement" +and differenceName = Name.fromString "difference" +and emptyName = Name.fromString "empty" +and memberName = Name.fromString "member" +and insertName = Name.fromString "insert" +and intersectName = Name.fromString "intersect" +and singletonName = Name.fromString "singleton" +and subsetName = Name.fromString "subset" +and symmetricDifferenceName = Name.fromString "symmetricDifference" +and unionName = Name.fromString "union" +and universeName = Name.fromString "universe"; + +local + (* Support *) + + fun eltN {size = N} = + let + fun f 0 acc = acc + | f x acc = f (x div 2) (acc + 1) + in + f N ~1 + end; + + fun posN i = Word.<< (0w1, Word.fromInt i); + + fun univN sz = Word.- (posN (eltN sz), 0w1); + + fun setN sz x = Word.andb (Word.fromInt x, univN sz); + + (* Functions *) + + fun cardFn sz x = + let + fun f 0w0 acc = acc + | f s acc = + let + val acc = if Word.andb (s,0w1) = 0w0 then acc else acc + 1 + in + f (Word.>> (s,0w1)) acc + end + in + SOME (f (setN sz x) 0) + end; + + fun complementFn sz x = SOME (Word.toInt (Word.xorb (univN sz, setN sz x))); - val univ = IntSet.fromList (interval 0 M) + fun differenceFn sz x y = + let + val x = setN sz x + and y = setN sz y + in + SOME (Word.toInt (Word.andb (x, Word.notb y))) + end; + + fun emptyFn _ = SOME 0; + + fun insertFn sz x y = + let + val x = x mod eltN sz + and y = setN sz y + in + SOME (Word.toInt (Word.orb (posN x, y))) + end; + + fun intersectFn sz x y = + SOME (Word.toInt (Word.andb (setN sz x, setN sz y))); + + fun singletonFn sz x = + let + val x = x mod eltN sz + in + SOME (Word.toInt (posN x)) + end; + + fun symmetricDifferenceFn sz x y = + let + val x = setN sz x + and y = setN sz y + in + SOME (Word.toInt (Word.xorb (x,y))) + end; + + fun unionFn sz x y = + SOME (Word.toInt (Word.orb (setN sz x, setN sz y))); + + fun universeFn sz = SOME (Word.toInt (univN sz)); + + (* Relations *) + + fun memberRel sz x y = + let + val x = x mod eltN sz + and y = setN sz y + in + SOME (Word.andb (posN x, y) <> 0w0) + end; - val mk_set = + fun subsetRel sz x y = + let + val x = setN sz x + and y = setN sz y + in + SOME (Word.andb (x, Word.notb y) = 0w0) + end; +in + val setFixed = + let + val fns = + NameArityMap.fromList + [((cardName,1), fixed1 cardFn), + ((complementName,1), fixed1 complementFn), + ((differenceName,2), fixed2 differenceFn), + ((emptyName,0), fixed0 emptyFn), + ((insertName,2), fixed2 insertFn), + ((intersectName,2), fixed2 intersectFn), + ((singletonName,1), fixed1 singletonFn), + ((symmetricDifferenceName,2), fixed2 symmetricDifferenceFn), + ((unionName,2), fixed2 unionFn), + ((universeName,0), fixed0 universeFn)] + + val rels = + NameArityMap.fromList + [((memberName,2), fixed2 memberRel), + ((subsetName,2), fixed2 subsetRel)] + in + Fixed + {functions = fns, + relations = rels} + end; +end; + +(* Lists *) + +val appendName = Name.fromString "@" +and consName = Name.fromString "::" +and lengthName = Name.fromString "length" +and nilName = Name.fromString "nil" +and nullName = Name.fromString "null" +and tailName = Name.fromString "tail"; + +local + val baseFix = + let + val fix = unionFixed projectionFixed overflowFixed + + val sucFn = getFunctionFixed fix (sucName,1) + + fun suc2Fn sz _ x = sucFn sz [x] + in + insertFunctionFixed fix ((sucName,2), fixed2 suc2Fn) + end; + + val fixMap = + {functionMap = NameArityMap.fromList + [((appendName,2),addName), + ((consName,2),sucName), + ((lengthName,1), projectionName 1), + ((nilName,0), numeralName 0), + ((tailName,1),preName)], + relationMap = NameArityMap.fromList + [((nullName,1),isZeroName)]}; + +in + val listFixed = mapFixed fixMap baseFix; +end; + +(* ------------------------------------------------------------------------- *) +(* Valuations. *) +(* ------------------------------------------------------------------------- *) + +datatype valuation = Valuation of element NameMap.map; + +val emptyValuation = Valuation (NameMap.new ()); + +fun insertValuation (Valuation m) v_i = Valuation (NameMap.insert m v_i); + +fun peekValuation (Valuation m) v = NameMap.peek m v; + +fun constantValuation i = + let + fun add (v,V) = insertValuation V (v,i) + in + NameSet.foldl add emptyValuation + end; + +val zeroValuation = constantValuation zeroElement; + +fun getValuation V v = + case peekValuation V v of + SOME i => i + | NONE => raise Error "Model.getValuation: incomplete valuation"; + +fun randomValuation {size = N} vs = + let + fun f (v,V) = insertValuation V (v, Portable.randomInt N) + in + NameSet.foldl f emptyValuation vs + end; + +fun incrementValuation N vars = + let + fun inc vs V = + case vs of + [] => NONE + | v :: vs => + let + val (carry,i) = + case incrementElement N (getValuation V v) of + SOME i => (false,i) + | NONE => (true,zeroElement) + + val V = insertValuation V (v,i) + in + if carry then inc vs V else SOME V + end + in + inc (NameSet.toList vars) + end; + +fun foldValuation N vars f = + let + val inc = incrementValuation N vars + + fun fold V acc = let - fun f _ s 0 = s - | f k s x = - let - val s = if x mod 2 = 0 then s else IntSet.add s k - in - f (k + 1) s (x div 2) - end + val acc = f (V,acc) in - f 0 IntSet.empty + case inc V of + NONE => acc + | SOME V => fold V acc end - fun dest_set s = - let - fun f 0 x = x - | f k x = - let - val k = k - 1 - in - f k (if IntSet.member k s then 2 * x + 1 else 2 * x) - end + val zero = zeroValuation vars + in + fold zero + end; + +(* ------------------------------------------------------------------------- *) +(* A type of random finite mapping Z^n -> Z. *) +(* ------------------------------------------------------------------------- *) + +val UNKNOWN = ~1; + +datatype table = + ForgetfulTable + | ArrayTable of int Array.array; + +fun newTable N arity = + case elementListSpace {size = N} arity of + NONE => ForgetfulTable + | SOME space => ArrayTable (Array.array (space,UNKNOWN)); - val x = case IntSet.findr (K true) s of NONE => 0 | SOME k => f k 1 - in - if x < N then SOME x else NONE - end +local + fun randomResult R = Portable.randomInt R; +in + fun lookupTable N R table elts = + case table of + ForgetfulTable => randomResult R + | ArrayTable a => + let + val i = elementListIndex {size = N} elts + + val r = Array.sub (a,i) + in + if r <> UNKNOWN then r + else + let + val r = randomResult R + + val () = Array.update (a,i,r) + in + r + end + end; +end; + +fun updateTable N table (elts,r) = + case table of + ForgetfulTable => () + | ArrayTable a => + let + val i = elementListIndex {size = N} elts + + val () = Array.update (a,i,r) + in + () + end; - fun functions ("empty",[]) = dest_set IntSet.empty - | functions ("univ",[]) = dest_set univ - | functions ("union",[x,y]) = - dest_set (IntSet.union (mk_set x) (mk_set y)) - | functions ("intersect",[x,y]) = - dest_set (IntSet.intersect (mk_set x) (mk_set y)) - | functions ("compl",[x]) = - dest_set (IntSet.difference univ (mk_set x)) - | functions ("card",[x]) = SOME (IntSet.size (mk_set x)) - | functions _ = NONE - - fun relations ("in",[x,y]) = SOME (IntSet.member (x mod M) (mk_set y)) - | relations ("subset",[x,y]) = - SOME (IntSet.subset (mk_set x) (mk_set y)) - | relations _ = NONE +(* ------------------------------------------------------------------------- *) +(* A type of random finite mappings name * arity -> Z^arity -> Z. *) +(* ------------------------------------------------------------------------- *) + +datatype tables = + Tables of + {domainSize : int, + rangeSize : int, + tableMap : table NameArityMap.map ref}; + +fun newTables N R = + Tables + {domainSize = N, + rangeSize = R, + tableMap = ref (NameArityMap.new ())}; + +fun getTables tables n_a = + let + val Tables {domainSize = N, rangeSize = _, tableMap = tm} = tables + + val ref m = tm in - {functions = functions, relations = relations} + case NameArityMap.peek m n_a of + SOME t => t + | NONE => + let + val (_,a) = n_a + + val t = newTable N a + + val m = NameArityMap.insert m (n_a,t) + + val () = tm := m + in + t + end + end; + +fun lookupTables tables (n,elts) = + let + val Tables {domainSize = N, rangeSize = R, ...} = tables + + val a = length elts + + val table = getTables tables (n,a) + in + lookupTable N R table elts + end; + +fun updateTables tables ((n,elts),r) = + let + val Tables {domainSize = N, ...} = tables + + val a = length elts + + val table = getTables tables (n,a) + in + updateTable N table (elts,r) end; (* ------------------------------------------------------------------------- *) @@ -392,166 +926,153 @@ datatype model = Model of {size : int, - fixed : fixedModel, - functions : (Term.functionName * int list, int) Map.map ref, - relations : (Atom.relationName * int list, bool) Map.map ref}; + fixedFunctions : (element list -> element option) NameArityMap.map, + fixedRelations : (element list -> bool option) NameArityMap.map, + randomFunctions : tables, + randomRelations : tables}; -local - fun cmp ((n1,l1),(n2,l2)) = - case String.compare (n1,n2) of - LESS => LESS - | EQUAL => lexCompare Int.compare (l1,l2) - | GREATER => GREATER; -in - fun new {size = N, fixed} = +fun new {size = N, fixed} = + let + val Fixed {functions = fns, relations = rels} = fixed + + val fixFns = NameArityMap.transform (fn f => f {size = N}) fns + and fixRels = NameArityMap.transform (fn r => r {size = N}) rels + + val rndFns = newTables N N + and rndRels = newTables N 2 + in Model {size = N, - fixed = fixed {size = N}, - functions = ref (Map.new cmp), - relations = ref (Map.new cmp)}; -end; + fixedFunctions = fixFns, + fixedRelations = fixRels, + randomFunctions = rndFns, + randomRelations = rndRels} + end; + +fun size (Model {size = N, ...}) = N; + +fun peekFixedFunction M (n,elts) = + let + val Model {fixedFunctions = fixFns, ...} = M + in + case NameArityMap.peek fixFns (n, length elts) of + NONE => NONE + | SOME fixFn => fixFn elts + end; + +fun isFixedFunction M n_elts = Option.isSome (peekFixedFunction M n_elts); -fun size (Model {size = s, ...}) = s; +fun peekFixedRelation M (n,elts) = + let + val Model {fixedRelations = fixRels, ...} = M + in + case NameArityMap.peek fixRels (n, length elts) of + NONE => NONE + | SOME fixRel => fixRel elts + end; + +fun isFixedRelation M n_elts = Option.isSome (peekFixedRelation M n_elts); + +(* A default model *) + +val defaultSize = 8; + +val defaultFixed = + unionListFixed + [basicFixed, + projectionFixed, + modularFixed, + setFixed, + listFixed]; + +val default = {size = defaultSize, fixed = defaultFixed}; (* ------------------------------------------------------------------------- *) -(* Valuations. *) +(* Taking apart terms to interpret them. *) (* ------------------------------------------------------------------------- *) -type valuation = int NameMap.map; - -val valuationEmpty : valuation = NameMap.new (); - -fun valuationRandom {size = N} vs = - let - fun f (v,V) = NameMap.insert V (v, Portable.randomInt N) - in - NameSet.foldl f valuationEmpty vs - end; - -fun valuationFold {size = N} vs f = - let - val vs = NameSet.toList vs - - fun inc [] _ = NONE - | inc (v :: l) V = - case NameMap.peek V v of - NONE => raise Bug "Model.valuationFold" - | SOME k => - let - val k = if k = N - 1 then 0 else k + 1 - val V = NameMap.insert V (v,k) - in - if k = 0 then inc l V else SOME V - end - - val zero = foldl (fn (v,V) => NameMap.insert V (v,0)) valuationEmpty vs - - fun fold V acc = - let - val acc = f (V,acc) - in - case inc vs V of NONE => acc | SOME V => fold V acc - end - in - fold zero - end; +fun destTerm tm = + case tm of + Term.Var _ => tm + | Term.Fn f_tms => + case Term.stripApp tm of + (_,[]) => tm + | (v as Term.Var _, tms) => Term.Fn (Term.appName, v :: tms) + | (Term.Fn (f,tms), tms') => Term.Fn (f, tms @ tms'); (* ------------------------------------------------------------------------- *) (* Interpreting terms and formulas in the model. *) (* ------------------------------------------------------------------------- *) +fun interpretFunction M n_elts = + case peekFixedFunction M n_elts of + SOME r => r + | NONE => + let + val Model {randomFunctions = rndFns, ...} = M + in + lookupTables rndFns n_elts + end; + +fun interpretRelation M n_elts = + case peekFixedRelation M n_elts of + SOME r => r + | NONE => + let + val Model {randomRelations = rndRels, ...} = M + in + intToBool (lookupTables rndRels n_elts) + end; + fun interpretTerm M V = let - val Model {size = N, fixed, functions, ...} = M - val {functions = fixed_functions, ...} = fixed - - fun interpret (Term.Var v) = - (case NameMap.peek V v of - NONE => raise Error "Model.interpretTerm: incomplete valuation" - | SOME i => i) - | interpret (tm as Term.Fn f_tms) = - let - val (f,tms) = - case Term.stripComb tm of - (_,[]) => f_tms - | (v as Term.Var _, tms) => (".", v :: tms) - | (Term.Fn (f,tms), tms') => (f, tms @ tms') - val elts = map interpret tms - val f_elts = (f,elts) - val ref funcs = functions - in - case Map.peek funcs f_elts of - SOME k => k - | NONE => - let - val k = - case fixed_functions f_elts of - SOME k => k - | NONE => Portable.randomInt N - - val () = functions := Map.insert funcs (f_elts,k) - in - k - end - end; + fun interpret tm = + case destTerm tm of + Term.Var v => getValuation V v + | Term.Fn (f,tms) => interpretFunction M (f, map interpret tms) in interpret end; fun interpretAtom M V (r,tms) = - let - val Model {fixed,relations,...} = M - val {relations = fixed_relations, ...} = fixed - - val elts = map (interpretTerm M V) tms - val r_elts = (r,elts) - val ref rels = relations - in - case Map.peek rels r_elts of - SOME b => b - | NONE => - let - val b = - case fixed_relations r_elts of - SOME b => b - | NONE => Portable.randomBool () - - val () = relations := Map.insert rels (r_elts,b) - in - b - end - end; + interpretRelation M (r, map (interpretTerm M V) tms); fun interpretFormula M = let - val Model {size = N, ...} = M + val N = size M - fun interpret _ Formula.True = true - | interpret _ Formula.False = false - | interpret V (Formula.Atom atm) = interpretAtom M V atm - | interpret V (Formula.Not p) = not (interpret V p) - | interpret V (Formula.Or (p,q)) = interpret V p orelse interpret V q - | interpret V (Formula.And (p,q)) = interpret V p andalso interpret V q - | interpret V (Formula.Imp (p,q)) = - interpret V (Formula.Or (Formula.Not p, q)) - | interpret V (Formula.Iff (p,q)) = interpret V p = interpret V q - | interpret V (Formula.Forall (v,p)) = interpret' V v p N - | interpret V (Formula.Exists (v,p)) = - interpret V (Formula.Not (Formula.Forall (v, Formula.Not p))) - and interpret' _ _ _ 0 = true - | interpret' V v p i = + fun interpret V fm = + case fm of + Formula.True => true + | Formula.False => false + | Formula.Atom atm => interpretAtom M V atm + | Formula.Not p => not (interpret V p) + | Formula.Or (p,q) => interpret V p orelse interpret V q + | Formula.And (p,q) => interpret V p andalso interpret V q + | Formula.Imp (p,q) => interpret V (Formula.Or (Formula.Not p, q)) + | Formula.Iff (p,q) => interpret V p = interpret V q + | Formula.Forall (v,p) => interpret' V p v N + | Formula.Exists (v,p) => + interpret V (Formula.Not (Formula.Forall (v, Formula.Not p))) + + and interpret' V fm v i = + i = 0 orelse let val i = i - 1 - val V' = NameMap.insert V (v,i) + val V' = insertValuation V (v,i) in - interpret V' p andalso interpret' V v p i + interpret V' fm andalso interpret' V fm v i end in interpret end; -fun interpretLiteral M V (true,atm) = interpretAtom M V atm - | interpretLiteral M V (false,atm) = not (interpretAtom M V atm); +fun interpretLiteral M V (pol,atm) = + let + val b = interpretAtom M V atm + in + if pol then b else not b + end; fun interpretClause M V cl = LiteralSet.exists (interpretLiteral M V) cl; @@ -560,32 +1081,198 @@ (* Note: if it's cheaper, a systematic check will be performed instead. *) (* ------------------------------------------------------------------------- *) -local - fun checkGen freeVars interpret {maxChecks} M x = - let - val Model {size = N, ...} = M - - fun score (V,{T,F}) = - if interpret M V x then {T = T + 1, F = F} else {T = T, F = F + 1} +fun check interpret {maxChecks} M fv x = + let + val N = size M + + fun score (V,{T,F}) = + if interpret M V x then {T = T + 1, F = F} else {T = T, F = F + 1} + + fun randomCheck acc = score (randomValuation {size = N} fv, acc) + + val maxChecks = + case maxChecks of + NONE => maxChecks + | SOME m => + case expInt N (NameSet.size fv) of + SOME n => if n <= m then NONE else maxChecks + | NONE => maxChecks + in + case maxChecks of + SOME m => funpow m randomCheck {T = 0, F = 0} + | NONE => foldValuation {size = N} fv score {T = 0, F = 0} + end; + +fun checkAtom maxChecks M atm = + check interpretAtom maxChecks M (Atom.freeVars atm) atm; + +fun checkFormula maxChecks M fm = + check interpretFormula maxChecks M (Formula.freeVars fm) fm; + +fun checkLiteral maxChecks M lit = + check interpretLiteral maxChecks M (Literal.freeVars lit) lit; + +fun checkClause maxChecks M cl = + check interpretClause maxChecks M (LiteralSet.freeVars cl) cl; + +(* ------------------------------------------------------------------------- *) +(* Updating the model. *) +(* ------------------------------------------------------------------------- *) + +fun updateFunction M func_elts_elt = + let + val Model {randomFunctions = rndFns, ...} = M - val vs = freeVars x + val () = updateTables rndFns func_elts_elt + in + () + end; + +fun updateRelation M (rel_elts,pol) = + let + val Model {randomRelations = rndRels, ...} = M + + val () = updateTables rndRels (rel_elts, boolToInt pol) + in + () + end; + +(* ------------------------------------------------------------------------- *) +(* A type of terms with interpretations embedded in the subterms. *) +(* ------------------------------------------------------------------------- *) + +datatype modelTerm = + ModelVar + | ModelFn of Term.functionName * modelTerm list * int list; - fun randomCheck acc = score (valuationRandom {size = N} vs, acc) +fun modelTerm M V = + let + fun modelTm tm = + case destTerm tm of + Term.Var v => (ModelVar, getValuation V v) + | Term.Fn (f,tms) => + let + val (tms,xs) = unzip (map modelTm tms) + in + (ModelFn (f,tms,xs), interpretFunction M (f,xs)) + end + in + modelTm + end; + +(* ------------------------------------------------------------------------- *) +(* Perturbing the model. *) +(* ------------------------------------------------------------------------- *) + +datatype perturbation = + FunctionPerturbation of (Term.functionName * element list) * element + | RelationPerturbation of (Atom.relationName * element list) * bool; - val small = - intExp N (NameSet.size vs) <= maxChecks handle Overflow => false +fun perturb M pert = + case pert of + FunctionPerturbation func_elts_elt => updateFunction M func_elts_elt + | RelationPerturbation rel_elts_pol => updateRelation M rel_elts_pol; + +local + fun pertTerm _ [] _ acc = acc + | pertTerm M target tm acc = + case tm of + ModelVar => acc + | ModelFn (func,tms,xs) => + let + fun onTarget ys = mem (interpretFunction M (func,ys)) target + + val func_xs = (func,xs) + + val acc = + if isFixedFunction M func_xs then acc + else + let + fun add (y,acc) = FunctionPerturbation (func_xs,y) :: acc + in + foldl add acc target + end + in + pertTerms M onTarget tms xs acc + end + + and pertTerms M onTarget = + let + val N = size M + + fun filterElements pred = + let + fun filt 0 acc = acc + | filt i acc = + let + val i = i - 1 + val acc = if pred i then i :: acc else acc + in + filt i acc + end + in + filt N [] + end + + fun pert _ [] [] acc = acc + | pert ys (tm :: tms) (x :: xs) acc = + let + fun pred y = + y <> x andalso onTarget (List.revAppend (ys, y :: xs)) + + val target = filterElements pred + + val acc = pertTerm M target tm acc + in + pert (x :: ys) tms xs acc + end + | pert _ _ _ _ = raise Bug "Model.pertTerms.pert" in - if small then valuationFold {size = N} vs score {T = 0, F = 0} - else funpow maxChecks randomCheck {T = 0, F = 0} + pert [] end; + + fun pertAtom M V target (rel,tms) acc = + let + fun onTarget ys = interpretRelation M (rel,ys) = target + + val (tms,xs) = unzip (map (modelTerm M V) tms) + + val rel_xs = (rel,xs) + + val acc = + if isFixedRelation M rel_xs then acc + else RelationPerturbation (rel_xs,target) :: acc + in + pertTerms M onTarget tms xs acc + end; + + fun pertLiteral M V ((pol,atm),acc) = pertAtom M V pol atm acc; + + fun pertClause M V cl acc = LiteralSet.foldl (pertLiteral M V) acc cl; + + fun pickPerturb M perts = + if null perts then () + else perturb M (List.nth (perts, Portable.randomInt (length perts))); in - val checkAtom = checkGen Atom.freeVars interpretAtom; - - val checkFormula = checkGen Formula.freeVars interpretFormula; + fun perturbTerm M V (tm,target) = + pickPerturb M (pertTerm M target (fst (modelTerm M V tm)) []); - val checkLiteral = checkGen Literal.freeVars interpretLiteral; + fun perturbAtom M V (atm,target) = + pickPerturb M (pertAtom M V target atm []); - val checkClause = checkGen LiteralSet.freeVars interpretClause; + fun perturbLiteral M V lit = pickPerturb M (pertLiteral M V (lit,[])); + + fun perturbClause M V cl = pickPerturb M (pertClause M V cl []); end; +(* ------------------------------------------------------------------------- *) +(* Pretty printing. *) +(* ------------------------------------------------------------------------- *) + +fun pp M = + Print.program + [Print.addString "Model{", + Print.ppInt (size M), + Print.addString "}"]; + end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Name.sig --- a/src/Tools/Metis/src/Name.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Name.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,15 +1,45 @@ (* ========================================================================= *) (* NAMES *) -(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Name = sig -type name = string +(* ------------------------------------------------------------------------- *) +(* A type of names. *) +(* ------------------------------------------------------------------------- *) + +type name + +(* ------------------------------------------------------------------------- *) +(* A total ordering. *) +(* ------------------------------------------------------------------------- *) val compare : name * name -> order -val pp : name Parser.pp +val equal : name -> name -> bool + +(* ------------------------------------------------------------------------- *) +(* Fresh names. *) +(* ------------------------------------------------------------------------- *) + +val newName : unit -> name + +val newNames : int -> name list + +val variantPrime : (name -> bool) -> name -> name + +val variantNum : (name -> bool) -> name -> name + +(* ------------------------------------------------------------------------- *) +(* Parsing and pretty printing. *) +(* ------------------------------------------------------------------------- *) + +val pp : name Print.pp + +val toString : name -> string + +val fromString : string -> name end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Name.sml --- a/src/Tools/Metis/src/Name.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Name.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,85 +1,98 @@ (* ========================================================================= *) (* NAMES *) -(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Name :> Name = struct +open Useful; + +(* ------------------------------------------------------------------------- *) +(* A type of names. *) +(* ------------------------------------------------------------------------- *) + type name = string; +(* ------------------------------------------------------------------------- *) +(* A total ordering. *) +(* ------------------------------------------------------------------------- *) + val compare = String.compare; -val pp = Parser.ppString; +fun equal n1 n2 = n1 = n2; + +(* ------------------------------------------------------------------------- *) +(* Fresh variables. *) +(* ------------------------------------------------------------------------- *) + +local + val prefix = "_"; + + fun numName i = mkPrefix prefix (Int.toString i); +in + fun newName () = numName (newInt ()); + + fun newNames n = map numName (newInts n); +end; + +fun variantPrime acceptable = + let + fun variant n = if acceptable n then n else variant (n ^ "'") + in + variant + end; + +local + fun isDigitOrPrime #"'" = true + | isDigitOrPrime c = Char.isDigit c; +in + fun variantNum acceptable n = + if acceptable n then n + else + let + val n = stripSuffix isDigitOrPrime n + + fun variant i = + let + val n_i = n ^ Int.toString i + in + if acceptable n_i then n_i else variant (i + 1) + end + in + variant 0 + end; +end; + +(* ------------------------------------------------------------------------- *) +(* Parsing and pretty printing. *) +(* ------------------------------------------------------------------------- *) + +val pp = Print.ppString; + +fun toString s : string = s; + +fun fromString s : name = s; end structure NameOrdered = struct type t = Name.name val compare = Name.compare end +structure NameMap = KeyMap (NameOrdered); + structure NameSet = struct local - structure S = ElementSet (NameOrdered); + structure S = ElementSet (NameMap); in open S; end; val pp = - Parser.ppMap + Print.ppMap toList - (Parser.ppBracket "{" "}" (Parser.ppSequence "," Name.pp)); + (Print.ppBracket "{" "}" (Print.ppOpList "," Name.pp)); end - -structure NameMap = KeyMap (NameOrdered); - -structure NameArity = -struct - -type nameArity = Name.name * int; - -fun name ((n,_) : nameArity) = n; - -fun arity ((_,i) : nameArity) = i; - -fun nary i n_i = arity n_i = i; - -val nullary = nary 0 -and unary = nary 1 -and binary = nary 2 -and ternary = nary 3; - -fun compare ((n1,i1),(n2,i2)) = - case Name.compare (n1,n2) of - LESS => LESS - | EQUAL => Int.compare (i1,i2) - | GREATER => GREATER; - -val pp = Parser.ppMap (fn (n,i) => n ^ "/" ^ Int.toString i) Parser.ppString; - -end - -structure NameArityOrdered = -struct type t = NameArity.nameArity val compare = NameArity.compare end - -structure NameAritySet = -struct - - local - structure S = ElementSet (NameArityOrdered); - in - open S; - end; - - val allNullary = all NameArity.nullary; - - val pp = - Parser.ppMap - toList - (Parser.ppBracket "{" "}" (Parser.ppSequence "," NameArity.pp)); - -end - -structure NameArityMap = KeyMap (NameArityOrdered); diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/NameArity.sig --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/Metis/src/NameArity.sig Mon Sep 13 21:24:10 2010 +0200 @@ -0,0 +1,47 @@ +(* ========================================================================= *) +(* NAME/ARITY PAIRS *) +(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) +(* ========================================================================= *) + +signature NameArity = +sig + +(* ------------------------------------------------------------------------- *) +(* A type of name/arity pairs. *) +(* ------------------------------------------------------------------------- *) + +type nameArity = Name.name * int + +val name : nameArity -> Name.name + +val arity : nameArity -> int + +(* ------------------------------------------------------------------------- *) +(* Testing for different arities. *) +(* ------------------------------------------------------------------------- *) + +val nary : int -> nameArity -> bool + +val nullary : nameArity -> bool + +val unary : nameArity -> bool + +val binary : nameArity -> bool + +val ternary : nameArity -> bool + +(* ------------------------------------------------------------------------- *) +(* A total ordering. *) +(* ------------------------------------------------------------------------- *) + +val compare : nameArity * nameArity -> order + +val equal : nameArity -> nameArity -> bool + +(* ------------------------------------------------------------------------- *) +(* Parsing and pretty printing. *) +(* ------------------------------------------------------------------------- *) + +val pp : nameArity Print.pp + +end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/NameArity.sml --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/Metis/src/NameArity.sml Mon Sep 13 21:24:10 2010 +0200 @@ -0,0 +1,91 @@ +(* ========================================================================= *) +(* NAME/ARITY PAIRS *) +(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) +(* ========================================================================= *) + +structure NameArity :> NameArity = +struct + +(* ------------------------------------------------------------------------- *) +(* A type of name/arity pairs. *) +(* ------------------------------------------------------------------------- *) + +type nameArity = Name.name * int; + +fun name ((n,_) : nameArity) = n; + +fun arity ((_,i) : nameArity) = i; + +(* ------------------------------------------------------------------------- *) +(* Testing for different arities. *) +(* ------------------------------------------------------------------------- *) + +fun nary i n_i = arity n_i = i; + +val nullary = nary 0 +and unary = nary 1 +and binary = nary 2 +and ternary = nary 3; + +(* ------------------------------------------------------------------------- *) +(* A total ordering. *) +(* ------------------------------------------------------------------------- *) + +fun compare ((n1,i1),(n2,i2)) = + case Name.compare (n1,n2) of + LESS => LESS + | EQUAL => Int.compare (i1,i2) + | GREATER => GREATER; + +fun equal (n1,i1) (n2,i2) = i1 = i2 andalso Name.equal n1 n2; + +(* ------------------------------------------------------------------------- *) +(* Parsing and pretty printing. *) +(* ------------------------------------------------------------------------- *) + +fun pp (n,i) = + Print.blockProgram Print.Inconsistent 0 + [Name.pp n, + Print.addString "/", + Print.ppInt i]; + +end + +structure NameArityOrdered = +struct type t = NameArity.nameArity val compare = NameArity.compare end + +structure NameArityMap = +struct + + local + structure S = KeyMap (NameArityOrdered); + in + open S; + end; + + fun compose m1 m2 = + let + fun pk ((_,a),n) = peek m2 (n,a) + in + mapPartial pk m1 + end; + +end + +structure NameAritySet = +struct + + local + structure S = ElementSet (NameArityMap); + in + open S; + end; + + val allNullary = all NameArity.nullary; + + val pp = + Print.ppMap + toList + (Print.ppBracket "{" "}" (Print.ppOpList "," NameArity.pp)); + +end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Normalize.sig --- a/src/Tools/Metis/src/Normalize.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Normalize.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* NORMALIZING FORMULAS *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2009 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Normalize = @@ -13,9 +13,43 @@ val nnf : Formula.formula -> Formula.formula (* ------------------------------------------------------------------------- *) +(* Conjunctive normal form derivations. *) +(* ------------------------------------------------------------------------- *) + +type thm + +datatype inference = + Axiom of Formula.formula + | Definition of string * Formula.formula + | Simplify of thm * thm list + | Conjunct of thm + | Specialize of thm + | Skolemize of thm + | Clausify of thm + +val mkAxiom : Formula.formula -> thm + +val destThm : thm -> Formula.formula * inference + +val proveThms : + thm list -> (Formula.formula * inference * Formula.formula list) list + +val toStringInference : inference -> string + +val ppInference : inference Print.pp + +(* ------------------------------------------------------------------------- *) (* Conjunctive normal form. *) (* ------------------------------------------------------------------------- *) -val cnf : Formula.formula -> Formula.formula list +type cnf + +val initialCnf : cnf + +val addCnf : thm -> cnf -> (Thm.clause * thm) list * cnf + +val proveCnf : thm list -> (Thm.clause * thm) list + +val cnf : Formula.formula -> Thm.clause list end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Normalize.sml --- a/src/Tools/Metis/src/Normalize.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Normalize.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* NORMALIZING FORMULAS *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Normalize :> Normalize = @@ -9,38 +9,102 @@ open Useful; (* ------------------------------------------------------------------------- *) +(* Constants. *) +(* ------------------------------------------------------------------------- *) + +val prefix = "FOFtoCNF"; + +val skolemPrefix = "skolem" ^ prefix; + +val definitionPrefix = "definition" ^ prefix; + +(* ------------------------------------------------------------------------- *) +(* Storing huge real numbers as their log. *) +(* ------------------------------------------------------------------------- *) + +datatype logReal = LogReal of real; + +fun compareLogReal (LogReal logX, LogReal logY) = + Real.compare (logX,logY); + +val zeroLogReal = LogReal ~1.0; + +val oneLogReal = LogReal 0.0; + +local + fun isZero logX = logX < 0.0; + + (* Assume logX >= logY >= 0.0 *) + fun add logX logY = logX + Math.ln (1.0 + Math.exp (logY - logX)); +in + fun isZeroLogReal (LogReal logX) = isZero logX; + + fun multiplyLogReal (LogReal logX) (LogReal logY) = + if isZero logX orelse isZero logY then zeroLogReal + else LogReal (logX + logY); + + fun addLogReal (lx as LogReal logX) (ly as LogReal logY) = + if isZero logX then ly + else if isZero logY then lx + else if logX < logY then LogReal (add logY logX) + else LogReal (add logX logY); + + fun withinRelativeLogReal logDelta (LogReal logX) (LogReal logY) = + isZero logX orelse + (not (isZero logY) andalso logX < logY + logDelta); +end; + +fun toStringLogReal (LogReal logX) = Real.toString logX; + +(* ------------------------------------------------------------------------- *) (* Counting the clauses that would be generated by conjunctive normal form. *) (* ------------------------------------------------------------------------- *) -datatype count = Count of {positive : real, negative : real}; +val countLogDelta = 0.01; + +datatype count = Count of {positive : logReal, negative : logReal}; -fun positive (Count {positive = p, ...}) = p; - -fun negative (Count {negative = n, ...}) = n; +fun countCompare (count1,count2) = + let + val Count {positive = p1, negative = _} = count1 + and Count {positive = p2, negative = _} = count2 + in + compareLogReal (p1,p2) + end; fun countNegate (Count {positive = p, negative = n}) = Count {positive = n, negative = p}; -fun countEqualish count1 count2 = +fun countLeqish count1 count2 = let - val Count {positive = p1, negative = n1} = count1 - and Count {positive = p2, negative = n2} = count2 + val Count {positive = p1, negative = _} = count1 + and Count {positive = p2, negative = _} = count2 in - Real.abs (p1 - p2) < 0.5 andalso Real.abs (n1 - n2) < 0.5 + withinRelativeLogReal countLogDelta p1 p2 end; -val countTrue = Count {positive = 0.0, negative = 1.0}; +(*MetisDebug +fun countEqualish count1 count2 = + countLeqish count1 count2 andalso + countLeqish count2 count1; -val countFalse = Count {positive = 1.0, negative = 0.0}; +fun countEquivish count1 count2 = + countEqualish count1 count2 andalso + countEqualish (countNegate count1) (countNegate count2); +*) -val countLiteral = Count {positive = 1.0, negative = 1.0}; +val countTrue = Count {positive = zeroLogReal, negative = oneLogReal}; + +val countFalse = Count {positive = oneLogReal, negative = zeroLogReal}; + +val countLiteral = Count {positive = oneLogReal, negative = oneLogReal}; fun countAnd2 (count1,count2) = let val Count {positive = p1, negative = n1} = count1 and Count {positive = p2, negative = n2} = count2 - val p = p1 + p2 - and n = n1 * n2 + val p = addLogReal p1 p2 + and n = multiplyLogReal n1 n2 in Count {positive = p, negative = n} end; @@ -49,25 +113,36 @@ let val Count {positive = p1, negative = n1} = count1 and Count {positive = p2, negative = n2} = count2 - val p = p1 * p2 - and n = n1 + n2 + val p = multiplyLogReal p1 p2 + and n = addLogReal n1 n2 in Count {positive = p, negative = n} end; -(*** Is this associative? ***) +(* Whether countXor2 is associative or not is an open question. *) + fun countXor2 (count1,count2) = let val Count {positive = p1, negative = n1} = count1 and Count {positive = p2, negative = n2} = count2 - val p = p1 * p2 + n1 * n2 - and n = p1 * n2 + n1 * p2 + val p = addLogReal (multiplyLogReal p1 p2) (multiplyLogReal n1 n2) + and n = addLogReal (multiplyLogReal p1 n2) (multiplyLogReal n1 p2) in Count {positive = p, negative = n} end; fun countDefinition body_count = countXor2 (countLiteral,body_count); +val countToString = + let + val rToS = toStringLogReal + in + fn Count {positive = p, negative = n} => + "(+" ^ rToS p ^ ",-" ^ rToS n ^ ")" + end; + +val ppCount = Print.ppMap countToString Print.ppString; + (* ------------------------------------------------------------------------- *) (* A type of normalized formula. *) (* ------------------------------------------------------------------------- *) @@ -83,41 +158,43 @@ | Forall of NameSet.set * count * NameSet.set * formula; fun compare f1_f2 = - case f1_f2 of - (True,True) => EQUAL - | (True,_) => LESS - | (_,True) => GREATER - | (False,False) => EQUAL - | (False,_) => LESS - | (_,False) => GREATER - | (Literal (_,l1), Literal (_,l2)) => Literal.compare (l1,l2) - | (Literal _, _) => LESS - | (_, Literal _) => GREATER - | (And (_,_,s1), And (_,_,s2)) => Set.compare (s1,s2) - | (And _, _) => LESS - | (_, And _) => GREATER - | (Or (_,_,s1), Or (_,_,s2)) => Set.compare (s1,s2) - | (Or _, _) => LESS - | (_, Or _) => GREATER - | (Xor (_,_,p1,s1), Xor (_,_,p2,s2)) => - (case boolCompare (p1,p2) of - LESS => LESS - | EQUAL => Set.compare (s1,s2) - | GREATER => GREATER) - | (Xor _, _) => LESS - | (_, Xor _) => GREATER - | (Exists (_,_,n1,f1), Exists (_,_,n2,f2)) => - (case NameSet.compare (n1,n2) of - LESS => LESS - | EQUAL => compare (f1,f2) - | GREATER => GREATER) - | (Exists _, _) => LESS - | (_, Exists _) => GREATER - | (Forall (_,_,n1,f1), Forall (_,_,n2,f2)) => - (case NameSet.compare (n1,n2) of - LESS => LESS - | EQUAL => compare (f1,f2) - | GREATER => GREATER); + if Portable.pointerEqual f1_f2 then EQUAL + else + case f1_f2 of + (True,True) => EQUAL + | (True,_) => LESS + | (_,True) => GREATER + | (False,False) => EQUAL + | (False,_) => LESS + | (_,False) => GREATER + | (Literal (_,l1), Literal (_,l2)) => Literal.compare (l1,l2) + | (Literal _, _) => LESS + | (_, Literal _) => GREATER + | (And (_,_,s1), And (_,_,s2)) => Set.compare (s1,s2) + | (And _, _) => LESS + | (_, And _) => GREATER + | (Or (_,_,s1), Or (_,_,s2)) => Set.compare (s1,s2) + | (Or _, _) => LESS + | (_, Or _) => GREATER + | (Xor (_,_,p1,s1), Xor (_,_,p2,s2)) => + (case boolCompare (p1,p2) of + LESS => LESS + | EQUAL => Set.compare (s1,s2) + | GREATER => GREATER) + | (Xor _, _) => LESS + | (_, Xor _) => GREATER + | (Exists (_,_,n1,f1), Exists (_,_,n2,f2)) => + (case NameSet.compare (n1,n2) of + LESS => LESS + | EQUAL => compare (f1,f2) + | GREATER => GREATER) + | (Exists _, _) => LESS + | (_, Exists _) => GREATER + | (Forall (_,_,n1,f1), Forall (_,_,n2,f2)) => + (case NameSet.compare (n1,n2) of + LESS => LESS + | EQUAL => compare (f1,f2) + | GREATER => GREATER); val empty = Set.empty compare; @@ -161,7 +238,7 @@ | polarity (Exists _) = true | polarity (Forall _) = false; -(*DEBUG +(*MetisDebug val polarity = fn f => let val res1 = compare (f, negate f) = LESS @@ -255,7 +332,7 @@ end end; -val AndList = foldl And2 True; +val AndList = List.foldl And2 True; val AndSet = Set.foldl And2 True; @@ -291,7 +368,7 @@ end end; -val OrList = foldl Or2 False; +val OrList = List.foldl Or2 False; val OrSet = Set.foldl Or2 False; @@ -301,12 +378,13 @@ and s2 = case f2 of And (_,_,s) => s | _ => singleton f2 fun g x1 (x2,acc) = And2 (Or2 (x1,x2), acc) + fun f (x1,acc) = Set.foldl (g x1) acc s2 in Set.foldl f True s1 end; -val pushOrList = foldl pushOr2 False; +val pushOrList = List.foldl pushOr2 False; local fun normalize fm = @@ -344,7 +422,7 @@ end; end; -val XorList = foldl Xor2 False; +val XorList = List.foldl Xor2 False; val XorSet = Set.foldl Xor2 False; @@ -406,7 +484,7 @@ exists init_fm end; -fun ExistsList (vs,f) = foldl Exists1 f vs; +fun ExistsList (vs,f) = List.foldl Exists1 f vs; fun ExistsSet (n,f) = NameSet.foldl Exists1 f n; @@ -444,10 +522,12 @@ forall init_fm end; -fun ForallList (vs,f) = foldl Forall1 f vs; +fun ForallList (vs,f) = List.foldl Forall1 f vs; fun ForallSet (n,f) = NameSet.foldl Forall1 f n; +fun generalize f = ForallSet (freeVars f, f); + local fun subst_fv fvSub = let @@ -580,9 +660,20 @@ val toFormula = form; end; -val pp = Parser.ppMap toFormula Formula.pp; +fun toLiteral (Literal (_,lit)) = lit + | toLiteral _ = raise Error "Normalize.toLiteral"; -val toString = Parser.toString pp; +local + fun addLiteral (l,s) = LiteralSet.add s (toLiteral l); +in + fun toClause False = LiteralSet.empty + | toClause (Or (_,_,s)) = Set.foldl addLiteral LiteralSet.empty s + | toClause l = LiteralSet.singleton (toLiteral l); +end; + +val pp = Print.ppMap toFormula Formula.pp; + +val toString = Print.toString pp; (* ------------------------------------------------------------------------- *) (* Negation normal form. *) @@ -591,224 +682,30 @@ fun nnf fm = toFormula (fromFormula fm); (* ------------------------------------------------------------------------- *) -(* Simplifying with definitions. *) -(* ------------------------------------------------------------------------- *) - -datatype simplify = - Simplify of - {formula : (formula,formula) Map.map, - andSet : (formula Set.set * formula) list, - orSet : (formula Set.set * formula) list, - xorSet : (formula Set.set * formula) list}; - -val simplifyEmpty = - Simplify {formula = Map.new compare, andSet = [], orSet = [], xorSet = []}; - -local - fun simpler fm s = - Set.size s <> 1 orelse - case Set.pick s of - True => false - | False => false - | Literal _ => false - | _ => true; - - fun addSet set_defs body_def = - let - fun def_body_size (body,_) = Set.size body - - val body_size = def_body_size body_def - - val (body,_) = body_def - - fun add acc [] = List.revAppend (acc,[body_def]) - | add acc (l as (bd as (b,_)) :: bds) = - case Int.compare (def_body_size bd, body_size) of - LESS => List.revAppend (acc, body_def :: l) - | EQUAL => if Set.equal b body then List.revAppend (acc,l) - else add (bd :: acc) bds - | GREATER => add (bd :: acc) bds - in - add [] set_defs - end; - - fun add simp (body,False) = add simp (negate body, True) - | add simp (True,_) = simp - | add (Simplify {formula,andSet,orSet,xorSet}) (And (_,_,s), def) = - let - val andSet = addSet andSet (s,def) - and orSet = addSet orSet (negateSet s, negate def) - in - Simplify - {formula = formula, - andSet = andSet, orSet = orSet, xorSet = xorSet} - end - | add (Simplify {formula,andSet,orSet,xorSet}) (Or (_,_,s), def) = - let - val orSet = addSet orSet (s,def) - and andSet = addSet andSet (negateSet s, negate def) - in - Simplify - {formula = formula, - andSet = andSet, orSet = orSet, xorSet = xorSet} - end - | add simp (Xor (_,_,p,s), def) = - let - val simp = addXorSet simp (s, applyPolarity p def) - in - case def of - True => - let - fun addXorLiteral (fm as Literal _, simp) = - let - val s = Set.delete s fm - in - if not (simpler fm s) then simp - else addXorSet simp (s, applyPolarity (not p) fm) - end - | addXorLiteral (_,simp) = simp - in - Set.foldl addXorLiteral simp s - end - | _ => simp - end - | add (simp as Simplify {formula,andSet,orSet,xorSet}) (body,def) = - if Map.inDomain body formula then simp - else - let - val formula = Map.insert formula (body,def) - val formula = Map.insert formula (negate body, negate def) - in - Simplify - {formula = formula, - andSet = andSet, orSet = orSet, xorSet = xorSet} - end - - and addXorSet (simp as Simplify {formula,andSet,orSet,xorSet}) (s,def) = - if Set.size s = 1 then add simp (Set.pick s, def) - else - let - val xorSet = addSet xorSet (s,def) - in - Simplify - {formula = formula, - andSet = andSet, orSet = orSet, xorSet = xorSet} - end; -in - fun simplifyAdd simp fm = add simp (fm,True); -end; - -local - fun simplifySet set_defs set = - let - fun pred (s,_) = Set.subset s set - in - case List.find pred set_defs of - NONE => NONE - | SOME (s,f) => SOME (Set.add (Set.difference set s) f) - end; -in - fun simplify (Simplify {formula,andSet,orSet,xorSet}) = - let - fun simp fm = simp_top (simp_sub fm) - - and simp_top (changed_fm as (_, And (_,_,s))) = - (case simplifySet andSet s of - NONE => changed_fm - | SOME s => simp_top (true, AndSet s)) - | simp_top (changed_fm as (_, Or (_,_,s))) = - (case simplifySet orSet s of - NONE => changed_fm - | SOME s => simp_top (true, OrSet s)) - | simp_top (changed_fm as (_, Xor (_,_,p,s))) = - (case simplifySet xorSet s of - NONE => changed_fm - | SOME s => simp_top (true, XorPolaritySet (p,s))) - | simp_top (changed_fm as (_,fm)) = - (case Map.peek formula fm of - NONE => changed_fm - | SOME fm => simp_top (true,fm)) - - and simp_sub fm = - case fm of - And (_,_,s) => - let - val l = Set.transform simp s - val changed = List.exists fst l - val fm = if changed then AndList (map snd l) else fm - in - (changed,fm) - end - | Or (_,_,s) => - let - val l = Set.transform simp s - val changed = List.exists fst l - val fm = if changed then OrList (map snd l) else fm - in - (changed,fm) - end - | Xor (_,_,p,s) => - let - val l = Set.transform simp s - val changed = List.exists fst l - val fm = if changed then XorPolarityList (p, map snd l) else fm - in - (changed,fm) - end - | Exists (_,_,n,f) => - let - val (changed,f) = simp f - val fm = if changed then ExistsSet (n,f) else fm - in - (changed,fm) - end - | Forall (_,_,n,f) => - let - val (changed,f) = simp f - val fm = if changed then ForallSet (n,f) else fm - in - (changed,fm) - end - | _ => (false,fm); - in - fn fm => snd (simp fm) - end; -end; - -(*TRACE2 -val simplify = fn simp => fn fm => - let - val fm' = simplify simp fm - val () = if compare (fm,fm') = EQUAL then () - else (Parser.ppTrace pp "Normalize.simplify: fm" fm; - Parser.ppTrace pp "Normalize.simplify: fm'" fm') - in - fm' - end; -*) - -(* ------------------------------------------------------------------------- *) (* Basic conjunctive normal form. *) (* ------------------------------------------------------------------------- *) val newSkolemFunction = let - val counter : int NameMap.map ref = ref (NameMap.new ()) + val counter : int StringMap.map ref = ref (StringMap.new ()) in - fn n => CRITICAL (fn () => - let - val ref m = counter - val i = Option.getOpt (NameMap.peek m n, 0) - val () = counter := NameMap.insert m (n, i + 1) - in - "skolem_" ^ n ^ (if i = 0 then "" else "_" ^ Int.toString i) - end) + fn n => + let + val ref m = counter + val s = Name.toString n + val i = Option.getOpt (StringMap.peek m s, 0) + val () = counter := StringMap.insert m (s, i + 1) + val i = if i = 0 then "" else "_" ^ Int.toString i + val s = skolemPrefix ^ "_" ^ s ^ i + in + Name.fromString s + end end; fun skolemize fv bv fm = let val fv = NameSet.transform Term.Var fv - + fun mk (v,s) = Subst.insert s (v, Term.Fn (newSkolemFunction v, fv)) in subst (NameSet.foldl mk Subst.empty bv) fm @@ -828,7 +725,7 @@ in (NameSet.add a v', Subst.insert s (v, Term.Var v')) end - + val avoid = NameSet.union (NameSet.union avoid fv) bv val (_,sub) = NameSet.foldl ren (avoid,Subst.empty) captured @@ -837,35 +734,43 @@ end end; - fun cnf avoid fm = -(*TRACE5 + fun cnfFm avoid fm = +(*MetisTrace5 let - val fm' = cnf' avoid fm - val () = Parser.ppTrace pp "Normalize.cnf: fm" fm - val () = Parser.ppTrace pp "Normalize.cnf: fm'" fm' + val fm' = cnfFm' avoid fm + val () = Print.trace pp "Normalize.cnfFm: fm" fm + val () = Print.trace pp "Normalize.cnfFm: fm'" fm' in fm' end - and cnf' avoid fm = + and cnfFm' avoid fm = *) case fm of True => True | False => False | Literal _ => fm - | And (_,_,s) => AndList (Set.transform (cnf avoid) s) - | Or (_,_,s) => pushOrList (snd (Set.foldl cnfOr (avoid,[]) s)) - | Xor _ => cnf avoid (pushXor fm) - | Exists (fv,_,n,f) => cnf avoid (skolemize fv n f) - | Forall (fv,_,n,f) => cnf avoid (rename avoid fv n f) + | And (_,_,s) => AndList (Set.transform (cnfFm avoid) s) + | Or (fv,_,s) => + let + val avoid = NameSet.union avoid fv + val (fms,_) = Set.foldl cnfOr ([],avoid) s + in + pushOrList fms + end + | Xor _ => cnfFm avoid (pushXor fm) + | Exists (fv,_,n,f) => cnfFm avoid (skolemize fv n f) + | Forall (fv,_,n,f) => cnfFm avoid (rename avoid fv n f) - and cnfOr (fm,(avoid,acc)) = + and cnfOr (fm,(fms,avoid)) = let - val fm = cnf avoid fm + val fm = cnfFm avoid fm + val fms = fm :: fms + val avoid = NameSet.union avoid (freeVars fm) in - (NameSet.union (freeVars fm) avoid, fm :: acc) + (fms,avoid) end; in - val basicCnf = cnf NameSet.empty; + val basicCnf = cnfFm NameSet.empty; end; (* ------------------------------------------------------------------------- *) @@ -873,7 +778,7 @@ (* ------------------------------------------------------------------------- *) local - type best = real * formula option; + type best = count * formula option; fun minBreak countClauses fm best = case fm of @@ -888,7 +793,7 @@ minBreakSet countClauses countXor2 countFalse XorSet s best | Exists (_,_,_,f) => minBreak countClauses f best | Forall (_,_,_,f) => minBreak countClauses f best - + and minBreakSet countClauses count2 count0 mkSet fmSet best = let fun cumulatives fms = @@ -912,7 +817,11 @@ val (c1,_,fms) = foldl fwd (count0,empty,[]) fms val (c2,_,fms) = foldl bwd (count0,empty,[]) fms - val _ = countEqualish c1 c2 orelse raise Bug "cumulativeCounts" +(*MetisDebug + val _ = countEquivish c1 c2 orelse + raise Bug ("cumulativeCounts: c1 = " ^ countToString c1 ^ + ", c2 = " ^ countToString c2) +*) in fms end @@ -920,6 +829,7 @@ fun breakSing ((c1,_,fm,c2,_),best) = let val cFms = count2 (c1,c2) + fun countCls cFm = countClauses (count2 (cFms,cFm)) in minBreak countCls fm best @@ -931,13 +841,13 @@ if Set.null s1 then best else let - val cDef = countDefinition (count2 (c1, count fm)) + val cDef = countDefinition (countXor2 (c1, count fm)) val cFm = count2 (countLiteral,c2) - val cl = positive cDef + countClauses cFm - val better = cl < bcl - 0.5 + val cl = countAnd2 (cDef, countClauses cFm) + val noBetter = countLeqish bcl cl in - if better then (cl, SOME (mkSet (Set.add s1 fm))) - else best + if noBetter then best + else (cl, SOME (mkSet (Set.add s1 fm))) end in fn ((c1,s1,fm,c2,s2),best) => @@ -948,14 +858,14 @@ fun breakSet measure best = let - val fms = sortMap (measure o count) Real.compare fms + val fms = sortMap (measure o count) countCompare fms in foldl breakSet1 best (cumulatives fms) end val best = foldl breakSing best (cumulatives fms) - val best = breakSet positive best - val best = breakSet negative best + val best = breakSet I best + val best = breakSet countNegate best val best = breakSet countClauses best in best @@ -963,21 +873,20 @@ in fun minimumDefinition fm = let - val countClauses = positive - val cl = countClauses (count fm) + val cl = count fm in - if cl < 1.5 then NONE + if countLeqish cl countLiteral then NONE else let - val (cl',def) = minBreak countClauses fm (cl,NONE) -(*TRACE1 + val (cl',def) = minBreak I fm (cl,NONE) +(*MetisTrace1 val () = case def of NONE => () | SOME d => - Parser.ppTrace pp ("defCNF: before = " ^ Real.toString cl ^ - ", after = " ^ Real.toString cl' ^ - ", definition") d + Print.trace pp ("defCNF: before = " ^ countToString cl ^ + ", after = " ^ countToString cl' ^ + ", definition") d *) in def @@ -986,77 +895,492 @@ end; (* ------------------------------------------------------------------------- *) -(* Conjunctive normal form. *) +(* Conjunctive normal form derivations. *) +(* ------------------------------------------------------------------------- *) + +datatype thm = Thm of formula * inference + +and inference = + Axiom of Formula.formula + | Definition of string * Formula.formula + | Simplify of thm * thm list + | Conjunct of thm + | Specialize of thm + | Skolemize of thm + | Clausify of thm; + +fun parentsInference inf = + case inf of + Axiom _ => [] + | Definition _ => [] + | Simplify (th,ths) => th :: ths + | Conjunct th => [th] + | Specialize th => [th] + | Skolemize th => [th] + | Clausify th => [th]; + +fun compareThm (Thm (fm1,_), Thm (fm2,_)) = compare (fm1,fm2); + +fun parentsThm (Thm (_,inf)) = parentsInference inf; + +fun mkAxiom fm = Thm (fromFormula fm, Axiom fm); + +fun destThm (Thm (fm,inf)) = (toFormula fm, inf); + +local + val emptyProved : (thm,Formula.formula) Map.map = Map.new compareThm; + + fun isProved proved th = Map.inDomain th proved; + + fun isUnproved proved th = not (isProved proved th); + + fun lookupProved proved th = + case Map.peek proved th of + SOME fm => fm + | NONE => raise Bug "Normalize.lookupProved"; + + fun prove acc proved ths = + case ths of + [] => rev acc + | th :: ths' => + if isProved proved th then prove acc proved ths' + else + let + val pars = parentsThm th + + val deps = List.filter (isUnproved proved) pars + in + if null deps then + let + val (fm,inf) = destThm th + + val fms = map (lookupProved proved) pars + + val acc = (fm,inf,fms) :: acc + + val proved = Map.insert proved (th,fm) + in + prove acc proved ths' + end + else + let + val ths = deps @ ths + in + prove acc proved ths + end + end; +in + val proveThms = prove [] emptyProved; +end; + +fun toStringInference inf = + case inf of + Axiom _ => "Axiom" + | Definition _ => "Definition" + | Simplify _ => "Simplify" + | Conjunct _ => "Conjunct" + | Specialize _ => "Specialize" + | Skolemize _ => "Skolemize" + | Clausify _ => "Clausify"; + +val ppInference = Print.ppMap toStringInference Print.ppString; + +(* ------------------------------------------------------------------------- *) +(* Simplifying with definitions. *) +(* ------------------------------------------------------------------------- *) + +datatype simplify = + Simp of + {formula : (formula, formula * thm) Map.map, + andSet : (formula Set.set * formula * thm) list, + orSet : (formula Set.set * formula * thm) list, + xorSet : (formula Set.set * formula * thm) list}; + +val simplifyEmpty = + Simp + {formula = Map.new compare, + andSet = [], + orSet = [], + xorSet = []}; + +local + fun simpler fm s = + Set.size s <> 1 orelse + case Set.pick s of + True => false + | False => false + | Literal _ => false + | _ => true; + + fun addSet set_defs body_def = + let + fun def_body_size (body,_,_) = Set.size body + + val body_size = def_body_size body_def + + val (body,_,_) = body_def + + fun add acc [] = List.revAppend (acc,[body_def]) + | add acc (l as (bd as (b,_,_)) :: bds) = + case Int.compare (def_body_size bd, body_size) of + LESS => List.revAppend (acc, body_def :: l) + | EQUAL => + if Set.equal b body then List.revAppend (acc,l) + else add (bd :: acc) bds + | GREATER => add (bd :: acc) bds + in + add [] set_defs + end; + + fun add simp (body,False,th) = add simp (negate body, True, th) + | add simp (True,_,_) = simp + | add (Simp {formula,andSet,orSet,xorSet}) (And (_,_,s), def, th) = + let + val andSet = addSet andSet (s,def,th) + and orSet = addSet orSet (negateSet s, negate def, th) + in + Simp + {formula = formula, + andSet = andSet, + orSet = orSet, + xorSet = xorSet} + end + | add (Simp {formula,andSet,orSet,xorSet}) (Or (_,_,s), def, th) = + let + val orSet = addSet orSet (s,def,th) + and andSet = addSet andSet (negateSet s, negate def, th) + in + Simp + {formula = formula, + andSet = andSet, + orSet = orSet, + xorSet = xorSet} + end + | add simp (Xor (_,_,p,s), def, th) = + let + val simp = addXorSet simp (s, applyPolarity p def, th) + in + case def of + True => + let + fun addXorLiteral (fm as Literal _, simp) = + let + val s = Set.delete s fm + in + if not (simpler fm s) then simp + else addXorSet simp (s, applyPolarity (not p) fm, th) + end + | addXorLiteral (_,simp) = simp + in + Set.foldl addXorLiteral simp s + end + | _ => simp + end + | add (simp as Simp {formula,andSet,orSet,xorSet}) (body,def,th) = + if Map.inDomain body formula then simp + else + let + val formula = Map.insert formula (body,(def,th)) + val formula = Map.insert formula (negate body, (negate def, th)) + in + Simp + {formula = formula, + andSet = andSet, + orSet = orSet, + xorSet = xorSet} + end + + and addXorSet (simp as Simp {formula,andSet,orSet,xorSet}) (s,def,th) = + if Set.size s = 1 then add simp (Set.pick s, def, th) + else + let + val xorSet = addSet xorSet (s,def,th) + in + Simp + {formula = formula, + andSet = andSet, + orSet = orSet, + xorSet = xorSet} + end; +in + fun simplifyAdd simp (th as Thm (fm,_)) = add simp (fm,True,th); +end; + +local + fun simplifySet set_defs set = + let + fun pred (s,_,_) = Set.subset s set + in + case List.find pred set_defs of + NONE => NONE + | SOME (s,f,th) => + let + val set = Set.add (Set.difference set s) f + in + SOME (set,th) + end + end; +in + fun simplify (Simp {formula,andSet,orSet,xorSet}) = + let + fun simp fm inf = + case simp_sub fm inf of + NONE => simp_top fm inf + | SOME (fm,inf) => try_simp_top fm inf + + and try_simp_top fm inf = + case simp_top fm inf of + NONE => SOME (fm,inf) + | x => x + + and simp_top fm inf = + case fm of + And (_,_,s) => + (case simplifySet andSet s of + NONE => NONE + | SOME (s,th) => + let + val fm = AndSet s + val inf = th :: inf + in + try_simp_top fm inf + end) + | Or (_,_,s) => + (case simplifySet orSet s of + NONE => NONE + | SOME (s,th) => + let + val fm = OrSet s + val inf = th :: inf + in + try_simp_top fm inf + end) + | Xor (_,_,p,s) => + (case simplifySet xorSet s of + NONE => NONE + | SOME (s,th) => + let + val fm = XorPolaritySet (p,s) + val inf = th :: inf + in + try_simp_top fm inf + end) + | _ => + (case Map.peek formula fm of + NONE => NONE + | SOME (fm,th) => + let + val inf = th :: inf + in + try_simp_top fm inf + end) + + and simp_sub fm inf = + case fm of + And (_,_,s) => + (case simp_set s inf of + NONE => NONE + | SOME (l,inf) => SOME (AndList l, inf)) + | Or (_,_,s) => + (case simp_set s inf of + NONE => NONE + | SOME (l,inf) => SOME (OrList l, inf)) + | Xor (_,_,p,s) => + (case simp_set s inf of + NONE => NONE + | SOME (l,inf) => SOME (XorPolarityList (p,l), inf)) + | Exists (_,_,n,f) => + (case simp f inf of + NONE => NONE + | SOME (f,inf) => SOME (ExistsSet (n,f), inf)) + | Forall (_,_,n,f) => + (case simp f inf of + NONE => NONE + | SOME (f,inf) => SOME (ForallSet (n,f), inf)) + | _ => NONE + + and simp_set s inf = + let + val (changed,l,inf) = Set.foldr simp_set_elt (false,[],inf) s + in + if changed then SOME (l,inf) else NONE + end + + and simp_set_elt (fm,(changed,l,inf)) = + case simp fm inf of + NONE => (changed, fm :: l, inf) + | SOME (fm,inf) => (true, fm :: l, inf) + in + fn th as Thm (fm,_) => + case simp fm [] of + SOME (fm,ths) => + let + val inf = Simplify (th,ths) + in + Thm (fm,inf) + end + | NONE => th + end; +end; + +(*MetisTrace2 +val simplify = fn simp => fn th as Thm (fm,_) => + let + val th' as Thm (fm',_) = simplify simp th + val () = if compare (fm,fm') = EQUAL then () + else (Print.trace pp "Normalize.simplify: fm" fm; + Print.trace pp "Normalize.simplify: fm'" fm') + in + th' + end; +*) + +(* ------------------------------------------------------------------------- *) +(* Definitions. *) (* ------------------------------------------------------------------------- *) val newDefinitionRelation = let val counter : int ref = ref 0 in - fn () => CRITICAL (fn () => - let - val ref i = counter - val () = counter := i + 1 - in - "defCNF_" ^ Int.toString i - end) + fn () => + let + val ref i = counter + val () = counter := i + 1 + in + definitionPrefix ^ "_" ^ Int.toString i + end end; fun newDefinition def = let val fv = freeVars def - val atm = (newDefinitionRelation (), NameSet.transform Term.Var fv) - val lit = Literal (fv,(true,atm)) + val rel = newDefinitionRelation () + val atm = (Name.fromString rel, NameSet.transform Term.Var fv) + val fm = Formula.Iff (Formula.Atom atm, toFormula def) + val fm = Formula.setMkForall (fv,fm) + val inf = Definition (rel,fm) + val lit = Literal (fv,(false,atm)) + val fm = Xor2 (lit,def) in - Xor2 (lit,def) + Thm (fm,inf) end; +(* ------------------------------------------------------------------------- *) +(* Definitional conjunctive normal form. *) +(* ------------------------------------------------------------------------- *) + +datatype cnf = + ConsistentCnf of simplify + | InconsistentCnf; + +val initialCnf = ConsistentCnf simplifyEmpty; + local - fun def_cnf acc [] = acc - | def_cnf acc ((prob,simp,fms) :: probs) = - def_cnf_problem acc prob simp fms probs + fun def_cnf_inconsistent th = + let + val cls = [(LiteralSet.empty,th)] + in + (cls,InconsistentCnf) + end; - and def_cnf_problem acc prob _ [] probs = def_cnf (prob :: acc) probs - | def_cnf_problem acc prob simp (fm :: fms) probs = - def_cnf_formula acc prob simp (simplify simp fm) fms probs + fun def_cnf_clause inf (fm,acc) = + let + val cl = toClause fm + val th = Thm (fm,inf) + in + (cl,th) :: acc + end +(*MetisDebug + handle Error err => + (Print.trace pp "Normalize.addCnf.def_cnf_clause: fm" fm; + raise Bug ("Normalize.addCnf.def_cnf_clause: " ^ err)); +*) - and def_cnf_formula acc prob simp fm fms probs = + fun def_cnf cls simp ths = + case ths of + [] => (cls, ConsistentCnf simp) + | th :: ths => def_cnf_formula cls simp (simplify simp th) ths + + and def_cnf_formula cls simp (th as Thm (fm,_)) ths = case fm of - True => def_cnf_problem acc prob simp fms probs - | False => def_cnf acc probs - | And (_,_,s) => def_cnf_problem acc prob simp (Set.toList s @ fms) probs + True => def_cnf cls simp ths + | False => def_cnf_inconsistent th + | And (_,_,s) => + let + fun add (f,z) = Thm (f, Conjunct th) :: z + in + def_cnf cls simp (Set.foldr add ths s) + end | Exists (fv,_,n,f) => - def_cnf_formula acc prob simp (skolemize fv n f) fms probs - | Forall (_,_,_,f) => def_cnf_formula acc prob simp f fms probs + let + val th = Thm (skolemize fv n f, Skolemize th) + in + def_cnf_formula cls simp th ths + end + | Forall (_,_,_,f) => + let + val th = Thm (f, Specialize th) + in + def_cnf_formula cls simp th ths + end | _ => case minimumDefinition fm of - NONE => + SOME def => let - val prob = fm :: prob - and simp = simplifyAdd simp fm + val ths = th :: ths + val th = newDefinition def in - def_cnf_problem acc prob simp fms probs + def_cnf_formula cls simp th ths end - | SOME def => + | NONE => let - val def_fm = newDefinition def - and fms = fm :: fms + val simp = simplifyAdd simp th + + val fm = basicCnf fm + + val inf = Clausify th in - def_cnf_formula acc prob simp def_fm fms probs + case fm of + True => def_cnf cls simp ths + | False => def_cnf_inconsistent (Thm (fm,inf)) + | And (_,_,s) => + let + val inf = Conjunct (Thm (fm,inf)) + val cls = Set.foldl (def_cnf_clause inf) cls s + in + def_cnf cls simp ths + end + | fm => def_cnf (def_cnf_clause inf (fm,cls)) simp ths end; +in + fun addCnf th cnf = + case cnf of + ConsistentCnf simp => def_cnf [] simp [th] + | InconsistentCnf => ([],cnf); +end; - fun cnf_prob prob = toFormula (AndList (map basicCnf prob)); -in - fun cnf fm = +local + fun add (th,(cls,cnf)) = let - val fm = fromFormula fm -(*TRACE2 - val () = Parser.ppTrace pp "Normalize.cnf: fm" fm -*) - val probs = def_cnf [] [([],simplifyEmpty,[fm])] + val (cls',cnf) = addCnf th cnf in - map cnf_prob probs + (cls' @ cls, cnf) + end; +in + fun proveCnf ths = + let + val (cls,_) = List.foldl add ([],initialCnf) ths + in + rev cls end; end; +fun cnf fm = + let + val cls = proveCnf [mkAxiom fm] + in + map fst cls + end; + end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Options.sig --- a/src/Tools/Metis/src/Options.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Options.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* PROCESSING COMMAND LINE OPTIONS *) -(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Options = diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Options.sml --- a/src/Tools/Metis/src/Options.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Options.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* PROCESSING COMMAND LINE OPTIONS *) -(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Options :> Options = @@ -116,9 +116,10 @@ description = "no more options", processor = fn _ => raise Fail "basicOptions: --"}, {switches = ["-?","-h","--help"], arguments = [], - description = "display all options and exit", + description = "display option information and exit", processor = fn _ => raise OptionExit - {message = SOME "displaying all options", usage = true, success = true}}, + {message = SOME "displaying option information", + usage = true, success = true}}, {switches = ["-v", "--version"], arguments = [], description = "display version information", processor = fn _ => raise Fail "basicOptions: -v, --version"}]; @@ -127,8 +128,9 @@ (* All the command line options of a program *) (* ------------------------------------------------------------------------- *) -type allOptions = {name : string, version : string, header : string, - footer : string, options : opt list}; +type allOptions = + {name : string, version : string, header : string, + footer : string, options : opt list}; (* ------------------------------------------------------------------------- *) (* Usage information *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Ordered.sig --- a/src/Tools/Metis/src/Ordered.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Ordered.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* 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 = diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Ordered.sml --- a/src/Tools/Metis/src/Ordered.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Ordered.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,10 +1,23 @@ (* ========================================================================= *) (* 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; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/PP.sig --- a/src/Tools/Metis/src/PP.sig Mon Sep 13 16:44:20 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,224 +0,0 @@ -(* ========================================================================= *) -(* 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 = PP.mk_ppstream {consumer = - fn s => TextIO.output(TextIO.stdOut, s), - linewidth = 72, - flush = - fn () => TextIO.flushOut TextIO.stdOut}; - - fun ppb b = (ppbool ppstrm b; 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/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 (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)))))); -*) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/PP.sml --- a/src/Tools/Metis/src/PP.sml Mon Sep 13 16:44:20 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,617 +0,0 @@ -(* ========================================================================= *) -(* 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 ref, - back: int ref, - size: int} (* fixed size of element array *) -with - - fun is_empty (QUEUE{front=ref ~1, back=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=ref ~1, back=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 ref -with - fun mk_indent_stack() = Istack (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 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 ref, (* Processed *) - Ublocks : block_info list ref} (* Unprocessed *) - | E of {Pend : int ref, Uend : int ref} - | BR of {Distance_to_next_break : int 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 ref -> unit, (* increment circular buffer index *) - space_left : int ref, (* remaining columns on page *) - left_index : int ref, (* insertion index *) - right_index : int ref, (* output index *) - left_sum : int ref, (* size of strings and spaces inserted *) - right_sum : int 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 = ref linewidth, - left_index = ref 0, right_index = ref 0, - left_sum = ref 0, right_sum = 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 = ref [], Ublocks = ref []}) = - raise Fail "PP-error: print_BB" - | print_BB (PPS{the_indent_stack,linewidth,space_left=ref sp_left,...}, - {Pblocks as ref({How_to_indent=CONSISTENT,Block_size, - Block_offset}::rst), - Ublocks=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=ref sp_left,...}, - {Pblocks as ref({Block_size,Block_offset,...}::rst),Ublocks=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=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 = ref 0, Uend = 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 = ref [], Ublocks = ref []}) => true - | (BB _) => false - | (E {Pend = ref 0, Uend = 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 = ref [], Ublocks = ref []} => - raise Fail "PP-error: BB_size" - | {Pblocks as ref(_::_),Ublocks=ref[]} => POS - | {Ublocks, ...} => last_size (!Ublocks)) - | token_size (E{Pend = ref 0, Uend = ref 0}) = - raise Fail "PP-error: token_size.E" - | token_size (E{Pend = 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 = ref [], Ublocks = ref []}) => - (update(the_token_buffer,!left_index, - initial_token_value); - ++left_index) - | (BB _) => () - | (E {Pend = ref 0, Uend = 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 = ref (~(!right_sum)), - Block_offset = offset, - How_to_indent = style}::(!Ublocks) - | _ => (update(the_token_buffer, !right_index, - BB{Pblocks = ref [], - Ublocks = ref [{Block_size = 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 = ref 1, Uend = 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 = ref 1, Pend = 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 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 = 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 ref[{Block_size,...}:block_info],...}) = - (pop_bottom_delim_stack dstack; - Block_size := INFINITY) - | set (_,BB {Ublocks = 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 = 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Parse.sig --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/Metis/src/Parse.sig Mon Sep 13 21:24:10 2010 +0200 @@ -0,0 +1,113 @@ +(* ========================================================================= *) +(* PARSING *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) +(* ========================================================================= *) + +signature Parse = +sig + +(* ------------------------------------------------------------------------- *) +(* A "cannot parse" exception. *) +(* ------------------------------------------------------------------------- *) + +exception NoParse + +(* ------------------------------------------------------------------------- *) +(* Recursive descent parsing combinators. *) +(* ------------------------------------------------------------------------- *) + +(* + Recommended fixities: + + infixr 9 >>++ + infixr 8 ++ + infixr 7 >> + infixr 6 || +*) + +val error : 'a -> 'b * 'a + +val ++ : ('a -> 'b * 'a) * ('a -> 'c * 'a) -> 'a -> ('b * 'c) * 'a + +val >> : ('a -> 'b * 'a) * ('b -> 'c) -> 'a -> 'c * 'a + +val >>++ : ('a -> 'b * 'a) * ('b -> 'a -> 'c * 'a) -> 'a -> 'c * 'a + +val || : ('a -> 'b * 'a) * ('a -> 'b * 'a) -> 'a -> 'b * 'a + +val first : ('a -> 'b * 'a) list -> 'a -> 'b * 'a + +val mmany : ('s -> 'a -> 's * 'a) -> 's -> 'a -> 's * 'a + +val many : ('a -> 'b * 'a) -> 'a -> 'b list * 'a + +val atLeastOne : ('a -> 'b * 'a) -> 'a -> 'b list * 'a + +val nothing : 'a -> unit * 'a + +val optional : ('a -> 'b * 'a) -> 'a -> 'b option * 'a + +(* ------------------------------------------------------------------------- *) +(* Stream-based parsers. *) +(* ------------------------------------------------------------------------- *) + +type ('a,'b) parser = 'a Stream.stream -> 'b * 'a Stream.stream + +val maybe : ('a -> 'b option) -> ('a,'b) parser + +val finished : ('a,unit) parser + +val some : ('a -> bool) -> ('a,'a) parser + +val any : ('a,'a) parser + +(* ------------------------------------------------------------------------- *) +(* Parsing whole streams. *) +(* ------------------------------------------------------------------------- *) + +val fromStream : ('a,'b) parser -> 'a Stream.stream -> 'b + +val fromList : ('a,'b) parser -> 'a list -> 'b + +val everything : ('a, 'b list) parser -> 'a Stream.stream -> 'b Stream.stream + +(* ------------------------------------------------------------------------- *) +(* Parsing lines of text. *) +(* ------------------------------------------------------------------------- *) + +val initialize : + {lines : string Stream.stream} -> + {chars : char list Stream.stream, + parseErrorLocation : unit -> string} + +val exactChar : char -> (char,unit) parser + +val exactCharList : char list -> (char,unit) parser + +val exactString : string -> (char,unit) parser + +val escapeString : {escape : char list} -> (char,string) parser + +val manySpace : (char,unit) parser + +val atLeastOneSpace : (char,unit) parser + +val fromString : (char,'a) parser -> string -> 'a + +(* ------------------------------------------------------------------------- *) +(* Infix operators. *) +(* ------------------------------------------------------------------------- *) + +val parseInfixes : + Print.infixes -> (string * 'a * 'a -> 'a) -> (string,'a) parser -> + (string,'a) parser + +(* ------------------------------------------------------------------------- *) +(* Quotations. *) +(* ------------------------------------------------------------------------- *) + +type 'a quotation = 'a frag list + +val parseQuotation : ('a -> string) -> (string -> 'b) -> 'a quotation -> 'b + +end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Parse.sml --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/Metis/src/Parse.sml Mon Sep 13 21:24:10 2010 +0200 @@ -0,0 +1,265 @@ +(* ========================================================================= *) +(* PARSING *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) +(* ========================================================================= *) + +structure Parse :> Parse = +struct + +open Useful; + +infixr 9 >>++ +infixr 8 ++ +infixr 7 >> +infixr 6 || + +(* ------------------------------------------------------------------------- *) +(* A "cannot parse" exception. *) +(* ------------------------------------------------------------------------- *) + +exception NoParse; + +(* ------------------------------------------------------------------------- *) +(* Recursive descent parsing combinators. *) +(* ------------------------------------------------------------------------- *) + +val error : 'a -> 'b * 'a = fn _ => raise NoParse; + +fun op ++ (parser1,parser2) input = + let + val (result1,input) = parser1 input + val (result2,input) = parser2 input + in + ((result1,result2),input) + end; + +fun op >> (parser : 'a -> 'b * 'a, treatment) input = + let + val (result,input) = parser input + in + (treatment result, input) + end; + +fun op >>++ (parser,treatment) input = + let + val (result,input) = parser input + in + treatment result input + end; + +fun op || (parser1,parser2) input = + parser1 input handle NoParse => parser2 input; + +fun first [] _ = raise NoParse + | first (parser :: parsers) input = (parser || first parsers) input; + +fun mmany parser state input = + let + val (state,input) = parser state input + in + mmany parser state input + end + handle NoParse => (state,input); + +fun many parser = + let + fun sparser l = parser >> (fn x => x :: l) + in + mmany sparser [] >> rev + end; + +fun atLeastOne p = (p ++ many p) >> op::; + +fun nothing input = ((),input); + +fun optional p = (p >> SOME) || (nothing >> K NONE); + +(* ------------------------------------------------------------------------- *) +(* Stream-based parsers. *) +(* ------------------------------------------------------------------------- *) + +type ('a,'b) parser = 'a Stream.stream -> 'b * 'a Stream.stream + +fun maybe p Stream.Nil = raise NoParse + | maybe p (Stream.Cons (h,t)) = + case p h of SOME r => (r, t ()) | NONE => raise NoParse; + +fun finished Stream.Nil = ((), Stream.Nil) + | finished (Stream.Cons _) = raise NoParse; + +fun some p = maybe (fn x => if p x then SOME x else NONE); + +fun any input = some (K true) input; + +(* ------------------------------------------------------------------------- *) +(* Parsing whole streams. *) +(* ------------------------------------------------------------------------- *) + +fun fromStream parser input = + let + val (res,_) = (parser ++ finished >> fst) input + in + res + end; + +fun fromList parser l = fromStream parser (Stream.fromList l); + +fun everything parser = + let + fun parserOption input = + SOME (parser input) + handle e as NoParse => if Stream.null input then NONE else raise e + + fun parserList input = + case parserOption input of + NONE => Stream.Nil + | SOME (result,input) => + Stream.append (Stream.fromList result) (fn () => parserList input) + in + parserList + end; + +(* ------------------------------------------------------------------------- *) +(* Parsing lines of text. *) +(* ------------------------------------------------------------------------- *) + +fun initialize {lines} = + let + val lastLine = ref (~1,"","","") + + val chars = + let + fun saveLast line = + let + val ref (n,_,l2,l3) = lastLine + val () = lastLine := (n + 1, l2, l3, line) + in + explode line + end + in + Stream.memoize (Stream.map saveLast lines) + end + + fun parseErrorLocation () = + let + val ref (n,l1,l2,l3) = lastLine + in + (if n <= 0 then "at start" + else "around line " ^ Int.toString n) ^ + chomp (":\n" ^ l1 ^ l2 ^ l3) + end + in + {chars = chars, + parseErrorLocation = parseErrorLocation} + end; + +fun exactChar (c : char) = some (equal c) >> K (); + +fun exactCharList cs = + case cs of + [] => nothing + | c :: cs => (exactChar c ++ exactCharList cs) >> snd; + +fun exactString s = exactCharList (explode s); + +fun escapeString {escape} = + let + fun isEscape c = mem c escape + + fun isNormal c = + case c of + #"\\" => false + | #"\n" => false + | #"\t" => false + | _ => not (isEscape c) + + val escapeParser = + (exactChar #"\\" >> K #"\\") || + (exactChar #"n" >> K #"\n") || + (exactChar #"t" >> K #"\t") || + some isEscape + + val charParser = + ((exactChar #"\\" ++ escapeParser) >> snd) || + some isNormal + in + many charParser >> implode + end; + +local + val isSpace = Char.isSpace; + + val space = some isSpace; +in + val manySpace = many space >> K (); + + val atLeastOneSpace = atLeastOne space >> K (); +end; + +fun fromString parser s = fromList parser (explode s); + +(* ------------------------------------------------------------------------- *) +(* Infix operators. *) +(* ------------------------------------------------------------------------- *) + +fun parseGenInfix update sof toks parse inp = + let + val (e,rest) = parse inp + + val continue = + case rest of + Stream.Nil => NONE + | Stream.Cons (h_t as (h,_)) => + if StringSet.member h toks then SOME h_t else NONE + in + case continue of + NONE => (sof e, rest) + | SOME (h,t) => parseGenInfix update (update sof h e) toks parse (t ()) + end; + +local + fun add ({leftSpaces = _, token = t, rightSpaces = _}, s) = StringSet.add s t; + + fun parse leftAssoc toks con = + let + val update = + if leftAssoc then (fn f => fn t => fn a => fn b => con (t, f a, b)) + else (fn f => fn t => fn a => fn b => f (con (t, a, b))) + in + parseGenInfix update I toks + end; +in + fun parseLayeredInfixes {tokens,leftAssoc} = + let + val toks = List.foldl add StringSet.empty tokens + in + parse leftAssoc toks + end; +end; + +fun parseInfixes ops = + let + val layeredOps = Print.layerInfixes ops + + val iparsers = List.map parseLayeredInfixes layeredOps + in + fn con => fn subparser => foldl (fn (p,sp) => p con sp) subparser iparsers + end; + +(* ------------------------------------------------------------------------- *) +(* Quotations. *) +(* ------------------------------------------------------------------------- *) + +type 'a quotation = 'a frag list; + +fun parseQuotation printer parser quote = + let + fun expand (QUOTE q, s) = s ^ q + | expand (ANTIQUOTE a, s) = s ^ printer a + + val string = foldl expand "" quote + in + parser string + end; + +end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Parser.sig --- a/src/Tools/Metis/src/Parser.sig Mon Sep 13 16:44:20 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,144 +0,0 @@ -(* ========================================================================= *) -(* PARSING AND PRETTY PRINTING *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) -(* ========================================================================= *) - -signature Parser = -sig - -(* ------------------------------------------------------------------------- *) -(* Pretty printing for built-in types *) -(* ------------------------------------------------------------------------- *) - -type ppstream = PP.ppstream - -datatype breakStyle = Consistent | Inconsistent - -type 'a pp = ppstream -> 'a -> unit - -val lineLength : int ref - -val beginBlock : ppstream -> breakStyle -> int -> unit - -val endBlock : ppstream -> unit - -val addString : ppstream -> string -> unit - -val addBreak : ppstream -> int * int -> unit - -val addNewline : ppstream -> unit - -val ppMap : ('a -> 'b) -> 'b pp -> 'a pp - -val ppBracket : string -> string -> 'a pp -> 'a pp - -val ppSequence : string -> 'a pp -> 'a list pp - -val ppBinop : string -> 'a pp -> 'b pp -> ('a * 'b) pp - -val ppChar : char pp - -val ppString : string pp - -val ppUnit : unit pp - -val ppBool : bool pp - -val ppInt : int pp - -val ppReal : real pp - -val ppOrder : order pp - -val ppList : 'a pp -> 'a list pp - -val ppOption : 'a pp -> 'a option pp - -val ppPair : 'a pp -> 'b pp -> ('a * 'b) pp - -val ppTriple : 'a pp -> 'b pp -> 'c pp -> ('a * 'b * 'c) pp - -val toString : 'a pp -> 'a -> string (* Uses !lineLength *) - -val fromString : ('a -> string) -> 'a pp - -val ppTrace : 'a pp -> string -> 'a -> unit - -(* ------------------------------------------------------------------------- *) -(* Recursive descent parsing combinators *) -(* ------------------------------------------------------------------------- *) - -(* Generic parsers - -Recommended fixities: - infixr 9 >>++ - infixr 8 ++ - infixr 7 >> - infixr 6 || -*) - -exception NoParse - -val error : 'a -> 'b * 'a - -val ++ : ('a -> 'b * 'a) * ('a -> 'c * 'a) -> 'a -> ('b * 'c) * 'a - -val >> : ('a -> 'b * 'a) * ('b -> 'c) -> 'a -> 'c * 'a - -val >>++ : ('a -> 'b * 'a) * ('b -> 'a -> 'c * 'a) -> 'a -> 'c * 'a - -val || : ('a -> 'b * 'a) * ('a -> 'b * 'a) -> 'a -> 'b * 'a - -val first : ('a -> 'b * 'a) list -> 'a -> 'b * 'a - -val mmany : ('s -> 'a -> 's * 'a) -> 's -> 'a -> 's * 'a - -val many : ('a -> 'b * 'a) -> 'a -> 'b list * 'a - -val atLeastOne : ('a -> 'b * 'a) -> 'a -> 'b list * 'a - -val nothing : 'a -> unit * 'a - -val optional : ('a -> 'b * 'a) -> 'a -> 'b option * 'a - -(* Stream based parsers *) - -type ('a,'b) parser = 'a Stream.stream -> 'b * 'a Stream.stream - -val everything : ('a, 'b list) parser -> 'a Stream.stream -> 'b Stream.stream - -val maybe : ('a -> 'b option) -> ('a,'b) parser - -val finished : ('a,unit) parser - -val some : ('a -> bool) -> ('a,'a) parser - -val any : ('a,'a) parser - -val exact : ''a -> (''a,''a) parser - -(* ------------------------------------------------------------------------- *) -(* Infix operators *) -(* ------------------------------------------------------------------------- *) - -type infixities = {token : string, precedence : int, leftAssoc : bool} list - -val infixTokens : infixities -> string list - -val parseInfixes : - infixities -> (string * 'a * 'a -> 'a) -> (string,'a) parser -> - (string,'a) parser - -val ppInfixes : - infixities -> ('a -> (string * 'a * 'a) option) -> ('a * bool) pp -> - ('a * bool) pp - -(* ------------------------------------------------------------------------- *) -(* Quotations *) -(* ------------------------------------------------------------------------- *) - -type 'a quotation = 'a frag list - -val parseQuotation : ('a -> string) -> (string -> 'b) -> 'a quotation -> 'b - -end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Parser.sml --- a/src/Tools/Metis/src/Parser.sml Mon Sep 13 16:44:20 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,397 +0,0 @@ -(* ========================================================================= *) -(* PARSER COMBINATORS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) -(* ========================================================================= *) - -structure Parser :> Parser = -struct - -infixr 9 >>++ -infixr 8 ++ -infixr 7 >> -infixr 6 || - -(* ------------------------------------------------------------------------- *) -(* Helper functions. *) -(* ------------------------------------------------------------------------- *) - -exception Bug = Useful.Bug; - -val trace = Useful.trace -and equal = Useful.equal -and I = Useful.I -and K = Useful.K -and C = Useful.C -and fst = Useful.fst -and snd = Useful.snd -and pair = Useful.pair -and curry = Useful.curry -and funpow = Useful.funpow -and mem = Useful.mem -and sortMap = Useful.sortMap; - -(* ------------------------------------------------------------------------- *) -(* Pretty printing for built-in types *) -(* ------------------------------------------------------------------------- *) - -type ppstream = PP.ppstream - -datatype breakStyle = Consistent | Inconsistent - -type 'a pp = PP.ppstream -> 'a -> unit; - -val lineLength = ref 75; - -fun beginBlock pp Consistent = PP.begin_block pp PP.CONSISTENT - | beginBlock pp Inconsistent = PP.begin_block pp PP.INCONSISTENT; - -val endBlock = PP.end_block -and addString = PP.add_string -and addBreak = PP.add_break -and addNewline = PP.add_newline; - -fun ppMap f ppA (ppstrm : PP.ppstream) x : unit = ppA ppstrm (f x); - -fun ppBracket l r ppA pp a = - let - val ln = size l - in - beginBlock pp Inconsistent ln; - if ln = 0 then () else addString pp l; - ppA pp a; - if r = "" then () else addString pp r; - endBlock pp - end; - -fun ppSequence sep ppA pp = - let - fun ppX x = (addString pp sep; addBreak pp (1,0); ppA pp x) - in - fn [] => () - | h :: t => - (beginBlock pp Inconsistent 0; - ppA pp h; - app ppX t; - endBlock pp) - end; - -fun ppBinop s ppA ppB pp (a,b) = - (beginBlock pp Inconsistent 0; - ppA pp a; - if s = "" then () else addString pp s; - addBreak pp (1,0); - ppB pp b; - endBlock pp); - -fun ppTrinop ab bc ppA ppB ppC pp (a,b,c) = - (beginBlock pp Inconsistent 0; - ppA pp a; - if ab = "" then () else addString pp ab; - addBreak pp (1,0); - ppB pp b; - if bc = "" then () else addString pp bc; - addBreak pp (1,0); - ppC pp c; - endBlock pp); - -(* Pretty-printers for common types *) - -fun ppString pp s = - (beginBlock pp Inconsistent 0; - addString pp s; - endBlock pp); - -val ppUnit = ppMap (fn () => "()") ppString; - -val ppChar = ppMap str ppString; - -val ppBool = ppMap (fn true => "true" | false => "false") ppString; - -val ppInt = ppMap Int.toString ppString; - -val ppReal = ppMap Real.toString ppString; - -val ppOrder = - let - fun f LESS = "Less" - | f EQUAL = "Equal" - | f GREATER = "Greater" - in - ppMap f ppString - end; - -fun ppList ppA = ppBracket "[" "]" (ppSequence "," ppA); - -fun ppOption _ pp NONE = ppString pp "-" - | ppOption ppA pp (SOME a) = ppA pp a; - -fun ppPair ppA ppB = ppBracket "(" ")" (ppBinop "," ppA ppB); - -fun ppTriple ppA ppB ppC = ppBracket "(" ")" (ppTrinop "," "," ppA ppB ppC); - -(* Keep eta expanded to evaluate lineLength when supplied with a *) -fun toString ppA a = PP.pp_to_string (!lineLength) ppA a; - -fun fromString toS = ppMap toS ppString; - -fun ppTrace ppX nameX x = - trace (toString (ppBinop " =" ppString ppX) (nameX,x) ^ "\n"); - -(* ------------------------------------------------------------------------- *) -(* Generic. *) -(* ------------------------------------------------------------------------- *) - -exception NoParse; - -val error : 'a -> 'b * 'a = fn _ => raise NoParse; - -fun op ++ (parser1,parser2) input = - let - val (result1,input) = parser1 input - val (result2,input) = parser2 input - in - ((result1,result2),input) - end; - -fun op >> (parser : 'a -> 'b * 'a, treatment) input = - let - val (result,input) = parser input - in - (treatment result, input) - end; - -fun op >>++ (parser,treatment) input = - let - val (result,input) = parser input - in - treatment result input - end; - -fun op || (parser1,parser2) input = - parser1 input handle NoParse => parser2 input; - -fun first [] _ = raise NoParse - | first (parser :: parsers) input = (parser || first parsers) input; - -fun mmany parser state input = - let - val (state,input) = parser state input - in - mmany parser state input - end - handle NoParse => (state,input); - -fun many parser = - let - fun sparser l = parser >> (fn x => x :: l) - in - mmany sparser [] >> rev - end; - -fun atLeastOne p = (p ++ many p) >> op::; - -fun nothing input = ((),input); - -fun optional p = (p >> SOME) || (nothing >> K NONE); - -(* ------------------------------------------------------------------------- *) -(* Stream-based. *) -(* ------------------------------------------------------------------------- *) - -type ('a,'b) parser = 'a Stream.stream -> 'b * 'a Stream.stream - -fun everything parser = - let - fun f input = - let - val (result,input) = parser input - in - Stream.append (Stream.fromList result) (fn () => f input) - end - handle NoParse => - if Stream.null input then Stream.NIL else raise NoParse - in - f - end; - -fun maybe p Stream.NIL = raise NoParse - | maybe p (Stream.CONS (h,t)) = - case p h of SOME r => (r, t ()) | NONE => raise NoParse; - -fun finished Stream.NIL = ((), Stream.NIL) - | finished (Stream.CONS _) = raise NoParse; - -fun some p = maybe (fn x => if p x then SOME x else NONE); - -fun any input = some (K true) input; - -fun exact tok = some (fn item => item = tok); - -(* ------------------------------------------------------------------------- *) -(* Parsing and pretty-printing for infix operators. *) -(* ------------------------------------------------------------------------- *) - -type infixities = {token : string, precedence : int, leftAssoc : bool} list; - -local - fun unflatten ({token,precedence,leftAssoc}, ([],_)) = - ([(leftAssoc, [token])], precedence) - | unflatten ({token,precedence,leftAssoc}, ((a,l) :: dealt, p)) = - if p = precedence then - let - val _ = leftAssoc = a orelse - raise Bug "infix parser/printer: mixed assocs" - in - ((a, token :: l) :: dealt, p) - end - else - ((leftAssoc,[token]) :: (a,l) :: dealt, precedence); -in - fun layerOps infixes = - let - val infixes = sortMap #precedence Int.compare infixes - val (parsers,_) = foldl unflatten ([],0) infixes - in - parsers - end; -end; - -local - fun chop (#" " :: chs) = let val (n,l) = chop chs in (n + 1, l) end - | chop chs = (0,chs); - - fun nspaces n = funpow n (curry op^ " ") ""; - - fun spacify tok = - let - val chs = explode tok - val (r,chs) = chop (rev chs) - val (l,chs) = chop (rev chs) - in - ((l,r), implode chs) - end; - - fun lrspaces (l,r) = - (if l = 0 then K () else C addString (nspaces l), - if r = 0 then K () else C addBreak (r, 0)); -in - fun opSpaces s = let val (l_r,s) = spacify s in (lrspaces l_r, s) end; - - val opClean = snd o spacify; -end; - -val infixTokens : infixities -> string list = - List.map (fn {token,...} => opClean token); - -fun parseGenInfix update sof toks parse inp = - let - val (e, rest) = parse inp - - val continue = - case rest of - Stream.NIL => NONE - | Stream.CONS (h, t) => if mem h toks then SOME (h, t) else NONE - in - case continue of - NONE => (sof e, rest) - | SOME (h,t) => parseGenInfix update (update sof h e) toks parse (t ()) - end; - -fun parseLeftInfix toks con = - parseGenInfix (fn f => fn t => fn a => fn b => con (t, f a, b)) I toks; - -fun parseRightInfix toks con = - parseGenInfix (fn f => fn t => fn a => fn b => f (con (t, a, b))) I toks; - -fun parseInfixes ops = - let - fun layeredOp (x,y) = (x, List.map opClean y) - - val layeredOps = List.map layeredOp (layerOps ops) - - fun iparser (a,t) = (if a then parseLeftInfix else parseRightInfix) t - - val iparsers = List.map iparser layeredOps - in - fn con => fn subparser => foldl (fn (p,sp) => p con sp) subparser iparsers - end; - -fun ppGenInfix left toks = - let - val spc = List.map opSpaces toks - in - fn dest => fn ppSub => - let - fun dest' tm = - case dest tm of - NONE => NONE - | SOME (t, a, b) => - Option.map (pair (a,b)) (List.find (equal t o snd) spc) - - open PP - - fun ppGo pp (tmr as (tm,r)) = - case dest' tm of - NONE => ppSub pp tmr - | SOME ((a,b),((lspc,rspc),tok)) => - ((if left then ppGo else ppSub) pp (a,true); - lspc pp; addString pp tok; rspc pp; - (if left then ppSub else ppGo) pp (b,r)) - in - fn pp => fn tmr as (tm,_) => - case dest' tm of - NONE => ppSub pp tmr - | SOME _ => (beginBlock pp Inconsistent 0; ppGo pp tmr; endBlock pp) - end - end; - -fun ppLeftInfix toks = ppGenInfix true toks; - -fun ppRightInfix toks = ppGenInfix false toks; - -fun ppInfixes ops = - let - val layeredOps = layerOps ops - - val toks = List.concat (List.map (List.map opClean o snd) layeredOps) - - fun iprinter (a,t) = (if a then ppLeftInfix else ppRightInfix) t - - val iprinters = List.map iprinter layeredOps - in - fn dest => fn ppSub => - let - fun printer sub = foldl (fn (ip,p) => ip dest p) sub iprinters - - fun isOp t = case dest t of SOME (x,_,_) => mem x toks | _ => false - - open PP - - fun subpr pp (tmr as (tm,_)) = - if isOp tm then - (beginBlock pp Inconsistent 1; addString pp "("; - printer subpr pp (tm, false); addString pp ")"; endBlock pp) - else ppSub pp tmr - in - fn pp => fn tmr => - (beginBlock pp Inconsistent 0; printer subpr pp tmr; endBlock pp) - end - end; - -(* ------------------------------------------------------------------------- *) -(* Quotations *) -(* ------------------------------------------------------------------------- *) - -type 'a quotation = 'a frag list; - -fun parseQuotation printer parser quote = - let - fun expand (QUOTE q, s) = s ^ q - | expand (ANTIQUOTE a, s) = s ^ printer a - - val string = foldl expand "" quote - in - parser string - end; - -end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Portable.sig --- a/src/Tools/Metis/src/Portable.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Portable.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* 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 = @@ -25,12 +25,6 @@ val time : ('a -> 'b) -> 'a -> 'b (* ------------------------------------------------------------------------- *) -(* Critical section markup (multiprocessing) *) -(* ------------------------------------------------------------------------- *) - -val CRITICAL: (unit -> 'a) -> 'a - -(* ------------------------------------------------------------------------- *) (* Generating random values. *) (* ------------------------------------------------------------------------- *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/PortableIsabelle.sml --- a/src/Tools/Metis/src/PortableIsabelle.sml Mon Sep 13 16:44:20 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,47 +0,0 @@ -(* ========================================================================= *) -(* Isabelle ML SPECIFIC FUNCTIONS *) -(* ========================================================================= *) - -structure Portable :> Portable = -struct - -(* ------------------------------------------------------------------------- *) -(* The ML implementation. *) -(* ------------------------------------------------------------------------- *) - -val ml = ml_system; - -(* ------------------------------------------------------------------------- *) -(* 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; - -(* ------------------------------------------------------------------------- *) -(* 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; - -(* ------------------------------------------------------------------------- *) -(* Quotations a la Moscow ML. *) -(* ------------------------------------------------------------------------- *) - -datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/PortableMlton.sml --- a/src/Tools/Metis/src/PortableMlton.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/PortableMlton.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* MLTON SPECIFIC FUNCTIONS *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Portable :> Portable = @@ -57,12 +57,6 @@ end; (* ------------------------------------------------------------------------- *) -(* Critical section markup (multiprocessing) *) -(* ------------------------------------------------------------------------- *) - -fun CRITICAL e = e (); (*dummy*) - -(* ------------------------------------------------------------------------- *) (* Generating random values. *) (* ------------------------------------------------------------------------- *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/PortableMosml.sml --- a/src/Tools/Metis/src/PortableMosml.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/PortableMosml.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* MOSCOW ML SPECIFIC FUNCTIONS *) -(* Copyright (c) 2002-2004 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Portable :> Portable = @@ -29,12 +29,6 @@ val time = Mosml.time; (* ------------------------------------------------------------------------- *) -(* Critical section markup (multiprocessing) *) -(* ------------------------------------------------------------------------- *) - -fun CRITICAL e = e (); (*dummy*) - -(* ------------------------------------------------------------------------- *) (* Generating random values. *) (* ------------------------------------------------------------------------- *) @@ -98,6 +92,24 @@ modifyi f (a,0,NONE) end; +fun OS_Process_isSuccess s = s = OS.Process.success; + +fun String_isSuffix p s = + let + val sizeP = size p + and sizeS = size s + in + sizeP <= sizeS andalso + String.extract (s, sizeS - sizeP, NONE) = p + end; + +fun Substring_full s = + let + open Substring + in + all s + end; + fun TextIO_inputLine h = let open TextIO diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/PortablePolyml.sml --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/Metis/src/PortablePolyml.sml Mon Sep 13 21:24:10 2010 +0200 @@ -0,0 +1,77 @@ +(* ========================================================================= *) +(* POLY/ML SPECIFIC FUNCTIONS *) +(* Copyright (c) 2008 Joe Hurd, distributed under the BSD License *) +(* ========================================================================= *) + +structure Portable :> Portable = +struct + +(* ------------------------------------------------------------------------- *) +(* The ML implementation. *) +(* ------------------------------------------------------------------------- *) + +val ml = "polyml"; + +(* ------------------------------------------------------------------------- *) +(* Pointer equality using the run-time system. *) +(* ------------------------------------------------------------------------- *) + +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.nextWord; + +val randomBool = Random.nextBool; + +val randomInt = Random.nextInt; + +val randomReal = Random.nextReal; + +end + +(* ------------------------------------------------------------------------- *) +(* Quotations a la Moscow ML. *) +(* ------------------------------------------------------------------------- *) + +datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Print.sig --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/Metis/src/Print.sig Mon Sep 13 21:24:10 2010 +0200 @@ -0,0 +1,157 @@ +(* ========================================================================= *) +(* PRETTY-PRINTING *) +(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *) +(* ========================================================================= *) + +signature Print = +sig + +(* ------------------------------------------------------------------------- *) +(* A type of pretty-printers. *) +(* ------------------------------------------------------------------------- *) + +datatype breakStyle = Consistent | Inconsistent + +datatype ppStep = + BeginBlock of breakStyle * int + | EndBlock + | AddString of string + | AddBreak of int + | AddNewline + +type ppstream = ppStep Stream.stream + +type 'a pp = 'a -> ppstream + +(* ------------------------------------------------------------------------- *) +(* Pretty-printer primitives. *) +(* ------------------------------------------------------------------------- *) + +val beginBlock : breakStyle -> int -> ppstream + +val endBlock : ppstream + +val addString : string -> ppstream + +val addBreak : int -> ppstream + +val addNewline : ppstream + +val skip : ppstream + +val sequence : ppstream -> ppstream -> ppstream + +val duplicate : int -> ppstream -> ppstream + +val program : ppstream list -> ppstream + +val stream : ppstream Stream.stream -> ppstream + +val block : breakStyle -> int -> ppstream -> ppstream + +val blockProgram : breakStyle -> int -> ppstream list -> ppstream + +val bracket : string -> string -> ppstream -> ppstream + +val field : string -> ppstream -> ppstream + +val record : (string * ppstream) list -> ppstream + +(* ------------------------------------------------------------------------- *) +(* Pretty-printer combinators. *) +(* ------------------------------------------------------------------------- *) + +val ppMap : ('a -> 'b) -> 'b pp -> 'a pp + +val ppBracket : string -> string -> 'a pp -> 'a pp + +val ppOp : string -> ppstream + +val ppOp2 : string -> 'a pp -> 'b pp -> ('a * 'b) pp + +val ppOp3 : string -> string -> 'a pp -> 'b pp -> 'c pp -> ('a * 'b * 'c) pp + +val ppOpList : string -> 'a pp -> 'a list pp + +val ppOpStream : string -> 'a pp -> 'a Stream.stream pp + +(* ------------------------------------------------------------------------- *) +(* Pretty-printers for common types. *) +(* ------------------------------------------------------------------------- *) + +val ppChar : char pp + +val ppString : string pp + +val ppEscapeString : {escape : char list} -> string pp + +val ppUnit : unit pp + +val ppBool : bool pp + +val ppInt : int pp + +val ppPrettyInt : int pp + +val ppReal : real pp + +val ppPercent : real pp + +val ppOrder : order pp + +val ppList : 'a pp -> 'a list pp + +val ppStream : 'a pp -> 'a Stream.stream pp + +val ppOption : 'a pp -> 'a option pp + +val ppPair : 'a pp -> 'b pp -> ('a * 'b) pp + +val ppTriple : 'a pp -> 'b pp -> 'c pp -> ('a * 'b * 'c) pp + +val ppBreakStyle : breakStyle pp + +val ppPpStep : ppStep pp + +val ppPpstream : ppstream pp + +(* ------------------------------------------------------------------------- *) +(* Pretty-printing infix operators. *) +(* ------------------------------------------------------------------------- *) + +datatype infixes = + Infixes of + {token : string, + precedence : int, + leftAssoc : bool} list + +val tokensInfixes : infixes -> StringSet.set + +val layerInfixes : + infixes -> + {tokens : {leftSpaces : int, token : string, rightSpaces : int} list, + leftAssoc : bool} list + +val ppInfixes : + infixes -> ('a -> (string * 'a * 'a) option) -> ('a * bool) pp -> + ('a * bool) pp + +(* ------------------------------------------------------------------------- *) +(* Executing pretty-printers to generate lines. *) +(* ------------------------------------------------------------------------- *) + +val execute : {lineLength : int} -> ppstream -> string Stream.stream + +(* ------------------------------------------------------------------------- *) +(* Executing pretty-printers with a global line length. *) +(* ------------------------------------------------------------------------- *) + +val lineLength : int ref + +val toString : 'a pp -> 'a -> string + +val toStream : 'a pp -> 'a -> string Stream.stream + +val trace : 'a pp -> string -> 'a -> unit + +end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Print.sml --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/Metis/src/Print.sml Mon Sep 13 21:24:10 2010 +0200 @@ -0,0 +1,1137 @@ +(* ========================================================================= *) +(* PRETTY-PRINTING *) +(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *) +(* ========================================================================= *) + +structure Print :> Print = +struct + +open Useful; + +(* ------------------------------------------------------------------------- *) +(* Constants. *) +(* ------------------------------------------------------------------------- *) + +val initialLineLength = 75; + +(* ------------------------------------------------------------------------- *) +(* Helper functions. *) +(* ------------------------------------------------------------------------- *) + +fun revAppend xs s = + case xs of + [] => s () + | x :: xs => revAppend xs (K (Stream.Cons (x,s))); + +fun revConcat strm = + case strm of + Stream.Nil => Stream.Nil + | Stream.Cons (h,t) => revAppend h (revConcat o t); + +local + fun join current prev = (prev ^ "\n", current); +in + fun joinNewline strm = + case strm of + Stream.Nil => Stream.Nil + | Stream.Cons (h,t) => Stream.maps join Stream.singleton h (t ()); +end; + +local + fun calcSpaces n = nChars #" " n; + + val cachedSpaces = Vector.tabulate (initialLineLength,calcSpaces); +in + fun nSpaces n = + if n < initialLineLength then Vector.sub (cachedSpaces,n) + else calcSpaces n; +end; + +(* ------------------------------------------------------------------------- *) +(* A type of pretty-printers. *) +(* ------------------------------------------------------------------------- *) + +datatype breakStyle = Consistent | Inconsistent; + +datatype ppStep = + BeginBlock of breakStyle * int + | EndBlock + | AddString of string + | AddBreak of int + | AddNewline; + +type ppstream = ppStep Stream.stream; + +type 'a pp = 'a -> ppstream; + +fun breakStyleToString style = + case style of + Consistent => "Consistent" + | Inconsistent => "Inconsistent"; + +fun ppStepToString step = + case step of + BeginBlock _ => "BeginBlock" + | EndBlock => "EndBlock" + | AddString _ => "AddString" + | AddBreak _ => "AddBreak" + | AddNewline => "AddNewline"; + +(* ------------------------------------------------------------------------- *) +(* Pretty-printer primitives. *) +(* ------------------------------------------------------------------------- *) + +fun beginBlock style indent = Stream.singleton (BeginBlock (style,indent)); + +val endBlock = Stream.singleton EndBlock; + +fun addString s = Stream.singleton (AddString s); + +fun addBreak spaces = Stream.singleton (AddBreak spaces); + +val addNewline = Stream.singleton AddNewline; + +val skip : ppstream = Stream.Nil; + +fun sequence pp1 pp2 : ppstream = Stream.append pp1 (K pp2); + +local + fun dup pp n () = if n = 1 then pp else Stream.append pp (dup pp (n - 1)); +in + fun duplicate n pp = if n = 0 then skip else dup pp n (); +end; + +val program : ppstream list -> ppstream = Stream.concatList; + +val stream : ppstream Stream.stream -> ppstream = Stream.concat; + +fun block style indent pp = program [beginBlock style indent, pp, endBlock]; + +fun blockProgram style indent pps = block style indent (program pps); + +fun bracket l r pp = + blockProgram Inconsistent (size l) + [addString l, + pp, + addString r]; + +fun field f pp = + blockProgram Inconsistent 2 + [addString (f ^ " ="), + addBreak 1, + pp]; + +val record = + let + val sep = sequence (addString ",") (addBreak 1) + + fun recordField (f,pp) = field f pp + + fun sepField f = sequence sep (recordField f) + + fun fields [] = [] + | fields (f :: fs) = recordField f :: map sepField fs + in + bracket "{" "}" o blockProgram Consistent 0 o fields + end; + +(* ------------------------------------------------------------------------- *) +(* Pretty-printer combinators. *) +(* ------------------------------------------------------------------------- *) + +fun ppMap f ppB a : ppstream = ppB (f a); + +fun ppBracket l r ppA a = bracket l r (ppA a); + +fun ppOp s = sequence (if s = "" then skip else addString s) (addBreak 1); + +fun ppOp2 ab ppA ppB (a,b) = + blockProgram Inconsistent 0 + [ppA a, + ppOp ab, + ppB b]; + +fun ppOp3 ab bc ppA ppB ppC (a,b,c) = + blockProgram Inconsistent 0 + [ppA a, + ppOp ab, + ppB b, + ppOp bc, + ppC c]; + +fun ppOpList s ppA = + let + fun ppOpA a = sequence (ppOp s) (ppA a) + in + fn [] => skip + | h :: t => blockProgram Inconsistent 0 (ppA h :: map ppOpA t) + end; + +fun ppOpStream s ppA = + let + fun ppOpA a = sequence (ppOp s) (ppA a) + in + fn Stream.Nil => skip + | Stream.Cons (h,t) => + blockProgram Inconsistent 0 + [ppA h, + Stream.concat (Stream.map ppOpA (t ()))] + end; + +(* ------------------------------------------------------------------------- *) +(* Pretty-printers for common types. *) +(* ------------------------------------------------------------------------- *) + +fun ppChar c = addString (str c); + +val ppString = addString; + +fun ppEscapeString {escape} = + let + val escapeMap = map (fn c => (c, "\\" ^ str c)) escape + + fun escapeChar c = + case c of + #"\\" => "\\\\" + | #"\n" => "\\n" + | #"\t" => "\\t" + | _ => + case List.find (equal c o fst) escapeMap of + SOME (_,s) => s + | NONE => str c + in + fn s => addString (String.translate escapeChar s) + end; + +val ppUnit : unit pp = K (addString "()"); + +fun ppBool b = addString (if b then "true" else "false"); + +fun ppInt i = addString (Int.toString i); + +local + val ppNeg = addString "~" + and ppSep = addString "," + and ppZero = addString "0" + and ppZeroZero = addString "00"; + + fun ppIntBlock i = + if i < 10 then sequence ppZeroZero (ppInt i) + else if i < 100 then sequence ppZero (ppInt i) + else ppInt i; + + fun ppIntBlocks i = + if i < 1000 then ppInt i + else sequence (ppIntBlocks (i div 1000)) + (sequence ppSep (ppIntBlock (i mod 1000))); +in + fun ppPrettyInt i = + if i < 0 then sequence ppNeg (ppIntBlocks (~i)) + else ppIntBlocks i; +end; + +fun ppReal r = addString (Real.toString r); + +fun ppPercent p = addString (percentToString p); + +fun ppOrder x = + addString + (case x of + LESS => "Less" + | EQUAL => "Equal" + | GREATER => "Greater"); + +fun ppList ppA = ppBracket "[" "]" (ppOpList "," ppA); + +fun ppStream ppA = ppBracket "[" "]" (ppOpStream "," ppA); + +fun ppOption ppA ao = + case ao of + SOME a => ppA a + | NONE => addString "-"; + +fun ppPair ppA ppB = ppBracket "(" ")" (ppOp2 "," ppA ppB); + +fun ppTriple ppA ppB ppC = ppBracket "(" ")" (ppOp3 "," "," ppA ppB ppC); + +fun ppBreakStyle style = addString (breakStyleToString style); + +fun ppPpStep step = + let + val cmd = ppStepToString step + in + blockProgram Inconsistent 2 + (addString cmd :: + (case step of + BeginBlock style_indent => + [addBreak 1, + ppPair ppBreakStyle ppInt style_indent] + | EndBlock => [] + | AddString s => + [addBreak 1, + addString ("\"" ^ s ^ "\"")] + | AddBreak n => + [addBreak 1, + ppInt n] + | AddNewline => [])) + end; + +val ppPpstream = ppStream ppPpStep; + +(* ------------------------------------------------------------------------- *) +(* Pretty-printing infix operators. *) +(* ------------------------------------------------------------------------- *) + +datatype infixes = + Infixes of + {token : string, + precedence : int, + leftAssoc : bool} list; + +local + fun chop l = + case l of + #" " :: l => let val (n,l) = chop l in (n + 1, l) end + | _ => (0,l); +in + fun opSpaces tok = + let + val tok = explode tok + val (r,tok) = chop (rev tok) + val (l,tok) = chop (rev tok) + val tok = implode tok + in + {leftSpaces = l, token = tok, rightSpaces = r} + end; +end; + +fun ppOpSpaces {leftSpaces,token,rightSpaces} = + let + val leftSpacesToken = + if leftSpaces = 0 then token else nSpaces leftSpaces ^ token + in + sequence + (addString leftSpacesToken) + (addBreak rightSpaces) + end; + +local + fun new t l acc = {tokens = [opSpaces t], leftAssoc = l} :: acc; + + fun add t l acc = + case acc of + [] => raise Bug "Print.layerInfixOps.layer" + | {tokens = ts, leftAssoc = l'} :: acc => + if l = l' then {tokens = opSpaces t :: ts, leftAssoc = l} :: acc + else raise Bug "Print.layerInfixOps: mixed assocs"; + + fun layer ({token = t, precedence = p, leftAssoc = l}, (acc,p')) = + let + val acc = if p = p' then add t l acc else new t l acc + in + (acc,p) + end; +in + fun layerInfixes (Infixes i) = + case sortMap #precedence Int.compare i of + [] => [] + | {token = t, precedence = p, leftAssoc = l} :: i => + let + val acc = new t l [] + + val (acc,_) = List.foldl layer (acc,p) i + in + acc + end; +end; + +val tokensLayeredInfixes = + let + fun addToken ({leftSpaces = _, token = t, rightSpaces = _}, s) = + StringSet.add s t + + fun addTokens ({tokens = t, leftAssoc = _}, s) = + List.foldl addToken s t + in + List.foldl addTokens StringSet.empty + end; + +val tokensInfixes = tokensLayeredInfixes o layerInfixes; + +local + val mkTokenMap = + let + fun add (token,m) = + let + val {leftSpaces = _, token = t, rightSpaces = _} = token + in + StringMap.insert m (t, ppOpSpaces token) + end + in + List.foldl add (StringMap.new ()) + end; +in + fun ppGenInfix {tokens,leftAssoc} = + let + val tokenMap = mkTokenMap tokens + in + fn dest => fn ppSub => + let + fun dest' tm = + case dest tm of + NONE => NONE + | SOME (t,a,b) => + case StringMap.peek tokenMap t of + NONE => NONE + | SOME p => SOME (p,a,b) + + fun ppGo (tmr as (tm,r)) = + case dest' tm of + NONE => ppSub tmr + | SOME (p,a,b) => + program + [(if leftAssoc then ppGo else ppSub) (a,true), + p, + (if leftAssoc then ppSub else ppGo) (b,r)] + in + fn tmr as (tm,_) => + if Option.isSome (dest' tm) then + block Inconsistent 0 (ppGo tmr) + else + ppSub tmr + end + end; +end + +fun ppInfixes ops = + let + val layeredOps = layerInfixes ops + + val toks = tokensLayeredInfixes layeredOps + + val iprinters = List.map ppGenInfix layeredOps + in + fn dest => fn ppSub => + let + fun printer sub = foldl (fn (ip,p) => ip dest p) sub iprinters + + fun isOp t = + case dest t of + SOME (x,_,_) => StringSet.member x toks + | NONE => false + + fun subpr (tmr as (tm,_)) = + if isOp tm then + blockProgram Inconsistent 1 + [addString "(", + printer subpr (tm,false), + addString ")"] + else + ppSub tmr + in + fn tmr => block Inconsistent 0 (printer subpr tmr) + end + end; + +(* ------------------------------------------------------------------------- *) +(* Executing pretty-printers to generate lines. *) +(* ------------------------------------------------------------------------- *) + +datatype blockBreakStyle = + InconsistentBlock + | ConsistentBlock + | BreakingBlock; + +datatype block = + Block of + {indent : int, + style : blockBreakStyle, + size : int, + chunks : chunk list} + +and chunk = + StringChunk of {size : int, string : string} + | BreakChunk of int + | BlockChunk of block; + +datatype state = + State of + {blocks : block list, + lineIndent : int, + lineSize : int}; + +val initialIndent = 0; + +val initialStyle = Inconsistent; + +fun liftStyle style = + case style of + Inconsistent => InconsistentBlock + | Consistent => ConsistentBlock; + +fun breakStyle style = + case style of + ConsistentBlock => BreakingBlock + | _ => style; + +fun sizeBlock (Block {size,...}) = size; + +fun sizeChunk chunk = + case chunk of + StringChunk {size,...} => size + | BreakChunk spaces => spaces + | BlockChunk block => sizeBlock block; + +val splitChunks = + let + fun split _ [] = NONE + | split acc (chunk :: chunks) = + case chunk of + BreakChunk _ => SOME (rev acc, chunks) + | _ => split (chunk :: acc) chunks + in + split [] + end; + +val sizeChunks = List.foldl (fn (c,z) => sizeChunk c + z) 0; + +local + fun render acc [] = acc + | render acc (chunk :: chunks) = + case chunk of + StringChunk {string = s, ...} => render (acc ^ s) chunks + | BreakChunk n => render (acc ^ nSpaces n) chunks + | BlockChunk (Block {chunks = l, ...}) => + render acc (List.revAppend (l,chunks)); +in + fun renderChunks indent chunks = render (nSpaces indent) (rev chunks); + + fun renderChunk indent chunk = renderChunks indent [chunk]; +end; + +fun isEmptyBlock block = + let + val Block {indent = _, style = _, size, chunks} = block + + val empty = null chunks + +(*BasicDebug + val _ = not empty orelse size = 0 orelse + raise Bug "Print.isEmptyBlock: bad size" +*) + in + empty + end; + +fun checkBlock ind block = + let + val Block {indent, style = _, size, chunks} = block + val _ = indent >= ind orelse raise Bug "Print.checkBlock: bad indents" + val size' = checkChunks indent chunks + val _ = size = size' orelse raise Bug "Print.checkBlock: wrong size" + in + size + end + +and checkChunks ind chunks = + case chunks of + [] => 0 + | chunk :: chunks => checkChunk ind chunk + checkChunks ind chunks + +and checkChunk ind chunk = + case chunk of + StringChunk {size,...} => size + | BreakChunk spaces => spaces + | BlockChunk block => checkBlock ind block; + +val checkBlocks = + let + fun check ind blocks = + case blocks of + [] => 0 + | block :: blocks => + let + val Block {indent,...} = block + in + checkBlock ind block + check indent blocks + end + in + check initialIndent o rev + end; + +val initialBlock = + let + val indent = initialIndent + val style = liftStyle initialStyle + val size = 0 + val chunks = [] + in + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + end; + +val initialState = + let + val blocks = [initialBlock] + val lineIndent = initialIndent + val lineSize = 0 + in + State + {blocks = blocks, + lineIndent = lineIndent, + lineSize = lineSize} + end; + +(*BasicDebug +fun checkState state = + (let + val State {blocks, lineIndent = _, lineSize} = state + val lineSize' = checkBlocks blocks + val _ = lineSize = lineSize' orelse + raise Error "wrong lineSize" + in + () + end + handle Error err => raise Bug err) + handle Bug bug => raise Bug ("Print.checkState: " ^ bug); +*) + +fun isFinalState state = + let + val State {blocks,lineIndent,lineSize} = state + in + case blocks of + [] => raise Bug "Print.isFinalState: no block" + | [block] => isEmptyBlock block + | _ :: _ :: _ => false + end; + +local + fun renderBreak lineIndent (chunks,lines) = + let + val line = renderChunks lineIndent chunks + + val lines = line :: lines + in + lines + end; + + fun renderBreaks lineIndent lineIndent' breaks lines = + case rev breaks of + [] => raise Bug "Print.renderBreaks" + | c :: cs => + let + val lines = renderBreak lineIndent (c,lines) + in + List.foldl (renderBreak lineIndent') lines cs + end; + + fun splitAllChunks cumulatingChunks = + let + fun split chunks = + case splitChunks chunks of + SOME (prefix,chunks) => prefix :: split chunks + | NONE => [List.concat (chunks :: cumulatingChunks)] + in + split + end; + + fun mkBreak style cumulatingChunks chunks = + case splitChunks chunks of + NONE => NONE + | SOME (chunks,broken) => + let + val breaks = + case style of + InconsistentBlock => + [List.concat (broken :: cumulatingChunks)] + | _ => splitAllChunks cumulatingChunks broken + in + SOME (breaks,chunks) + end; + + fun naturalBreak blocks = + case blocks of + [] => Right ([],[]) + | block :: blocks => + case naturalBreak blocks of + Left (breaks,blocks,lineIndent,lineSize) => + let + val Block {size,...} = block + + val blocks = block :: blocks + + val lineSize = lineSize + size + in + Left (breaks,blocks,lineIndent,lineSize) + end + | Right (cumulatingChunks,blocks) => + let + val Block {indent,style,size,chunks} = block + + val style = breakStyle style + in + case mkBreak style cumulatingChunks chunks of + SOME (breaks,chunks) => + let + val size = sizeChunks chunks + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + + val blocks = block :: blocks + + val lineIndent = indent + + val lineSize = size + in + Left (breaks,blocks,lineIndent,lineSize) + end + | NONE => + let + val cumulatingChunks = chunks :: cumulatingChunks + + val size = 0 + + val chunks = [] + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + + val blocks = block :: blocks + in + Right (cumulatingChunks,blocks) + end + end; + + fun forceBreakBlock cumulatingChunks block = + let + val Block {indent, style, size = _, chunks} = block + + val style = breakStyle style + + val break = + case mkBreak style cumulatingChunks chunks of + SOME (breaks,chunks) => + let + val lineSize = sizeChunks chunks + val lineIndent = indent + in + SOME (breaks,chunks,lineIndent,lineSize) + end + | NONE => forceBreakChunks cumulatingChunks chunks + in + case break of + SOME (breaks,chunks,lineIndent,lineSize) => + let + val size = lineSize + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + in + SOME (breaks,block,lineIndent,lineSize) + end + | NONE => NONE + end + + and forceBreakChunks cumulatingChunks chunks = + case chunks of + [] => NONE + | chunk :: chunks => + case forceBreakChunk (chunks :: cumulatingChunks) chunk of + SOME (breaks,chunk,lineIndent,lineSize) => + let + val chunks = [chunk] + in + SOME (breaks,chunks,lineIndent,lineSize) + end + | NONE => + case forceBreakChunks cumulatingChunks chunks of + SOME (breaks,chunks,lineIndent,lineSize) => + let + val chunks = chunk :: chunks + + val lineSize = lineSize + sizeChunk chunk + in + SOME (breaks,chunks,lineIndent,lineSize) + end + | NONE => NONE + + and forceBreakChunk cumulatingChunks chunk = + case chunk of + StringChunk _ => NONE + | BreakChunk _ => raise Bug "Print.forceBreakChunk: BreakChunk" + | BlockChunk block => + case forceBreakBlock cumulatingChunks block of + SOME (breaks,block,lineIndent,lineSize) => + let + val chunk = BlockChunk block + in + SOME (breaks,chunk,lineIndent,lineSize) + end + | NONE => NONE; + + fun forceBreak cumulatingChunks blocks' blocks = + case blocks of + [] => NONE + | block :: blocks => + let + val cumulatingChunks = + case cumulatingChunks of + [] => raise Bug "Print.forceBreak: null cumulatingChunks" + | _ :: cumulatingChunks => cumulatingChunks + + val blocks' = + case blocks' of + [] => raise Bug "Print.forceBreak: null blocks'" + | _ :: blocks' => blocks' + in + case forceBreakBlock cumulatingChunks block of + SOME (breaks,block,lineIndent,lineSize) => + let + val blocks = block :: blocks' + in + SOME (breaks,blocks,lineIndent,lineSize) + end + | NONE => + case forceBreak cumulatingChunks blocks' blocks of + SOME (breaks,blocks,lineIndent,lineSize) => + let + val blocks = block :: blocks + + val Block {size,...} = block + + val lineSize = lineSize + size + in + SOME (breaks,blocks,lineIndent,lineSize) + end + | NONE => NONE + end; + + fun normalize lineLength lines state = + let + val State {blocks,lineIndent,lineSize} = state + in + if lineIndent + lineSize <= lineLength then (lines,state) + else + let + val break = + case naturalBreak blocks of + Left break => SOME break + | Right (c,b) => forceBreak c b blocks + in + case break of + SOME (breaks,blocks,lineIndent',lineSize) => + let + val lines = renderBreaks lineIndent lineIndent' breaks lines + + val state = + State + {blocks = blocks, + lineIndent = lineIndent', + lineSize = lineSize} + in + normalize lineLength lines state + end + | NONE => (lines,state) + end + end; + +(*BasicDebug + val normalize = fn lineLength => fn lines => fn state => + let + val () = checkState state + in + normalize lineLength lines state + end + handle Bug bug => + raise Bug ("Print.normalize: before normalize:\n" ^ bug) +*) + + fun executeBeginBlock (style,ind) lines state = + let + val State {blocks,lineIndent,lineSize} = state + + val Block {indent,...} = + case blocks of + [] => raise Bug "Print.executeBeginBlock: no block" + | block :: _ => block + + val indent = indent + ind + + val style = liftStyle style + + val size = 0 + + val chunks = [] + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + + val blocks = block :: blocks + + val state = + State + {blocks = blocks, + lineIndent = lineIndent, + lineSize = lineSize} + in + (lines,state) + end; + + fun executeEndBlock lines state = + let + val State {blocks,lineIndent,lineSize} = state + + val (lineSize,blocks) = + case blocks of + [] => raise Bug "Print.executeEndBlock: no block" + | topBlock :: blocks => + let + val Block + {indent = topIndent, + style = topStyle, + size = topSize, + chunks = topChunks} = topBlock + in + case topChunks of + [] => (lineSize,blocks) + | headTopChunks :: tailTopChunks => + let + val (lineSize,topSize,topChunks) = + case headTopChunks of + BreakChunk spaces => + let + val lineSize = lineSize - spaces + and topSize = topSize - spaces + and topChunks = tailTopChunks + in + (lineSize,topSize,topChunks) + end + | _ => (lineSize,topSize,topChunks) + + val topBlock = + Block + {indent = topIndent, + style = topStyle, + size = topSize, + chunks = topChunks} + in + case blocks of + [] => raise Error "Print.executeEndBlock: no block" + | block :: blocks => + let + val Block {indent,style,size,chunks} = block + + val size = size + topSize + + val chunks = BlockChunk topBlock :: chunks + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + + val blocks = block :: blocks + in + (lineSize,blocks) + end + end + end + + val state = + State + {blocks = blocks, + lineIndent = lineIndent, + lineSize = lineSize} + in + (lines,state) + end; + + fun executeAddString lineLength s lines state = + let + val State {blocks,lineIndent,lineSize} = state + + val n = size s + + val blocks = + case blocks of + [] => raise Bug "Print.executeAddString: no block" + | Block {indent,style,size,chunks} :: blocks => + let + val size = size + n + + val chunk = StringChunk {size = n, string = s} + + val chunks = chunk :: chunks + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + + val blocks = block :: blocks + in + blocks + end + + val lineSize = lineSize + n + + val state = + State + {blocks = blocks, + lineIndent = lineIndent, + lineSize = lineSize} + in + normalize lineLength lines state + end; + + fun executeAddBreak lineLength spaces lines state = + let + val State {blocks,lineIndent,lineSize} = state + + val (blocks,lineSize) = + case blocks of + [] => raise Bug "Print.executeAddBreak: no block" + | Block {indent,style,size,chunks} :: blocks' => + case chunks of + [] => (blocks,lineSize) + | chunk :: chunks' => + let + val spaces = + case style of + BreakingBlock => lineLength + 1 + | _ => spaces + + val size = size + spaces + + val chunks = + case chunk of + BreakChunk k => BreakChunk (k + spaces) :: chunks' + | _ => BreakChunk spaces :: chunks + + val block = + Block + {indent = indent, + style = style, + size = size, + chunks = chunks} + + val blocks = block :: blocks' + + val lineSize = lineSize + spaces + in + (blocks,lineSize) + end + + val state = + State + {blocks = blocks, + lineIndent = lineIndent, + lineSize = lineSize} + in + normalize lineLength lines state + end; + + fun executeBigBreak lineLength lines state = + executeAddBreak lineLength (lineLength + 1) lines state; + + fun executeAddNewline lineLength lines state = + let + val (lines,state) = executeAddString lineLength "" lines state + val (lines,state) = executeBigBreak lineLength lines state + in + executeAddString lineLength "" lines state + end; + + fun final lineLength lines state = + let + val lines = + if isFinalState state then lines + else + let + val (lines,state) = executeBigBreak lineLength lines state + +(*BasicDebug + val _ = isFinalState state orelse raise Bug "Print.final" +*) + in + lines + end + in + if null lines then Stream.Nil else Stream.singleton lines + end; +in + fun execute {lineLength} = + let + fun advance step state = + let + val lines = [] + in + case step of + BeginBlock style_ind => executeBeginBlock style_ind lines state + | EndBlock => executeEndBlock lines state + | AddString s => executeAddString lineLength s lines state + | AddBreak spaces => executeAddBreak lineLength spaces lines state + | AddNewline => executeAddNewline lineLength lines state + end + +(*BasicDebug + val advance = fn step => fn state => + let + val (lines,state) = advance step state + val () = checkState state + in + (lines,state) + end + handle Bug bug => + raise Bug ("Print.advance: after " ^ ppStepToString step ^ + " command:\n" ^ bug) +*) + in + revConcat o Stream.maps advance (final lineLength []) initialState + end; +end; + +(* ------------------------------------------------------------------------- *) +(* Executing pretty-printers with a global line length. *) +(* ------------------------------------------------------------------------- *) + +val lineLength = ref initialLineLength; + +fun toStream ppA a = + Stream.map (fn s => s ^ "\n") + (execute {lineLength = !lineLength} (ppA a)); + +fun toString ppA a = + case execute {lineLength = !lineLength} (ppA a) of + Stream.Nil => "" + | Stream.Cons (h,t) => Stream.foldl (fn (s,z) => z ^ "\n" ^ s) h (t ()); + +fun trace ppX nameX x = + Useful.trace (toString (ppOp2 " =" ppString ppX) (nameX,x) ^ "\n"); + +end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Problem.sig --- a/src/Tools/Metis/src/Problem.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Problem.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) -(* SOME SAMPLE PROBLEMS TO TEST PROOF PROCEDURES *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* CNF PROBLEMS *) +(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Problem = @@ -10,16 +10,22 @@ (* Problems. *) (* ------------------------------------------------------------------------- *) -type problem = Thm.clause list +type problem = + {axioms : Thm.clause list, + conjecture : Thm.clause list} val size : problem -> {clauses : int, literals : int, symbols : int, typedSymbols : int} -val fromGoal : Formula.formula -> problem list +val freeVars : problem -> NameSet.set + +val toClauses : problem -> Thm.clause list -val toClauses : problem -> Formula.formula +val toFormula : problem -> Formula.formula + +val toGoal : problem -> Formula.formula val toString : problem -> string diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Problem.sml --- a/src/Tools/Metis/src/Problem.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Problem.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) -(* SOME SAMPLE PROBLEMS TO TEST PROOF PROCEDURES *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* CNF PROBLEMS *) +(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Problem :> Problem = @@ -12,83 +12,56 @@ (* Problems. *) (* ------------------------------------------------------------------------- *) -type problem = Thm.clause list; +type problem = + {axioms : Thm.clause list, + conjecture : Thm.clause list}; -fun size cls = - {clauses = length cls, - literals = foldl (fn (cl,n) => n + LiteralSet.size cl) 0 cls, - symbols = foldl (fn (cl,n) => n + LiteralSet.symbols cl) 0 cls, - typedSymbols = foldl (fn (cl,n) => n + LiteralSet.typedSymbols cl) 0 cls}; +fun toClauses {axioms,conjecture} = axioms @ conjecture; -fun checkFormula {models,checks} exp fm = +fun size prob = let - fun check 0 = true - | check n = - let - val N = 3 + Portable.randomInt 3 - val M = Model.new {size = N, fixed = Model.fixedPure} - val {T,F} = Model.checkFormula {maxChecks = checks} M fm - in - (if exp then F = 0 else T = 0) andalso check (n - 1) - end + fun lits (cl,n) = n + LiteralSet.size cl + + fun syms (cl,n) = n + LiteralSet.symbols cl + + fun typedSyms (cl,n) = n + LiteralSet.typedSymbols cl + + val cls = toClauses prob in - check models + {clauses = length cls, + literals = foldl lits 0 cls, + symbols = foldl syms 0 cls, + typedSymbols = foldl typedSyms 0 cls} end; -val checkGoal = checkFormula {models = 10, checks = 100} true; - -val checkClauses = checkFormula {models = 10, checks = 100} false; - -fun fromGoal goal = - let - fun fromLiterals fms = LiteralSet.fromList (map Literal.fromFormula fms) - - fun fromClause fm = fromLiterals (Formula.stripDisj fm) - - fun fromCnf fm = map fromClause (Formula.stripConj fm) +fun freeVars {axioms,conjecture} = + NameSet.union + (LiteralSet.freeVarsList axioms) + (LiteralSet.freeVarsList conjecture); -(*DEBUG - val () = if checkGoal goal then () - else raise Error "goal failed the finite model test" -*) - - val refute = Formula.Not (Formula.generalize goal) - -(*TRACE2 - val () = Parser.ppTrace Formula.pp "Problem.fromGoal: refute" refute -*) - - val cnfs = Normalize.cnf refute +local + fun clauseToFormula cl = + Formula.listMkDisj (LiteralSet.transform Literal.toFormula cl); +in + fun toFormula prob = + Formula.listMkConj (map clauseToFormula (toClauses prob)); -(* - val () = - let - fun check fm = - let -(*TRACE2 - val () = Parser.ppTrace Formula.pp "Problem.fromGoal: cnf" fm -*) - in - if checkClauses fm then () - else raise Error "cnf failed the finite model test" - end - in - app check cnfs - end -*) - in - map fromCnf cnfs - end; + fun toGoal {axioms,conjecture} = + let + val clToFm = Formula.generalize o clauseToFormula + val clsToFm = Formula.listMkConj o map clToFm -fun toClauses cls = - let - fun formulize cl = - Formula.listMkDisj (LiteralSet.transform Literal.toFormula cl) - in - Formula.listMkConj (map formulize cls) - end; + val fm = Formula.False + val fm = + if null conjecture then fm + else Formula.Imp (clsToFm conjecture, fm) + val fm = Formula.Imp (clsToFm axioms, fm) + in + fm + end; +end; -fun toString cls = Formula.toString (toClauses cls); +fun toString prob = Formula.toString (toFormula prob); (* ------------------------------------------------------------------------- *) (* Categorizing problems. *) @@ -117,8 +90,10 @@ equality : equality, horn : horn}; -fun categorize cls = +fun categorize prob = let + val cls = toClauses prob + val rels = let fun f (cl,set) = NameAritySet.union set (LiteralSet.relations cl) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Proof.sig --- a/src/Tools/Metis/src/Proof.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Proof.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* PROOFS IN FIRST ORDER LOGIC *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Proof = @@ -39,14 +39,22 @@ val proof : Thm.thm -> proof (* ------------------------------------------------------------------------- *) +(* Free variables. *) +(* ------------------------------------------------------------------------- *) + +val freeIn : Term.var -> proof -> bool + +val freeVars : proof -> NameSet.set + +(* ------------------------------------------------------------------------- *) (* Printing. *) (* ------------------------------------------------------------------------- *) -val ppInference : inference Parser.pp +val ppInference : inference Print.pp val inferenceToString : inference -> string -val pp : proof Parser.pp +val pp : proof Print.pp val toString : proof -> string diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Proof.sml --- a/src/Tools/Metis/src/Proof.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Proof.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* PROOFS IN FIRST ORDER LOGIC *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Proof :> Proof = @@ -34,120 +34,117 @@ | inferenceType (Equality _) = Thm.Equality; local - fun ppAssume pp atm = (Parser.addBreak pp (1,0); Atom.pp pp atm); + fun ppAssume atm = Print.sequence (Print.addBreak 1) (Atom.pp atm); - fun ppSubst ppThm pp (sub,thm) = - (Parser.addBreak pp (1,0); - Parser.beginBlock pp Parser.Inconsistent 1; - Parser.addString pp "{"; - Parser.ppBinop " =" Parser.ppString Subst.pp pp ("sub",sub); - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Parser.ppBinop " =" Parser.ppString ppThm pp ("thm",thm); - Parser.addString pp "}"; - Parser.endBlock pp); + fun ppSubst ppThm (sub,thm) = + Print.sequence (Print.addBreak 1) + (Print.blockProgram Print.Inconsistent 1 + [Print.addString "{", + Print.ppOp2 " =" Print.ppString Subst.pp ("sub",sub), + Print.addString ",", + Print.addBreak 1, + Print.ppOp2 " =" Print.ppString ppThm ("thm",thm), + Print.addString "}"]); - fun ppResolve ppThm pp (res,pos,neg) = - (Parser.addBreak pp (1,0); - Parser.beginBlock pp Parser.Inconsistent 1; - Parser.addString pp "{"; - Parser.ppBinop " =" Parser.ppString Atom.pp pp ("res",res); - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Parser.ppBinop " =" Parser.ppString ppThm pp ("pos",pos); - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Parser.ppBinop " =" Parser.ppString ppThm pp ("neg",neg); - Parser.addString pp "}"; - Parser.endBlock pp); + fun ppResolve ppThm (res,pos,neg) = + Print.sequence (Print.addBreak 1) + (Print.blockProgram Print.Inconsistent 1 + [Print.addString "{", + Print.ppOp2 " =" Print.ppString Atom.pp ("res",res), + Print.addString ",", + Print.addBreak 1, + Print.ppOp2 " =" Print.ppString ppThm ("pos",pos), + Print.addString ",", + Print.addBreak 1, + Print.ppOp2 " =" Print.ppString ppThm ("neg",neg), + Print.addString "}"]); - fun ppRefl pp tm = (Parser.addBreak pp (1,0); Term.pp pp tm); + fun ppRefl tm = Print.sequence (Print.addBreak 1) (Term.pp tm); - fun ppEquality pp (lit,path,res) = - (Parser.addBreak pp (1,0); - Parser.beginBlock pp Parser.Inconsistent 1; - Parser.addString pp "{"; - Parser.ppBinop " =" Parser.ppString Literal.pp pp ("lit",lit); - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Parser.ppBinop " =" Parser.ppString Term.ppPath pp ("path",path); - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Parser.ppBinop " =" Parser.ppString Term.pp pp ("res",res); - Parser.addString pp "}"; - Parser.endBlock pp); + fun ppEquality (lit,path,res) = + Print.sequence (Print.addBreak 1) + (Print.blockProgram Print.Inconsistent 1 + [Print.addString "{", + Print.ppOp2 " =" Print.ppString Literal.pp ("lit",lit), + Print.addString ",", + Print.addBreak 1, + Print.ppOp2 " =" Print.ppString Term.ppPath ("path",path), + Print.addString ",", + Print.addBreak 1, + Print.ppOp2 " =" Print.ppString Term.pp ("res",res), + Print.addString "}"]); - fun ppInf ppAxiom ppThm pp inf = + fun ppInf ppAxiom ppThm inf = let val infString = Thm.inferenceTypeToString (inferenceType inf) in - Parser.beginBlock pp Parser.Inconsistent (size infString + 1); - Parser.ppString pp infString; - case inf of - Axiom cl => ppAxiom pp cl - | Assume x => ppAssume pp x - | Subst x => ppSubst ppThm pp x - | Resolve x => ppResolve ppThm pp x - | Refl x => ppRefl pp x - | Equality x => ppEquality pp x; - Parser.endBlock pp + Print.block Print.Inconsistent 2 + (Print.sequence + (Print.addString infString) + (case inf of + Axiom cl => ppAxiom cl + | Assume x => ppAssume x + | Subst x => ppSubst ppThm x + | Resolve x => ppResolve ppThm x + | Refl x => ppRefl x + | Equality x => ppEquality x)) end; - fun ppAxiom pp cl = - (Parser.addBreak pp (1,0); - Parser.ppMap - LiteralSet.toList - (Parser.ppBracket "{" "}" (Parser.ppSequence "," Literal.pp)) pp cl); + fun ppAxiom cl = + Print.sequence + (Print.addBreak 1) + (Print.ppMap + LiteralSet.toList + (Print.ppBracket "{" "}" (Print.ppOpList "," Literal.pp)) cl); in val ppInference = ppInf ppAxiom Thm.pp; - fun pp p prf = + fun pp prf = let fun thmString n = "(" ^ Int.toString n ^ ")" - + val prf = enumerate prf - fun ppThm p th = + fun ppThm th = + Print.addString let val cl = Thm.clause th fun pred (_,(th',_)) = LiteralSet.equal (Thm.clause th') cl in case List.find pred prf of - NONE => Parser.addString p "(?)" - | SOME (n,_) => Parser.addString p (thmString n) + NONE => "(?)" + | SOME (n,_) => thmString n end fun ppStep (n,(th,inf)) = let val s = thmString n in - Parser.beginBlock p Parser.Consistent (1 + size s); - Parser.addString p (s ^ " "); - Thm.pp p th; - Parser.addBreak p (2,0); - Parser.ppBracket "[" "]" (ppInf (K (K ())) ppThm) p inf; - Parser.endBlock p; - Parser.addNewline p + Print.sequence + (Print.blockProgram Print.Consistent (1 + size s) + [Print.addString (s ^ " "), + Thm.pp th, + Print.addBreak 2, + Print.ppBracket "[" "]" (ppInf (K Print.skip) ppThm) inf]) + Print.addNewline end in - Parser.beginBlock p Parser.Consistent 0; - Parser.addString p "START OF PROOF"; - Parser.addNewline p; - app ppStep prf; - Parser.addString p "END OF PROOF"; - Parser.addNewline p; - Parser.endBlock p + Print.blockProgram Print.Consistent 0 + [Print.addString "START OF PROOF", + Print.addNewline, + Print.program (map ppStep prf), + Print.addString "END OF PROOF"] end -(*DEBUG +(*MetisDebug handle Error err => raise Bug ("Proof.pp: shouldn't fail:\n" ^ err); *) end; -val inferenceToString = Parser.toString ppInference; +val inferenceToString = Print.toString ppInference; -val toString = Parser.toString pp; +val toString = Print.toString pp; (* ------------------------------------------------------------------------- *) (* Reconstructing single inferences. *) @@ -172,9 +169,9 @@ let fun recon [] = let -(*TRACE3 - val () = Parser.ppTrace LiteralSet.pp "reconstructSubst: cl" cl - val () = Parser.ppTrace LiteralSet.pp "reconstructSubst: cl'" cl' +(*MetisTrace3 + val () = Print.trace LiteralSet.pp "reconstructSubst: cl" cl + val () = Print.trace LiteralSet.pp "reconstructSubst: cl'" cl' *) in raise Bug "can't reconstruct Subst rule" @@ -194,7 +191,7 @@ in Subst.normalize (recon [(LiteralSet.toList cl, Subst.empty)]) end -(*DEBUG +(*MetisDebug handle Error err => raise Bug ("Proof.recontructSubst: shouldn't fail:\n" ^ err); *) @@ -214,32 +211,37 @@ if not (LiteralSet.null lits) then LiteralSet.pick lits else raise Bug "can't reconstruct Resolve rule" end) -(*DEBUG +(*MetisDebug handle Error err => raise Bug ("Proof.recontructResolvant: shouldn't fail:\n" ^ err); *) fun reconstructEquality cl = let -(*TRACE3 - val () = Parser.ppTrace LiteralSet.pp "Proof.reconstructEquality: cl" cl +(*MetisTrace3 + val () = Print.trace LiteralSet.pp "Proof.reconstructEquality: cl" cl *) fun sync s t path (f,a) (f',a') = - if f <> f' orelse length a <> length a' then NONE + if not (Name.equal f f' andalso length a = length a') then NONE else - case List.filter (op<> o snd) (enumerate (zip a a')) of - [(i,(tm,tm'))] => - let - val path = i :: path - in - if tm = s andalso tm' = t then SOME (rev path) - else - case (tm,tm') of - (Term.Fn f_a, Term.Fn f_a') => sync s t path f_a f_a' - | _ => NONE - end - | _ => NONE + let + val itms = enumerate (zip a a') + in + case List.filter (not o uncurry Term.equal o snd) itms of + [(i,(tm,tm'))] => + let + val path = i :: path + in + if Term.equal tm s andalso Term.equal tm' t then + SOME (rev path) + else + case (tm,tm') of + (Term.Fn f_a, Term.Fn f_a') => sync s t path f_a f_a' + | _ => NONE + end + | _ => NONE + end fun recon (neq,(pol,atm),(pol',atm')) = if pol = pol' then NONE @@ -248,9 +250,9 @@ val (s,t) = Literal.destNeq neq val path = - if s <> t then sync s t [] atm atm' - else if atm <> atm' then NONE - else Atom.find (equal s) atm + if not (Term.equal s t) then sync s t [] atm atm' + else if not (Atom.equal atm atm') then NONE + else Atom.find (Term.equal s) atm in case path of SOME path => SOME ((pol',atm),path,t) @@ -264,10 +266,10 @@ | ([l1],[l2]) => [(l1,l1,l2),(l1,l2,l1)] | _ => raise Bug "reconstructEquality: malformed" -(*TRACE3 +(*MetisTrace3 val ppCands = - Parser.ppList (Parser.ppTriple Literal.pp Literal.pp Literal.pp) - val () = Parser.ppTrace ppCands + Print.ppList (Print.ppTriple Literal.pp Literal.pp Literal.pp) + val () = Print.trace ppCands "Proof.reconstructEquality: candidates" candidates *) in @@ -275,7 +277,7 @@ SOME info => info | NONE => raise Bug "can't reconstruct Equality rule" end -(*DEBUG +(*MetisDebug handle Error err => raise Bug ("Proof.recontructEquality: shouldn't fail:\n" ^ err); *) @@ -304,25 +306,25 @@ in fun thmToInference th = let -(*TRACE3 - val () = Parser.ppTrace Thm.pp "Proof.thmToInference: th" th +(*MetisTrace3 + val () = Print.trace Thm.pp "Proof.thmToInference: th" th *) val cl = Thm.clause th val thmInf = Thm.inference th -(*TRACE3 - val ppThmInf = Parser.ppPair Thm.ppInferenceType (Parser.ppList Thm.pp) - val () = Parser.ppTrace ppThmInf "Proof.thmToInference: thmInf" thmInf +(*MetisTrace3 + val ppThmInf = Print.ppPair Thm.ppInferenceType (Print.ppList Thm.pp) + val () = Print.trace ppThmInf "Proof.thmToInference: thmInf" thmInf *) val inf = reconstruct cl thmInf -(*TRACE3 - val () = Parser.ppTrace ppInference "Proof.thmToInference: inf" inf +(*MetisTrace3 + val () = Print.trace ppInference "Proof.thmToInference: inf" inf *) -(*DEBUG +(*MetisDebug val () = let val th' = inferenceToThm inf @@ -340,7 +342,7 @@ in inf end -(*DEBUG +(*MetisDebug handle Error err => raise Bug ("Proof.thmToInference: shouldn't fail:\n" ^ err); *) @@ -351,39 +353,100 @@ (* ------------------------------------------------------------------------- *) local - fun thmCompare (th1,th2) = - LiteralSet.compare (Thm.clause th1, Thm.clause th2); + val emptyThms : Thm.thm LiteralSetMap.map = LiteralSetMap.new (); + + fun addThms (th,ths) = + let + val cl = Thm.clause th + in + if LiteralSetMap.inDomain cl ths then ths + else + let + val (_,pars) = Thm.inference th + val ths = List.foldl addThms ths pars + in + if LiteralSetMap.inDomain cl ths then ths + else LiteralSetMap.insert ths (cl,th) + end + end; + + fun mkThms th = addThms (th,emptyThms); - fun buildProof (th,(m,l)) = - if Map.inDomain th m then (m,l) - else - let - val (_,deps) = Thm.inference th - val (m,l) = foldl buildProof (m,l) deps - in - if Map.inDomain th m then (m,l) - else - let - val l = (th, thmToInference th) :: l - in - (Map.insert m (th,l), l) - end - end; + fun addProof (th,(ths,acc)) = + let + val cl = Thm.clause th + in + case LiteralSetMap.peek ths cl of + NONE => (ths,acc) + | SOME th => + let + val (_,pars) = Thm.inference th + val (ths,acc) = List.foldl addProof (ths,acc) pars + val ths = LiteralSetMap.delete ths cl + val acc = (th, thmToInference th) :: acc + in + (ths,acc) + end + end; + + fun mkProof ths th = + let + val (ths,acc) = addProof (th,(ths,[])) +(*MetisTrace4 + val () = Print.trace Print.ppInt "Proof.proof: unnecessary clauses" (LiteralSetMap.size ths) +*) + in + rev acc + end; in fun proof th = let -(*TRACE3 - val () = Parser.ppTrace Thm.pp "Proof.proof: th" th +(*MetisTrace3 + val () = Print.trace Thm.pp "Proof.proof: th" th *) - val (m,_) = buildProof (th, (Map.new thmCompare, [])) -(*TRACE3 - val () = Parser.ppTrace Parser.ppInt "Proof.proof: size" (Map.size m) + val ths = mkThms th + val infs = mkProof ths th +(*MetisTrace3 + val () = Print.trace Print.ppInt "Proof.proof: size" (length infs) *) in - case Map.peek m th of - SOME l => rev l - | NONE => raise Bug "Proof.proof" + infs end; end; +(* ------------------------------------------------------------------------- *) +(* Free variables. *) +(* ------------------------------------------------------------------------- *) + +fun freeIn v = + let + fun free th_inf = + case th_inf of + (_, Axiom lits) => LiteralSet.freeIn v lits + | (_, Assume atm) => Atom.freeIn v atm + | (th, Subst _) => Thm.freeIn v th + | (_, Resolve _) => false + | (_, Refl tm) => Term.freeIn v tm + | (_, Equality (lit,_,tm)) => + Literal.freeIn v lit orelse Term.freeIn v tm + in + List.exists free + end; + +val freeVars = + let + fun inc (th_inf,set) = + NameSet.union set + (case th_inf of + (_, Axiom lits) => LiteralSet.freeVars lits + | (_, Assume atm) => Atom.freeVars atm + | (th, Subst _) => Thm.freeVars th + | (_, Resolve _) => NameSet.empty + | (_, Refl tm) => Term.freeVars tm + | (_, Equality (lit,_,tm)) => + NameSet.union (Literal.freeVars lit) (Term.freeVars tm)) + in + List.foldl inc NameSet.empty + end; + end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Random.sig --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/Metis/src/Random.sig Mon Sep 13 21:24:10 2010 +0200 @@ -0,0 +1,20 @@ +(* 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; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Random.sml --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/Metis/src/Random.sml Mon Sep 13 21:24:10 2010 +0200 @@ -0,0 +1,48 @@ +(* 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.*)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; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/RandomMap.sml --- a/src/Tools/Metis/src/RandomMap.sml Mon Sep 13 16:44:20 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,623 +0,0 @@ -(* ========================================================================= *) -(* FINITE MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES *) -(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *) -(* ========================================================================= *) - -structure RandomMap :> Map = -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,'b) tree = - E - | T of - {size : int, - priority : priority, - left : ('a,'b) tree, - key : 'a, - value : 'b, - right : ('a,'b) tree}; - -type ('a,'b) node = - {size : int, - priority : priority, - left : ('a,'b) tree, - key : 'a, - value : 'b, - right : ('a,'b) tree}; - -datatype ('a,'b) map = Map of ('a * 'a -> order) * ('a,'b) tree; - -(* ------------------------------------------------------------------------- *) -(* Random priorities. *) -(* ------------------------------------------------------------------------- *) - -local - val randomPriority = randomWord; - - val priorityOrder = Word.compare; -in - fun treeSingleton (key,value) = - T {size = 1, priority = randomPriority (), - left = E, key = key, value = value, right = E}; - - fun nodePriorityOrder cmp (x1 : ('a,'b) node, x2 : ('a,'b) 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 (m as Map (cmp,tree)) = - (let - val _ = checkSizes tree - val _ = checkSorted cmp NONE tree - val _ = checkPriorities cmp tree - in - m - end - handle Error err => raise Bug err) - handle Bug bug => raise Bug (s ^ "\nRandomMap.checkWellformed: " ^ bug); -end; - -(* ------------------------------------------------------------------------- *) -(* Basic operations. *) -(* ------------------------------------------------------------------------- *) - -fun comparison (Map (cmp,_)) = cmp; - -fun new cmp = Map (cmp,E); - -fun treeSize E = 0 - | treeSize (T {size = s, ...}) = s; - -fun size (Map (_,tree)) = treeSize tree; - -fun mkT p l k v r = - T {size = treeSize l + 1 + treeSize r, priority = p, - left = l, key = k, value = v, right = r}; - -fun singleton cmp key_value = Map (cmp, treeSingleton key_value); - -local - fun treePeek cmp E pkey = NONE - | treePeek cmp (T {left,key,value,right,...}) pkey = - case cmp (pkey,key) of - LESS => treePeek cmp left pkey - | EQUAL => SOME value - | GREATER => treePeek cmp right pkey -in - fun peek (Map (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, value = v2, right = r2, ...} = x2 - in - mkT p2 (treeAppend cmp t1 l2) k2 v2 r2 - end - | EQUAL => raise Bug "RandomSet.treeAppend: equal keys" - | GREATER => - let - val {priority = p1, - left = l1, key = k1, value = v1, right = r1, ...} = x1 - in - mkT p1 l1 k1 v1 (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,value,...} : ('a,'b) node) :: xs) t = - mkLeft xs (mkT priority left key value t); - - fun mkRight [] t = t - | mkRight (({priority,key,value,right,...} : ('a,'b) node) :: xs) t = - mkRight xs (mkT priority t key value 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,value,right,...}) = - case cmp (pkey,key) of - LESS => treePart cmp pkey lefts (x :: rights) left - | EQUAL => (mkLeft lefts left, SOME (key,value), 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 f (t1 as T x1) (t2 as T x2) = - let - val {priority = p1, - left = l1, key = k1, value = v1, right = r1, ...} = x1 - val (l2,k2_v2,r2) = nodePartition cmp x2 k1 - val (l1,l2) = treeCombineRemove cmp f l1 l2 - and (r1,r2) = treeCombineRemove cmp f r1 r2 - in - case k2_v2 of - NONE => - if treeSize l2 + treeSize r2 = #size x2 then (t1,t2) - else (mkT p1 l1 k1 v1 r1, treeAppend cmp l2 r2) - | SOME (k2,v2) => - case f (v1,v2) of - NONE => (treeAppend cmp l1 r1, treeAppend cmp l2 r2) - | SOME v => (mkT p1 l1 k2 v 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, value = v1, 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 v1 r - end; -in - fun union f (m1 as Map (cmp,t1)) (Map (_,t2)) = - if pointerEqual (t1,t2) then m1 - else - let - val (t1,t2) = treeCombineRemove cmp f t1 t2 - in - Map (cmp, treeUnionDisjoint cmp t1 t2) - end; -end; - -(*DEBUG -val union = fn f => fn t1 => fn t2 => - checkWellformed "RandomMap.union: result" - (union f (checkWellformed "RandomMap.union: input 1" t1) - (checkWellformed "RandomMap.union: input 2" t2)); -*) - -(* intersect is a simple case of the union algorithm. *) - -local - fun treeIntersect _ _ _ E = E - | treeIntersect _ _ E _ = E - | treeIntersect cmp f (t1 as T x1) (t2 as T x2) = - let - val {priority = p1, - left = l1, key = k1, value = v1, right = r1, ...} = x1 - val (l2,k2_v2,r2) = nodePartition cmp x2 k1 - val l = treeIntersect cmp f l1 l2 - and r = treeIntersect cmp f r1 r2 - in - case k2_v2 of - NONE => treeAppend cmp l r - | SOME (k2,v2) => - case f (v1,v2) of - NONE => treeAppend cmp l r - | SOME v => mkT p1 l k2 v r - end; -in - fun intersect f (m1 as Map (cmp,t1)) (Map (_,t2)) = - if pointerEqual (t1,t2) then m1 - else Map (cmp, treeIntersect cmp f t1 t2); -end; - -(*DEBUG -val intersect = fn f => fn t1 => fn t2 => - checkWellformed "RandomMap.intersect: result" - (intersect f (checkWellformed "RandomMap.intersect: input 1" t1) - (checkWellformed "RandomMap.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 "RandomMap.delete: element not found" - | treeDelete cmp (T {priority,left,key,value,right,...}) dkey = - case cmp (dkey,key) of - LESS => mkT priority (treeDelete cmp left dkey) key value right - | EQUAL => treeAppend cmp left right - | GREATER => mkT priority left key value (treeDelete cmp right dkey); -in - fun delete (Map (cmp,tree)) key = Map (cmp, treeDelete cmp tree key); -end; - -(*DEBUG -val delete = fn t => fn x => - checkWellformed "RandomMap.delete: result" - (delete (checkWellformed "RandomMap.delete: input" t) x); -*) - -(* Set difference on domains *) - -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, value = v1, right = r1} = x1 - val (l2,k2_v2,r2) = nodePartition cmp x2 k1 - val l = treeDifference cmp l1 l2 - and r = treeDifference cmp r1 r2 - in - if Option.isSome k2_v2 then treeAppend cmp l r - else if treeSize l + treeSize r + 1 = s1 then t1 - else mkT p1 l k1 v1 r - end; -in - fun difference (Map (cmp,tree1)) (Map (_,tree2)) = - Map (cmp, treeDifference cmp tree1 tree2); -end; - -(*DEBUG -val difference = fn t1 => fn t2 => - checkWellformed "RandomMap.difference: result" - (difference (checkWellformed "RandomMap.difference: input 1" t1) - (checkWellformed "RandomMap.difference: input 2" t2)); -*) - -(* subsetDomain is mainly used when using maps as sets. *) - -local - fun treeSubsetDomain _ E _ = true - | treeSubsetDomain _ _ E = false - | treeSubsetDomain 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_v2,r2) = nodePartition cmp x2 k1 - in - Option.isSome k2_v2 andalso - treeSubsetDomain cmp l1 l2 andalso - treeSubsetDomain cmp r1 r2 - end - end; -in - fun subsetDomain (Map (cmp,tree1)) (Map (_,tree2)) = - pointerEqual (tree1,tree2) orelse - treeSubsetDomain cmp tree1 tree2; -end; - -(* Map equality *) - -local - fun treeEqual _ _ E E = true - | treeEqual _ _ E _ = false - | treeEqual _ _ _ E = false - | treeEqual cmp veq (t1 as T x1) (T x2) = - let - val {size = s1, left = l1, key = k1, value = v1, right = r1, ...} = x1 - and {size = s2, ...} = x2 - in - s1 = s2 andalso - let - val (l2,k2_v2,r2) = nodePartition cmp x2 k1 - in - (case k2_v2 of NONE => false | SOME (_,v2) => veq v1 v2) andalso - treeEqual cmp veq l1 l2 andalso - treeEqual cmp veq r1 r2 - end - end; -in - fun equal veq (Map (cmp,tree1)) (Map (_,tree2)) = - pointerEqual (tree1,tree2) orelse - treeEqual cmp veq tree1 tree2; -end; - -(* mapPartial is the basic function for preserving the tree structure. *) -(* It applies the argument function to the elements *in order*. *) - -local - fun treeMapPartial cmp _ E = E - | treeMapPartial cmp f (T {priority,left,key,value,right,...}) = - let - val left = treeMapPartial cmp f left - and value' = f (key,value) - and right = treeMapPartial cmp f right - in - case value' of - NONE => treeAppend cmp left right - | SOME value => mkT priority left key value right - end; -in - fun mapPartial f (Map (cmp,tree)) = Map (cmp, treeMapPartial cmp f tree); -end; - -(* map is a primitive function for efficiency reasons. *) -(* It also applies the argument function to the elements *in order*. *) - -local - fun treeMap _ E = E - | treeMap f (T {size,priority,left,key,value,right}) = - let - val left = treeMap f left - and value = f (key,value) - and right = treeMap f right - in - T {size = size, priority = priority, left = left, - key = key, value = value, right = right} - end; -in - fun map f (Map (cmp,tree)) = Map (cmp, treeMap f tree); -end; - -(* nth picks the nth smallest key/value (counting from 0). *) - -local - fun treeNth E _ = raise Subscript - | treeNth (T {left,key,value,right,...}) n = - let - val k = treeSize left - in - if n = k then (key,value) - else if n < k then treeNth left n - else treeNth right (n - (k + 1)) - end; -in - fun nth (Map (_,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 ('key,'a) iterator = - LR of ('key * 'a) * ('key,'a) tree * ('key,'a) tree list - | RL of ('key * 'a) * ('key,'a) tree * ('key,'a) tree list; - -fun mkLR [] = NONE - | mkLR (T {key,value,right,...} :: l) = SOME (LR ((key,value),right,l)) - | mkLR (E :: _) = raise Bug "RandomMap.mkLR"; - -fun mkRL [] = NONE - | mkRL (T {key,value,left,...} :: l) = SOME (RL ((key,value),left,l)) - | mkRL (E :: _) = raise Bug "RandomMap.mkRL"; - -fun mkIterator (Map (_,tree)) = mkLR (leftSpine tree []); - -fun mkRevIterator (Map (_,tree)) = mkRL (rightSpine tree []); - -fun readIterator (LR (key_value,_,_)) = key_value - | readIterator (RL (key_value,_,_)) = key_value; - -fun advanceIterator (LR (_,next,l)) = mkLR (leftSpine next l) - | advanceIterator (RL (_,next,l)) = mkRL (rightSpine next l); - -(* ------------------------------------------------------------------------- *) -(* Derived operations. *) -(* ------------------------------------------------------------------------- *) - -fun null m = size m = 0; - -fun get m key = - case peek m key of - NONE => raise Error "RandomMap.get: element not found" - | SOME value => value; - -fun inDomain key m = Option.isSome (peek m key); - -fun insert m key_value = - union (SOME o snd) m (singleton (comparison m) key_value); - -(*DEBUG -val insert = fn m => fn x => - checkWellformed "RandomMap.insert: result" - (insert (checkWellformed "RandomMap.insert: input" m) x); -*) - -local - fun fold _ NONE acc = acc - | fold f (SOME iter) acc = - let - val (key,value) = readIterator iter - in - fold f (advanceIterator iter) (f (key,value,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_value = readIterator iter - in - if pred key_value then SOME key_value - 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_value = readIterator iter - in - case f key_value 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 fromList cmp l = List.foldl (fn (k_v,m) => insert m k_v) (new cmp) l; - -fun insertList m l = union (SOME o snd) m (fromList (comparison m) l); - -fun filter p = - let - fun f (key_value as (_,value)) = - if p key_value then SOME value else NONE - in - mapPartial f - end; - -fun app f m = foldl (fn (key,value,()) => f (key,value)) () m; - -fun transform f = map (fn (_,value) => f value); - -fun toList m = foldr (fn (key,value,l) => (key,value) :: l) [] m; - -fun domain m = foldr (fn (key,_,l) => key :: l) [] m; - -fun exists p m = Option.isSome (findl p m); - -fun all p m = not (exists (not o p) m); - -fun random m = - case size m of - 0 => raise Error "RandomMap.random: empty" - | n => nth m (randomInt n); - -local - fun iterCompare _ _ NONE NONE = EQUAL - | iterCompare _ _ NONE (SOME _) = LESS - | iterCompare _ _ (SOME _) NONE = GREATER - | iterCompare kcmp vcmp (SOME i1) (SOME i2) = - keyIterCompare kcmp vcmp (readIterator i1) (readIterator i2) i1 i2 - - and keyIterCompare kcmp vcmp (k1,v1) (k2,v2) i1 i2 = - case kcmp (k1,k2) of - LESS => LESS - | EQUAL => - (case vcmp (v1,v2) of - LESS => LESS - | EQUAL => - iterCompare kcmp vcmp (advanceIterator i1) (advanceIterator i2) - | GREATER => GREATER) - | GREATER => GREATER; -in - fun compare vcmp (m1,m2) = - if pointerEqual (m1,m2) then EQUAL - else - case Int.compare (size m1, size m2) of - LESS => LESS - | EQUAL => - iterCompare (comparison m1) vcmp (mkIterator m1) (mkIterator m2) - | GREATER => GREATER; -end; - -fun equalDomain m1 m2 = equal (K (K true)) m1 m2; - -fun toString m = "<" ^ (if null m then "" else Int.toString (size m)) ^ ">"; - -end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/RandomSet.sml --- a/src/Tools/Metis/src/RandomSet.sml Mon Sep 13 16:44:20 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,607 +0,0 @@ -(* ========================================================================= *) -(* 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Resolution.sig --- a/src/Tools/Metis/src/Resolution.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Resolution.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* THE RESOLUTION PROOF PROCEDURE *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Resolution = @@ -22,13 +22,15 @@ val default : parameters -val new : parameters -> Thm.thm list -> resolution +val new : + parameters -> {axioms : Thm.thm list, conjecture : Thm.thm list} -> + resolution val active : resolution -> Active.active val waiting : resolution -> Waiting.waiting -val pp : resolution Parser.pp +val pp : resolution Print.pp (* ------------------------------------------------------------------------- *) (* The main proof loop. *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Resolution.sml --- a/src/Tools/Metis/src/Resolution.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Resolution.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* THE RESOLUTION PROOF PROCEDURE *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Resolution :> Resolution = @@ -9,7 +9,7 @@ open Useful; (* ------------------------------------------------------------------------- *) -(* Parameters. *) +(* A type of resolution proof procedures. *) (* ------------------------------------------------------------------------- *) type parameters = @@ -44,11 +44,11 @@ fun waiting (Resolution {waiting = w, ...}) = w; val pp = - Parser.ppMap + Print.ppMap (fn Resolution {active,waiting,...} => "Resolution(" ^ Int.toString (Active.size active) ^ "<-" ^ Int.toString (Waiting.size waiting) ^ ")") - Parser.ppString; + Print.ppString; (* ------------------------------------------------------------------------- *) (* The main proof loop. *) @@ -65,21 +65,21 @@ fun iterate resolution = let val Resolution {parameters,active,waiting} = resolution -(*TRACE2 - val () = Parser.ppTrace Active.pp "Resolution.iterate: active" active - val () = Parser.ppTrace Waiting.pp "Resolution.iterate: waiting" waiting +(*MetisTrace2 + val () = Print.trace Active.pp "Resolution.iterate: active" active + val () = Print.trace Waiting.pp "Resolution.iterate: waiting" waiting *) in case Waiting.remove waiting of NONE => - Decided (Satisfiable (map Clause.thm (Active.saturated active))) + Decided (Satisfiable (map Clause.thm (Active.saturation active))) | SOME ((d,cl),waiting) => if Clause.isContradiction cl then Decided (Contradiction (Clause.thm cl)) else let -(*TRACE1 - val () = Parser.ppTrace Clause.pp "Resolution.iterate: cl" cl +(*MetisTrace1 + val () = Print.trace Clause.pp "Resolution.iterate: cl" cl *) val (active,cls) = Active.add active cl val waiting = Waiting.add waiting (d,cls) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Rewrite.sig --- a/src/Tools/Metis/src/Rewrite.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Rewrite.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,17 +1,29 @@ (* ========================================================================= *) (* ORDERED REWRITING FOR FIRST ORDER TERMS *) -(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Rewrite = sig (* ------------------------------------------------------------------------- *) -(* A type of rewrite systems. *) +(* Orientations of equations. *) (* ------------------------------------------------------------------------- *) datatype orient = LeftToRight | RightToLeft +val toStringOrient : orient -> string + +val ppOrient : orient Print.pp + +val toStringOrientOption : orient option -> string + +val ppOrientOption : orient option Print.pp + +(* ------------------------------------------------------------------------- *) +(* A type of rewrite systems. *) +(* ------------------------------------------------------------------------- *) + type reductionOrder = Term.term * Term.term -> order option type equationId = int @@ -34,7 +46,7 @@ val toString : rewrite -> string -val pp : rewrite Parser.pp +val pp : rewrite Print.pp (* ------------------------------------------------------------------------- *) (* Add equations into the system. *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Rewrite.sml --- a/src/Tools/Metis/src/Rewrite.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Rewrite.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* ORDERED REWRITING FOR FIRST ORDER TERMS *) -(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Rewrite :> Rewrite = @@ -9,11 +9,29 @@ open Useful; (* ------------------------------------------------------------------------- *) -(* A type of rewrite systems. *) +(* Orientations of equations. *) (* ------------------------------------------------------------------------- *) datatype orient = LeftToRight | RightToLeft; +fun toStringOrient ort = + case ort of + LeftToRight => "-->" + | RightToLeft => "<--"; + +val ppOrient = Print.ppMap toStringOrient Print.ppString; + +fun toStringOrientOption orto = + case orto of + SOME ort => toStringOrient ort + | NONE => "<->"; + +val ppOrientOption = Print.ppMap toStringOrientOption Print.ppString; + +(* ------------------------------------------------------------------------- *) +(* A type of rewrite systems. *) +(* ------------------------------------------------------------------------- *) + type reductionOrder = Term.term * Term.term -> order option; type equationId = int; @@ -59,72 +77,63 @@ fun equations (Rewrite {known,...}) = IntMap.foldr (fn (_,(eqn,_),eqns) => eqn :: eqns) [] known; -val pp = Parser.ppMap equations (Parser.ppList Rule.ppEquation); - -(*DEBUG -local - fun orientOptionToString ort = - case ort of - SOME LeftToRight => "-->" - | SOME RightToLeft => "<--" - | NONE => "<->"; +val pp = Print.ppMap equations (Print.ppList Rule.ppEquation); - open Parser; - - fun ppEq p ((x_y,_),ort) = - ppBinop (" " ^ orientOptionToString ort) Term.pp Term.pp p x_y; +(*MetisTrace1 +local + fun ppEq ((x_y,_),ort) = + Print.ppOp2 (" " ^ toStringOrientOption ort) Term.pp Term.pp x_y; - fun ppField f ppA p a = - (beginBlock p Inconsistent 2; - addString p (f ^ " ="); - addBreak p (1,0); - ppA p a; - endBlock p); + fun ppField f ppA a = + Print.blockProgram Print.Inconsistent 2 + [Print.addString (f ^ " ="), + Print.addBreak 1, + ppA a]; val ppKnown = - ppField "known" (ppMap IntMap.toList (ppList (ppPair ppInt ppEq))); + ppField "known" + (Print.ppMap IntMap.toList + (Print.ppList (Print.ppPair Print.ppInt ppEq))); val ppRedexes = - ppField - "redexes" - (TermNet.pp - (ppPair ppInt (ppMap (orientOptionToString o SOME) ppString))); + ppField "redexes" + (TermNet.pp (Print.ppPair Print.ppInt ppOrient)); val ppSubterms = - ppField - "subterms" + ppField "subterms" (TermNet.pp - (ppMap + (Print.ppMap (fn (i,l,p) => (i, (if l then 0 else 1) :: p)) - (ppPair ppInt Term.ppPath))); + (Print.ppPair Print.ppInt Term.ppPath))); - val ppWaiting = ppField "waiting" (ppMap (IntSet.toList) (ppList ppInt)); + val ppWaiting = + ppField "waiting" + (Print.ppMap (IntSet.toList) (Print.ppList Print.ppInt)); in - fun pp p (Rewrite {known,redexes,subterms,waiting,...}) = - (beginBlock p Inconsistent 2; - addString p "Rewrite"; - addBreak p (1,0); - beginBlock p Inconsistent 1; - addString p "{"; - ppKnown p known; -(*TRACE5 - addString p ","; - addBreak p (1,0); - ppRedexes p redexes; - addString p ","; - addBreak p (1,0); - ppSubterms p subterms; - addString p ","; - addBreak p (1,0); - ppWaiting p waiting; + fun pp (Rewrite {known,redexes,subterms,waiting,...}) = + Print.blockProgram Print.Inconsistent 2 + [Print.addString "Rewrite", + Print.addBreak 1, + Print.blockProgram Print.Inconsistent 1 + [Print.addString "{", + ppKnown known, +(*MetisTrace5 + Print.addString ",", + Print.addBreak 1, + ppRedexes redexes, + Print.addString ",", + Print.addBreak 1, + ppSubterms subterms, + Print.addString ",", + Print.addBreak 1, + ppWaiting waiting, *) - endBlock p; - addString p "}"; - endBlock p); + Print.skip], + Print.addString "}"] end; *) -val toString = Parser.toString pp; +val toString = Print.toString pp; (* ------------------------------------------------------------------------- *) (* Debug functions. *) @@ -137,7 +146,7 @@ NONE => false | SOME sub => order (tm, Subst.subst (Subst.normalize sub) r) = SOME GREATER - + fun knownRed tm (eqnId,(eqn,ort)) = eqnId <> id andalso ((ort <> SOME RightToLeft andalso eqnRed eqn tm) orelse @@ -191,8 +200,8 @@ Rewrite {order = order, known = known, redexes = redexes, subterms = subterms, waiting = waiting} -(*TRACE5 - val () = Parser.ppTrace pp "Rewrite.add: result" rw +(*MetisTrace5 + val () = Print.trace pp "Rewrite.add: result" rw *) in rw @@ -256,17 +265,18 @@ NONE => raise Error "incomparable" | SOME LESS => let - val sub = Subst.fromList [("x",l),("y",r)] - val th = Thm.subst sub Rule.symmetry + val th = Rule.symmetryRule l r in - fn tm => if tm = r then (l,th) else raise Error "mkNeqConv: RL" + fn tm => + if Term.equal tm r then (l,th) else raise Error "mkNeqConv: RL" end | SOME EQUAL => raise Error "irreflexive" | SOME GREATER => let val th = Thm.assume lit in - fn tm => if tm = l then (r,th) else raise Error "mkNeqConv: LR" + fn tm => + if Term.equal tm l then (r,th) else raise Error "mkNeqConv: LR" end end; @@ -321,14 +331,14 @@ | SOME lit => (false,lit) val (lit',litTh) = literule lit in - if lit = lit' then eqn + if Literal.equal lit lit' then eqn else (Literal.destEq lit', if strongEqn then th else if not (Thm.negateMember lit litTh) then litTh else Thm.resolve lit th litTh) end -(*DEBUG +(*MetisDebug handle Error err => raise Error ("Rewrite.rewriteIdEqn':\n" ^ err); *) @@ -341,7 +351,7 @@ val neq = neqConvsDelete neq lit val (lit',litTh) = mk_literule neq lit in - if lit = lit' then acc + if Literal.equal lit lit' then acc else let val th = Thm.resolve lit th litTh @@ -377,15 +387,15 @@ fun rewriteIdRule' order known redexes id th = rewriteIdLiteralsRule' order known redexes id (Thm.clause th) th; -(*DEBUG +(*MetisDebug val rewriteIdRule' = fn order => fn known => fn redexes => fn id => fn th => let -(*TRACE6 - val () = Parser.ppTrace Thm.pp "Rewrite.rewriteIdRule': th" th +(*MetisTrace6 + val () = Print.trace Thm.pp "Rewrite.rewriteIdRule': th" th *) val result = rewriteIdRule' order known redexes id th -(*TRACE6 - val () = Parser.ppTrace Thm.pp "Rewrite.rewriteIdRule': result" result +(*MetisTrace6 + val () = Print.trace Thm.pp "Rewrite.rewriteIdRule': result" result *) val _ = not (thmReducible order known id result) orelse raise Bug "rewriteIdRule: should be normalized" @@ -431,8 +441,8 @@ end; fun sameRedexes NONE _ _ = false - | sameRedexes (SOME LeftToRight) (l0,_) (l,_) = l0 = l - | sameRedexes (SOME RightToLeft) (_,r0) (_,r) = r0 = r; + | sameRedexes (SOME LeftToRight) (l0,_) (l,_) = Term.equal l0 l + | sameRedexes (SOME RightToLeft) (_,r0) (_,r) = Term.equal r0 r; fun redexResidues NONE (l,r) = [(l,r,false),(r,l,false)] | redexResidues (SOME LeftToRight) (l,r) = [(l,r,true)] @@ -455,13 +465,13 @@ else raise Error "order" end end - + fun addRed lr ((id',left,path),todo) = if id <> id' andalso not (IntSet.member id' todo) andalso can (checkValidRewr lr id' left) path then IntSet.add todo id' else todo - + fun findRed (lr as (l,_,_), todo) = List.foldl (addRed lr) todo (TermNet.matched subterms l) in @@ -473,7 +483,13 @@ val (eq0,_) = eqn0 val Rewrite {order,known,redexes,subterms,waiting} = rw val eqn as (eq,_) = rewriteIdEqn' order known redexes id eqn0 - val identical = eq = eq0 + val identical = + let + val (l0,r0) = eq0 + and (l,r) = eq + in + Term.equal l l0 andalso Term.equal r r0 + end val same_redexes = identical orelse sameRedexes ort0 eq0 eq val rpl = if same_redexes then rpl else IntSet.add rpl id val spl = if new orelse identical then spl else IntSet.add spl id @@ -543,7 +559,7 @@ case IntMap.peek known id of NONE => reds | SOME eqn_ort => addRedexes id eqn_ort reds - + val redexes = TermNet.filter filt redexes val redexes = IntSet.foldl addReds redexes rpl in @@ -560,7 +576,7 @@ case IntMap.peek known id of NONE => subtms | SOME (eqn,_) => addSubterms id eqn subtms - + val subterms = TermNet.filter filt subterms val subterms = IntSet.foldl addSubtms subterms spl in @@ -569,18 +585,21 @@ in fun rebuild rpl spl rw = let -(*TRACE5 - val ppPl = Parser.ppMap IntSet.toList (Parser.ppList Parser.ppInt) - val () = Parser.ppTrace ppPl "Rewrite.rebuild: rpl" rpl - val () = Parser.ppTrace ppPl "Rewrite.rebuild: spl" spl +(*MetisTrace5 + val ppPl = Print.ppMap IntSet.toList (Print.ppList Print.ppInt) + val () = Print.trace ppPl "Rewrite.rebuild: rpl" rpl + val () = Print.trace ppPl "Rewrite.rebuild: spl" spl *) val Rewrite {order,known,redexes,subterms,waiting} = rw val redexes = cleanRedexes known redexes rpl val subterms = cleanSubterms known subterms spl in Rewrite - {order = order, known = known, redexes = redexes, - subterms = subterms, waiting = waiting} + {order = order, + known = known, + redexes = redexes, + subterms = subterms, + waiting = waiting} end; end; @@ -608,17 +627,17 @@ if isReduced rw then (rw,[]) else reduceAcc (IntSet.empty,IntSet.empty,IntSet.empty,rw,IntSet.empty); -(*DEBUG +(*MetisDebug val reduce' = fn rw => let -(*TRACE4 - val () = Parser.ppTrace pp "Rewrite.reduce': rw" rw +(*MetisTrace4 + val () = Print.trace pp "Rewrite.reduce': rw" rw *) val Rewrite {known,order,...} = rw val result as (Rewrite {known = known', ...}, _) = reduce' rw -(*TRACE4 - val ppResult = Parser.ppPair pp (Parser.ppList Parser.ppInt) - val () = Parser.ppTrace ppResult "Rewrite.reduce': result" result +(*MetisTrace4 + val ppResult = Print.ppPair pp (Print.ppList Print.ppInt) + val () = Print.trace ppResult "Rewrite.reduce': result" result *) val ths = map (fn (id,((_,th),_)) => (id,th)) (IntMap.toList known') val _ = diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Rule.sig --- a/src/Tools/Metis/src/Rule.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Rule.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Rule = @@ -13,7 +13,7 @@ type equation = (Term.term * Term.term) * Thm.thm -val ppEquation : equation Parser.pp +val ppEquation : equation Print.pp val equationToString : equation -> string @@ -143,6 +143,8 @@ (* x = x *) (* ------------------------------------------------------------------------- *) +val reflexivityRule : Term.term -> Thm.thm + val reflexivity : Thm.thm (* ------------------------------------------------------------------------- *) @@ -151,6 +153,8 @@ (* ~(x = y) \/ y = x *) (* ------------------------------------------------------------------------- *) +val symmetryRule : Term.term -> Term.term -> Thm.thm + val symmetry : Thm.thm (* ------------------------------------------------------------------------- *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Rule.sml --- a/src/Tools/Metis/src/Rule.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Rule.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Rule :> Rule = @@ -9,12 +9,33 @@ open Useful; (* ------------------------------------------------------------------------- *) +(* Variable names. *) +(* ------------------------------------------------------------------------- *) + +val xVarName = Name.fromString "x"; +val xVar = Term.Var xVarName; + +val yVarName = Name.fromString "y"; +val yVar = Term.Var yVarName; + +val zVarName = Name.fromString "z"; +val zVar = Term.Var zVarName; + +fun xIVarName i = Name.fromString ("x" ^ Int.toString i); +fun xIVar i = Term.Var (xIVarName i); + +fun yIVarName i = Name.fromString ("y" ^ Int.toString i); +fun yIVar i = Term.Var (yIVarName i); + +(* ------------------------------------------------------------------------- *) (* *) (* --------- reflexivity *) (* x = x *) (* ------------------------------------------------------------------------- *) -val reflexivity = Thm.refl (Term.Var "x"); +fun reflexivityRule x = Thm.refl x; + +val reflexivity = reflexivityRule xVar; (* ------------------------------------------------------------------------- *) (* *) @@ -22,17 +43,17 @@ (* ~(x = y) \/ y = x *) (* ------------------------------------------------------------------------- *) -val symmetry = +fun symmetryRule x y = let - val x = Term.Var "x" - and y = Term.Var "y" - val reflTh = reflexivity + val reflTh = reflexivityRule x val reflLit = Thm.destUnit reflTh val eqTh = Thm.equality reflLit [0] y in Thm.resolve reflLit reflTh eqTh end; +val symmetry = symmetryRule xVar yVar; + (* ------------------------------------------------------------------------- *) (* *) (* --------------------------------- transitivity *) @@ -41,12 +62,9 @@ val transitivity = let - val x = Term.Var "x" - and y = Term.Var "y" - and z = Term.Var "z" - val eqTh = Thm.equality (Literal.mkEq (y,z)) [0] x + val eqTh = Thm.equality (Literal.mkEq (yVar,zVar)) [0] xVar in - Thm.resolve (Literal.mkEq (y,x)) symmetry eqTh + Thm.resolve (Literal.mkEq (yVar,xVar)) symmetry eqTh end; (* ------------------------------------------------------------------------- *) @@ -59,10 +77,10 @@ let val (x,y) = Literal.destEq lit in - if x = y then th + if Term.equal x y then th else let - val sub = Subst.fromList [("x",x),("y",y)] + val sub = Subst.fromList [(xVarName,x),(yVarName,y)] val symTh = Thm.subst sub symmetry in Thm.resolve lit th symTh @@ -76,9 +94,9 @@ type equation = (Term.term * Term.term) * Thm.thm; -fun ppEquation pp (eqn as (_,th)) = Thm.pp pp th; +fun ppEquation (_,th) = Thm.pp th; -fun equationToString x = Parser.toString ppEquation x; +val equationToString = Print.toString ppEquation; fun equationLiteral (t_u,th) = let @@ -90,7 +108,7 @@ fun reflEqn t = ((t,t), Thm.refl t); fun symEqn (eqn as ((t,u), th)) = - if t = u then eqn + if Term.equal t u then eqn else ((u,t), case equationLiteral eqn of @@ -98,9 +116,9 @@ | NONE => th); fun transEqn (eqn1 as ((x,y), th1)) (eqn2 as ((_,z), th2)) = - if x = y then eqn2 - else if y = z then eqn1 - else if x = z then reflEqn x + if Term.equal x y then eqn2 + else if Term.equal y z then eqn1 + else if Term.equal x z then reflEqn x else ((x,z), case equationLiteral eqn1 of @@ -110,7 +128,7 @@ NONE => th2 | SOME y_z => let - val sub = Subst.fromList [("x",x),("y",y),("z",z)] + val sub = Subst.fromList [(xVarName,x),(yVarName,y),(zVarName,z)] val th = Thm.subst sub transitivity val th = Thm.resolve x_y th1 th val th = Thm.resolve y_z th2 th @@ -118,7 +136,7 @@ th end); -(*DEBUG +(*MetisDebug val transEqn = fn eqn1 => fn eqn2 => transEqn eqn1 eqn2 handle Error err => @@ -149,7 +167,7 @@ handle Error err => (print (s ^ ": " ^ Term.toString tm ^ " --> Error: " ^ err ^ "\n"); raise Error (s ^ ": " ^ err)); - + fun thenConvTrans tm (tm',th1) (tm'',th2) = let val eqn1 = ((tm,tm'),th1) @@ -189,7 +207,7 @@ | everyConv (conv :: convs) tm = thenConv conv (everyConv convs) tm; fun rewrConv (eqn as ((x,y), eqTh)) path tm = - if x = y then allConv tm + if Term.equal x y then allConv tm else if null path then (y,eqTh) else let @@ -206,7 +224,7 @@ (tm',th) end; -(*DEBUG +(*MetisDebug val rewrConv = fn eqn as ((x,y),eqTh) => fn path => fn tm => rewrConv eqn path tm handle Error err => @@ -250,7 +268,7 @@ f end; -(*DEBUG +(*MetisDebug val repeatTopDownConv = fn conv => fn tm => repeatTopDownConv conv tm handle Error err => raise Error ("repeatTopDownConv: " ^ err); @@ -273,9 +291,9 @@ val res1 as (lit',th1) = literule1 lit val res2 as (lit'',th2) = literule2 lit' in - if lit = lit' then res2 - else if lit' = lit'' then res1 - else if lit = lit'' then allLiterule lit + if Literal.equal lit lit' then res2 + else if Literal.equal lit' lit'' then res1 + else if Literal.equal lit lit'' then allLiterule lit else (lit'', if not (Thm.member lit' th1) then th1 @@ -309,7 +327,7 @@ thenLiterule literule (everyLiterule literules) lit; fun rewrLiterule (eqn as ((x,y),eqTh)) path lit = - if x = y then allLiterule lit + if Term.equal x y then allLiterule lit else let val th = Thm.equality lit path y @@ -322,7 +340,7 @@ (lit',th) end; -(*DEBUG +(*MetisDebug val rewrLiterule = fn eqn => fn path => fn lit => rewrLiterule eqn path lit handle Error err => @@ -384,12 +402,12 @@ let val (lit',litTh) = literule lit in - if lit = lit' then th + if Literal.equal lit lit' then th else if not (Thm.negateMember lit litTh) then litTh else Thm.resolve lit th litTh end; -(*DEBUG +(*MetisDebug val literalRule = fn literule => fn lit => fn th => literalRule literule lit th handle Error err => @@ -412,7 +430,7 @@ fun allLiteralsRule literule th = literalsRule literule (Thm.clause th) th; fun convRule conv = allLiteralsRule (allArgumentsLiterule conv); - + (* ------------------------------------------------------------------------- *) (* *) (* ---------------------------------------------- functionCongruence (f,n) *) @@ -422,8 +440,8 @@ fun functionCongruence (f,n) = let - val xs = List.tabulate (n, fn i => Term.Var ("x" ^ Int.toString i)) - and ys = List.tabulate (n, fn i => Term.Var ("y" ^ Int.toString i)) + val xs = List.tabulate (n,xIVar) + and ys = List.tabulate (n,yIVar) fun cong ((i,yi),(th,lit)) = let @@ -449,8 +467,8 @@ fun relationCongruence (R,n) = let - val xs = List.tabulate (n, fn i => Term.Var ("x" ^ Int.toString i)) - and ys = List.tabulate (n, fn i => Term.Var ("y" ^ Int.toString i)) + val xs = List.tabulate (n,xIVar) + and ys = List.tabulate (n,yIVar) fun cong ((i,yi),(th,lit)) = let @@ -477,10 +495,10 @@ let val (x,y) = Literal.destNeq lit in - if x = y then th + if Term.equal x y then th else let - val sub = Subst.fromList [("x",y),("y",x)] + val sub = Subst.fromList [(xVarName,y),(yVarName,x)] val symTh = Thm.subst sub symmetry in Thm.resolve lit th symTh @@ -545,10 +563,12 @@ fun expand lit = let val (x,y) = Literal.destNeq lit + val _ = Term.isTypedVar x orelse Term.isTypedVar y orelse + raise Error "Rule.expandAbbrevs: no vars" + val _ = not (Term.equal x y) orelse + raise Error "Rule.expandAbbrevs: equal vars" in - if (Term.isTypedVar x orelse Term.isTypedVar y) andalso x <> y then - Subst.unify Subst.empty x y - else raise Error "expand" + Subst.unify Subst.empty x y end; in fun expandAbbrevs th = @@ -603,8 +623,8 @@ FactorEdge of Atom.atom * Atom.atom | ReflEdge of Term.term * Term.term; - fun ppEdge p (FactorEdge atm_atm') = Parser.ppPair Atom.pp Atom.pp p atm_atm' - | ppEdge p (ReflEdge tm_tm') = Parser.ppPair Term.pp Term.pp p tm_tm'; + fun ppEdge (FactorEdge atm_atm') = Print.ppPair Atom.pp Atom.pp atm_atm' + | ppEdge (ReflEdge tm_tm') = Print.ppPair Term.pp Term.pp tm_tm'; datatype joinStatus = Joined @@ -679,7 +699,7 @@ end | init_edges acc apart ((sub,edge) :: sub_edges) = let -(*DEBUG +(*MetisDebug val () = if not (Subst.null sub) then () else raise Bug "Rule.factor.init_edges: empty subst" *) @@ -732,21 +752,21 @@ in fun factor' cl = let -(*TRACE6 - val () = Parser.ppTrace LiteralSet.pp "Rule.factor': cl" cl +(*MetisTrace6 + val () = Print.trace LiteralSet.pp "Rule.factor': cl" cl *) val edges = mk_edges [] [] (LiteralSet.toList cl) -(*TRACE6 - val ppEdgesSize = Parser.ppMap length Parser.ppInt - val ppEdgel = Parser.ppList ppEdge - val ppEdges = Parser.ppList (Parser.ppTriple ppEdgel Subst.pp ppEdgel) - val () = Parser.ppTrace ppEdgesSize "Rule.factor': |edges|" edges - val () = Parser.ppTrace ppEdges "Rule.factor': edges" edges +(*MetisTrace6 + val ppEdgesSize = Print.ppMap length Print.ppInt + val ppEdgel = Print.ppList ppEdge + val ppEdges = Print.ppList (Print.ppTriple ppEdgel Subst.pp ppEdgel) + val () = Print.trace ppEdgesSize "Rule.factor': |edges|" edges + val () = Print.trace ppEdges "Rule.factor': edges" edges *) val result = fact [] edges -(*TRACE6 - val ppResult = Parser.ppList Subst.pp - val () = Parser.ppTrace ppResult "Rule.factor': result" result +(*MetisTrace6 + val ppResult = Print.ppList Subst.pp + val () = Print.trace ppResult "Rule.factor': result" result *) in result diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Set.sig --- a/src/Tools/Metis/src/Set.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Set.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,36 +1,72 @@ (* ========================================================================= *) -(* 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 *) +(* ------------------------------------------------------------------------- *) +(* Removing. *) +(* ------------------------------------------------------------------------- *) + +val delete : 'elt set -> 'elt -> 'elt set (* must be present *) + +val remove : 'elt set -> 'elt -> 'elt set -(* Union and intersect prefer elements in the second 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 @@ -44,22 +80,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 @@ -72,27 +110,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 *) +(* ------------------------------------------------------------------------- *) +(* Converting to and from maps. *) +(* ------------------------------------------------------------------------- *) -val random : 'elt set -> 'elt (* raises Empty *) +type ('elt,'a) map = ('elt,'a) Map.map -val deletePick : 'elt set -> 'elt * 'elt set (* raises Empty *) +val mapPartial : ('elt -> 'a option) -> 'elt set -> ('elt,'a) map -val deleteRandom : 'elt set -> 'elt * 'elt set (* raises Empty *) +val map : ('elt -> 'a) -> 'elt set -> ('elt,'a) map + +val domain : ('elt,'a) map -> 'elt set -val compare : 'elt set * 'elt set -> order - -val close : ('elt set -> 'elt set) -> 'elt set -> 'elt set +(* ------------------------------------------------------------------------- *) +(* Pretty-printing. *) +(* ------------------------------------------------------------------------- *) val toString : 'elt set -> string diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Set.sml --- a/src/Tools/Metis/src/Set.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Set.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,331 @@ (* ========================================================================= *) -(* 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 *) (* ========================================================================= *) -structure Set = RandomSet; +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 diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Sharing.sig --- a/src/Tools/Metis/src/Sharing.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Sharing.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,16 +1,18 @@ (* ========================================================================= *) (* PRESERVING SHARING OF ML VALUES *) -(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Sharing = sig (* ------------------------------------------------------------------------- *) -(* Pointer equality. *) +(* Option operations. *) (* ------------------------------------------------------------------------- *) -val pointerEqual : 'a * 'a -> bool +val mapOption : ('a -> 'a) -> 'a option -> 'a option + +val mapsOption : ('a -> 's -> 'a * 's) -> 'a option -> 's -> 'a option * 's (* ------------------------------------------------------------------------- *) (* List operations. *) @@ -18,6 +20,12 @@ val map : ('a -> 'a) -> 'a list -> 'a list +val revMap : ('a -> 'a) -> 'a list -> 'a list + +val maps : ('a -> 's -> 'a * 's) -> 'a list -> 's -> 'a list * 's + +val revMaps : ('a -> 's -> 'a * 's) -> 'a list -> 's -> 'a list * 's + val updateNth : int * 'a -> 'a list -> 'a list val setify : ''a list -> ''a list diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Sharing.sml --- a/src/Tools/Metis/src/Sharing.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Sharing.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* PRESERVING SHARING OF ML VALUES *) -(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Sharing :> Sharing = @@ -8,13 +8,31 @@ infix == +val op== = Portable.pointerEqual; + (* ------------------------------------------------------------------------- *) -(* Pointer equality. *) +(* Option operations. *) (* ------------------------------------------------------------------------- *) -val pointerEqual = Portable.pointerEqual; +fun mapOption f xo = + case xo of + SOME x => + let + val y = f x + in + if x == y then xo else SOME y + end + | NONE => xo; -val op== = pointerEqual; +fun mapsOption f xo acc = + case xo of + SOME x => + let + val (y,acc) = f x acc + in + if x == y then (xo,acc) else (SOME y, acc) + end + | NONE => (xo,acc); (* ------------------------------------------------------------------------- *) (* List operations. *) @@ -22,18 +40,79 @@ fun map f = let - fun m _ a_b [] = List.revAppend a_b - | m ys a_b (x :: xs) = - let - val y = f x - val ys = y :: ys - in - m ys (if x == y then a_b else (ys,xs)) xs - end + fun m ys ys_xs xs = + case xs of + [] => List.revAppend ys_xs + | x :: xs => + let + val y = f x + val ys = y :: ys + val ys_xs = if x == y then ys_xs else (ys,xs) + in + m ys ys_xs xs + end + in + fn xs => m [] ([],xs) xs + end; + +fun maps f = + let + fun m acc ys ys_xs xs = + case xs of + [] => (List.revAppend ys_xs, acc) + | x :: xs => + let + val (y,acc) = f x acc + val ys = y :: ys + val ys_xs = if x == y then ys_xs else (ys,xs) + in + m acc ys ys_xs xs + end in - fn l => m [] ([],l) l + fn xs => fn acc => m acc [] ([],xs) xs end; +local + fun revTails acc xs = + case xs of + [] => acc + | x :: xs' => revTails ((x,xs) :: acc) xs'; +in + fun revMap f = + let + fun m ys same xxss = + case xxss of + [] => ys + | (x,xs) :: xxss => + let + val y = f x + val same = same andalso x == y + val ys = if same then xs else y :: ys + in + m ys same xxss + end + in + fn xs => m [] true (revTails [] xs) + end; + + fun revMaps f = + let + fun m acc ys same xxss = + case xxss of + [] => (ys,acc) + | (x,xs) :: xxss => + let + val (y,acc) = f x acc + val same = same andalso x == y + val ys = if same then xs else y :: ys + in + m acc ys same xxss + end + in + fn xs => fn acc => m acc [] true (revTails [] xs) + end; +end; + fun updateNth (n,x) l = let val (a,b) = Useful.revDivide l n diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Stream.sig --- a/src/Tools/Metis/src/Stream.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Stream.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,22 +1,22 @@ (* ========================================================================= *) (* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Stream = sig (* ------------------------------------------------------------------------- *) -(* The stream type *) +(* The stream type. *) (* ------------------------------------------------------------------------- *) -datatype 'a stream = NIL | CONS of 'a * (unit -> 'a stream) +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 () = Stream.CONS (4,s4) in s4 () end; *) +(* val stream4 = let fun s4 () = Stream.Cons (4,s4) in s4 () end; *) (* ------------------------------------------------------------------------- *) -(* Stream constructors *) +(* Stream constructors. *) (* ------------------------------------------------------------------------- *) val repeat : 'a -> 'a stream @@ -26,7 +26,7 @@ val funpows : ('a -> 'a) -> 'a -> 'a stream (* ------------------------------------------------------------------------- *) -(* Stream versions of standard list operations: these should all terminate *) +(* Stream versions of standard list operations: these should all terminate. *) (* ------------------------------------------------------------------------- *) val cons : 'a -> (unit -> 'a stream) -> 'a stream @@ -45,7 +45,8 @@ val map : ('a -> 'b) -> 'a stream -> 'b stream -val maps : ('a -> 's -> 'b * 's) -> 's -> '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 @@ -56,7 +57,7 @@ val drop : int -> 'a stream -> 'a stream (* raises Subscript *) (* ------------------------------------------------------------------------- *) -(* Stream versions of standard list operations: these might not terminate *) +(* Stream versions of standard list operations: these might not terminate. *) (* ------------------------------------------------------------------------- *) val length : 'a stream -> int @@ -73,14 +74,26 @@ val mapPartial : ('a -> 'b option) -> 'a stream -> 'b stream -val mapsPartial : ('a -> 's -> 'b option * 's) -> 's -> '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 *) +(* 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 diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Stream.sml --- a/src/Tools/Metis/src/Stream.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Stream.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Stream :> Stream = @@ -13,60 +13,61 @@ val funpow = Useful.funpow; (* ------------------------------------------------------------------------- *) -(* The datatype declaration encapsulates all the primitive operations *) +(* The stream type. *) (* ------------------------------------------------------------------------- *) datatype 'a stream = - NIL - | CONS of 'a * (unit -> 'a stream); + Nil + | Cons of 'a * (unit -> 'a stream); (* ------------------------------------------------------------------------- *) -(* Stream constructors *) +(* Stream constructors. *) (* ------------------------------------------------------------------------- *) -fun repeat x = let fun rep () = CONS (x,rep) in rep () end; +fun repeat x = let fun rep () = Cons (x,rep) in rep () end; -fun count n = CONS (n, fn () => count (n + 1)); +fun count n = Cons (n, fn () => count (n + 1)); -fun funpows f x = CONS (x, fn () => funpows f (f x)); +fun funpows f x = Cons (x, fn () => funpows f (f x)); (* ------------------------------------------------------------------------- *) -(* Stream versions of standard list operations: these should all terminate *) +(* Stream versions of standard list operations: these should all terminate. *) (* ------------------------------------------------------------------------- *) -fun cons h t = CONS (h,t); +fun cons h t = Cons (h,t); -fun null NIL = true | null (CONS _) = false; +fun null Nil = true + | null (Cons _) = false; -fun hd NIL = raise Empty - | hd (CONS (h,_)) = h; +fun hd Nil = raise Empty + | hd (Cons (h,_)) = h; -fun tl NIL = raise Empty - | tl (CONS (_,t)) = t (); +fun tl Nil = raise Empty + | tl (Cons (_,t)) = t (); fun hdTl s = (hd s, tl s); -fun singleton s = CONS (s, K NIL); +fun singleton s = Cons (s, K Nil); -fun append NIL s = s () - | append (CONS (h,t)) s = CONS (h, fn () => append (t ()) s); +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, fn () => m (t ())) + fun m Nil = Nil + | m (Cons (h,t)) = Cons (f h, m o t) in m end; -fun maps f = +fun maps f g = let - fun mm _ NIL = NIL - | mm s (CONS (x, xs)) = + fun mm s Nil = g s + | mm s (Cons (x,xs)) = let - val (y, s') = f x s + val (y,s') = f x s in - CONS (y, fn () => mm s' (xs ())) + Cons (y, mm s' o xs) end in mm @@ -74,102 +75,129 @@ 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 ())) + 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 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 *) +(* Stream versions of standard list operations: these might not terminate. *) (* ------------------------------------------------------------------------- *) local - fun len n NIL = n - | len n (CONS (_,t)) = len (n + 1) (t ()); + 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 ()) + 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 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 ()) + 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 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)) = + fun mp Nil = Nil + | mp (Cons (h,t)) = case f h of NONE => mp (t ()) - | SOME h' => CONS (h', fn () => 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 mapsPartial f = +fun mapConcat f = let - fun mm _ NIL = NIL - | mm s (CONS (x, xs)) = + 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 (yo, s') = f x s - val t = mm s' o xs + val (l,s) = f h s in - case yo of NONE => t () | SOME y => CONS (y, t) + append l (fn () => mc s (t ())) end in - mm + mc end; (* ------------------------------------------------------------------------- *) -(* Stream operations *) +(* Stream operations. *) (* ------------------------------------------------------------------------- *) -fun memoize NIL = NIL - | memoize (CONS (h,t)) = CONS (h, Lazy.memoize (fn () => memoize (t ()))); +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 ()); + 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 fromList [] = Nil + | fromList (x :: xs) = Cons (x, fn () => fromList xs); + +fun listConcat s = concat (map fromList s); fun toString s = implode (toList s); @@ -181,8 +209,8 @@ 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 ())) + fun toFile Nil = () + | toFile (Cons (x,y)) = (TextIO.output (h,x); toFile (y ())) val () = toFile s in @@ -197,8 +225,8 @@ fun strm () = case TextIO.inputLine h of - NONE => (close h; NIL) - | SOME s => CONS (s,strm) + NONE => (close h; Nil) + | SOME s => Cons (s,strm) in memoize (strm ()) end; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Subst.sig --- a/src/Tools/Metis/src/Subst.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Subst.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC SUBSTITUTIONS *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Subst = @@ -28,8 +28,6 @@ val singleton : Term.var * Term.term -> subst -val union : subst -> subst -> subst - val toList : subst -> (Term.var * Term.term) list val fromList : (Term.var * Term.term) list -> subst @@ -38,7 +36,7 @@ val foldr : (Term.var * Term.term * 'a -> 'a) -> 'a -> subst -> 'a -val pp : subst Parser.pp +val pp : subst Print.pp val toString : subst -> string @@ -71,7 +69,13 @@ val compose : subst -> subst -> subst (* ------------------------------------------------------------------------- *) -(* Substitutions can be inverted iff they are renaming substitutions. *) +(* Creating the union of two compatible substitutions. *) +(* ------------------------------------------------------------------------- *) + +val union : subst -> subst -> subst (* raises Error *) + +(* ------------------------------------------------------------------------- *) +(* Substitutions can be inverted iff they are renaming substitutions. *) (* ------------------------------------------------------------------------- *) val invert : subst -> subst (* raises Error *) @@ -85,6 +89,22 @@ val freshVars : NameSet.set -> subst (* ------------------------------------------------------------------------- *) +(* Free variables. *) +(* ------------------------------------------------------------------------- *) + +val redexes : subst -> NameSet.set + +val residueFreeVars : subst -> NameSet.set + +val freeVars : subst -> NameSet.set + +(* ------------------------------------------------------------------------- *) +(* Functions. *) +(* ------------------------------------------------------------------------- *) + +val functions : subst -> NameAritySet.set + +(* ------------------------------------------------------------------------- *) (* Matching for first order logic terms. *) (* ------------------------------------------------------------------------- *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Subst.sml --- a/src/Tools/Metis/src/Subst.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Subst.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC SUBSTITUTIONS *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Subst :> Subst = @@ -30,16 +30,6 @@ fun singleton v_tm = insert empty v_tm; -local - fun compatible (tm1,tm2) = - if tm1 = tm2 then SOME tm1 else raise Error "Subst.union: incompatible"; -in - fun union (s1 as Subst m1) (s2 as Subst m2) = - if NameMap.null m1 then s2 - else if NameMap.null m2 then s1 - else Subst (NameMap.union compatible m1 m2); -end; - fun toList (Subst m) = NameMap.toList m; fun fromList l = Subst (NameMap.fromList l); @@ -48,12 +38,12 @@ fun foldr f b (Subst m) = NameMap.foldr f b m; -fun pp ppstrm sub = - Parser.ppBracket "<[" "]>" - (Parser.ppSequence "," (Parser.ppBinop " |->" Parser.ppString Term.pp)) - ppstrm (toList sub); +fun pp sub = + Print.ppBracket "<[" "]>" + (Print.ppOpList "," (Print.ppOp2 " |->" Name.pp Term.pp)) + (toList sub); -val toString = Parser.toString pp; +val toString = Print.toString pp; (* ------------------------------------------------------------------------- *) (* Normalizing removes identity substitutions. *) @@ -78,13 +68,13 @@ let fun tmSub (tm as Term.Var v) = (case peek sub v of - SOME tm' => if Sharing.pointerEqual (tm,tm') then tm else tm' + SOME tm' => if Portable.pointerEqual (tm,tm') then tm else tm' | NONE => tm) | tmSub (tm as Term.Fn (f,args)) = let val args' = Sharing.map tmSub args in - if Sharing.pointerEqual (args,args') then tm + if Portable.pointerEqual (args,args') then tm else Term.Fn (f,args') end in @@ -127,7 +117,22 @@ end; (* ------------------------------------------------------------------------- *) -(* Substitutions can be inverted iff they are renaming substitutions. *) +(* Creating the union of two compatible substitutions. *) +(* ------------------------------------------------------------------------- *) + +local + fun compatible ((_,tm1),(_,tm2)) = + if Term.equal tm1 tm2 then SOME tm1 + else raise Error "Subst.union: incompatible"; +in + fun union (s1 as Subst m1) (s2 as Subst m2) = + if NameMap.null m1 then s2 + else if NameMap.null m2 then s1 + else Subst (NameMap.union compatible m1 m2); +end; + +(* ------------------------------------------------------------------------- *) +(* Substitutions can be inverted iff they are renaming substitutions. *) (* ------------------------------------------------------------------------- *) local @@ -153,6 +158,42 @@ end; (* ------------------------------------------------------------------------- *) +(* Free variables. *) +(* ------------------------------------------------------------------------- *) + +val redexes = + let + fun add (v,_,s) = NameSet.add s v + in + foldl add NameSet.empty + end; + +val residueFreeVars = + let + fun add (_,t,s) = NameSet.union s (Term.freeVars t) + in + foldl add NameSet.empty + end; + +val freeVars = + let + fun add (v,t,s) = NameSet.union (NameSet.add s v) (Term.freeVars t) + in + foldl add NameSet.empty + end; + +(* ------------------------------------------------------------------------- *) +(* Functions. *) +(* ------------------------------------------------------------------------- *) + +val functions = + let + fun add (_,t,s) = NameAritySet.union s (Term.functions t) + in + foldl add NameAritySet.empty + end; + +(* ------------------------------------------------------------------------- *) (* Matching for first order logic terms. *) (* ------------------------------------------------------------------------- *) @@ -164,13 +205,13 @@ case peek sub v of NONE => insert sub (v,tm) | SOME tm' => - if tm = tm' then sub + if Term.equal tm tm' then sub else raise Error "Subst.match: incompatible matches" in matchList sub rest end | matchList sub ((Term.Fn (f1,args1), Term.Fn (f2,args2)) :: rest) = - if f1 = f2 andalso length args1 = length args2 then + if Name.equal f1 f2 andalso length args1 = length args2 then matchList sub (zip args1 args2 @ rest) else raise Error "Subst.match: different structure" | matchList _ _ = raise Error "Subst.match: functions can't match vars"; @@ -197,7 +238,7 @@ | SOME tm' => solve' sub tm' tm rest) | solve' sub tm1 (tm2 as Term.Var _) rest = solve' sub tm2 tm1 rest | solve' sub (Term.Fn (f1,args1)) (Term.Fn (f2,args2)) rest = - if f1 = f2 andalso length args1 = length args2 then + if Name.equal f1 f2 andalso length args1 = length args2 then solve sub (zip args1 args2 @ rest) else raise Error "Subst.unify: different structure"; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Subsume.sig --- a/src/Tools/Metis/src/Subsume.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Subsume.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Subsume = @@ -20,7 +20,7 @@ val filter : ('a -> bool) -> 'a subsume -> 'a subsume -val pp : 'a subsume Parser.pp +val pp : 'a subsume Print.pp val toString : 'a subsume -> string diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Subsume.sml --- a/src/Tools/Metis/src/Subsume.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Subsume.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES *) -(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Subsume :> Subsume = @@ -159,7 +159,7 @@ fun toString subsume = "Subsume{" ^ Int.toString (size subsume) ^ "}"; -fun pp p = Parser.ppMap toString Parser.ppString p; +fun pp subsume = Print.ppMap toString Print.ppString subsume; (* ------------------------------------------------------------------------- *) (* Subsumption checking. *) @@ -274,19 +274,19 @@ genSubsumes pred subsume (SOME (LiteralSet.size cl)) cl; end; -(*TRACE4 +(*MetisTrace4 val subsumes = fn pred => fn subsume => fn cl => let val ppCl = LiteralSet.pp val ppSub = Subst.pp - val () = Parser.ppTrace ppCl "Subsume.subsumes: cl" cl + val () = Print.trace ppCl "Subsume.subsumes: cl" cl val result = subsumes pred subsume cl val () = case result of NONE => trace "Subsume.subsumes: not subsumed\n" | SOME (cl,sub,_) => - (Parser.ppTrace ppCl "Subsume.subsumes: subsuming cl" cl; - Parser.ppTrace ppSub "Subsume.subsumes: subsuming sub" sub) + (Print.trace ppCl "Subsume.subsumes: subsuming cl" cl; + Print.trace ppSub "Subsume.subsumes: subsuming sub" sub) in result end; @@ -295,14 +295,14 @@ let val ppCl = LiteralSet.pp val ppSub = Subst.pp - val () = Parser.ppTrace ppCl "Subsume.strictlySubsumes: cl" cl + val () = Print.trace ppCl "Subsume.strictlySubsumes: cl" cl val result = strictlySubsumes pred subsume cl val () = case result of NONE => trace "Subsume.subsumes: not subsumed\n" | SOME (cl,sub,_) => - (Parser.ppTrace ppCl "Subsume.subsumes: subsuming cl" cl; - Parser.ppTrace ppSub "Subsume.subsumes: subsuming sub" sub) + (Print.trace ppCl "Subsume.subsumes: subsuming cl" cl; + Print.trace ppSub "Subsume.subsumes: subsuming sub" sub) in result end; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Term.sig --- a/src/Tools/Metis/src/Term.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Term.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC TERMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Term = @@ -80,6 +80,8 @@ val compare : term * term -> order +val equal : term -> term -> bool + (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) @@ -94,7 +96,7 @@ val find : (term -> bool) -> term -> path option -val ppPath : path Parser.pp +val ppPath : path Print.pp val pathToString : path -> string @@ -106,6 +108,8 @@ val freeVars : term -> NameSet.set +val freeVarsList : term list -> NameSet.set + (* ------------------------------------------------------------------------- *) (* Fresh variables. *) (* ------------------------------------------------------------------------- *) @@ -122,6 +126,10 @@ (* Special support for terms with type annotations. *) (* ------------------------------------------------------------------------- *) +val hasTypeFunctionName : functionName + +val hasTypeFunction : function + val isTypedVar : term -> bool val typedSymbols : term -> int @@ -132,15 +140,17 @@ (* Special support for terms with an explicit function application operator. *) (* ------------------------------------------------------------------------- *) -val mkComb : term * term -> term +val appName : Name.name -val destComb : term -> term * term +val mkApp : term * term -> term + +val destApp : term -> term * term -val isComb : term -> bool +val isApp : term -> bool -val listMkComb : term * term list -> term +val listMkApp : term * term list -> term -val stripComb : term -> term * term list +val stripApp : term -> term * term list (* ------------------------------------------------------------------------- *) (* Parsing and pretty printing. *) @@ -148,23 +158,23 @@ (* Infix symbols *) -val infixes : Parser.infixities ref +val infixes : Print.infixes ref (* The negation symbol *) -val negation : Name.name ref +val negation : string ref (* Binder symbols *) -val binders : Name.name list ref +val binders : string list ref (* Bracket symbols *) -val brackets : (Name.name * Name.name) list ref +val brackets : (string * string) list ref (* Pretty printing *) -val pp : term Parser.pp +val pp : term Print.pp val toString : term -> string @@ -172,6 +182,6 @@ val fromString : string -> term -val parse : term Parser.quotation -> term +val parse : term Parse.quotation -> term end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Term.sml --- a/src/Tools/Metis/src/Term.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Term.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* FIRST ORDER LOGIC TERMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Term :> Term = @@ -9,24 +9,6 @@ open Useful; (* ------------------------------------------------------------------------- *) -(* Helper functions. *) -(* ------------------------------------------------------------------------- *) - -fun stripSuffix pred s = - let - fun f 0 = "" - | f n = - let - val n' = n - 1 - in - if pred (String.sub (s,n')) then f n' - else String.substring (s,0,n) - end - in - f (size s) - end; - -(* ------------------------------------------------------------------------- *) (* A type of first order logic terms. *) (* ------------------------------------------------------------------------- *) @@ -53,7 +35,7 @@ val isVar = can destVar; -fun equalVar v (Var v') = v = v' +fun equalVar v (Var v') = Name.equal v v' | equalVar _ _ = false; (* Functions *) @@ -102,7 +84,7 @@ fun mkBinop f (a,b) = Fn (f,[a,b]); fun destBinop f (Fn (x,[a,b])) = - if x = f then (a,b) else raise Error "Term.destBinop: wrong binop" + if Name.equal x f then (a,b) else raise Error "Term.destBinop: wrong binop" | destBinop _ _ = raise Error "Term.destBinop: not a binop"; fun isBinop f = can (destBinop f); @@ -129,27 +111,37 @@ local fun cmp [] [] = EQUAL - | cmp (Var _ :: _) (Fn _ :: _) = LESS - | cmp (Fn _ :: _) (Var _ :: _) = GREATER - | cmp (Var v1 :: tms1) (Var v2 :: tms2) = - (case Name.compare (v1,v2) of - LESS => LESS - | EQUAL => cmp tms1 tms2 - | GREATER => GREATER) - | cmp (Fn (f1,a1) :: tms1) (Fn (f2,a2) :: tms2) = - (case Name.compare (f1,f2) of - LESS => LESS - | EQUAL => - (case Int.compare (length a1, length a2) of - LESS => LESS - | EQUAL => cmp (a1 @ tms1) (a2 @ tms2) - | GREATER => GREATER) - | GREATER => GREATER) + | cmp (tm1 :: tms1) (tm2 :: tms2) = + let + val tm1_tm2 = (tm1,tm2) + in + if Portable.pointerEqual tm1_tm2 then cmp tms1 tms2 + else + case tm1_tm2 of + (Var v1, Var v2) => + (case Name.compare (v1,v2) of + LESS => LESS + | EQUAL => cmp tms1 tms2 + | GREATER => GREATER) + | (Var _, Fn _) => LESS + | (Fn _, Var _) => GREATER + | (Fn (f1,a1), Fn (f2,a2)) => + (case Name.compare (f1,f2) of + LESS => LESS + | EQUAL => + (case Int.compare (length a1, length a2) of + LESS => LESS + | EQUAL => cmp (a1 @ tms1) (a2 @ tms2) + | GREATER => GREATER) + | GREATER => GREATER) + end | cmp _ _ = raise Bug "Term.compare"; in fun compare (tm1,tm2) = cmp [tm1] [tm2]; end; +fun equal tm1 tm2 = compare (tm1,tm2) = EQUAL; + (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) @@ -178,7 +170,7 @@ fun subterms tm = subtms [([],tm)] []; end; -fun replace tm ([],res) = if res = tm then tm else res +fun replace tm ([],res) = if equal res tm then tm else res | replace tm (h :: t, res) = case tm of Var _ => raise Error "Term.replace: Var" @@ -189,7 +181,7 @@ val arg = List.nth (tms,h) val arg' = replace arg (t,res) in - if Sharing.pointerEqual (arg',arg) then tm + if Portable.pointerEqual (arg',arg) then tm else Fn (func, updateNth (h,arg') tms) end; @@ -211,9 +203,9 @@ fn tm => search [([],tm)] end; -val ppPath = Parser.ppList Parser.ppInt; +val ppPath = Print.ppList Print.ppInt; -val pathToString = Parser.toString ppPath; +val pathToString = Print.toString ppPath; (* ------------------------------------------------------------------------- *) (* Free variables. *) @@ -221,7 +213,7 @@ local fun free _ [] = false - | free v (Var w :: tms) = v = w orelse free v tms + | free v (Var w :: tms) = Name.equal v w orelse free v tms | free v (Fn (_,args) :: tms) = free v (args @ tms); in fun freeIn v tm = free v [tm]; @@ -232,77 +224,100 @@ | free vs (Var v :: tms) = free (NameSet.add vs v) tms | free vs (Fn (_,args) :: tms) = free vs (args @ tms); in - fun freeVars tm = free NameSet.empty [tm]; + val freeVarsList = free NameSet.empty; + + fun freeVars tm = freeVarsList [tm]; end; (* ------------------------------------------------------------------------- *) (* Fresh variables. *) (* ------------------------------------------------------------------------- *) -local - val prefix = "_"; +fun newVar () = Var (Name.newName ()); - fun numVar i = Var (mkPrefix prefix (Int.toString i)); -in - fun newVar () = numVar (newInt ()); - - fun newVars n = map numVar (newInts n); -end; +fun newVars n = map Var (Name.newNames n); -fun variantPrime avoid = - let - fun f v = if NameSet.member v avoid then f (v ^ "'") else v - in - f - end; +local + fun avoidAcceptable avoid n = not (NameSet.member n avoid); +in + fun variantPrime avoid = Name.variantPrime (avoidAcceptable avoid); -fun variantNum avoid v = - if not (NameSet.member v avoid) then v - else - let - val v = stripSuffix Char.isDigit v - - fun f n = - let - val v_n = v ^ Int.toString n - in - if NameSet.member v_n avoid then f (n + 1) else v_n - end - in - f 0 - end; + fun variantNum avoid = Name.variantNum (avoidAcceptable avoid); +end; (* ------------------------------------------------------------------------- *) (* Special support for terms with type annotations. *) (* ------------------------------------------------------------------------- *) -fun isTypedVar (Var _) = true - | isTypedVar (Fn (":", [Var _, _])) = true - | isTypedVar (Fn _) = false; +val hasTypeFunctionName = Name.fromString ":"; + +val hasTypeFunction = (hasTypeFunctionName,2); + +fun destFnHasType ((f,a) : functionName * term list) = + if not (Name.equal f hasTypeFunctionName) then + raise Error "Term.destFnHasType" + else + case a of + [tm,ty] => (tm,ty) + | _ => raise Error "Term.destFnHasType"; + +val isFnHasType = can destFnHasType; + +fun isTypedVar tm = + case tm of + Var _ => true + | Fn func => + case total destFnHasType func of + SOME (Var _, _) => true + | _ => false; local fun sz n [] = n - | sz n (Var _ :: tms) = sz (n + 1) tms - | sz n (Fn (":",[tm,_]) :: tms) = sz n (tm :: tms) - | sz n (Fn (_,args) :: tms) = sz (n + 1) (args @ tms); + | sz n (tm :: tms) = + case tm of + Var _ => sz (n + 1) tms + | Fn func => + case total destFnHasType func of + SOME (tm,_) => sz n (tm :: tms) + | NONE => + let + val (_,a) = func + in + sz (n + 1) (a @ tms) + end; in fun typedSymbols tm = sz 0 [tm]; end; local fun subtms [] acc = acc - | subtms ((_, Var _) :: rest) acc = subtms rest acc - | subtms ((_, Fn (":", [Var _, _])) :: rest) acc = subtms rest acc - | subtms ((path, tm as Fn func) :: rest) acc = - let - fun f (n,arg) = (n :: path, arg) + | subtms ((path,tm) :: rest) acc = + case tm of + Var _ => subtms rest acc + | Fn func => + case total destFnHasType func of + SOME (t,_) => + (case t of + Var _ => subtms rest acc + | Fn _ => + let + val acc = (rev path, tm) :: acc + val rest = (0 :: path, t) :: rest + in + subtms rest acc + end) + | NONE => + let + fun f (n,arg) = (n :: path, arg) - val acc = (rev path, tm) :: acc - in - case func of - (":",[arg,_]) => subtms ((0 :: path, arg) :: rest) acc - | (_,args) => subtms (map f (enumerate args) @ rest) acc - end; + val (_,args) = func + + val acc = (rev path, tm) :: acc + + val rest = map f (enumerate args) @ rest + in + subtms rest acc + end; in fun nonVarTypedSubterms tm = subtms [([],tm)] []; end; @@ -311,20 +326,37 @@ (* Special support for terms with an explicit function application operator. *) (* ------------------------------------------------------------------------- *) -fun mkComb (f,a) = Fn (".",[f,a]); +val appName = Name.fromString "."; + +fun mkFnApp (fTm,aTm) = (appName, [fTm,aTm]); + +fun mkApp f_a = Fn (mkFnApp f_a); -fun destComb (Fn (".",[f,a])) = (f,a) - | destComb _ = raise Error "destComb"; +fun destFnApp ((f,a) : Name.name * term list) = + if not (Name.equal f appName) then raise Error "Term.destFnApp" + else + case a of + [fTm,aTm] => (fTm,aTm) + | _ => raise Error "Term.destFnApp"; + +val isFnApp = can destFnApp; -val isComb = can destComb; +fun destApp tm = + case tm of + Var _ => raise Error "Term.destApp" + | Fn func => destFnApp func; -fun listMkComb (f,l) = foldl mkComb f l; +val isApp = can destApp; + +fun listMkApp (f,l) = foldl mkApp f l; local - fun strip tms (Fn (".",[f,a])) = strip (a :: tms) f - | strip tms tm = (tm,tms); + fun strip tms tm = + case total destApp tm of + SOME (f,a) => strip (a :: tms) f + | NONE => (tm,tms); in - fun stripComb tm = strip [] tm; + fun stripApp tm = strip [] tm; end; (* ------------------------------------------------------------------------- *) @@ -333,185 +365,204 @@ (* Operators parsed and printed infix *) -val infixes : Parser.infixities ref = ref - [(* ML symbols *) - {token = " / ", precedence = 7, leftAssoc = true}, - {token = " div ", precedence = 7, leftAssoc = true}, - {token = " mod ", precedence = 7, leftAssoc = true}, - {token = " * ", precedence = 7, leftAssoc = true}, - {token = " + ", precedence = 6, leftAssoc = true}, - {token = " - ", precedence = 6, leftAssoc = true}, - {token = " ^ ", precedence = 6, leftAssoc = true}, - {token = " @ ", precedence = 5, leftAssoc = false}, - {token = " :: ", precedence = 5, leftAssoc = false}, - {token = " = ", precedence = 4, leftAssoc = true}, - {token = " <> ", precedence = 4, leftAssoc = true}, - {token = " <= ", precedence = 4, leftAssoc = true}, - {token = " < ", precedence = 4, leftAssoc = true}, - {token = " >= ", precedence = 4, leftAssoc = true}, - {token = " > ", precedence = 4, leftAssoc = true}, - {token = " o ", precedence = 3, leftAssoc = true}, - {token = " -> ", precedence = 2, leftAssoc = false}, (* inferred prec *) - {token = " : ", precedence = 1, leftAssoc = false}, (* inferred prec *) - {token = ", ", precedence = 0, leftAssoc = false}, (* inferred prec *) +val infixes = + (ref o Print.Infixes) + [(* ML symbols *) + {token = " / ", precedence = 7, leftAssoc = true}, + {token = " div ", precedence = 7, leftAssoc = true}, + {token = " mod ", precedence = 7, leftAssoc = true}, + {token = " * ", precedence = 7, leftAssoc = true}, + {token = " + ", precedence = 6, leftAssoc = true}, + {token = " - ", precedence = 6, leftAssoc = true}, + {token = " ^ ", precedence = 6, leftAssoc = true}, + {token = " @ ", precedence = 5, leftAssoc = false}, + {token = " :: ", precedence = 5, leftAssoc = false}, + {token = " = ", precedence = 4, leftAssoc = true}, + {token = " <> ", precedence = 4, leftAssoc = true}, + {token = " <= ", precedence = 4, leftAssoc = true}, + {token = " < ", precedence = 4, leftAssoc = true}, + {token = " >= ", precedence = 4, leftAssoc = true}, + {token = " > ", precedence = 4, leftAssoc = true}, + {token = " o ", precedence = 3, leftAssoc = true}, + {token = " -> ", precedence = 2, leftAssoc = false}, (* inferred prec *) + {token = " : ", precedence = 1, leftAssoc = false}, (* inferred prec *) + {token = ", ", precedence = 0, leftAssoc = false}, (* inferred prec *) - (* Logical connectives *) - {token = " /\\ ", precedence = ~1, leftAssoc = false}, - {token = " \\/ ", precedence = ~2, leftAssoc = false}, - {token = " ==> ", precedence = ~3, leftAssoc = false}, - {token = " <=> ", precedence = ~4, leftAssoc = false}, + (* Logical connectives *) + {token = " /\\ ", precedence = ~1, leftAssoc = false}, + {token = " \\/ ", precedence = ~2, leftAssoc = false}, + {token = " ==> ", precedence = ~3, leftAssoc = false}, + {token = " <=> ", precedence = ~4, leftAssoc = false}, - (* Other symbols *) - {token = " . ", precedence = 9, leftAssoc = true}, (* function app *) - {token = " ** ", precedence = 8, leftAssoc = true}, - {token = " ++ ", precedence = 6, leftAssoc = true}, - {token = " -- ", precedence = 6, leftAssoc = true}, - {token = " == ", precedence = 4, leftAssoc = true}]; + (* Other symbols *) + {token = " . ", precedence = 9, leftAssoc = true}, (* function app *) + {token = " ** ", precedence = 8, leftAssoc = true}, + {token = " ++ ", precedence = 6, leftAssoc = true}, + {token = " -- ", precedence = 6, leftAssoc = true}, + {token = " == ", precedence = 4, leftAssoc = true}]; (* The negation symbol *) -val negation : Name.name ref = ref "~"; +val negation : string ref = ref "~"; (* Binder symbols *) -val binders : Name.name list ref = ref ["\\","!","?","?!"]; +val binders : string list ref = ref ["\\","!","?","?!"]; (* Bracket symbols *) -val brackets : (Name.name * Name.name) list ref = ref [("[","]"),("{","}")]; +val brackets : (string * string) list ref = ref [("[","]"),("{","}")]; (* Pretty printing *) -local - open Parser; -in - fun pp inputPpstrm inputTerm = - let - val quants = !binders - and iOps = !infixes - and neg = !negation - and bracks = !brackets +fun pp inputTerm = + let + val quants = !binders + and iOps = !infixes + and neg = !negation + and bracks = !brackets + + val bracks = map (fn (b1,b2) => (b1 ^ b2, b1, b2)) bracks + + val bTokens = map #2 bracks @ map #3 bracks + + val iTokens = Print.tokensInfixes iOps - val bracks = map (fn (b1,b2) => (b1 ^ b2, b1, b2)) bracks + fun destI tm = + case tm of + Fn (f,[a,b]) => + let + val f = Name.toString f + in + if StringSet.member f iTokens then SOME (f,a,b) else NONE + end + | _ => NONE + + val iPrinter = Print.ppInfixes iOps destI + + val specialTokens = + StringSet.addList iTokens (neg :: quants @ ["$","(",")"] @ bTokens) + + fun vName bv s = StringSet.member s bv - val bTokens = map #2 bracks @ map #3 bracks + fun checkVarName bv n = + let + val s = Name.toString n + in + if vName bv s then s else "$" ^ s + end - val iTokens = infixTokens iOps + fun varName bv = Print.ppMap (checkVarName bv) Print.ppString - fun destI (Fn (f,[a,b])) = - if mem f iTokens then SOME (f,a,b) else NONE - | destI _ = NONE + fun checkFunctionName bv n = + let + val s = Name.toString n + in + if StringSet.member s specialTokens orelse vName bv s then + "(" ^ s ^ ")" + else + s + end - val iPrinter = ppInfixes iOps destI + fun functionName bv = Print.ppMap (checkFunctionName bv) Print.ppString + + fun isI tm = Option.isSome (destI tm) - val specialTokens = neg :: quants @ ["$","(",")"] @ bTokens @ iTokens - - fun vName bv s = NameSet.member s bv + fun stripNeg tm = + case tm of + Fn (f,[a]) => + if Name.toString f <> neg then (0,tm) + else let val (n,tm) = stripNeg a in (n + 1, tm) end + | _ => (0,tm) - fun checkVarName bv s = if vName bv s then s else "$" ^ s - - fun varName bv = ppMap (checkVarName bv) ppString + val destQuant = + let + fun dest q (Fn (q', [Var v, body])) = + if Name.toString q' <> q then NONE + else + (case dest q body of + NONE => SOME (q,v,[],body) + | SOME (_,v',vs,body) => SOME (q, v, v' :: vs, body)) + | dest _ _ = NONE + in + fn tm => Useful.first (fn q => dest q tm) quants + end - fun checkFunctionName bv s = - if mem s specialTokens orelse vName bv s then "(" ^ s ^ ")" else s + fun isQuant tm = Option.isSome (destQuant tm) - fun functionName bv = ppMap (checkFunctionName bv) ppString + fun destBrack (Fn (b,[tm])) = + let + val s = Name.toString b + in + case List.find (fn (n,_,_) => n = s) bracks of + NONE => NONE + | SOME (_,b1,b2) => SOME (b1,tm,b2) + end + | destBrack _ = NONE - fun isI tm = Option.isSome (destI tm) + fun isBrack tm = Option.isSome (destBrack tm) - fun stripNeg (tm as Fn (f,[a])) = - if f <> neg then (0,tm) - else let val (n,tm) = stripNeg a in (n + 1, tm) end - | stripNeg tm = (0,tm) + fun functionArgument bv tm = + Print.sequence + (Print.addBreak 1) + (if isBrack tm then customBracket bv tm + else if isVar tm orelse isConst tm then basic bv tm + else bracket bv tm) - val destQuant = + and basic bv (Var v) = varName bv v + | basic bv (Fn (f,args)) = + Print.blockProgram Print.Inconsistent 2 + (functionName bv f :: map (functionArgument bv) args) + + and customBracket bv tm = + case destBrack tm of + SOME (b1,tm,b2) => Print.ppBracket b1 b2 (term bv) tm + | NONE => basic bv tm + + and innerQuant bv tm = + case destQuant tm of + NONE => term bv tm + | SOME (q,v,vs,tm) => let - fun dest q (Fn (q', [Var v, body])) = - if q <> q' then NONE - else - (case dest q body of - NONE => SOME (q,v,[],body) - | SOME (_,v',vs,body) => SOME (q, v, v' :: vs, body)) - | dest _ _ = NONE + val bv = StringSet.addList bv (map Name.toString (v :: vs)) in - fn tm => Useful.first (fn q => dest q tm) quants + Print.program + [Print.addString q, + varName bv v, + Print.program + (map (Print.sequence (Print.addBreak 1) o varName bv) vs), + Print.addString ".", + Print.addBreak 1, + innerQuant bv tm] end - fun isQuant tm = Option.isSome (destQuant tm) - - fun destBrack (Fn (b,[tm])) = - (case List.find (fn (n,_,_) => n = b) bracks of - NONE => NONE - | SOME (_,b1,b2) => SOME (b1,tm,b2)) - | destBrack _ = NONE - - fun isBrack tm = Option.isSome (destBrack tm) - - fun functionArgument bv ppstrm tm = - (addBreak ppstrm (1,0); - if isBrack tm then customBracket bv ppstrm tm - else if isVar tm orelse isConst tm then basic bv ppstrm tm - else bracket bv ppstrm tm) - - and basic bv ppstrm (Var v) = varName bv ppstrm v - | basic bv ppstrm (Fn (f,args)) = - (beginBlock ppstrm Inconsistent 2; - functionName bv ppstrm f; - app (functionArgument bv ppstrm) args; - endBlock ppstrm) - - and customBracket bv ppstrm tm = - case destBrack tm of - SOME (b1,tm,b2) => ppBracket b1 b2 (term bv) ppstrm tm - | NONE => basic bv ppstrm tm + and quantifier bv tm = + if not (isQuant tm) then customBracket bv tm + else Print.block Print.Inconsistent 2 (innerQuant bv tm) - and innerQuant bv ppstrm tm = - case destQuant tm of - NONE => term bv ppstrm tm - | SOME (q,v,vs,tm) => - let - val bv = NameSet.addList (NameSet.add bv v) vs - in - addString ppstrm q; - varName bv ppstrm v; - app (fn v => (addBreak ppstrm (1,0); varName bv ppstrm v)) vs; - addString ppstrm "."; - addBreak ppstrm (1,0); - innerQuant bv ppstrm tm - end + and molecule bv (tm,r) = + let + val (n,tm) = stripNeg tm + in + Print.blockProgram Print.Inconsistent n + [Print.duplicate n (Print.addString neg), + if isI tm orelse (r andalso isQuant tm) then bracket bv tm + else quantifier bv tm] + end - and quantifier bv ppstrm tm = - if not (isQuant tm) then customBracket bv ppstrm tm - else - (beginBlock ppstrm Inconsistent 2; - innerQuant bv ppstrm tm; - endBlock ppstrm) + and term bv tm = iPrinter (molecule bv) (tm,false) - and molecule bv ppstrm (tm,r) = - let - val (n,tm) = stripNeg tm - in - beginBlock ppstrm Inconsistent n; - funpow n (fn () => addString ppstrm neg) (); - if isI tm orelse (r andalso isQuant tm) then bracket bv ppstrm tm - else quantifier bv ppstrm tm; - endBlock ppstrm - end + and bracket bv tm = Print.ppBracket "(" ")" (term bv) tm + in + term StringSet.empty + end inputTerm; - and term bv ppstrm tm = iPrinter (molecule bv) ppstrm (tm,false) - - and bracket bv ppstrm tm = ppBracket "(" ")" (term bv) ppstrm tm - in - term NameSet.empty - end inputPpstrm inputTerm; -end; - -fun toString tm = Parser.toString pp tm; +val toString = Print.toString pp; (* Parsing *) local - open Parser; + open Parse; infixr 9 >>++ infixr 8 ++ @@ -531,7 +582,7 @@ val symbolToken = let fun isNeg c = str c = !negation - + val symbolChars = explode "<>=-*+/\\?@|!$%&#^:;~" fun isSymbol c = mem c symbolChars @@ -572,42 +623,50 @@ fun possibleVarName "" = false | possibleVarName s = isAlphaNum (String.sub (s,0)) - fun vName bv s = NameSet.member s bv + fun vName bv s = StringSet.member s bv - val iTokens = infixTokens iOps + val iTokens = Print.tokensInfixes iOps - val iParser = parseInfixes iOps (fn (f,a,b) => Fn (f,[a,b])) + val iParser = + parseInfixes iOps (fn (f,a,b) => Fn (Name.fromString f, [a,b])) - val specialTokens = neg :: quants @ ["$"] @ bTokens @ iTokens + val specialTokens = + StringSet.addList iTokens (neg :: quants @ ["$"] @ bTokens) fun varName bv = - Parser.some (vName bv) || - (exact "$" ++ some possibleVarName) >> (fn (_,s) => s) + some (vName bv) || + (some (Useful.equal "$") ++ some possibleVarName) >> snd - fun fName bv s = not (mem s specialTokens) andalso not (vName bv s) + fun fName bv s = + not (StringSet.member s specialTokens) andalso not (vName bv s) fun functionName bv = - Parser.some (fName bv) || - (exact "(" ++ any ++ exact ")") >> (fn (_,(s,_)) => s) + some (fName bv) || + (some (Useful.equal "(") ++ any ++ some (Useful.equal ")")) >> + (fn (_,(s,_)) => s) fun basic bv tokens = let - val var = varName bv >> Var + val var = varName bv >> (Var o Name.fromString) - val const = functionName bv >> (fn f => Fn (f,[])) + val const = + functionName bv >> (fn f => Fn (Name.fromString f, [])) fun bracket (ab,a,b) = - (exact a ++ term bv ++ exact b) >> - (fn (_,(tm,_)) => if ab = "()" then tm else Fn (ab,[tm])) + (some (Useful.equal a) ++ term bv ++ some (Useful.equal b)) >> + (fn (_,(tm,_)) => + if ab = "()" then tm else Fn (Name.fromString ab, [tm])) fun quantifier q = let - fun bind (v,t) = Fn (q, [Var v, t]) + fun bind (v,t) = + Fn (Name.fromString q, [Var (Name.fromString v), t]) in - (exact q ++ atLeastOne (some possibleVarName) ++ - exact ".") >>++ + (some (Useful.equal q) ++ + atLeastOne (some possibleVarName) ++ + some (Useful.equal ".")) >>++ (fn (_,(vs,_)) => - term (NameSet.addList bv vs) >> + term (StringSet.addList bv vs) >> (fn body => foldr bind body vs)) end in @@ -619,18 +678,20 @@ and molecule bv tokens = let - val negations = many (exact neg) >> length + val negations = many (some (Useful.equal neg)) >> length val function = - (functionName bv ++ many (basic bv)) >> Fn || basic bv + (functionName bv ++ many (basic bv)) >> + (fn (f,args) => Fn (Name.fromString f, args)) || + basic bv in (negations ++ function) >> - (fn (n,tm) => funpow n (fn t => Fn (neg,[t])) tm) + (fn (n,tm) => funpow n (fn t => Fn (Name.fromString neg, [t])) tm) end tokens and term bv tokens = iParser (molecule bv) tokens in - term NameSet.empty + term StringSet.empty end inputStream; in fun fromString input = @@ -643,15 +704,14 @@ in case Stream.toList terms of [tm] => tm - | _ => raise Error "Syntax.stringToTerm" + | _ => raise Error "Term.fromString" end; end; local - val antiquotedTermToString = - Parser.toString (Parser.ppBracket "(" ")" pp); + val antiquotedTermToString = Print.toString (Print.ppBracket "(" ")" pp); in - val parse = Parser.parseQuotation antiquotedTermToString fromString; + val parse = Parse.parseQuotation antiquotedTermToString fromString; end; end @@ -659,6 +719,6 @@ structure TermOrdered = struct type t = Term.term val compare = Term.compare end -structure TermSet = ElementSet (TermOrdered); +structure TermMap = KeyMap (TermOrdered); -structure TermMap = KeyMap (TermOrdered); +structure TermSet = ElementSet (TermMap); diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/TermNet.sig --- a/src/Tools/Metis/src/TermNet.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/TermNet.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature TermNet = @@ -32,7 +32,7 @@ val toString : 'a termNet -> string -val pp : 'a Parser.pp -> 'a termNet Parser.pp +val pp : 'a Print.pp -> 'a termNet Print.pp (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/TermNet.sml --- a/src/Tools/Metis/src/TermNet.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/TermNet.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure TermNet :> TermNet = @@ -9,29 +9,65 @@ open Useful; (* ------------------------------------------------------------------------- *) -(* Quotient terms *) +(* Anonymous variables. *) +(* ------------------------------------------------------------------------- *) + +val anonymousName = Name.fromString "_"; +val anonymousVar = Term.Var anonymousName; + +(* ------------------------------------------------------------------------- *) +(* Quotient terms. *) (* ------------------------------------------------------------------------- *) -datatype qterm = VAR | FN of NameArity.nameArity * qterm list; +datatype qterm = + Var + | Fn of NameArity.nameArity * qterm list; + +local + fun cmp [] = EQUAL + | cmp (q1_q2 :: qs) = + if Portable.pointerEqual q1_q2 then cmp qs + else + case q1_q2 of + (Var,Var) => EQUAL + | (Var, Fn _) => LESS + | (Fn _, Var) => GREATER + | (Fn f1, Fn f2) => fnCmp f1 f2 qs -fun termToQterm (Term.Var _) = VAR - | termToQterm (Term.Fn (f,l)) = FN ((f, length l), map termToQterm l); + and fnCmp (n1,q1) (n2,q2) qs = + case NameArity.compare (n1,n2) of + LESS => LESS + | EQUAL => cmp (zip q1 q2 @ qs) + | GREATER => GREATER; +in + fun compareQterm q1_q2 = cmp [q1_q2]; + + fun compareFnQterm (f1,f2) = fnCmp f1 f2 []; +end; + +fun equalQterm q1 q2 = compareQterm (q1,q2) = EQUAL; + +fun equalFnQterm f1 f2 = compareFnQterm (f1,f2) = EQUAL; + +fun termToQterm (Term.Var _) = Var + | termToQterm (Term.Fn (f,l)) = Fn ((f, length l), map termToQterm l); local fun qm [] = true - | qm ((VAR,_) :: rest) = qm rest - | qm ((FN _, VAR) :: _) = false - | qm ((FN (f,a), FN (g,b)) :: rest) = f = g andalso qm (zip a b @ rest); + | qm ((Var,_) :: rest) = qm rest + | qm ((Fn _, Var) :: _) = false + | qm ((Fn (f,a), Fn (g,b)) :: rest) = + NameArity.equal f g andalso qm (zip a b @ rest); in fun matchQtermQterm qtm qtm' = qm [(qtm,qtm')]; end; local fun qm [] = true - | qm ((VAR,_) :: rest) = qm rest - | qm ((FN _, Term.Var _) :: _) = false - | qm ((FN ((f,n),a), Term.Fn (g,b)) :: rest) = - f = g andalso n = length b andalso qm (zip a b @ rest); + | qm ((Var,_) :: rest) = qm rest + | qm ((Fn _, Term.Var _) :: _) = false + | qm ((Fn ((f,n),a), Term.Fn (g,b)) :: rest) = + Name.equal f g andalso n = length b andalso qm (zip a b @ rest); in fun matchQtermTerm qtm tm = qm [(qtm,tm)]; end; @@ -41,26 +77,27 @@ | qn qsub ((Term.Var v, qtm) :: rest) = (case NameMap.peek qsub v of NONE => qn (NameMap.insert qsub (v,qtm)) rest - | SOME qtm' => if qtm = qtm' then qn qsub rest else NONE) - | qn _ ((Term.Fn _, VAR) :: _) = NONE - | qn qsub ((Term.Fn (f,a), FN ((g,n),b)) :: rest) = - if f = g andalso length a = n then qn qsub (zip a b @ rest) else NONE; + | SOME qtm' => if equalQterm qtm qtm' then qn qsub rest else NONE) + | qn _ ((Term.Fn _, Var) :: _) = NONE + | qn qsub ((Term.Fn (f,a), Fn ((g,n),b)) :: rest) = + if Name.equal f g andalso length a = n then qn qsub (zip a b @ rest) + else NONE; in fun matchTermQterm qsub tm qtm = qn qsub [(tm,qtm)]; end; local - fun qv VAR x = x - | qv x VAR = x - | qv (FN (f,a)) (FN (g,b)) = + fun qv Var x = x + | qv x Var = x + | qv (Fn (f,a)) (Fn (g,b)) = let - val _ = f = g orelse raise Error "TermNet.qv" + val _ = NameArity.equal f g orelse raise Error "TermNet.qv" in - FN (f, zipwith qv a b) + Fn (f, zipWith qv a b) end; fun qu qsub [] = qsub - | qu qsub ((VAR, _) :: rest) = qu qsub rest + | qu qsub ((Var, _) :: rest) = qu qsub rest | qu qsub ((qtm, Term.Var v) :: rest) = let val qtm = @@ -68,8 +105,8 @@ in qu (NameMap.insert qsub (v,qtm)) rest end - | qu qsub ((FN ((f,n),a), Term.Fn (g,b)) :: rest) = - if f = g andalso n = length b then qu qsub (zip a b @ rest) + | qu qsub ((Fn ((f,n),a), Term.Fn (g,b)) :: rest) = + if Name.equal f g andalso n = length b then qu qsub (zip a b @ rest) else raise Error "TermNet.qu"; in fun unifyQtermQterm qtm qtm' = total (qv qtm) qtm'; @@ -78,10 +115,10 @@ end; local - fun qtermToTerm VAR = Term.Var "" - | qtermToTerm (FN ((f,_),l)) = Term.Fn (f, map qtermToTerm l); + fun qtermToTerm Var = anonymousVar + | qtermToTerm (Fn ((f,_),l)) = Term.Fn (f, map qtermToTerm l); in - val ppQterm = Parser.ppMap qtermToTerm Term.pp; + val ppQterm = Print.ppMap qtermToTerm Term.pp; end; (* ------------------------------------------------------------------------- *) @@ -91,22 +128,22 @@ type parameters = {fifo : bool}; datatype 'a net = - RESULT of 'a list - | SINGLE of qterm * 'a net - | MULTIPLE of 'a net option * 'a net NameArityMap.map; + Result of 'a list + | Single of qterm * 'a net + | Multiple of 'a net option * 'a net NameArityMap.map; -datatype 'a termNet = NET of parameters * int * (int * (int * 'a) net) option; +datatype 'a termNet = Net of parameters * int * (int * (int * 'a) net) option; (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) -fun new parm = NET (parm,0,NONE); +fun new parm = Net (parm,0,NONE); local - fun computeSize (RESULT l) = length l - | computeSize (SINGLE (_,n)) = computeSize n - | computeSize (MULTIPLE (vs,fs)) = + fun computeSize (Result l) = length l + | computeSize (Single (_,n)) = computeSize n + | computeSize (Multiple (vs,fs)) = NameArityMap.foldl (fn (_,n,acc) => acc + computeSize n) (case vs of SOME n => computeSize n | NONE => 0) @@ -116,38 +153,38 @@ | netSize (SOME n) = SOME (computeSize n, n); end; -fun size (NET (_,_,NONE)) = 0 - | size (NET (_, _, SOME (i,_))) = i; +fun size (Net (_,_,NONE)) = 0 + | size (Net (_, _, SOME (i,_))) = i; fun null net = size net = 0; -fun singles qtms a = foldr SINGLE a qtms; +fun singles qtms a = foldr Single a qtms; local fun pre NONE = (0,NONE) | pre (SOME (i,n)) = (i, SOME n); - fun add (RESULT l) [] (RESULT l') = RESULT (l @ l') - | add a (input1 as qtm :: qtms) (SINGLE (qtm',n)) = - if qtm = qtm' then SINGLE (qtm, add a qtms n) - else add a input1 (add n [qtm'] (MULTIPLE (NONE, NameArityMap.new ()))) - | add a (VAR :: qtms) (MULTIPLE (vs,fs)) = - MULTIPLE (SOME (oadd a qtms vs), fs) - | add a (FN (f,l) :: qtms) (MULTIPLE (vs,fs)) = + fun add (Result l) [] (Result l') = Result (l @ l') + | add a (input1 as qtm :: qtms) (Single (qtm',n)) = + if equalQterm qtm qtm' then Single (qtm, add a qtms n) + else add a input1 (add n [qtm'] (Multiple (NONE, NameArityMap.new ()))) + | add a (Var :: qtms) (Multiple (vs,fs)) = + Multiple (SOME (oadd a qtms vs), fs) + | add a (Fn (f,l) :: qtms) (Multiple (vs,fs)) = let val n = NameArityMap.peek fs f in - MULTIPLE (vs, NameArityMap.insert fs (f, oadd a (l @ qtms) n)) + Multiple (vs, NameArityMap.insert fs (f, oadd a (l @ qtms) n)) end | add _ _ _ = raise Bug "TermNet.insert: Match" and oadd a qtms NONE = singles qtms a | oadd a qtms (SOME n) = add a qtms n; - fun ins a qtm (i,n) = SOME (i + 1, oadd (RESULT [a]) [qtm] n); + fun ins a qtm (i,n) = SOME (i + 1, oadd (Result [a]) [qtm] n); in - fun insert (NET (p,k,n)) (tm,a) = - NET (p, k + 1, ins (k,a) (termToQterm tm) (pre n)) + fun insert (Net (p,k,n)) (tm,a) = + Net (p, k + 1, ins (k,a) (termToQterm tm) (pre n)) handle Error _ => raise Bug "TermNet.insert: should never fail"; end; @@ -155,26 +192,26 @@ fun filter pred = let - fun filt (RESULT l) = + fun filt (Result l) = (case List.filter (fn (_,a) => pred a) l of [] => NONE - | l => SOME (RESULT l)) - | filt (SINGLE (qtm,n)) = + | l => SOME (Result l)) + | filt (Single (qtm,n)) = (case filt n of NONE => NONE - | SOME n => SOME (SINGLE (qtm,n))) - | filt (MULTIPLE (vs,fs)) = + | SOME n => SOME (Single (qtm,n))) + | filt (Multiple (vs,fs)) = let val vs = Option.mapPartial filt vs val fs = NameArityMap.mapPartial (fn (_,n) => filt n) fs in if not (Option.isSome vs) andalso NameArityMap.null fs then NONE - else SOME (MULTIPLE (vs,fs)) + else SOME (Multiple (vs,fs)) end in - fn net as NET (_,_,NONE) => net - | NET (p, k, SOME (_,n)) => NET (p, k, netSize (filt n)) + fn net as Net (_,_,NONE) => net + | Net (p, k, SOME (_,n)) => Net (p, k, netSize (filt n)) end handle Error _ => raise Bug "TermNet.filter: should never fail"; @@ -189,7 +226,7 @@ let val (a,qtms) = revDivide qtms n in - addQterm (FN (f,a)) (ks,fs,qtms) + addQterm (Fn (f,a)) (ks,fs,qtms) end | norm stack = stack @@ -203,7 +240,7 @@ and addFn (f as (_,n)) (ks,fs,qtms) = norm (n :: ks, f :: fs, qtms); in val stackEmpty = ([],[],[]); - + val stackAddQterm = addQterm; val stackAddFn = addFn; @@ -216,16 +253,16 @@ fun fold _ acc [] = acc | fold inc acc ((0,stack,net) :: rest) = fold inc (inc (stackValue stack, net, acc)) rest - | fold inc acc ((n, stack, SINGLE (qtm,net)) :: rest) = + | fold inc acc ((n, stack, Single (qtm,net)) :: rest) = fold inc acc ((n - 1, stackAddQterm qtm stack, net) :: rest) - | fold inc acc ((n, stack, MULTIPLE (v,fns)) :: rest) = + | fold inc acc ((n, stack, Multiple (v,fns)) :: rest) = let val n = n - 1 val rest = case v of NONE => rest - | SOME net => (n, stackAddQterm VAR stack, net) :: rest + | SOME net => (n, stackAddQterm Var stack, net) :: rest fun getFns (f as (_,k), net, x) = (k + n, stackAddFn f stack, net) :: x @@ -240,11 +277,11 @@ fun foldEqualTerms pat inc acc = let fun fold ([],net) = inc (pat,net,acc) - | fold (pat :: pats, SINGLE (qtm,net)) = - if pat = qtm then fold (pats,net) else acc - | fold (VAR :: pats, MULTIPLE (v,_)) = + | fold (pat :: pats, Single (qtm,net)) = + if equalQterm pat qtm then fold (pats,net) else acc + | fold (Var :: pats, Multiple (v,_)) = (case v of NONE => acc | SOME net => fold (pats,net)) - | fold (FN (f,a) :: pats, MULTIPLE (_,fns)) = + | fold (Fn (f,a) :: pats, Multiple (_,fns)) = (case NameArityMap.peek fns f of NONE => acc | SOME net => fold (a @ pats, net)) @@ -257,20 +294,20 @@ fun fold _ acc [] = acc | fold inc acc (([],stack,net) :: rest) = fold inc (inc (stackValue stack, net, acc)) rest - | fold inc acc ((VAR :: pats, stack, net) :: rest) = + | fold inc acc ((Var :: pats, stack, net) :: rest) = let fun harvest (qtm,n,l) = (pats, stackAddQterm qtm stack, n) :: l in fold inc acc (foldTerms harvest rest net) end - | fold inc acc ((pat :: pats, stack, SINGLE (qtm,net)) :: rest) = + | fold inc acc ((pat :: pats, stack, Single (qtm,net)) :: rest) = (case unifyQtermQterm pat qtm of NONE => fold inc acc rest | SOME qtm => fold inc acc ((pats, stackAddQterm qtm stack, net) :: rest)) | fold inc acc - (((pat as FN (f,a)) :: pats, stack, MULTIPLE (v,fns)) :: rest) = + (((pat as Fn (f,a)) :: pats, stack, Multiple (v,fns)) :: rest) = let val rest = case v of @@ -307,10 +344,10 @@ local fun mat acc [] = acc - | mat acc ((RESULT l, []) :: rest) = mat (l @ acc) rest - | mat acc ((SINGLE (qtm,n), tm :: tms) :: rest) = + | mat acc ((Result l, []) :: rest) = mat (l @ acc) rest + | mat acc ((Single (qtm,n), tm :: tms) :: rest) = mat acc (if matchQtermTerm qtm tm then (n,tms) :: rest else rest) - | mat acc ((MULTIPLE (vs,fs), tm :: tms) :: rest) = + | mat acc ((Multiple (vs,fs), tm :: tms) :: rest) = let val rest = case vs of NONE => rest | SOME n => (n,tms) :: rest @@ -326,8 +363,8 @@ end | mat _ _ = raise Bug "TermNet.match: Match"; in - fun match (NET (_,_,NONE)) _ = [] - | match (NET (p, _, SOME (_,n))) tm = + fun match (Net (_,_,NONE)) _ = [] + | match (Net (p, _, SOME (_,n))) tm = finally p (mat [] [(n,[tm])]) handle Error _ => raise Bug "TermNet.match: should never fail"; end; @@ -339,16 +376,16 @@ fun seenInc qsub tms (_,net,rest) = (qsub,net,tms) :: rest; fun mat acc [] = acc - | mat acc ((_, RESULT l, []) :: rest) = mat (l @ acc) rest - | mat acc ((qsub, SINGLE (qtm,net), tm :: tms) :: rest) = + | mat acc ((_, Result l, []) :: rest) = mat (l @ acc) rest + | mat acc ((qsub, Single (qtm,net), tm :: tms) :: rest) = (case matchTermQterm qsub tm qtm of NONE => mat acc rest | SOME qsub => mat acc ((qsub,net,tms) :: rest)) - | mat acc ((qsub, net as MULTIPLE _, Term.Var v :: tms) :: rest) = + | mat acc ((qsub, net as Multiple _, Term.Var v :: tms) :: rest) = (case NameMap.peek qsub v of NONE => mat acc (foldTerms (unseenInc qsub v tms) rest net) | SOME qtm => mat acc (foldEqualTerms qtm (seenInc qsub tms) rest net)) - | mat acc ((qsub, MULTIPLE (_,fns), Term.Fn (f,a) :: tms) :: rest) = + | mat acc ((qsub, Multiple (_,fns), Term.Fn (f,a) :: tms) :: rest) = let val rest = case NameArityMap.peek fns (f, length a) of @@ -359,8 +396,8 @@ end | mat _ _ = raise Bug "TermNet.matched.mat"; in - fun matched (NET (_,_,NONE)) _ = [] - | matched (NET (parm, _, SOME (_,net))) tm = + fun matched (Net (_,_,NONE)) _ = [] + | matched (Net (parm, _, SOME (_,net))) tm = finally parm (mat [] [(NameMap.new (), net, [tm])]) handle Error _ => raise Bug "TermNet.matched: should never fail"; end; @@ -370,16 +407,16 @@ (NameMap.insert qsub (v,qtm), net, tms) :: rest; fun mat acc [] = acc - | mat acc ((_, RESULT l, []) :: rest) = mat (l @ acc) rest - | mat acc ((qsub, SINGLE (qtm,net), tm :: tms) :: rest) = + | mat acc ((_, Result l, []) :: rest) = mat (l @ acc) rest + | mat acc ((qsub, Single (qtm,net), tm :: tms) :: rest) = (case unifyQtermTerm qsub qtm tm of NONE => mat acc rest | SOME qsub => mat acc ((qsub,net,tms) :: rest)) - | mat acc ((qsub, net as MULTIPLE _, Term.Var v :: tms) :: rest) = + | mat acc ((qsub, net as Multiple _, Term.Var v :: tms) :: rest) = (case NameMap.peek qsub v of NONE => mat acc (foldTerms (inc qsub v tms) rest net) | SOME qtm => mat acc (foldUnifiableTerms qtm (inc qsub v tms) rest net)) - | mat acc ((qsub, MULTIPLE (v,fns), Term.Fn (f,a) :: tms) :: rest) = + | mat acc ((qsub, Multiple (v,fns), Term.Fn (f,a) :: tms) :: rest) = let val rest = case v of NONE => rest | SOME net => (qsub,net,tms) :: rest @@ -392,8 +429,8 @@ end | mat _ _ = raise Bug "TermNet.unify.mat"; in - fun unify (NET (_,_,NONE)) _ = [] - | unify (NET (parm, _, SOME (_,net))) tm = + fun unify (Net (_,_,NONE)) _ = [] + | unify (Net (parm, _, SOME (_,net))) tm = finally parm (mat [] [(NameMap.new (), net, [tm])]) handle Error _ => raise Bug "TermNet.unify: should never fail"; end; @@ -403,16 +440,16 @@ (* ------------------------------------------------------------------------- *) local - fun inc (qtm, RESULT l, acc) = + fun inc (qtm, Result l, acc) = foldl (fn ((n,a),acc) => (n,(qtm,a)) :: acc) acc l | inc _ = raise Bug "TermNet.pp.inc"; - - fun toList (NET (_,_,NONE)) = [] - | toList (NET (parm, _, SOME (_,net))) = + + fun toList (Net (_,_,NONE)) = [] + | toList (Net (parm, _, SOME (_,net))) = finally parm (foldTerms inc [] net); in fun pp ppA = - Parser.ppMap toList (Parser.ppList (Parser.ppBinop " |->" ppQterm ppA)); + Print.ppMap toList (Print.ppList (Print.ppOp2 " |->" ppQterm ppA)); end; end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Thm.sig --- a/src/Tools/Metis/src/Thm.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Thm.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) -(* A LOGICAL KERNEL FOR FIRST ORDER CLAUSES *) -(* Copyright (c) 2001-2004 Joe Hurd, distributed under the BSD License *) +(* A LOGICAL KERNEL FOR FIRST ORDER CLAUSAL THEOREMS *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Thm = @@ -10,6 +10,12 @@ (* An abstract type of first order logic theorems. *) (* ------------------------------------------------------------------------- *) +type thm + +(* ------------------------------------------------------------------------- *) +(* Theorem destructors. *) +(* ------------------------------------------------------------------------- *) + type clause = LiteralSet.set datatype inferenceType = @@ -21,14 +27,8 @@ | Refl | Equality -type thm - type inference = inferenceType * thm list -(* ------------------------------------------------------------------------- *) -(* Theorem destructors. *) -(* ------------------------------------------------------------------------- *) - val clause : thm -> clause val inference : thm -> inference @@ -79,11 +79,11 @@ (* Pretty-printing. *) (* ------------------------------------------------------------------------- *) -val ppInferenceType : inferenceType Parser.pp +val ppInferenceType : inferenceType Print.pp val inferenceTypeToString : inferenceType -> string -val pp : thm Parser.pp +val pp : thm Print.pp val toString : thm -> string diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Thm.sml --- a/src/Tools/Metis/src/Thm.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Thm.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) -(* A LOGICAL KERNEL FOR FIRST ORDER CLAUSES *) -(* Copyright (c) 2001-2004 Joe Hurd, distributed under the BSD License *) +(* A LOGICAL KERNEL FOR FIRST ORDER CLAUSAL THEOREMS *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Thm :> Thm = @@ -85,13 +85,9 @@ (* Free variables. *) (* ------------------------------------------------------------------------- *) -fun freeIn v (Thm (cl,_)) = LiteralSet.exists (Literal.freeIn v) cl; +fun freeIn v (Thm (cl,_)) = LiteralSet.freeIn v cl; -local - fun free (lit,set) = NameSet.union (Literal.freeVars lit) set; -in - fun freeVars (Thm (cl,_)) = LiteralSet.foldl free NameSet.empty cl; -end; +fun freeVars (Thm (cl,_)) = LiteralSet.freeVars cl; (* ------------------------------------------------------------------------- *) (* Pretty-printing. *) @@ -105,26 +101,21 @@ | inferenceTypeToString Refl = "Refl" | inferenceTypeToString Equality = "Equality"; -fun ppInferenceType ppstrm inf = - Parser.ppString ppstrm (inferenceTypeToString inf); +fun ppInferenceType inf = + Print.ppString (inferenceTypeToString inf); local fun toFormula th = Formula.listMkDisj (map Literal.toFormula (LiteralSet.toList (clause th))); in - fun pp ppstrm th = - let - open PP - in - begin_block ppstrm INCONSISTENT 3; - add_string ppstrm "|- "; - Formula.pp ppstrm (toFormula th); - end_block ppstrm - end; + fun pp th = + Print.blockProgram Print.Inconsistent 3 + [Print.addString "|- ", + Formula.pp (toFormula th)]; end; -val toString = Parser.toString pp; +val toString = Print.toString pp; (* ------------------------------------------------------------------------- *) (* Primitive rules of inference. *) @@ -157,7 +148,7 @@ let val cl' = LiteralSet.subst sub cl in - if Sharing.pointerEqual (cl,cl') then th + if Portable.pointerEqual (cl,cl') then th else case inf of (Subst,_) => Thm (cl',inf) @@ -181,7 +172,7 @@ Thm (LiteralSet.union cl1' cl2', (Resolve,[th1,th2])) end; -(*DEBUG +(*MetisDebug val resolve = fn lit => fn pos => fn neg => resolve lit pos neg handle Error err => diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Tptp.sig --- a/src/Tools/Metis/src/Tptp.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Tptp.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,18 +1,73 @@ (* ========================================================================= *) -(* THE TPTP PROBLEM FILE FORMAT (TPTP v2) *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* THE TPTP PROBLEM FILE FORMAT *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Tptp = sig (* ------------------------------------------------------------------------- *) -(* Mapping TPTP functions and relations to different names. *) +(* Mapping to and from TPTP variable, function and relation names. *) +(* ------------------------------------------------------------------------- *) + +type mapping + +val defaultMapping : mapping + +val mkMapping : + {functionMapping : {name : Name.name, arity : int, tptp : string} list, + relationMapping : {name : Name.name, arity : int, tptp : string} list} -> + mapping + +val addVarSetMapping : mapping -> NameSet.set -> mapping + +(* ------------------------------------------------------------------------- *) +(* Interpreting TPTP functions and relations in a finite model. *) +(* ------------------------------------------------------------------------- *) + +val defaultFixedMap : Model.fixedMap + +val defaultModel : Model.parameters + +val ppFixedMap : mapping -> Model.fixedMap Print.pp + +(* ------------------------------------------------------------------------- *) +(* TPTP roles. *) (* ------------------------------------------------------------------------- *) -val functionMapping : {name : string, arity : int, tptp : string} list ref +datatype role = + AxiomRole + | ConjectureRole + | DefinitionRole + | NegatedConjectureRole + | PlainRole + | TheoremRole + | OtherRole of string; + +val isCnfConjectureRole : role -> bool + +val isFofConjectureRole : role -> bool + +val toStringRole : role -> string + +val fromStringRole : string -> role -val relationMapping : {name : string, arity : int, tptp : string} list ref +val ppRole : role Print.pp + +(* ------------------------------------------------------------------------- *) +(* SZS statuses. *) +(* ------------------------------------------------------------------------- *) + +datatype status = + CounterSatisfiableStatus + | TheoremStatus + | SatisfiableStatus + | UnknownStatus + | UnsatisfiableStatus + +val toStringStatus : status -> string + +val ppStatus : status Print.pp (* ------------------------------------------------------------------------- *) (* TPTP literals. *) @@ -22,66 +77,152 @@ Boolean of bool | Literal of Literal.literal -val negate : literal -> literal +val negateLiteral : literal -> literal + +val functionsLiteral : literal -> NameAritySet.set + +val relationLiteral : literal -> Atom.relation option + +val freeVarsLiteral : literal -> NameSet.set -val literalFunctions : literal -> NameAritySet.set +(* ------------------------------------------------------------------------- *) +(* TPTP formula names. *) +(* ------------------------------------------------------------------------- *) + +datatype formulaName = + FormulaName of string + +val ppFormulaName : formulaName Print.pp -val literalRelation : literal -> Atom.relation option +(* ------------------------------------------------------------------------- *) +(* TPTP formula bodies. *) +(* ------------------------------------------------------------------------- *) + +datatype formulaBody = + CnfFormulaBody of literal list + | FofFormulaBody of Formula.formula + +(* ------------------------------------------------------------------------- *) +(* TPTP formula sources. *) +(* ------------------------------------------------------------------------- *) -val literalFreeVars : literal -> NameSet.set +datatype formulaSource = + NoFormulaSource + | StripFormulaSource of + {inference : string, + parents : formulaName list} + | NormalizeFormulaSource of + {inference : Normalize.inference, + parents : formulaName list} + | ProofFormulaSource of + {inference : Proof.inference, + parents : formulaName list} (* ------------------------------------------------------------------------- *) (* TPTP formulas. *) (* ------------------------------------------------------------------------- *) datatype formula = - CnfFormula of {name : string, role : string, clause : literal list} - | FofFormula of {name : string, role : string, formula : Formula.formula} + Formula of + {name : formulaName, + role : role, + body : formulaBody, + source : formulaSource} + +val nameFormula : formula -> formulaName + +val roleFormula : formula -> role -val formulaFunctions : formula -> NameAritySet.set +val bodyFormula : formula -> formulaBody + +val sourceFormula : formula -> formulaSource + +val functionsFormula : formula -> NameAritySet.set -val formulaRelations : formula -> NameAritySet.set +val relationsFormula : formula -> NameAritySet.set -val formulaFreeVars : formula -> NameSet.set +val freeVarsFormula : formula -> NameSet.set + +val freeVarsListFormula : formula list -> NameSet.set -val formulaIsConjecture : formula -> bool +val isCnfConjectureFormula : formula -> bool +val isFofConjectureFormula : formula -> bool +val isConjectureFormula : formula -> bool -val ppFormula : formula Parser.pp +(* ------------------------------------------------------------------------- *) +(* Clause information. *) +(* ------------------------------------------------------------------------- *) + +datatype clauseSource = + CnfClauseSource of formulaName * literal list + | FofClauseSource of Normalize.thm -val parseFormula : char Stream.stream -> formula Stream.stream +type 'a clauseInfo = 'a LiteralSetMap.map + +type clauseNames = formulaName clauseInfo + +type clauseRoles = role clauseInfo -val formulaToString : formula -> string +type clauseSources = clauseSource clauseInfo + +val noClauseNames : clauseNames -val formulaFromString : string -> formula +val noClauseRoles : clauseRoles + +val noClauseSources : clauseSources (* ------------------------------------------------------------------------- *) (* TPTP problems. *) (* ------------------------------------------------------------------------- *) -datatype goal = - Cnf of Problem.problem - | Fof of Formula.formula +type comments = string list + +type includes = string list -type problem = {comments : string list, formulas : formula list} +datatype problem = + Problem of + {comments : comments, + includes : includes, + formulas : formula list} -val read : {filename : string} -> problem - -val write : {filename : string} -> problem -> unit - +val hasCnfConjecture : problem -> bool +val hasFofConjecture : problem -> bool val hasConjecture : problem -> bool -val toGoal : problem -> goal +val freeVars : problem -> NameSet.set + +val mkProblem : + {comments : comments, + includes : includes, + names : clauseNames, + roles : clauseRoles, + problem : Problem.problem} -> problem -val fromProblem : Problem.problem -> problem +val normalize : + problem -> + {subgoal : Formula.formula * formulaName list, + problem : Problem.problem, + sources : clauseSources} list + +val goal : problem -> Formula.formula -val prove : {filename : string} -> bool +val read : {mapping : mapping, filename : string} -> problem + +val write : + {problem : problem, + mapping : mapping, + filename : string} -> unit + +val prove : {filename : string, mapping : mapping} -> bool (* ------------------------------------------------------------------------- *) (* TSTP proofs. *) (* ------------------------------------------------------------------------- *) -val ppProof : Proof.proof Parser.pp - -val proofToString : Proof.proof -> string +val fromProof : + {problem : problem, + proofs : {subgoal : Formula.formula * formulaName list, + sources : clauseSources, + refutation : Thm.thm} list} -> formula list end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Tptp.sml --- a/src/Tools/Metis/src/Tptp.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Tptp.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) -(* THE TPTP PROBLEM FILE FORMAT (TPTP v2) *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* THE TPTP PROBLEM FILE FORMAT *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Tptp :> Tptp = @@ -9,11 +9,95 @@ open Useful; (* ------------------------------------------------------------------------- *) -(* Constants. *) +(* Default TPTP function and relation name mapping. *) +(* ------------------------------------------------------------------------- *) + +val defaultFunctionMapping = + [(* Mapping TPTP functions to infix symbols *) + {name = "~", arity = 1, tptp = "negate"}, + {name = "*", arity = 2, tptp = "multiply"}, + {name = "/", arity = 2, tptp = "divide"}, + {name = "+", arity = 2, tptp = "add"}, + {name = "-", arity = 2, tptp = "subtract"}, + {name = "::", arity = 2, tptp = "cons"}, + {name = "@", arity = 2, tptp = "append"}, + {name = ",", arity = 2, tptp = "pair"}, + (* Expanding HOL symbols to TPTP alphanumerics *) + {name = ":", arity = 2, tptp = "has_type"}, + {name = ".", arity = 2, tptp = "apply"}]; + +val defaultRelationMapping = + [(* Mapping TPTP relations to infix symbols *) + {name = "=", arity = 2, tptp = "="}, (* this preserves the = symbol *) + {name = "==", arity = 2, tptp = "equalish"}, + {name = "<=", arity = 2, tptp = "less_equal"}, + {name = "<", arity = 2, tptp = "less_than"}, + {name = ">=", arity = 2, tptp = "greater_equal"}, + {name = ">", arity = 2, tptp = "greater_than"}, + (* Expanding HOL symbols to TPTP alphanumerics *) + {name = "{}", arity = 1, tptp = "bool"}]; + +(* ------------------------------------------------------------------------- *) +(* Interpreting TPTP functions and relations in a finite model. *) (* ------------------------------------------------------------------------- *) -val ROLE_NEGATED_CONJECTURE = "negated_conjecture" -and ROLE_CONJECTURE = "conjecture"; +val defaultFunctionModel = + [{name = "~", arity = 1, model = Model.negName}, + {name = "*", arity = 2, model = Model.multName}, + {name = "/", arity = 2, model = Model.divName}, + {name = "+", arity = 2, model = Model.addName}, + {name = "-", arity = 2, model = Model.subName}, + {name = "::", arity = 2, model = Model.consName}, + {name = "@", arity = 2, model = Model.appendName}, + {name = ":", arity = 2, model = Term.hasTypeFunctionName}, + {name = "additive_identity", arity = 0, model = Model.numeralName 0}, + {name = "app", arity = 2, model = Model.appendName}, + {name = "complement", arity = 1, model = Model.complementName}, + {name = "difference", arity = 2, model = Model.differenceName}, + {name = "divide", arity = 2, model = Model.divName}, + {name = "empty_set", arity = 0, model = Model.emptyName}, + {name = "identity", arity = 0, model = Model.numeralName 1}, + {name = "identity_map", arity = 1, model = Model.projectionName 1}, + {name = "intersection", arity = 2, model = Model.intersectName}, + {name = "minus", arity = 1, model = Model.negName}, + {name = "multiplicative_identity", arity = 0, model = Model.numeralName 1}, + {name = "n0", arity = 0, model = Model.numeralName 0}, + {name = "n1", arity = 0, model = Model.numeralName 1}, + {name = "n2", arity = 0, model = Model.numeralName 2}, + {name = "n3", arity = 0, model = Model.numeralName 3}, + {name = "n4", arity = 0, model = Model.numeralName 4}, + {name = "n5", arity = 0, model = Model.numeralName 5}, + {name = "n6", arity = 0, model = Model.numeralName 6}, + {name = "n7", arity = 0, model = Model.numeralName 7}, + {name = "n8", arity = 0, model = Model.numeralName 8}, + {name = "n9", arity = 0, model = Model.numeralName 9}, + {name = "nil", arity = 0, model = Model.nilName}, + {name = "null_class", arity = 0, model = Model.emptyName}, + {name = "singleton", arity = 1, model = Model.singletonName}, + {name = "successor", arity = 1, model = Model.sucName}, + {name = "symmetric_difference", arity = 2, + model = Model.symmetricDifferenceName}, + {name = "union", arity = 2, model = Model.unionName}, + {name = "universal_class", arity = 0, model = Model.universeName}]; + +val defaultRelationModel = + [{name = "=", arity = 2, model = Atom.eqRelationName}, + {name = "==", arity = 2, model = Atom.eqRelationName}, + {name = "<=", arity = 2, model = Model.leName}, + {name = "<", arity = 2, model = Model.ltName}, + {name = ">=", arity = 2, model = Model.geName}, + {name = ">", arity = 2, model = Model.gtName}, + {name = "divides", arity = 2, model = Model.dividesName}, + {name = "element_of_set", arity = 2, model = Model.memberName}, + {name = "equal", arity = 2, model = Atom.eqRelationName}, + {name = "equal_elements", arity = 2, model = Atom.eqRelationName}, + {name = "equal_sets", arity = 2, model = Atom.eqRelationName}, + {name = "equivalent", arity = 2, model = Atom.eqRelationName}, + {name = "less", arity = 2, model = Model.ltName}, + {name = "less_or_equal", arity = 2, model = Model.leName}, + {name = "member", arity = 2, model = Model.memberName}, + {name = "subclass", arity = 2, model = Model.subsetName}, + {name = "subset", arity = 2, model = Model.subsetName}]; (* ------------------------------------------------------------------------- *) (* Helper functions. *) @@ -29,97 +113,451 @@ n > 0 andalso hp (String.sub (s,0)) andalso ct (n - 1) end; -(* ------------------------------------------------------------------------- *) -(* Mapping TPTP functions and relations to different names. *) -(* ------------------------------------------------------------------------- *) - -val functionMapping = ref - [(* Mapping TPTP functions to infix symbols *) - {name = "*", arity = 2, tptp = "multiply"}, - {name = "/", arity = 2, tptp = "divide"}, - {name = "+", arity = 2, tptp = "add"}, - {name = "-", arity = 2, tptp = "subtract"}, - {name = "::", arity = 2, tptp = "cons"}, - {name = ",", arity = 2, tptp = "pair"}, - (* Expanding HOL symbols to TPTP alphanumerics *) - {name = ":", arity = 2, tptp = "has_type"}, - {name = ".", arity = 2, tptp = "apply"}, - {name = "<=", arity = 0, tptp = "less_equal"}]; - -val relationMapping = ref - [(* Mapping TPTP relations to infix symbols *) - {name = "=", arity = 2, tptp = "="}, - {name = "==", arity = 2, tptp = "equalish"}, - {name = "<=", arity = 2, tptp = "less_equal"}, - {name = "<", arity = 2, tptp = "less_than"}, - (* Expanding HOL symbols to TPTP alphanumerics *) - {name = "{}", arity = 1, tptp = "bool"}]; - -fun mappingToTptp x = - let - fun mk ({name,arity,tptp},m) = NameArityMap.insert m ((name,arity),tptp) - in - foldl mk (NameArityMap.new ()) x - end; - -fun mappingFromTptp x = +fun stripSuffix pred s = let - fun mk ({name,arity,tptp},m) = NameArityMap.insert m ((tptp,arity),name) + fun f 0 = "" + | f n = + let + val n' = n - 1 + in + if pred (String.sub (s,n')) then f n' + else String.substring (s,0,n) + end in - foldl mk (NameArityMap.new ()) x - end; - -fun findMapping mapping (name_arity as (n,_)) = - Option.getOpt (NameArityMap.peek mapping name_arity, n); - -fun mapTerm functionMap = - let - val mapName = findMapping functionMap - - fun mapTm (tm as Term.Var _) = tm - | mapTm (Term.Fn (f,a)) = Term.Fn (mapName (f, length a), map mapTm a) - in - mapTm + f (size s) end; -fun mapAtom {functionMap,relationMap} (p,a) = - (findMapping relationMap (p, length a), map (mapTerm functionMap) a); - -fun mapFof maps = +fun variant avoid s = + if not (StringSet.member s avoid) then s + else + let + val s = stripSuffix Char.isDigit s + + fun var i = + let + val s_i = s ^ Int.toString i + in + if not (StringSet.member s_i avoid) then s_i else var (i + 1) + end + in + var 0 + end; + +(* ------------------------------------------------------------------------- *) +(* Mapping to legal TPTP names. *) +(* ------------------------------------------------------------------------- *) + +local + fun nonEmptyPred p l = + case l of + [] => false + | c :: cs => p (c,cs); + + fun existsPred l x = List.exists (fn p => p x) l; + + fun isTptpChar #"_" = true + | isTptpChar c = Char.isAlphaNum c; + + fun isTptpName l s = nonEmptyPred (existsPred l) (explode s); + + fun isRegular (c,cs) = + Char.isLower c andalso List.all isTptpChar cs; + + fun isNumber (c,cs) = + Char.isDigit c andalso List.all Char.isDigit cs; + + fun isDefined (c,cs) = + c = #"$" andalso nonEmptyPred isRegular cs; + + fun isSystem (c,cs) = + c = #"$" andalso nonEmptyPred isDefined cs; +in + fun mkTptpVarName s = + let + val s = + case List.filter isTptpChar (explode s) of + [] => [#"X"] + | l as c :: cs => + if Char.isUpper c then l + else if Char.isLower c then Char.toUpper c :: cs + else #"X" :: l + in + implode s + end; + + val isTptpConstName = isTptpName [isRegular,isNumber,isDefined,isSystem] + and isTptpFnName = isTptpName [isRegular,isDefined,isSystem] + and isTptpPropName = isTptpName [isRegular,isDefined,isSystem] + and isTptpRelName = isTptpName [isRegular,isDefined,isSystem]; + + val isTptpFormulaName = isTptpName [isRegular,isNumber]; +end; + +(* ------------------------------------------------------------------------- *) +(* Mapping to legal TPTP variable names. *) +(* ------------------------------------------------------------------------- *) + +datatype varToTptp = VarToTptp of StringSet.set * string NameMap.map; + +val emptyVarToTptp = VarToTptp (StringSet.empty, NameMap.new ()); + +fun addVarToTptp vm v = let - open Formula - - fun form True = True - | form False = False - | form (Atom a) = Atom (mapAtom maps a) - | form (Not p) = Not (form p) - | form (And (p,q)) = And (form p, form q) - | form (Or (p,q)) = Or (form p, form q) - | form (Imp (p,q)) = Imp (form p, form q) - | form (Iff (p,q)) = Iff (form p, form q) - | form (Forall (v,p)) = Forall (v, form p) - | form (Exists (v,p)) = Exists (v, form p) + val VarToTptp (avoid,mapping) = vm in - form + if NameMap.inDomain v mapping then vm + else + let + val s = variant avoid (mkTptpVarName (Name.toString v)) + + val avoid = StringSet.add avoid s + and mapping = NameMap.insert mapping (v,s) + in + VarToTptp (avoid,mapping) + end + end; + +local + fun add (v,vm) = addVarToTptp vm v; +in + val addListVarToTptp = List.foldl add; + + val addSetVarToTptp = NameSet.foldl add; +end; + +val fromListVarToTptp = addListVarToTptp emptyVarToTptp; + +val fromSetVarToTptp = addSetVarToTptp emptyVarToTptp; + +fun getVarToTptp vm v = + let + val VarToTptp (_,mapping) = vm + in + case NameMap.peek mapping v of + SOME s => s + | NONE => raise Bug "Tptp.getVarToTptp: unknown var" end; (* ------------------------------------------------------------------------- *) -(* Comments. *) +(* Mapping from TPTP variable names. *) +(* ------------------------------------------------------------------------- *) + +fun getVarFromTptp s = Name.fromString s; + +(* ------------------------------------------------------------------------- *) +(* Mapping to TPTP function and relation names. *) +(* ------------------------------------------------------------------------- *) + +datatype nameToTptp = NameToTptp of string NameArityMap.map; + +local + val emptyNames : string NameArityMap.map = NameArityMap.new (); + + fun addNames ({name,arity,tptp},mapping) = + NameArityMap.insert mapping ((name,arity),tptp); + + val fromListNames = List.foldl addNames emptyNames; +in + fun mkNameToTptp mapping = NameToTptp (fromListNames mapping); +end; + +local + fun escapeChar c = + case c of + #"\\" => "\\\\" + | #"'" => "\\'" + | #"\n" => "\\n" + | #"\t" => "\\t" + | _ => str c; + + val escapeString = String.translate escapeChar; +in + fun singleQuote s = "'" ^ escapeString s ^ "'"; +end; + +fun getNameToTptp isTptp s = if isTptp s then s else singleQuote s; + +fun getNameArityToTptp isZeroTptp isPlusTptp (NameToTptp mapping) na = + case NameArityMap.peek mapping na of + SOME s => s + | NONE => + let + val (n,a) = na + val isTptp = if a = 0 then isZeroTptp else isPlusTptp + val s = Name.toString n + in + getNameToTptp isTptp s + end; + +(* ------------------------------------------------------------------------- *) +(* Mapping from TPTP function and relation names. *) +(* ------------------------------------------------------------------------- *) + +datatype nameFromTptp = NameFromTptp of (string * int, Name.name) Map.map; + +local + val stringArityCompare = prodCompare String.compare Int.compare; + + val emptyStringArityMap = Map.new stringArityCompare; + + fun addStringArityMap ({name,arity,tptp},mapping) = + Map.insert mapping ((tptp,arity),name); + + val fromListStringArityMap = + List.foldl addStringArityMap emptyStringArityMap; +in + fun mkNameFromTptp mapping = NameFromTptp (fromListStringArityMap mapping); +end; + +fun getNameFromTptp (NameFromTptp mapping) sa = + case Map.peek mapping sa of + SOME n => n + | NONE => + let + val (s,_) = sa + in + Name.fromString s + end; + +(* ------------------------------------------------------------------------- *) +(* Mapping to and from TPTP variable, function and relation names. *) (* ------------------------------------------------------------------------- *) -fun mkComment "" = "%" - | mkComment line = "% " ^ line; - -fun destComment "" = "" - | destComment l = +datatype mapping = + Mapping of + {varTo : varToTptp, + fnTo : nameToTptp, + relTo : nameToTptp, + fnFrom : nameFromTptp, + relFrom : nameFromTptp}; + +fun mkMapping mapping = + let + val {functionMapping,relationMapping} = mapping + + val varTo = emptyVarToTptp + val fnTo = mkNameToTptp functionMapping + val relTo = mkNameToTptp relationMapping + + val fnFrom = mkNameFromTptp functionMapping + val relFrom = mkNameFromTptp relationMapping + in + Mapping + {varTo = varTo, + fnTo = fnTo, + relTo = relTo, + fnFrom = fnFrom, + relFrom = relFrom} + end; + +fun addVarListMapping mapping vs = + let + val Mapping + {varTo, + fnTo, + relTo, + fnFrom, + relFrom} = mapping + + val varTo = addListVarToTptp varTo vs + in + Mapping + {varTo = varTo, + fnTo = fnTo, + relTo = relTo, + fnFrom = fnFrom, + relFrom = relFrom} + end; + +fun addVarSetMapping mapping vs = + let + val Mapping + {varTo, + fnTo, + relTo, + fnFrom, + relFrom} = mapping + + val varTo = addSetVarToTptp varTo vs + in + Mapping + {varTo = varTo, + fnTo = fnTo, + relTo = relTo, + fnFrom = fnFrom, + relFrom = relFrom} + end; + +fun varToTptp mapping v = + let + val Mapping {varTo,...} = mapping + in + getVarToTptp varTo v + end; + +fun fnToTptp mapping fa = + let + val Mapping {fnTo,...} = mapping + in + getNameArityToTptp isTptpConstName isTptpFnName fnTo fa + end; + +fun relToTptp mapping ra = + let + val Mapping {relTo,...} = mapping + in + getNameArityToTptp isTptpPropName isTptpRelName relTo ra + end; + +fun varFromTptp (_ : mapping) v = getVarFromTptp v; + +fun fnFromTptp mapping fa = + let + val Mapping {fnFrom,...} = mapping + in + getNameFromTptp fnFrom fa + end; + +fun relFromTptp mapping ra = + let + val Mapping {relFrom,...} = mapping + in + getNameFromTptp relFrom ra + end; + +val defaultMapping = let - val _ = String.sub (l,0) = #"%" orelse raise Error "destComment" - val n = if size l >= 2 andalso String.sub (l,1) = #" " then 2 else 1 + fun lift {name,arity,tptp} = + {name = Name.fromString name, arity = arity, tptp = tptp} + + val functionMapping = map lift defaultFunctionMapping + and relationMapping = map lift defaultRelationMapping + + val mapping = + {functionMapping = functionMapping, + relationMapping = relationMapping} + in + mkMapping mapping + end; + +(* ------------------------------------------------------------------------- *) +(* Interpreting TPTP functions and relations in a finite model. *) +(* ------------------------------------------------------------------------- *) + +fun mkFixedMap funcModel relModel = + let + fun mkEntry {name,arity,model} = ((Name.fromString name, arity), model) + + fun mkMap l = NameArityMap.fromList (map mkEntry l) + in + {functionMap = mkMap funcModel, + relationMap = mkMap relModel} + end; + +val defaultFixedMap = mkFixedMap defaultFunctionModel defaultRelationModel; + +val defaultModel = + let + val {size = N, fixed = fix} = Model.default + + val fix = Model.mapFixed defaultFixedMap fix in - String.extract (l,n,NONE) + {size = N, fixed = fix} end; -val isComment = can destComment; +local + fun toTptpMap toTptp = + let + fun add ((src,arity),dest,m) = + let + val src = Name.fromString (toTptp (src,arity)) + in + NameArityMap.insert m ((src,arity),dest) + end + in + fn m => NameArityMap.foldl add (NameArityMap.new ()) m + end; + + fun toTptpFixedMap mapping fixMap = + let + val {functionMap = fnMap, relationMap = relMap} = fixMap + + val fnMap = toTptpMap (fnToTptp mapping) fnMap + and relMap = toTptpMap (relToTptp mapping) relMap + in + {functionMap = fnMap, + relationMap = relMap} + end; +in + fun ppFixedMap mapping fixMap = + Model.ppFixedMap (toTptpFixedMap mapping fixMap); +end; + +(* ------------------------------------------------------------------------- *) +(* TPTP roles. *) +(* ------------------------------------------------------------------------- *) + +datatype role = + AxiomRole + | ConjectureRole + | DefinitionRole + | NegatedConjectureRole + | PlainRole + | TheoremRole + | OtherRole of string; + +fun isCnfConjectureRole role = + case role of + NegatedConjectureRole => true + | _ => false; + +fun isFofConjectureRole role = + case role of + ConjectureRole => true + | _ => false; + +fun toStringRole role = + case role of + AxiomRole => "axiom" + | ConjectureRole => "conjecture" + | DefinitionRole => "definition" + | NegatedConjectureRole => "negated_conjecture" + | PlainRole => "plain" + | TheoremRole => "theorem" + | OtherRole s => s; + +fun fromStringRole s = + case s of + "axiom" => AxiomRole + | "conjecture" => ConjectureRole + | "definition" => DefinitionRole + | "negated_conjecture" => NegatedConjectureRole + | "plain" => PlainRole + | "theorem" => TheoremRole + | _ => OtherRole s; + +val ppRole = Print.ppMap toStringRole Print.ppString; + +(* ------------------------------------------------------------------------- *) +(* SZS statuses. *) +(* ------------------------------------------------------------------------- *) + +datatype status = + CounterSatisfiableStatus + | TheoremStatus + | SatisfiableStatus + | UnknownStatus + | UnsatisfiableStatus; + +fun toStringStatus status = + case status of + CounterSatisfiableStatus => "CounterSatisfiable" + | TheoremStatus => "Theorem" + | SatisfiableStatus => "Satisfiable" + | UnknownStatus => "Unknown" + | UnsatisfiableStatus => "Unsatisfiable"; + +val ppStatus = Print.ppMap toStringStatus Print.ppString; (* ------------------------------------------------------------------------- *) (* TPTP literals. *) @@ -129,14 +567,29 @@ Boolean of bool | Literal of Literal.literal; -fun negate (Boolean b) = (Boolean (not b)) - | negate (Literal l) = (Literal (Literal.negate l)); - -fun literalFunctions (Boolean _) = NameAritySet.empty - | literalFunctions (Literal lit) = Literal.functions lit; - -fun literalRelation (Boolean _) = NONE - | literalRelation (Literal lit) = SOME (Literal.relation lit); +fun destLiteral lit = + case lit of + Literal l => l + | _ => raise Error "Tptp.destLiteral"; + +fun isBooleanLiteral lit = + case lit of + Boolean _ => true + | _ => false; + +fun equalBooleanLiteral b lit = + case lit of + Boolean b' => b = b' + | _ => false; + +fun negateLiteral (Boolean b) = (Boolean (not b)) + | negateLiteral (Literal l) = (Literal (Literal.negate l)); + +fun functionsLiteral (Boolean _) = NameAritySet.empty + | functionsLiteral (Literal lit) = Literal.functions lit; + +fun relationLiteral (Boolean _) = NONE + | relationLiteral (Literal lit) = SOME (Literal.relation lit); fun literalToFormula (Boolean true) = Formula.True | literalToFormula (Boolean false) = Formula.False @@ -146,107 +599,174 @@ | literalFromFormula Formula.False = Boolean false | literalFromFormula fm = Literal (Literal.fromFormula fm); -fun literalFreeVars (Boolean _) = NameSet.empty - | literalFreeVars (Literal lit) = Literal.freeVars lit; +fun freeVarsLiteral (Boolean _) = NameSet.empty + | freeVarsLiteral (Literal lit) = Literal.freeVars lit; fun literalSubst sub lit = case lit of Boolean _ => lit | Literal l => Literal (Literal.subst sub l); -fun mapLiteral maps lit = - case lit of - Boolean _ => lit - | Literal (p,a) => Literal (p, mapAtom maps a); - -fun destLiteral (Literal l) = l - | destLiteral _ = raise Error "destLiteral"; - (* ------------------------------------------------------------------------- *) (* Printing formulas using TPTP syntax. *) (* ------------------------------------------------------------------------- *) -val ppVar = Parser.ppString; - -local - fun term pp (Term.Var v) = ppVar pp v - | term pp (Term.Fn (c,[])) = Parser.addString pp c - | term pp (Term.Fn (f,tms)) = - (Parser.beginBlock pp Parser.Inconsistent 2; - Parser.addString pp (f ^ "("); - Parser.ppSequence "," term pp tms; - Parser.addString pp ")"; - Parser.endBlock pp); -in - fun ppTerm pp tm = - (Parser.beginBlock pp Parser.Inconsistent 0; - term pp tm; - Parser.endBlock pp); -end; - -fun ppAtom pp atm = ppTerm pp (Term.Fn atm); +fun ppVar mapping v = + let + val s = varToTptp mapping v + in + Print.addString s + end; + +fun ppFnName mapping fa = Print.addString (fnToTptp mapping fa); + +fun ppConst mapping c = ppFnName mapping (c,0); + +fun ppTerm mapping = + let + fun term tm = + case tm of + Term.Var v => ppVar mapping v + | Term.Fn (f,tms) => + case length tms of + 0 => ppConst mapping f + | a => + Print.blockProgram Print.Inconsistent 2 + [ppFnName mapping (f,a), + Print.addString "(", + Print.ppOpList "," term tms, + Print.addString ")"] + in + Print.block Print.Inconsistent 0 o term + end; + +fun ppRelName mapping ra = Print.addString (relToTptp mapping ra); + +fun ppProp mapping p = ppRelName mapping (p,0); + +fun ppAtom mapping (r,tms) = + case length tms of + 0 => ppProp mapping r + | a => + Print.blockProgram Print.Inconsistent 2 + [ppRelName mapping (r,a), + Print.addString "(", + Print.ppOpList "," (ppTerm mapping) tms, + Print.addString ")"]; local - open Formula; - - fun fof pp (fm as And _) = assoc_binary pp ("&", stripConj fm) - | fof pp (fm as Or _) = assoc_binary pp ("|", stripDisj fm) - | fof pp (Imp a_b) = nonassoc_binary pp ("=>",a_b) - | fof pp (Iff a_b) = nonassoc_binary pp ("<=>",a_b) - | fof pp fm = unitary pp fm - - and nonassoc_binary pp (s,a_b) = - Parser.ppBinop (" " ^ s) unitary unitary pp a_b - - and assoc_binary pp (s,l) = Parser.ppSequence (" " ^ s) unitary pp l - - and unitary pp fm = - if isForall fm then quantified pp ("!", stripForall fm) - else if isExists fm then quantified pp ("?", stripExists fm) - else if atom pp fm then () - else if isNeg fm then - let - fun pr () = (Parser.addString pp "~"; Parser.addBreak pp (1,0)) - val (n,fm) = Formula.stripNeg fm - in - Parser.beginBlock pp Parser.Inconsistent 2; - funpow n pr (); - unitary pp fm; - Parser.endBlock pp - end - else - (Parser.beginBlock pp Parser.Inconsistent 1; - Parser.addString pp "("; - fof pp fm; - Parser.addString pp ")"; - Parser.endBlock pp) - - and quantified pp (q,(vs,fm)) = - (Parser.beginBlock pp Parser.Inconsistent 2; - Parser.addString pp (q ^ " "); - Parser.beginBlock pp Parser.Inconsistent (String.size q); - Parser.addString pp "["; - Parser.ppSequence "," ppVar pp vs; - Parser.addString pp "] :"; - Parser.endBlock pp; - Parser.addBreak pp (1,0); - unitary pp fm; - Parser.endBlock pp) - - and atom pp True = (Parser.addString pp "$true"; true) - | atom pp False = (Parser.addString pp "$false"; true) - | atom pp fm = - case total destEq fm of - SOME a_b => (Parser.ppBinop " =" ppTerm ppTerm pp a_b; true) - | NONE => - case total destNeq fm of - SOME a_b => (Parser.ppBinop " !=" ppTerm ppTerm pp a_b; true) - | NONE => case fm of Atom atm => (ppAtom pp atm; true) | _ => false; + val neg = Print.sequence (Print.addString "~") (Print.addBreak 1); + + fun fof mapping fm = + case fm of + Formula.And _ => assoc_binary mapping ("&", Formula.stripConj fm) + | Formula.Or _ => assoc_binary mapping ("|", Formula.stripDisj fm) + | Formula.Imp a_b => nonassoc_binary mapping ("=>",a_b) + | Formula.Iff a_b => nonassoc_binary mapping ("<=>",a_b) + | _ => unitary mapping fm + + and nonassoc_binary mapping (s,a_b) = + Print.ppOp2 (" " ^ s) (unitary mapping) (unitary mapping) a_b + + and assoc_binary mapping (s,l) = Print.ppOpList (" " ^ s) (unitary mapping) l + + and unitary mapping fm = + case fm of + Formula.True => Print.addString "$true" + | Formula.False => Print.addString "$false" + | Formula.Forall _ => quantified mapping ("!", Formula.stripForall fm) + | Formula.Exists _ => quantified mapping ("?", Formula.stripExists fm) + | Formula.Not _ => + (case total Formula.destNeq fm of + SOME a_b => Print.ppOp2 " !=" (ppTerm mapping) (ppTerm mapping) a_b + | NONE => + let + val (n,fm) = Formula.stripNeg fm + in + Print.blockProgram Print.Inconsistent 2 + [Print.duplicate n neg, + unitary mapping fm] + end) + | Formula.Atom atm => + (case total Formula.destEq fm of + SOME a_b => Print.ppOp2 " =" (ppTerm mapping) (ppTerm mapping) a_b + | NONE => ppAtom mapping atm) + | _ => + Print.blockProgram Print.Inconsistent 1 + [Print.addString "(", + fof mapping fm, + Print.addString ")"] + + and quantified mapping (q,(vs,fm)) = + let + val mapping = addVarListMapping mapping vs + in + Print.blockProgram Print.Inconsistent 2 + [Print.addString q, + Print.addString " ", + Print.blockProgram Print.Inconsistent (String.size q) + [Print.addString "[", + Print.ppOpList "," (ppVar mapping) vs, + Print.addString "] :"], + Print.addBreak 1, + unitary mapping fm] + end; in - fun ppFof pp fm = - (Parser.beginBlock pp Parser.Inconsistent 0; - fof pp fm; - Parser.endBlock pp); + fun ppFof mapping fm = Print.block Print.Inconsistent 0 (fof mapping fm); +end; + +(* ------------------------------------------------------------------------- *) +(* Lexing TPTP files. *) +(* ------------------------------------------------------------------------- *) + +datatype token = + AlphaNum of string + | Punct of char + | Quote of string; + +fun isAlphaNum #"_" = true + | isAlphaNum c = Char.isAlphaNum c; + +local + open Parse; + + infixr 9 >>++ + infixr 8 ++ + infixr 7 >> + infixr 6 || + + val alphaNumToken = atLeastOne (some isAlphaNum) >> (AlphaNum o implode); + + val punctToken = + let + val punctChars = "<>=-*+/\\?@|!$%&#^:;~()[]{}.," + in + some (Char.contains punctChars) >> Punct + end; + + val quoteToken = + let + val escapeParser = + some (equal #"'") >> singleton || + some (equal #"\\") >> singleton + + fun stopOn #"'" = true + | stopOn #"\n" = true + | stopOn _ = false + + val quotedParser = + some (equal #"\\") ++ escapeParser >> op:: || + some (not o stopOn) >> singleton + in + exactChar #"'" ++ many quotedParser ++ exactChar #"'" >> + (fn (_,(l,_)) => Quote (implode (List.concat l))) + end; + + val lexToken = alphaNumToken || punctToken || quoteToken; + + val space = many (some Char.isSpace) >> K (); +in + val lexer = (space ++ lexToken ++ space) >> (fn ((),(tok,())) => tok); end; (* ------------------------------------------------------------------------- *) @@ -257,7 +777,7 @@ val clauseFunctions = let - fun funcs (lit,acc) = NameAritySet.union (literalFunctions lit) acc + fun funcs (lit,acc) = NameAritySet.union (functionsLiteral lit) acc in foldl funcs NameAritySet.empty end; @@ -265,7 +785,7 @@ val clauseRelations = let fun rels (lit,acc) = - case literalRelation lit of + case relationLiteral lit of NONE => acc | SOME r => NameAritySet.add acc r in @@ -274,15 +794,13 @@ val clauseFreeVars = let - fun fvs (lit,acc) = NameSet.union (literalFreeVars lit) acc + fun fvs (lit,acc) = NameSet.union (freeVarsLiteral lit) acc in foldl fvs NameSet.empty end; fun clauseSubst sub lits = map (literalSubst sub) lits; -fun mapClause maps lits = map (mapLiteral maps) lits; - fun clauseToFormula lits = Formula.listMkDisj (map literalToFormula lits); fun clauseFromFormula fm = map literalFromFormula (Formula.stripDisj fm); @@ -293,115 +811,475 @@ fun clauseFromThm th = clauseFromLiteralSet (Thm.clause th); -val ppClause = Parser.ppMap clauseToFormula ppFof; +fun ppClause mapping = Print.ppMap clauseToFormula (ppFof mapping); + +(* ------------------------------------------------------------------------- *) +(* TPTP formula names. *) +(* ------------------------------------------------------------------------- *) + +datatype formulaName = FormulaName of string; + +datatype formulaNameSet = FormulaNameSet of formulaName Set.set; + +fun compareFormulaName (FormulaName s1, FormulaName s2) = + String.compare (s1,s2); + +fun toTptpFormulaName (FormulaName s) = + getNameToTptp isTptpFormulaName s; + +val ppFormulaName = Print.ppMap toTptpFormulaName Print.ppString; + +val emptyFormulaNameSet = FormulaNameSet (Set.empty compareFormulaName); + +fun memberFormulaNameSet n (FormulaNameSet s) = Set.member n s; + +fun addFormulaNameSet (FormulaNameSet s) n = FormulaNameSet (Set.add s n); + +fun addListFormulaNameSet (FormulaNameSet s) l = + FormulaNameSet (Set.addList s l); + +(* ------------------------------------------------------------------------- *) +(* TPTP formula bodies. *) +(* ------------------------------------------------------------------------- *) + +datatype formulaBody = + CnfFormulaBody of literal list + | FofFormulaBody of Formula.formula; + +fun destCnfFormulaBody body = + case body of + CnfFormulaBody x => x + | _ => raise Error "destCnfFormulaBody"; + +val isCnfFormulaBody = can destCnfFormulaBody; + +fun destFofFormulaBody body = + case body of + FofFormulaBody x => x + | _ => raise Error "destFofFormulaBody"; + +val isFofFormulaBody = can destFofFormulaBody; + +fun formulaBodyFunctions body = + case body of + CnfFormulaBody cl => clauseFunctions cl + | FofFormulaBody fm => Formula.functions fm; + +fun formulaBodyRelations body = + case body of + CnfFormulaBody cl => clauseRelations cl + | FofFormulaBody fm => Formula.relations fm; + +fun formulaBodyFreeVars body = + case body of + CnfFormulaBody cl => clauseFreeVars cl + | FofFormulaBody fm => Formula.freeVars fm; + +fun ppFormulaBody mapping body = + case body of + CnfFormulaBody cl => ppClause mapping cl + | FofFormulaBody fm => ppFof mapping (Formula.generalize fm); + +(* ------------------------------------------------------------------------- *) +(* TPTP formula sources. *) +(* ------------------------------------------------------------------------- *) + +datatype formulaSource = + NoFormulaSource + | StripFormulaSource of + {inference : string, + parents : formulaName list} + | NormalizeFormulaSource of + {inference : Normalize.inference, + parents : formulaName list} + | ProofFormulaSource of + {inference : Proof.inference, + parents : formulaName list}; + +fun isNoFormulaSource source = + case source of + NoFormulaSource => true + | _ => false; + +fun functionsFormulaSource source = + case source of + NoFormulaSource => NameAritySet.empty + | StripFormulaSource _ => NameAritySet.empty + | NormalizeFormulaSource data => + let + val {inference = inf, parents = _} = data + in + case inf of + Normalize.Axiom fm => Formula.functions fm + | Normalize.Definition (_,fm) => Formula.functions fm + | _ => NameAritySet.empty + end + | ProofFormulaSource data => + let + val {inference = inf, parents = _} = data + in + case inf of + Proof.Axiom cl => LiteralSet.functions cl + | Proof.Assume atm => Atom.functions atm + | Proof.Subst (sub,_) => Subst.functions sub + | Proof.Resolve (atm,_,_) => Atom.functions atm + | Proof.Refl tm => Term.functions tm + | Proof.Equality (lit,_,tm) => + NameAritySet.union (Literal.functions lit) (Term.functions tm) + end; + +fun relationsFormulaSource source = + case source of + NoFormulaSource => NameAritySet.empty + | StripFormulaSource _ => NameAritySet.empty + | NormalizeFormulaSource data => + let + val {inference = inf, parents = _} = data + in + case inf of + Normalize.Axiom fm => Formula.relations fm + | Normalize.Definition (_,fm) => Formula.relations fm + | _ => NameAritySet.empty + end + | ProofFormulaSource data => + let + val {inference = inf, parents = _} = data + in + case inf of + Proof.Axiom cl => LiteralSet.relations cl + | Proof.Assume atm => NameAritySet.singleton (Atom.relation atm) + | Proof.Subst _ => NameAritySet.empty + | Proof.Resolve (atm,_,_) => NameAritySet.singleton (Atom.relation atm) + | Proof.Refl tm => NameAritySet.empty + | Proof.Equality (lit,_,_) => + NameAritySet.singleton (Literal.relation lit) + end; + +fun freeVarsFormulaSource source = + case source of + NoFormulaSource => NameSet.empty + | StripFormulaSource _ => NameSet.empty + | NormalizeFormulaSource data => NameSet.empty + | ProofFormulaSource data => + let + val {inference = inf, parents = _} = data + in + case inf of + Proof.Axiom cl => LiteralSet.freeVars cl + | Proof.Assume atm => Atom.freeVars atm + | Proof.Subst (sub,_) => Subst.freeVars sub + | Proof.Resolve (atm,_,_) => Atom.freeVars atm + | Proof.Refl tm => Term.freeVars tm + | Proof.Equality (lit,_,tm) => + NameSet.union (Literal.freeVars lit) (Term.freeVars tm) + end; + +local + val GEN_INFERENCE = "inference" + and GEN_INTRODUCED = "introduced"; + + fun nameStrip inf = inf; + + fun ppStrip mapping inf = Print.skip; + + fun nameNormalize inf = + case inf of + Normalize.Axiom _ => "canonicalize" + | Normalize.Definition _ => "canonicalize" + | Normalize.Simplify _ => "simplify" + | Normalize.Conjunct _ => "conjunct" + | Normalize.Specialize _ => "specialize" + | Normalize.Skolemize _ => "skolemize" + | Normalize.Clausify _ => "clausify"; + + fun ppNormalize mapping inf = Print.skip; + + fun nameProof inf = + case inf of + Proof.Axiom _ => "canonicalize" + | Proof.Assume _ => "assume" + | Proof.Subst _ => "subst" + | Proof.Resolve _ => "resolve" + | Proof.Refl _ => "refl" + | Proof.Equality _ => "equality"; + + local + fun ppTermInf mapping = ppTerm mapping; + + fun ppAtomInf mapping atm = + case total Atom.destEq atm of + SOME (a,b) => ppAtom mapping (Name.fromString "$equal", [a,b]) + | NONE => ppAtom mapping atm; + + fun ppLiteralInf mapping (pol,atm) = + Print.sequence + (if pol then Print.skip else Print.addString "~ ") + (ppAtomInf mapping atm); + in + fun ppProofTerm mapping = + Print.ppBracket "$fot(" ")" (ppTermInf mapping); + + fun ppProofAtom mapping = + Print.ppBracket "$cnf(" ")" (ppAtomInf mapping); + + fun ppProofLiteral mapping = + Print.ppBracket "$cnf(" ")" (ppLiteralInf mapping); + end; + + val ppProofVar = ppVar; + + val ppProofPath = Term.ppPath; + + fun ppProof mapping inf = + Print.blockProgram Print.Inconsistent 1 + [Print.addString "[", + (case inf of + Proof.Axiom _ => Print.skip + | Proof.Assume atm => ppProofAtom mapping atm + | Proof.Subst _ => Print.skip + | Proof.Resolve (atm,_,_) => ppProofAtom mapping atm + | Proof.Refl tm => ppProofTerm mapping tm + | Proof.Equality (lit,path,tm) => + Print.program + [ppProofLiteral mapping lit, + Print.addString ",", + Print.addBreak 1, + ppProofPath path, + Print.addString ",", + Print.addBreak 1, + ppProofTerm mapping tm]), + Print.addString "]"]; + + val ppParent = ppFormulaName; + + fun ppProofSubst mapping = + Print.ppMap Subst.toList + (Print.ppList + (Print.ppBracket "bind(" ")" + (Print.ppOp2 "," (ppProofVar mapping) + (ppProofTerm mapping)))); + + fun ppProofParent mapping (p,s) = + if Subst.null s then ppParent p + else Print.ppOp2 " :" ppParent (ppProofSubst mapping) (p,s); +in + fun ppFormulaSource mapping source = + case source of + NoFormulaSource => Print.skip + | StripFormulaSource {inference,parents} => + let + val gen = GEN_INFERENCE + + val name = nameStrip inference + in + Print.blockProgram Print.Inconsistent (size gen + 1) + [Print.addString gen, + Print.addString "(", + Print.addString name, + Print.addString ",", + Print.addBreak 1, + Print.ppBracket "[" "]" (ppStrip mapping) inference, + Print.addString ",", + Print.addBreak 1, + Print.ppList ppParent parents, + Print.addString ")"] + end + | NormalizeFormulaSource {inference,parents} => + let + val gen = GEN_INFERENCE + + val name = nameNormalize inference + in + Print.blockProgram Print.Inconsistent (size gen + 1) + [Print.addString gen, + Print.addString "(", + Print.addString name, + Print.addString ",", + Print.addBreak 1, + Print.ppBracket "[" "]" (ppNormalize mapping) inference, + Print.addString ",", + Print.addBreak 1, + Print.ppList ppParent parents, + Print.addString ")"] + end + | ProofFormulaSource {inference,parents} => + let + val isTaut = null parents + + val gen = if isTaut then GEN_INTRODUCED else GEN_INFERENCE + + val name = nameProof inference + + val parents = + let + val sub = + case inference of + Proof.Subst (s,_) => s + | _ => Subst.empty + in + map (fn parent => (parent,sub)) parents + end + in + Print.blockProgram Print.Inconsistent (size gen + 1) + ([Print.addString gen, + Print.addString "("] @ + (if isTaut then + [Print.addString "tautology", + Print.addString ",", + Print.addBreak 1, + Print.blockProgram Print.Inconsistent 1 + [Print.addString "[", + Print.addString name, + Print.addString ",", + Print.addBreak 1, + ppProof mapping inference, + Print.addString "]"]] + else + [Print.addString name, + Print.addString ",", + Print.addBreak 1, + ppProof mapping inference, + Print.addString ",", + Print.addBreak 1, + Print.ppList (ppProofParent mapping) parents]) @ + [Print.addString ")"]) + end +end; (* ------------------------------------------------------------------------- *) (* TPTP formulas. *) (* ------------------------------------------------------------------------- *) datatype formula = - CnfFormula of {name : string, role : string, clause : clause} - | FofFormula of {name : string, role : string, formula : Formula.formula}; - -fun destCnfFormula (CnfFormula x) = x - | destCnfFormula _ = raise Error "destCnfFormula"; + Formula of + {name : formulaName, + role : role, + body : formulaBody, + source : formulaSource}; + +fun nameFormula (Formula {name,...}) = name; + +fun roleFormula (Formula {role,...}) = role; + +fun bodyFormula (Formula {body,...}) = body; + +fun sourceFormula (Formula {source,...}) = source; + +fun destCnfFormula fm = destCnfFormulaBody (bodyFormula fm); val isCnfFormula = can destCnfFormula; -fun destFofFormula (FofFormula x) = x - | destFofFormula _ = raise Error "destFofFormula"; +fun destFofFormula fm = destFofFormulaBody (bodyFormula fm); val isFofFormula = can destFofFormula; -fun formulaFunctions (CnfFormula {clause,...}) = clauseFunctions clause - | formulaFunctions (FofFormula {formula,...}) = Formula.functions formula; - -fun formulaRelations (CnfFormula {clause,...}) = clauseRelations clause - | formulaRelations (FofFormula {formula,...}) = Formula.relations formula; - -fun formulaFreeVars (CnfFormula {clause,...}) = clauseFreeVars clause - | formulaFreeVars (FofFormula {formula,...}) = Formula.freeVars formula; - -fun mapFormula maps (CnfFormula {name,role,clause}) = - CnfFormula {name = name, role = role, clause = mapClause maps clause} - | mapFormula maps (FofFormula {name,role,formula}) = - FofFormula {name = name, role = role, formula = mapFof maps formula}; +fun functionsFormula fm = + let + val bodyFns = formulaBodyFunctions (bodyFormula fm) + and sourceFns = functionsFormulaSource (sourceFormula fm) + in + NameAritySet.union bodyFns sourceFns + end; + +fun relationsFormula fm = + let + val bodyRels = formulaBodyRelations (bodyFormula fm) + and sourceRels = relationsFormulaSource (sourceFormula fm) + in + NameAritySet.union bodyRels sourceRels + end; + +fun freeVarsFormula fm = + let + val bodyFvs = formulaBodyFreeVars (bodyFormula fm) + and sourceFvs = freeVarsFormulaSource (sourceFormula fm) + in + NameSet.union bodyFvs sourceFvs + end; + +val freeVarsListFormula = + let + fun add (fm,vs) = NameSet.union vs (freeVarsFormula fm) + in + List.foldl add NameSet.empty + end; val formulasFunctions = let - fun funcs (fm,acc) = NameAritySet.union (formulaFunctions fm) acc + fun funcs (fm,acc) = NameAritySet.union (functionsFormula fm) acc in foldl funcs NameAritySet.empty end; val formulasRelations = let - fun rels (fm,acc) = NameAritySet.union (formulaRelations fm) acc + fun rels (fm,acc) = NameAritySet.union (relationsFormula fm) acc in foldl rels NameAritySet.empty end; -fun formulaIsConjecture (CnfFormula {role,...}) = role = ROLE_NEGATED_CONJECTURE - | formulaIsConjecture (FofFormula {role,...}) = role = ROLE_CONJECTURE; +fun isCnfConjectureFormula fm = + case fm of + Formula {role, body = CnfFormulaBody _, ...} => isCnfConjectureRole role + | _ => false; + +fun isFofConjectureFormula fm = + case fm of + Formula {role, body = FofFormulaBody _, ...} => isFofConjectureRole role + | _ => false; + +fun isConjectureFormula fm = + isCnfConjectureFormula fm orelse + isFofConjectureFormula fm; + +(* Parsing and pretty-printing *) + +fun ppFormula mapping fm = + let + val Formula {name,role,body,source} = fm + + val gen = + case body of + CnfFormulaBody _ => "cnf" + | FofFormulaBody _ => "fof" + in + Print.blockProgram Print.Inconsistent (size gen + 1) + ([Print.addString gen, + Print.addString "(", + ppFormulaName name, + Print.addString ",", + Print.addBreak 1, + ppRole role, + Print.addString ",", + Print.addBreak 1, + Print.blockProgram Print.Consistent 1 + [Print.addString "(", + ppFormulaBody mapping body, + Print.addString ")"]] @ + (if isNoFormulaSource source then [] + else + [Print.addString ",", + Print.addBreak 1, + ppFormulaSource mapping source]) @ + [Print.addString ")."]) + end; + +fun formulaToString mapping = Print.toString (ppFormula mapping); local - open Parser; - + open Parse; + + infixr 9 >>++ infixr 8 ++ infixr 7 >> infixr 6 || - datatype token = - AlphaNum of string - | Punct of char - | Quote of string; - - fun isAlphaNum #"_" = true - | isAlphaNum c = Char.isAlphaNum c; - - local - val alphaNumToken = atLeastOne (some isAlphaNum) >> (AlphaNum o implode); - - val punctToken = - let - val punctChars = "<>=-*+/\\?@|!$%&#^:;~()[]{}.," - in - some (Char.contains punctChars) >> Punct - end; - - val quoteToken = - let - val escapeParser = - exact #"'" >> singleton || - exact #"\\" >> singleton - - fun stopOn #"'" = true - | stopOn #"\n" = true - | stopOn _ = false - - val quotedParser = - exact #"\\" ++ escapeParser >> op:: || - some (not o stopOn) >> singleton - in - exact #"'" ++ many quotedParser ++ exact #"'" >> - (fn (_,(l,_)) => Quote (implode (List.concat l))) - end; - - val lexToken = alphaNumToken || punctToken || quoteToken; - - val space = many (some Char.isSpace) >> K (); - in - val lexer = (space ++ lexToken ++ space) >> (fn ((),(tok,())) => tok); - end; - fun someAlphaNum p = maybe (fn AlphaNum s => if p s then SOME s else NONE | _ => NONE); fun alphaNumParser s = someAlphaNum (equal s) >> K (); - fun isLower s = Char.isLower (String.sub (s,0)); - - val lowerParser = someAlphaNum isLower; + val lowerParser = someAlphaNum (fn s => Char.isLower (String.sub (s,0))); val upperParser = someAlphaNum (fn s => Char.isUpper (String.sub (s,0))); @@ -414,12 +1292,7 @@ fun punctParser c = somePunct (equal c) >> K (); - fun quoteParser p = - let - fun q s = if p s then s else "'" ^ s ^ "'" - in - maybe (fn Quote s => SOME (q s) | _ => NONE) - end; + val quoteParser = maybe (fn Quote s => SOME s | _ => NONE); local fun f [] = raise Bug "symbolParser" @@ -436,16 +1309,19 @@ punctParser #"$" ++ punctParser #"$" ++ someAlphaNum (K true) >> (fn ((),((),s)) => "$$" ^ s); - val nameParser = stringParser || numberParser || quoteParser (K false); - - val roleParser = lowerParser; + val nameParser = + (stringParser || numberParser || quoteParser) >> FormulaName; + + val roleParser = lowerParser >> fromStringRole; local fun isProposition s = isHdTlString Char.isLower isAlphaNum s; in val propositionParser = someAlphaNum isProposition || - definedParser || systemParser || quoteParser isProposition; + definedParser || + systemParser || + quoteParser; end; local @@ -453,17 +1329,20 @@ in val functionParser = someAlphaNum isFunction || - definedParser || systemParser || quoteParser isFunction; + definedParser || + systemParser || + quoteParser; end; local - fun isConstant s = - isHdTlString Char.isLower isAlphaNum s orelse - isHdTlString Char.isDigit Char.isDigit s; + fun isConstant s = isHdTlString Char.isLower isAlphaNum s; in val constantParser = someAlphaNum isConstant || - definedParser || systemParser || quoteParser isConstant; + definedParser || + numberParser || + systemParser || + quoteParser; end; val varParser = upperParser; @@ -474,536 +1353,819 @@ punctParser #"]") >> (fn ((),(h,(t,()))) => h :: t); - fun termParser input = - ((functionArgumentsParser >> Term.Fn) || - nonFunctionArgumentsTermParser) input - - and functionArgumentsParser input = - ((functionParser ++ punctParser #"(" ++ termParser ++ - many ((punctParser #"," ++ termParser) >> snd) ++ - punctParser #")") >> - (fn (f,((),(t,(ts,())))) => (f, t :: ts))) input - - and nonFunctionArgumentsTermParser input = - ((varParser >> Term.Var) || - (constantParser >> (fn n => Term.Fn (n,[])))) input - - val binaryAtomParser = - ((punctParser #"=" ++ termParser) >> - (fn ((),r) => fn l => Literal.mkEq (l,r))) || - ((symbolParser "!=" ++ termParser) >> - (fn ((),r) => fn l => Literal.mkNeq (l,r))); - - val maybeBinaryAtomParser = - optional binaryAtomParser >> - (fn SOME f => (fn a => f (Term.Fn a)) - | NONE => (fn a => (true,a))); - - val literalAtomParser = - ((functionArgumentsParser ++ maybeBinaryAtomParser) >> - (fn (a,f) => f a)) || - ((nonFunctionArgumentsTermParser ++ binaryAtomParser) >> - (fn (a,f) => f a)) || - (propositionParser >> - (fn n => (true,(n,[])))); - - val atomParser = - literalAtomParser >> - (fn (pol,("$true",[])) => Boolean pol - | (pol,("$false",[])) => Boolean (not pol) - | (pol,("$equal",[a,b])) => Literal (pol, Atom.mkEq (a,b)) - | lit => Literal lit); - - val literalParser = - ((punctParser #"~" ++ atomParser) >> (negate o snd)) || - atomParser; - - val disjunctionParser = - (literalParser ++ many ((punctParser #"|" ++ literalParser) >> snd)) >> - (fn (h,t) => h :: t); - - val clauseParser = - ((punctParser #"(" ++ disjunctionParser ++ punctParser #")") >> - (fn ((),(c,())) => c)) || - disjunctionParser; - -(* - An exact transcription of the fof_formula syntax from - - TPTP-v3.2.0/Documents/SyntaxBNF, - - fun fofFormulaParser input = - (binaryFormulaParser || unitaryFormulaParser) input - - and binaryFormulaParser input = - (nonAssocBinaryFormulaParser || assocBinaryFormulaParser) input - - and nonAssocBinaryFormulaParser input = - ((unitaryFormulaParser ++ binaryConnectiveParser ++ - unitaryFormulaParser) >> - (fn (f,(c,g)) => c (f,g))) input - - and binaryConnectiveParser input = - ((symbolParser "<=>" >> K Formula.Iff) || - (symbolParser "=>" >> K Formula.Imp) || - (symbolParser "<=" >> K (fn (f,g) => Formula.Imp (g,f))) || - (symbolParser "<~>" >> K (Formula.Not o Formula.Iff)) || - (symbolParser "~|" >> K (Formula.Not o Formula.Or)) || - (symbolParser "~&" >> K (Formula.Not o Formula.And))) input - - and assocBinaryFormulaParser input = - (orFormulaParser || andFormulaParser) input - - and orFormulaParser input = - ((unitaryFormulaParser ++ - atLeastOne ((punctParser #"|" ++ unitaryFormulaParser) >> snd)) >> - (fn (f,fs) => Formula.listMkDisj (f :: fs))) input - - and andFormulaParser input = - ((unitaryFormulaParser ++ - atLeastOne ((punctParser #"&" ++ unitaryFormulaParser) >> snd)) >> - (fn (f,fs) => Formula.listMkConj (f :: fs))) input - - and unitaryFormulaParser input = - (quantifiedFormulaParser || - unaryFormulaParser || - ((punctParser #"(" ++ fofFormulaParser ++ punctParser #")") >> - (fn ((),(f,())) => f)) || - (atomParser >> - (fn Boolean b => Formula.mkBoolean b - | Literal l => Literal.toFormula l))) input - - and quantifiedFormulaParser input = - ((quantifierParser ++ varListParser ++ punctParser #":" ++ - unitaryFormulaParser) >> - (fn (q,(v,((),f))) => q (v,f))) input - - and quantifierParser input = - ((punctParser #"!" >> K Formula.listMkForall) || - (punctParser #"?" >> K Formula.listMkExists)) input - - and unaryFormulaParser input = - ((unaryConnectiveParser ++ unitaryFormulaParser) >> - (fn (c,f) => c f)) input - - and unaryConnectiveParser input = - (punctParser #"~" >> K Formula.Not) input; -*) - -(* - This version is supposed to be equivalent to the spec version above, - but uses closures to avoid reparsing prefixes. -*) - - fun fofFormulaParser input = - ((unitaryFormulaParser ++ optional binaryFormulaParser) >> - (fn (f,NONE) => f | (f, SOME t) => t f)) input - - and binaryFormulaParser input = - (nonAssocBinaryFormulaParser || assocBinaryFormulaParser) input - - and nonAssocBinaryFormulaParser input = - ((binaryConnectiveParser ++ unitaryFormulaParser) >> - (fn (c,g) => fn f => c (f,g))) input - - and binaryConnectiveParser input = - ((symbolParser "<=>" >> K Formula.Iff) || - (symbolParser "=>" >> K Formula.Imp) || - (symbolParser "<=" >> K (fn (f,g) => Formula.Imp (g,f))) || - (symbolParser "<~>" >> K (Formula.Not o Formula.Iff)) || - (symbolParser "~|" >> K (Formula.Not o Formula.Or)) || - (symbolParser "~&" >> K (Formula.Not o Formula.And))) input - - and assocBinaryFormulaParser input = - (orFormulaParser || andFormulaParser) input - - and orFormulaParser input = - (atLeastOne ((punctParser #"|" ++ unitaryFormulaParser) >> snd) >> - (fn fs => fn f => Formula.listMkDisj (f :: fs))) input - - and andFormulaParser input = - (atLeastOne ((punctParser #"&" ++ unitaryFormulaParser) >> snd) >> - (fn fs => fn f => Formula.listMkConj (f :: fs))) input - - and unitaryFormulaParser input = - (quantifiedFormulaParser || - unaryFormulaParser || - ((punctParser #"(" ++ fofFormulaParser ++ punctParser #")") >> - (fn ((),(f,())) => f)) || - (atomParser >> - (fn Boolean b => Formula.mkBoolean b - | Literal l => Literal.toFormula l))) input - - and quantifiedFormulaParser input = - ((quantifierParser ++ varListParser ++ punctParser #":" ++ - unitaryFormulaParser) >> - (fn (q,(v,((),f))) => q (v,f))) input - - and quantifierParser input = - ((punctParser #"!" >> K Formula.listMkForall) || - (punctParser #"?" >> K Formula.listMkExists)) input - - and unaryFormulaParser input = - ((unaryConnectiveParser ++ unitaryFormulaParser) >> - (fn (c,f) => c f)) input + fun mkVarName mapping v = varFromTptp mapping v; + + fun mkVar mapping v = + let + val v = mkVarName mapping v + in + Term.Var v + end + + fun mkFn mapping (f,tms) = + let + val f = fnFromTptp mapping (f, length tms) + in + Term.Fn (f,tms) + end; + + fun mkConst mapping c = mkFn mapping (c,[]); + + fun mkAtom mapping (r,tms) = + let + val r = relFromTptp mapping (r, length tms) + in + (r,tms) + end; + + fun termParser mapping input = + let + val fnP = functionArgumentsParser mapping >> mkFn mapping + val nonFnP = nonFunctionArgumentsTermParser mapping + in + fnP || nonFnP + end input + + and functionArgumentsParser mapping input = + let + val commaTmP = (punctParser #"," ++ termParser mapping) >> snd + in + (functionParser ++ punctParser #"(" ++ termParser mapping ++ + many commaTmP ++ punctParser #")") >> + (fn (f,((),(t,(ts,())))) => (f, t :: ts)) + end input + + and nonFunctionArgumentsTermParser mapping input = + let + val varP = varParser >> mkVar mapping + val constP = constantParser >> mkConst mapping + in + varP || constP + end input; + + fun binaryAtomParser mapping tm input = + let + val eqP = + (punctParser #"=" ++ termParser mapping) >> + (fn ((),r) => (true,("$equal",[tm,r]))) + + val neqP = + (symbolParser "!=" ++ termParser mapping) >> + (fn ((),r) => (false,("$equal",[tm,r]))) + in + eqP || neqP + end input; + + fun maybeBinaryAtomParser mapping (s,tms) input = + let + val tm = mkFn mapping (s,tms) + in + optional (binaryAtomParser mapping tm) >> + (fn SOME lit => lit + | NONE => (true,(s,tms))) + end input; + + fun literalAtomParser mapping input = + let + val fnP = + functionArgumentsParser mapping >>++ + maybeBinaryAtomParser mapping + + val nonFnP = + nonFunctionArgumentsTermParser mapping >>++ + binaryAtomParser mapping + + val propP = propositionParser >> (fn s => (true,(s,[]))) + in + fnP || nonFnP || propP + end input; + + fun atomParser mapping input = + let + fun mk (pol,rel) = + case rel of + ("$true",[]) => Boolean pol + | ("$false",[]) => Boolean (not pol) + | ("$equal",[l,r]) => Literal (pol, Atom.mkEq (l,r)) + | (r,tms) => Literal (pol, mkAtom mapping (r,tms)) + in + literalAtomParser mapping >> mk + end input; + + fun literalParser mapping input = + let + val negP = + (punctParser #"~" ++ atomParser mapping) >> + (negateLiteral o snd) + + val posP = atomParser mapping + in + negP || posP + end input; + + fun disjunctionParser mapping input = + let + val orLitP = (punctParser #"|" ++ literalParser mapping) >> snd + in + (literalParser mapping ++ many orLitP) >> (fn (h,t) => h :: t) + end input; + + fun clauseParser mapping input = + let + val disjP = disjunctionParser mapping + + val bracketDisjP = + (punctParser #"(" ++ disjP ++ punctParser #")") >> + (fn ((),(c,())) => c) + in + bracketDisjP || disjP + end input; + + val binaryConnectiveParser = + (symbolParser "<=>" >> K Formula.Iff) || + (symbolParser "=>" >> K Formula.Imp) || + (symbolParser "<=" >> K (fn (f,g) => Formula.Imp (g,f))) || + (symbolParser "<~>" >> K (Formula.Not o Formula.Iff)) || + (symbolParser "~|" >> K (Formula.Not o Formula.Or)) || + (symbolParser "~&" >> K (Formula.Not o Formula.And)); + + val quantifierParser = + (punctParser #"!" >> K Formula.listMkForall) || + (punctParser #"?" >> K Formula.listMkExists); + + fun fofFormulaParser mapping input = + let + fun mk (f,NONE) = f + | mk (f, SOME t) = t f + in + (unitaryFormulaParser mapping ++ + optional (binaryFormulaParser mapping)) >> mk + end input + + and binaryFormulaParser mapping input = + let + val nonAssocP = nonAssocBinaryFormulaParser mapping + + val assocP = assocBinaryFormulaParser mapping + in + nonAssocP || assocP + end input + + and nonAssocBinaryFormulaParser mapping input = + let + fun mk (c,g) f = c (f,g) + in + (binaryConnectiveParser ++ unitaryFormulaParser mapping) >> mk + end input + + and assocBinaryFormulaParser mapping input = + let + val orP = orFormulaParser mapping + + val andP = andFormulaParser mapping + in + orP || andP + end input + + and orFormulaParser mapping input = + let + val orFmP = (punctParser #"|" ++ unitaryFormulaParser mapping) >> snd + in + atLeastOne orFmP >> + (fn fs => fn f => Formula.listMkDisj (f :: fs)) + end input + + and andFormulaParser mapping input = + let + val andFmP = (punctParser #"&" ++ unitaryFormulaParser mapping) >> snd + in + atLeastOne andFmP >> + (fn fs => fn f => Formula.listMkConj (f :: fs)) + end input + + and unitaryFormulaParser mapping input = + let + val quantP = quantifiedFormulaParser mapping + + val unaryP = unaryFormulaParser mapping + + val brackP = + (punctParser #"(" ++ fofFormulaParser mapping ++ + punctParser #")") >> + (fn ((),(f,())) => f) + + val atomP = + atomParser mapping >> + (fn Boolean b => Formula.mkBoolean b + | Literal l => Literal.toFormula l) + in + quantP || + unaryP || + brackP || + atomP + end input + + and quantifiedFormulaParser mapping input = + let + fun mk (q,(vs,((),f))) = q (map (mkVarName mapping) vs, f) + in + (quantifierParser ++ varListParser ++ punctParser #":" ++ + unitaryFormulaParser mapping) >> mk + end input + + and unaryFormulaParser mapping input = + let + fun mk (c,f) = c f + in + (unaryConnectiveParser ++ unitaryFormulaParser mapping) >> mk + end input and unaryConnectiveParser input = (punctParser #"~" >> K Formula.Not) input; - val cnfParser = - (alphaNumParser "cnf" ++ punctParser #"(" ++ - nameParser ++ punctParser #"," ++ - roleParser ++ punctParser #"," ++ - clauseParser ++ punctParser #")" ++ - punctParser #".") >> - (fn ((),((),(n,((),(r,((),(c,((),())))))))) => - CnfFormula {name = n, role = r, clause = c}); - - val fofParser = - (alphaNumParser "fof" ++ punctParser #"(" ++ - nameParser ++ punctParser #"," ++ - roleParser ++ punctParser #"," ++ - fofFormulaParser ++ punctParser #")" ++ - punctParser #".") >> - (fn ((),((),(n,((),(r,((),(f,((),())))))))) => - FofFormula {name = n, role = r, formula = f}); - - val formulaParser = cnfParser || fofParser; + fun cnfParser mapping input = + let + fun mk ((),((),(name,((),(role,((),(cl,((),())))))))) = + let + val body = CnfFormulaBody cl + val source = NoFormulaSource + in + Formula + {name = name, + role = role, + body = body, + source = source} + end + in + (alphaNumParser "cnf" ++ punctParser #"(" ++ + nameParser ++ punctParser #"," ++ + roleParser ++ punctParser #"," ++ + clauseParser mapping ++ punctParser #")" ++ + punctParser #".") >> mk + end input; + + fun fofParser mapping input = + let + fun mk ((),((),(name,((),(role,((),(fm,((),())))))))) = + let + val body = FofFormulaBody fm + val source = NoFormulaSource + in + Formula + {name = name, + role = role, + body = body, + source = source} + end + in + (alphaNumParser "fof" ++ punctParser #"(" ++ + nameParser ++ punctParser #"," ++ + roleParser ++ punctParser #"," ++ + fofFormulaParser mapping ++ punctParser #")" ++ + punctParser #".") >> mk + end input; +in + fun formulaParser mapping input = + let + val cnfP = cnfParser mapping + + val fofP = fofParser mapping + in + cnfP || fofP + end input; +end; + +(* ------------------------------------------------------------------------- *) +(* Include declarations. *) +(* ------------------------------------------------------------------------- *) + +fun ppInclude i = + Print.blockProgram Print.Inconsistent 2 + [Print.addString "include('", + Print.addString i, + Print.addString "')."]; + +val includeToString = Print.toString ppInclude; + +local + open Parse; + + infixr 9 >>++ + infixr 8 ++ + infixr 7 >> + infixr 6 || + + val filenameParser = maybe (fn Quote s => SOME s | _ => NONE); +in + val includeParser = + (some (equal (AlphaNum "include")) ++ + some (equal (Punct #"(")) ++ + filenameParser ++ + some (equal (Punct #")")) ++ + some (equal (Punct #"."))) >> + (fn (_,(_,(f,(_,_)))) => f); +end; + +(* ------------------------------------------------------------------------- *) +(* Parsing TPTP files. *) +(* ------------------------------------------------------------------------- *) + +datatype declaration = + IncludeDeclaration of string + | FormulaDeclaration of formula; + +val partitionDeclarations = + let + fun part (d,(il,fl)) = + case d of + IncludeDeclaration i => (i :: il, fl) + | FormulaDeclaration f => (il, f :: fl) + in + fn l => List.foldl part ([],[]) (rev l) + end; + +local + open Parse; + + infixr 9 >>++ + infixr 8 ++ + infixr 7 >> + infixr 6 || + + fun declarationParser mapping = + (includeParser >> IncludeDeclaration) || + (formulaParser mapping >> FormulaDeclaration); fun parseChars parser chars = let - val tokens = Parser.everything (lexer >> singleton) chars - in - Parser.everything (parser >> singleton) tokens - end; - - fun canParseString parser s = - let - val chars = Stream.fromString s + val tokens = Parse.everything (lexer >> singleton) chars in - case Stream.toList (parseChars parser chars) of - [_] => true - | _ => false - end - handle NoParse => false; -in - val parseFormula = parseChars formulaParser; - - val isTptpRelation = canParseString functionParser - and isTptpProposition = canParseString propositionParser - and isTptpFunction = canParseString functionParser - and isTptpConstant = canParseString constantParser; -end; - -fun formulaFromString s = - case Stream.toList (parseFormula (Stream.fromList (explode s))) of - [fm] => fm - | _ => raise Parser.NoParse; - -local - local - fun explodeAlpha s = List.filter Char.isAlpha (explode s); - in - fun normTptpName s n = - case explodeAlpha n of - [] => s - | c :: cs => implode (Char.toLower c :: cs); - - fun normTptpVar s n = - case explodeAlpha n of - [] => s - | c :: cs => implode (Char.toUpper c :: cs); - end; - - fun normTptpFunc (n,0) = if isTptpConstant n then n else normTptpName "c" n - | normTptpFunc (n,_) = if isTptpFunction n then n else normTptpName "f" n; - - fun normTptpRel (n,0) = if isTptpProposition n then n else normTptpName "p" n - | normTptpRel (n,_) = if isTptpRelation n then n else normTptpName "r" n; - - fun mkMap set norm mapping = - let - val mapping = mappingToTptp mapping - - fun mk (n_r,(a,m)) = - case NameArityMap.peek mapping n_r of - SOME t => (a, NameArityMap.insert m (n_r,t)) - | NONE => - let - val t = norm n_r - val (n,_) = n_r - val t = if t = n then n else Term.variantNum a t - in - (NameSet.add a t, NameArityMap.insert m (n_r,t)) - end - - val avoid = - let - fun mk ((n,r),s) = - let - val n = Option.getOpt (NameArityMap.peek mapping (n,r), n) - in - NameSet.add s n - end - in - NameAritySet.foldl mk NameSet.empty set - end - in - snd (NameAritySet.foldl mk (avoid, NameArityMap.new ()) set) + Parse.everything (parser >> singleton) tokens end; - - fun mkTptpVar a v = Term.variantNum a (normTptpVar "V" v); - - fun isTptpVar v = mkTptpVar NameSet.empty v = v; - - fun alphaFormula fm = - let - fun addVar v a s = - let - val v' = mkTptpVar a v - val a = NameSet.add a v' - and s = if v = v' then s else Subst.insert s (v, Term.Var v') - in - (v',(a,s)) - end - - fun initVar (v,(a,s)) = snd (addVar v a s) - - open Formula - - fun alpha _ _ True = True - | alpha _ _ False = False - | alpha _ s (Atom atm) = Atom (Atom.subst s atm) - | alpha a s (Not p) = Not (alpha a s p) - | alpha a s (And (p,q)) = And (alpha a s p, alpha a s q) - | alpha a s (Or (p,q)) = Or (alpha a s p, alpha a s q) - | alpha a s (Imp (p,q)) = Imp (alpha a s p, alpha a s q) - | alpha a s (Iff (p,q)) = Iff (alpha a s p, alpha a s q) - | alpha a s (Forall (v,p)) = - let val (v,(a,s)) = addVar v a s in Forall (v, alpha a s p) end - | alpha a s (Exists (v,p)) = - let val (v,(a,s)) = addVar v a s in Exists (v, alpha a s p) end - - val fvs = formulaFreeVars fm - val (avoid,fvs) = NameSet.partition isTptpVar fvs - val (avoid,sub) = NameSet.foldl initVar (avoid,Subst.empty) fvs -(*TRACE5 - val () = Parser.ppTrace Subst.pp "Tptp.alpha: sub" sub -*) - in - case fm of - CnfFormula {name,role,clause} => - CnfFormula {name = name, role = role, clause = clauseSubst sub clause} - | FofFormula {name,role,formula} => - FofFormula {name = name, role = role, formula = alpha avoid sub formula} - end; - - fun formulaToTptp maps fm = alphaFormula (mapFormula maps fm); in - fun formulasToTptp formulas = - let - val funcs = formulasFunctions formulas - and rels = formulasRelations formulas - - val functionMap = mkMap funcs normTptpFunc (!functionMapping) - and relationMap = mkMap rels normTptpRel (!relationMapping) - - val maps = {functionMap = functionMap, relationMap = relationMap} - in - map (formulaToTptp maps) formulas - end; + fun parseDeclaration mapping = parseChars (declarationParser mapping); end; -fun formulasFromTptp formulas = +(* ------------------------------------------------------------------------- *) +(* Clause information. *) +(* ------------------------------------------------------------------------- *) + +datatype clauseSource = + CnfClauseSource of formulaName * literal list + | FofClauseSource of Normalize.thm; + +type 'a clauseInfo = 'a LiteralSetMap.map; + +type clauseNames = formulaName clauseInfo; + +type clauseRoles = role clauseInfo; + +type clauseSources = clauseSource clauseInfo; + +val noClauseNames : clauseNames = LiteralSetMap.new (); + +val allClauseNames : clauseNames -> formulaNameSet = let - val functionMap = mappingFromTptp (!functionMapping) - and relationMap = mappingFromTptp (!relationMapping) - - val maps = {functionMap = functionMap, relationMap = relationMap} + fun add (_,n,s) = addFormulaNameSet s n in - map (mapFormula maps) formulas + LiteralSetMap.foldl add emptyFormulaNameSet end; -local - fun ppGen ppX pp (gen,name,role,x) = - (Parser.beginBlock pp Parser.Inconsistent (size gen + 1); - Parser.addString pp (gen ^ "(" ^ name ^ ","); - Parser.addBreak pp (1,0); - Parser.addString pp (role ^ ","); - Parser.addBreak pp (1,0); - Parser.beginBlock pp Parser.Consistent 1; - Parser.addString pp "("; - ppX pp x; - Parser.addString pp ")"; - Parser.endBlock pp; - Parser.addString pp ")."; - Parser.endBlock pp); -in - fun ppFormula pp (CnfFormula {name,role,clause}) = - ppGen ppClause pp ("cnf",name,role,clause) - | ppFormula pp (FofFormula {name,role,formula}) = - ppGen ppFof pp ("fof",name,role,formula); -end; - -val formulaToString = Parser.toString ppFormula; +val noClauseRoles : clauseRoles = LiteralSetMap.new (); + +val noClauseSources : clauseSources = LiteralSetMap.new (); + +(* ------------------------------------------------------------------------- *) +(* Comments. *) +(* ------------------------------------------------------------------------- *) + +fun mkLineComment "" = "%" + | mkLineComment line = "% " ^ line; + +fun destLineComment cs = + case cs of + [] => "" + | #"%" :: #" " :: rest => implode rest + | #"%" :: rest => implode rest + | _ => raise Error "Tptp.destLineComment"; + +val isLineComment = can destLineComment; (* ------------------------------------------------------------------------- *) (* TPTP problems. *) (* ------------------------------------------------------------------------- *) -datatype goal = - Cnf of Problem.problem - | Fof of Formula.formula; - -type problem = {comments : string list, formulas : formula list}; +type comments = string list; + +type includes = string list; + +datatype problem = + Problem of + {comments : comments, + includes : includes, + formulas : formula list}; + +fun hasCnfConjecture (Problem {formulas,...}) = + List.exists isCnfConjectureFormula formulas; + +fun hasFofConjecture (Problem {formulas,...}) = + List.exists isFofConjectureFormula formulas; + +fun hasConjecture (Problem {formulas,...}) = + List.exists isConjectureFormula formulas; + +fun freeVars (Problem {formulas,...}) = freeVarsListFormula formulas; local - fun stripComments acc strm = - case strm of - Stream.NIL => (rev acc, Stream.NIL) - | Stream.CONS (line,rest) => - case total destComment line of - SOME s => stripComments (s :: acc) (rest ()) - | NONE => (rev acc, Stream.filter (not o isComment) strm); -in - fun read {filename} = + fun bump n avoid = + let + val s = FormulaName (Int.toString n) + in + if memberFormulaNameSet s avoid then bump (n + 1) avoid + else (s, n, addFormulaNameSet avoid s) + end; + + fun fromClause defaultRole names roles cl (n,avoid) = let - val lines = Stream.fromTextFile {filename = filename} - - val lines = Stream.map chomp lines - - val (comments,lines) = stripComments [] lines - - val chars = Stream.concat (Stream.map Stream.fromString lines) - - val formulas = Stream.toList (parseFormula chars) - - val formulas = formulasFromTptp formulas + val (name,n,avoid) = + case LiteralSetMap.peek names cl of + SOME name => (name,n,avoid) + | NONE => bump n avoid + + val role = Option.getOpt (LiteralSetMap.peek roles cl, defaultRole) + + val body = CnfFormulaBody (clauseFromLiteralSet cl) + + val source = NoFormulaSource + + val formula = + Formula + {name = name, + role = role, + body = body, + source = source} in - {comments = comments, formulas = formulas} + (formula,(n,avoid)) + end; +in + fun mkProblem {comments,includes,names,roles,problem} = + let + fun fromCl defaultRole = fromClause defaultRole names roles + + val {axioms,conjecture} = problem + + val n_avoid = (0, allClauseNames names) + + val (axiomFormulas,n_avoid) = maps (fromCl AxiomRole) axioms n_avoid + + val (conjectureFormulas,_) = + maps (fromCl NegatedConjectureRole) conjecture n_avoid + + val formulas = axiomFormulas @ conjectureFormulas + in + Problem + {comments = comments, + includes = includes, + formulas = formulas} end; end; -(* Quick testing -installPP Term.pp; -installPP Formula.pp; -val () = Term.isVarName := (fn s => Char.isUpper (String.sub (s,0))); -val TPTP_DIR = "/Users/Joe/ptr/tptp/tptp/"; -val num1 = read {filename = TPTP_DIR ^ "NUM/NUM001-1.tptp"}; -val lcl9 = read {filename = TPTP_DIR ^ "LCL/LCL009-1.tptp"}; -val set11 = read {filename = TPTP_DIR ^ "SET/SET011+3.tptp"}; -val swc128 = read {filename = TPTP_DIR ^ "SWC/SWC128-1.tptp"}; -*) +type normalization = + {problem : Problem.problem, + sources : clauseSources}; + +val initialNormalization : normalization = + {problem = {axioms = [], conjecture = []}, + sources = noClauseSources}; + +datatype problemGoal = + NoGoal + | CnfGoal of (formulaName * clause) list + | FofGoal of (formulaName * Formula.formula) list; local - fun mkCommentLine comment = mkComment comment ^ "\n"; - - fun formulaStream [] () = Stream.NIL - | formulaStream (h :: t) () = - Stream.CONS ("\n" ^ formulaToString h, formulaStream t); -in - fun write {filename} {comments,formulas} = - Stream.toTextFile - {filename = filename} - (Stream.append - (Stream.map mkCommentLine (Stream.fromList comments)) - (formulaStream (formulasToTptp formulas))); -end; - -(* ------------------------------------------------------------------------- *) -(* Converting TPTP problems to goal formulas. *) -(* ------------------------------------------------------------------------- *) - -fun isCnfProblem ({formulas,...} : problem) = - let - val cnf = List.exists isCnfFormula formulas - and fof = List.exists isFofFormula formulas - in - case (cnf,fof) of - (false,false) => raise Error "TPTP problem has no formulas" - | (true,true) => raise Error "TPTP problem has both cnf and fof formulas" - | (cnf,_) => cnf - end; - -fun hasConjecture ({formulas,...} : problem) = - List.exists formulaIsConjecture formulas; - -local - fun cnfFormulaToClause (CnfFormula {clause,...}) = - if mem (Boolean true) clause then NONE + fun partitionFormula (formula,(cnfAxioms,fofAxioms,cnfGoals,fofGoals)) = + let + val Formula {name,role,body,...} = formula + in + case body of + CnfFormulaBody cl => + if isCnfConjectureRole role then + let + val cnfGoals = (name,cl) :: cnfGoals + in + (cnfAxioms,fofAxioms,cnfGoals,fofGoals) + end + else + let + val cnfAxioms = (name,cl) :: cnfAxioms + in + (cnfAxioms,fofAxioms,cnfGoals,fofGoals) + end + | FofFormulaBody fm => + if isFofConjectureRole role then + let + val fofGoals = (name,fm) :: fofGoals + in + (cnfAxioms,fofAxioms,cnfGoals,fofGoals) + end + else + let + val fofAxioms = (name,fm) :: fofAxioms + in + (cnfAxioms,fofAxioms,cnfGoals,fofGoals) + end + end; + + fun partitionFormulas fms = + let + val (cnfAxioms,fofAxioms,cnfGoals,fofGoals) = + List.foldl partitionFormula ([],[],[],[]) fms + + val goal = + case (rev cnfGoals, rev fofGoals) of + ([],[]) => NoGoal + | (cnfGoals,[]) => CnfGoal cnfGoals + | ([],fofGoals) => FofGoal fofGoals + | (_ :: _, _ :: _) => + raise Error "TPTP problem has both cnf and fof conjecture formulas" + in + {cnfAxioms = rev cnfAxioms, + fofAxioms = rev fofAxioms, + goal = goal} + end; + + fun addClauses role clauses acc : normalization = + let + fun addClause (cl_src,sources) = + LiteralSetMap.insert sources cl_src + + val {problem,sources} : normalization = acc + val {axioms,conjecture} = problem + + val cls = map fst clauses + val (axioms,conjecture) = + if isCnfConjectureRole role then (axioms, cls @ conjecture) + else (cls @ axioms, conjecture) + + val problem = {axioms = axioms, conjecture = conjecture} + and sources = List.foldl addClause sources clauses + in + {problem = problem, + sources = sources} + end; + + fun addCnf role ((name,clause),(norm,cnf)) = + if List.exists (equalBooleanLiteral true) clause then (norm,cnf) else let - val lits = List.mapPartial (total destLiteral) clause + val cl = List.mapPartial (total destLiteral) clause + val cl = LiteralSet.fromList cl + + val src = CnfClauseSource (name,clause) + + val norm = addClauses role [(cl,src)] norm in - SOME (LiteralSet.fromList lits) - end - | cnfFormulaToClause (FofFormula _) = raise Bug "cnfFormulaToClause"; - - fun fofFormulaToGoal (FofFormula {formula,role,...}, {axioms,goals}) = + (norm,cnf) + end; + + val addCnfAxiom = addCnf AxiomRole; + + val addCnfGoal = addCnf NegatedConjectureRole; + + fun addFof role (th,(norm,cnf)) = let - val fm = Formula.generalize formula + fun sourcify (cl,inf) = (cl, FofClauseSource inf) + + val (clauses,cnf) = Normalize.addCnf th cnf + val clauses = map sourcify clauses + val norm = addClauses role clauses norm + in + (norm,cnf) + end; + + fun addFofAxiom ((_,fm),acc) = + addFof AxiomRole (Normalize.mkAxiom fm, acc); + + fun normProblem subgoal (norm,_) = + let + val {problem,sources} = norm + val {axioms,conjecture} = problem + val problem = {axioms = rev axioms, conjecture = rev conjecture} in - if role = ROLE_CONJECTURE then - {axioms = axioms, goals = fm :: goals} - else - {axioms = fm :: axioms, goals = goals} - end - | fofFormulaToGoal (CnfFormula _, _) = raise Bug "fofFormulaToGoal"; -in - fun toGoal (prob as {formulas,...}) = - if isCnfProblem prob then - Cnf (List.mapPartial cnfFormulaToClause formulas) - else - Fof - let - val axioms_goals = {axioms = [], goals = []} - val axioms_goals = List.foldl fofFormulaToGoal axioms_goals formulas - in - case axioms_goals of - {axioms, goals = []} => - Formula.Imp (Formula.listMkConj axioms, Formula.False) - | {axioms = [], goals} => Formula.listMkConj goals - | {axioms,goals} => - Formula.Imp (Formula.listMkConj axioms, Formula.listMkConj goals) - end; -end; - -local - fun fromClause cl n = + {subgoal = subgoal, + problem = problem, + sources = sources} + end; + + val normProblemFalse = normProblem (Formula.False,[]); + + fun splitProblem acc = let - val name = "clause_" ^ Int.toString n - val role = ROLE_NEGATED_CONJECTURE - val clause = clauseFromLiteralSet cl + fun mk parents subgoal = + let + val subgoal = Formula.generalize subgoal + + val th = Normalize.mkAxiom (Formula.Not subgoal) + + val acc = addFof NegatedConjectureRole (th,acc) + in + normProblem (subgoal,parents) acc + end + + fun split (name,goal) = + let + val subgoals = Formula.splitGoal goal + val subgoals = + if null subgoals then [Formula.True] else subgoals + + val parents = [name] + in + map (mk parents) subgoals + end in - (CnfFormula {name = name, role = role, clause = clause}, n + 1) + fn goals => List.concat (map split goals) + end; + + fun clausesToGoal cls = + let + val cls = map (Formula.generalize o clauseToFormula o snd) cls + in + Formula.listMkConj cls + end; + + fun formulasToGoal fms = + let + val fms = map (Formula.generalize o snd) fms + in + Formula.listMkConj fms end; in - fun fromProblem prob = + fun goal (Problem {formulas,...}) = let - val comments = [] - and (formulas,_) = maps fromClause prob 0 + val {cnfAxioms,fofAxioms,goal} = partitionFormulas formulas + + val fm = + case goal of + NoGoal => Formula.False + | CnfGoal cls => Formula.Imp (clausesToGoal cls, Formula.False) + | FofGoal goals => formulasToGoal goals + + val fm = + if null fofAxioms then fm + else Formula.Imp (formulasToGoal fofAxioms, fm) + + val fm = + if null cnfAxioms then fm + else Formula.Imp (clausesToGoal cnfAxioms, fm) in - {comments = comments, formulas = formulas} + fm + end; + + fun normalize (Problem {formulas,...}) = + let + val {cnfAxioms,fofAxioms,goal} = partitionFormulas formulas + + val acc = (initialNormalization, Normalize.initialCnf) + val acc = List.foldl addCnfAxiom acc cnfAxioms + val acc = List.foldl addFofAxiom acc fofAxioms + in + case goal of + NoGoal => [normProblemFalse acc] + | CnfGoal cls => [normProblemFalse (List.foldl addCnfGoal acc cls)] + | FofGoal goals => splitProblem acc goals end; end; local - fun refute cls = + datatype blockComment = + OutsideBlockComment + | EnteringBlockComment + | InsideBlockComment + | LeavingBlockComment; + + fun stripLineComments acc strm = + case strm of + Stream.Nil => (rev acc, Stream.Nil) + | Stream.Cons (line,rest) => + case total destLineComment line of + SOME s => stripLineComments (s :: acc) (rest ()) + | NONE => (rev acc, Stream.filter (not o isLineComment) strm); + + fun advanceBlockComment c state = + case state of + OutsideBlockComment => + if c = #"/" then (Stream.Nil, EnteringBlockComment) + else (Stream.singleton c, OutsideBlockComment) + | EnteringBlockComment => + if c = #"*" then (Stream.Nil, InsideBlockComment) + else if c = #"/" then (Stream.singleton #"/", EnteringBlockComment) + else (Stream.fromList [#"/",c], OutsideBlockComment) + | InsideBlockComment => + if c = #"*" then (Stream.Nil, LeavingBlockComment) + else (Stream.Nil, InsideBlockComment) + | LeavingBlockComment => + if c = #"/" then (Stream.Nil, OutsideBlockComment) + else if c = #"*" then (Stream.Nil, LeavingBlockComment) + else (Stream.Nil, InsideBlockComment); + + fun eofBlockComment state = + case state of + OutsideBlockComment => Stream.Nil + | EnteringBlockComment => Stream.singleton #"/" + | _ => raise Error "EOF inside a block comment"; + + val stripBlockComments = + Stream.mapsConcat advanceBlockComment eofBlockComment + OutsideBlockComment; +in + fun read {mapping,filename} = let - val res = Resolution.new Resolution.default (map Thm.axiom cls) + (* Estimating parse error line numbers *) + + val lines = Stream.fromTextFile {filename = filename} + + val {chars,parseErrorLocation} = Parse.initialize {lines = lines} in - case Resolution.loop res of + (let + (* The character stream *) + + val (comments,chars) = stripLineComments [] chars + + val chars = Parse.everything Parse.any chars + + val chars = stripBlockComments chars + + (* The declaration stream *) + + val declarations = Stream.toList (parseDeclaration mapping chars) + + val (includes,formulas) = partitionDeclarations declarations + in + Problem + {comments = comments, + includes = includes, + formulas = formulas} + end + handle Parse.NoParse => raise Error "parse error") + handle Error err => + raise Error ("error in TPTP file \"" ^ filename ^ "\" " ^ + parseErrorLocation () ^ "\n" ^ err) + end; +end; + +local + val newline = Stream.singleton "\n"; + + fun spacer top = if top then Stream.Nil else newline; + + fun mkComment comment = mkLineComment comment ^ "\n"; + + fun mkInclude inc = includeToString inc ^ "\n"; + + fun formulaStream _ _ [] = Stream.Nil + | formulaStream mapping top (h :: t) = + Stream.append + (Stream.concatList + [spacer top, + Stream.singleton (formulaToString mapping h), + newline]) + (fn () => formulaStream mapping false t); +in + fun write {problem,mapping,filename} = + let + val Problem {comments,includes,formulas} = problem + + val includesTop = null comments + val formulasTop = includesTop andalso null includes + in + Stream.toTextFile + {filename = filename} + (Stream.concatList + [Stream.map mkComment (Stream.fromList comments), + spacer includesTop, + Stream.map mkInclude (Stream.fromList includes), + formulaStream mapping formulasTop formulas]) + end; +end; + +local + fun refute {axioms,conjecture} = + let + val axioms = map Thm.axiom axioms + and conjecture = map Thm.axiom conjecture + val problem = {axioms = axioms, conjecture = conjecture} + val resolution = Resolution.new Resolution.default problem + in + case Resolution.loop resolution of Resolution.Contradiction _ => true | Resolution.Satisfiable _ => false end; in fun prove filename = let - val tptp = read filename - val problems = - case toGoal tptp of - Cnf prob => [prob] - | Fof goal => Problem.fromGoal goal + val problem = read filename + val problems = map #problem (normalize problem) in List.all refute problems end; @@ -1014,132 +2176,384 @@ (* ------------------------------------------------------------------------- *) local - fun ppAtomInfo pp atm = - case total Atom.destEq atm of - SOME (a,b) => ppAtom pp ("$equal",[a,b]) - | NONE => ppAtom pp atm; - - fun ppLiteralInfo pp (pol,atm) = - if pol then ppAtomInfo pp atm - else - (Parser.beginBlock pp Parser.Inconsistent 2; - Parser.addString pp "~"; - Parser.addBreak pp (1,0); - ppAtomInfo pp atm; - Parser.endBlock pp); - - val ppAssumeInfo = Parser.ppBracket "(" ")" ppAtomInfo; - - val ppSubstInfo = - Parser.ppMap - Subst.toList - (Parser.ppSequence "," - (Parser.ppBracket "[" "]" - (Parser.ppBinop "," ppVar (Parser.ppBracket "(" ")" ppTerm)))); - - val ppResolveInfo = Parser.ppBracket "(" ")" ppAtomInfo; - - val ppReflInfo = Parser.ppBracket "(" ")" ppTerm; - - fun ppEqualityInfo pp (lit,path,res) = - (Parser.ppBracket "(" ")" ppLiteralInfo pp lit; - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Term.ppPath pp path; - Parser.addString pp ","; - Parser.addBreak pp (1,0); - Parser.ppBracket "(" ")" ppTerm pp res); - - fun ppInfInfo pp inf = + fun newName avoid prefix = + let + fun bump i = + let + val name = FormulaName (prefix ^ Int.toString i) + val i = i + 1 + in + if memberFormulaNameSet name avoid then bump i else (name,i) + end + in + bump + end; + + fun lookupClauseSource sources cl = + case LiteralSetMap.peek sources cl of + SOME src => src + | NONE => raise Bug "Tptp.lookupClauseSource"; + + fun lookupFormulaName fmNames fm = + case FormulaMap.peek fmNames fm of + SOME name => name + | NONE => raise Bug "Tptp.lookupFormulaName"; + + fun lookupClauseName clNames cl = + case LiteralSetMap.peek clNames cl of + SOME name => name + | NONE => raise Bug "Tptp.lookupClauseName"; + + fun lookupClauseSourceName sources fmNames cl = + case lookupClauseSource sources cl of + CnfClauseSource (name,_) => name + | FofClauseSource th => + let + val (fm,_) = Normalize.destThm th + in + lookupFormulaName fmNames fm + end; + + fun collectProofDeps sources ((_,inf),names_ths) = case inf of - Proof.Axiom _ => raise Bug "ppInfInfo" - | Proof.Assume atm => ppAssumeInfo pp atm - | Proof.Subst (sub,_) => ppSubstInfo pp sub - | Proof.Resolve (res,_,_) => ppResolveInfo pp res - | Proof.Refl tm => ppReflInfo pp tm - | Proof.Equality x => ppEqualityInfo pp x; -in - fun ppProof p prf = - let - fun thmString n = Int.toString n - - val prf = enumerate prf - - fun ppThm p th = + Proof.Axiom cl => + let + val (names,ths) = names_ths + in + case lookupClauseSource sources cl of + CnfClauseSource (name,_) => let - val cl = Thm.clause th - - fun pred (_,(th',_)) = LiteralSet.equal (Thm.clause th') cl + val names = addFormulaNameSet names name in - case List.find pred prf of - NONE => Parser.addString p "(?)" - | SOME (n,_) => Parser.addString p (thmString n) + (names,ths) + end + | FofClauseSource th => + let + val ths = th :: ths + in + (names,ths) end - - fun ppInf p inf = - let - val name = Thm.inferenceTypeToString (Proof.inferenceType inf) - val name = String.map Char.toLower name - in - Parser.addString p (name ^ ","); - Parser.addBreak p (1,0); - Parser.ppBracket "[" "]" ppInfInfo p inf; - case Proof.parents inf of - [] => () - | ths => - (Parser.addString p ","; - Parser.addBreak p (1,0); - Parser.ppList ppThm p ths) - end - - fun ppTaut p inf = - (Parser.addString p "tautology,"; - Parser.addBreak p (1,0); - Parser.ppBracket "[" "]" ppInf p inf) - - fun ppStepInfo p (n,(th,inf)) = - let - val is_axiom = case inf of Proof.Axiom _ => true | _ => false - val name = thmString n - val role = - if is_axiom then "axiom" - else if Thm.isContradiction th then "theorem" - else "plain" - val cl = clauseFromThm th - in - Parser.addString p (name ^ ","); - Parser.addBreak p (1,0); - Parser.addString p (role ^ ","); - Parser.addBreak p (1,0); - Parser.ppBracket "(" ")" ppClause p cl; - if is_axiom then () - else - let - val is_tautology = null (Proof.parents inf) - in - Parser.addString p ","; - Parser.addBreak p (1,0); - if is_tautology then - Parser.ppBracket "introduced(" ")" ppTaut p inf - else - Parser.ppBracket "inference(" ")" ppInf p inf - end - end - - fun ppStep p step = - (Parser.ppBracket "cnf(" ")" ppStepInfo p step; - Parser.addString p "."; - Parser.addNewline p) + end + | _ => names_ths; + + fun collectNormalizeDeps ((_,inf,_),fofs_defs) = + case inf of + Normalize.Axiom fm => + let + val (fofs,defs) = fofs_defs + val fofs = FormulaSet.add fofs fm + in + (fofs,defs) + end + | Normalize.Definition n_d => + let + val (fofs,defs) = fofs_defs + val defs = StringMap.insert defs n_d + in + (fofs,defs) + end + | _ => fofs_defs; + + fun collectSubgoalProofDeps subgoalProof (names,fofs,defs) = + let + val {subgoal,sources,refutation} = subgoalProof + + val names = addListFormulaNameSet names (snd subgoal) + + val proof = Proof.proof refutation + + val (names,ths) = + List.foldl (collectProofDeps sources) (names,[]) proof + + val normalization = Normalize.proveThms (rev ths) + + val (fofs,defs) = + List.foldl collectNormalizeDeps (fofs,defs) normalization + + val subgoalProof = + {subgoal = subgoal, + normalization = normalization, + sources = sources, + proof = proof} + in + (subgoalProof,(names,fofs,defs)) + end; + + fun addProblemFormula names fofs (formula,(avoid,formulas,fmNames)) = + let + val name = nameFormula formula + + val avoid = addFormulaNameSet avoid name + + val (formulas,fmNames) = + if memberFormulaNameSet name names then + (formula :: formulas, fmNames) + else + case bodyFormula formula of + CnfFormulaBody _ => (formulas,fmNames) + | FofFormulaBody fm => + if not (FormulaSet.member fm fofs) then (formulas,fmNames) + else (formula :: formulas, FormulaMap.insert fmNames (fm,name)) + in + (avoid,formulas,fmNames) + end; + + fun addDefinitionFormula avoid (_,def,(formulas,i,fmNames)) = + let + val (name,i) = newName avoid "definition_" i + + val role = DefinitionRole + + val body = FofFormulaBody def + + val source = NoFormulaSource + + val formula = + Formula + {name = name, + role = role, + body = body, + source = source} + + val formulas = formula :: formulas + + val fmNames = FormulaMap.insert fmNames (def,name) + in + (formulas,i,fmNames) + end; + + fun addSubgoalFormula avoid subgoalProof (formulas,i) = + let + val {subgoal,normalization,sources,proof} = subgoalProof + + val (fm,pars) = subgoal + + val (name,i) = newName avoid "subgoal_" i + + val number = i - 1 + + val (subgoal,formulas) = + if null pars then (NONE,formulas) + else + let + val role = PlainRole + + val body = FofFormulaBody fm + + val source = + StripFormulaSource + {inference = "strip", + parents = pars} + + val formula = + Formula + {name = name, + role = role, + body = body, + source = source} + in + (SOME (name,fm), formula :: formulas) + end + + val subgoalProof = + {number = number, + subgoal = subgoal, + normalization = normalization, + sources = sources, + proof = proof} + in + (subgoalProof,(formulas,i)) + end; + + fun mkNormalizeFormulaSource fmNames inference fms = + let + val fms = + case inference of + Normalize.Axiom fm => fm :: fms + | Normalize.Definition (_,fm) => fm :: fms + | _ => fms + + val parents = map (lookupFormulaName fmNames) fms + in + NormalizeFormulaSource + {inference = inference, + parents = parents} + end; + + fun mkProofFormulaSource sources fmNames clNames inference = + let + val parents = + case inference of + Proof.Axiom cl => [lookupClauseSourceName sources fmNames cl] + | _ => + let + val cls = map Thm.clause (Proof.parents inference) + in + map (lookupClauseName clNames) cls + end in - Parser.beginBlock p Parser.Consistent 0; - app (ppStep p) prf; - Parser.endBlock p + ProofFormulaSource + {inference = inference, + parents = parents} + end; + + fun addNormalizeFormula avoid prefix ((fm,inf,fms),acc) = + let + val (formulas,i,fmNames) = acc + + val (name,i) = newName avoid prefix i + + val role = PlainRole + + val body = FofFormulaBody fm + + val source = mkNormalizeFormulaSource fmNames inf fms + + val formula = + Formula + {name = name, + role = role, + body = body, + source = source} + + val formulas = formula :: formulas + + val fmNames = FormulaMap.insert fmNames (fm,name) + in + (formulas,i,fmNames) + end; + + fun isSameClause sources formulas inf = + case inf of + Proof.Axiom cl => + (case lookupClauseSource sources cl of + CnfClauseSource (name,lits) => + if List.exists isBooleanLiteral lits then NONE + else SOME name + | _ => NONE) + | _ => NONE; + + fun addProofFormula avoid sources fmNames prefix ((th,inf),acc) = + let + val (formulas,i,clNames) = acc + + val cl = Thm.clause th + in + case isSameClause sources formulas inf of + SOME name => + let + val clNames = LiteralSetMap.insert clNames (cl,name) + in + (formulas,i,clNames) + end + | NONE => + let + val (name,i) = newName avoid prefix i + + val role = PlainRole + + val body = CnfFormulaBody (clauseFromLiteralSet cl) + + val source = mkProofFormulaSource sources fmNames clNames inf + + val formula = + Formula + {name = name, + role = role, + body = body, + source = source} + + val formulas = formula :: formulas + + val clNames = LiteralSetMap.insert clNames (cl,name) + in + (formulas,i,clNames) + end + end; + + fun addSubgoalProofFormulas avoid fmNames (subgoalProof,formulas) = + let + val {number,subgoal,normalization,sources,proof} = subgoalProof + + val (formulas,fmNames) = + case subgoal of + NONE => (formulas,fmNames) + | SOME (name,fm) => + let + val source = + StripFormulaSource + {inference = "negate", + parents = [name]} + + val prefix = "negate_" ^ Int.toString number ^ "_" + + val (name,_) = newName avoid prefix 0 + + val role = PlainRole + + val fm = Formula.Not fm + + val body = FofFormulaBody fm + + val formula = + Formula + {name = name, + role = role, + body = body, + source = source} + + val formulas = formula :: formulas + + val fmNames = FormulaMap.insert fmNames (fm,name) + in + (formulas,fmNames) + end + + val prefix = "normalize_" ^ Int.toString number ^ "_" + val (formulas,_,fmNames) = + List.foldl (addNormalizeFormula avoid prefix) + (formulas,0,fmNames) normalization + + val prefix = "refute_" ^ Int.toString number ^ "_" + val clNames : formulaName LiteralSetMap.map = LiteralSetMap.new () + val (formulas,_,_) = + List.foldl (addProofFormula avoid sources fmNames prefix) + (formulas,0,clNames) proof + in + formulas + end; +in + fun fromProof {problem,proofs} = + let + val names = emptyFormulaNameSet + and fofs = FormulaSet.empty + and defs : Formula.formula StringMap.map = StringMap.new () + + val (proofs,(names,fofs,defs)) = + maps collectSubgoalProofDeps proofs (names,fofs,defs) + + val Problem {formulas,...} = problem + + val fmNames : formulaName FormulaMap.map = FormulaMap.new () + val (avoid,formulas,fmNames) = + List.foldl (addProblemFormula names fofs) + (emptyFormulaNameSet,[],fmNames) formulas + + val (formulas,_,fmNames) = + StringMap.foldl (addDefinitionFormula avoid) + (formulas,0,fmNames) defs + + val (proofs,(formulas,_)) = + maps (addSubgoalFormula avoid) proofs (formulas,0) + + val formulas = + List.foldl (addSubgoalProofFormulas avoid fmNames) formulas proofs + in + rev formulas end -(*DEBUG - handle Error err => raise Bug ("Tptp.ppProof: shouldn't fail:\n" ^ err); +(*MetisDebug + handle Error err => raise Bug ("Tptp.fromProof: shouldn't fail:\n" ^ err); *) end; -val proofToString = Parser.toString ppProof; - end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Units.sig --- a/src/Tools/Metis/src/Units.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Units.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* A STORE FOR UNIT THEOREMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Units = @@ -24,7 +24,7 @@ val toString : units -> string -val pp : units Parser.pp +val pp : units Print.pp (* ------------------------------------------------------------------------- *) (* Add units into the store. *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Units.sml --- a/src/Tools/Metis/src/Units.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Units.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* A STORE FOR UNIT THEOREMS *) -(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Units :> Units = @@ -26,7 +26,7 @@ fun toString units = "U{" ^ Int.toString (size units) ^ "}"; -val pp = Parser.ppMap toString Parser.ppString; +val pp = Print.ppMap toString Print.ppString; (* ------------------------------------------------------------------------- *) (* Add units into the store. *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Useful.sig --- a/src/Tools/Metis/src/Useful.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Useful.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* 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 = @@ -14,8 +14,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 @@ -46,10 +44,6 @@ val exp : ('a * 'a -> 'a) -> 'a -> int -> 'a -> 'a -val equal : ''a -> ''a -> bool - -val notEqual : ''a -> ''a -> bool - (* ------------------------------------------------------------------------- *) (* Pairs. *) (* ------------------------------------------------------------------------- *) @@ -83,6 +77,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. *) (* ------------------------------------------------------------------------- *) @@ -96,15 +117,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 @@ -114,6 +131,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 *) @@ -145,23 +180,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. *) (* ------------------------------------------------------------------------- *) @@ -209,12 +227,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. *) (* ------------------------------------------------------------------------- *) @@ -271,9 +301,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. *) @@ -281,6 +313,8 @@ val try : ('a -> 'b) -> 'a -> 'b +val chat : string -> unit + val warn : string -> unit val die : string -> 'exit diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Useful.sml --- a/src/Tools/Metis/src/Useful.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Useful.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,42 +1,61 @@ (* ========================================================================= *) (* 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 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 bugToString (Bug message) = "\nBug: " ^ message ^ "\n" - | bugToString _ = raise Bug "bugToString: not a Bug 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 = ref print; -fun trace message = !tracePrint message; +fun trace mesg = !tracePrint mesg; (* ------------------------------------------------------------------------- *) -(* Combinators *) +(* Combinators. *) (* ------------------------------------------------------------------------- *) fun C f x y = f y x; @@ -60,12 +79,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; @@ -83,7 +98,7 @@ val op## = fn (f,g) => fn (x,y) => (f x, g y); (* ------------------------------------------------------------------------- *) -(* State transformers *) +(* State transformers. *) (* ------------------------------------------------------------------------- *) val unit : 'a -> 's -> 'a * 's = pair; @@ -97,119 +112,22 @@ fun mwhile c b = let fun f a = if c a then bind (b a) f else unit a in f end; (* ------------------------------------------------------------------------- *) -(* Lists *) +(* Equality. *) (* ------------------------------------------------------------------------- *) -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))); +val equal = fn x => fn y => x = y; -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); +val notEqual = fn x => fn y => x <> y; -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 = +fun listEqual xEq = let - val (a,b) = revDivide l n + fun xsEq [] [] = true + | xsEq (x1 :: xs1) (x2 :: xs2) = xEq x1 x2 andalso xsEq xs1 xs2 + | xsEq _ _ = false in - case b of [] => raise Subscript | _ :: t => List.revAppend (a, x :: t) + xsEq 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; - (* ------------------------------------------------------------------------- *) (* Comparisons. *) (* ------------------------------------------------------------------------- *) @@ -244,11 +162,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. *) (* ------------------------------------------------------------------------- *) @@ -301,7 +404,7 @@ | l as [_] => l | h :: t => mergePairs (findRuns [] h [] t) end; - + fun sortMap _ _ [] = [] | sortMap _ _ (l as [_]) = l | sortMap f cmp xs = @@ -339,50 +442,50 @@ 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 = 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 + val ps = ps @ [i] + and n = n - 1 in - l - end); + if n = 0 then ps else calcPrimes ps n (i + 1) + end; - fun primesUpTo n = CRITICAL (fn () => + val primesList = ref [2]; +in + fun primes n = 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) + val ref ps = primesList + + val k = n - length ps in - f 0 (!primesList) - end); + 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. *) (* ------------------------------------------------------------------------- *) @@ -449,7 +552,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 @@ -472,23 +576,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. *) (* ------------------------------------------------------------------------- *) @@ -507,13 +646,20 @@ else padding ^ entry ^ row end in - zipwith pad column + 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)); +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. *) @@ -556,22 +702,22 @@ local val generator = 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 = @@ -601,17 +747,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 @@ -622,11 +794,13 @@ end; (* ------------------------------------------------------------------------- *) -(* Profiling *) +(* Profiling and error reporting. *) (* ------------------------------------------------------------------------- *) +fun chat s = TextIO.output (TextIO.stdErr, s ^ "\n"); + local - fun err x s = TextIO.output (TextIO.stdErr, x ^ ": " ^ s ^ "\n"); + fun err x s = chat (x ^ ": " ^ s); in fun try f x = f x handle e as Error _ => (err "try" (errorToString e); raise e) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Waiting.sig --- a/src/Tools/Metis/src/Waiting.sig Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Waiting.sig Mon Sep 13 21:24:10 2010 +0200 @@ -1,21 +1,47 @@ (* ========================================================================= *) (* THE WAITING SET OF CLAUSES *) -(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Waiting = sig (* ------------------------------------------------------------------------- *) -(* A type of waiting sets of clauses. *) +(* The parameters control the order that clauses are removed from the *) +(* waiting set: clauses are assigned a weight and removed in strict weight *) +(* order, with smaller weights being removed before larger weights. *) +(* *) +(* The weight of a clause is defined to be *) +(* *) +(* d * s^symbolsWeight * v^variablesWeight * l^literalsWeight * m *) +(* *) +(* where *) +(* *) +(* d = the derivation distance of the clause from the axioms *) +(* s = the number of symbols in the clause *) +(* v = the number of distinct variables in the clause *) +(* l = the number of literals in the clause *) +(* m = the truth of the clause wrt the models *) (* ------------------------------------------------------------------------- *) +type weight = real + +type modelParameters = + {model : Model.parameters, + initialPerturbations : int, + maxChecks : int option, + perturbations : int, + weight : weight} + type parameters = - {symbolsWeight : real, - literalsWeight : real, - modelsWeight : real, - modelChecks : int, - models : Model.parameters list} + {symbolsWeight : weight, + variablesWeight : weight, + literalsWeight : weight, + models : modelParameters list} + +(* ------------------------------------------------------------------------- *) +(* A type of waiting sets of clauses. *) +(* ------------------------------------------------------------------------- *) type waiting @@ -27,11 +53,14 @@ val default : parameters -val new : parameters -> Clause.clause list -> waiting +val new : + parameters -> + {axioms : Clause.clause list, + conjecture : Clause.clause list} -> waiting val size : waiting -> int -val pp : waiting Parser.pp +val pp : waiting Print.pp (* ------------------------------------------------------------------------- *) (* Adding new clauses. *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/Waiting.sml --- a/src/Tools/Metis/src/Waiting.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/Waiting.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* THE WAITING SET OF CLAUSES *) -(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Waiting :> Waiting = @@ -12,35 +12,23 @@ (* A type of waiting sets of clauses. *) (* ------------------------------------------------------------------------- *) -(* The parameter type controls the heuristics for clause selection. *) -(* Increasing any of the *Weight parameters will favour clauses with low *) -(* values of that field. *) - -(* Note that there is an extra parameter of inference distance from the *) -(* starting axioms (a.k.a. time) which has a fixed weight of 1, so all *) -(* the other parameters should be set relative to this baseline. *) +type weight = real; -(* The first two parameters, symbolsWeight and literalsWeight, control the *) -(* time:weight ratio, i.e., whether to favour clauses derived in a few *) -(* steps from the axioms (time) or whether to favour small clauses (weight). *) -(* Small can be a combination of low number of symbols (the symbolWeight *) -(* parameter) or literals (the literalsWeight parameter). *) - -(* modelsWeight controls the semantic guidance. Increasing this parameter *) -(* favours clauses that return false more often when interpreted *) -(* modelChecks times over the given list of models. *) +type modelParameters = + {model : Model.parameters, + initialPerturbations : int, + maxChecks : int option, + perturbations : int, + weight : weight} type parameters = - {symbolsWeight : real, - literalsWeight : real, - modelsWeight : real, - modelChecks : int, - models : Model.parameters list}; + {symbolsWeight : weight, + variablesWeight : weight, + literalsWeight : weight, + models : modelParameters list}; type distance = real; -type weight = real; - datatype waiting = Waiting of {parameters : parameters, @@ -51,69 +39,145 @@ (* Basic operations. *) (* ------------------------------------------------------------------------- *) +val defaultModels : modelParameters list = + [{model = Model.default, + initialPerturbations = 100, + maxChecks = SOME 20, + perturbations = 0, + weight = 1.0}]; + val default : parameters = {symbolsWeight = 1.0, - literalsWeight = 0.0, - modelsWeight = 0.0, - modelChecks = 20, - models = []}; + literalsWeight = 1.0, + variablesWeight = 1.0, + models = defaultModels}; fun size (Waiting {clauses,...}) = Heap.size clauses; val pp = - Parser.ppMap + Print.ppMap (fn w => "Waiting{" ^ Int.toString (size w) ^ "}") - Parser.ppString; + Print.ppString; -(*DEBUG +(*MetisDebug val pp = - Parser.ppMap + Print.ppMap (fn Waiting {clauses,...} => map (fn (w,(_,cl)) => (w, Clause.id cl, cl)) (Heap.toList clauses)) - (Parser.ppList (Parser.ppTriple Parser.ppReal Parser.ppInt Clause.pp)); + (Print.ppList (Print.ppTriple Print.ppReal Print.ppInt Clause.pp)); *) (* ------------------------------------------------------------------------- *) +(* Perturbing the models. *) +(* ------------------------------------------------------------------------- *) + +type modelClause = NameSet.set * Thm.clause; + +fun mkModelClause cl = + let + val lits = Clause.literals cl + val fvs = LiteralSet.freeVars lits + in + (fvs,lits) + end; + +val mkModelClauses = map mkModelClause; + +fun perturbModel M cls = + if null cls then K () + else + let + val N = {size = Model.size M} + + fun perturbClause (fv,cl) = + let + val V = Model.randomValuation N fv + in + if Model.interpretClause M V cl then () + else Model.perturbClause M V cl + end + + fun perturbClauses () = app perturbClause cls + in + fn n => funpow n perturbClauses () + end; + +fun initialModel axioms conjecture parm = + let + val {model,initialPerturbations,...} : modelParameters = parm + val m = Model.new model + val () = perturbModel m conjecture initialPerturbations + val () = perturbModel m axioms initialPerturbations + in + m + end; + +fun checkModels parms models (fv,cl) = + let + fun check ((parm,model),z) = + let + val {maxChecks,weight,...} : modelParameters = parm + val n = {maxChecks = maxChecks} + val {T,F} = Model.check Model.interpretClause n model fv cl + in + Math.pow (1.0 + Real.fromInt T / Real.fromInt (T + F), weight) * z + end + in + List.foldl check 1.0 (zip parms models) + end; + +fun perturbModels parms models cls = + let + fun perturb (parm,model) = + let + val {perturbations,...} : modelParameters = parm + in + perturbModel model cls perturbations + end + in + app perturb (zip parms models) + end; + +(* ------------------------------------------------------------------------- *) (* Clause weights. *) (* ------------------------------------------------------------------------- *) local fun clauseSymbols cl = Real.fromInt (LiteralSet.typedSymbols cl); + fun clauseVariables cl = + Real.fromInt (NameSet.size (LiteralSet.freeVars cl) + 1); + fun clauseLiterals cl = Real.fromInt (LiteralSet.size cl); - fun clauseSat modelChecks models cl = + fun clausePriority cl = 1e~12 * Real.fromInt (Clause.id cl); +in + fun clauseWeight (parm : parameters) mods dist mcl cl = let - fun g {T,F} = (Real.fromInt T / Real.fromInt (T + F)) + 1.0 - fun f (m,z) = g (Model.checkClause {maxChecks = modelChecks} m cl) * z - in - foldl f 1.0 models - end; - - fun priority cl = 1e~12 * Real.fromInt (Clause.id cl); -in - fun clauseWeight (parm : parameters) models dist cl = - let -(*TRACE3 - val () = Parser.ppTrace Clause.pp "Waiting.clauseWeight: cl" cl +(*MetisTrace3 + val () = Print.trace Clause.pp "Waiting.clauseWeight: cl" cl *) - val {symbolsWeight,literalsWeight,modelsWeight,modelChecks,...} = parm + val {symbolsWeight,variablesWeight,literalsWeight,models,...} = parm val lits = Clause.literals cl val symbolsW = Math.pow (clauseSymbols lits, symbolsWeight) + val variablesW = Math.pow (clauseVariables lits, variablesWeight) val literalsW = Math.pow (clauseLiterals lits, literalsWeight) - val modelsW = Math.pow (clauseSat modelChecks models lits, modelsWeight) -(*TRACE4 + val modelsW = checkModels models mods mcl +(*MetisTrace4 val () = trace ("Waiting.clauseWeight: dist = " ^ Real.toString dist ^ "\n") val () = trace ("Waiting.clauseWeight: symbolsW = " ^ Real.toString symbolsW ^ "\n") + val () = trace ("Waiting.clauseWeight: variablesW = " ^ + Real.toString variablesW ^ "\n") val () = trace ("Waiting.clauseWeight: literalsW = " ^ Real.toString literalsW ^ "\n") val () = trace ("Waiting.clauseWeight: modelsW = " ^ Real.toString modelsW ^ "\n") *) - val weight = dist * symbolsW * literalsW * modelsW + priority cl -(*TRACE3 + val weight = dist * symbolsW * variablesW * literalsW * modelsW + val weight = weight + clausePriority cl +(*MetisTrace3 val () = trace ("Waiting.clauseWeight: weight = " ^ Real.toString weight ^ "\n") *) @@ -126,29 +190,39 @@ (* Adding new clauses. *) (* ------------------------------------------------------------------------- *) -fun add waiting (_,[]) = waiting - | add waiting (dist,cls) = +fun add' waiting dist mcls cls = let -(*TRACE3 - val () = Parser.ppTrace pp "Waiting.add: waiting" waiting - val () = Parser.ppTrace (Parser.ppList Clause.pp) "Waiting.add: cls" cls -*) - val Waiting {parameters,clauses,models} = waiting + val {models = modelParameters, ...} = parameters val dist = dist + Math.ln (Real.fromInt (length cls)) - val weight = clauseWeight parameters models dist + fun addCl ((mcl,cl),acc) = + let + val weight = clauseWeight parameters models dist mcl cl + in + Heap.add acc (weight,(dist,cl)) + end - fun f (cl,acc) = Heap.add acc (weight cl, (dist,cl)) - - val clauses = foldl f clauses cls + val clauses = List.foldl addCl clauses (zip mcls cls) - val waiting = - Waiting {parameters = parameters, clauses = clauses, models = models} + val () = perturbModels modelParameters models mcls + in + Waiting {parameters = parameters, clauses = clauses, models = models} + end; -(*TRACE3 - val () = Parser.ppTrace pp "Waiting.add: waiting" waiting +fun add waiting (_,[]) = waiting + | add waiting (dist,cls) = + let +(*MetisTrace3 + val () = Print.trace pp "Waiting.add: waiting" waiting + val () = Print.trace (Print.ppList Clause.pp) "Waiting.add: cls" cls +*) + + val waiting = add' waiting dist (mkModelClauses cls) cls + +(*MetisTrace3 + val () = Print.trace pp "Waiting.add: waiting" waiting *) in waiting @@ -157,15 +231,24 @@ local fun cmp ((w1,_),(w2,_)) = Real.compare (w1,w2); - fun empty parameters = + fun empty parameters axioms conjecture = let + val {models = modelParameters, ...} = parameters val clauses = Heap.new cmp - and models = map Model.new (#models parameters) + and models = map (initialModel axioms conjecture) modelParameters in Waiting {parameters = parameters, clauses = clauses, models = models} end; in - fun new parameters cls = add (empty parameters) (0.0,cls); + fun new parameters {axioms,conjecture} = + let + val mAxioms = mkModelClauses axioms + and mConjecture = mkModelClauses conjecture + + val waiting = empty parameters mAxioms mConjecture + in + add' waiting 0.0 (mAxioms @ mConjecture) (axioms @ conjecture) + end; end; (* ------------------------------------------------------------------------- *) diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/metis.sml --- a/src/Tools/Metis/src/metis.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/metis.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,7 +1,7 @@ (* ========================================================================= *) (* METIS FIRST ORDER PROVER *) (* *) -(* Copyright (c) 2001-2007 Joe Hurd *) +(* Copyright (c) 2001 Joe Hurd *) (* *) (* Metis is free software; you can redistribute it and/or modify *) (* it under the terms of the GNU General Public License as published by *) @@ -21,11 +21,15 @@ open Useful; (* ------------------------------------------------------------------------- *) -(* The program name. *) +(* The program name and version. *) (* ------------------------------------------------------------------------- *) val PROGRAM = "metis"; +val VERSION = "2.2"; + +val versionString = PROGRAM^" "^VERSION^" (release 20100825)"^"\n"; + (* ------------------------------------------------------------------------- *) (* Program options. *) (* ------------------------------------------------------------------------- *) @@ -34,7 +38,11 @@ val TEST = ref false; -val ITEMS = ["name","goal","clauses","size","category","proof","saturated"]; +val TPTP : string option ref = ref NONE; + +val ITEMS = ["name","goal","clauses","size","category","proof","saturation"]; + +val extended_items = "all" :: ITEMS; val show_items = map (fn s => (s, ref false)) ITEMS; @@ -43,6 +51,8 @@ NONE => raise Bug ("item " ^ s ^ " not found") | SOME (_,r) => r; +fun show_set b = app (fn (_,r) => r := b) show_items; + fun showing s = not (!QUIET) andalso (s = "status" orelse !(show_ref s)); fun notshowing s = not (showing s); @@ -51,9 +61,15 @@ fun notshowing_any () = not (showing_any ()); -fun show s = case show_ref s of r => r := true; +fun show "all" = show_set true + | show s = case show_ref s of r => r := true; -fun hide s = case show_ref s of r => r := false; +fun hide "all" = show_set false + | hide s = case show_ref s of r => r := false; + +(* ------------------------------------------------------------------------- *) +(* Process command line arguments and environment variables. *) +(* ------------------------------------------------------------------------- *) local open Useful Options; @@ -62,11 +78,15 @@ [{switches = ["--show"], arguments = ["ITEM"], description = "show ITEM (see below for list)", processor = - beginOpt (enumOpt ITEMS endOpt) (fn _ => fn s => show s)}, + beginOpt (enumOpt extended_items endOpt) (fn _ => fn s => show s)}, {switches = ["--hide"], arguments = ["ITEM"], description = "hide ITEM (see below for list)", processor = - beginOpt (enumOpt ITEMS endOpt) (fn _ => fn s => hide s)}, + beginOpt (enumOpt extended_items endOpt) (fn _ => fn s => hide s)}, + {switches = ["--tptp"], arguments = ["DIR"], + description = "specify the TPTP installation directory", + processor = + beginOpt (stringOpt endOpt) (fn _ => fn s => TPTP := SOME s)}, {switches = ["-q","--quiet"], arguments = [], description = "Run quietly; indicate provability with return value", processor = beginOpt endOpt (fn _ => QUIET := true)}, @@ -75,16 +95,12 @@ processor = beginOpt endOpt (fn _ => TEST := true)}]; end; -val VERSION = "2.0"; - -val versionString = "Metis "^VERSION^" (release 20071110)"^"\n"; - val programOptions = {name = PROGRAM, version = versionString, header = "usage: "^PROGRAM^" [option ...] problem.tptp ...\n" ^ "Proves the input TPTP problem files.\n", - footer = "Possible ITEMs are {" ^ join "," ITEMS ^ "}.\n" ^ + footer = "Possible ITEMs are {" ^ join "," extended_items ^ "}.\n" ^ "Problems can be read from standard input using the " ^ "special - filename.\n", options = specialOptions @ Options.basicOptions}; @@ -94,27 +110,56 @@ fun fail mesg = Options.fail programOptions mesg; fun usage mesg = Options.usage programOptions mesg; -val (opts,work) = - Options.processOptions programOptions (CommandLine.arguments ()); +fun processOptions () = + let + val args = CommandLine.arguments () + + val (_,work) = Options.processOptions programOptions args -val () = if null work then usage "no input problem files" else (); + val () = + case !TPTP of + SOME _ => () + | NONE => TPTP := OS.Process.getEnv "TPTP" + in + work + end; (* ------------------------------------------------------------------------- *) (* The core application. *) (* ------------------------------------------------------------------------- *) +(*MetisDebug +val next_cnf = + let + val cnf_counter = ref 0 + in + fn () => + let + val ref cnf_count = cnf_counter + val () = cnf_counter := cnf_count + 1 + in + cnf_count + end + end; +*) + local fun display_sep () = if notshowing_any () then () - else print (nChars #"-" (!Parser.lineLength) ^ "\n"); + else print (nChars #"-" (!Print.lineLength) ^ "\n"); fun display_name filename = if notshowing "name" then () else print ("Problem: " ^ filename ^ "\n\n"); - fun display_goal goal = + fun display_goal tptp = if notshowing "goal" then () - else print ("Goal:\n" ^ Formula.toString goal ^ "\n\n"); + else + let + val goal = Tptp.goal tptp + in + print ("Goal:\n" ^ Formula.toString goal ^ "\n\n") + end; fun display_clauses cls = if notshowing "clauses" then () @@ -126,7 +171,7 @@ let fun plural 1 s = "1 " ^ s | plural n s = Int.toString n ^ " " ^ s ^ "s" - + val {clauses,literals,symbols,typedSymbols} = Problem.size cls in print @@ -136,7 +181,7 @@ plural symbols "symbol" ^ ", " ^ plural typedSymbols "typed symbol" ^ ".\n\n") end; - + fun display_category cls = if notshowing "category" then () else @@ -146,41 +191,115 @@ print ("Category: " ^ Problem.categoryToString cat ^ ".\n\n") end; - fun display_proof filename th = - if notshowing "proof" then () - else - (print ("SZS output start CNFRefutation for " ^ filename ^ "\n"); - print (Tptp.proofToString (Proof.proof th)); - print ("SZS output end CNFRefutation for " ^ filename ^ "\n\n")); + local + fun display_proof_start filename = + print ("\nSZS output start CNFRefutation for " ^ filename ^ "\n"); + + fun display_proof_body problem proofs = + let + val comments = [] + + val includes = [] + + val formulas = + Tptp.fromProof + {problem = problem, + proofs = proofs} + + val proof = + Tptp.Problem + {comments = comments, + includes = includes, + formulas = formulas} + + val mapping = Tptp.defaultMapping + val mapping = Tptp.addVarSetMapping mapping (Tptp.freeVars proof) - fun display_saturated filename ths = - if notshowing "saturated" then () + val filename = "-" + in + Tptp.write + {problem = proof, + mapping = mapping, + filename = filename} + end; + + fun display_proof_end filename = + print ("SZS output end CNFRefutation for " ^ filename ^ "\n\n"); + in + fun display_proof filename problem proofs = + if notshowing "proof" then () + else + let + val () = display_proof_start filename + val () = display_proof_body problem proofs + val () = display_proof_end filename + in + () + end; + end; + + fun display_saturation filename ths = + if notshowing "saturation" then () else let -(*DEBUG - val () = Tptp.write {filename = "saturated.tptp"} - (Tptp.fromProblem (map Thm.clause ths)) +(*MetisDebug + val () = + let + val problem = + Tptp.mkProblem + {comments = ["Saturation clause set for " ^ filename], + includes = [], + names = Tptp.noClauseNames, + roles = Tptp.noClauseRoles, + problem = {axioms = [], + conjecture = map Thm.clause ths}} + + val mapping = + Tptp.addVarSetMapping Tptp.defaultMapping + (Tptp.freeVars problem) + in + Tptp.write + {problem = problem, + mapping = mapping, + filename = "saturation.tptp"} + end *) - val () = print ("SZS output start Saturated for " ^ filename ^ "\n") + val () = print ("\nSZS output start Saturation for " ^ filename ^ "\n") val () = app (fn th => print (Thm.toString th ^ "\n")) ths - val () = print ("SZS output end Saturated for " ^ filename ^ "\n\n") + val () = print ("SZS output end Saturation for " ^ filename ^ "\n\n") in () end; - fun display_result filename result = - case result of - Resolution.Contradiction th => display_proof filename th - | Resolution.Satisfiable ths => display_saturated filename ths; - fun display_status filename status = if notshowing "status" then () - else print ("SZS status " ^ status ^ " for " ^ filename ^ "\n"); + else print ("SZS status " ^ Tptp.toStringStatus status ^ + " for " ^ filename ^ "\n"); fun display_problem filename cls = let -(*DEBUG - val () = Tptp.write {filename = "cnf.tptp"} (Tptp.fromProblem cls) +(*MetisDebug + val () = + let + val problem = + Tptp.mkProblem + {comments = ["CNF clauses for " ^ filename], + includes = [], + names = Tptp.noClauseNames, + roles = Tptp.noClauseRoles, + problem = cls} + + val mapping = + Tptp.addVarSetMapping Tptp.defaultMapping + (Tptp.freeVars problem) + + val filename = "cnf_" ^ Int.toString (next_cnf ()) ^ ".tptp" + in + Tptp.write + {problem = problem, + mapping = mapping, + filename = filename} + end *) val () = display_clauses cls val () = display_size cls @@ -189,66 +308,178 @@ () end; - fun display_problems filename problems = - List.app (display_problem filename) problems; + fun mkTptpFilename filename = + case !TPTP of + NONE => filename + | SOME tptp => + let + val tptp = stripSuffix (equal #"/") tptp + in + tptp ^ "/" ^ filename + end; + + fun readIncludes mapping seen formulas includes = + case includes of + [] => formulas + | inc :: includes => + if StringSet.member inc seen then + readIncludes mapping seen formulas includes + else + let + val seen = StringSet.add seen inc + + val filename = mkTptpFilename inc + + val Tptp.Problem {includes = i, formulas = f, ...} = + Tptp.read {filename = filename, mapping = mapping} + + val formulas = f @ formulas - fun refute cls = - Resolution.loop (Resolution.new Resolution.default (map Thm.axiom cls)); + val includes = List.revAppend (i,includes) + in + readIncludes mapping seen formulas includes + end; + + fun read mapping filename = + let + val problem = Tptp.read {filename = filename, mapping = mapping} - fun refutable filename cls = + val Tptp.Problem {comments,includes,formulas} = problem + in + if null includes then problem + else + let + val seen = StringSet.empty + + val includes = rev includes + + val formulas = readIncludes mapping seen formulas includes + in + Tptp.Problem + {comments = comments, + includes = [], + formulas = formulas} + end + end; + + val resolutionParameters = let - val () = display_problem filename cls + val {active, + waiting} = Resolution.default + + val waiting = + let + val {symbolsWeight, + variablesWeight, + literalsWeight, + models} = waiting + + val models = + case models of + [{model = _, + initialPerturbations, + maxChecks, + perturbations, + weight}] => + let + val model = Tptp.defaultModel + in + [{model = model, + initialPerturbations = initialPerturbations, + maxChecks = maxChecks, + perturbations = perturbations, + weight = weight}] + end + | _ => raise Bug "resolutionParameters.waiting.models" + in + {symbolsWeight = symbolsWeight, + variablesWeight = variablesWeight, + literalsWeight = literalsWeight, + models = models} + end + in + {active = active, + waiting = waiting} + end; + + fun refute {axioms,conjecture} = + let + val axioms = map Thm.axiom axioms + and conjecture = map Thm.axiom conjecture + val problem = {axioms = axioms, conjecture = conjecture} + val resolution = Resolution.new resolutionParameters problem in - case refute cls of - Resolution.Contradiction th => (display_proof filename th; true) - | Resolution.Satisfiable ths => (display_saturated filename ths; false) + Resolution.loop resolution end; + + fun refuteAll filename tptp probs acc = + case probs of + [] => + let + val status = + if !TEST then Tptp.UnknownStatus + else if Tptp.hasFofConjecture tptp then Tptp.TheoremStatus + else Tptp.UnsatisfiableStatus + + val () = display_status filename status + + val () = + if !TEST then () + else display_proof filename tptp (rev acc) + in + true + end + | prob :: probs => + let + val {subgoal,problem,sources} = prob + + val () = display_problem filename problem + in + if !TEST then refuteAll filename tptp probs acc + else + case refute problem of + Resolution.Contradiction th => + let + val subgoalProof = + {subgoal = subgoal, + sources = sources, + refutation = th} + + val acc = subgoalProof :: acc + in + refuteAll filename tptp probs acc + end + | Resolution.Satisfiable ths => + let + val status = + if Tptp.hasFofConjecture tptp then + Tptp.CounterSatisfiableStatus + else + Tptp.SatisfiableStatus + + val () = display_status filename status + val () = display_saturation filename ths + in + false + end + end; in - fun prove filename = + fun prove mapping filename = let val () = display_sep () val () = display_name filename - val tptp = Tptp.read {filename = filename} + val tptp = read mapping filename + val () = display_goal tptp + val problems = Tptp.normalize tptp in - case Tptp.toGoal tptp of - Tptp.Cnf prob => - let - val () = display_problem filename prob - in - if !TEST then - (display_status filename "Unknown"; - true) - else - case refute prob of - Resolution.Contradiction th => - (display_status filename "Unsatisfiable"; - if showing "proof" then print "\n" else (); - display_proof filename th; - true) - | Resolution.Satisfiable ths => - (display_status filename "Satisfiable"; - if showing "saturated" then print "\n" else (); - display_saturated filename ths; - false) - end - | Tptp.Fof goal => - let - val () = display_goal goal - val problems = Problem.fromGoal goal - val result = - if !TEST then (display_problems filename problems; true) - else List.all (refutable filename) problems - val status = - if !TEST then "Unknown" - else if Tptp.hasConjecture tptp then - if result then "Theorem" else "CounterSatisfiable" - else - if result then "Unsatisfiable" else "Satisfiable" - val () = display_status filename status - in - result - end + refuteAll filename tptp problems [] end; + + fun proveAll mapping filenames = + List.all + (if !QUIET then prove mapping + else fn filename => prove mapping filename orelse true) + filenames; end; (* ------------------------------------------------------------------------- *) @@ -257,13 +488,15 @@ val () = let -(*DEBUG - val () = print "Running in DEBUG mode.\n" -*) - val success = List.all prove work - val return = not (!QUIET) orelse success + val work = processOptions () + + val () = if null work then usage "no input problem files" else () + + val mapping = Tptp.defaultMapping + + val success = proveAll mapping work in - exit {message = NONE, usage = false, success = return} + exit {message = NONE, usage = false, success = success} end handle Error s => die (PROGRAM^" failed:\n" ^ s) | Bug s => die ("BUG found in "^PROGRAM^" program:\n" ^ s); diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/problems.sml --- a/src/Tools/Metis/src/problems.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/problems.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* SOME SAMPLE PROBLEMS TO TEST PROOF PROCEDURES *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) (* ========================================================================= *) @@ -16,16 +16,25 @@ (* Helper functions. *) (* ========================================================================= *) -fun mkProblem description (problem : problem) : problem = - let - val {name,comments,goal} = problem - val comments = if null comments then [] else "" :: comments - val comments = "Collection: " ^ description :: comments - in - {name = name, comments = comments, goal = goal} - end; +local + fun mkCollection collection = "Collection: " ^ collection; -fun mkProblems description problems = map (mkProblem description) problems; + fun mkProblem collection description (problem : problem) : problem = + let + val {name,comments,goal} = problem + val comments = if null comments then [] else "" :: comments + val comments = "Description: " ^ description :: comments + val comments = mkCollection collection :: comments + in + {name = name, comments = comments, goal = goal} + end; +in + fun isCollection collection {name = _, comments, goal = _} = + Useful.mem (mkCollection collection) comments; + + fun mkProblems collection description problems = + map (mkProblem collection description) problems; +end; (* ========================================================================= *) (* The collection of problems. *) @@ -37,7 +46,7 @@ (* Problems without equality. *) (* ========================================================================= *) -mkProblems "Problems without equality from various sources" [ +mkProblems "nonequality" "Problems without equality from various sources" [ (* ------------------------------------------------------------------------- *) (* Trivia (some of which demonstrate ex-bugs in provers). *) @@ -69,6 +78,27 @@ (!x. ?y. p x y) /\ (!x. ?y. q x y) /\ (!x y z. p x y /\ q y z ==> r x z) ==> !x. ?y. r x y`}, +{name = "TOBIAS_NIPKOW", + comments = [], + goal = ` +(!x y. p x y ==> f x = f y) /\ (!x. f (g x) = f x) /\ p (g a) (g b) ==> +f a = f b`}, + +{name = "SLEDGEHAMMER", + comments = ["From Tobias Nipkow: A ==> A takes 1 minute in sledgehammer."], + goal = ` +(!x y z t. + x @ y = z @ t <=> + ?u. x = z @ u /\ u @ y = t \/ x @ u = z /\ y = u @ t) ==> +!x y z t. + x @ y = z @ t <=> ?u. x = z @ u /\ u @ y = t \/ x @ u = z /\ y = u @ t`}, + +{name = "SPLITTING_UNSOUNDNESS", + comments = ["A trivial example to illustrate a bug spotted by", + "Geoff Sutcliffe in Dec 2008."], + goal = ` +(!x. p x /\ q x ==> F) ==> !x. p x ==> !x. q x ==> F`}, + (* ------------------------------------------------------------------------- *) (* Propositional Logic. *) (* ------------------------------------------------------------------------- *) @@ -185,6 +215,11 @@ goal = ` (lit ==> clause) ==> (lit \/ clause <=> clause)`}, +{name = "SPLIT_NOT_IFF", + comments = ["A way to split a goal that looks like ~(p <=> q)"], + goal = ` +~(p <=> q) <=> (p ==> ~q) /\ (q ==> ~p)`}, + (* ------------------------------------------------------------------------- *) (* Monadic Predicate Logic. *) (* ------------------------------------------------------------------------- *) @@ -542,9 +577,8 @@ ?x y. p0 x /\ p0 y /\ ?z. q1 z /\ r y z /\ r x y`}, {name = "MODEL_COMPLETENESS", - comments = -["An incestuous example used to establish completeness", - "characterization. [JRH]"], + comments = ["An incestuous example used to establish completeness", + "characterization. [JRH]"], goal = ` (!w x. sentence x ==> holds w x \/ holds w (not x)) /\ (!w x. ~(holds w x /\ holds w (not x))) ==> @@ -562,7 +596,7 @@ (* Problems with equality. *) (* ========================================================================= *) -mkProblems "Equality problems from various sources" [ +mkProblems "equality" "Equality problems from various sources" [ (* ------------------------------------------------------------------------- *) (* Trivia (some of which demonstrate ex-bugs in the prover). *) @@ -639,9 +673,9 @@ f (g (h c)) = h c /\ g (h c) = b /\ f b = a /\ (!x. ~(a = h x)) ==> F`}, {name = "EQUALITY_ORDERING", - comments = -["Positive resolution saturates if equality literals are ordered like other", - "literals, instead of considering their left and right sides."], + comments = ["Positive resolution saturates if equality literals are", + "ordered like other literals, instead of considering their", + "left and right sides."], goal = ` p a /\ q a /\ q b /\ r b /\ (~p c \/ c = a) /\ (~r c \/ c = b) /\ (!x. ~q x \/ p x \/ r x) /\ (~p c \/ ~q c) /\ (~q c \/ ~r c) /\ @@ -677,7 +711,7 @@ comments = ["The Melham problem after an inverse skolemization step."], goal = ` (!x y. g x = g y ==> f x = f y) ==> !y. ?w. !x. y = g x ==> w = f x`}, - + {name = "CONGRUENCE_CLOSURE_EXAMPLE", comments = ["The example always given for congruence closure."], goal = ` @@ -736,6 +770,41 @@ (!x y. divides x y <=> ?z. z * x = y) ==> !x y z. divides x y ==> divides x (z * y)`}, +{name = "XOR_COUNT_COMMUTATIVE", + comments = ["The xor literal counting function in Normalize is commutative."], + goal = ` +(!x y. x + y = y + x) /\ (!x y z. x + (y + z) = x + y + z) /\ +(!x y. x * y = y * x) /\ (!x y z. x * (y * z) = x * y * z) /\ +pl = p1 * p2 + n1 * n2 /\ nl = p1 * n2 + n1 * p2 /\ +pr = p2 * p1 + n2 * n1 /\ nr = p2 * n1 + n2 * p1 ==> pl = pr /\ nl = nr`}, + +{name = "XOR_COUNT_ASSOCIATIVE", + comments = ["The xor literal counting function in Normalize is associative."], + goal = ` +(!x y. x + y = y + x) /\ (!x y z. x + (y + z) = x + y + z) /\ +(!x y. x * y = y * x) /\ (!x y z. x * (y * z) = x * y * z) /\ +px = p1 * p2 + n1 * n2 /\ nx = p1 * n2 + n1 * p2 /\ +pl = px * p3 + nx * n3 /\ nl = px * n3 + nx * p3 /\ +py = p2 * p3 + n2 * n3 /\ ny = p2 * n3 + n2 * p3 /\ +pr = p1 * py + n1 * ny /\ nr = p1 * ny + n1 * py ==> pl = pr /\ nl = nr`}, + +{name = "REVERSE_REVERSE", + comments = ["Proving the goal", + " !l. finite l ==> rev (rev l) = l", + "after first generalizing it to", + " !l k. finite l /\\ finite k ==> rev (rev l @ k) = rev k @ l", + "and then applying list induction."], + goal = ` +finite nil /\ (!h t. finite (h :: t) <=> finite t) /\ +(!l1 l2. finite (l1 @ l2) <=> finite l1 /\ finite l2) /\ +(!l. nil @ l = l) /\ (!h t l. (h :: t) @ l = h :: t @ l) /\ +rev nil = nil /\ (!h t. rev (h :: t) = rev t @ h :: nil) /\ +(!l. l @ nil = l) /\ (!l1 l2 l3. l1 @ l2 @ l3 = (l1 @ l2) @ l3) ==> +(!k. finite k ==> rev (rev nil @ k) = rev k @ nil) /\ +!t. + finite t ==> (!k. finite k ==> rev (rev t @ k) = rev k @ t) ==> + !h k. finite k ==> rev (rev (h :: t) @ k) = rev k @ h :: t`}, + (* ------------------------------------------------------------------------- *) (* Group theory examples. *) (* ------------------------------------------------------------------------- *) @@ -770,15 +839,34 @@ (!x y z. x * (y * z) = x * y * z) /\ (!x. e * x = x) /\ (!x. x * e = x) /\ (!x. x * x = e) ==> !x y. x * y = y * x`}, +{name = "DOUBLE_DISTRIB", + comments = ["From a John Harrison post to hol-info on 2008-04-15"], + goal = ` +(!x y z. x * y * z = x * z * (y * z)) /\ +(!x y z. z * (x * y) = z * x * (z * y)) ==> +!a b c. a * b * (c * a) = a * c * (b * a)`}, + (* ------------------------------------------------------------------------- *) (* Ring theory examples. *) (* ------------------------------------------------------------------------- *) +{name = "CONWAY_2", + comments = ["From a John Harrison post to hol-info on 2008-04-15"], + goal = ` +(!x. 0 + x = x) /\ (!x y. x + y = y + x) /\ +(!x y z. x + (y + z) = x + y + z) /\ (!x. 1 * x = x) /\ (!x. x * 1 = x) /\ +(!x y z. x * (y * z) = x * y * z) /\ (!x. 0 * x = 0) /\ (!x. x * 0 = 0) /\ +(!x y z. x * (y + z) = x * y + x * z) /\ +(!x y z. (x + y) * z = x * z + y * z) /\ +(!x y. star (x * y) = 1 + x * star (y * x) * y) /\ +(!x y. star (x + y) = star (star (x) * y) * star (x)) ==> +!a. star (star (star (star (a)))) = star (star (star (a)))`}, + {name = "JACOBSON_2", comments = [], goal = ` -(!x. 0 + x = x) /\ (!x. x + 0 = x) /\ (!x. n x + x = 0) /\ -(!x. x + n x = 0) /\ (!x y z. x + (y + z) = x + y + z) /\ +(!x. 0 + x = x) /\ (!x. x + 0 = x) /\ (!x. ~x + x = 0) /\ +(!x. x + ~x = 0) /\ (!x y z. x + (y + z) = x + y + z) /\ (!x y. x + y = y + x) /\ (!x y z. x * (y * z) = x * y * z) /\ (!x y z. x * (y + z) = x * y + x * z) /\ (!x y z. (x + y) * z = x * z + y * z) /\ (!x. x * x = x) ==> @@ -787,21 +875,57 @@ {name = "JACOBSON_3", comments = [], goal = ` -(!x. 0 + x = x) /\ (!x. x + 0 = x) /\ (!x. n x + x = 0) /\ -(!x. x + n x = 0) /\ (!x y z. x + (y + z) = x + y + z) /\ +(!x. 0 + x = x) /\ (!x. x + 0 = x) /\ (!x. ~x + x = 0) /\ +(!x. x + ~x = 0) /\ (!x y z. x + (y + z) = x + y + z) /\ (!x y. x + y = y + x) /\ (!x y z. x * (y * z) = x * y * z) /\ (!x y z. x * (y + z) = x * y + x * z) /\ (!x y z. (x + y) * z = x * z + y * z) /\ (!x. x * (x * x) = x) ==> -!x y. x * y = y * x`} +!x y. x * y = y * x`}, + +(* ------------------------------------------------------------------------- *) +(* Set theory examples. *) +(* ------------------------------------------------------------------------- *) + +{name = "UNION_2_SUBSET", + comments = [], + goal = ` +(!x y. subset x y ==> subset y x ==> x = y) /\ +(!x y. (!z. member z x ==> member z y) ==> subset x y) /\ +(!x y z. member z (x + y) <=> member z x \/ member z y) ==> +!x y. subset (x + y) (y + x)`}, -] @ +{name = "UNION_2", + comments = [], + goal = ` +(!x y. subset x y ==> subset y x ==> x = y) /\ +(!x y. (!z. member z x ==> member z y) ==> subset x y) /\ +(!x y z. member z (x + y) <=> member z x \/ member z y) ==> +!x y. x + y = y + x`}, + +{name = "UNION_3_SUBSET", + comments = ["From an email from Tobias Nipkow, 4 Nov 2008.", + "The Isabelle version of metis diverges on this goal"], + goal = ` +(!x y. subset x y ==> subset y x ==> x = y) /\ +(!x y. (!z. member z x ==> member z y) ==> subset x y) /\ +(!x y z. member z (x + y) <=> member z x \/ member z y) ==> +!x y z. subset (x + y + z) (z + y + x)`}, + +{name = "UNION_3", + comments = ["From an email from Tobias Nipkow, 28 Oct 2008.", + "The Isabelle version of metis diverges on this goal"], + goal = ` +(!x y. subset x y ==> subset y x ==> x = y) /\ +(!x y. (!z. member z x ==> member z y) ==> subset x y) /\ +(!x y z. member z (x + y) <=> member z x \/ member z y) ==> +!x y z. x + y + z = z + y + x`}] @ (* ========================================================================= *) (* Some sample problems from the TPTP archive. *) (* Note: for brevity some relation/function names have been shortened. *) (* ========================================================================= *) -mkProblems "Sample problems from the TPTP collection" +mkProblems "tptp" "Sample problems from the TPTP collection" [ @@ -1219,7 +1343,7 @@ (* Some problems from HOL. *) (* ========================================================================= *) -mkProblems "HOL subgoals sent to MESON_TAC" [ +mkProblems "hol" "HOL subgoals sent to MESON_TAC" [ {name = "Coder_4_0", comments = [], @@ -1384,7 +1508,7 @@ (!x y z. divides x y \/ ~divides x z \/ ~divides x (z + y)) ==> ~(x + z <= x) /\ divides c (d * SUC x) /\ divides c (d * SUC (x + z)) /\ ~divides c (d * z) ==> F`}, - + {name = "gcd_20", comments = [], goal = ` diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/problems2tptp.sml --- a/src/Tools/Metis/src/problems2tptp.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/problems2tptp.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,6 +1,6 @@ (* ========================================================================= *) (* SOME SAMPLE PROBLEMS TO TEST PROOF PROCEDURES *) -(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *) +(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) open Useful; @@ -39,6 +39,8 @@ dups names end; +fun listProblem {name, comments = _, goal = _} = print (name ^ "\n"); + fun outputProblem outputDir {name,comments,goal} = let val filename = name ^ ".tptp" @@ -58,26 +60,65 @@ (if null comment_footer then [] else "" :: comment_footer) @ [comment_bar] - val goal = Formula.parse goal + val includes = [] + val formulas = - [Tptp.FofFormula {name = "goal", role = "conjecture", formula = goal}] + let + val name = Tptp.FormulaName "goal" + val role = Tptp.ConjectureRole + val body = Tptp.FofFormulaBody (Formula.parse goal) + val source = Tptp.NoFormulaSource + in + [Tptp.Formula + {name = name, + role = role, + body = body, + source = source}] + end - val problem = {comments = comments, formulas = formulas} + val problem = + Tptp.Problem + {comments = comments, + includes = includes, + formulas = formulas} + + val mapping = Tptp.defaultMapping + + val () = + Tptp.write + {problem = problem, + mapping = mapping, + filename = filename} in - Tptp.write {filename = filename} problem + () end; (* ------------------------------------------------------------------------- *) (* Program options. *) (* ------------------------------------------------------------------------- *) +datatype mode = OutputMode | ListMode; + +val MODE : mode ref = ref OutputMode; + +val COLLECTION : string option ref = ref NONE; + val OUTPUT_DIRECTORY : string option ref = ref NONE; local open Useful Options; in val specialOptions = - [{switches = ["--output"], arguments = ["DIR"], + [{switches = ["--collection"], arguments = ["C"], + description = "restrict to the problems in collection C", + processor = + beginOpt + (stringOpt endOpt) + (fn _ => fn c => COLLECTION := SOME c)}, + {switches = ["--list"], arguments = [], + description = "just list the problems", + processor = beginOpt endOpt (fn _ => MODE := ListMode)}, + {switches = ["--output-dir"], arguments = ["DIR"], description = "the output directory", processor = beginOpt @@ -112,9 +153,17 @@ val () = let + val problems = + case !COLLECTION of + NONE => problems + | SOME c => List.filter (isCollection c) problems + val () = checkProblems problems - val () = app (outputProblem (!OUTPUT_DIRECTORY)) problems + val () = + case !MODE of + ListMode => app listProblem problems + | OutputMode => app (outputProblem (!OUTPUT_DIRECTORY)) problems in succeed () end diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/Metis/src/selftest.sml --- a/src/Tools/Metis/src/selftest.sml Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/Metis/src/selftest.sml Mon Sep 13 21:24:10 2010 +0200 @@ -1,10 +1,10 @@ (* ========================================================================= *) (* METIS TESTS *) -(* Copyright (c) 2004-2006 Joe Hurd *) +(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *) (* ========================================================================= *) (* ------------------------------------------------------------------------- *) -(* Dummy versions of Moscow ML declarations to stop MLton barfing. *) +(* Dummy versions of Moscow ML declarations to stop real compilers barfing. *) (* ------------------------------------------------------------------------- *) (*mlton @@ -12,7 +12,13 @@ val quietdec = ref true; val loadPath = ref ([] : string list); val load = fn (_ : string) => (); -val installPP = fn (_ : 'a Parser.pp) => (); +*) + +(*polyml +val quotation = ref true; +val quietdec = ref true; +val loadPath = ref ([] : string list); +val load = fn (_ : string) => (); *) (* ------------------------------------------------------------------------- *) @@ -20,23 +26,12 @@ (* ------------------------------------------------------------------------- *) val () = loadPath := !loadPath @ ["../bin/mosml"]; -val () = app load ["Tptp"]; +val () = app load ["Options"]; open Useful; -val () = installPP Term.pp; -val () = installPP Formula.pp; -val () = installPP Subst.pp; -val () = installPP Thm.pp; -val () = installPP Rewrite.pp; -val () = installPP Clause.pp; - val time = Portable.time; -(*DEBUG - val () = print "Running in DEBUG mode.\n" -*) - (* ------------------------------------------------------------------------- *) (* Problem data. *) (* ------------------------------------------------------------------------- *) @@ -57,48 +52,56 @@ | partialOrderToString NONE = "NONE"; fun SAY s = - print - ("-------------------------------------" ^ - "-------------------------------------\n" ^ s ^ "\n\n"); + print + ("-------------------------------------" ^ + "-------------------------------------\n" ^ s ^ "\n\n"); -fun printval p x = (print (PP.pp_to_string 79 p x ^ "\n\n"); x); +fun printval p x = (print (Print.toString p x ^ "\n\n"); x); + +fun mkCl p th = Clause.mk {parameters = p, id = Clause.newId (), thm = th}; -val pvBool = printval Parser.ppBool -and pvPo = printval (Parser.ppMap partialOrderToString Parser.ppString) +val pvBool = printval Print.ppBool +and pvPo = printval (Print.ppMap partialOrderToString Print.ppString) and pvFm = printval Formula.pp -and pvFms = printval (Parser.ppList Formula.pp) +and pvFms = printval (Print.ppList Formula.pp) and pvThm = printval Thm.pp -and pvEqn : Rule.equation -> Rule.equation = printval (Parser.ppMap snd Thm.pp) -and pvNet = printval (LiteralNet.pp Parser.ppInt) +and pvEqn : Rule.equation -> Rule.equation = printval (Print.ppMap snd Thm.pp) +and pvNet = printval (LiteralNet.pp Print.ppInt) and pvRw = printval Rewrite.pp and pvU = printval Units.pp and pvLits = printval LiteralSet.pp and pvCl = printval Clause.pp -and pvCls = printval (Parser.ppList Clause.pp); +and pvCls = printval (Print.ppList Clause.pp) +and pvM = printval Model.pp; -val T = Term.parse +val NV = Name.fromString +and NF = Name.fromString +and NR = Name.fromString; +val V = Term.Var o NV +and C = (fn c => Term.Fn (NF c, [])) +and T = Term.parse and A = Atom.parse and L = Literal.parse and F = Formula.parse and S = Subst.fromList; -val AX = Thm.axiom o LiteralSet.fromList o map L; -fun CL q = - Clause.mk {parameters = Clause.default, id = Clause.newId (), thm = AX q}; +val LS = LiteralSet.fromList o map L; +val AX = Thm.axiom o LS; +val CL = mkCl Clause.default o AX; val Q = (fn th => (Thm.destUnitEq th, th)) o AX o singleton and U = (fn th => (Thm.destUnit th, th)) o AX o singleton; -fun test_fun p r a = - if r = a then p a ^ "\n" else +fun test_fun eq p r a = + if eq r a then p a ^ "\n" else (print ("\n\n" ^ "test: should have\n-->" ^ p r ^ "<--\n\n" ^ "test: actually have\n-->" ^ p a ^ "<--\n\n"); raise Fail "test: failed a test"); -fun test p r a = print (test_fun p r a ^ "\n"); +fun test eq p r a = print (test_fun eq p r a ^ "\n"); -val test_tm = test Term.toString o Term.parse; +val test_tm = test Term.equal Term.toString o Term.parse; -val test_fm = test Formula.toString o Formula.parse; +val test_fm = test Formula.equal Formula.toString o Formula.parse; fun test_id p f a = test p a (f a); @@ -108,46 +111,16 @@ fun unquote (QUOTE q) = q | unquote (ANTIQUOTE _) = raise Fail "unquote"; -(*** -fun quick_prove slv goal = - let - val {thms,hyps} = Thm.clauses goal - val solv = initialize slv - in - (printval (pp_map Option.isSome pp_bool) o (time o try) refute) - (solv {limit = unlimited, thms = thms, hyps = hyps}) - end; - -val meson_prove = - quick_prove (Solver.apply_sos_filter Solver.all_negative meson); -val resolution_prove = quick_prove resolution; -val metis_prove = quick_prove metis; - -fun quick_solve slv n ruls = - printval (pp_list (pp_list pp_thm)) o - (time o try) - (solve - (initialize slv {limit = unlimited, thms = axiomatize ruls, hyps = []}) n); - -val meson_solve = quick_solve meson 1; -val prolog_solve = quick_solve prolog 1; -val resolution_solve = quick_solve resolution 1; -val metis_solve = quick_solve metis 1; - -val pfm = printval pp_formula; -val pfms = printval (pp_list pp_formula); -***) - (* ------------------------------------------------------------------------- *) val () = SAY "The parser and pretty-printer"; (* ------------------------------------------------------------------------- *) fun prep l = (chop_newline o String.concat o map unquote) l; -fun mini_print n = withRef (Parser.lineLength,n) Formula.toString; +fun mini_print n fm = withRef (Print.lineLength,n) Formula.toString fm; fun testlen_pp n q = - (fn s => test_fun I s ((mini_print n o Formula.fromString) s)) + (fn s => test_fun equal I s ((mini_print n o Formula.fromString) s)) (prep q); fun test_pp q = print (testlen_pp 40 q ^ "\n"); @@ -281,6 +254,11 @@ (f (f x y) (f x y)))`; val () = test_pp ` +(!x. + extremely__long__predicate__name) /\ +F`; + +val () = test_pp ` (!x. x = x) /\ (!x y. ~(x = y) \/ y = x) /\ (!x y z. @@ -296,32 +274,45 @@ val () = SAY "Substitution"; (* ------------------------------------------------------------------------- *) -val () = test I "x" (Term.variantPrime (NameSet.fromList ["y","z" ]) "x"); +val () = + test Name.equal Name.toString (NV"x") + (Term.variantPrime (NameSet.fromList [NV"y",NV"z" ]) (NV"x")); -val () = test I "x'" (Term.variantPrime (NameSet.fromList ["x","y" ]) "x"); +val () = + test Name.equal Name.toString (NV"x'") + (Term.variantPrime (NameSet.fromList [NV"x",NV"y" ]) (NV"x")); -val () = test I "x''" (Term.variantPrime (NameSet.fromList ["x","x'"]) "x"); +val () = + test Name.equal Name.toString (NV"x''") + (Term.variantPrime (NameSet.fromList [NV"x",NV"x'"]) (NV"x")); -val () = test I "x" (Term.variantNum (NameSet.fromList ["y","z"]) "x"); +val () = + test Name.equal Name.toString (NV"x") + (Term.variantNum (NameSet.fromList [NV"y",NV"z"]) (NV"x")); -val () = test I "x0" (Term.variantNum (NameSet.fromList ["x","y"]) "x"); +val () = + test Name.equal Name.toString (NV"x0") + (Term.variantNum (NameSet.fromList [NV"x",NV"y"]) (NV"x")); -val () = test I "x1" (Term.variantNum (NameSet.fromList ["x","x0"]) "x"); +val () = + test Name.equal Name.toString (NV"x1") + (Term.variantNum (NameSet.fromList [NV"x",NV"x0"]) (NV"x")); val () = test_fm `!x. x = $z` - (Formula.subst (S [("y", Term.Var "z")]) (F`!x. x = $y`)); + (Formula.subst (S [(NV"y", V"z")]) (F`!x. x = $y`)); val () = test_fm `!x'. x' = $x` - (Formula.subst (S [("y", Term.Var "x")]) (F`!x. x = $y`)); + (Formula.subst (S [(NV"y", V"x")]) (F`!x. x = $y`)); val () = test_fm `!x' x''. x' = $x ==> x' = x''` - (Formula.subst (S [("y", Term.Var "x")]) (F`!x x'. x = $y ==> x = x'`)); + (Formula.subst (S [(NV"y", V"x")]) + (F`!x x'. x = $y ==> x = x'`)); (* ------------------------------------------------------------------------- *) val () = SAY "Unification"; @@ -330,24 +321,24 @@ fun unify_and_apply tm1 tm2 = Subst.subst (Subst.unify Subst.empty tm1 tm2) tm1; -val () = test_tm `c` (unify_and_apply (Term.Var "x") (Term.Fn ("c", []))); +val () = test_tm `c` (unify_and_apply (V"x") (C"c")); -val () = test_tm `c` (unify_and_apply (Term.Fn ("c", [])) (Term.Var "x")); +val () = test_tm `c` (unify_and_apply (C"c") (V"x")); val () = test_tm `f c` (unify_and_apply - (Term.Fn ("f", [Term.Var "x"])) - (Term.Fn ("f", [Term.Fn ("c", [])]))); - + (Term.Fn (NF"f", [V"x"])) + (Term.Fn (NF"f", [C"c"]))); + val () = test_tm `f 0 0 0` (unify_and_apply (T`f 0 $x $x`) (T`f $y $y $z`)); fun f x y = (printval Subst.pp (Atom.unify Subst.empty x y); ()); -val () = f ("P", [Term.Var "x"]) ("P", [Term.Var "x"]); +val () = f (NR"P", [V"x"]) (NR"P", [V"x"]); -val () = f ("P", [Term.Var "x"]) ("P", [Term.Fn ("c",[])]); +val () = f (NR"P", [V"x"]) (NR"P", [C"c"]); val () = f (A`P c_x`) (A`P $x`); @@ -364,7 +355,7 @@ val th0 = Rule.relationCongruence Atom.eqRelation; val th1 = - Thm.subst (S [("y0",T`$x`),("y1",T`$y`),("x1",T`$z`),("x0",T`$x`)]) th0; + Thm.subst (S [(NV"y0",T`$x`),(NV"y1",T`$y`),(NV"x1",T`$z`),(NV"x0",T`$x`)]) th0; val th2 = Thm.resolve (L`$x = $x`) Rule.reflexivity th1; val th3 = Rule.symNeq (L`~($z = $y)`) th2; val _ = printval Proof.pp (Proof.proof th3); @@ -372,8 +363,8 @@ (* Testing the elimination of redundancies in proofs *) val th0 = Rule.reflexivity; -val th1 = Thm.subst (S [("x", Term.Fn ("f", [Term.Var "y"]))]) th0; -val th2 = Thm.subst (S [("y", Term.mkConst "c")]) th1; +val th1 = Thm.subst (S [(NV"x", Term.Fn (NF"f", [V"y"]))]) th0; +val th2 = Thm.subst (S [(NV"y", C"c")]) th1; val _ = printval Proof.pp (Proof.proof th2); (* ------------------------------------------------------------------------- *) @@ -389,7 +380,7 @@ (* Testing the rewrConv conversion *) val (x_y as (x,y), eqTh) = pvEqn (Q`e * (i $z * $z) = e`); -val tm = Term.Fn ("f",[x]); +val tm = Term.Fn (NF"f",[x]); val path : int list = [0]; val reflTh = Thm.refl tm; val reflLit = Thm.destUnit reflTh; @@ -434,6 +425,7 @@ val x = pvPo (kboCmp (T`f (g a)`, T`f (h $x)`)); val x = pvPo (kboCmp (T`f (h a)`, T`f (g $x)`)); val x = pvPo (kboCmp (T`f $y`, T`f (g a b c)`)); +val x = pvPo (kboCmp (T`$x * $y + $x * $z`, T`$x * ($y + $z)`)); (* ------------------------------------------------------------------------- *) val () = SAY "Rewriting"; @@ -512,7 +504,17 @@ val () = SAY "Conjunctive normal form"; (* ------------------------------------------------------------------------- *) -val cnf = pvFms o Normalize.cnf o Formula.Not o Formula.generalize o F; +local + fun clauseToFormula cl = + Formula.listMkDisj (LiteralSet.transform Literal.toFormula cl); +in + fun clausesToFormula cls = Formula.listMkConj (map clauseToFormula cls); +end; + +val cnf' = pvFm o clausesToFormula o Normalize.cnf o F; + +val cnf = pvFm o clausesToFormula o Normalize.cnf o + Formula.Not o Formula.generalize o F; val _ = cnf `p \/ ~p`; val _ = cnf `~((p /\ (q \/ r /\ s)) /\ (~p \/ ~q \/ ~s))`; @@ -522,97 +524,545 @@ val _ = cnf `~(!x. ?y. x < y ==> !v. ?w. x * v < y * w)`; val _ = cnf `~(!x. P x ==> (?y z. Q y \/ ~(?z. P z /\ Q z)))`; val _ = cnf `~(?x y. x + y = 2)`; +val _ = cnf' `(!x. p x) \/ (!y. r $x y)`; val _ = cnf `(!x. P x ==> (!x. Q x)) /\ ((!x. Q x \/ R x) ==> (?x. Q x /\ R x)) /\ ((?x. R x) ==> (!x. L x ==> M x)) ==> (!x. P x /\ L x ==> M x)`; -(* verbose -use "../test/large-problem.sml"; -val large_problem = time F large_problem_quotation; -val large_refute = time (Formula.Not o Formula.generalize) large_problem; -val _ = time Normalize.cnf large_refute; - -User: 0.256 System: 0.002 GC: 0.060 Real: 0.261 (* Parsing *) -User: 0.017 System: 0.000 GC: 0.000 Real: 0.017 (* Negation *) -User: 0.706 System: 0.004 GC: 0.050 Real: 0.724 (* CNF *) -*) - -(*** (* ------------------------------------------------------------------------- *) val () = SAY "Finite models"; (* ------------------------------------------------------------------------- *) -val pv = printval M.pp_model; -fun f m fm = - let - val PRINT_TIMING_INFO = false - val TIME_PER_SAMPLE = false - val RANDOM_SAMPLES = 1000 - val timex_fn = if PRINT_TIMING_INFO then timed_many else timed - val timey_fn = if PRINT_TIMING_INFO then timed_many else timed - val (tx,i) = timex_fn (M.checkn m fm) RANDOM_SAMPLES - val tx = if TIME_PER_SAMPLE then tx / Real.fromInt RANDOM_SAMPLES else tx - val rx = Real.round (100.0 * Real.fromInt i / Real.fromInt RANDOM_SAMPLES) - val (ty,(j,n)) = timey_fn (M.count m) fm - val ty = if TIME_PER_SAMPLE then ty / Real.fromInt n else ty - val ry = Real.round (100.0 * Real.fromInt j / Real.fromInt n) - val () = - if not PRINT_TIMING_INFO then () else - print ("random sample time = " ^ real_to_string tx ^ "s\n" ^ - "exhaustive search time = " ^ real_to_string ty ^ "s\n") - in - print (formula_to_string fm ^ " random sampling = " ^ int_to_string rx ^ - "% exhaustive search = " ^ int_to_string ry ^ "%\n\n") - end; +fun checkModelClause M cl = + let + val randomSamples = 100 + + fun addRandomSample {T,F} = + let + val {T = T', F = F'} = Model.checkClause {maxChecks = SOME 1} M cl + in + {T = T + T', F = F + F'} + end + + val {T,F} = funpow randomSamples addRandomSample {T = 0, F = 0} + val rx = Real.fromInt T / Real.fromInt (T + F) + + val {T,F} = Model.checkClause {maxChecks = NONE} M cl + val ry = Real.fromInt T / Real.fromInt (T + F) + in + [Formula.toString (LiteralSet.disjoin cl), + " | random sampling = " ^ percentToString rx, + " | exhaustive = " ^ percentToString ry] + end; + +local + val format = + [{leftAlign = true, padChar = #" "}, + {leftAlign = true, padChar = #" "}, + {leftAlign = true, padChar = #" "}]; +in + fun checkModel M cls = + let + val table = map (checkModelClause M) cls + + val rows = alignTable format table -val group_axioms = map Syntax.parseFormula - [`e * x = x`, `i x * x = e`, `x * y * z = x * (y * z)`]; + val () = print (join "\n" rows ^ "\n\n") + in + () + end; +end; + +fun perturbModel M cls n = + let + val N = {size = Model.size M} -val group_thms = map Syntax.parseFormula - [`x * e = x`, `x * i x = e`, `i (i x) = x`]; + fun perturbClause (fv,cl) = + let + val V = Model.randomValuation N fv + in + if Model.interpretClause M V cl then () + else Model.perturbClause M V cl + end -val m = pv (M.new M.defaults); -val () = app (f m) (group_axioms @ group_thms); -val m = pv (M.perturb group_axioms 1000 m); -val () = app (f m) (group_axioms @ group_thms); + val cls = map (fn cl => (LiteralSet.freeVars cl, cl)) cls + + fun perturbClauses () = app perturbClause cls -(* Given the multiplication, can perturbations find inverse and identity? *) -val gfix = M.map_fix (fn "*" => SOME "+" | _ => NONE) M.modulo_fix; -val gparm = M.update_fix (M.fix_merge gfix) o M.update_size (K 10); -val m = pv (M.new (gparm M.defaults)); -val () = app (f m) (group_axioms @ group_thms); -val m = pv (M.perturb group_axioms 1000 m); -val () = app (f m) (group_axioms @ group_thms); -val () = print ("e = " ^ M.term_to_string m (Syntax.parseTerm `e`) ^ "\n\n"); -val () = print ("i x =\n" ^ M.term_to_string m (Syntax.parseTerm `i x`) ^ "\n"); -val () = print ("x * y =\n" ^ M.term_to_string m (Syntax.parseTerm `x * y`) ^ "\n"); -val () = print ("x = y =\n"^M.formula_to_string m (Syntax.parseFormula `x = y`)^"\n"); + val () = funpow n perturbClauses () + in + M + end; + +val groupAxioms = + [LS[`0 + $x = $x`], + LS[`~$x + $x = 0`], + LS[`$x + $y + $z = $x + ($y + $z)`]]; + +val groupThms = + [LS[`$x + 0 = $x`], + LS[`$x + ~$x = 0`], + LS[`~~$x = $x`]]; + +fun newM fixed = Model.new {size = 8, fixed = fixed}; +val M = pvM (newM Model.basicFixed); +val () = checkModel M (groupAxioms @ groupThms); +val M = pvM (perturbModel M groupAxioms 1000); +val () = checkModel M (groupAxioms @ groupThms); +val M = pvM (newM (Model.unionFixed Model.modularFixed Model.basicFixed)); +val () = checkModel M (groupAxioms @ groupThms); (* ------------------------------------------------------------------------- *) -val () = SAY "Completion engine"; +val () = SAY "Checking the standard model"; (* ------------------------------------------------------------------------- *) -val pv = printval C.pp_completion; -fun wght ("i",1) = 0 | wght ("*",2) = 2 | wght _ = 1; -fun prec (("i",1),("i",1)) = EQUAL - | prec (_,("i",1)) = LESS - | prec (("i",1),_) = GREATER - | prec ((f,m),(g,n)) = - if m < n then LESS else if m > n then GREATER else String.compare (f,g); -val c_parm = {weight = wght, precedence = prec, precision = 3}; -val c_emp = C.empty (T.empty c_parm); -val add = try (foldl (fn (q,r) => C.add (axiom [q]) r) c_emp); +fun ppPercentClause (r,cl) = + let + val ind = 6 + + val p = percentToString r + + val fm = LiteralSet.disjoin cl + in + Print.blockProgram Print.Consistent ind + [Print.addString p, + Print.addString (nChars #" " (ind - size p)), + Formula.pp fm] + end; + +val standardModel = Model.new Model.default; + +fun checkStandardModelClause cl = + let + val {T,F} = Model.checkClause {maxChecks = SOME 1000} standardModel cl + val r = Real.fromInt T / Real.fromInt (T + F) + in + (r,cl) + end; + +val pvPCl = printval ppPercentClause + +(* Equality *) + +val cl = LS[`$x = $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~($x = $y)`,`$y = $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~($x = $y)`,`~($y = $z)`,`$x = $z`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Projections *) + +val cl = LS[`project1 $x1 = $x1`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project1 $x1 $x2 = $x1`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project2 $x1 $x2 = $x2`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project1 $x1 $x2 $x3 = $x1`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project2 $x1 $x2 $x3 = $x2`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project3 $x1 $x2 $x3 = $x3`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project1 $x1 $x2 $x3 $x4 = $x1`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project2 $x1 $x2 $x3 $x4 = $x2`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project3 $x1 $x2 $x3 $x4 = $x3`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project4 $x1 $x2 $x3 $x4 = $x4`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project1 $x1 $x2 $x3 $x4 $x5 = $x1`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project2 $x1 $x2 $x3 $x4 $x5 = $x2`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project3 $x1 $x2 $x3 $x4 $x5 = $x3`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project4 $x1 $x2 $x3 $x4 $x5 = $x4`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project5 $x1 $x2 $x3 $x4 $x5 = $x5`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project1 $x1 $x2 $x3 $x4 $x5 $x6 = $x1`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project2 $x1 $x2 $x3 $x4 $x5 $x6 = $x2`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project3 $x1 $x2 $x3 $x4 $x5 $x6 = $x3`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project4 $x1 $x2 $x3 $x4 $x5 $x6 = $x4`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project5 $x1 $x2 $x3 $x4 $x5 $x6 = $x5`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project6 $x1 $x2 $x3 $x4 $x5 $x6 = $x6`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project1 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x1`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project2 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x2`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project3 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x3`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project4 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x4`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project5 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x5`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project6 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x6`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project7 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x7`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project1 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x1`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project2 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x2`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project3 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x3`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project4 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x4`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project5 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x5`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project6 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x6`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project7 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x7`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project8 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x8`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project1 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x1`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project2 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x2`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project3 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x3`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project4 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x4`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project5 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x5`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project6 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x6`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project7 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x7`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project8 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x8`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`project9 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x9`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Arithmetic *) + +(* Zero *) +val cl = LS[`~isZero $x`,`$x = 0`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`isZero $x`,`~($x = 0)`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Positive numerals *) +val cl = LS[`0 + 1 = 1`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`1 + 1 = 2`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`2 + 1 = 3`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`3 + 1 = 4`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`4 + 1 = 5`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`5 + 1 = 6`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`6 + 1 = 7`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`7 + 1 = 8`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`8 + 1 = 9`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`9 + 1 = 10`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Negative numerals *) +val cl = LS[`~1 = negative1`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~2 = negative2`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~3 = negative3`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~4 = negative4`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~5 = negative5`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~6 = negative6`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~7 = negative7`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~8 = negative8`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~9 = negative9`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~10 = negative10`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Addition *) +val cl = LS[`0 + $x = $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$x + $y = $y + $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$x + ($y + $z) = ($x + $y) + $z`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Negation *) +val cl = LS[`~$x + $x = 0`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~~$x = $x`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Subtraction *) +val cl = LS[`$x - $y = $x + ~$y`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Successor *) +val cl = LS[`suc $x = $x + 1`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Predecessor *) +val cl = LS[`pre $x = $x - 1`]; +val _ = pvPCl (checkStandardModelClause cl); -val c = pv (add [`f (f x) = g x`]); -val c = pv (funpow 2 C.deduce c); +(* Ordering *) +val cl = LS[`$x <= $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~($x <= $y)`,`~($y <= $z)`,`$x <= $z`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~($x <= $y)`,`~($y <= $x)`,`$x = $y`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`0 <= $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~($x >= $y)`,`$y <= $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$x >= $y`,`~($y <= $x)`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$x > $y`,`$x <= $y`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~($x > $y)`,`~($x <= $y)`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$x < $y`,`$y <= $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~($x < $y)`,`~($y <= $x)`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$x = 0`,`~($x <= $y)`,`~$y <= ~$x`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Multiplication *) +val cl = LS[`1 * $x = $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`0 * $x = 0`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$x * $y = $y * $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$x * ($y * $z) = ($x * $y) * $z`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$x * ($y + $z) = ($x * $y) + ($x * $z)`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$x * ~$y = ~($x * $y)`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Division *) +val cl = LS[`$y = 0`,`$x mod $y < $y`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$y * ($x div $y) + $x mod $y = $x`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Exponentiation *) +val cl = LS[`exp $x 0 = 1`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$y = 0`,`exp $x $y = $x * exp $x (pre $y)`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Divides *) +val cl = LS[`divides $x $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~(divides $x $y)`,`~(divides $y $z)`,`divides $x $z`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~(divides $x $y)`,`~(divides $y $x)`,`$x = $y`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`divides 1 $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`divides $x 0`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Even and odd *) +val cl = LS[`even 0`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$x = 0`,`~(even (pre $x))`,`odd $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$x = 0`,`~(odd (pre $x))`,`even $x`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Sets *) + +(* The empty set *) +val cl = LS[`~member $x empty`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* The universal set *) +val cl = LS[`member $x universe`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Complement *) +val cl = LS[`member $x $y`,`member $x (complement $y)`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~(member $x $y)`,`~member $x (complement $y)`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`complement (complement $x) = $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`complement empty = universe`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`complement universe = empty`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* The subset relation *) +val cl = LS[`subset $x $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~subset $x $y`,`~subset $y $z`,`subset $x $z`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~subset $x $y`,`~subset $y $x`,`$x = $y`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`subset empty $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`subset $x universe`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~subset $x $y`,`subset (complement $y) (complement $x)`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~member $x $y`,`~subset $y $z`,`member $x $z`]; +val _ = pvPCl (checkStandardModelClause cl); -val c = pv (add [`x * y * z = x * (y * z)`, `1 * x = x`, `i x * x = 1`]); -val c = pv (funpow 44 C.deduce c); +(* Union *) +val cl = LS[`union $x $y = union $y $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`union $x (union $y $z) = union (union $x $y) $z`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`union empty $x = $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`union universe $x = universe`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`subset $x (union $x $y)`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~member $x (union $y $z)`,`member $x $y`,`member $x $z`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Intersection *) +val cl = LS[`intersect $x $y = + complement (union (complement $x) (complement $y))`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`subset (intersect $x $y) $x`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Difference *) +val cl = LS[`difference $x $y = intersect $x (complement $y)`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Symmetric difference *) +val cl = LS[`symmetricDifference $x $y = + union (difference $x $y) (difference $y $x)`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Insert *) +val cl = LS[`member $x (insert $x $y)`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Singleton *) +val cl = LS[`singleton $x = (insert $x empty)`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Cardinality *) +val cl = LS[`card empty = 0`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`member $x $y`,`card (insert $x $y) = suc (card $y)`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Lists *) + +(* Nil *) +val cl = LS[`null nil`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`~null $x`, `$x = nil`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Cons *) +val cl = LS[`~(nil = $x :: $y)`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Append *) +val cl = LS[`$x @ ($y @ $z) = ($x @ $y) @ $z`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`nil @ $x = $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`$x @ nil = $x`]; +val _ = pvPCl (checkStandardModelClause cl); -val c = pv (add [`x*y * z = x * (y*z)`, `1 * x = x`, `x * 1 = x`, `x * x = 1`]); -val c = pv (funpow 4 C.deduce c); -***) +(* Length *) +val cl = LS[`length nil = 0`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`length ($x :: $y) >= length $y`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`length ($x @ $y) >= length $x`]; +val _ = pvPCl (checkStandardModelClause cl); +val cl = LS[`length ($x @ $y) >= length $y`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* Tail *) +val cl = LS[`null $x`,`suc (length (tail $x)) = length $x`]; +val _ = pvPCl (checkStandardModelClause cl); + +(* ------------------------------------------------------------------------- *) +val () = SAY "Clauses"; +(* ------------------------------------------------------------------------- *) + +val cl = pvCl (CL[`P $x`,`P $y`]); +val _ = pvLits (Clause.largestLiterals cl); +val _ = pvCls (Clause.factor cl); +val cl = pvCl (CL[`P $x`,`~P (f $x)`]); +val _ = pvLits (Clause.largestLiterals cl); +val cl = pvCl (CL[`$x = $y`,`f $y = f $x`]); +val _ = pvLits (Clause.largestLiterals cl); +val cl = pvCl (CL[`$x = f $y`,`f $x = $y`]); +val _ = pvLits (Clause.largestLiterals cl); +val cl = pvCl (CL[`s = a`,`s = b`,`h b c`]); +val _ = pvLits (Clause.largestLiterals cl); +val cl = pvCl (CL[`a = a`,`a = b`,`h b c`]); +val _ = pvLits (Clause.largestLiterals cl); + +(* Test cases contributed by Larry Paulson *) + +local + val lnFnName = Name.fromString "ln" + and expFnName = Name.fromString "exp" + and divFnName = Name.fromString "/" + + val leRelName = Name.fromString "<"; + + fun weight na = + case na of + (n,1) => + if Name.equal n lnFnName then 500 + else if Name.equal n expFnName then 500 + else 1 + | (n,2) => + if Name.equal n divFnName then 50 + else if Name.equal n leRelName then 20 + else 1 + | _ => 1; + + val ordering = + {weight = weight, precedence = #precedence KnuthBendixOrder.default}; + + val clauseParameters = + {ordering = ordering, + orderLiterals = Clause.UnsignedLiteralOrder, + orderTerms = true}; +in + val LcpCL = mkCl clauseParameters o AX; +end; + +val cl = pvCl (LcpCL[`~($y <= (2 + (2 * $x + pow $x 2)) / 2)`, `~(0 <= $x)`, + `$y <= exp $x`]); +val _ = pvLits (Clause.largestLiterals cl); + (* ------------------------------------------------------------------------- *) val () = SAY "Syntax checking the problem sets"; (* ------------------------------------------------------------------------- *) @@ -625,7 +1075,7 @@ fun quot fm = let - fun f (v,s) = Subst.insert s (v, Term.Var "_") + fun f (v,s) = Subst.insert s (v,V"_") val sub = NameSet.foldl f Subst.empty (Formula.freeVars fm) in @@ -645,19 +1095,19 @@ val g = prep goal val p = Formula.fromString g - handle Parser.NoParse => + handle Parse.NoParse => raise Error ("failed to parse problem " ^ name) - + val () = case List.find (fn (n,_) => n = name) acc of NONE => () | SOME _ => same name val () = - case List.find (fn (_,x) => x = p) acc of NONE => () + case List.find (fn (_,x) => Formula.equal x p) acc of NONE => () | SOME (n,_) => dup n name val _ = - test_fun I g (mini_print (!Parser.lineLength) p) + test_fun equal I g (mini_print (!Print.lineLength) p) handle e => (print ("Error in problem " ^ name ^ "\n\n"); raise e) in (name,p) :: acc @@ -673,283 +1123,31 @@ val () = SAY "Parsing TPTP problems"; (* ------------------------------------------------------------------------- *) -val TPTP_DIR = "../data/problems/all"; - -fun tptp d f = +fun tptp f = let val () = print ("parsing " ^ f ^ "... ") - val goal = - case Tptp.toGoal (Tptp.read {filename = d ^ "/" ^ f}) of - Tptp.Fof goal => goal - | Tptp.Cnf prob => Problem.toClauses prob + val filename = "tptp/" ^ f ^ ".tptp" + val mapping = Tptp.defaultMapping + val goal = Tptp.goal (Tptp.read {filename = filename, mapping = mapping}) val () = print "ok\n" in pvFm goal end; -val Agatha = tptp TPTP_DIR "PUZ001-1.tptp"; -val _ = tptp "tptp" "NUMBERED_FORMULAS.tptp"; -val _ = tptp "tptp" "DEFINED_TERMS.tptp"; -val _ = tptp "tptp" "SYSTEM_TERMS.tptp"; -val _ = tptp "tptp" "QUOTED_TERMS.tptp"; -val _ = tptp "tptp" "QUOTED_TERMS_IDENTITY.tptp"; -val _ = tptp "tptp" "QUOTED_TERMS_SPECIAL.tptp"; - -(* ------------------------------------------------------------------------- *) -val () = SAY "Clauses"; -(* ------------------------------------------------------------------------- *) - -val cl = pvCl (CL[`P $x`,`P $y`]); -val _ = pvLits (Clause.largestLiterals cl); -val _ = pvCls (Clause.factor cl); -val cl = pvCl (CL[`$x = $y`,`f $y = f $x`]); -val _ = pvLits (Clause.largestLiterals cl); -val cl = pvCl (CL[`$x = f $y`,`f $x = $y`]); -val _ = pvLits (Clause.largestLiterals cl); -val cl = pvCl (CL[`s = a`,`s = b`,`h b c`]); -val _ = pvLits (Clause.largestLiterals cl); -val cl = pvCl (CL[`a = a`,`a = b`,`h b c`]); -val _ = pvLits (Clause.largestLiterals cl); +val _ = tptp "PUZ001-1"; +val _ = tptp "NUMBERED_FORMULAS"; +val _ = tptp "DEFINED_TERMS"; +val _ = tptp "SYSTEM_TERMS"; +val _ = tptp "QUOTED_TERMS"; +val _ = tptp "QUOTED_TERMS_IDENTITY"; +val _ = tptp "QUOTED_TERMS_DIFFERENT"; +val _ = tptp "QUOTED_TERMS_SPECIAL"; +val _ = tptp "RENAMING_VARIABLES"; +val _ = tptp "MIXED_PROBLEM"; +val _ = tptp "BLOCK_COMMENTS"; (* ------------------------------------------------------------------------- *) -(* Exporting problems to an external FOL datatype. *) -(* ------------------------------------------------------------------------- *) - -(* -printDepth := 10000000; - -datatype xterm = - Fun of string * xterm list -| Var of string; - -datatype xformula = - All of xterm list * xformula -| Exi of xterm list * xformula -| Iff of xformula * xformula -| Imp of xformula * xformula -| And of xformula * xformula -| Or of xformula * xformula -| Not of xformula -| Tm of xterm -| BoolT -| BoolF -| Box; (*which can be ignored entirely*) - -fun xterm (Term.Var v) = Var v - | xterm (Term.Fn (f,a)) = Fun (f, map xterm a); - -fun xformula Term.True = BoolT - | xformula Term.False = BoolF - | xformula (Term.Atom tm) = Tm (xterm tm) - | xformula (Term.Not p) = Not (xformula p) - | xformula (Term.And (p,q)) = And (xformula p, xformula q) - | xformula (Term.Or (p,q)) = Or (xformula p, xformula q) - | xformula (Term.Imp (p,q)) = Imp (xformula p, xformula q) - | xformula (Term.Iff (p,q)) = Iff (xformula p, xformula q) - | xformula fm = - (case strip_exists fm of ([],_) => - (case strip_forall fm of ([],_) => raise Fail "xformula: can't identify" - | (vs,p) => All (map Var vs, xformula p)) - | (vs,p) => Exi (map Var vs, xformula p)); - -fun xproblem {name, goal : thing quotation} = - {name = name, goal = xformula (Syntax.parseFormula goal)}; - -val xset = map xproblem; - -val xnonequality = xset Problem.nonequality; -*) - -(*** -(* ------------------------------------------------------------------------- *) -val () = SAY "Problems for provers"; -(* ------------------------------------------------------------------------- *) - -(* Non-equality *) - -val p59 = pfm (Syntax.parseFormula `(!x. P x <=> ~P (f x)) ==> ?x. P x /\ ~P (f x)`); - -val p39 = pfm (Syntax.parseFormula `~?x. !y. P y x <=> ~P y y`); - -(* Equality *) - -val p48 = (pfm o Syntax.parseFormula) - `(a = b \/ c = d) /\ (a = c \/ b = d) ==> a = d \/ b = c`; - -val cong = (pfm o Syntax.parseFormula) - `(!x. f (f (f (f (f x)))) = x) /\ (!x. f (f (f x)) = x) ==> !x. f x = x`; - -(* Impossible problems to test failure modes *) - -val square = (pfm o Syntax.parseFormula) - `sq 0 should_be_zero_here /\ - (!x. sq x x ==> sq 0 (s x)) /\ (!x y. sq x y ==> sq (s x) y) ==> - sq 0 (s (s (s (s (s (s (s (s (s (s (s (s 0))))))))))))`; - -(* ------------------------------------------------------------------------- *) -val () = SAY "Problems for solvers"; +val () = SAY "The TPTP finite model"; (* ------------------------------------------------------------------------- *) -val fib_rules = (pfm o Syntax.parseFormula) - `(!x. x + 0 = x) /\ (!x y z. x + y = z ==> x + suc y = suc z) /\ - fib 0 = 0 /\ fib (suc 0) = suc 0 /\ - (!x y z w. - fib x = y /\ fib (suc x) = z /\ y + z = w ==> fib (suc (suc x)) = w)`; - -val fib_query = pfms [Syntax.parseFormula `fib x = suc (suc y)`]; - -val agatha_rules = (pfm o Syntax.parseFormula) - `lives agatha /\ lives butler /\ lives charles /\ - (killed agatha agatha \/ killed butler agatha \/ killed charles agatha) /\ - (!x y. killed x y ==> hates x y /\ ~richer x y) /\ - (!x. hates agatha x ==> ~hates charles x) /\ - hates agatha agatha /\ hates agatha charles /\ - (!x. lives x /\ ~richer x agatha ==> hates butler x) /\ - (!x. hates agatha x ==> hates butler x) /\ - (!x. ~hates x agatha \/ ~hates x butler \/ ~hates x charles)`; - -val agatha_query1 = pfms [Syntax.parseFormula `killed x agatha`]; -val agatha_query2 = pfms [Syntax.parseFormula `~killed x agatha`]; -val agatha_query3 = pfms (map Syntax.parseFormula [`killed x agatha`, `lives x`]); - -val subset_rules = (pfm o Syntax.parseFormula) - `subset nil nil /\ - (!v x y. subset x y ==> subset (v :: x) (v :: y)) /\ - (!v x y. subset x y ==> subset x (v :: y))`; - -val subset_query1 = pfms [Syntax.parseFormula `subset x (0 :: 1 :: 2 :: nil)`]; -val subset_query2 = pfms [Syntax.parseFormula `subset (0 :: 1 :: 2 :: nil) x`]; - -val matchorder_rules = (pfm o Syntax.parseFormula) - `p 0 3 /\ - (!x. p x 4) /\ - (!x. p x 3 ==> p (s (s (s x))) 3) /\ - (!x. p (s x) 3 ==> p x 3)`; - -val matchorder_query = pfms [Syntax.parseFormula `p (s 0) 3`]; - -val ackermann_rules = (pfm o Syntax.parseFormula) - `(!x. ack 0 x = s x) /\ - (!x y. ack x (s 0) = y ==> ack (s x) 0 = y) /\ - (!x y z. ack (s x) y = w /\ ack x w = z ==> ack (s x) (s y) = z)`; - -val ackermann_query = pfms [Syntax.parseFormula `ack (s (s (s 0))) 0 = x`]; - -(* ------------------------------------------------------------------------- *) -(* Debugging Central. *) -(* ------------------------------------------------------------------------- *) - -(* -val () = Useful.trace_level := 4; -val () = Clause.show_constraint := true; - -local - open Resolution; - val to_parm = Termorder.update_precision I Termorder.defaults; - val cl_parm = {term_order = true, literal_order = true, - order_stickiness = 0, termorder_parm = to_parm}; -in - val tres_prove = (quick_prove o resolution') - ("tres", - {clause_parm = cl_parm, - set_parm = Clauseset.defaults, - sos_parm = Support.defaults}); -end; - -val prob = Syntax.parseFormula ` - (!x. x = x) /\ (!x y z v. x + y <= z + v \/ ~(x <= z) \/ ~(y <= v)) /\ - (!x. x + 0 = x) /\ (!x. x + ~x = 0) /\ - (!x y z. x + (y + z) = x + y + z) ==> - ~d <= 0 /\ c + d <= i /\ ~(c <= i) ==> F`; -val prob = Syntax.parseFormula (get std "P21"); -val prob = Syntax.parseFormula (get std "ROB002-1"); -val prob = Syntax.parseFormula (get std "KLEIN_GROUP_COMMUTATIVE"); -val prob = Syntax.parseFormula (get hol "pred_set_1"); - -(* -(cnf o Not o generalize) prob; -stop; -*) - -tres_prove prob; -stop; - -val SOME th = meson_prove prob; -print (proof_to_string' 70 (proof th)); - -stop; -*) - -(* ------------------------------------------------------------------------- *) -val () = SAY "Meson prover"; -(* ------------------------------------------------------------------------- *) - -val meson_prove_p59 = meson_prove p59; -val meson_prove_p39 = meson_prove p39; - -val meson_prove_p48 = meson_prove p48; -val meson_prove_cong = meson_prove cong; - -val meson_prove_square = meson_prove square; - -(* ------------------------------------------------------------------------- *) -val () = SAY "Meson solver"; -(* ------------------------------------------------------------------------- *) - -val meson_solve_fib = meson_solve fib_rules fib_query; - -val meson_solve_agatha1 = meson_solve agatha_rules agatha_query1; -val meson_solve_agatha2 = meson_solve agatha_rules agatha_query2; -val meson_solve_agatha3 = meson_solve agatha_rules agatha_query3; - -val meson_solve_subset1 = meson_solve subset_rules subset_query1; -val meson_solve_subset2 = meson_solve subset_rules subset_query2; - -val meson_solve_matchorder = meson_solve matchorder_rules matchorder_query; - -val meson_solve_ackermann = meson_solve ackermann_rules ackermann_query; - -(* ------------------------------------------------------------------------- *) -val () = SAY "Prolog solver"; -(* ------------------------------------------------------------------------- *) - -val prolog_solve_subset1 = prolog_solve subset_rules subset_query1; -val prolog_solve_subset2 = prolog_solve subset_rules subset_query2; - -val prolog_solve_matchorder = prolog_solve matchorder_rules matchorder_query; - -(* ------------------------------------------------------------------------- *) -val () = SAY "Resolution prover"; -(* ------------------------------------------------------------------------- *) - -val resolution_prove_p59 = resolution_prove p59; -val resolution_prove_p39 = resolution_prove p39; - -val resolution_prove_p48 = resolution_prove p48; -val resolution_prove_cong = resolution_prove cong; - -(* It would appear that resolution can't detect that this is unprovable -val resolution_prove_square = resolution_prove square; *) - -(* Printing proofs -load "Problem"; -val p21 = Syntax.parseFormula (Problem.get Problem.std "P21"); -val p21_cnf = cnf (Not (generalize p21)); -val SOME th = resolution_prove p21; -print (proof_to_string' 70 (proof th)); -*) - -(* ------------------------------------------------------------------------- *) -val () = SAY "Metis prover"; -(* ------------------------------------------------------------------------- *) - -val metis_prove_p59 = metis_prove p59; -val metis_prove_p39 = metis_prove p39; - -val metis_prove_p48 = metis_prove p48; -val metis_prove_cong = metis_prove cong; - -(* Poor delta is terribly slow at giving up on impossible problems - and in any case resolution can't detect that this is unprovable. -val metis_prove_square = metis_prove square; *) -***) +val _ = printval (Tptp.ppFixedMap Tptp.defaultMapping) Tptp.defaultFixedMap; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/auto_counterexample.ML --- a/src/Tools/auto_counterexample.ML Mon Sep 13 16:44:20 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,59 +0,0 @@ -(* Title: Tools/auto_counterexample.ML - Author: Jasmin Blanchette, TU Muenchen - -Counterexample Search Unit (do not abbreviate!). -*) - -signature AUTO_COUNTEREXAMPLE = -sig - val time_limit: int Unsynchronized.ref - - val register_tool : - string * (Proof.state -> bool * Proof.state) -> theory -> theory -end; - -structure Auto_Counterexample : AUTO_COUNTEREXAMPLE = -struct - -(* preferences *) - -val time_limit = Unsynchronized.ref 2500 - -val _ = - ProofGeneralPgip.add_preference Preferences.category_tracing - (Preferences.nat_pref time_limit - "auto-counterexample-time-limit" - "Time limit for automatic counterexample search (in milliseconds).") - - -(* configuration *) - -structure Data = Theory_Data -( - type T = (string * (Proof.state -> bool * Proof.state)) list - val empty = [] - val extend = I - fun merge data : T = AList.merge (op =) (K true) data -) - -val register_tool = Data.map o AList.update (op =) - - -(* automatic testing *) - -val _ = Context.>> (Specification.add_theorem_hook (fn interact => fn state => - case !time_limit of - 0 => state - | ms => - TimeLimit.timeLimit (Time.fromMilliseconds ms) - (fn () => - if interact andalso not (!Toplevel.quiet) then - fold_rev (fn (_, tool) => fn (true, state) => (true, state) - | (false, state) => tool state) - (Data.get (Proof.theory_of state)) (false, state) |> snd - else - state) () - handle TimeLimit.TimeOut => - (warning "Automatic counterexample search timed out."; state))) - -end; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/auto_solve.ML --- a/src/Tools/auto_solve.ML Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/auto_solve.ML Mon Sep 13 21:24:10 2010 +0200 @@ -4,16 +4,15 @@ Check whether a newly stated theorem can be solved directly by an existing theorem. Duplicate lemmas can be detected in this way. -The implementation is based in part on Berghofer and Haftmann's -quickcheck.ML. It relies critically on the Find_Theorems solves +The implementation relies critically on the Find_Theorems solves feature. *) signature AUTO_SOLVE = sig val auto : bool Unsynchronized.ref - val auto_time_limit : int Unsynchronized.ref - val limit : int Unsynchronized.ref + val max_solutions : int Unsynchronized.ref + val setup : theory -> theory end; structure Auto_Solve : AUTO_SOLVE = @@ -22,31 +21,23 @@ (* preferences *) val auto = Unsynchronized.ref false; -val auto_time_limit = Unsynchronized.ref 2500; -val limit = Unsynchronized.ref 5; - -val _ = - ProofGeneralPgip.add_preference Preferences.category_tracing - (Preferences.nat_pref auto_time_limit - "auto-solve-time-limit" - "Time limit for seeking automatic solutions (in milliseconds)."); +val max_solutions = Unsynchronized.ref 5; val _ = ProofGeneralPgip.add_preference Preferences.category_tracing (setmp_noncritical auto true (fn () => Preferences.bool_pref auto - "auto-solve" - "Try to solve newly declared lemmas with existing theorems.") ()); + "auto-solve" "Try to solve conjectures with existing theorems.") ()); (* hook *) -val _ = Context.>> (Specification.add_theorem_hook (fn int => fn state => +fun auto_solve state = let val ctxt = Proof.context_of state; val crits = [(true, Find_Theorems.Solves)]; - fun find g = snd (Find_Theorems.find_theorems ctxt (SOME g) (SOME (! limit)) false crits); + fun find g = snd (Find_Theorems.find_theorems ctxt (SOME g) (SOME (! max_solutions)) false crits); fun prt_result (goal, results) = let @@ -74,28 +65,16 @@ in if null results then NONE else SOME results end); fun go () = - let - val res = - TimeLimit.timeLimit (Time.fromMilliseconds (! auto_time_limit)) - (try seek_against_goal) (); - in - (case res of - SOME (SOME results) => - state |> Proof.goal_message - (fn () => Pretty.chunks - [Pretty.str "", - Pretty.markup Markup.hilite - (separate (Pretty.brk 0) (map prt_result results))]) - | _ => state) - end handle TimeLimit.TimeOut => (warning "Auto solve: timeout"; state); - in - if int andalso ! auto andalso not (! Toplevel.quiet) - then go () - else state - end)); + (case try seek_against_goal () of + SOME (SOME results) => + (true, state |> Proof.goal_message + (fn () => Pretty.chunks + [Pretty.str "", + Pretty.markup Markup.hilite + (separate (Pretty.brk 0) (map prt_result results))])) + | _ => (false, state)); + in if not (!auto) then (false, state) else go () end; + +val setup = Auto_Tools.register_tool ("solve", auto_solve) end; - -val auto_solve = Auto_Solve.auto; -val auto_solve_time_limit = Auto_Solve.auto_time_limit; - diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/auto_tools.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/auto_tools.ML Mon Sep 13 21:24:10 2010 +0200 @@ -0,0 +1,58 @@ +(* Title: Tools/auto_tools.ML + Author: Jasmin Blanchette, TU Muenchen + +Manager for tools that should be run automatically on newly entered conjectures. +*) + +signature AUTO_TOOLS = +sig + val time_limit: int Unsynchronized.ref + + val register_tool : + string * (Proof.state -> bool * Proof.state) -> theory -> theory +end; + +structure Auto_Tools : AUTO_TOOLS = +struct + +(* preferences *) + +val time_limit = Unsynchronized.ref 4000 + +val _ = + ProofGeneralPgip.add_preference Preferences.category_tracing + (Preferences.nat_pref time_limit + "auto-tools-time-limit" + "Time limit for automatic tools (in milliseconds).") + + +(* configuration *) + +structure Data = Theory_Data +( + type T = (string * (Proof.state -> bool * Proof.state)) list + val empty = [] + val extend = I + fun merge data : T = AList.merge (op =) (K true) data +) + +val register_tool = Data.map o AList.update (op =) + + +(* automatic testing *) + +val _ = Context.>> (Specification.add_theorem_hook (fn interact => fn state => + case !time_limit of + 0 => state + | ms => + TimeLimit.timeLimit (Time.fromMilliseconds ms) + (fn () => + if interact andalso not (!Toplevel.quiet) then + fold_rev (fn (_, tool) => fn (true, state) => (true, state) + | (false, state) => tool state) + (Data.get (Proof.theory_of state)) (false, state) |> snd + else + state) () + handle TimeLimit.TimeOut => state)) + +end; diff -r 41ce0b56d858 -r 7f11d833d65b src/Tools/quickcheck.ML --- a/src/Tools/quickcheck.ML Mon Sep 13 16:44:20 2010 +0200 +++ b/src/Tools/quickcheck.ML Mon Sep 13 21:24:10 2010 +0200 @@ -45,7 +45,7 @@ (setmp_noncritical auto true (fn () => Preferences.bool_pref auto "auto-quickcheck" - "Whether to run Quickcheck automatically.") ()); + "Run Quickcheck automatically.") ()); (* quickcheck report *) @@ -320,7 +320,7 @@ Pretty.mark Markup.hilite (pretty_counterex ctxt cex)])) state) end -val setup = Auto_Counterexample.register_tool ("quickcheck", auto_quickcheck) +val setup = Auto_Tools.register_tool ("quickcheck", auto_quickcheck) (* Isar commands *)