diff --git a/src/HOL/Tools/Sledgehammer/async_manager_legacy.ML b/src/HOL/Tools/Sledgehammer/async_manager_legacy.ML --- a/src/HOL/Tools/Sledgehammer/async_manager_legacy.ML +++ b/src/HOL/Tools/Sledgehammer/async_manager_legacy.ML @@ -1,155 +1,149 @@ (* Title: HOL/Tools/Sledgehammer/async_manager_legacy.ML Author: Fabian Immler, TU Muenchen Author: Makarius Author: Jasmin Blanchette, TU Muenchen Central manager for asynchronous diagnosis tool threads. Proof General legacy! *) signature ASYNC_MANAGER_LEGACY = sig val thread : string -> Time.time -> Time.time -> string * string -> (unit -> bool * string) -> unit val has_running_threads : string -> bool end; structure Async_Manager_Legacy : ASYNC_MANAGER_LEGACY = struct fun make_thread interrupts body = - Thread.fork - (fn () => - Runtime.debugging NONE body () handle exn => - if Exn.is_interrupt exn then () - else writeln ("## INTERNAL ERROR ##\n" ^ Runtime.exn_message exn), - Isabelle_Thread.attributes - {name = "async_manager", stack_limit = NONE, interrupts = interrupts}); + Isabelle_Thread.fork {name = "async_manager", stack_limit = NONE, interrupts = interrupts} body; fun implode_message (workers, work) = space_implode " " (Try.serial_commas "and" workers) ^ work structure Thread_Heap = Heap ( - type elem = Time.time * Thread.thread + type elem = Time.time * Isabelle_Thread.T fun ord ((a, _), (b, _)) = Time.compare (a, b) ) -fun lookup_thread xs = AList.lookup Thread.equal xs -fun delete_thread xs = AList.delete Thread.equal xs -fun update_thread xs = AList.update Thread.equal xs +fun lookup_thread xs = AList.lookup Isabelle_Thread.equal xs +fun delete_thread xs = AList.delete Isabelle_Thread.equal xs +fun update_thread xs = AList.update Isabelle_Thread.equal xs type state = - {manager: Thread.thread option, + {manager: Isabelle_Thread.T option, timeout_heap: Thread_Heap.T, active: - (Thread.thread + (Isabelle_Thread.T * (string * Time.time * Time.time * (string * string))) list, - canceling: (Thread.thread * (string * Time.time * (string * string))) list, + canceling: (Isabelle_Thread.T * (string * Time.time * (string * string))) list, messages: (bool * (string * (string * string))) list} fun make_state manager timeout_heap active canceling messages : state = {manager = manager, timeout_heap = timeout_heap, active = active, canceling = canceling, messages = messages} val global_state = Synchronized.var "async_manager" (make_state NONE Thread_Heap.empty [] [] []) fun unregister (urgent, message) thread = Synchronized.change global_state (fn state as {manager, timeout_heap, active, canceling, messages} => (case lookup_thread active thread of SOME (tool, _, _, desc as (worker, its_desc)) => let val active' = delete_thread thread active val now = Time.now () val canceling' = (thread, (tool, now, desc)) :: canceling val message' = (worker, its_desc ^ (if message = "" then "" else "\n" ^ message)) val messages' = (urgent, (tool, message')) :: messages in make_state manager timeout_heap active' canceling' messages' end | NONE => state)) val min_wait_time = seconds 0.3 val max_wait_time = seconds 10.0 fun print_new_messages () = Synchronized.change_result global_state (fn {manager, timeout_heap, active, canceling, messages} => messages |> List.partition (fn (urgent, _) => (null active andalso null canceling) orelse urgent) ||> make_state manager timeout_heap active canceling) |> map (fn (_, (tool, (worker, work))) => ((tool, work), worker)) |> AList.group (op =) |> List.app (fn ((_, ""), _) => () | ((tool, work), workers) => tool ^ ": " ^ implode_message (workers |> sort_distinct string_ord, work) |> writeln) fun check_thread_manager () = Synchronized.change global_state (fn state as {manager, timeout_heap, active, canceling, messages} => - if (case manager of SOME thread => Thread.isActive thread | NONE => false) then state + if (case manager of SOME thread => Isabelle_Thread.is_active thread | NONE => false) then state else let val manager = SOME (make_thread false (fn () => let fun time_limit timeout_heap = (case try Thread_Heap.min timeout_heap of NONE => Time.now () + max_wait_time | SOME (time, _) => time) (*action: find threads whose timeout is reached, and interrupt canceling threads*) fun action {manager, timeout_heap, active, canceling, messages} = let val (timeout_threads, timeout_heap') = - Thread_Heap.upto (Time.now (), Thread.self ()) timeout_heap + Thread_Heap.upto (Time.now (), Isabelle_Thread.self ()) timeout_heap in if null timeout_threads andalso null canceling then NONE else let val _ = List.app (Isabelle_Thread.interrupt_unsynchronized o #1) canceling - val canceling' = filter (Thread.isActive o #1) canceling + val canceling' = filter (Isabelle_Thread.is_active o #1) canceling val state' = make_state manager timeout_heap' active canceling' messages in SOME (map #2 timeout_threads, state') end end in while Synchronized.change_result global_state (fn state as {timeout_heap, active, canceling, messages, ...} => if null active andalso null canceling andalso null messages then (false, make_state NONE timeout_heap active canceling messages) else (true, state)) do (Synchronized.timed_access global_state (SOME o time_limit o #timeout_heap) action |> these |> List.app (unregister (false, "Timed out")); print_new_messages (); (* give threads some time to respond to interrupt *) OS.Process.sleep min_wait_time) end)) in make_state manager timeout_heap active canceling messages end) fun register tool birth_time death_time desc thread = (Synchronized.change global_state (fn {manager, timeout_heap, active, canceling, messages} => let val timeout_heap' = Thread_Heap.insert (death_time, thread) timeout_heap val active' = update_thread (thread, (tool, birth_time, death_time, desc)) active val state' = make_state manager timeout_heap' active' canceling messages in state' end); check_thread_manager ()) fun thread tool birth_time death_time desc f = (make_thread true (fn () => let - val self = Thread.self () + val self = Isabelle_Thread.self () val _ = register tool birth_time death_time desc self in unregister (f ()) self end); ()) fun has_running_threads tool = exists (fn (_, (tool', _, _, _)) => tool' = tool) (#active (Synchronized.value global_state)) end; diff --git a/src/Pure/Concurrent/event_timer.ML b/src/Pure/Concurrent/event_timer.ML --- a/src/Pure/Concurrent/event_timer.ML +++ b/src/Pure/Concurrent/event_timer.ML @@ -1,181 +1,181 @@ (* Title: Pure/Concurrent/event_timer.ML Author: Makarius Initiate event after given point in time. Note: events are run as synchronized action within a dedicated thread and should finish quickly without further ado. *) signature EVENT_TIMER = sig type request val eq_request: request * request -> bool val request: {physical: bool} -> Time.time -> (unit -> unit) -> request val cancel: request -> bool val future: {physical: bool} -> Time.time -> unit future val shutdown: unit -> unit end; structure Event_Timer: EVENT_TIMER = struct (* type request *) datatype request = Request of {id: int, gc_start: Time.time option, time: Time.time, event: unit -> unit}; val new_id = Counter.make (); fun new_request physical time event = Request { id = new_id (), gc_start = if physical then NONE else SOME (ML_Heap.gc_now ()), time = time, event = event}; fun eq_request (Request {id = id1, ...}, Request {id = id2, ...}) = id1 = id2; fun request_gc_start (Request {gc_start, ...}) = gc_start; fun request_time (Request {time, ...}) = time; fun request_event (Request {event, ...}) = event; (* type requests *) structure Requests = Table(type key = Time.time val ord = Time.compare); type requests = request list Requests.table; fun add req_time req (requests: requests) = Requests.cons_list (req_time, req) requests; fun del req (requests: requests) = let val old = requests |> Requests.get_first (fn (key, reqs) => reqs |> get_first (fn req' => if eq_request (req', req) then SOME (key, req') else NONE)); in (case old of NONE => (false, requests) | SOME req' => (true, Requests.remove_list eq_request req' requests)) end; fun next_time (requests: requests) = Option.map #1 (Requests.min requests); fun next_event (now, gc_now) (requests: requests) = (case Requests.min requests of NONE => NONE | SOME (req_time, reqs) => if now < req_time then NONE else let val (reqs', req) = split_last reqs; val requests' = if null reqs' then Requests.delete req_time requests else Requests.update (req_time, reqs') requests; val req_time' = (case request_gc_start req of NONE => req_time | SOME gc_start => request_time req + (gc_now - gc_start)); in if now < req_time' then (fn () => (), add req_time' req requests') else (request_event req, requests') end |> SOME); (* global state *) datatype status = Normal | Shutdown_Req | Shutdown_Ack; datatype state = - State of {requests: requests, status: status, manager: Thread.thread option}; + State of {requests: requests, status: status, manager: Isabelle_Thread.T option}; fun make_state (requests, status, manager) = State {requests = requests, status = status, manager = manager}; val normal_state = make_state (Requests.empty, Normal, NONE); val shutdown_ack_state = make_state (Requests.empty, Shutdown_Ack, NONE); fun is_shutdown s (State {requests, status, manager}) = Requests.is_empty requests andalso status = s andalso is_none manager; fun is_shutdown_req (State {requests, status, ...}) = Requests.is_empty requests andalso status = Shutdown_Req; val state = Synchronized.var "Event_Timer.state" normal_state; (* manager thread *) fun manager_loop () = if Synchronized.timed_access state (fn State {requests, ...} => next_time requests) (fn st as State {requests, status, manager} => (case next_event (Time.now (), ML_Heap.gc_now ()) requests of SOME (event, requests') => let val _ = Exn.capture event (); val state' = make_state (requests', status, manager); in SOME (true, state') end | NONE => if is_shutdown_req st then SOME (false, shutdown_ack_state) else NONE)) <> SOME false then manager_loop () else (); fun manager_check manager = - if is_some manager andalso Thread.isActive (the manager) then manager + if is_some manager andalso Isabelle_Thread.is_active (the manager) then manager else SOME (Isabelle_Thread.fork {name = "event_timer", stack_limit = NONE, interrupts = false} manager_loop); fun shutdown () = Thread_Attributes.uninterruptible (fn restore_attributes => fn () => if Synchronized.change_result state (fn st as State {requests, manager, ...} => if is_shutdown Normal st then (false, st) else if is_shutdown Shutdown_Ack st orelse is_shutdown_req st then raise Fail "Concurrent attempt to shutdown event timer" else (true, make_state (requests, Shutdown_Req, manager_check manager))) then restore_attributes (fn () => Synchronized.guarded_access state (fn st => if is_shutdown Shutdown_Ack st then SOME ((), normal_state) else NONE)) () handle exn => if Exn.is_interrupt exn then Synchronized.change state (fn State {requests, manager, ...} => make_state (requests, Normal, manager)) else () else ()) (); (* main operations *) fun request {physical} time event = Synchronized.change_result state (fn State {requests, status, manager} => let val req = new_request physical time event; val requests' = add time req requests; val manager' = manager_check manager; in (req, make_state (requests', status, manager')) end); fun cancel req = Synchronized.change_result state (fn State {requests, status, manager} => let val (canceled, requests') = del req requests; val manager' = manager_check manager; in (canceled, make_state (requests', status, manager')) end); fun future physical time = Thread_Attributes.uninterruptible (fn _ => fn () => let val req: request Single_Assignment.var = Single_Assignment.var "req"; fun abort () = ignore (cancel (Single_Assignment.await req)); val promise: unit future = Future.promise_name "event_timer" abort; val _ = Single_Assignment.assign req (request physical time (Future.fulfill promise)); in promise end) (); end; diff --git a/src/Pure/Concurrent/future.ML b/src/Pure/Concurrent/future.ML --- a/src/Pure/Concurrent/future.ML +++ b/src/Pure/Concurrent/future.ML @@ -1,715 +1,717 @@ (* Title: Pure/Concurrent/future.ML Author: Makarius Value-oriented parallel execution via futures and promises. *) signature FUTURE = sig type task = Task_Queue.task type group = Task_Queue.group val new_group: group option -> group val worker_task: unit -> task option val worker_group: unit -> group option val the_worker_group: unit -> group val worker_subgroup: unit -> group type 'a future val task_of: 'a future -> task val peek: 'a future -> 'a Exn.result option val is_finished: 'a future -> bool val interruptible_task: ('a -> 'b) -> 'a -> 'b val cancel_group: group -> unit val cancel: 'a future -> unit val error_message: Position.T -> (serial * string) * string option -> unit val identify_result: Position.T -> 'a Exn.result -> 'a Exn.result type params = {name: string, group: group option, deps: task list, pri: int, interrupts: bool} val default_params: params val forks: params -> (unit -> 'a) list -> 'a future list val fork: (unit -> 'a) -> 'a future val get_result: 'a future -> 'a Exn.result val join_results: 'a future list -> 'a Exn.result list val join_result: 'a future -> 'a Exn.result val joins: 'a future list -> 'a list val join: 'a future -> 'a val forked_results: {name: string, deps: Task_Queue.task list} -> (unit -> 'a) list -> 'a Exn.result list val task_context: string -> group -> ('a -> 'b) -> 'a -> 'b val enabled: unit -> bool val relevant: 'a list -> bool val proofs_enabled: int -> bool val proofs_enabled_timing: Time.time -> bool val value_result: 'a Exn.result -> 'a future val value: 'a -> 'a future val cond_forks: params -> (unit -> 'a) list -> 'a future list val map: ('a -> 'b) -> 'a future -> 'b future val promise_name: string -> (unit -> unit) -> 'a future val promise: (unit -> unit) -> 'a future val fulfill_result: 'a future -> 'a Exn.result -> unit val fulfill: 'a future -> 'a -> unit val snapshot: group list -> task list val shutdown: unit -> unit end; structure Future: FUTURE = struct (** future values **) type task = Task_Queue.task; type group = Task_Queue.group; val new_group = Task_Queue.new_group; (* identifiers *) local val worker_task_var = Thread_Data.var () : task Thread_Data.var; in fun worker_task () = Thread_Data.get worker_task_var; fun setmp_worker_task task f x = Thread_Data.setmp worker_task_var (SOME task) f x; end; val worker_group = Option.map Task_Queue.group_of_task o worker_task; fun the_worker_group () = (case worker_group () of SOME group => group | NONE => raise Fail "Missing worker thread context"); fun worker_subgroup () = new_group (worker_group ()); fun worker_joining e = (case worker_task () of NONE => e () | SOME task => Task_Queue.joining task e); fun worker_waiting deps e = (case worker_task () of NONE => e () | SOME task => Task_Queue.waiting task deps e); (* datatype future *) type 'a result = 'a Exn.result Single_Assignment.var; datatype 'a future = Value of 'a Exn.result | Future of {promised: bool, task: task, result: 'a result}; fun task_of (Value _) = Task_Queue.dummy_task | task_of (Future {task, ...}) = task; fun peek (Value res) = SOME res | peek (Future {result, ...}) = Single_Assignment.peek result; fun is_finished x = is_some (peek x); val _ = ML_system_pp (fn depth => fn pretty => fn x => (case peek x of NONE => PolyML.PrettyString "" | SOME (Exn.Exn _) => PolyML.PrettyString "" | SOME (Exn.Res y) => pretty (y, depth))); (** scheduling **) (* synchronization *) -val scheduler_event = ConditionVar.conditionVar (); -val work_available = ConditionVar.conditionVar (); -val work_finished = ConditionVar.conditionVar (); +val scheduler_event = Thread.ConditionVar.conditionVar (); +val work_available = Thread.ConditionVar.conditionVar (); +val work_finished = Thread.ConditionVar.conditionVar (); local - val lock = Mutex.mutex (); + val lock = Thread.Mutex.mutex (); in fun SYNCHRONIZED name = Multithreading.synchronized name lock; fun wait cond = (*requires SYNCHRONIZED*) Multithreading.sync_wait NONE cond lock; fun wait_timeout timeout cond = (*requires SYNCHRONIZED*) Multithreading.sync_wait (SOME (Time.now () + timeout)) cond lock; fun signal cond = (*requires SYNCHRONIZED*) - ConditionVar.signal cond; + Thread.ConditionVar.signal cond; fun broadcast cond = (*requires SYNCHRONIZED*) - ConditionVar.broadcast cond; + Thread.ConditionVar.broadcast cond; end; (* global state *) val queue = Unsynchronized.ref Task_Queue.empty; val next = Unsynchronized.ref 0; -val scheduler = Unsynchronized.ref (NONE: Thread.thread option); +val scheduler = Unsynchronized.ref (NONE: Isabelle_Thread.T option); val canceled = Unsynchronized.ref ([]: group list); val do_shutdown = Unsynchronized.ref false; val max_workers = Unsynchronized.ref 0; val max_active = Unsynchronized.ref 0; val status_ticks = Unsynchronized.ref 0; val last_round = Unsynchronized.ref Time.zeroTime; val next_round = seconds 0.05; datatype worker_state = Working | Waiting | Sleeping; -val workers = Unsynchronized.ref ([]: (Thread.thread * worker_state Unsynchronized.ref) list); +val workers = Unsynchronized.ref ([]: (Isabelle_Thread.T * worker_state Unsynchronized.ref) list); fun count_workers state = (*requires SYNCHRONIZED*) fold (fn (_, state_ref) => fn i => if ! state_ref = state then i + 1 else i) (! workers) 0; (* ML statistics *) fun ML_statistics () = (*requires SYNCHRONIZED*) let val {ready, pending, running, passive, urgent} = Task_Queue.status (! queue); val workers_total = length (! workers); val workers_active = count_workers Working; val workers_waiting = count_workers Waiting; in ML_Statistics.set {tasks_ready = ready, tasks_pending = pending, tasks_running = running, tasks_passive = passive, tasks_urgent = urgent, workers_total = workers_total, workers_active = workers_active, workers_waiting = workers_waiting} end; (* cancellation primitives *) fun cancel_now group = (*requires SYNCHRONIZED*) let val running = Task_Queue.cancel (! queue) group; val _ = running |> List.app (fn thread => if Isabelle_Thread.is_self thread then () else Isabelle_Thread.interrupt_unsynchronized thread); in running end; fun cancel_all () = (*requires SYNCHRONIZED*) let val (groups, threads) = Task_Queue.cancel_all (! queue); val _ = List.app Isabelle_Thread.interrupt_unsynchronized threads; in groups end; fun cancel_later group = (*requires SYNCHRONIZED*) (Unsynchronized.change canceled (insert Task_Queue.eq_group group); broadcast scheduler_event); fun interruptible_task f x = Thread_Attributes.with_attributes (if is_some (worker_task ()) then Thread_Attributes.private_interrupts else Thread_Attributes.public_interrupts) (fn _ => f x) before Thread_Attributes.expose_interrupt (); (* worker threads *) fun worker_exec (task, jobs) = let val group = Task_Queue.group_of_task task; val valid = not (Task_Queue.is_canceled group); val ok = Task_Queue.running task (fn () => setmp_worker_task task (fn () => fold (fn job => fn ok => job valid andalso ok) jobs true) ()); val _ = if ! Multithreading.trace >= 2 then Output.try_protocol_message (Markup.task_statistics :: Task_Queue.task_statistics task) [] else (); val _ = SYNCHRONIZED "finish" (fn () => let val maximal = Unsynchronized.change_result queue (Task_Queue.finish task); val test = Exn.capture Thread_Attributes.expose_interrupt (); val _ = if ok andalso not (Exn.is_interrupt_exn test) then () else if null (cancel_now group) then () else cancel_later group; val _ = broadcast work_finished; val _ = if maximal then () else signal work_available; in () end); in () end; fun worker_wait worker_state cond = (*requires SYNCHRONIZED*) - (case AList.lookup Thread.equal (! workers) (Thread.self ()) of + (case AList.lookup Isabelle_Thread.equal (! workers) (Isabelle_Thread.self ()) of SOME state => Unsynchronized.setmp state worker_state wait cond | NONE => wait cond); fun worker_next () = (*requires SYNCHRONIZED*) if length (! workers) > ! max_workers then - (Unsynchronized.change workers (AList.delete Thread.equal (Thread.self ())); + (Unsynchronized.change workers (AList.delete Isabelle_Thread.equal (Isabelle_Thread.self ())); signal work_available; NONE) else let val urgent_only = count_workers Working > ! max_active in - (case Unsynchronized.change_result queue (Task_Queue.dequeue (Thread.self ()) urgent_only) of + (case Unsynchronized.change_result queue + (Task_Queue.dequeue (Isabelle_Thread.self ()) urgent_only) of NONE => (worker_wait Sleeping work_available; worker_next ()) | some => (signal work_available; some)) end; fun worker_loop name = (case SYNCHRONIZED name (fn () => worker_next ()) of NONE => () | SOME work => (worker_exec work; worker_loop name)); fun worker_start name = (*requires SYNCHRONIZED*) let val worker = Isabelle_Thread.fork {name = "worker", stack_limit = Isabelle_Thread.stack_limit (), interrupts = false} (fn () => worker_loop name); in Unsynchronized.change workers (cons (worker, Unsynchronized.ref Working)) end handle Fail msg => Multithreading.tracing 0 (fn () => "SCHEDULER: " ^ msg); (* scheduler *) fun scheduler_end () = (*requires SYNCHRONIZED*) (ML_statistics (); scheduler := NONE); fun scheduler_next () = (*requires SYNCHRONIZED*) let val now = Time.now (); val tick = ! last_round + next_round <= now; val _ = if tick then last_round := now else (); (* runtime status *) val _ = if tick then Unsynchronized.change status_ticks (fn i => i + 1) else (); val _ = if tick andalso ! status_ticks mod (if ! Multithreading.trace >= 1 then 2 else 5) = 0 then ML_statistics () else (); val _ = - if not tick orelse forall (Thread.isActive o #1) (! workers) then () + if not tick orelse forall (Isabelle_Thread.is_active o #1) (! workers) then () else let - val (alive, dead) = List.partition (Thread.isActive o #1) (! workers); + val (alive, dead) = List.partition (Isabelle_Thread.is_active o #1) (! workers); val _ = workers := alive; in Multithreading.tracing 0 (fn () => "SCHEDULER: disposed " ^ string_of_int (length dead) ^ " dead worker threads") end; (* worker pool adjustments *) val max_active0 = ! max_active; val max_workers0 = ! max_workers; val workers_waiting = count_workers Waiting; val m = if ! do_shutdown andalso Task_Queue.all_passive (! queue) then 0 else Multithreading.max_threads (); val _ = max_active := m; val _ = max_workers := Int.max (2 * m, if workers_waiting > 0 then workers_waiting + 1 else 0); val missing = ! max_workers - length (! workers); val _ = if missing > 0 then funpow missing (fn () => ignore (worker_start ("worker " ^ string_of_int (Unsynchronized.inc next)))) () else (); val _ = if ! max_active = max_active0 andalso ! max_workers = max_workers0 then () else signal work_available; (* canceled groups *) val _ = if null (! canceled) then () else (Multithreading.tracing 1 (fn () => string_of_int (length (! canceled)) ^ " canceled groups"); Unsynchronized.change canceled (filter_out (null o cancel_now)); signal work_available); (* delay loop *) val _ = Exn.release (wait_timeout next_round scheduler_event); (* shutdown *) val continue = not (! do_shutdown andalso null (! workers)); val _ = if continue then () else scheduler_end (); val _ = broadcast scheduler_event; in continue end handle exn => (Multithreading.tracing 1 (fn () => "SCHEDULER: " ^ General.exnMessage exn); if Exn.is_interrupt exn then (List.app cancel_later (cancel_all ()); signal work_available; true) else (scheduler_end (); Exn.reraise exn)); fun scheduler_loop () = (while Thread_Attributes.with_attributes (Thread_Attributes.sync_interrupts Thread_Attributes.public_interrupts) (fn _ => SYNCHRONIZED "scheduler" (fn () => scheduler_next ())) do (); last_round := Time.zeroTime); fun scheduler_active () = (*requires SYNCHRONIZED*) - (case ! scheduler of NONE => false | SOME thread => Thread.isActive thread); + (case ! scheduler of NONE => false | SOME thread => Isabelle_Thread.is_active thread); fun scheduler_check () = (*requires SYNCHRONIZED*) (do_shutdown := false; if scheduler_active () then () else scheduler := SOME (Isabelle_Thread.fork {name = "scheduler", stack_limit = NONE, interrupts = false} scheduler_loop)); (** futures **) (* cancel *) fun cancel_group_unsynchronized group = (*requires SYNCHRONIZED*) let val _ = if null (cancel_now group) then () else cancel_later group; val _ = signal work_available; val _ = scheduler_check (); in () end; fun cancel_group group = SYNCHRONIZED "cancel_group" (fn () => cancel_group_unsynchronized group); fun cancel x = cancel_group (Task_Queue.group_of_task (task_of x)); (* results *) fun error_message pos ((serial, msg), exec_id) = Position.setmp_thread_data pos (fn () => let val id = Position.id_of pos in if is_none id orelse is_none exec_id orelse id = exec_id then Output.error_message' (serial, msg) else () end) (); fun identify_result pos res = res |> Exn.map_exn (fn exn => let val exec_id = (case Position.id_of pos of NONE => [] | SOME id => [(Markup.exec_idN, id)]) in Par_Exn.identify exec_id exn end); fun assign_result group result res = let val _ = Single_Assignment.assign result res handle exn as Fail _ => (case Single_Assignment.peek result of SOME (Exn.Exn e) => Exn.reraise (if Exn.is_interrupt e then e else exn) | _ => Exn.reraise exn); val ok = (case the (Single_Assignment.peek result) of Exn.Exn exn => (SYNCHRONIZED "cancel" (fn () => Task_Queue.cancel_group group exn); false) | Exn.Res _ => true); in ok end; (* future jobs *) fun future_job group atts (e: unit -> 'a) = let val result = Single_Assignment.var "future" : 'a result; val pos = Position.thread_data (); val context = Context.get_generic_context (); fun job ok = let val res = if ok then Exn.capture (fn () => Thread_Attributes.with_attributes atts (fn _ => (Position.setmp_thread_data pos o Context.setmp_generic_context context) e ())) () else Exn.interrupt_exn; in assign_result group result (identify_result pos res) end; in (result, job) end; (* fork *) type params = {name: string, group: group option, deps: task list, pri: int, interrupts: bool}; val default_params: params = {name = "", group = NONE, deps = [], pri = 0, interrupts = true}; fun forks ({name, group, deps, pri, interrupts}: params) es = if null es then [] else let val grp = (case group of NONE => worker_subgroup () | SOME grp => grp); fun enqueue e queue = let val atts = if interrupts then Thread_Attributes.private_interrupts else Thread_Attributes.no_interrupts; val (result, job) = future_job grp atts e; val (task, queue') = Task_Queue.enqueue name grp deps pri job queue; val future = Future {promised = false, task = task, result = result}; in (future, queue') end; in SYNCHRONIZED "enqueue" (fn () => let val (futures, queue') = fold_map enqueue es (! queue); val _ = queue := queue'; val minimal = forall (not o Task_Queue.known_task queue') deps; val _ = if minimal then signal work_available else (); val _ = scheduler_check (); in futures end) end; fun fork e = (singleton o forks) {name = "fork", group = NONE, deps = [], pri = 0, interrupts = true} e; (* join *) fun get_result x = (case peek x of NONE => Exn.Exn (Fail "Unfinished future") | SOME res => if Exn.is_interrupt_exn res then (case Task_Queue.group_status (Task_Queue.group_of_task (task_of x)) of [] => res | exns => Exn.Exn (Par_Exn.make exns)) else res); local fun join_next atts deps = (*requires SYNCHRONIZED*) if null deps then NONE else - (case Unsynchronized.change_result queue (Task_Queue.dequeue_deps (Thread.self ()) deps) of + (case Unsynchronized.change_result queue + (Task_Queue.dequeue_deps (Isabelle_Thread.self ()) deps) of (NONE, []) => NONE | (NONE, deps') => (worker_waiting deps' (fn () => Thread_Attributes.with_attributes atts (fn _ => Exn.release (worker_wait Waiting work_finished))); join_next atts deps') | (SOME work, deps') => SOME (work, deps')); fun join_loop atts deps = (case SYNCHRONIZED "join" (fn () => join_next atts deps) of NONE => () | SOME (work, deps') => (worker_joining (fn () => worker_exec work); join_loop atts deps')); in fun join_results xs = let val _ = if forall is_finished xs then () else if is_some (worker_task ()) then Thread_Attributes.with_attributes Thread_Attributes.no_interrupts (fn orig_atts => join_loop orig_atts (map task_of xs)) else xs |> List.app (fn Value _ => () | Future {result, ...} => ignore (Single_Assignment.await result)); in map get_result xs end; end; fun join_result x = singleton join_results x; fun joins xs = Par_Exn.release_all (join_results xs); fun join x = Exn.release (join_result x); (* forked results: nested parallel evaluation *) fun forked_results {name, deps} es = Thread_Attributes.uninterruptible (fn restore_attributes => fn () => let val (group, pri) = (case worker_task () of SOME task => (new_group (SOME (Task_Queue.group_of_task task)), Task_Queue.pri_of_task task) | NONE => (new_group NONE, 0)); val futures = forks {name = name, group = SOME group, deps = deps, pri = pri, interrupts = true} es; in restore_attributes join_results futures handle exn => (if Exn.is_interrupt exn then cancel_group group else (); Exn.reraise exn) end) (); (* task context for running thread *) fun task_context name group f x = Thread_Attributes.with_attributes Thread_Attributes.no_interrupts (fn orig_atts => let val (result, job) = future_job group orig_atts (fn () => f x); val task = SYNCHRONIZED "enroll" (fn () => - Unsynchronized.change_result queue (Task_Queue.enroll (Thread.self ()) name group)); + Unsynchronized.change_result queue (Task_Queue.enroll (Isabelle_Thread.self ()) name group)); val _ = worker_exec (task, [job]); in (case Single_Assignment.peek result of NONE => raise Fail "Missing task context result" | SOME res => Exn.release res) end); (* scheduling parameters *) fun enabled () = let val threads = Multithreading.max_threads () in threads > 1 andalso let val lim = threads * Options.default_int "parallel_limit" in lim <= 0 orelse SYNCHRONIZED "enabled" (fn () => Task_Queue.total_jobs (! queue) < lim) end end; val relevant = (fn [] => false | [_] => false | _ => enabled ()); fun proofs_enabled n = ! Multithreading.parallel_proofs >= n andalso is_some (worker_task ()) andalso enabled (); fun proofs_enabled_timing t = proofs_enabled 1 andalso Time.toReal t >= Options.default_real "parallel_subproofs_threshold"; (* fast-path operations -- bypass task queue if possible *) fun value_result (res: 'a Exn.result) = let val task = Task_Queue.dummy_task; val group = Task_Queue.group_of_task task; val result = Single_Assignment.var "value" : 'a result; val _ = assign_result group result (identify_result (Position.thread_data ()) res); in Future {promised = false, task = task, result = result} end; fun value x = value_result (Exn.Res x); fun cond_forks args es = if enabled () then forks args es else map (fn e => value_result (Exn.interruptible_capture e ())) es; fun map_future f x = if is_finished x then value_result (Exn.interruptible_capture (f o join) x) else let val task = task_of x; val group = Task_Queue.group_of_task task; val (result, job) = future_job group Thread_Attributes.private_interrupts (fn () => f (join x)); val extended = SYNCHRONIZED "extend" (fn () => (case Task_Queue.extend task job (! queue) of SOME queue' => (queue := queue'; true) | NONE => false)); in if extended then Future {promised = false, task = task, result = result} else (singleton o cond_forks) {name = "map_future", group = SOME group, deps = [task], pri = Task_Queue.pri_of_task task, interrupts = true} (fn () => f (join x)) end; (* promised futures -- fulfilled by external means *) fun promise_name name abort : 'a future = let val group = worker_subgroup (); val result = Single_Assignment.var "promise" : 'a result; fun assign () = assign_result group result Exn.interrupt_exn handle Fail _ => true | exn => if Exn.is_interrupt exn then raise Fail "Concurrent attempt to fulfill promise" else Exn.reraise exn; fun job () = Thread_Attributes.with_attributes Thread_Attributes.no_interrupts (fn _ => Exn.release (Exn.capture assign () before abort ())); val task = SYNCHRONIZED "enqueue_passive" (fn () => Unsynchronized.change_result queue (Task_Queue.enqueue_passive group name job)); in Future {promised = true, task = task, result = result} end; fun promise abort = promise_name "passive" abort; fun fulfill_result (Future {promised = true, task, result}) res = let val group = Task_Queue.group_of_task task; val pos = Position.thread_data (); fun job ok = assign_result group result (if ok then identify_result pos res else Exn.interrupt_exn); val _ = Thread_Attributes.with_attributes Thread_Attributes.no_interrupts (fn _ => let val passive_job = SYNCHRONIZED "fulfill_result" (fn () => Unsynchronized.change_result queue - (Task_Queue.dequeue_passive (Thread.self ()) task)); + (Task_Queue.dequeue_passive (Isabelle_Thread.self ()) task)); in (case passive_job of SOME true => worker_exec (task, [job]) | SOME false => () | NONE => ignore (job (not (Task_Queue.is_canceled group)))) end); val _ = if is_some (Single_Assignment.peek result) then () else worker_waiting [task] (fn () => ignore (Single_Assignment.await result)); in () end | fulfill_result _ _ = raise Fail "Not a promised future"; fun fulfill x res = fulfill_result x (Exn.Res res); (* snapshot: current tasks of groups *) fun snapshot [] = [] | snapshot groups = SYNCHRONIZED "snapshot" (fn () => Task_Queue.group_tasks (! queue) groups); (* shutdown *) fun shutdown () = if is_some (worker_task ()) then raise Fail "Cannot shutdown while running as worker thread" else SYNCHRONIZED "shutdown" (fn () => while scheduler_active () do (do_shutdown := true; Multithreading.tracing 1 (fn () => "SHUTDOWN: wait"); wait_timeout next_round scheduler_event)); (*final declarations of this structure!*) val map = map_future; end; type 'a future = 'a Future.future; diff --git a/src/Pure/Concurrent/isabelle_thread.ML b/src/Pure/Concurrent/isabelle_thread.ML --- a/src/Pure/Concurrent/isabelle_thread.ML +++ b/src/Pure/Concurrent/isabelle_thread.ML @@ -1,80 +1,108 @@ (* Title: Pure/Concurrent/isabelle_thread.ML Author: Makarius Isabelle-specific thread management. *) signature ISABELLE_THREAD = sig - val is_self: Thread.thread -> bool - val get_name: unit -> string + type T + val get_thread: T -> Thread.Thread.thread + val get_name: T -> string + val get_id: T -> int + val equal: T * T -> bool + val print: T -> string + val self: unit -> T + val is_self: T -> bool val stack_limit: unit -> int option type params = {name: string, stack_limit: int option, interrupts: bool} - val attributes: params -> Thread.threadAttribute list - val fork: params -> (unit -> unit) -> Thread.thread - val join: Thread.thread -> unit - val interrupt_unsynchronized: Thread.thread -> unit + val attributes: params -> Thread.Thread.threadAttribute list + val fork: params -> (unit -> unit) -> T + val is_active: T -> bool + val join: T -> unit + val interrupt_unsynchronized: T -> unit end; structure Isabelle_Thread: ISABELLE_THREAD = struct +(* abstract type *) + +abstype T = T of {thread: Thread.Thread.thread, name: string, id: int} +with + val make = T; + fun dest (T args) = args; +end; + +val get_thread = #thread o dest; +val get_name = #name o dest; +val get_id = #id o dest; + +val equal = Thread.Thread.equal o apply2 get_thread; + +fun print t = + (case get_name t of "" => "ML" | a => "Isabelle." ^ a) ^ + "-" ^ string_of_int (get_id t); + + (* self *) -fun is_self thread = Thread.equal (Thread.self (), thread); - - -(* unique name *) +val make_id = Counter.make (); local - val name_var = Thread_Data.var () : string Thread_Data.var; - val count = Counter.make (); + val self_var = Thread_Data.var () : T Thread_Data.var; in -fun get_name () = - (case Thread_Data.get name_var of - SOME name => name - | NONE => raise Fail "Isabelle-specific thread required"); +fun set_self t = Thread_Data.put self_var (SOME t); -fun set_name base = - Thread_Data.put name_var (SOME ("Isabelle." ^ base ^ "-" ^ string_of_int (count ()))); +fun self () = + (case Thread_Data.get self_var of + SOME t => t + | NONE => + let val t = make {thread = Thread.Thread.self (), name = "", id = make_id ()} + in set_self t; t end); + +fun is_self t = equal (t, self ()); end; (* fork *) fun stack_limit () = let val threads_stack_limit = Real.floor (Options.default_real "threads_stack_limit" * 1024.0 * 1024.0 * 1024.0); in if threads_stack_limit <= 0 then NONE else SOME threads_stack_limit end; type params = {name: string, stack_limit: int option, interrupts: bool}; fun attributes ({stack_limit, interrupts, ...}: params) = - Thread.MaximumMLStack stack_limit :: + Thread.Thread.MaximumMLStack stack_limit :: Thread_Attributes.convert_attributes (if interrupts then Thread_Attributes.public_interrupts else Thread_Attributes.no_interrupts); fun fork (params: params) body = - Thread.fork (fn () => - Exn.trace General.exnMessage tracing (fn () => - (set_name (#name params); body ()) - handle exn => if Exn.is_interrupt exn then () (*sic!*) else Exn.reraise exn), - attributes params); + let + val name = #name params; + val id = make_id (); + fun main () = (set_self (make {thread = Thread.Thread.self (), name = name, id = id}); body ()); + val thread = Thread.Thread.fork (main, attributes params); + in make {thread = thread, name = name, id = id} end; (* join *) -fun join thread = - while Thread.isActive thread +val is_active = Thread.Thread.isActive o get_thread; + +fun join t = + while is_active t do OS.Process.sleep (seconds 0.1); (* interrupt *) -fun interrupt_unsynchronized thread = - Thread.interrupt thread handle Thread _ => (); +fun interrupt_unsynchronized t = + Thread.Thread.interrupt (get_thread t) handle Thread.Thread _ => (); end; diff --git a/src/Pure/Concurrent/multithreading.ML b/src/Pure/Concurrent/multithreading.ML --- a/src/Pure/Concurrent/multithreading.ML +++ b/src/Pure/Concurrent/multithreading.ML @@ -1,109 +1,110 @@ (* Title: Pure/Concurrent/multithreading.ML Author: Makarius Multithreading in Poly/ML (cf. polyml/basis/Thread.sml). *) signature MULTITHREADING = sig val max_threads: unit -> int val max_threads_update: int -> unit val parallel_proofs: int ref - val sync_wait: Time.time option -> ConditionVar.conditionVar -> Mutex.mutex -> bool Exn.result + val sync_wait: Time.time option -> Thread.ConditionVar.conditionVar -> Thread.Mutex.mutex -> + bool Exn.result val trace: int ref val tracing: int -> (unit -> string) -> unit val tracing_time: bool -> Time.time -> (unit -> string) -> unit - val synchronized: string -> Mutex.mutex -> (unit -> 'a) -> 'a + val synchronized: string -> Thread.Mutex.mutex -> (unit -> 'a) -> 'a end; structure Multithreading: MULTITHREADING = struct (* max_threads *) local fun num_processors () = - (case Thread.numPhysicalProcessors () of + (case Thread.Thread.numPhysicalProcessors () of SOME n => n - | NONE => Thread.numProcessors ()); + | NONE => Thread.Thread.numProcessors ()); fun max_threads_result m = if Thread_Data.is_virtual then 1 else if m > 0 then m else Int.min (Int.max (num_processors (), 1), 8); val max_threads_state = ref 1; in fun max_threads () = ! max_threads_state; fun max_threads_update m = max_threads_state := max_threads_result m; end; (* parallel_proofs *) val parallel_proofs = ref 1; (* synchronous wait *) fun sync_wait time cond lock = Thread_Attributes.with_attributes (Thread_Attributes.sync_interrupts (Thread_Attributes.get_attributes ())) (fn _ => (case time of - SOME t => Exn.Res (ConditionVar.waitUntil (cond, lock, t)) - | NONE => (ConditionVar.wait (cond, lock); Exn.Res true)) + SOME t => Exn.Res (Thread.ConditionVar.waitUntil (cond, lock, t)) + | NONE => (Thread.ConditionVar.wait (cond, lock); Exn.Res true)) handle exn => Exn.Exn exn); (* tracing *) val trace = ref 0; fun tracing level msg = if ! trace < level then () else Thread_Attributes.uninterruptible (fn _ => fn () => (TextIO.output (TextIO.stdErr, (">>> " ^ msg () ^ "\n")); TextIO.flushOut TextIO.stdErr) handle _ (*sic*) => ()) (); fun tracing_time detailed time = tracing (if not detailed then 5 else if time >= seconds 1.0 then 1 else if time >= seconds 0.1 then 2 else if time >= seconds 0.01 then 3 else if time >= seconds 0.001 then 4 else 5); (* synchronized evaluation *) fun synchronized name lock e = Exn.release (Thread_Attributes.uninterruptible (fn restore_attributes => fn () => if ! trace > 0 then let val immediate = - if Mutex.trylock lock then true + if Thread.Mutex.trylock lock then true else let val _ = tracing 5 (fn () => name ^ ": locking ..."); val timer = Timer.startRealTimer (); - val _ = Mutex.lock lock; + val _ = Thread.Mutex.lock lock; val time = Timer.checkRealTimer timer; val _ = tracing_time true time (fn () => name ^ ": locked after " ^ Time.toString time); in false end; val result = Exn.capture (restore_attributes e) (); val _ = if immediate then () else tracing 5 (fn () => name ^ ": unlocking ..."); - val _ = Mutex.unlock lock; + val _ = Thread.Mutex.unlock lock; in result end else let - val _ = Mutex.lock lock; + val _ = Thread.Mutex.lock lock; val result = Exn.capture (restore_attributes e) (); - val _ = Mutex.unlock lock; + val _ = Thread.Mutex.unlock lock; in result end) ()); end; diff --git a/src/Pure/Concurrent/single_assignment.ML b/src/Pure/Concurrent/single_assignment.ML --- a/src/Pure/Concurrent/single_assignment.ML +++ b/src/Pure/Concurrent/single_assignment.ML @@ -1,69 +1,70 @@ (* Title: Pure/Concurrent/single_assignment.ML Author: Makarius Single-assignment variables with locking/signalling. *) signature SINGLE_ASSIGNMENT = sig type 'a var val var: string -> 'a var val peek: 'a var -> 'a option val await: 'a var -> 'a val assign: 'a var -> 'a -> unit end; structure Single_Assignment: SINGLE_ASSIGNMENT = struct datatype 'a state = Set of 'a - | Unset of {lock: Mutex.mutex, cond: ConditionVar.conditionVar}; + | Unset of {lock: Thread.Mutex.mutex, cond: Thread.ConditionVar.conditionVar}; fun init_state () = - Unset {lock = Mutex.mutex (), cond = ConditionVar.conditionVar ()}; + Unset {lock = Thread.Mutex.mutex (), cond = Thread.ConditionVar.conditionVar ()}; abstype 'a var = Var of {name: string, state: 'a state Unsynchronized.ref} with fun var name = Var {name = name, state = Unsynchronized.ref (init_state ())}; fun peek (Var {name, state}) = (case ! state of Set x => SOME x | Unset {lock, ...} => Multithreading.synchronized name lock (fn () => (case ! state of Set x => SOME x | Unset _ => NONE))); fun await (Var {name, state}) = (case ! state of Set x => x | Unset {lock, cond} => Multithreading.synchronized name lock (fn () => let fun wait () = (case ! state of Unset _ => (case Multithreading.sync_wait NONE cond lock of Exn.Res _ => wait () | Exn.Exn exn => Exn.reraise exn) | Set x => x); in wait () end)); fun assign_fail name = raise Fail ("Duplicate assignment to " ^ name); fun assign (Var {name, state}) x = (case ! state of Set _ => assign_fail name | Unset {lock, cond} => Multithreading.synchronized name lock (fn () => (case ! state of Set _ => assign_fail name | Unset _ => Thread_Attributes.uninterruptible (fn _ => fn () => - (state := Set x; RunCall.clearMutableBit state; ConditionVar.broadcast cond)) ()))); + (state := Set x; RunCall.clearMutableBit state; + Thread.ConditionVar.broadcast cond)) ()))); end; end; diff --git a/src/Pure/Concurrent/synchronized.ML b/src/Pure/Concurrent/synchronized.ML --- a/src/Pure/Concurrent/synchronized.ML +++ b/src/Pure/Concurrent/synchronized.ML @@ -1,122 +1,122 @@ (* Title: Pure/Concurrent/synchronized.ML Author: Fabian Immler and Makarius Synchronized variables. *) signature SYNCHRONIZED = sig type 'a var val var: string -> 'a -> 'a var val value: 'a var -> 'a val assign: 'a var -> 'a -> unit val timed_access: 'a var -> ('a -> Time.time option) -> ('a -> ('b * 'a) option) -> 'b option val guarded_access: 'a var -> ('a -> ('b * 'a) option) -> 'b val change_result: 'a var -> ('a -> 'b * 'a) -> 'b val change: 'a var -> ('a -> 'a) -> unit type 'a cache val cache: (unit -> 'a) -> 'a cache val cache_peek: 'a cache -> 'a option val cache_eval: 'a cache -> 'a end; structure Synchronized: SYNCHRONIZED = struct (* state variable *) datatype 'a state = Immutable of 'a - | Mutable of {lock: Mutex.mutex, cond: ConditionVar.conditionVar, content: 'a}; + | Mutable of {lock: Thread.Mutex.mutex, cond: Thread.ConditionVar.conditionVar, content: 'a}; fun init_state x = - Mutable {lock = Mutex.mutex (), cond = ConditionVar.conditionVar (), content = x}; + Mutable {lock = Thread.Mutex.mutex (), cond = Thread.ConditionVar.conditionVar (), content = x}; fun immutable_fail name = raise Fail ("Illegal access to immutable value " ^ name); abstype 'a var = Var of {name: string, state: 'a state Unsynchronized.ref} with fun var name x = Var {name = name, state = Unsynchronized.ref (init_state x)}; fun value (Var {name, state}) = (case ! state of Immutable x => x | Mutable {lock, ...} => Multithreading.synchronized name lock (fn () => (case ! state of Immutable x => x | Mutable {content, ...} => content))); fun assign (Var {name, state}) x = (case ! state of Immutable _ => immutable_fail name | Mutable {lock, cond, ...} => Multithreading.synchronized name lock (fn () => (case ! state of Immutable _ => immutable_fail name | Mutable _ => Thread_Attributes.uninterruptible (fn _ => fn () => (state := Immutable x; RunCall.clearMutableBit state; - ConditionVar.broadcast cond)) ()))); + Thread.ConditionVar.broadcast cond)) ()))); (* synchronized access *) fun timed_access (Var {name, state}) time_limit f = (case ! state of Immutable _ => immutable_fail name | Mutable {lock, cond, ...} => Multithreading.synchronized name lock (fn () => let fun try_change () = (case ! state of Immutable _ => immutable_fail name | Mutable {content = x, ...} => (case f x of NONE => (case Multithreading.sync_wait (time_limit x) cond lock of Exn.Res true => try_change () | Exn.Res false => NONE | Exn.Exn exn => Exn.reraise exn) | SOME (y, x') => Thread_Attributes.uninterruptible (fn _ => fn () => (state := Mutable {lock = lock, cond = cond, content = x'}; - ConditionVar.broadcast cond; SOME y)) ())); + Thread.ConditionVar.broadcast cond; SOME y)) ())); in try_change () end)); fun guarded_access var f = the (timed_access var (fn _ => NONE) f); (* unconditional change *) fun change_result var f = guarded_access var (SOME o f); fun change var f = change_result var (fn x => ((), f x)); end; (* cached evaluation via weak_ref *) abstype 'a cache = Cache of {expr: unit -> 'a, var: 'a Unsynchronized.weak_ref option var} with fun cache expr = Cache {expr = expr, var = var "Synchronized.cache" NONE}; fun cache_peek (Cache {var, ...}) = Option.mapPartial Unsynchronized.weak_peek (value var); fun cache_eval (Cache {expr, var}) = change_result var (fn state => (case Option.mapPartial Unsynchronized.weak_peek state of SOME result => (result, state) | NONE => let val result = expr () in (result, SOME (Unsynchronized.weak_ref result)) end)); end; end; diff --git a/src/Pure/Concurrent/task_queue.ML b/src/Pure/Concurrent/task_queue.ML --- a/src/Pure/Concurrent/task_queue.ML +++ b/src/Pure/Concurrent/task_queue.ML @@ -1,436 +1,436 @@ (* Title: Pure/Concurrent/task_queue.ML Author: Makarius Ordered queue of grouped tasks. *) signature TASK_QUEUE = sig type group val new_group: group option -> group val group_id: group -> int val eq_group: group * group -> bool val cancel_group: group -> exn -> unit val is_canceled: group -> bool val group_status: group -> exn list val str_of_group: group -> string val str_of_groups: group -> string val urgent_pri: int type task val dummy_task: task val group_of_task: task -> group val name_of_task: task -> string val pri_of_task: task -> int val eq_task: task * task -> bool val str_of_task: task -> string val str_of_task_groups: task -> string val task_statistics: task -> Properties.T val running: task -> (unit -> 'a) -> 'a val joining: task -> (unit -> 'a) -> 'a val waiting: task -> task list -> (unit -> 'a) -> 'a type queue val empty: queue val group_tasks: queue -> group list -> task list val known_task: queue -> task -> bool val all_passive: queue -> bool val total_jobs: queue -> int val status: queue -> {ready: int, pending: int, running: int, passive: int, urgent: int} - val cancel: queue -> group -> Thread.thread list - val cancel_all: queue -> group list * Thread.thread list + val cancel: queue -> group -> Isabelle_Thread.T list + val cancel_all: queue -> group list * Isabelle_Thread.T list val finish: task -> queue -> bool * queue - val enroll: Thread.thread -> string -> group -> queue -> task * queue + val enroll: Isabelle_Thread.T -> string -> group -> queue -> task * queue val enqueue_passive: group -> string -> (unit -> bool) -> queue -> task * queue val enqueue: string -> group -> task list -> int -> (bool -> bool) -> queue -> task * queue val extend: task -> (bool -> bool) -> queue -> queue option - val dequeue_passive: Thread.thread -> task -> queue -> bool option * queue - val dequeue: Thread.thread -> bool -> queue -> (task * (bool -> bool) list) option * queue - val dequeue_deps: Thread.thread -> task list -> queue -> + val dequeue_passive: Isabelle_Thread.T -> task -> queue -> bool option * queue + val dequeue: Isabelle_Thread.T -> bool -> queue -> (task * (bool -> bool) list) option * queue + val dequeue_deps: Isabelle_Thread.T -> task list -> queue -> (((task * (bool -> bool) list) option * task list) * queue) end; structure Task_Queue: TASK_QUEUE = struct val new_id = Counter.make (); (** nested groups of tasks **) (* groups *) abstype group = Group of {parent: group option, id: int, status: exn option Unsynchronized.ref} with fun make_group (parent, id, status) = Group {parent = parent, id = id, status = status}; fun new_group parent = make_group (parent, new_id (), Unsynchronized.ref NONE); fun group_id (Group {id, ...}) = id; fun eq_group (group1, group2) = group_id group1 = group_id group2; fun fold_groups f (g as Group {parent = NONE, ...}) a = f g a | fold_groups f (g as Group {parent = SOME group, ...}) a = fold_groups f group (f g a); (* group status *) local fun is_canceled_unsynchronized (Group {parent, status, ...}) = is_some (! status) orelse (case parent of NONE => false | SOME group => is_canceled_unsynchronized group); fun group_status_unsynchronized (Group {parent, status, ...}) = the_list (! status) @ (case parent of NONE => [] | SOME group => group_status_unsynchronized group); -val lock = Mutex.mutex (); +val lock = Thread.Mutex.mutex (); fun SYNCHRONIZED e = Multithreading.synchronized "group_status" lock e; in fun cancel_group (Group {status, ...}) exn = SYNCHRONIZED (fn () => Unsynchronized.change status (fn exns => SOME (Par_Exn.make (exn :: the_list exns)))); fun is_canceled group = SYNCHRONIZED (fn () => is_canceled_unsynchronized group); fun group_status group = SYNCHRONIZED (fn () => group_status_unsynchronized group); end; fun str_of_group group = (is_canceled group ? enclose "(" ")") (string_of_int (group_id group)); fun str_of_groups group = space_implode "/" (map str_of_group (rev (fold_groups cons group []))); end; (* tasks *) val urgent_pri = 1000; type timing = Time.time * Time.time * string list; (*run, wait, wait dependencies*) val timing_start = (Time.zeroTime, Time.zeroTime, []): timing; fun new_timing () = if ! Multithreading.trace < 2 then NONE else SOME (Synchronized.var "timing" timing_start); abstype task = Task of {group: group, name: string, id: int, pri: int option, timing: timing Synchronized.var option, pos: Position.T} with val dummy_task = Task {group = new_group NONE, name = "", id = 0, pri = NONE, timing = NONE, pos = Position.none}; fun new_task group name pri = Task {group = group, name = name, id = new_id (), pri = pri, timing = new_timing (), pos = Position.thread_data ()}; fun group_of_task (Task {group, ...}) = group; fun name_of_task (Task {name, ...}) = name; fun id_of_task (Task {id, ...}) = id; fun pri_of_task (Task {pri, ...}) = the_default 0 pri; fun eq_task (task1, task2) = id_of_task task1 = id_of_task task2; fun str_of_task (Task {name, id, ...}) = if name = "" then string_of_int id else string_of_int id ^ " (" ^ name ^ ")"; fun str_of_task_groups task = str_of_task task ^ " in " ^ str_of_groups (group_of_task task); fun update_timing update (Task {timing, ...}) e = Thread_Attributes.uninterruptible (fn restore_attributes => fn () => let val start = Time.now (); val result = Exn.capture (restore_attributes e) (); val t = Time.now () - start; val _ = (case timing of NONE => () | SOME var => Synchronized.change var (update t)); in Exn.release result end) (); fun task_ord (Task {id = id1, pri = pri1, ...}, Task {id = id2, pri = pri2, ...}) = prod_ord (rev_order o option_ord int_ord) int_ord ((pri1, id1), (pri2, id2)); fun task_statistics (Task {name, id, timing, pos, ...}) = let val (run, wait, wait_deps) = (case timing of NONE => timing_start | SOME var => Synchronized.value var); fun micros time = string_of_int (Time.toNanoseconds time div 1000); in [("now", Value.print_real (Time.toReal (Time.now ()))), ("task_name", name), ("task_id", Value.print_int id), ("run", micros run), ("wait", micros wait), ("wait_deps", commas wait_deps)] @ Position.properties_of pos end; end; structure Tasks = Set(type key = task val ord = task_ord); structure Task_Graph = Graph(Tasks.Key); (* timing *) fun running task = update_timing (fn t => fn (a, b, ds) => (a + t, b, ds)) task; fun joining task = update_timing (fn t => fn (a, b, ds) => (a - t, b, ds)) task; fun waiting task deps = update_timing (fn t => fn (a, b, ds) => (a - t, b + t, if ! Multithreading.trace > 0 then fold (insert (op =) o name_of_task) deps ds else ds)) task; (** queue of jobs and groups **) (* known group members *) type groups = Tasks.T Inttab.table; fun get_tasks (groups: groups) gid = the_default Tasks.empty (Inttab.lookup groups gid); fun add_task (gid, task) groups = Inttab.update (gid, Tasks.insert task (get_tasks groups gid)) groups; fun del_task (gid, task) groups = let val tasks = Tasks.remove task (get_tasks groups gid) in if Tasks.is_empty tasks then Inttab.delete_safe gid groups else Inttab.update (gid, tasks) groups end; (* job dependency graph *) datatype job = Job of (bool -> bool) list | - Running of Thread.thread | + Running of Isabelle_Thread.T | Passive of unit -> bool; type jobs = job Task_Graph.T; fun get_job (jobs: jobs) task = Task_Graph.get_node jobs task; fun set_job task job (jobs: jobs) = Task_Graph.map_node task (K job) jobs; fun add_job task dep (jobs: jobs) = Task_Graph.add_edge (dep, task) jobs handle Task_Graph.UNDEF _ => jobs; (* queue *) datatype queue = Queue of {groups: groups, jobs: jobs, urgent: int, total: int}; fun make_queue groups jobs urgent total = Queue {groups = groups, jobs = jobs, urgent = urgent, total = total}; val empty = make_queue Inttab.empty Task_Graph.empty 0 0; fun group_tasks (Queue {groups, ...}) gs = fold (fn g => fn tasks => Tasks.merge (tasks, get_tasks groups (group_id g))) gs Tasks.empty |> Tasks.dest; fun known_task (Queue {jobs, ...}) task = can (Task_Graph.get_entry jobs) task; (* job status *) fun ready_job (task, (Job list, (deps, _))) = if Task_Graph.Keys.is_empty deps then SOME (task, rev list) else NONE | ready_job (task, (Passive abort, (deps, _))) = if Task_Graph.Keys.is_empty deps andalso is_canceled (group_of_task task) then SOME (task, [fn _ => abort ()]) else NONE | ready_job _ = NONE; fun ready_job_urgent false = ready_job | ready_job_urgent true = (fn entry as (task, _) => if pri_of_task task >= urgent_pri then ready_job entry else NONE); fun active_job (task, (Running _, _)) = SOME (task, []) | active_job arg = ready_job arg; fun all_passive (Queue {jobs, ...}) = is_none (Task_Graph.get_first active_job jobs); fun total_jobs (Queue {total, ...}) = total; (* queue status *) fun status (Queue {jobs, urgent, ...}) = let val (x, y, z, w) = Task_Graph.fold (fn (_, (job, (deps, _))) => fn (x, y, z, w) => (case job of Job _ => if Task_Graph.Keys.is_empty deps then (x + 1, y, z, w) else (x, y + 1, z, w) | Running _ => (x, y, z + 1, w) | Passive _ => (x, y, z, w + 1))) jobs (0, 0, 0, 0); in {ready = x, pending = y, running = z, passive = w, urgent = urgent} end; (** task queue operations **) (* cancel -- peers and sub-groups *) fun cancel (Queue {groups, jobs, ...}) group = let val _ = cancel_group group Exn.Interrupt; val running = Tasks.fold (fn task => - (case get_job jobs task of Running thread => insert Thread.equal thread | _ => I)) + (case get_job jobs task of Running thread => insert Isabelle_Thread.equal thread | _ => I)) (get_tasks groups (group_id group)) []; in running end; fun cancel_all (Queue {jobs, ...}) = let fun cancel_job (task, (job, _)) (groups, running) = let val group = group_of_task task; val _ = cancel_group group Exn.Interrupt; in (case job of - Running t => (insert eq_group group groups, insert Thread.equal t running) + Running t => (insert eq_group group groups, insert Isabelle_Thread.equal t running) | _ => (groups, running)) end; val running = Task_Graph.fold cancel_job jobs ([], []); in running end; (* finish *) fun finish task (Queue {groups, jobs, urgent, total}) = let val group = group_of_task task; val groups' = fold_groups (fn g => del_task (group_id g, task)) group groups; val jobs' = Task_Graph.del_node task jobs; val total' = total - 1; val maximal = Task_Graph.is_maximal jobs task; in (maximal, make_queue groups' jobs' urgent total') end; (* enroll *) fun enroll thread name group (Queue {groups, jobs, urgent, total}) = let val task = new_task group name NONE; val groups' = fold_groups (fn g => add_task (group_id g, task)) group groups; val jobs' = jobs |> Task_Graph.new_node (task, Running thread); val total' = total + 1; in (task, make_queue groups' jobs' urgent total') end; (* enqueue *) fun enqueue_passive group name abort (Queue {groups, jobs, urgent, total}) = let val task = new_task group name NONE; val groups' = fold_groups (fn g => add_task (group_id g, task)) group groups; val jobs' = jobs |> Task_Graph.new_node (task, Passive abort); val total' = total + 1; in (task, make_queue groups' jobs' urgent total') end; fun enqueue name group deps pri job (Queue {groups, jobs, urgent, total}) = let val task = new_task group name (SOME pri); val groups' = fold_groups (fn g => add_task (group_id g, task)) group groups; val jobs' = jobs |> Task_Graph.new_node (task, Job [job]) |> fold (add_job task) deps; val urgent' = if pri >= urgent_pri then urgent + 1 else urgent; val total' = total + 1; in (task, make_queue groups' jobs' urgent' total') end; fun extend task job (Queue {groups, jobs, urgent, total}) = (case try (get_job jobs) task of SOME (Job list) => SOME (make_queue groups (set_job task (Job (job :: list)) jobs) urgent total) | _ => NONE); (* dequeue *) fun dequeue_passive thread task (queue as Queue {groups, jobs, urgent, total}) = (case try (get_job jobs) task of SOME (Passive _) => let val jobs' = set_job task (Running thread) jobs in (SOME true, make_queue groups jobs' urgent total) end | SOME _ => (SOME false, queue) | NONE => (NONE, queue)); fun dequeue thread urgent_only (queue as Queue {groups, jobs, urgent, total}) = if not urgent_only orelse urgent > 0 then (case Task_Graph.get_first (ready_job_urgent urgent_only) jobs of SOME (result as (task, _)) => let val jobs' = set_job task (Running thread) jobs; val urgent' = if pri_of_task task >= urgent_pri then urgent - 1 else urgent; in (SOME result, make_queue groups jobs' urgent' total) end | NONE => (NONE, queue)) else (NONE, queue); (* dequeue wrt. dynamic dependencies *) fun dequeue_deps thread deps (queue as Queue {groups, jobs, urgent, total}) = let fun ready [] rest = (NONE, rev rest) | ready (task :: tasks) rest = (case try (Task_Graph.get_entry jobs) task of NONE => ready tasks rest | SOME (_, entry) => (case ready_job (task, entry) of NONE => ready tasks (task :: rest) | some => (some, fold cons rest tasks))); fun ready_dep _ [] = NONE | ready_dep seen (task :: tasks) = if Tasks.member seen task then ready_dep seen tasks else let val entry as (_, (ds, _)) = #2 (Task_Graph.get_entry jobs task) in (case ready_job (task, entry) of NONE => ready_dep (Tasks.insert task seen) (Task_Graph.Keys.dest ds @ tasks) | some => some) end; fun result (res as (task, _)) deps' = let val jobs' = set_job task (Running thread) jobs; val urgent' = if pri_of_task task >= urgent_pri then urgent - 1 else urgent; in ((SOME res, deps'), make_queue groups jobs' urgent' total) end; in (case ready deps [] of (SOME res, deps') => result res deps' | (NONE, deps') => (case ready_dep Tasks.empty deps' of SOME res => result res deps' | NONE => ((NONE, deps'), queue))) end; (* toplevel pretty printing *) val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o Pretty.str o str_of_task); val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o Pretty.str o str_of_group); end; diff --git a/src/Pure/Concurrent/timeout.ML b/src/Pure/Concurrent/timeout.ML --- a/src/Pure/Concurrent/timeout.ML +++ b/src/Pure/Concurrent/timeout.ML @@ -1,61 +1,61 @@ (* Title: Pure/Concurrent/timeout.ML Author: Makarius Execution with relative timeout: - timeout specification < 1ms means no timeout - actual timeout is subject to system option "timeout_scale" *) signature TIMEOUT = sig exception TIMEOUT of Time.time val ignored: Time.time -> bool val scale: unit -> real val scale_time: Time.time -> Time.time val end_time: Time.time -> Time.time val apply: Time.time -> ('a -> 'b) -> 'a -> 'b val apply_physical: Time.time -> ('a -> 'b) -> 'a -> 'b val print: Time.time -> string end; structure Timeout: TIMEOUT = struct exception TIMEOUT of Time.time; fun ignored timeout = timeout < Time.fromMilliseconds 1; fun scale () = Options.default_real "timeout_scale"; fun scale_time t = Time.scale (scale ()) t; fun end_time timeout = Time.now () + scale_time timeout; fun apply' {physical, timeout} f x = if ignored timeout then f x else Thread_Attributes.with_attributes Thread_Attributes.no_interrupts (fn orig_atts => let - val self = Thread.self (); + val self = Isabelle_Thread.self (); val start = Time.now (); val request = Event_Timer.request {physical = physical} (start + scale_time timeout) (fn () => Isabelle_Thread.interrupt_unsynchronized self); val result = Exn.capture (fn () => Thread_Attributes.with_attributes orig_atts (fn _ => f x)) (); val stop = Time.now (); val was_timeout = not (Event_Timer.cancel request); val test = Exn.capture Thread_Attributes.expose_interrupt (); in if was_timeout andalso (Exn.is_interrupt_exn result orelse Exn.is_interrupt_exn test) then raise TIMEOUT (stop - start) else (Exn.release test; Exn.release result) end); fun apply timeout f x = apply' {physical = false, timeout = timeout} f x; fun apply_physical timeout f x = apply' {physical = true, timeout = timeout} f x; fun print t = "Timeout after " ^ Value.print_time t ^ "s"; end; diff --git a/src/Pure/ML/ml_init.ML b/src/Pure/ML/ml_init.ML --- a/src/Pure/ML/ml_init.ML +++ b/src/Pure/ML/ml_init.ML @@ -1,41 +1,39 @@ (* Title: Pure/ML/ml_init.ML Author: Makarius Initial ML environment. *) val seconds = Time.fromReal; val _ = List.app ML_Name_Space.forget_val ["isSome", "getOpt", "valOf", "foldl", "foldr", "print", "explode", "concat"]; val ord = SML90.ord; val chr = SML90.chr; val raw_explode = SML90.explode; val implode = String.concat; val pointer_eq = PolyML.pointerEq; val error_depth = PolyML.error_depth; -open Thread; - datatype illegal = Interrupt; structure Basic_Exn: BASIC_EXN = Exn; open Basic_Exn; structure String = struct open String; fun substring (s, i, n) = if n = 1 then String.str (String.sub (s, i)) else String.substring (s, i, n); end; structure Substring = struct open Substring; val empty = full ""; end; diff --git a/src/Pure/ML/ml_pp.ML b/src/Pure/ML/ml_pp.ML --- a/src/Pure/ML/ml_pp.ML +++ b/src/Pure/ML/ml_pp.ML @@ -1,111 +1,116 @@ (* Title: Pure/ML/ml_pp.ML Author: Makarius ML toplevel pretty-printing for logical entities. *) structure ML_PP: sig end = struct val _ = + ML_system_pp (fn _ => fn _ => fn t => + PolyML.PrettyString ("")); + +val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o Proof_Display.pp_context); val _ = ML_system_pp (fn depth => fn _ => ML_Pretty.to_polyml o Pretty.to_ML depth o Proof_Display.pp_thm Theory.get_pure); val _ = ML_system_pp (fn depth => fn _ => ML_Pretty.to_polyml o Pretty.to_ML depth o Proof_Display.pp_cterm Theory.get_pure); val _ = ML_system_pp (fn depth => fn _ => ML_Pretty.to_polyml o Pretty.to_ML depth o Proof_Display.pp_ctyp Theory.get_pure); val _ = ML_system_pp (fn depth => fn _ => ML_Pretty.to_polyml o Pretty.to_ML depth o Proof_Display.pp_typ Theory.get_pure); local fun prt_app name prt = Pretty.block [Pretty.str (name ^ " "), prt]; fun prt_apps name = Pretty.enum "," (name ^ " (") ")"; fun prt_term parens (dp: FixedInt.int) t = if dp <= 0 then Pretty.str "..." else (case t of _ $ _ => op :: (strip_comb t) |> map_index (fn (i, u) => prt_term true (dp - FixedInt.fromInt i - 1) u) |> Pretty.separate " $" |> (if parens then Pretty.enclose "(" ")" else Pretty.block) | Abs (a, T, b) => prt_apps "Abs" [Pretty.from_polyml (ML_system_pretty (a, dp - 1)), Pretty.from_polyml (ML_system_pretty (T, dp - 2)), prt_term false (dp - 3) b] | Const a => prt_app "Const" (Pretty.from_polyml (ML_system_pretty (a, dp - 1))) | Free a => prt_app "Free" (Pretty.from_polyml (ML_system_pretty (a, dp - 1))) | Var a => prt_app "Var" (Pretty.from_polyml (ML_system_pretty (a, dp - 1))) | Bound a => prt_app "Bound" (Pretty.from_polyml (ML_system_pretty (a, dp - 1)))); in val _ = ML_system_pp (fn depth => fn _ => Pretty.to_polyml o prt_term false depth); local fun prt_proof parens dp prf = if dp <= 0 then Pretty.str "..." else (case prf of _ % _ => prt_proofs parens dp prf | _ %% _ => prt_proofs parens dp prf | Abst (a, T, b) => prt_apps "Abst" [Pretty.from_polyml (ML_system_pretty (a, dp - 1)), Pretty.from_polyml (ML_system_pretty (T, dp - 2)), prt_proof false (dp - 3) b] | AbsP (a, t, b) => prt_apps "AbsP" [Pretty.from_polyml (ML_system_pretty (a, dp - 1)), Pretty.from_polyml (ML_system_pretty (t, dp - 2)), prt_proof false (dp - 3) b] | Hyp t => prt_app "Hyp" (prt_term true (dp - 1) t) | MinProof => Pretty.str "MinProof" | PBound a => prt_app "PBound" (Pretty.from_polyml (ML_system_pretty (a, dp - 1))) | PAxm a => prt_app "PAxm" (Pretty.from_polyml (ML_system_pretty (a, dp - 1))) | PClass a => prt_app "PClass" (Pretty.from_polyml (ML_system_pretty (a, dp - 1))) | Oracle a => prt_app "Oracle" (Pretty.from_polyml (ML_system_pretty (a, dp - 1))) | PThm a => prt_app "PThm" (Pretty.from_polyml (ML_system_pretty (a, dp - 1)))) and prt_proofs parens dp prf = let val (head, args) = strip_proof prf []; val prts = head (dp - 1) :: flat (map_index (fn (i, prt) => prt (dp - FixedInt.fromInt i - 2)) args); in if parens then Pretty.enclose "(" ")" prts else Pretty.block prts end and strip_proof (p % t) res = strip_proof p ((fn d => [Pretty.str " %", Pretty.brk 1, Pretty.from_polyml (ML_system_pretty (t, d))]) :: res) | strip_proof (p %% q) res = strip_proof p ((fn d => [Pretty.str " %%", Pretty.brk 1, prt_proof true d q]) :: res) | strip_proof p res = (fn d => prt_proof true d p, res); in val _ = ML_system_pp (fn depth => fn _ => fn prf => ML_Pretty.to_polyml (Pretty.to_ML ~1 (prt_proof false depth prf))); end; end; end; diff --git a/src/Pure/ROOT.ML b/src/Pure/ROOT.ML --- a/src/Pure/ROOT.ML +++ b/src/Pure/ROOT.ML @@ -1,373 +1,372 @@ (* Title: Pure/ROOT.ML Author: Makarius Main entry point for the Isabelle/Pure bootstrap process. Note: When this file is open in the Prover IDE, the ML files of Isabelle/Pure can be explored interactively. This is a separate copy of Pure within Pure: it does not affect the running logic session. *) chapter "Isabelle/Pure bootstrap"; ML_file "ML/ml_name_space.ML"; section "Bootstrap phase 0: Poly/ML setup"; ML_file "ML/ml_init.ML"; ML_file "ML/ml_system.ML"; ML_file "General/basics.ML"; ML_file "General/string.ML"; ML_file "General/vector.ML"; ML_file "General/array.ML"; ML_file "General/symbol_explode.ML"; ML_file "Concurrent/multithreading.ML"; ML_file "Concurrent/unsynchronized.ML"; ML_file "Concurrent/synchronized.ML"; ML_file "Concurrent/counter.ML"; ML_file "ML/ml_heap.ML"; ML_file "ML/ml_print_depth0.ML"; ML_file "ML/ml_pretty.ML"; ML_file "ML/ml_compiler0.ML"; section "Bootstrap phase 1: towards ML within position context"; subsection "Library of general tools"; ML_file "library.ML"; ML_file "General/print_mode.ML"; ML_file "General/integer.ML"; ML_file "General/alist.ML"; ML_file "General/table.ML"; ML_file "General/set.ML"; ML_file "General/random.ML"; ML_file "General/value.ML"; ML_file "General/properties.ML"; ML_file "General/output.ML"; ML_file "PIDE/markup.ML"; ML_file "General/utf8.ML"; ML_file "General/scan.ML"; ML_file "General/source.ML"; ML_file "General/symbol.ML"; ML_file "General/position.ML"; ML_file "General/symbol_pos.ML"; ML_file "General/input.ML"; ML_file "General/comment.ML"; ML_file "General/antiquote.ML"; ML_file "ML/ml_lex.ML"; ML_file "ML/ml_compiler1.ML"; section "Bootstrap phase 2: towards ML within Isar context"; subsection "Library of general tools"; ML_file "General/stack.ML"; ML_file "General/queue.ML"; ML_file "General/heap.ML"; ML_file "General/same.ML"; ML_file "General/ord_list.ML"; ML_file "General/balanced_tree.ML"; ML_file "General/linear_set.ML"; ML_file "General/change_table.ML"; ML_file "General/buffer.ML"; ML_file "General/pretty.ML"; ML_file "General/rat.ML"; ML_file "PIDE/xml.ML"; ML_file "General/path.ML"; ML_file "General/url.ML"; ML_file "System/bash.ML"; ML_file "General/file_stream.ML"; ML_file "General/bytes.ML"; ML_file "General/file.ML"; ML_file "General/long_name.ML"; ML_file "General/binding.ML"; ML_file "General/seq.ML"; ML_file "General/time.ML"; ML_file "General/timing.ML"; ML_file "General/sha1.ML"; ML_file "PIDE/yxml.ML"; ML_file "ML/ml_profiling.ML"; ML_file "PIDE/byte_message.ML"; ML_file "PIDE/protocol_message.ML"; ML_file "PIDE/document_id.ML"; ML_file "General/socket_io.ML"; ML_file "General/graph.ML"; ML_file "System/options.ML"; subsection "Fundamental structures"; ML_file "name.ML"; ML_file "term.ML"; ML_file "context.ML"; ML_file "config.ML"; ML_file "context_position.ML"; ML_file "soft_type_system.ML"; subsection "Concurrency within the ML runtime"; ML_file "ML/exn_properties.ML"; ML_file_no_debug "ML/exn_debugger.ML"; ML_file "Concurrent/thread_data_virtual.ML"; ML_file "Concurrent/isabelle_thread.ML"; ML_file "Concurrent/single_assignment.ML"; ML_file "Concurrent/par_exn.ML"; ML_file "Concurrent/task_queue.ML"; ML_file "Concurrent/future.ML"; ML_file "Concurrent/event_timer.ML"; ML_file "Concurrent/timeout.ML"; ML_file "Concurrent/lazy.ML"; ML_file "Concurrent/par_list.ML"; ML_file "Concurrent/mailbox.ML"; ML_file "Concurrent/cache.ML"; ML_file "PIDE/active.ML"; ML_file "Thy/export.ML"; subsection "Inner syntax"; ML_file "Syntax/type_annotation.ML"; ML_file "Syntax/term_position.ML"; ML_file "Syntax/lexicon.ML"; ML_file "Syntax/ast.ML"; ML_file "Syntax/syntax_ext.ML"; ML_file "Syntax/parser.ML"; ML_file "Syntax/syntax_trans.ML"; ML_file "Syntax/mixfix.ML"; ML_file "Syntax/printer.ML"; ML_file "Syntax/syntax.ML"; subsection "Core of tactical proof system"; ML_file "term_ord.ML"; ML_file "term_items.ML"; ML_file "term_subst.ML"; ML_file "General/completion.ML"; ML_file "General/name_space.ML"; ML_file "sorts.ML"; ML_file "type.ML"; ML_file "logic.ML"; ML_file "Syntax/simple_syntax.ML"; ML_file "net.ML"; ML_file "item_net.ML"; ML_file "envir.ML"; ML_file "consts.ML"; ML_file "term_xml.ML"; ML_file "primitive_defs.ML"; ML_file "sign.ML"; ML_file "defs.ML"; ML_file "term_sharing.ML"; ML_file "pattern.ML"; ML_file "unify.ML"; ML_file "theory.ML"; ML_file "proofterm.ML"; ML_file "thm.ML"; ML_file "cterm_items.ML"; ML_file "more_pattern.ML"; ML_file "more_unify.ML"; ML_file "more_thm.ML"; ML_file "facts.ML"; ML_file "thm_name.ML"; ML_file "global_theory.ML"; ML_file "pure_thy.ML"; ML_file "drule.ML"; ML_file "morphism.ML"; ML_file "variable.ML"; ML_file "conv.ML"; ML_file "goal_display.ML"; ML_file "tactical.ML"; ML_file "search.ML"; ML_file "tactic.ML"; ML_file "raw_simplifier.ML"; ML_file "conjunction.ML"; ML_file "assumption.ML"; subsection "Isar -- Intelligible Semi-Automated Reasoning"; (*ML support and global execution*) ML_file "ML/ml_syntax.ML"; ML_file "ML/ml_env.ML"; ML_file "ML/ml_options.ML"; ML_file "ML/ml_print_depth.ML"; ML_file_no_debug "Isar/runtime.ML"; ML_file "PIDE/execution.ML"; ML_file "ML/ml_compiler.ML"; ML_file "skip_proof.ML"; ML_file "goal.ML"; (*outer syntax*) ML_file "Isar/keyword.ML"; ML_file "Isar/token.ML"; ML_file "Isar/parse.ML"; ML_file "Thy/document_source.ML"; ML_file "Thy/thy_header.ML"; ML_file "Thy/document_marker.ML"; (*proof context*) ML_file "Isar/object_logic.ML"; ML_file "Isar/rule_cases.ML"; ML_file "Isar/auto_bind.ML"; ML_file "type_infer.ML"; ML_file "Syntax/local_syntax.ML"; ML_file "Isar/proof_context.ML"; ML_file "type_infer_context.ML"; ML_file "Syntax/syntax_phases.ML"; ML_file "Isar/args.ML"; (*theory specifications*) ML_file "Isar/local_defs.ML"; ML_file "Isar/local_theory.ML"; ML_file "Isar/entity.ML"; ML_file "PIDE/command_span.ML"; ML_file "Thy/thy_element.ML"; ML_file "Thy/markdown.ML"; ML_file "Thy/latex.ML"; (*ML with context and antiquotations*) ML_file "ML/ml_context.ML"; ML_file "ML/ml_antiquotation.ML"; ML_file "ML/ml_compiler2.ML"; ML_file "ML/ml_antiquotations.ML"; section "Bootstrap phase 3: towards theory Pure and final ML toplevel setup"; (*basic proof engine*) ML_file "par_tactical.ML"; ML_file "context_tactic.ML"; ML_file "Isar/attrib.ML"; ML_file "Isar/proof_display.ML"; ML_file "Isar/context_rules.ML"; ML_file "Isar/method.ML"; ML_file "Isar/proof.ML"; ML_file "Isar/element.ML"; ML_file "Isar/obtain.ML"; ML_file "Isar/subgoal.ML"; ML_file "Isar/calculation.ML"; (*local theories and targets*) ML_file "Isar/locale.ML"; ML_file "Isar/generic_target.ML"; ML_file "Isar/bundle.ML"; ML_file "Isar/overloading.ML"; ML_file "axclass.ML"; ML_file "Isar/class.ML"; ML_file "Isar/named_target.ML"; ML_file "Isar/expression.ML"; ML_file "Isar/interpretation.ML"; ML_file "Isar/class_declaration.ML"; ML_file "Isar/target_context.ML"; ML_file "Isar/experiment.ML"; ML_file "simplifier.ML"; ML_file "Tools/plugin.ML"; (*executable theory content*) ML_file "Isar/code.ML"; (*specifications*) ML_file "Isar/spec_rules.ML"; ML_file "Isar/specification.ML"; ML_file "Isar/parse_spec.ML"; ML_file "Isar/typedecl.ML"; (*toplevel transactions*) ML_file "Isar/proof_node.ML"; ML_file "Isar/toplevel.ML"; (*proof term operations*) ML_file "Proof/proof_rewrite_rules.ML"; ML_file "Proof/proof_syntax.ML"; ML_file "Proof/proof_checker.ML"; ML_file "Proof/extraction.ML"; (*Isabelle system*) ML_file "PIDE/protocol_command.ML"; ML_file "System/java.ML"; ML_file "System/scala.ML"; ML_file "System/process_result.ML"; ML_file "System/isabelle_system.ML"; (*theory documents*) ML_file "Thy/term_style.ML"; ML_file "Isar/outer_syntax.ML"; ML_file "Thy/document_antiquotation.ML"; ML_file "Thy/document_output.ML"; ML_file "Thy/document_antiquotations.ML"; ML_file "General/graph_display.ML"; ML_file "pure_syn.ML"; ML_file "PIDE/resources.ML"; ML_file "PIDE/command.ML"; ML_file "PIDE/query_operation.ML"; ML_file "Thy/thy_info.ML"; ML_file "thm_deps.ML"; ML_file "Thy/export_theory.ML"; ML_file "Thy/sessions.ML"; ML_file "PIDE/session.ML"; ML_file "PIDE/document.ML"; (*ML add-ons*) ML_file "ML/ml_pp.ML"; ML_file "ML/ml_thms.ML"; ML_file "ML/ml_instantiate.ML"; ML_file "ML/ml_file.ML"; ML_file "ML/ml_pid.ML"; (*theory and proof operations*) ML_file "Isar/isar_cmd.ML"; subsection "Isabelle/Isar system"; ML_file "System/command_line.ML"; ML_file "System/message_channel.ML"; ML_file "System/isabelle_process.ML"; ML_file "System/scala_compiler.ML"; ML_file "System/isabelle_tool.ML"; ML_file "Thy/bibtex.ML"; ML_file "PIDE/protocol.ML"; ML_file "General/output_primitives_virtual.ML"; subsection "Miscellaneous tools and packages for Pure Isabelle"; ML_file "General/base64.ML"; ML_file "General/xz.ML"; ML_file "General/zstd.ML"; ML_file "Tools/prismjs.ML"; ML_file "Tools/build.ML"; ML_file "Tools/named_thms.ML"; ML_file "Tools/print_operation.ML"; ML_file "Tools/rail.ML"; ML_file "Tools/rule_insts.ML"; ML_file "Tools/thy_deps.ML"; ML_file "Tools/class_deps.ML"; ML_file "Tools/find_theorems.ML"; ML_file "Tools/find_consts.ML"; ML_file "Tools/simplifier_trace.ML"; ML_file_no_debug "Tools/debugger.ML"; ML_file "Tools/named_theorems.ML"; ML_file "Tools/doc.ML"; ML_file "Tools/jedit.ML"; ML_file "Tools/ghc.ML"; ML_file "Tools/generated_files.ML"; - diff --git a/src/Pure/System/message_channel.ML b/src/Pure/System/message_channel.ML --- a/src/Pure/System/message_channel.ML +++ b/src/Pure/System/message_channel.ML @@ -1,45 +1,45 @@ (* Title: Pure/System/message_channel.ML Author: Makarius Preferably asynchronous channel for Isabelle messages. *) signature MESSAGE_CHANNEL = sig type T val shutdown: T -> unit val message: T -> string -> Properties.T -> XML.body list -> unit val make: BinIO.outstream -> T end; structure Message_Channel: MESSAGE_CHANNEL = struct datatype message = Shutdown | Message of XML.body list; -datatype T = Message_Channel of {mbox: message Mailbox.T, thread: Thread.thread}; +datatype T = Message_Channel of {mbox: message Mailbox.T, thread: Isabelle_Thread.T}; fun shutdown (Message_Channel {mbox, thread}) = (Mailbox.send mbox Shutdown; Mailbox.await_empty mbox; Isabelle_Thread.join thread); fun message (Message_Channel {mbox, ...}) name raw_props chunks = let val robust_props = map (apply2 YXML.embed_controls) raw_props; val header = [XML.Elem ((name, robust_props), [])]; in Mailbox.send mbox (Message (header :: chunks)) end; fun make stream = let val mbox = Mailbox.create (); fun loop () = Mailbox.receive NONE mbox |> dispatch and dispatch [] = loop () | dispatch (Shutdown :: _) = () | dispatch (Message chunks :: rest) = (Byte_Message.write_message_yxml stream chunks; dispatch rest); val thread = Isabelle_Thread.fork {name = "message_channel", stack_limit = NONE, interrupts = false} loop; in Message_Channel {mbox = mbox, thread = thread} end; end; diff --git a/src/Pure/Tools/debugger.ML b/src/Pure/Tools/debugger.ML --- a/src/Pure/Tools/debugger.ML +++ b/src/Pure/Tools/debugger.ML @@ -1,295 +1,296 @@ (* Title: Pure/Tools/debugger.ML Author: Makarius Interactive debugger for Isabelle/ML. *) signature DEBUGGER = sig val writeln_message: string -> unit val warning_message: string -> unit val error_message: string -> unit end; structure Debugger: DEBUGGER = struct (** global state **) (* output messages *) fun output_message kind msg = if msg = "" then () else Output.protocol_message - (Markup.debugger_output (Isabelle_Thread.get_name ())) + (Markup.debugger_output (Isabelle_Thread.print (Isabelle_Thread.self ()))) [[XML.Text (Markup.markup (kind, Markup.serial_properties (serial ())) msg)]]; val writeln_message = output_message Markup.writelnN; val warning_message = output_message Markup.warningN; val error_message = output_message Markup.errorN; fun error_wrapper e = e () handle exn => if Exn.is_interrupt exn then Exn.reraise exn else error_message (Runtime.exn_message exn); (* thread input *) val thread_input = Synchronized.var "Debugger.state" (NONE: string list Queue.T Symtab.table option); fun init_input () = Synchronized.change thread_input (K (SOME Symtab.empty)); fun exit_input () = Synchronized.change thread_input (K NONE); fun input thread_name msg = if null msg then error "Empty input" else Synchronized.change thread_input (Option.map (Symtab.map_default (thread_name, Queue.empty) (Queue.enqueue msg))); fun get_input thread_name = Synchronized.guarded_access thread_input (fn NONE => SOME ([], NONE) | SOME input => (case Symtab.lookup input thread_name of NONE => NONE | SOME queue => let val (msg, queue') = Queue.dequeue queue; val input' = if Queue.is_empty queue' then Symtab.delete_safe thread_name input else Symtab.update (thread_name, queue') input; in SOME (msg, SOME input') end)); (* global break *) local val break = Synchronized.var "Debugger.break" false; in fun is_break () = Synchronized.value break; fun set_break b = Synchronized.change break (K b); end; (** thread state **) (* stack frame during debugging *) val debugging_var = Thread_Data.var () : PolyML.DebuggerInterface.debugState list Thread_Data.var; fun get_debugging () = the_default [] (Thread_Data.get debugging_var); val is_debugging = not o null o get_debugging; fun with_debugging e = - Thread_Data.setmp debugging_var (SOME (PolyML.DebuggerInterface.debugState (Thread.self ()))) e (); + Thread_Data.setmp debugging_var + (SOME (PolyML.DebuggerInterface.debugState (Thread.Thread.self ()))) e (); fun the_debug_state thread_name index = (case get_debugging () of [] => error ("Missing debugger information for thread " ^ quote thread_name) | states => if index < 0 orelse index >= length states then error ("Bad debugger stack index " ^ signed_string_of_int index ^ " for thread " ^ quote thread_name) else nth states index); (* flags for single-stepping *) datatype stepping = Stepping of bool * int; (*stepping enabled, stack depth limit*) val stepping_var = Thread_Data.var () : stepping Thread_Data.var; fun get_stepping () = the_default (Stepping (false, ~1)) (Thread_Data.get stepping_var); fun put_stepping stepping = Thread_Data.put stepping_var (SOME (Stepping stepping)); fun is_stepping () = let - val stack = PolyML.DebuggerInterface.debugState (Thread.self ()); + val stack = PolyML.DebuggerInterface.debugState (Thread.Thread.self ()); val Stepping (stepping, depth) = get_stepping (); in stepping andalso (depth < 0 orelse length stack <= depth) end; fun continue () = put_stepping (false, ~1); fun step () = put_stepping (true, ~1); fun step_over () = put_stepping (true, length (get_debugging ())); fun step_out () = put_stepping (true, length (get_debugging ()) - 1); (** eval ML **) local val context_attempts = map ML_Lex.read ["Context.put_generic_context (SOME (Context.Theory ML_context))", "Context.put_generic_context (SOME (Context.Proof ML_context))", "Context.put_generic_context (SOME ML_context)"]; fun environment SML = if SML then ML_Env.SML else ML_Env.Isabelle; fun operations SML = if SML then ML_Env.SML_operations else ML_Env.Isabelle_operations; fun evaluate {SML, verbose} = ML_Context.eval {environment = environment SML, redirect = false, verbose = verbose, debug = SOME false, writeln = writeln_message, warning = warning_message} Position.none; fun eval_setup thread_name index SML context = context |> Context_Position.set_visible_generic false |> ML_Env.add_name_space (environment SML) (PolyML.DebuggerInterface.debugNameSpace (the_debug_state thread_name index)); fun eval_context thread_name index SML toks = let val context = Context.the_generic_context (); val context1 = if SML orelse forall (fn Antiquote.Text tok => ML_Lex.is_improper tok | _ => false) toks then context else let val context' = context |> eval_setup thread_name index SML |> ML_Context.exec (fn () => evaluate {SML = SML, verbose = true} (ML_Lex.read "val ML_context = " @ toks)); fun try_exec toks = try (ML_Context.exec (fn () => evaluate {SML = false, verbose = false} toks)) context'; in (case get_first try_exec context_attempts of SOME context2 => context2 | NONE => error "Malformed context: expected type theory, Proof.context, Context.generic") end; in context1 |> eval_setup thread_name index SML end; in fun eval thread_name index SML txt1 txt2 = let val (toks1, toks2) = apply2 (#read_source (operations SML) o Input.string) (txt1, txt2); val context = eval_context thread_name index SML toks1; in Context.setmp_generic_context (SOME context) (evaluate {SML = SML, verbose = true}) toks2 end; fun print_vals thread_name index SML txt = let val toks = #read_source (operations SML) (Input.string txt) val context = eval_context thread_name index SML toks; val space = PolyML.DebuggerInterface.debugNameSpace (the_debug_state thread_name index); fun print x = Pretty.from_polyml (PolyML.NameSpace.Values.printWithType (x, FixedInt.fromInt (ML_Print_Depth.get_print_depth ()), SOME space)); fun print_all () = #allVal (PolyML.DebuggerInterface.debugLocalNameSpace (the_debug_state thread_name index)) () |> sort_by #1 |> map (Pretty.item o single o print o #2) |> Pretty.chunks |> Pretty.string_of |> writeln_message; in Context.setmp_generic_context (SOME context) print_all () end; end; (** debugger loop **) local fun debugger_state thread_name = Output.protocol_message (Markup.debugger_state thread_name) [get_debugging () |> map (fn st => (Position.properties_of (Exn_Properties.position_of_polyml_location (PolyML.DebuggerInterface.debugLocation st)), PolyML.DebuggerInterface.debugFunction st)) |> let open XML.Encode in list (pair properties string) end]; fun debugger_command thread_name = (case get_input thread_name of [] => (continue (); false) | ["continue"] => (continue (); false) | ["step"] => (step (); false) | ["step_over"] => (step_over (); false) | ["step_out"] => (step_out (); false) | ["eval", index, SML, txt1, txt2] => (error_wrapper (fn () => eval thread_name (Value.parse_int index) (Value.parse_bool SML) txt1 txt2); true) | ["print_vals", index, SML, txt] => (error_wrapper (fn () => print_vals thread_name (Value.parse_int index) (Value.parse_bool SML) txt); true) | bad => (Output.system_message ("Debugger: bad input " ^ ML_Syntax.print_list ML_Syntax.print_string bad); true)); in fun debugger_loop thread_name = Thread_Attributes.uninterruptible (fn restore_attributes => fn () => let fun loop () = (debugger_state thread_name; if debugger_command thread_name then loop () else ()); val result = Exn.capture (restore_attributes with_debugging) loop; val _ = debugger_state thread_name; in Exn.release result end) (); end; (** protocol commands **) val _ = Protocol_Command.define "Debugger.init" (fn [] => (init_input (); PolyML.DebuggerInterface.setOnBreakPoint (SOME (fn (_, break) => if not (is_debugging ()) andalso (! break orelse is_break () orelse is_stepping ()) then - (case try Isabelle_Thread.get_name () of + (case try (Isabelle_Thread.print o Isabelle_Thread.self) () of SOME thread_name => debugger_loop thread_name | NONE => ()) else ())))); val _ = Protocol_Command.define "Debugger.exit" (fn [] => (PolyML.DebuggerInterface.setOnBreakPoint NONE; exit_input ())); val _ = Protocol_Command.define "Debugger.break" (fn [b] => set_break (Value.parse_bool b)); val _ = Protocol_Command.define "Debugger.breakpoint" (fn [node_name, id0, breakpoint0, breakpoint_state0] => let val id = Value.parse_int id0; val breakpoint = Value.parse_int breakpoint0; val breakpoint_state = Value.parse_bool breakpoint_state0; fun err () = error ("Bad exec for command " ^ Value.print_int id); in (case Document.command_exec (Document.state ()) node_name id of SOME (eval, _) => if Command.eval_finished eval then let val st = Command.eval_result_state eval; val ctxt = Toplevel.presentation_context st; in (case ML_Env.get_breakpoint (Context.Proof ctxt) breakpoint of SOME (b, _) => b := breakpoint_state | NONE => err ()) end else err () | NONE => err ()) end); val _ = Protocol_Command.define "Debugger.input" (fn thread_name :: msg => input thread_name msg); end; diff --git a/src/Tools/Metis/PortableIsabelle.sml b/src/Tools/Metis/PortableIsabelle.sml --- a/src/Tools/Metis/PortableIsabelle.sml +++ b/src/Tools/Metis/PortableIsabelle.sml @@ -1,30 +1,30 @@ (* Title: Tools/Metis/PortableIsabelle.sml Author: Jasmin Blanchette, TU Muenchen Copyright 2010 Isabelle-specific setup for Metis. Based on "src/PortablePolyml.sml". *) structure Portable :> Portable = struct val ml = "isabelle" fun pointerEqual (x : 'a, y : 'a) = pointer_eq (x, y) local - val lock = Mutex.mutex (); + val lock = Thread.Mutex.mutex (); in fun critical e () = Multithreading.synchronized "metis" lock e end; val randomWord = Random.nextWord val randomBool = Random.nextBool val randomInt = Random.nextInt val randomReal = Random.nextReal fun time f x = f x (* dummy *) end datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a diff --git a/src/Tools/Metis/metis.ML b/src/Tools/Metis/metis.ML --- a/src/Tools/Metis/metis.ML +++ b/src/Tools/Metis/metis.ML @@ -1,22013 +1,22013 @@ (* This file was generated by the "make_metis" script. The BSD License is used with Joe Leslie-Hurd's kind permission. Extract from a September 13, 2010 email written by him: 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 *) (* GENERATED FILE -- DO NOT EDIT -- GENERATED FILE -- DO NOT EDIT *) (******************************************************************) (**** Original file: src/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 Metis_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: src/Random.sml ****) (* 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 Metis_Random :> Metis_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; (**** Original file: src/Portable.sig ****) (* ========================================================================= *) (* ML COMPILER SPECIFIC FUNCTIONS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Portable = sig (* ------------------------------------------------------------------------- *) (* The ML implementation. *) (* ------------------------------------------------------------------------- *) val ml : string (* ------------------------------------------------------------------------- *) (* Pointer equality using the run-time system. *) (* ------------------------------------------------------------------------- *) val pointerEqual : 'a * 'a -> bool (* ------------------------------------------------------------------------- *) (* Marking critical sections of code. *) (* ------------------------------------------------------------------------- *) val critical : (unit -> 'a) -> unit -> 'a (* ------------------------------------------------------------------------- *) (* Generating random values. *) (* ------------------------------------------------------------------------- *) val randomBool : unit -> bool val randomInt : int -> int (* n |-> [0,n) *) val randomReal : unit -> real (* () |-> [0,1] *) val randomWord : unit -> Word.word (* ------------------------------------------------------------------------- *) (* Timing function applications. *) (* ------------------------------------------------------------------------- *) val time : ('a -> 'b) -> 'a -> 'b end (**** Original file: PortableIsabelle.sml ****) (* Title: Tools/Metis/PortableIsabelle.sml Author: Jasmin Blanchette, TU Muenchen Copyright 2010 Isabelle-specific setup for Metis. Based on "src/PortablePolyml.sml". *) structure Metis_Portable :> Metis_Portable = struct val ml = "isabelle" fun pointerEqual (x : 'a, y : 'a) = pointer_eq (x, y) local - val lock = Mutex.mutex (); + val lock = Thread.Mutex.mutex (); in fun critical e () = Multithreading.synchronized "metis" lock e end; val randomWord = Metis_Random.nextWord val randomBool = Metis_Random.nextBool val randomInt = Metis_Random.nextInt val randomReal = Metis_Random.nextReal fun time f x = f x (* dummy *) end datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a (**** Original file: src/Useful.sig ****) (* ========================================================================= *) (* ML UTILITY FUNCTIONS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Useful = sig (* ------------------------------------------------------------------------- *) (* Exceptions. *) (* ------------------------------------------------------------------------- *) exception Error of string exception Bug of string val total : ('a -> 'b) -> 'a -> 'b option val can : ('a -> 'b) -> 'a -> bool (* ------------------------------------------------------------------------- *) (* Tracing. *) (* ------------------------------------------------------------------------- *) val tracePrint : (string -> unit) Unsynchronized.ref val trace : string -> unit (* ------------------------------------------------------------------------- *) (* Combinators. *) (* ------------------------------------------------------------------------- *) val C : ('a -> 'b -> 'c) -> 'b -> 'a -> 'c val I : 'a -> 'a val K : 'a -> 'b -> 'a val S : ('a -> 'b -> 'c) -> ('a -> 'b) -> 'a -> 'c val W : ('a -> 'a -> 'b) -> 'a -> 'b val funpow : int -> ('a -> 'a) -> 'a -> 'a val exp : ('a * 'a -> 'a) -> 'a -> int -> 'a -> 'a (* ------------------------------------------------------------------------- *) (* Pairs. *) (* ------------------------------------------------------------------------- *) val fst : 'a * 'b -> 'a val snd : 'a * 'b -> 'b val pair : 'a -> 'b -> 'a * 'b val swap : 'a * 'b -> 'b * 'a val curry : ('a * 'b -> 'c) -> 'a -> 'b -> 'c val uncurry : ('a -> 'b -> 'c) -> 'a * 'b -> 'c val ## : ('a -> 'c) * ('b -> 'd) -> 'a * 'b -> 'c * 'd (* ------------------------------------------------------------------------- *) (* State transformers. *) (* ------------------------------------------------------------------------- *) val unit : 'a -> 's -> 'a * 's val bind : ('s -> 'a * 's) -> ('a -> 's -> 'b * 's) -> 's -> 'b * 's val mmap : ('a -> 'b) -> ('s -> 'a * 's) -> 's -> 'b * 's val mjoin : ('s -> ('s -> 'a * 's) * 's) -> 's -> 'a * 's 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. *) (* ------------------------------------------------------------------------- *) val cons : 'a -> 'a list -> 'a list val hdTl : 'a list -> 'a * 'a list val append : 'a list -> 'a list -> 'a list val singleton : 'a -> 'a list val first : ('a -> 'b option) -> 'a list -> 'b 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 zipWith : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list val zip : 'a list -> 'b list -> ('a * 'b) list val unzip : ('a * 'b) list -> 'a list * 'b list val cartwith : ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list 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 *) val updateNth : int * 'a -> 'a list -> 'a list (* Subscript *) val deleteNth : int -> 'a list -> 'a list (* Subscript *) (* ------------------------------------------------------------------------- *) (* Sets implemented with lists. *) (* ------------------------------------------------------------------------- *) val mem : ''a -> ''a list -> bool val insert : ''a -> ''a list -> ''a list val delete : ''a -> ''a list -> ''a list val setify : ''a list -> ''a list (* removes duplicates *) val union : ''a list -> ''a list -> ''a list val intersect : ''a list -> ''a list -> ''a list val difference : ''a list -> ''a list -> ''a list val subset : ''a list -> ''a list -> bool val distinct : ''a list -> bool (* ------------------------------------------------------------------------- *) (* Sorting and searching. *) (* ------------------------------------------------------------------------- *) val minimum : ('a * 'a -> order) -> 'a list -> 'a * 'a list (* Empty *) val maximum : ('a * 'a -> order) -> 'a list -> 'a * 'a list (* Empty *) val merge : ('a * 'a -> order) -> 'a list -> 'a list -> 'a list val sort : ('a * 'a -> order) -> 'a list -> 'a list val sortMap : ('a -> 'b) -> ('b * 'b -> order) -> 'a list -> 'a list (* ------------------------------------------------------------------------- *) (* Integers. *) (* ------------------------------------------------------------------------- *) val interval : int -> int -> int list val divides : int -> int -> bool val gcd : int -> int -> int (* Primes *) type sieve val initSieve : sieve val maxSieve : sieve -> int val primesSieve : sieve -> int list val incSieve : sieve -> bool * sieve val nextSieve : sieve -> int * sieve val primes : int -> int list val primesUpTo : int -> int list (* ------------------------------------------------------------------------- *) (* Strings. *) (* ------------------------------------------------------------------------- *) val rot : int -> char -> char val charToInt : char -> int option val charFromInt : int -> char option val nChars : char -> int -> string val chomp : string -> string val trim : string -> string val join : string -> string list -> string 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. *) (* ------------------------------------------------------------------------- *) type columnAlignment = {leftAlign : bool, padChar : char} val alignColumn : columnAlignment -> string list -> string list -> string list val alignTable : columnAlignment list -> string list list -> string list (* ------------------------------------------------------------------------- *) (* Reals. *) (* ------------------------------------------------------------------------- *) val percentToString : real -> string val pos : real -> real val log2 : real -> real (* Domain *) (* ------------------------------------------------------------------------- *) (* Sum datatype. *) (* ------------------------------------------------------------------------- *) datatype ('a,'b) sum = Left of 'a | Right of 'b val destLeft : ('a,'b) sum -> 'a val isLeft : ('a,'b) sum -> bool val destRight : ('a,'b) sum -> 'b val isRight : ('a,'b) sum -> bool (* ------------------------------------------------------------------------- *) (* Metis_Useful impure features. *) (* ------------------------------------------------------------------------- *) val newInt : unit -> int val newInts : int -> int list val withRef : 'r Unsynchronized.ref * 'r -> ('a -> 'b) -> 'a -> 'b val cloneArray : 'a Array.array -> 'a Array.array (* ------------------------------------------------------------------------- *) (* The environment. *) (* ------------------------------------------------------------------------- *) val host : unit -> string val time : unit -> string val date : unit -> string val readDirectory : {directory : string} -> {filename : string} list val readTextFile : {filename : string} -> string val writeTextFile : {contents : string, filename : string} -> unit (* ------------------------------------------------------------------------- *) (* Profiling and error reporting. *) (* ------------------------------------------------------------------------- *) val try : ('a -> 'b) -> 'a -> 'b val chat : string -> unit (* stdout *) val chide : string -> unit (* stderr *) val warn : string -> unit val die : string -> 'exit val timed : ('a -> 'b) -> 'a -> real * 'b val timedMany : ('a -> 'b) -> 'a -> real * 'b val executionTime : unit -> real (* Wall clock execution time *) end (**** Original file: src/Useful.sml ****) (* ========================================================================= *) (* ML UTILITY FUNCTIONS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Useful :> Metis_Useful = struct (* ------------------------------------------------------------------------- *) (* Exceptions. *) (* ------------------------------------------------------------------------- *) exception Error of string; exception Bug of string; 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; (* ------------------------------------------------------------------------- *) (* Tracing. *) (* ------------------------------------------------------------------------- *) local val traceOut = TextIO.stdOut; fun tracePrintFn mesg = let val () = TextIO.output (traceOut,mesg) val () = TextIO.flushOut traceOut in () end; in val tracePrint = Unsynchronized.ref tracePrintFn; end; fun trace mesg = !tracePrint mesg; (* ------------------------------------------------------------------------- *) (* Combinators. *) (* ------------------------------------------------------------------------- *) fun C f x y = f y x; fun I x = x; fun K x y = x; fun S f g x = f x (g x); fun W f x = f x x; fun funpow 0 _ x = x | funpow n f x = funpow (n - 1) f (f x); fun exp m = let fun f _ 0 z = z | f x y z = f (m (x,x)) (y div 2) (if y mod 2 = 0 then z else m (z,x)) in f end; (* ------------------------------------------------------------------------- *) (* Pairs. *) (* ------------------------------------------------------------------------- *) fun fst (x,_) = x; fun snd (_,y) = y; fun pair x y = (x,y); fun swap (x,y) = (y,x); fun curry f x y = f (x,y); fun uncurry f (x,y) = f x y; val op## = fn (f,g) => fn (x,y) => (f x, g y); (* ------------------------------------------------------------------------- *) (* State transformers. *) (* ------------------------------------------------------------------------- *) val unit : 'a -> 's -> 'a * 's = pair; fun bind f (g : 'a -> 's -> 'b * 's) = uncurry g o f; fun mmap f (m : 's -> 'a * 's) = bind m (unit o f); fun mjoin (f : 's -> ('s -> 'a * 's) * 's) = bind f I; fun mwhile c b = let fun f a = if c a then bind (b a) f else unit a in f end; (* ------------------------------------------------------------------------- *) (* 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. *) (* ------------------------------------------------------------------------- *) fun mapCompare f cmp (a,b) = cmp (f a, f b); fun revCompare cmp x_y = case cmp x_y of LESS => GREATER | EQUAL => EQUAL | GREATER => LESS; fun prodCompare xCmp yCmp ((x1,y1),(x2,y2)) = case xCmp (x1,x2) of LESS => LESS | EQUAL => yCmp (y1,y2) | GREATER => GREATER; fun lexCompare cmp = let fun lex ([],[]) = EQUAL | lex ([], _ :: _) = LESS | lex (_ :: _, []) = GREATER | lex (x :: xs, y :: ys) = case cmp (x,y) of LESS => LESS | EQUAL => lex (xs,ys) | GREATER => GREATER in lex end; fun optionCompare _ (NONE,NONE) = EQUAL | optionCompare _ (NONE,_) = LESS | optionCompare _ (_,NONE) = GREATER | optionCompare cmp (SOME x, SOME y) = cmp (x,y); 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 => List.rev (z [] xs ys) end; fun zip xs ys = zipWith pair xs ys; local fun inc ((x,y),(xs,ys)) = (x :: xs, y :: ys); in fun unzip ab = List.foldl inc ([],[]) (List.rev ab); end; 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' = List.rev xs in aux xs' [] xs' (List.rev ys) end end; fun cart xs ys = cartwith pair xs ys; fun takeWhile p = let fun f acc [] = List.rev acc | f acc (x :: xs) = if p x then f (x :: acc) xs else List.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 [] = (List.rev acc, []) | f acc (l as x :: xs) = if p x then f (x :: acc) xs else (List.rev acc, l) in f [] end; fun groups f = let fun group acc row x l = case l of [] => let val acc = if List.null row then acc else List.rev row :: acc in List.rev acc end | h :: t => let val (eor,x) = f (h,x) in if eor then group (List.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), List.map snd l); in fun groupsByFst l = List.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 (List.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; local fun inc (v,x) = if mem v x then x else v :: x; in fun setify s = List.rev (List.foldl inc [] s); end; fun union s t = let fun inc (v,x) = if mem v t then x else v :: x in List.foldl inc t (List.rev s) end; fun intersect s t = let fun inc (v,x) = if mem v t then v :: x else x in List.foldl inc [] (List.rev s) end; fun difference s t = let fun inc (v,x) = if mem v t then x else v :: x in List.foldl inc [] (List.rev s) end; 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. *) (* ------------------------------------------------------------------------- *) (* Finding the minimum and maximum element of a list, wrt some order. *) fun minimum cmp = let fun min (l,m,r) _ [] = (m, List.revAppend (l,r)) | min (best as (_,m,_)) l (x :: r) = min (case cmp (x,m) of LESS => (l,x,r) | _ => best) (x :: l) r in fn [] => raise Empty | h :: t => min ([],h,t) [h] t end; fun maximum cmp = minimum (revCompare cmp); (* Merge (for the following merge-sort, but generally useful too). *) fun merge cmp = let fun mrg acc [] ys = List.revAppend (acc,ys) | mrg acc xs [] = List.revAppend (acc,xs) | mrg acc (xs as x :: xt) (ys as y :: yt) = (case cmp (x,y) of GREATER => mrg (y :: acc) xs yt | _ => mrg (x :: acc) xt ys) in mrg [] end; (* Merge sort (stable). *) fun sort cmp = let fun findRuns acc r rs [] = List.rev (List.rev (r :: rs) :: acc) | findRuns acc r rs (x :: xs) = case cmp (r,x) of GREATER => findRuns (List.rev (r :: rs) :: acc) x [] xs | _ => findRuns acc x (r :: rs) xs fun mergeAdj acc [] = List.rev acc | mergeAdj acc (xs as [_]) = List.revAppend (acc,xs) | mergeAdj acc (x :: y :: xs) = mergeAdj (merge cmp x y :: acc) xs fun mergePairs [xs] = xs | mergePairs l = mergePairs (mergeAdj [] l) in fn [] => [] | l as [_] => l | h :: t => mergePairs (findRuns [] h [] t) end; fun sortMap _ _ [] = [] | sortMap _ _ (l as [_]) = l | sortMap f cmp xs = let fun ncmp ((m,_),(n,_)) = cmp (m,n) val nxs = List.map (fn x => (f x, x)) xs val nys = sort ncmp nxs in List.map snd nys end; (* ------------------------------------------------------------------------- *) (* Integers. *) (* ------------------------------------------------------------------------- *) fun interval m 0 = [] | interval m len = m :: interval (m + 1) (len - 1); fun divides _ 0 = true | divides 0 _ = false | divides a b = b mod (Int.abs a) = 0; local fun hcf 0 n = n | hcf 1 _ = 1 | hcf m n = hcf (n mod m) m; in fun gcd m n = let val m = Int.abs m and n = Int.abs n in if m < n then hcf m n else hcf n m end; end; (* Primes *) datatype sieve = Sieve of {max : int, primes : (int * (int * int)) list}; val initSieve = let val n = 1 and ps = [] in Sieve {max = n, primes = ps} end; fun maxSieve (Sieve {max = n, ...}) = n; fun primesSieve (Sieve {primes = ps, ...}) = List.map fst ps; fun incSieve sieve = let val n = maxSieve sieve + 1 fun add i ps = case ps of [] => (true,[(n,(0,0))]) | (p,(k,j)) :: ps => let val k = (k + i) mod p val j = j + i in if k = 0 then (false, (p,(k,j)) :: ps) else let val (b,ps) = add j ps in (b, (p,(k,0)) :: ps) end end val Sieve {primes = ps, ...} = sieve val (b,ps) = add 1 ps val sieve = Sieve {max = n, primes = ps} in (b,sieve) end; fun nextSieve sieve = let val (b,sieve) = incSieve sieve in if b then (maxSieve sieve, sieve) else nextSieve sieve end; local fun inc s = let val (_,s) = incSieve s in s end; in fun primesUpTo m = if m <= 1 then [] else primesSieve (funpow (m - 1) inc initSieve); end; val primes = let fun next s n = if n <= 0 then [] else let val (p,s) = nextSieve s val n = n - 1 val ps = next s n in p :: ps end in next initSieve end; (* ------------------------------------------------------------------------- *) (* Strings. *) (* ------------------------------------------------------------------------- *) local fun len l = (length l, l) val upper = len (String.explode "ABCDEFGHIJKLMNOPQRSTUVWXYZ"); val lower = len (String.explode "abcdefghijklmnopqrstuvwxyz"); fun rotate (n,l) c k = List.nth (l, (k + Option.valOf (index (equal c) l)) mod n); in fun rot k c = if Char.isLower c then rotate lower c k else if Char.isUpper c then rotate upper c k else c; end; fun charToInt #"0" = SOME 0 | charToInt #"1" = SOME 1 | charToInt #"2" = SOME 2 | charToInt #"3" = SOME 3 | charToInt #"4" = SOME 4 | charToInt #"5" = SOME 5 | charToInt #"6" = SOME 6 | charToInt #"7" = SOME 7 | charToInt #"8" = SOME 8 | charToInt #"9" = SOME 9 | charToInt _ = NONE; fun charFromInt 0 = SOME #"0" | charFromInt 1 = SOME #"1" | charFromInt 2 = SOME #"2" | charFromInt 3 = SOME #"3" | charFromInt 4 = SOME #"4" | charFromInt 5 = SOME #"5" | charFromInt 6 = SOME #"6" | charFromInt 7 = SOME #"7" | charFromInt 8 = SOME #"8" | charFromInt 9 = SOME #"9" | charFromInt _ = NONE; fun nChars x = let fun dup 0 l = l | dup n l = dup (n - 1) (x :: l) in fn n => String.implode (dup n []) end; fun chomp s = let val n = size s in if n = 0 orelse String.sub (s, n - 1) <> #"\n" then s else String.substring (s, 0, n - 1) end; local fun chop l = case l of [] => [] | h :: t => if Char.isSpace h then chop t else l; in val trim = String.implode o chop o List.rev o chop o List.rev o String.explode; end; val join = String.concatWith; local fun match [] l = SOME l | match _ [] = NONE | match (x :: xs) (y :: ys) = if x = y then match xs ys else NONE; fun stringify acc [] = acc | stringify acc (h :: t) = stringify (String.implode h :: acc) t; in fun split sep = let val pat = String.explode sep fun div1 prev recent [] = stringify [] (List.rev recent :: prev) | div1 prev recent (l as h :: t) = case match pat l of NONE => div1 prev (h :: recent) t | SOME rest => div1 (List.rev recent :: prev) [] rest in fn s => div1 [] [] (String.explode s) end; end; 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 = if String.isPrefix p s then () else raise Error "destPrefix" val sizeP = size p in 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. *) (* ------------------------------------------------------------------------- *) type columnAlignment = {leftAlign : bool, padChar : char} fun alignColumn {leftAlign,padChar} column = let val (n,_) = maximum Int.compare (List.map size column) fun pad entry row = let val padding = nChars padChar (n - size entry) in if leftAlign then entry ^ padding ^ row else padding ^ entry ^ row end in zipWith pad column end; local fun alignTab aligns rows = case aligns of [] => List.map (K "") rows | [{leftAlign = true, padChar = #" "}] => List.map hd rows | align :: aligns => let val col = List.map hd rows and cols = alignTab aligns (List.map tl rows) in alignColumn align col cols end; in fun alignTable aligns rows = if List.null rows then [] else alignTab aligns rows; end; (* ------------------------------------------------------------------------- *) (* Reals. *) (* ------------------------------------------------------------------------- *) val realToString = Real.toString; fun percentToString x = Int.toString (Real.round (100.0 * x)) ^ "%"; fun pos r = Real.max (r,0.0); local val invLn2 = 1.0 / Math.ln 2.0; in fun log2 x = invLn2 * Math.ln x; end; (* ------------------------------------------------------------------------- *) (* Sums. *) (* ------------------------------------------------------------------------- *) datatype ('a,'b) sum = Left of 'a | Right of 'b fun destLeft (Left l) = l | destLeft _ = raise Error "destLeft"; fun isLeft (Left _) = true | isLeft (Right _) = false; fun destRight (Right r) = r | destRight _ = raise Error "destRight"; fun isRight (Left _) = false | isRight (Right _) = true; (* ------------------------------------------------------------------------- *) (* Metis_Useful impure features. *) (* ------------------------------------------------------------------------- *) local val generator = Unsynchronized.ref 0 fun newIntThunk () = let val n = !generator val () = generator := n + 1 in n end; fun newIntsThunk k () = let val n = !generator val () = generator := n + k in interval n k end; in fun newInt () = Metis_Portable.critical newIntThunk (); fun newInts k = if k <= 0 then [] else Metis_Portable.critical (newIntsThunk k) (); end; fun withRef (r,new) f x = let val old = !r val () = r := new val y = f x handle e => (r := old; raise e) val () = r := old in y end; fun cloneArray a = let fun index i = Array.sub (a,i) in Array.tabulate (Array.length a, index) end; (* ------------------------------------------------------------------------- *) (* Environment. *) (* ------------------------------------------------------------------------- *) fun host () = Option.getOpt (OS.Process.getEnv "HOSTNAME", "unknown"); fun time () = Date.fmt "%H:%M:%S" (Date.fromTimeLocal (Time.now ())); 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 List.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 {contents,filename} = let open TextIO val h = openOut filename val () = output (h,contents) val () = closeOut h in () end; (* ------------------------------------------------------------------------- *) (* Profiling and error reporting. *) (* ------------------------------------------------------------------------- *) fun chat s = TextIO.output (TextIO.stdOut, s ^ "\n"); fun chide s = TextIO.output (TextIO.stdErr, s ^ "\n"); local fun err x s = chide (x ^ ": " ^ s); in fun try f x = f x handle e as Error _ => (err "try" (errorToString e); raise e) | e as Bug _ => (err "try" (bugToString e); raise e) | e => (err "try" "strange exception raised"; raise e); val warn = err "WARNING"; fun die s = (err "\nFATAL ERROR" s; OS.Process.exit OS.Process.failure); end; fun timed f a = let val tmr = Timer.startCPUTimer () val res = f a val {usr,sys,...} = Timer.checkCPUTimer tmr in (Time.toReal usr + Time.toReal sys, res) end; local val MIN = 1.0; fun several n t f a = let val (t',res) = timed f a val t = t + t' val n = n + 1 in if t > MIN then (t / Real.fromInt n, res) else several n t f a end; in fun timedMany f a = several 0 0.0 f a end; val executionTime = let val startTime = Time.toReal (Time.now ()) in fn () => Time.toReal (Time.now ()) - startTime end; end (**** Original file: src/Lazy.sig ****) (* ========================================================================= *) (* SUPPORT FOR LAZY EVALUATION *) (* Copyright (c) 2007 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Lazy = sig type 'a lazy val quickly : 'a -> 'a lazy val delay : (unit -> 'a) -> 'a lazy val force : 'a lazy -> 'a val memoize : (unit -> 'a) -> unit -> 'a end (**** Original file: src/Lazy.sml ****) (* ========================================================================= *) (* SUPPORT FOR LAZY EVALUATION *) (* Copyright (c) 2007 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Lazy :> Metis_Lazy = struct datatype 'a thunk = Value of 'a | Thunk of unit -> 'a; datatype 'a lazy = Metis_Lazy of 'a thunk Unsynchronized.ref; fun quickly v = Metis_Lazy (Unsynchronized.ref (Value v)); fun delay f = Metis_Lazy (Unsynchronized.ref (Thunk f)); fun force (Metis_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 val t = delay f in fn () => force t end; end (**** Original file: src/Ordered.sig ****) (* ========================================================================= *) (* ORDERED TYPES *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Ordered = sig type t val compare : t * t -> order (* PROVIDES !x : t. compare (x,x) = EQUAL !x y : t. compare (x,y) = LESS <=> compare (y,x) = GREATER !x y : t. compare (x,y) = EQUAL ==> compare (y,x) = EQUAL !x y z : t. compare (x,y) = EQUAL ==> compare (x,z) = compare (y,z) !x y z : t. compare (x,y) = LESS andalso compare (y,z) = LESS ==> compare (x,z) = LESS !x y z : t. compare (x,y) = GREATER andalso compare (y,z) = GREATER ==> compare (x,z) = GREATER *) end (**** Original file: src/Ordered.sml ****) (* ========================================================================= *) (* ORDERED TYPES *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_IntOrdered = struct type t = int val compare = Int.compare end; structure Metis_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 Metis_StringOrdered = struct type t = string val compare = String.compare end; (**** Original file: src/Map.sig ****) (* ========================================================================= *) (* FINITE MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_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 (* ------------------------------------------------------------------------- *) (* Metis_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 (* ------------------------------------------------------------------------- *) (* Metis_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: src/Map.sml ****) (* ========================================================================= *) (* FINITE MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Map :> Metis_Map = struct (* ------------------------------------------------------------------------- *) (* Importing useful functionality. *) (* ------------------------------------------------------------------------- *) exception Bug = Metis_Useful.Bug; exception Error = Metis_Useful.Error; val pointerEqual = Metis_Portable.pointerEqual; val K = Metis_Useful.K; val randomInt = Metis_Portable.randomInt; val randomWord = Metis_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 "Metis_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 "Metis_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 "Metis_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 "Metis_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 "Metis_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 = LeftToRightIterator of ('key * 'value) * ('key,'value) tree * ('key,'value) node list | RightToLeftIterator of ('key * 'value) * ('key,'value) tree * ('key,'value) node list; fun fromSpineLeftToRightIterator nodes = case nodes of [] => NONE | Node {key,value,right,...} :: nodes => SOME (LeftToRightIterator ((key,value),right,nodes)); fun fromSpineRightToLeftIterator nodes = case nodes of [] => NONE | Node {key,value,left,...} :: nodes => SOME (RightToLeftIterator ((key,value),left,nodes)); fun addLeftToRightIterator nodes tree = fromSpineLeftToRightIterator (treeLeftSpine nodes tree); fun addRightToLeftIterator nodes tree = fromSpineRightToLeftIterator (treeRightSpine nodes tree); fun treeMkIterator tree = addLeftToRightIterator [] tree; fun treeMkRevIterator tree = addRightToLeftIterator [] tree; fun readIterator iter = case iter of LeftToRightIterator (key_value,_,_) => key_value | RightToLeftIterator (key_value,_,_) => key_value; fun advanceIterator iter = case iter of LeftToRightIterator (_,tree,nodes) => addLeftToRightIterator nodes tree | RightToLeftIterator (_,tree,nodes) => addRightToLeftIterator 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 = Metis_Map of ('key * 'key -> order) * ('key,'value) tree; (* ------------------------------------------------------------------------- *) (* Metis_Map debugging functions. *) (* ------------------------------------------------------------------------- *) (*BasicDebug fun checkInvariants s m = let val Metis_Map (compareKey,tree) = m val _ = treeCheckInvariants compareKey tree in m end handle Bug bug => raise Bug (s ^ "\n" ^ "Metis_Map.checkInvariants: " ^ bug); *) (* ------------------------------------------------------------------------- *) (* Constructors. *) (* ------------------------------------------------------------------------- *) fun new compareKey = let val tree = treeNew () in Metis_Map (compareKey,tree) end; fun singleton compareKey key_value = let val tree = treeSingleton key_value in Metis_Map (compareKey,tree) end; (* ------------------------------------------------------------------------- *) (* Metis_Map size. *) (* ------------------------------------------------------------------------- *) fun size (Metis_Map (_,tree)) = treeSize tree; fun null m = size m = 0; (* ------------------------------------------------------------------------- *) (* Querying. *) (* ------------------------------------------------------------------------- *) fun peekKey (Metis_Map (compareKey,tree)) key = treePeekKey compareKey key tree; fun peek (Metis_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 "Metis_Map.get: element not found" | SOME value => value; fun pick (Metis_Map (_,tree)) = treePick tree; fun nth (Metis_Map (_,tree)) n = treeNth n tree; fun random m = let val n = size m in if n = 0 then raise Bug "Metis_Map.random: empty" else nth m (randomInt n) end; (* ------------------------------------------------------------------------- *) (* Adding. *) (* ------------------------------------------------------------------------- *) fun insert (Metis_Map (compareKey,tree)) key_value = let val tree = treeInsert compareKey key_value tree in Metis_Map (compareKey,tree) end; (*BasicDebug val insert = fn m => fn kv => checkInvariants "Metis_Map.insert: result" (insert (checkInvariants "Metis_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 (Metis_Map (compareKey,tree)) dkey = let val tree = treeDelete compareKey dkey tree in Metis_Map (compareKey,tree) end; (*BasicDebug val delete = fn m => fn k => checkInvariants "Metis_Map.delete: result" (delete (checkInvariants "Metis_Map.delete: input" m) k); *) fun remove m key = if inDomain key m then delete m key else m; fun deletePick (Metis_Map (compareKey,tree)) = let val (key_value,tree) = treeDeletePick tree in (key_value, Metis_Map (compareKey,tree)) end; (*BasicDebug val deletePick = fn m => let val (kv,m) = deletePick (checkInvariants "Metis_Map.deletePick: input" m) in (kv, checkInvariants "Metis_Map.deletePick: result" m) end; *) fun deleteNth (Metis_Map (compareKey,tree)) n = let val (key_value,tree) = treeDeleteNth n tree in (key_value, Metis_Map (compareKey,tree)) end; (*BasicDebug val deleteNth = fn m => fn n => let val (kv,m) = deleteNth (checkInvariants "Metis_Map.deleteNth: input" m) n in (kv, checkInvariants "Metis_Map.deleteNth: result" m) end; *) fun deleteRandom m = let val n = size m in if n = 0 then raise Bug "Metis_Map.deleteRandom: empty" else deleteNth m (randomInt n) end; (* ------------------------------------------------------------------------- *) (* Joining (all join operations prefer keys in the second map). *) (* ------------------------------------------------------------------------- *) fun merge {first,second,both} (Metis_Map (compareKey,tree1)) (Metis_Map (_,tree2)) = let val tree = treeMerge compareKey first second both tree1 tree2 in Metis_Map (compareKey,tree) end; (*BasicDebug val merge = fn f => fn m1 => fn m2 => checkInvariants "Metis_Map.merge: result" (merge f (checkInvariants "Metis_Map.merge: input 1" m1) (checkInvariants "Metis_Map.merge: input 2" m2)); *) fun union f (Metis_Map (compareKey,tree1)) (Metis_Map (_,tree2)) = let fun f2 kv = f (kv,kv) val tree = treeUnion compareKey f f2 tree1 tree2 in Metis_Map (compareKey,tree) end; (*BasicDebug val union = fn f => fn m1 => fn m2 => checkInvariants "Metis_Map.union: result" (union f (checkInvariants "Metis_Map.union: input 1" m1) (checkInvariants "Metis_Map.union: input 2" m2)); *) fun intersect f (Metis_Map (compareKey,tree1)) (Metis_Map (_,tree2)) = let val tree = treeIntersect compareKey f tree1 tree2 in Metis_Map (compareKey,tree) end; (*BasicDebug val intersect = fn f => fn m1 => fn m2 => checkInvariants "Metis_Map.intersect: result" (intersect f (checkInvariants "Metis_Map.intersect: input 1" m1) (checkInvariants "Metis_Map.intersect: input 2" m2)); *) (* ------------------------------------------------------------------------- *) (* Iterators over maps. *) (* ------------------------------------------------------------------------- *) fun mkIterator (Metis_Map (_,tree)) = treeMkIterator tree; fun mkRevIterator (Metis_Map (_,tree)) = treeMkRevIterator tree; (* ------------------------------------------------------------------------- *) (* Mapping and folding. *) (* ------------------------------------------------------------------------- *) fun mapPartial f (Metis_Map (compareKey,tree)) = let val tree = treeMapPartial f tree in Metis_Map (compareKey,tree) end; (*BasicDebug val mapPartial = fn f => fn m => checkInvariants "Metis_Map.mapPartial: result" (mapPartial f (checkInvariants "Metis_Map.mapPartial: input" m)); *) fun map f (Metis_Map (compareKey,tree)) = let val tree = treeMap f tree in Metis_Map (compareKey,tree) end; (*BasicDebug val map = fn f => fn m => checkInvariants "Metis_Map.map: result" (map f (checkInvariants "Metis_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 Metis_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 Metis_Map (compareKey,_) = m1 val io1 = mkIterator m1 and io2 = mkIterator m2 in equalIterator (equalKey compareKey) equalValue io1 io2 end); (* ------------------------------------------------------------------------- *) (* Metis_Set operations on the domain. *) (* ------------------------------------------------------------------------- *) fun unionDomain (Metis_Map (compareKey,tree1)) (Metis_Map (_,tree2)) = let val tree = treeUnionDomain compareKey tree1 tree2 in Metis_Map (compareKey,tree) end; (*BasicDebug val unionDomain = fn m1 => fn m2 => checkInvariants "Metis_Map.unionDomain: result" (unionDomain (checkInvariants "Metis_Map.unionDomain: input 1" m1) (checkInvariants "Metis_Map.unionDomain: input 2" m2)); *) local fun uncurriedUnionDomain (m,acc) = unionDomain acc m; in fun unionListDomain ms = case ms of [] => raise Bug "Metis_Map.unionListDomain: no sets" | m :: ms => List.foldl uncurriedUnionDomain m ms; end; fun intersectDomain (Metis_Map (compareKey,tree1)) (Metis_Map (_,tree2)) = let val tree = treeIntersectDomain compareKey tree1 tree2 in Metis_Map (compareKey,tree) end; (*BasicDebug val intersectDomain = fn m1 => fn m2 => checkInvariants "Metis_Map.intersectDomain: result" (intersectDomain (checkInvariants "Metis_Map.intersectDomain: input 1" m1) (checkInvariants "Metis_Map.intersectDomain: input 2" m2)); *) local fun uncurriedIntersectDomain (m,acc) = intersectDomain acc m; in fun intersectListDomain ms = case ms of [] => raise Bug "Metis_Map.intersectListDomain: no sets" | m :: ms => List.foldl uncurriedIntersectDomain m ms; end; fun differenceDomain (Metis_Map (compareKey,tree1)) (Metis_Map (_,tree2)) = let val tree = treeDifferenceDomain compareKey tree1 tree2 in Metis_Map (compareKey,tree) end; (*BasicDebug val differenceDomain = fn m1 => fn m2 => checkInvariants "Metis_Map.differenceDomain: result" (differenceDomain (checkInvariants "Metis_Map.differenceDomain: input 1" m1) (checkInvariants "Metis_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 (Metis_Map (compareKey,tree1)) (Metis_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 (**** Original file: src/KeyMap.sig ****) (* ========================================================================= *) (* FINITE MAPS WITH A FIXED KEY TYPE *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_KeyMap = sig (* ------------------------------------------------------------------------- *) (* A type of map keys. *) (* ------------------------------------------------------------------------- *) type key val compareKey : key * key -> order val equalKey : key -> key -> bool (* ------------------------------------------------------------------------- *) (* A type of finite maps. *) (* ------------------------------------------------------------------------- *) type 'a map (* ------------------------------------------------------------------------- *) (* Constructors. *) (* ------------------------------------------------------------------------- *) val new : unit -> 'a map val singleton : key * 'a -> 'a map (* ------------------------------------------------------------------------- *) (* Metis_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 (* ------------------------------------------------------------------------- *) (* Metis_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: src/KeyMap.sml ****) (* ========================================================================= *) (* FINITE MAPS WITH A FIXED KEY TYPE *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) functor Metis_KeyMap (Key : Metis_Ordered) :> Metis_KeyMap where type key = Key.t = struct (* ------------------------------------------------------------------------- *) (* Importing from the input signature. *) (* ------------------------------------------------------------------------- *) type key = Key.t; val compareKey = Key.compare; (* ------------------------------------------------------------------------- *) (* Importing useful functionality. *) (* ------------------------------------------------------------------------- *) exception Bug = Metis_Useful.Bug; exception Error = Metis_Useful.Error; val pointerEqual = Metis_Portable.pointerEqual; val K = Metis_Useful.K; val randomInt = Metis_Portable.randomInt; val randomWord = Metis_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 "Metis_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. *) (* ------------------------------------------------------------------------- *) 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 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 "Metis_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 "Metis_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 "Metis_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 "Metis_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 = LeftToRightIterator of (key * 'value) * 'value tree * 'value node list | RightToLeftIterator of (key * 'value) * 'value tree * 'value node list; fun fromSpineLeftToRightIterator nodes = case nodes of [] => NONE | Node {key,value,right,...} :: nodes => SOME (LeftToRightIterator ((key,value),right,nodes)); fun fromSpineRightToLeftIterator nodes = case nodes of [] => NONE | Node {key,value,left,...} :: nodes => SOME (RightToLeftIterator ((key,value),left,nodes)); fun addLeftToRightIterator nodes tree = fromSpineLeftToRightIterator (treeLeftSpine nodes tree); fun addRightToLeftIterator nodes tree = fromSpineRightToLeftIterator (treeRightSpine nodes tree); fun treeMkIterator tree = addLeftToRightIterator [] tree; fun treeMkRevIterator tree = addRightToLeftIterator [] tree; fun readIterator iter = case iter of LeftToRightIterator (key_value,_,_) => key_value | RightToLeftIterator (key_value,_,_) => key_value; fun advanceIterator iter = case iter of LeftToRightIterator (_,tree,nodes) => addLeftToRightIterator nodes tree | RightToLeftIterator (_,tree,nodes) => addRightToLeftIterator 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 = Metis_Map of 'value tree; (* ------------------------------------------------------------------------- *) (* Metis_Map debugging functions. *) (* ------------------------------------------------------------------------- *) (*BasicDebug fun checkInvariants s m = let val Metis_Map tree = m val _ = treeCheckInvariants tree in m end handle Bug bug => raise Bug (s ^ "\n" ^ "Metis_KeyMap.checkInvariants: " ^ bug); *) (* ------------------------------------------------------------------------- *) (* Constructors. *) (* ------------------------------------------------------------------------- *) fun new () = let val tree = treeNew () in Metis_Map tree end; fun singleton key_value = let val tree = treeSingleton key_value in Metis_Map tree end; (* ------------------------------------------------------------------------- *) (* Metis_Map size. *) (* ------------------------------------------------------------------------- *) fun size (Metis_Map tree) = treeSize tree; fun null m = size m = 0; (* ------------------------------------------------------------------------- *) (* Querying. *) (* ------------------------------------------------------------------------- *) fun peekKey (Metis_Map tree) key = treePeekKey key tree; fun peek (Metis_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 "Metis_KeyMap.get: element not found" | SOME value => value; fun pick (Metis_Map tree) = treePick tree; fun nth (Metis_Map tree) n = treeNth n tree; fun random m = let val n = size m in if n = 0 then raise Bug "Metis_KeyMap.random: empty" else nth m (randomInt n) end; (* ------------------------------------------------------------------------- *) (* Adding. *) (* ------------------------------------------------------------------------- *) fun insert (Metis_Map tree) key_value = let val tree = treeInsert key_value tree in Metis_Map tree end; (*BasicDebug val insert = fn m => fn kv => checkInvariants "Metis_KeyMap.insert: result" (insert (checkInvariants "Metis_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 (Metis_Map tree) dkey = let val tree = treeDelete dkey tree in Metis_Map tree end; (*BasicDebug val delete = fn m => fn k => checkInvariants "Metis_KeyMap.delete: result" (delete (checkInvariants "Metis_KeyMap.delete: input" m) k); *) fun remove m key = if inDomain key m then delete m key else m; fun deletePick (Metis_Map tree) = let val (key_value,tree) = treeDeletePick tree in (key_value, Metis_Map tree) end; (*BasicDebug val deletePick = fn m => let val (kv,m) = deletePick (checkInvariants "Metis_KeyMap.deletePick: input" m) in (kv, checkInvariants "Metis_KeyMap.deletePick: result" m) end; *) fun deleteNth (Metis_Map tree) n = let val (key_value,tree) = treeDeleteNth n tree in (key_value, Metis_Map tree) end; (*BasicDebug val deleteNth = fn m => fn n => let val (kv,m) = deleteNth (checkInvariants "Metis_KeyMap.deleteNth: input" m) n in (kv, checkInvariants "Metis_KeyMap.deleteNth: result" m) end; *) fun deleteRandom m = let val n = size m in if n = 0 then raise Bug "Metis_KeyMap.deleteRandom: empty" else deleteNth m (randomInt n) end; (* ------------------------------------------------------------------------- *) (* Joining (all join operations prefer keys in the second map). *) (* ------------------------------------------------------------------------- *) fun merge {first,second,both} (Metis_Map tree1) (Metis_Map tree2) = let val tree = treeMerge first second both tree1 tree2 in Metis_Map tree end; (*BasicDebug val merge = fn f => fn m1 => fn m2 => checkInvariants "Metis_KeyMap.merge: result" (merge f (checkInvariants "Metis_KeyMap.merge: input 1" m1) (checkInvariants "Metis_KeyMap.merge: input 2" m2)); *) fun union f (Metis_Map tree1) (Metis_Map tree2) = let fun f2 kv = f (kv,kv) val tree = treeUnion f f2 tree1 tree2 in Metis_Map tree end; (*BasicDebug val union = fn f => fn m1 => fn m2 => checkInvariants "Metis_KeyMap.union: result" (union f (checkInvariants "Metis_KeyMap.union: input 1" m1) (checkInvariants "Metis_KeyMap.union: input 2" m2)); *) fun intersect f (Metis_Map tree1) (Metis_Map tree2) = let val tree = treeIntersect f tree1 tree2 in Metis_Map tree end; (*BasicDebug val intersect = fn f => fn m1 => fn m2 => checkInvariants "Metis_KeyMap.intersect: result" (intersect f (checkInvariants "Metis_KeyMap.intersect: input 1" m1) (checkInvariants "Metis_KeyMap.intersect: input 2" m2)); *) (* ------------------------------------------------------------------------- *) (* Iterators over maps. *) (* ------------------------------------------------------------------------- *) fun mkIterator (Metis_Map tree) = treeMkIterator tree; fun mkRevIterator (Metis_Map tree) = treeMkRevIterator tree; (* ------------------------------------------------------------------------- *) (* Mapping and folding. *) (* ------------------------------------------------------------------------- *) fun mapPartial f (Metis_Map tree) = let val tree = treeMapPartial f tree in Metis_Map tree end; (*BasicDebug val mapPartial = fn f => fn m => checkInvariants "Metis_KeyMap.mapPartial: result" (mapPartial f (checkInvariants "Metis_KeyMap.mapPartial: input" m)); *) fun map f (Metis_Map tree) = let val tree = treeMap f tree in Metis_Map tree end; (*BasicDebug val map = fn f => fn m => checkInvariants "Metis_KeyMap.map: result" (map f (checkInvariants "Metis_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 Metis_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 Metis_Map _ = m1 val io1 = mkIterator m1 and io2 = mkIterator m2 in equalIterator equalValue io1 io2 end); (* ------------------------------------------------------------------------- *) (* Metis_Set operations on the domain. *) (* ------------------------------------------------------------------------- *) fun unionDomain (Metis_Map tree1) (Metis_Map tree2) = let val tree = treeUnionDomain tree1 tree2 in Metis_Map tree end; (*BasicDebug val unionDomain = fn m1 => fn m2 => checkInvariants "Metis_KeyMap.unionDomain: result" (unionDomain (checkInvariants "Metis_KeyMap.unionDomain: input 1" m1) (checkInvariants "Metis_KeyMap.unionDomain: input 2" m2)); *) local fun uncurriedUnionDomain (m,acc) = unionDomain acc m; in fun unionListDomain ms = case ms of [] => raise Bug "Metis_KeyMap.unionListDomain: no sets" | m :: ms => List.foldl uncurriedUnionDomain m ms; end; fun intersectDomain (Metis_Map tree1) (Metis_Map tree2) = let val tree = treeIntersectDomain tree1 tree2 in Metis_Map tree end; (*BasicDebug val intersectDomain = fn m1 => fn m2 => checkInvariants "Metis_KeyMap.intersectDomain: result" (intersectDomain (checkInvariants "Metis_KeyMap.intersectDomain: input 1" m1) (checkInvariants "Metis_KeyMap.intersectDomain: input 2" m2)); *) local fun uncurriedIntersectDomain (m,acc) = intersectDomain acc m; in fun intersectListDomain ms = case ms of [] => raise Bug "Metis_KeyMap.intersectListDomain: no sets" | m :: ms => List.foldl uncurriedIntersectDomain m ms; end; fun differenceDomain (Metis_Map tree1) (Metis_Map tree2) = let val tree = treeDifferenceDomain tree1 tree2 in Metis_Map tree end; (*BasicDebug val differenceDomain = fn m1 => fn m2 => checkInvariants "Metis_KeyMap.differenceDomain: result" (differenceDomain (checkInvariants "Metis_KeyMap.differenceDomain: input 1" m1) (checkInvariants "Metis_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 (Metis_Map tree1) (Metis_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 Metis_IntMap = Metis_KeyMap (Metis_IntOrdered); structure Metis_IntPairMap = Metis_KeyMap (Metis_IntPairOrdered); structure Metis_StringMap = Metis_KeyMap (Metis_StringOrdered); (**** Original file: src/Set.sig ****) (* ========================================================================= *) (* FINITE SETS IMPLEMENTED WITH RANDOMLY BALANCED TREES *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Set = sig (* ------------------------------------------------------------------------- *) (* A type of finite sets. *) (* ------------------------------------------------------------------------- *) type 'elt set (* ------------------------------------------------------------------------- *) (* Constructors. *) (* ------------------------------------------------------------------------- *) val empty : ('elt * 'elt -> order) -> 'elt set val singleton : ('elt * 'elt -> order) -> 'elt -> 'elt set (* ------------------------------------------------------------------------- *) (* Metis_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 (* ------------------------------------------------------------------------- *) (* 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 val unionList : 'elt set list -> 'elt set val intersect : 'elt set -> 'elt set -> 'elt set val intersectList : 'elt set list -> 'elt set val difference : 'elt set -> 'elt set -> 'elt set val symmetricDifference : 'elt set -> 'elt set -> 'elt set (* ------------------------------------------------------------------------- *) (* Mapping and folding. *) (* ------------------------------------------------------------------------- *) val filter : ('elt -> bool) -> 'elt set -> 'elt set val partition : ('elt -> bool) -> 'elt set -> 'elt set * 'elt set 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 val firstl : ('elt -> 'a option) -> 'elt set -> 'a option val firstr : ('elt -> 'a option) -> 'elt set -> 'a option val exists : ('elt -> bool) -> 'elt set -> bool val all : ('elt -> bool) -> 'elt set -> bool 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 toList : 'elt set -> 'elt list val fromList : ('elt * 'elt -> order) -> 'elt list -> '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 (* ------------------------------------------------------------------------- *) (* Iterators over sets *) (* ------------------------------------------------------------------------- *) type 'elt iterator val mkIterator : 'elt set -> 'elt iterator option val mkRevIterator : 'elt set -> 'elt iterator option val readIterator : 'elt iterator -> 'elt val advanceIterator : 'elt iterator -> 'elt iterator option end (**** Original file: src/Set.sml ****) (* ========================================================================= *) (* FINITE SETS IMPLEMENTED WITH RANDOMLY BALANCED TREES *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Set :> Metis_Set = struct (* ------------------------------------------------------------------------- *) (* A type of finite sets. *) (* ------------------------------------------------------------------------- *) type ('elt,'a) map = ('elt,'a) Metis_Map.map; datatype 'elt set = Metis_Set of ('elt,unit) map; (* ------------------------------------------------------------------------- *) (* Converting to and from maps. *) (* ------------------------------------------------------------------------- *) fun dest (Metis_Set m) = m; fun mapPartial f = let fun mf (elt,()) = f elt in fn Metis_Set m => Metis_Map.mapPartial mf m end; fun map f = let fun mf (elt,()) = f elt in fn Metis_Set m => Metis_Map.map mf m end; fun domain m = Metis_Set (Metis_Map.transform (fn _ => ()) m); (* ------------------------------------------------------------------------- *) (* Constructors. *) (* ------------------------------------------------------------------------- *) fun empty cmp = Metis_Set (Metis_Map.new cmp); fun singleton cmp elt = Metis_Set (Metis_Map.singleton cmp (elt,())); (* ------------------------------------------------------------------------- *) (* Metis_Set size. *) (* ------------------------------------------------------------------------- *) fun null (Metis_Set m) = Metis_Map.null m; fun size (Metis_Set m) = Metis_Map.size m; (* ------------------------------------------------------------------------- *) (* Querying. *) (* ------------------------------------------------------------------------- *) fun peek (Metis_Set m) elt = case Metis_Map.peekKey m elt of SOME (elt,()) => SOME elt | NONE => NONE; fun member elt (Metis_Set m) = Metis_Map.inDomain elt m; fun pick (Metis_Set m) = let val (elt,_) = Metis_Map.pick m in elt end; fun nth (Metis_Set m) n = let val (elt,_) = Metis_Map.nth m n in elt end; fun random (Metis_Set m) = let val (elt,_) = Metis_Map.random m in elt end; (* ------------------------------------------------------------------------- *) (* Adding. *) (* ------------------------------------------------------------------------- *) fun add (Metis_Set m) elt = let val m = Metis_Map.insert m (elt,()) in Metis_Set m end; local fun uncurriedAdd (elt,set) = add set elt; in fun addList set = List.foldl uncurriedAdd set; end; (* ------------------------------------------------------------------------- *) (* Removing. *) (* ------------------------------------------------------------------------- *) fun delete (Metis_Set m) elt = let val m = Metis_Map.delete m elt in Metis_Set m end; fun remove (Metis_Set m) elt = let val m = Metis_Map.remove m elt in Metis_Set m end; fun deletePick (Metis_Set m) = let val ((elt,()),m) = Metis_Map.deletePick m in (elt, Metis_Set m) end; fun deleteNth (Metis_Set m) n = let val ((elt,()),m) = Metis_Map.deleteNth m n in (elt, Metis_Set m) end; fun deleteRandom (Metis_Set m) = let val ((elt,()),m) = Metis_Map.deleteRandom m in (elt, Metis_Set m) end; (* ------------------------------------------------------------------------- *) (* Joining. *) (* ------------------------------------------------------------------------- *) fun union (Metis_Set m1) (Metis_Set m2) = Metis_Set (Metis_Map.unionDomain m1 m2); fun unionList sets = let val ms = List.map dest sets in Metis_Set (Metis_Map.unionListDomain ms) end; fun intersect (Metis_Set m1) (Metis_Set m2) = Metis_Set (Metis_Map.intersectDomain m1 m2); fun intersectList sets = let val ms = List.map dest sets in Metis_Set (Metis_Map.intersectListDomain ms) end; fun difference (Metis_Set m1) (Metis_Set m2) = Metis_Set (Metis_Map.differenceDomain m1 m2); fun symmetricDifference (Metis_Set m1) (Metis_Set m2) = Metis_Set (Metis_Map.symmetricDifferenceDomain m1 m2); (* ------------------------------------------------------------------------- *) (* Mapping and folding. *) (* ------------------------------------------------------------------------- *) fun filter pred = let fun mpred (elt,()) = pred elt in fn Metis_Set m => Metis_Set (Metis_Map.filter mpred m) end; fun partition pred = let fun mpred (elt,()) = pred elt in fn Metis_Set m => let val (m1,m2) = Metis_Map.partition mpred m in (Metis_Set m1, Metis_Set m2) end end; fun app f = let fun mf (elt,()) = f elt in fn Metis_Set m => Metis_Map.app mf m end; fun foldl f = let fun mf (elt,(),acc) = f (elt,acc) in fn acc => fn Metis_Set m => Metis_Map.foldl mf acc m end; fun foldr f = let fun mf (elt,(),acc) = f (elt,acc) in fn acc => fn Metis_Set m => Metis_Map.foldr mf acc m end; (* ------------------------------------------------------------------------- *) (* Searching. *) (* ------------------------------------------------------------------------- *) fun findl p = let fun mp (elt,()) = p elt in fn Metis_Set m => case Metis_Map.findl mp m of SOME (elt,()) => SOME elt | NONE => NONE end; fun findr p = let fun mp (elt,()) = p elt in fn Metis_Set m => case Metis_Map.findr mp m of SOME (elt,()) => SOME elt | NONE => NONE end; fun firstl f = let fun mf (elt,()) = f elt in fn Metis_Set m => Metis_Map.firstl mf m end; fun firstr f = let fun mf (elt,()) = f elt in fn Metis_Set m => Metis_Map.firstr mf m end; fun exists p = let fun mp (elt,()) = p elt in fn Metis_Set m => Metis_Map.exists mp m end; fun all p = let fun mp (elt,()) = p elt in fn Metis_Set m => Metis_Map.all mp m end; fun count p = let fun mp (elt,()) = p elt in fn Metis_Set m => Metis_Map.count mp m end; (* ------------------------------------------------------------------------- *) (* Comparing. *) (* ------------------------------------------------------------------------- *) fun compareValue ((),()) = EQUAL; fun equalValue () () = true; fun compare (Metis_Set m1, Metis_Set m2) = Metis_Map.compare compareValue (m1,m2); fun equal (Metis_Set m1) (Metis_Set m2) = Metis_Map.equal equalValue m1 m2; fun subset (Metis_Set m1) (Metis_Set m2) = Metis_Map.subsetDomain m1 m2; fun disjoint (Metis_Set m1) (Metis_Set m2) = Metis_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 (Metis_Set m) = Metis_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) Metis_Map.iterator; fun mkIterator (Metis_Set m) = Metis_Map.mkIterator m; fun mkRevIterator (Metis_Set m) = Metis_Map.mkRevIterator m; fun readIterator iter = let val (elt,()) = Metis_Map.readIterator iter in elt end; fun advanceIterator iter = Metis_Map.advanceIterator iter; end (**** Original file: src/ElementSet.sig ****) (* ========================================================================= *) (* FINITE SETS WITH A FIXED ELEMENT TYPE *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_ElementSet = sig (* ------------------------------------------------------------------------- *) (* A type of set elements. *) (* ------------------------------------------------------------------------- *) type element val compareElement : element * element -> order val equalElement : element -> element -> bool (* ------------------------------------------------------------------------- *) (* A type of finite sets. *) (* ------------------------------------------------------------------------- *) type set (* ------------------------------------------------------------------------- *) (* Constructors. *) (* ------------------------------------------------------------------------- *) val empty : set val singleton : element -> set (* ------------------------------------------------------------------------- *) (* Metis_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 (* ------------------------------------------------------------------------- *) (* 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 val unionList : set list -> set val intersect : set -> set -> set val intersectList : set list -> set val difference : set -> set -> set val symmetricDifference : set -> set -> set (* ------------------------------------------------------------------------- *) (* Mapping and folding. *) (* ------------------------------------------------------------------------- *) val filter : (element -> bool) -> set -> set val partition : (element -> bool) -> set -> set * set 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 val firstl : (element -> 'a option) -> set -> 'a option val firstr : (element -> 'a option) -> set -> 'a option val exists : (element -> bool) -> set -> bool val all : (element -> bool) -> set -> bool 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 (* ------------------------------------------------------------------------- *) (* Pointwise operations. *) (* ------------------------------------------------------------------------- *) val lift : (element -> set) -> set -> set val closedAdd : (element -> set) -> set -> set -> set val close : (element -> set) -> set -> set (* ------------------------------------------------------------------------- *) (* Converting to and from lists. *) (* ------------------------------------------------------------------------- *) val transform : (element -> 'a) -> set -> 'a list val toList : set -> element list val fromList : element list -> 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 (* ------------------------------------------------------------------------- *) (* Depth-first search. *) (* ------------------------------------------------------------------------- *) datatype ordering = Linear of element list | Cycle of element list val preOrder : (element -> set) -> set -> ordering val postOrder : (element -> set) -> set -> ordering val preOrdered : (element -> set) -> element list -> bool val postOrdered : (element -> set) -> element list -> bool (* ------------------------------------------------------------------------- *) (* Strongly connected components. *) (* ------------------------------------------------------------------------- *) val preOrderSCC : (element -> set) -> set -> set list val postOrderSCC : (element -> set) -> set -> set list val preOrderedSCC : (element -> set) -> set list -> bool val postOrderedSCC : (element -> set) -> set list -> bool (* ------------------------------------------------------------------------- *) (* Pretty-printing. *) (* ------------------------------------------------------------------------- *) val toString : set -> string (* ------------------------------------------------------------------------- *) (* Iterators over sets *) (* ------------------------------------------------------------------------- *) type iterator val mkIterator : set -> iterator option val mkRevIterator : set -> iterator option val readIterator : iterator -> element val advanceIterator : iterator -> iterator option end (**** Original file: src/ElementSet.sml ****) (* ========================================================================= *) (* FINITE SETS WITH A FIXED ELEMENT TYPE *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) functor Metis_ElementSet ( KM : Metis_KeyMap ) :> Metis_ElementSet where type element = KM.key and type 'a map = 'a KM.map = struct (* ------------------------------------------------------------------------- *) (* A type of set elements. *) (* ------------------------------------------------------------------------- *) type element = KM.key; val compareElement = KM.compareKey; val equalElement = KM.equalKey; (* ------------------------------------------------------------------------- *) (* A type of finite sets. *) (* ------------------------------------------------------------------------- *) type 'a map = 'a KM.map; datatype set = Metis_Set of unit map; (* ------------------------------------------------------------------------- *) (* Converting to and from maps. *) (* ------------------------------------------------------------------------- *) fun dest (Metis_Set m) = m; fun mapPartial f = let fun mf (elt,()) = f elt in fn Metis_Set m => KM.mapPartial mf m end; fun map f = let fun mf (elt,()) = f elt in fn Metis_Set m => KM.map mf m end; fun domain m = Metis_Set (KM.transform (fn _ => ()) m); (* ------------------------------------------------------------------------- *) (* Constructors. *) (* ------------------------------------------------------------------------- *) val empty = Metis_Set (KM.new ()); fun singleton elt = Metis_Set (KM.singleton (elt,())); (* ------------------------------------------------------------------------- *) (* Metis_Set size. *) (* ------------------------------------------------------------------------- *) fun null (Metis_Set m) = KM.null m; fun size (Metis_Set m) = KM.size m; (* ------------------------------------------------------------------------- *) (* Querying. *) (* ------------------------------------------------------------------------- *) fun peek (Metis_Set m) elt = case KM.peekKey m elt of SOME (elt,()) => SOME elt | NONE => NONE; fun member elt (Metis_Set m) = KM.inDomain elt m; fun pick (Metis_Set m) = let val (elt,_) = KM.pick m in elt end; fun nth (Metis_Set m) n = let val (elt,_) = KM.nth m n in elt end; fun random (Metis_Set m) = let val (elt,_) = KM.random m in elt end; (* ------------------------------------------------------------------------- *) (* Adding. *) (* ------------------------------------------------------------------------- *) fun add (Metis_Set m) elt = let val m = KM.insert m (elt,()) in Metis_Set m end; local fun uncurriedAdd (elt,set) = add set elt; in fun addList set = List.foldl uncurriedAdd set; end; (* ------------------------------------------------------------------------- *) (* Removing. *) (* ------------------------------------------------------------------------- *) fun delete (Metis_Set m) elt = let val m = KM.delete m elt in Metis_Set m end; fun remove (Metis_Set m) elt = let val m = KM.remove m elt in Metis_Set m end; fun deletePick (Metis_Set m) = let val ((elt,()),m) = KM.deletePick m in (elt, Metis_Set m) end; fun deleteNth (Metis_Set m) n = let val ((elt,()),m) = KM.deleteNth m n in (elt, Metis_Set m) end; fun deleteRandom (Metis_Set m) = let val ((elt,()),m) = KM.deleteRandom m in (elt, Metis_Set m) end; (* ------------------------------------------------------------------------- *) (* Joining. *) (* ------------------------------------------------------------------------- *) fun union (Metis_Set m1) (Metis_Set m2) = Metis_Set (KM.unionDomain m1 m2); fun unionList sets = let val ms = List.map dest sets in Metis_Set (KM.unionListDomain ms) end; fun intersect (Metis_Set m1) (Metis_Set m2) = Metis_Set (KM.intersectDomain m1 m2); fun intersectList sets = let val ms = List.map dest sets in Metis_Set (KM.intersectListDomain ms) end; fun difference (Metis_Set m1) (Metis_Set m2) = Metis_Set (KM.differenceDomain m1 m2); fun symmetricDifference (Metis_Set m1) (Metis_Set m2) = Metis_Set (KM.symmetricDifferenceDomain m1 m2); (* ------------------------------------------------------------------------- *) (* Mapping and folding. *) (* ------------------------------------------------------------------------- *) fun filter pred = let fun mpred (elt,()) = pred elt in fn Metis_Set m => Metis_Set (KM.filter mpred m) end; fun partition pred = let fun mpred (elt,()) = pred elt in fn Metis_Set m => let val (m1,m2) = KM.partition mpred m in (Metis_Set m1, Metis_Set m2) end end; fun app f = let fun mf (elt,()) = f elt in fn Metis_Set m => KM.app mf m end; fun foldl f = let fun mf (elt,(),acc) = f (elt,acc) in fn acc => fn Metis_Set m => KM.foldl mf acc m end; fun foldr f = let fun mf (elt,(),acc) = f (elt,acc) in fn acc => fn Metis_Set m => KM.foldr mf acc m end; (* ------------------------------------------------------------------------- *) (* Searching. *) (* ------------------------------------------------------------------------- *) fun findl p = let fun mp (elt,()) = p elt in fn Metis_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 Metis_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 Metis_Set m => KM.firstl mf m end; fun firstr f = let fun mf (elt,()) = f elt in fn Metis_Set m => KM.firstr mf m end; fun exists p = let fun mp (elt,()) = p elt in fn Metis_Set m => KM.exists mp m end; fun all p = let fun mp (elt,()) = p elt in fn Metis_Set m => KM.all mp m end; fun count p = let fun mp (elt,()) = p elt in fn Metis_Set m => KM.count mp m end; (* ------------------------------------------------------------------------- *) (* Comparing. *) (* ------------------------------------------------------------------------- *) fun compareValue ((),()) = EQUAL; fun equalValue () () = true; fun compare (Metis_Set m1, Metis_Set m2) = KM.compare compareValue (m1,m2); fun equal (Metis_Set m1) (Metis_Set m2) = KM.equal equalValue m1 m2; fun subset (Metis_Set m1) (Metis_Set m2) = KM.subsetDomain m1 m2; fun disjoint (Metis_Set m1) (Metis_Set m2) = KM.disjointDomain m1 m2; (* ------------------------------------------------------------------------- *) (* Pointwise operations. *) (* ------------------------------------------------------------------------- *) fun lift f = let fun inc (elt,set) = union set (f elt) in foldl inc empty end; fun closedAdd f = let fun adds acc set = foldl check acc set and check (elt,acc) = if member elt acc then acc else expand (add acc elt) elt and expand acc elt = adds acc (f elt) in adds end; fun close f = closedAdd f empty; (* ------------------------------------------------------------------------- *) (* Converting to and from lists. *) (* ------------------------------------------------------------------------- *) fun transform f = let fun inc (x,l) = f x :: l in foldr inc [] end; fun toList (Metis_Set m) = KM.keys m; fun fromList elts = addList empty elts; (* ------------------------------------------------------------------------- *) (* Depth-first search. *) (* ------------------------------------------------------------------------- *) datatype ordering = Linear of element list | Cycle of element list; fun postOrdered children = let fun check acc elts = case elts of [] => true | elt :: elts => not (member elt acc) andalso let val acc = closedAdd children acc (singleton elt) in check acc elts end in check empty end; fun preOrdered children elts = postOrdered children (List.rev elts); local fun takeStackset elt = let fun notElement (e,_,_) = not (equalElement e elt) in Metis_Useful.takeWhile notElement end; fun consElement ((e,_,_),el) = e :: el; fun depthFirstSearch children = let fun traverse (dealt,dealtset) (stack,stackset) work = case work of [] => (case stack of [] => Linear dealt | (elt,work,stackset) :: stack => let val dealt = elt :: dealt val dealtset = add dealtset elt in traverse (dealt,dealtset) (stack,stackset) work end) | elt :: work => if member elt dealtset then traverse (dealt,dealtset) (stack,stackset) work else if member elt stackset then let val cycle = takeStackset elt stack val cycle = elt :: List.foldl consElement [elt] cycle in Cycle cycle end else let val stack = (elt,work,stackset) :: stack val stackset = add stackset elt val work = toList (children elt) in traverse (dealt,dealtset) (stack,stackset) work end val dealt = [] and dealtset = empty and stack = [] and stackset = empty in traverse (dealt,dealtset) (stack,stackset) end; in fun preOrder children roots = let val result = depthFirstSearch children (toList roots) (*BasicDebug val () = case result of Cycle _ => () | Linear l => let val () = if subset roots (fromList l) then () else raise Metis_Useful.Bug "Metis_ElementSet.preOrder: missing roots" val () = if preOrdered children l then () else raise Metis_Useful.Bug "Metis_ElementSet.preOrder: bad ordering" in () end *) in result end; fun postOrder children roots = case depthFirstSearch children (toList roots) of Linear l => let val l = List.rev l (*BasicDebug val () = if subset roots (fromList l) then () else raise Metis_Useful.Bug "Metis_ElementSet.postOrder: missing roots" val () = if postOrdered children l then () else raise Metis_Useful.Bug "Metis_ElementSet.postOrder: bad ordering" *) in Linear l end | cycle => cycle; end; (* ------------------------------------------------------------------------- *) (* Strongly connected components. *) (* ------------------------------------------------------------------------- *) fun postOrderedSCC children = let fun check acc eltsl = case eltsl of [] => true | elts :: eltsl => not (null elts) andalso disjoint elts acc andalso let fun addElt elt = closedAdd children acc (singleton elt) val (root,elts) = deletePick elts fun checkElt elt = member root (addElt elt) in all checkElt elts andalso let val acc = addElt root in subset elts acc andalso check acc eltsl end end in check empty end; fun preOrderedSCC children eltsl = postOrderedSCC children (List.rev eltsl); (* An implementation of Tarjan's algorithm: *) (* http://en.wikipedia.org/wiki/Tarjan%27s_strongly_connected_components_algorithm *) local datatype stackSCC = StackSCC of set * (element * set) list; val emptyStack = StackSCC (empty,[]); fun pushStack (StackSCC (elts,eltl)) elt = StackSCC (add elts elt, (elt,elts) :: eltl); fun inStack elt (StackSCC (elts,_)) = member elt elts; fun popStack root = let fun pop scc eltl = case eltl of [] => raise Metis_Useful.Bug "Metis_ElementSet.popStack" | (elt,elts) :: eltl => let val scc = add scc elt in if equalElement elt root then (scc, StackSCC (elts,eltl)) else pop scc eltl end in fn sccs => fn StackSCC (_,eltl) => let val (scc,stack) = pop empty eltl in (scc :: sccs, stack) end end; fun getIndex indices e : int = case KM.peek indices e of SOME i => i | NONE => raise Metis_Useful.Bug "Metis_ElementSet.getIndex"; fun isRoot indices lows e = getIndex indices e = getIndex lows e; fun reduceIndex indices (e,i) = let val j = getIndex indices e in if j <= i then indices else KM.insert indices (e,i) end; fun tarjan children = let fun dfsVertex sccs callstack index indices lows stack elt = let val indices = KM.insert indices (elt,index) and lows = KM.insert lows (elt,index) val index = index + 1 val stack = pushStack stack elt val chil = toList (children elt) in dfsSuccessors sccs callstack index indices lows stack elt chil end and dfsSuccessors sccs callstack index indices lows stack elt chil = case chil of [] => let val (sccs,stack) = if isRoot indices lows elt then popStack elt sccs stack else (sccs,stack) in case callstack of [] => (sccs,index,indices,lows) | (p,elts) :: callstack => let val lows = reduceIndex lows (p, getIndex lows elt) in dfsSuccessors sccs callstack index indices lows stack p elts end end | c :: chil => case KM.peek indices c of NONE => let val callstack = (elt,chil) :: callstack in dfsVertex sccs callstack index indices lows stack c end | SOME cind => let val lows = if inStack c stack then reduceIndex lows (elt,cind) else lows in dfsSuccessors sccs callstack index indices lows stack elt chil end fun dfsRoots sccs index indices lows elts = case elts of [] => sccs | elt :: elts => if KM.inDomain elt indices then dfsRoots sccs index indices lows elts else let val callstack = [] val (sccs,index,indices,lows) = dfsVertex sccs callstack index indices lows emptyStack elt in dfsRoots sccs index indices lows elts end val sccs = [] and index = 0 and indices = KM.new () and lows = KM.new () in dfsRoots sccs index indices lows end; in fun preOrderSCC children roots = let val result = tarjan children (toList roots) (*BasicDebug val () = if subset roots (unionList result) then () else raise Metis_Useful.Bug "Metis_ElementSet.preOrderSCC: missing roots" val () = if preOrderedSCC children result then () else raise Metis_Useful.Bug "Metis_ElementSet.preOrderSCC: bad ordering" *) in result end; fun postOrderSCC children roots = let val result = List.rev (tarjan children (toList roots)) (*BasicDebug val () = if subset roots (unionList result) then () else raise Metis_Useful.Bug "Metis_ElementSet.postOrderSCC: missing roots" val () = if postOrderedSCC children result then () else raise Metis_Useful.Bug "Metis_ElementSet.postOrderSCC: bad ordering" *) in result end; end; (* ------------------------------------------------------------------------- *) (* Pretty-printing. *) (* ------------------------------------------------------------------------- *) fun toString set = "{" ^ (if null set then "" else Int.toString (size set)) ^ "}"; (* ------------------------------------------------------------------------- *) (* Iterators over sets *) (* ------------------------------------------------------------------------- *) type iterator = unit KM.iterator; fun mkIterator (Metis_Set m) = KM.mkIterator m; fun mkRevIterator (Metis_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 Metis_IntSet = Metis_ElementSet (Metis_IntMap); structure Metis_IntPairSet = Metis_ElementSet (Metis_IntPairMap); structure Metis_StringSet = Metis_ElementSet (Metis_StringMap); (**** Original file: src/Sharing.sig ****) (* ========================================================================= *) (* PRESERVING SHARING OF ML VALUES *) (* Copyright (c) 2005 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Sharing = sig (* ------------------------------------------------------------------------- *) (* Option operations. *) (* ------------------------------------------------------------------------- *) val mapOption : ('a -> 'a) -> 'a option -> 'a option val mapsOption : ('a -> 's -> 'a * 's) -> 'a option -> 's -> 'a option * 's (* ------------------------------------------------------------------------- *) (* List operations. *) (* ------------------------------------------------------------------------- *) 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 (* ------------------------------------------------------------------------- *) (* Function caching. *) (* ------------------------------------------------------------------------- *) val cache : ('a * 'a -> order) -> ('a -> 'b) -> 'a -> 'b (* ------------------------------------------------------------------------- *) (* Hash consing. *) (* ------------------------------------------------------------------------- *) val hashCons : ('a * 'a -> order) -> 'a -> 'a end (**** Original file: src/Sharing.sml ****) (* ========================================================================= *) (* PRESERVING SHARING OF ML VALUES *) (* Copyright (c) 2005 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Sharing :> Metis_Sharing = struct infix == val op== = Metis_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. *) (* ------------------------------------------------------------------------- *) fun map f = let 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 val (a,b) = Metis_Useful.revDivide l n in case b of [] => raise Subscript | h :: t => if x == h then l else List.revAppend (a, x :: t) end; fun setify l = let val l' = Metis_Useful.setify l in if length l' = length l then l else l' end; (* ------------------------------------------------------------------------- *) (* Function caching. *) (* ------------------------------------------------------------------------- *) fun cache cmp f = let val cache = Unsynchronized.ref (Metis_Map.new cmp) in fn a => case Metis_Map.peek (!cache) a of SOME b => b | NONE => let val b = f a val () = cache := Metis_Map.insert (!cache) (a,b) in b end end; (* ------------------------------------------------------------------------- *) (* Hash consing. *) (* ------------------------------------------------------------------------- *) fun hashCons cmp = cache cmp Metis_Useful.I; end (**** Original file: src/Stream.sig ****) (* ========================================================================= *) (* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_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; *) (* ------------------------------------------------------------------------- *) (* Metis_Stream constructors. *) (* ------------------------------------------------------------------------- *) val repeat : 'a -> 'a stream val count : int -> int stream val funpows : ('a -> 'a) -> 'a -> 'a stream (* ------------------------------------------------------------------------- *) (* Metis_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 *) val unfold : ('b -> ('a * 'b) option) -> 'b -> 'a stream (* ------------------------------------------------------------------------- *) (* Metis_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 (* ------------------------------------------------------------------------- *) (* Metis_Stream operations. *) (* ------------------------------------------------------------------------- *) val primes : int stream 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: src/Stream.sml ****) (* ========================================================================= *) (* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Stream :> Metis_Stream = struct val K = Metis_Useful.K; val pair = Metis_Useful.pair; val funpow = Metis_Useful.funpow; (* ------------------------------------------------------------------------- *) (* The stream type. *) (* ------------------------------------------------------------------------- *) datatype 'a stream = Nil | Cons of 'a * (unit -> 'a stream); (* ------------------------------------------------------------------------- *) (* Metis_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)); (* ------------------------------------------------------------------------- *) (* Metis_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; fun unfold f = let fun next b () = case f b of NONE => Nil | SOME (a,b) => Cons (a, next b) in fn b => next b () end; (* ------------------------------------------------------------------------- *) (* Metis_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; (* ------------------------------------------------------------------------- *) (* Metis_Stream operations. *) (* ------------------------------------------------------------------------- *) val primes = let fun next s = SOME (Metis_Useful.nextSieve s) in unfold next Metis_Useful.initSieve end; fun memoize Nil = Nil | memoize (Cons (h,t)) = Cons (h, Metis_Lazy.memoize (fn () => memoize (t ()))); fun concatList [] = Nil | concatList (h :: t) = append h (fn () => concatList t); local fun toLst res Nil = List.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 = String.implode (toList s); fun fromString s = fromList (String.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 (**** Original file: src/Heap.sig ****) (* ========================================================================= *) (* A HEAP DATATYPE FOR ML *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Heap = sig type 'a heap val new : ('a * 'a -> order) -> 'a heap val add : 'a heap -> 'a -> 'a heap val null : 'a heap -> bool val top : 'a heap -> 'a (* raises Empty *) val remove : 'a heap -> 'a * 'a heap (* raises Empty *) val size : 'a heap -> int val app : ('a -> unit) -> 'a heap -> unit val toList : 'a heap -> 'a list val toStream : 'a heap -> 'a Metis_Stream.stream val toString : 'a heap -> string end (**** Original file: src/Heap.sml ****) (* ========================================================================= *) (* A HEAP DATATYPE FOR ML *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Heap :> Metis_Heap = struct (* Leftist heaps as in Purely Functional Data Structures, by Chris Okasaki *) datatype 'a node = E | T of int * 'a * 'a node * 'a node; datatype 'a heap = Metis_Heap of ('a * 'a -> order) * int * 'a node; fun rank E = 0 | rank (T (r,_,_,_)) = r; fun makeT (x,a,b) = if rank a >= rank b then T (rank b + 1, x, a, b) else T (rank a + 1, x, b, a); fun merge cmp = let fun mrg (h,E) = h | mrg (E,h) = h | mrg (h1 as T (_,x,a1,b1), h2 as T (_,y,a2,b2)) = case cmp (x,y) of GREATER => makeT (y, a2, mrg (h1,b2)) | _ => makeT (x, a1, mrg (b1,h2)) in mrg end; fun new cmp = Metis_Heap (cmp,0,E); fun add (Metis_Heap (f,n,a)) x = Metis_Heap (f, n + 1, merge f (T (1,x,E,E), a)); fun size (Metis_Heap (_, n, _)) = n; fun null h = size h = 0; fun top (Metis_Heap (_,_,E)) = raise Empty | top (Metis_Heap (_, _, T (_,x,_,_))) = x; fun remove (Metis_Heap (_,_,E)) = raise Empty | remove (Metis_Heap (f, n, T (_,x,a,b))) = (x, Metis_Heap (f, n - 1, merge f (a,b))); fun app f = let fun ap [] = () | ap (E :: rest) = ap rest | ap (T (_,d,a,b) :: rest) = (f d; ap (a :: b :: rest)) in fn Metis_Heap (_,_,a) => ap [a] end; fun toList h = if null h then [] else let val (x,h) = remove h in x :: toList h end; fun toStream h = if null h then Metis_Stream.Nil else let val (x,h) = remove h in Metis_Stream.Cons (x, fn () => toStream h) end; fun toString h = "Metis_Heap[" ^ (if null h then "" else Int.toString (size h)) ^ "]"; end (**** Original file: src/Print.sig ****) (* ========================================================================= *) (* PRETTY-PRINTING *) (* Copyright (c) 2008 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Print = sig (* ------------------------------------------------------------------------- *) (* Escaping strings. *) (* ------------------------------------------------------------------------- *) val escapeString : {escape : char list} -> string -> string (* ------------------------------------------------------------------------- *) (* A type of pretty-printers. *) (* ------------------------------------------------------------------------- *) type ppstream type 'a pp = 'a -> 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 ppPpstream : ppstream pp (* ------------------------------------------------------------------------- *) (* Pretty-printing blocks. *) (* ------------------------------------------------------------------------- *) datatype style = Consistent | Inconsistent datatype block = Block of {style : style, indent : int} val styleBlock : block -> style val indentBlock : block -> int val block : block -> ppstream -> ppstream val consistentBlock : int -> ppstream list -> ppstream val inconsistentBlock : int -> ppstream list -> ppstream (* ------------------------------------------------------------------------- *) (* Words are unbreakable strings annotated with their effective size. *) (* ------------------------------------------------------------------------- *) datatype word = Word of {word : string, size : int} val mkWord : string -> word val emptyWord : word val charWord : char -> word val ppWord : word pp val space : ppstream val spaces : int -> ppstream (* ------------------------------------------------------------------------- *) (* Possible line breaks. *) (* ------------------------------------------------------------------------- *) datatype break = Break of {size : int, extraIndent : int} val mkBreak : int -> break val ppBreak : break pp val break : ppstream val breaks : int -> ppstream (* ------------------------------------------------------------------------- *) (* Forced line breaks. *) (* ------------------------------------------------------------------------- *) val newline : ppstream val newlines : int -> 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 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 ppException : exn pp (* ------------------------------------------------------------------------- *) (* Pretty-printing infix operators. *) (* ------------------------------------------------------------------------- *) type token = string datatype assoc = LeftAssoc | NonAssoc | RightAssoc datatype infixes = Infixes of {token : token, precedence : int, assoc : assoc} list val tokensInfixes : infixes -> Metis_StringSet.set val layerInfixes : infixes -> {tokens : Metis_StringSet.set, assoc : assoc} list val ppInfixes : infixes -> ('a -> (token * 'a * 'a) option) -> ('a * token) pp -> ('a * bool) pp -> ('a * bool) pp (* ------------------------------------------------------------------------- *) (* Pretty-printer rendering. *) (* ------------------------------------------------------------------------- *) val render : {lineLength : int option} -> ppstream -> {indent : int, line : string} Metis_Stream.stream val toStringWithLineLength : {lineLength : int option} -> 'a pp -> 'a -> string val toStreamWithLineLength : {lineLength : int option} -> 'a pp -> 'a -> string Metis_Stream.stream val toLine : 'a pp -> 'a -> string (* ------------------------------------------------------------------------- *) (* Pretty-printer rendering 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: src/Print.sml ****) (* ========================================================================= *) (* PRETTY-PRINTING *) (* Copyright (c) 2008 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Print :> Metis_Print = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* Constants. *) (* ------------------------------------------------------------------------- *) val initialLineLength = 75; (* ------------------------------------------------------------------------- *) (* Helper functions. *) (* ------------------------------------------------------------------------- *) fun revAppend xs s = case xs of [] => s () | x :: xs => revAppend xs (K (Metis_Stream.Cons (x,s))); fun revConcat strm = case strm of Metis_Stream.Nil => Metis_Stream.Nil | Metis_Stream.Cons (h,t) => revAppend h (fn () => revConcat (t ())); local fun calcSpaces n = nChars #" " n; val cacheSize = 2 * initialLineLength; val cachedSpaces = Vector.tabulate (cacheSize,calcSpaces); in fun nSpaces n = if n < cacheSize then Vector.sub (cachedSpaces,n) else calcSpaces n; end; (* ------------------------------------------------------------------------- *) (* Escaping strings. *) (* ------------------------------------------------------------------------- *) fun escapeString {escape} = let val escapeMap = List.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 String.translate escapeChar end; (* ------------------------------------------------------------------------- *) (* Pretty-printing blocks. *) (* ------------------------------------------------------------------------- *) datatype style = Consistent | Inconsistent; datatype block = Block of {style : style, indent : int}; fun toStringStyle style = case style of Consistent => "Consistent" | Inconsistent => "Inconsistent"; fun isConsistentStyle style = case style of Consistent => true | Inconsistent => false; fun isInconsistentStyle style = case style of Inconsistent => true | Consistent => false; fun mkBlock style indent = Block {style = style, indent = indent}; val mkConsistentBlock = mkBlock Consistent; val mkInconsistentBlock = mkBlock Inconsistent; fun styleBlock (Block {style = x, ...}) = x; fun indentBlock (Block {indent = x, ...}) = x; fun isConsistentBlock block = isConsistentStyle (styleBlock block); fun isInconsistentBlock block = isInconsistentStyle (styleBlock block); (* ------------------------------------------------------------------------- *) (* Words are unbreakable strings annotated with their effective size. *) (* ------------------------------------------------------------------------- *) datatype word = Word of {word : string, size : int}; fun mkWord s = Word {word = s, size = String.size s}; val emptyWord = mkWord ""; fun charWord c = mkWord (str c); fun spacesWord i = Word {word = nSpaces i, size = i}; fun sizeWord (Word {size = x, ...}) = x; fun renderWord (Word {word = x, ...}) = x; (* ------------------------------------------------------------------------- *) (* Possible line breaks. *) (* ------------------------------------------------------------------------- *) datatype break = Break of {size : int, extraIndent : int}; fun mkBreak n = Break {size = n, extraIndent = 0}; fun sizeBreak (Break {size = x, ...}) = x; fun extraIndentBreak (Break {extraIndent = x, ...}) = x; fun renderBreak b = nSpaces (sizeBreak b); fun updateSizeBreak size break = let val Break {size = _, extraIndent} = break in Break {size = size, extraIndent = extraIndent} end; fun appendBreak break1 break2 = let (*BasicDebug val () = warn "merging consecutive pretty-printing breaks" *) val Break {size = size1, extraIndent = extraIndent1} = break1 and Break {size = size2, extraIndent = extraIndent2} = break2 val size = size1 + size2 and extraIndent = Int.max (extraIndent1,extraIndent2) in Break {size = size, extraIndent = extraIndent} end; (* ------------------------------------------------------------------------- *) (* A type of pretty-printers. *) (* ------------------------------------------------------------------------- *) datatype step = BeginBlock of block | EndBlock | AddWord of word | AddBreak of break | AddNewline; type ppstream = step Metis_Stream.stream; type 'a pp = 'a -> ppstream; fun toStringStep step = case step of BeginBlock _ => "BeginBlock" | EndBlock => "EndBlock" | AddWord _ => "Word" | AddBreak _ => "Break" | AddNewline => "Newline"; val skip : ppstream = Metis_Stream.Nil; fun sequence pp1 pp2 : ppstream = Metis_Stream.append pp1 (K pp2); local fun dup pp n () = if n = 1 then pp else Metis_Stream.append pp (dup pp (n - 1)); in fun duplicate n pp : ppstream = let (*BasicDebug val () = if 0 <= n then () else raise Bug "Metis_Print.duplicate" *) in if n = 0 then skip else dup pp n () end; end; val program : ppstream list -> ppstream = Metis_Stream.concatList; val stream : ppstream Metis_Stream.stream -> ppstream = Metis_Stream.concat; (* ------------------------------------------------------------------------- *) (* Pretty-printing blocks. *) (* ------------------------------------------------------------------------- *) local fun beginBlock b = Metis_Stream.singleton (BeginBlock b); val endBlock = Metis_Stream.singleton EndBlock; in fun block b pp = program [beginBlock b, pp, endBlock]; end; fun consistentBlock i pps = block (mkConsistentBlock i) (program pps); fun inconsistentBlock i pps = block (mkInconsistentBlock i) (program pps); (* ------------------------------------------------------------------------- *) (* Words are unbreakable strings annotated with their effective size. *) (* ------------------------------------------------------------------------- *) fun ppWord w = Metis_Stream.singleton (AddWord w); val space = ppWord (mkWord " "); fun spaces i = ppWord (spacesWord i); (* ------------------------------------------------------------------------- *) (* Possible line breaks. *) (* ------------------------------------------------------------------------- *) fun ppBreak b = Metis_Stream.singleton (AddBreak b); fun breaks i = ppBreak (mkBreak i); val break = breaks 1; (* ------------------------------------------------------------------------- *) (* Forced line breaks. *) (* ------------------------------------------------------------------------- *) val newline = Metis_Stream.singleton AddNewline; fun newlines i = duplicate i newline; (* ------------------------------------------------------------------------- *) (* Pretty-printer combinators. *) (* ------------------------------------------------------------------------- *) fun ppMap f ppB a : ppstream = ppB (f a); fun ppBracket' l r ppA a = let val n = sizeWord l in inconsistentBlock n [ppWord l, ppA a, ppWord r] end; fun ppOp' w = sequence (ppWord w) break; fun ppOp2' ab ppA ppB (a,b) = inconsistentBlock 0 [ppA a, ppOp' ab, ppB b]; fun ppOp3' ab bc ppA ppB ppC (a,b,c) = inconsistentBlock 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 => inconsistentBlock 0 (ppA h :: List.map ppOpA t) end; fun ppOpStream' s ppA = let fun ppOpA a = sequence (ppOp' s) (ppA a) in fn Metis_Stream.Nil => skip | Metis_Stream.Cons (h,t) => inconsistentBlock 0 [ppA h, Metis_Stream.concat (Metis_Stream.map ppOpA (t ()))] end; fun ppBracket l r = ppBracket' (mkWord l) (mkWord r); fun ppOp s = ppOp' (mkWord s); fun ppOp2 ab = ppOp2' (mkWord ab); fun ppOp3 ab bc = ppOp3' (mkWord ab) (mkWord bc); fun ppOpList s = ppOpList' (mkWord s); fun ppOpStream s = ppOpStream' (mkWord s); (* ------------------------------------------------------------------------- *) (* Pretty-printers for common types. *) (* ------------------------------------------------------------------------- *) fun ppChar c = ppWord (charWord c); fun ppString s = ppWord (mkWord s); fun ppEscapeString escape = ppMap (escapeString escape) ppString; local val pp = ppString "()"; in fun ppUnit () = pp; end; local val ppTrue = ppString "true" and ppFalse = ppString "false"; in fun ppBool b = if b then ppTrue else ppFalse; end; val ppInt = ppMap Int.toString ppString; local val ppNeg = ppString "~" and ppSep = ppString "," and ppZero = ppString "0" and ppZeroZero = ppString "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; val ppReal = ppMap Real.toString ppString; val ppPercent = ppMap percentToString ppString; local val ppLess = ppString "Less" and ppEqual = ppString "Equal" and ppGreater = ppString "Greater"; in fun ppOrder ord = case ord of LESS => ppLess | EQUAL => ppEqual | GREATER => ppGreater; end; local val left = mkWord "[" and right = mkWord "]" and sep = mkWord ","; in fun ppList ppX xs = ppBracket' left right (ppOpList' sep ppX) xs; fun ppStream ppX xs = ppBracket' left right (ppOpStream' sep ppX) xs; end; local val ppNone = ppString "-"; in fun ppOption ppX xo = case xo of SOME x => ppX x | NONE => ppNone; end; local val left = mkWord "(" and right = mkWord ")" and sep = mkWord ","; in fun ppPair ppA ppB = ppBracket' left right (ppOp2' sep ppA ppB); fun ppTriple ppA ppB ppC = ppBracket' left right (ppOp3' sep sep ppA ppB ppC); end; fun ppException e = ppString (exnMessage e); (* ------------------------------------------------------------------------- *) (* A type of pretty-printers. *) (* ------------------------------------------------------------------------- *) local val ppStepType = ppMap toStringStep ppString; val ppStyle = ppMap toStringStyle ppString; val ppBlockInfo = let val sep = mkWord " " in fn Block {style = s, indent = i} => program [ppStyle s, ppWord sep, ppInt i] end; val ppWordInfo = let val left = mkWord "\"" and right = mkWord "\"" and escape = {escape = [#"\""]} val pp = ppBracket' left right (ppEscapeString escape) in fn Word {word = s, size = n} => if size s = n then pp s else ppPair pp ppInt (s,n) end; val ppBreakInfo = let val sep = mkWord "+" in fn Break {size = n, extraIndent = k} => if k = 0 then ppInt n else program [ppInt n, ppWord sep, ppInt k] end; fun ppStep step = inconsistentBlock 2 (ppStepType step :: (case step of BeginBlock b => [break, ppBlockInfo b] | EndBlock => [] | AddWord w => [break, ppWordInfo w] | AddBreak b => [break, ppBreakInfo b] | AddNewline => [])); in val ppPpstream = ppStream ppStep; end; (* ------------------------------------------------------------------------- *) (* Pretty-printing infix operators. *) (* ------------------------------------------------------------------------- *) type token = string; datatype assoc = LeftAssoc | NonAssoc | RightAssoc; datatype infixes = Infixes of {token : token, precedence : int, assoc : assoc} list; local fun comparePrecedence (io1,io2) = let val {token = _, precedence = p1, assoc = _} = io1 and {token = _, precedence = p2, assoc = _} = io2 in Int.compare (p2,p1) end; fun equalAssoc a a' = case a of LeftAssoc => (case a' of LeftAssoc => true | _ => false) | NonAssoc => (case a' of NonAssoc => true | _ => false) | RightAssoc => (case a' of RightAssoc => true | _ => false); fun new t a acc = {tokens = Metis_StringSet.singleton t, assoc = a} :: acc; fun add t a' acc = case acc of [] => raise Bug "Metis_Print.layerInfixes: null" | {tokens = ts, assoc = a} :: acc => if equalAssoc a a' then {tokens = Metis_StringSet.add ts t, assoc = a} :: acc else raise Bug "Metis_Print.layerInfixes: mixed assocs"; fun layer ({token = t, precedence = p, assoc = a}, (acc,p')) = let val acc = if p = p' then add t a acc else new t a acc in (acc,p) end; in fun layerInfixes (Infixes ios) = case sort comparePrecedence ios of [] => [] | {token = t, precedence = p, assoc = a} :: ios => let val acc = new t a [] val (acc,_) = List.foldl layer (acc,p) ios in acc end; end; local fun add ({tokens = ts, assoc = _}, tokens) = Metis_StringSet.union ts tokens; in fun tokensLayeredInfixes l = List.foldl add Metis_StringSet.empty l; end; fun tokensInfixes ios = tokensLayeredInfixes (layerInfixes ios); fun destInfixOp dest tokens tm = case dest tm of NONE => NONE | s as SOME (t,a,b) => if Metis_StringSet.member t tokens then s else NONE; fun ppLayeredInfixes dest ppTok {tokens,assoc} ppLower ppSub = let fun isLayer t = Metis_StringSet.member t tokens fun ppTerm aligned (tm,r) = case dest tm of NONE => ppSub (tm,r) | SOME (t,a,b) => if aligned andalso isLayer t then ppLayer (tm,t,a,b,r) else ppLower (tm,t,a,b,r) and ppLeft tm_r = let val aligned = case assoc of LeftAssoc => true | _ => false in ppTerm aligned tm_r end and ppLayer (tm,t,a,b,r) = program [ppLeft (a,true), ppTok (tm,t), ppRight (b,r)] and ppRight tm_r = let val aligned = case assoc of RightAssoc => true | _ => false in ppTerm aligned tm_r end in fn tm_t_a_b_r as (_,t,_,_,_) => if isLayer t then inconsistentBlock 0 [ppLayer tm_t_a_b_r] else ppLower tm_t_a_b_r end; local val leftBrack = mkWord "(" and rightBrack = mkWord ")"; in fun ppInfixes ops = let val layers = layerInfixes ops val toks = tokensLayeredInfixes layers in fn dest => fn ppTok => fn ppSub => let fun destOp tm = destInfixOp dest toks tm fun ppInfix tm_t_a_b_r = ppLayers layers tm_t_a_b_r and ppLayers ls (tm,t,a,b,r) = case ls of [] => ppBracket' leftBrack rightBrack ppInfix (tm,t,a,b,false) | l :: ls => let val ppLower = ppLayers ls in ppLayeredInfixes destOp ppTok l ppLower ppSub (tm,t,a,b,r) end in fn (tm,r) => case destOp tm of SOME (t,a,b) => ppInfix (tm,t,a,b,r) | NONE => ppSub (tm,r) end end; end; (* ------------------------------------------------------------------------- *) (* A type of output lines. *) (* ------------------------------------------------------------------------- *) type line = {indent : int, line : string}; val emptyLine = let val indent = 0 and line = "" in {indent = indent, line = line} end; fun addEmptyLine lines = emptyLine :: lines; fun addLine lines indent line = {indent = indent, line = line} :: lines; (* ------------------------------------------------------------------------- *) (* Pretty-printer rendering. *) (* ------------------------------------------------------------------------- *) datatype chunk = WordChunk of word | BreakChunk of break | FrameChunk of frame and frame = Frame of {block : block, broken : bool, indent : int, size : int, chunks : chunk list}; datatype state = State of {lineIndent : int, lineSize : int, stack : frame list}; fun blockFrame (Frame {block = x, ...}) = x; fun brokenFrame (Frame {broken = x, ...}) = x; fun indentFrame (Frame {indent = x, ...}) = x; fun sizeFrame (Frame {size = x, ...}) = x; fun chunksFrame (Frame {chunks = x, ...}) = x; fun styleFrame frame = styleBlock (blockFrame frame); fun isConsistentFrame frame = isConsistentBlock (blockFrame frame); fun breakingFrame frame = isConsistentFrame frame andalso brokenFrame frame; fun sizeChunk chunk = case chunk of WordChunk w => sizeWord w | BreakChunk b => sizeBreak b | FrameChunk f => sizeFrame f; local fun add (c,acc) = sizeChunk c + acc; in fun sizeChunks cs = List.foldl add 0 cs; end; local fun flattenChunks acc chunks = case chunks of [] => acc | chunk :: chunks => flattenChunk acc chunk chunks and flattenChunk acc chunk chunks = case chunk of WordChunk w => flattenChunks (renderWord w :: acc) chunks | BreakChunk b => flattenChunks (renderBreak b :: acc) chunks | FrameChunk f => flattenFrame acc f chunks and flattenFrame acc frame chunks = flattenChunks acc (chunksFrame frame @ chunks); in fun renderChunks chunks = String.concat (flattenChunks [] chunks); end; fun addChunksLine lines indent chunks = addLine lines indent (renderChunks chunks); local fun add baseIndent ((extraIndent,chunks),lines) = addChunksLine lines (baseIndent + extraIndent) chunks; in fun addIndentChunksLines lines baseIndent indent_chunks = List.foldl (add baseIndent) lines indent_chunks; end; fun isEmptyFrame frame = let val chunks = chunksFrame frame val empty = List.null chunks (*BasicDebug val () = if not empty orelse sizeFrame frame = 0 then () else raise Bug "Metis_Print.isEmptyFrame: bad size" *) in empty end; local fun breakInconsistent blockIndent = let fun break chunks = case chunks of [] => NONE | chunk :: chunks => case chunk of BreakChunk b => let val pre = chunks and indent = blockIndent + extraIndentBreak b and post = [] in SOME (pre,indent,post) end | _ => case break chunks of SOME (pre,indent,post) => let val post = chunk :: post in SOME (pre,indent,post) end | NONE => NONE in break end; fun breakConsistent blockIndent = let fun break indent_chunks chunks = case breakInconsistent blockIndent chunks of NONE => (chunks,indent_chunks) | SOME (pre,indent,post) => break ((indent,post) :: indent_chunks) pre in break [] end; in fun breakFrame frame = let val Frame {block, broken = _, indent = _, size = _, chunks} = frame val blockIndent = indentBlock block in case breakInconsistent blockIndent chunks of NONE => NONE | SOME (pre,indent,post) => let val broken = true and size = sizeChunks post val frame = Frame {block = block, broken = broken, indent = indent, size = size, chunks = post} in case styleBlock block of Inconsistent => let val indent_chunks = [] in SOME (pre,indent_chunks,frame) end | Consistent => let val (pre,indent_chunks) = breakConsistent blockIndent pre in SOME (pre,indent_chunks,frame) end end end; end; fun removeChunksFrame frame = let val Frame {block, broken, indent, size = _, chunks} = frame in if broken andalso List.null chunks then NONE else let val frame = Frame {block = block, broken = true, indent = indent, size = 0, chunks = []} in SOME (chunks,frame) end end; val removeChunksFrames = let fun remove frames = case frames of [] => let val chunks = [] and frames = NONE and indent = 0 in (chunks,frames,indent) end | top :: rest => let val (chunks,rest',indent) = remove rest val indent = indent + indentFrame top val (chunks,top') = case removeChunksFrame top of NONE => (chunks,NONE) | SOME (topChunks,top) => (topChunks @ chunks, SOME top) val frames' = case (top',rest') of (NONE,NONE) => NONE | (SOME top, NONE) => SOME (top :: rest) | (NONE, SOME rest) => SOME (top :: rest) | (SOME top, SOME rest) => SOME (top :: rest) in (chunks,frames',indent) end in fn frames => let val (chunks,frames',indent) = remove frames val frames = Option.getOpt (frames',frames) in (chunks,frames,indent) end end; local fun breakUp lines lineIndent frames = case frames of [] => NONE | frame :: frames => case breakUp lines lineIndent frames of SOME (lines_indent,lineSize,frames) => let val lineSize = lineSize + sizeFrame frame and frames = frame :: frames in SOME (lines_indent,lineSize,frames) end | NONE => case breakFrame frame of NONE => NONE | SOME (frameChunks,indent_chunks,frame) => let val (chunks,frames,indent) = removeChunksFrames frames val chunks = frameChunks @ chunks val lines = addChunksLine lines lineIndent chunks val lines = addIndentChunksLines lines indent indent_chunks val lineIndent = indent + indentFrame frame val lineSize = sizeFrame frame val frames = frame :: frames in SOME ((lines,lineIndent),lineSize,frames) end; fun breakInsideChunk chunk = case chunk of WordChunk _ => NONE | BreakChunk _ => raise Bug "Metis_Print.breakInsideChunk" | FrameChunk frame => case breakFrame frame of SOME (pathChunks,indent_chunks,frame) => let val pathIndent = 0 and breakIndent = indentFrame frame in SOME (pathChunks,pathIndent,indent_chunks,breakIndent,frame) end | NONE => breakInsideFrame frame and breakInsideChunks chunks = case chunks of [] => NONE | chunk :: chunks => case breakInsideChunk chunk of SOME (pathChunks,pathIndent,indent_chunks,breakIndent,frame) => let val pathChunks = pathChunks @ chunks and chunks = [FrameChunk frame] in SOME (pathChunks,pathIndent,indent_chunks,breakIndent,chunks) end | NONE => case breakInsideChunks chunks of SOME (pathChunks,pathIndent,indent_chunks,breakIndent,chunks) => let val chunks = chunk :: chunks in SOME (pathChunks,pathIndent,indent_chunks,breakIndent,chunks) end | NONE => NONE and breakInsideFrame frame = let val Frame {block, broken = _, indent, size = _, chunks} = frame in case breakInsideChunks chunks of SOME (pathChunks,pathIndent,indent_chunks,breakIndent,chunks) => let val pathIndent = pathIndent + indent and broken = true and size = sizeChunks chunks val frame = Frame {block = block, broken = broken, indent = indent, size = size, chunks = chunks} in SOME (pathChunks,pathIndent,indent_chunks,breakIndent,frame) end | NONE => NONE end; fun breakInside lines lineIndent frames = case frames of [] => NONE | frame :: frames => case breakInsideFrame frame of SOME (pathChunks,pathIndent,indent_chunks,breakIndent,frame) => let val (chunks,frames,indent) = removeChunksFrames frames val chunks = pathChunks @ chunks and indent = indent + pathIndent val lines = addChunksLine lines lineIndent chunks val lines = addIndentChunksLines lines indent indent_chunks val lineIndent = indent + breakIndent val lineSize = sizeFrame frame val frames = frame :: frames in SOME ((lines,lineIndent),lineSize,frames) end | NONE => case breakInside lines lineIndent frames of SOME (lines_indent,lineSize,frames) => let val lineSize = lineSize + sizeFrame frame and frames = frame :: frames in SOME (lines_indent,lineSize,frames) end | NONE => NONE; in fun breakFrames lines lineIndent frames = case breakUp lines lineIndent frames of SOME ((lines,lineIndent),lineSize,frames) => SOME (lines,lineIndent,lineSize,frames) | NONE => case breakInside lines lineIndent frames of SOME ((lines,lineIndent),lineSize,frames) => SOME (lines,lineIndent,lineSize,frames) | NONE => NONE; end; (*BasicDebug fun checkChunk chunk = case chunk of WordChunk t => (false, sizeWord t) | BreakChunk b => (false, sizeBreak b) | FrameChunk b => checkFrame b and checkChunks chunks = case chunks of [] => (false,0) | chunk :: chunks => let val (bk,sz) = checkChunk chunk val (bk',sz') = checkChunks chunks in (bk orelse bk', sz + sz') end and checkFrame frame = let val Frame {block = _, broken, indent = _, size, chunks} = frame val (bk,sz) = checkChunks chunks val () = if size = sz then () else raise Bug "Metis_Print.checkFrame: wrong size" val () = if broken orelse not bk then () else raise Bug "Metis_Print.checkFrame: deep broken frame" in (broken,size) end; fun checkFrames frames = case frames of [] => (true,0) | frame :: frames => let val (bk,sz) = checkFrame frame val (bk',sz') = checkFrames frames val () = if bk' orelse not bk then () else raise Bug "Metis_Print.checkFrame: broken stack frame" in (bk, sz + sz') end; *) (*BasicDebug fun checkState state = (let val State {lineIndent,lineSize,stack} = state val () = if not (List.null stack) then () else raise Error "no stack" val (_,sz) = checkFrames stack val () = if lineSize = sz then () else raise Error "wrong lineSize" in () end handle Error err => raise Bug err) handle Bug bug => raise Bug ("Metis_Print.checkState: " ^ bug); *) fun isEmptyState state = let val State {lineSize,stack,...} = state in lineSize = 0 andalso List.all isEmptyFrame stack end; fun isFinalState state = let val State {stack,...} = state in case stack of [] => raise Bug "Metis_Print.isFinalState: empty stack" | [frame] => isEmptyFrame frame | _ :: _ :: _ => false end; local val initialBlock = let val indent = 0 and style = Inconsistent in Block {indent = indent, style = style} end; val initialFrame = let val block = initialBlock and broken = false and indent = 0 and size = 0 and chunks = [] in Frame {block = block, broken = broken, indent = indent, size = size, chunks = chunks} end; in val initialState = let val lineIndent = 0 and lineSize = 0 and stack = [initialFrame] in State {lineIndent = lineIndent, lineSize = lineSize, stack = stack} end; end; fun normalizeState lineLength lines state = let val State {lineIndent,lineSize,stack} = state val within = case lineLength of NONE => true | SOME len => lineIndent + lineSize <= len in if within then (lines,state) else case breakFrames lines lineIndent stack of NONE => (lines,state) | SOME (lines,lineIndent,lineSize,stack) => let (*BasicDebug val () = checkState state *) val state = State {lineIndent = lineIndent, lineSize = lineSize, stack = stack} in normalizeState lineLength lines state end end (*BasicDebug handle Bug bug => raise Bug ("Metis_Print.normalizeState:\n" ^ bug) *) local fun executeBeginBlock block lines state = let val State {lineIndent,lineSize,stack} = state val broken = false and indent = indentBlock block and size = 0 and chunks = [] val frame = Frame {block = block, broken = broken, indent = indent, size = size, chunks = chunks} val stack = frame :: stack val state = State {lineIndent = lineIndent, lineSize = lineSize, stack = stack} in (lines,state) end; fun executeEndBlock lines state = let val State {lineIndent,lineSize,stack} = state val (lineSize,stack) = case stack of [] => raise Bug "Metis_Print.executeEndBlock: empty stack" | topFrame :: stack => let val Frame {block = topBlock, broken = topBroken, indent = topIndent, size = topSize, chunks = topChunks} = topFrame val (lineSize,topSize,topChunks,topFrame) = case topChunks of BreakChunk break :: chunks => let (*BasicDebug val () = let val mesg = "ignoring a break at the end of a " ^ "pretty-printing block" in warn mesg end *) val n = sizeBreak break val lineSize = lineSize - n and topSize = topSize - n and topChunks = chunks val topFrame = Frame {block = topBlock, broken = topBroken, indent = topIndent, size = topSize, chunks = topChunks} in (lineSize,topSize,topChunks,topFrame) end | _ => (lineSize,topSize,topChunks,topFrame) in if List.null topChunks then (lineSize,stack) else case stack of [] => raise Error "Metis_Print.execute: too many end blocks" | frame :: stack => let val Frame {block, broken, indent, size, chunks} = frame val size = size + topSize val chunk = FrameChunk topFrame val chunks = chunk :: chunks val frame = Frame {block = block, broken = broken, indent = indent, size = size, chunks = chunks} val stack = frame :: stack in (lineSize,stack) end end val state = State {lineIndent = lineIndent, lineSize = lineSize, stack = stack} in (lines,state) end; fun executeAddWord lineLength word lines state = let val State {lineIndent,lineSize,stack} = state val n = sizeWord word val lineSize = lineSize + n val stack = case stack of [] => raise Bug "Metis_Print.executeAddWord: empty stack" | frame :: stack => let val Frame {block, broken, indent, size, chunks} = frame val size = size + n val chunk = WordChunk word val chunks = chunk :: chunks val frame = Frame {block = block, broken = broken, indent = indent, size = size, chunks = chunks} val stack = frame :: stack in stack end val state = State {lineIndent = lineIndent, lineSize = lineSize, stack = stack} in normalizeState lineLength lines state end; fun executeAddBreak lineLength break lines state = let val State {lineIndent,lineSize,stack} = state val (topFrame,restFrames) = case stack of [] => raise Bug "Metis_Print.executeAddBreak: empty stack" | topFrame :: restFrames => (topFrame,restFrames) val Frame {block = topBlock, broken = topBroken, indent = topIndent, size = topSize, chunks = topChunks} = topFrame in case topChunks of [] => (lines,state) | topChunk :: restTopChunks => let val (topChunks,n) = case topChunk of BreakChunk break' => let val break = appendBreak break' break val chunk = BreakChunk break val topChunks = chunk :: restTopChunks and n = sizeBreak break - sizeBreak break' in (topChunks,n) end | _ => let val chunk = BreakChunk break val topChunks = chunk :: topChunks and n = sizeBreak break in (topChunks,n) end val lineSize = lineSize + n val topSize = topSize + n val topFrame = Frame {block = topBlock, broken = topBroken, indent = topIndent, size = topSize, chunks = topChunks} val stack = topFrame :: restFrames val state = State {lineIndent = lineIndent, lineSize = lineSize, stack = stack} val lineLength = if breakingFrame topFrame then SOME ~1 else lineLength in normalizeState lineLength lines state end end; fun executeBigBreak lines state = let val lineLength = SOME ~1 and break = mkBreak 0 in executeAddBreak lineLength break lines state end; fun executeAddNewline lineLength lines state = if isEmptyState state then (addEmptyLine lines, state) else executeBigBreak lines state; fun executeEof lineLength lines state = if isFinalState state then (lines,state) else let val (lines,state) = executeBigBreak lines state (*BasicDebug val () = if isFinalState state then () else raise Bug "Metis_Print.executeEof: not a final state" *) in (lines,state) end; in fun render {lineLength} = let fun execute step state = let val lines = [] in case step of BeginBlock block => executeBeginBlock block lines state | EndBlock => executeEndBlock lines state | AddWord word => executeAddWord lineLength word lines state | AddBreak break => executeAddBreak lineLength break lines state | AddNewline => executeAddNewline lineLength lines state end (*BasicDebug val execute = fn step => fn state => let val (lines,state) = execute step state val () = checkState state in (lines,state) end handle Bug bug => raise Bug ("Metis_Print.execute: after " ^ toStringStep step ^ " command:\n" ^ bug) *) fun final state = let val lines = [] val (lines,state) = executeEof lineLength lines state (*BasicDebug val () = checkState state *) in if List.null lines then Metis_Stream.Nil else Metis_Stream.singleton lines end (*BasicDebug handle Bug bug => raise Bug ("Metis_Print.final: " ^ bug) *) in fn pps => let val lines = Metis_Stream.maps execute final initialState pps in revConcat lines end end; end; local fun inc {indent,line} acc = line :: nSpaces indent :: acc; fun incn (indent_line,acc) = inc indent_line ("\n" :: acc); in fun toStringWithLineLength len ppA a = case render len (ppA a) of Metis_Stream.Nil => "" | Metis_Stream.Cons (h,t) => let val lines = Metis_Stream.foldl incn (inc h []) (t ()) in String.concat (List.rev lines) end; end; local fun renderLine {indent,line} = nSpaces indent ^ line ^ "\n"; in fun toStreamWithLineLength len ppA a = Metis_Stream.map renderLine (render len (ppA a)); end; fun toLine ppA a = toStringWithLineLength {lineLength = NONE} ppA a; (* ------------------------------------------------------------------------- *) (* Pretty-printer rendering with a global line length. *) (* ------------------------------------------------------------------------- *) val lineLength = Unsynchronized.ref initialLineLength; fun toString ppA a = let val len = {lineLength = SOME (!lineLength)} in toStringWithLineLength len ppA a end; fun toStream ppA a = let val len = {lineLength = SOME (!lineLength)} in toStreamWithLineLength len ppA a end; fun trace ppX nameX = let fun ppNameX x = consistentBlock 2 [ppString nameX, ppString " =", break, ppX x] in fn x => Metis_Useful.trace (toString ppNameX x ^ "\n") end; end (**** Original file: src/Parse.sig ****) (* ========================================================================= *) (* PARSING *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_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 (* ------------------------------------------------------------------------- *) (* Metis_Stream-based parsers. *) (* ------------------------------------------------------------------------- *) type ('a,'b) parser = 'a Metis_Stream.stream -> 'b * 'a Metis_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 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 : Metis_Print.infixes -> (Metis_Print.token * 'a * 'a -> 'a) -> ('b,Metis_Print.token) parser -> ('b,'a) parser -> ('b,'a) parser (* ------------------------------------------------------------------------- *) (* Quotations. *) (* ------------------------------------------------------------------------- *) type 'a quotation = 'a frag list val parseQuotation : ('a -> string) -> (string -> 'b) -> 'a quotation -> 'b end (**** Original file: src/Parse.sml ****) (* ========================================================================= *) (* PARSING *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Parse :> Metis_Parse = struct open Metis_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 [] >> List.rev end; fun atLeastOne p = (p ++ many p) >> op::; fun nothing input = ((),input); fun optional p = (p >> SOME) || (nothing >> K NONE); (* ------------------------------------------------------------------------- *) (* Metis_Stream-based parsers. *) (* ------------------------------------------------------------------------- *) type ('a,'b) parser = 'a Metis_Stream.stream -> 'b * 'a Metis_Stream.stream fun maybe p Metis_Stream.Nil = raise NoParse | maybe p (Metis_Stream.Cons (h,t)) = case p h of SOME r => (r, t ()) | NONE => raise NoParse; fun finished Metis_Stream.Nil = ((), Metis_Stream.Nil) | finished (Metis_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 (Metis_Stream.fromList l); fun everything parser = let fun parserOption input = SOME (parser input) handle e as NoParse => if Metis_Stream.null input then NONE else raise e fun parserList input = case parserOption input of NONE => Metis_Stream.Nil | SOME (result,input) => Metis_Stream.append (Metis_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 String.explode line end in Metis_Stream.memoize (Metis_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 (String.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 >> String.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 (String.explode s); (* ------------------------------------------------------------------------- *) (* Infix operators. *) (* ------------------------------------------------------------------------- *) fun parseLayeredInfixes {tokens,assoc} mk tokParser subParser = let fun layerTokParser inp = let val tok_rest as (tok,_) = tokParser inp in if Metis_StringSet.member tok tokens then tok_rest else raise NoParse end fun layerMk (x,txs) = case assoc of Metis_Print.LeftAssoc => let fun inc ((t,y),z) = mk (t,z,y) in List.foldl inc x txs end | Metis_Print.NonAssoc => (case txs of [] => x | [(t,y)] => mk (t,x,y) | _ => raise NoParse) | Metis_Print.RightAssoc => (case List.rev txs of [] => x | tx :: txs => let fun inc ((t,y),(u,z)) = (t, mk (u,y,z)) val (t,y) = List.foldl inc tx txs in mk (t,x,y) end) val layerParser = subParser ++ many (layerTokParser ++ subParser) in layerParser >> layerMk end; fun parseInfixes ops = let val layeredOps = Metis_Print.layerInfixes ops val iparsers = List.map parseLayeredInfixes layeredOps in fn mk => fn tokParser => fn subParser => List.foldr (fn (p,sp) => p mk tokParser 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 = List.foldl expand "" quote in parser string end; end (**** Original file: src/Name.sig ****) (* ========================================================================= *) (* NAMES *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Name = sig (* ------------------------------------------------------------------------- *) (* A type of names. *) (* ------------------------------------------------------------------------- *) type name (* ------------------------------------------------------------------------- *) (* A total ordering. *) (* ------------------------------------------------------------------------- *) val compare : name * name -> order val equal : name -> name -> bool (* ------------------------------------------------------------------------- *) (* Fresh names. *) (* ------------------------------------------------------------------------- *) val newName : unit -> name val newNames : int -> name list val variantPrime : {avoid : name -> bool} -> name -> name val variantNum : {avoid : name -> bool} -> name -> name (* ------------------------------------------------------------------------- *) (* Parsing and pretty printing. *) (* ------------------------------------------------------------------------- *) val pp : name Metis_Print.pp val toString : name -> string val fromString : string -> name end (**** Original file: src/Name.sml ****) (* ========================================================================= *) (* NAMES *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Name :> Metis_Name = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* A type of names. *) (* ------------------------------------------------------------------------- *) type name = string; (* ------------------------------------------------------------------------- *) (* A total ordering. *) (* ------------------------------------------------------------------------- *) val compare = String.compare; 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 = List.map numName (newInts n); end; fun variantPrime {avoid} = let fun variant n = if avoid n then variant (n ^ "'") else n in variant end; local fun isDigitOrPrime c = c = #"'" orelse Char.isDigit c; in fun variantNum {avoid} n = if not (avoid n) then n else let val n = stripSuffix isDigitOrPrime n fun variant i = let val n_i = n ^ Int.toString i in if avoid n_i then variant (i + 1) else n_i end in variant 0 end; end; (* ------------------------------------------------------------------------- *) (* Parsing and pretty printing. *) (* ------------------------------------------------------------------------- *) val pp = Metis_Print.ppString; fun toString s : string = s; fun fromString s : name = s; end structure Metis_NameOrdered = struct type t = Metis_Name.name val compare = Metis_Name.compare end structure Metis_NameMap = Metis_KeyMap (Metis_NameOrdered); structure Metis_NameSet = struct local structure S = Metis_ElementSet (Metis_NameMap); in open S; end; val pp = Metis_Print.ppMap toList (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_Name.pp)); end (**** Original file: src/NameArity.sig ****) (* ========================================================================= *) (* NAME/ARITY PAIRS *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_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: src/NameArity.sml ****) (* ========================================================================= *) (* NAME/ARITY PAIRS *) (* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_NameArity :> Metis_NameArity = struct (* ------------------------------------------------------------------------- *) (* A type of name/arity pairs. *) (* ------------------------------------------------------------------------- *) type nameArity = Metis_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 Metis_Name.compare (n1,n2) of LESS => LESS | EQUAL => Int.compare (i1,i2) | GREATER => GREATER; fun equal (n1,i1) (n2,i2) = i1 = i2 andalso Metis_Name.equal n1 n2; (* ------------------------------------------------------------------------- *) (* Parsing and pretty printing. *) (* ------------------------------------------------------------------------- *) fun pp (n,i) = Metis_Print.inconsistentBlock 0 [Metis_Name.pp n, Metis_Print.ppString "/", Metis_Print.ppInt i]; end structure Metis_NameArityOrdered = struct type t = Metis_NameArity.nameArity val compare = Metis_NameArity.compare end structure Metis_NameArityMap = struct local structure S = Metis_KeyMap (Metis_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 Metis_NameAritySet = struct local structure S = Metis_ElementSet (Metis_NameArityMap); in open S; end; val allNullary = all Metis_NameArity.nullary; val pp = Metis_Print.ppMap toList (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_NameArity.pp)); end (**** Original file: src/Term.sig ****) (* ========================================================================= *) (* FIRST ORDER LOGIC TERMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Term = sig (* ------------------------------------------------------------------------- *) (* A type of first order logic terms. *) (* ------------------------------------------------------------------------- *) type var = Metis_Name.name type functionName = Metis_Name.name type function = functionName * int type const = functionName datatype term = Var of var | Fn of functionName * term list (* ------------------------------------------------------------------------- *) (* Constructors and destructors. *) (* ------------------------------------------------------------------------- *) (* Variables *) val destVar : term -> var val isVar : term -> bool val equalVar : var -> term -> bool (* Functions *) val destFn : term -> functionName * term list val isFn : term -> bool val fnName : term -> functionName val fnArguments : term -> term list val fnArity : term -> int val fnFunction : term -> function val functions : term -> Metis_NameAritySet.set val functionNames : term -> Metis_NameSet.set (* Constants *) val mkConst : const -> term val destConst : term -> const val isConst : term -> bool (* Binary functions *) val mkBinop : functionName -> term * term -> term val destBinop : functionName -> term -> term * term val isBinop : functionName -> term -> bool (* ------------------------------------------------------------------------- *) (* The size of a term in symbols. *) (* ------------------------------------------------------------------------- *) val symbols : term -> int (* ------------------------------------------------------------------------- *) (* A total comparison function for terms. *) (* ------------------------------------------------------------------------- *) val compare : term * term -> order val equal : term -> term -> bool (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) type path = int list val subterm : term -> path -> term val subterms : term -> (path * term) list val replace : term -> path * term -> term val find : (term -> bool) -> term -> path option val ppPath : path Metis_Print.pp val pathToString : path -> string (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) val freeIn : var -> term -> bool val freeVars : term -> Metis_NameSet.set val freeVarsList : term list -> Metis_NameSet.set (* ------------------------------------------------------------------------- *) (* Fresh variables. *) (* ------------------------------------------------------------------------- *) val newVar : unit -> term val newVars : int -> term list val variantPrime : Metis_NameSet.set -> var -> var val variantNum : Metis_NameSet.set -> var -> var (* ------------------------------------------------------------------------- *) (* Special support for terms with type annotations. *) (* ------------------------------------------------------------------------- *) val hasTypeFunctionName : functionName val hasTypeFunction : function val isTypedVar : term -> bool val typedSymbols : term -> int val nonVarTypedSubterms : term -> (path * term) list (* ------------------------------------------------------------------------- *) (* Special support for terms with an explicit function application operator. *) (* ------------------------------------------------------------------------- *) 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. *) (* ------------------------------------------------------------------------- *) (* Infix symbols *) val infixes : Metis_Print.infixes Unsynchronized.ref (* The negation symbol *) val negation : string Unsynchronized.ref (* Binder symbols *) val binders : string list Unsynchronized.ref (* Bracket symbols *) val brackets : (string * string) list Unsynchronized.ref (* Pretty printing *) val pp : term Metis_Print.pp val toString : term -> string (* Parsing *) val fromString : string -> term val parse : term Metis_Parse.quotation -> term end (**** Original file: src/Term.sml ****) (* ========================================================================= *) (* FIRST ORDER LOGIC TERMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Term :> Metis_Term = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* A type of first order logic terms. *) (* ------------------------------------------------------------------------- *) type var = Metis_Name.name; type functionName = Metis_Name.name; type function = functionName * int; type const = functionName; datatype term = Var of Metis_Name.name | Fn of Metis_Name.name * term list; (* ------------------------------------------------------------------------- *) (* Constructors and destructors. *) (* ------------------------------------------------------------------------- *) (* Variables *) fun destVar (Var v) = v | destVar (Fn _) = raise Error "destVar"; val isVar = can destVar; fun equalVar v (Var v') = Metis_Name.equal v v' | equalVar _ _ = false; (* Functions *) fun destFn (Fn f) = f | destFn (Var _) = raise Error "destFn"; val isFn = can destFn; fun fnName tm = fst (destFn tm); fun fnArguments tm = snd (destFn tm); fun fnArity tm = length (fnArguments tm); fun fnFunction tm = (fnName tm, fnArity tm); local fun func fs [] = fs | func fs (Var _ :: tms) = func fs tms | func fs (Fn (n,l) :: tms) = func (Metis_NameAritySet.add fs (n, length l)) (l @ tms); in fun functions tm = func Metis_NameAritySet.empty [tm]; end; local fun func fs [] = fs | func fs (Var _ :: tms) = func fs tms | func fs (Fn (n,l) :: tms) = func (Metis_NameSet.add fs n) (l @ tms); in fun functionNames tm = func Metis_NameSet.empty [tm]; end; (* Constants *) fun mkConst c = (Fn (c, [])); fun destConst (Fn (c, [])) = c | destConst _ = raise Error "destConst"; val isConst = can destConst; (* Binary functions *) fun mkBinop f (a,b) = Fn (f,[a,b]); fun destBinop f (Fn (x,[a,b])) = if Metis_Name.equal x f then (a,b) else raise Error "Metis_Term.destBinop: wrong binop" | destBinop _ _ = raise Error "Metis_Term.destBinop: not a binop"; fun isBinop f = can (destBinop f); (* ------------------------------------------------------------------------- *) (* The size of a term in symbols. *) (* ------------------------------------------------------------------------- *) val VAR_SYMBOLS = 1; val FN_SYMBOLS = 1; local fun sz n [] = n | sz n (Var _ :: tms) = sz (n + VAR_SYMBOLS) tms | sz n (Fn (func,args) :: tms) = sz (n + FN_SYMBOLS) (args @ tms); in fun symbols tm = sz 0 [tm]; end; (* ------------------------------------------------------------------------- *) (* A total comparison function for terms. *) (* ------------------------------------------------------------------------- *) local fun cmp [] [] = EQUAL | cmp (tm1 :: tms1) (tm2 :: tms2) = let val tm1_tm2 = (tm1,tm2) in if Metis_Portable.pointerEqual tm1_tm2 then cmp tms1 tms2 else case tm1_tm2 of (Var v1, Var v2) => (case Metis_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 Metis_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 "Metis_Term.compare"; in fun compare (tm1,tm2) = cmp [tm1] [tm2]; end; fun equal tm1 tm2 = compare (tm1,tm2) = EQUAL; (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) type path = int list; fun subterm tm [] = tm | subterm (Var _) (_ :: _) = raise Error "Metis_Term.subterm: Var" | subterm (Fn (_,tms)) (h :: t) = if h >= length tms then raise Error "Metis_Term.replace: Fn" else subterm (List.nth (tms,h)) t; local fun subtms [] acc = acc | subtms ((path,tm) :: rest) acc = let fun f (n,arg) = (n :: path, arg) val acc = (List.rev path, tm) :: acc in case tm of Var _ => subtms rest acc | Fn (_,args) => subtms (List.map f (enumerate args) @ rest) acc end; in fun subterms tm = subtms [([],tm)] []; end; fun replace tm ([],res) = if equal res tm then tm else res | replace tm (h :: t, res) = case tm of Var _ => raise Error "Metis_Term.replace: Var" | Fn (func,tms) => if h >= length tms then raise Error "Metis_Term.replace: Fn" else let val arg = List.nth (tms,h) val arg' = replace arg (t,res) in if Metis_Portable.pointerEqual (arg',arg) then tm else Fn (func, updateNth (h,arg') tms) end; fun find pred = let fun search [] = NONE | search ((path,tm) :: rest) = if pred tm then SOME (List.rev path) else case tm of Var _ => search rest | Fn (_,a) => let val subtms = List.map (fn (i,t) => (i :: path, t)) (enumerate a) in search (subtms @ rest) end in fn tm => search [([],tm)] end; val ppPath = Metis_Print.ppList Metis_Print.ppInt; val pathToString = Metis_Print.toString ppPath; (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) local fun free _ [] = false | free v (Var w :: tms) = Metis_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]; end; local fun free vs [] = vs | free vs (Var v :: tms) = free (Metis_NameSet.add vs v) tms | free vs (Fn (_,args) :: tms) = free vs (args @ tms); in val freeVarsList = free Metis_NameSet.empty; fun freeVars tm = freeVarsList [tm]; end; (* ------------------------------------------------------------------------- *) (* Fresh variables. *) (* ------------------------------------------------------------------------- *) fun newVar () = Var (Metis_Name.newName ()); fun newVars n = List.map Var (Metis_Name.newNames n); local fun avoid av n = Metis_NameSet.member n av; in fun variantPrime av = Metis_Name.variantPrime {avoid = avoid av}; fun variantNum av = Metis_Name.variantNum {avoid = avoid av}; end; (* ------------------------------------------------------------------------- *) (* Special support for terms with type annotations. *) (* ------------------------------------------------------------------------- *) val hasTypeFunctionName = Metis_Name.fromString ":"; val hasTypeFunction = (hasTypeFunctionName,2); fun destFnHasType ((f,a) : functionName * term list) = if not (Metis_Name.equal f hasTypeFunctionName) then raise Error "Metis_Term.destFnHasType" else case a of [tm,ty] => (tm,ty) | _ => raise Error "Metis_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 (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 ((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 = (List.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 = (List.rev path, tm) :: acc val rest = List.map f (enumerate args) @ rest in subtms rest acc end; in fun nonVarTypedSubterms tm = subtms [([],tm)] []; end; (* ------------------------------------------------------------------------- *) (* Special support for terms with an explicit function application operator. *) (* ------------------------------------------------------------------------- *) val appName = Metis_Name.fromString "."; fun mkFnApp (fTm,aTm) = (appName, [fTm,aTm]); fun mkApp f_a = Fn (mkFnApp f_a); fun destFnApp ((f,a) : Metis_Name.name * term list) = if not (Metis_Name.equal f appName) then raise Error "Metis_Term.destFnApp" else case a of [fTm,aTm] => (fTm,aTm) | _ => raise Error "Metis_Term.destFnApp"; val isFnApp = can destFnApp; fun destApp tm = case tm of Var _ => raise Error "Metis_Term.destApp" | Fn func => destFnApp func; val isApp = can destApp; fun listMkApp (f,l) = List.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; (* ------------------------------------------------------------------------- *) (* Parsing and pretty printing. *) (* ------------------------------------------------------------------------- *) (* Operators parsed and printed infix *) val infixes = (Unsynchronized.ref o Metis_Print.Infixes) [(* ML symbols *) {token = "/", precedence = 7, assoc = Metis_Print.LeftAssoc}, {token = "div", precedence = 7, assoc = Metis_Print.LeftAssoc}, {token = "mod", precedence = 7, assoc = Metis_Print.LeftAssoc}, {token = "*", precedence = 7, assoc = Metis_Print.LeftAssoc}, {token = "+", precedence = 6, assoc = Metis_Print.LeftAssoc}, {token = "-", precedence = 6, assoc = Metis_Print.LeftAssoc}, {token = "^", precedence = 6, assoc = Metis_Print.LeftAssoc}, {token = "@", precedence = 5, assoc = Metis_Print.RightAssoc}, {token = "::", precedence = 5, assoc = Metis_Print.RightAssoc}, {token = "=", precedence = 4, assoc = Metis_Print.NonAssoc}, {token = "<>", precedence = 4, assoc = Metis_Print.NonAssoc}, {token = "<=", precedence = 4, assoc = Metis_Print.NonAssoc}, {token = "<", precedence = 4, assoc = Metis_Print.NonAssoc}, {token = ">=", precedence = 4, assoc = Metis_Print.NonAssoc}, {token = ">", precedence = 4, assoc = Metis_Print.NonAssoc}, {token = "o", precedence = 3, assoc = Metis_Print.LeftAssoc}, {token = "->", precedence = 2, assoc = Metis_Print.RightAssoc}, {token = ":", precedence = 1, assoc = Metis_Print.NonAssoc}, {token = ",", precedence = 0, assoc = Metis_Print.RightAssoc}, (* Logical connectives *) {token = "/\\", precedence = ~1, assoc = Metis_Print.RightAssoc}, {token = "\\/", precedence = ~2, assoc = Metis_Print.RightAssoc}, {token = "==>", precedence = ~3, assoc = Metis_Print.RightAssoc}, {token = "<=>", precedence = ~4, assoc = Metis_Print.RightAssoc}, (* Other symbols *) {token = ".", precedence = 9, assoc = Metis_Print.LeftAssoc}, {token = "**", precedence = 8, assoc = Metis_Print.LeftAssoc}, {token = "++", precedence = 6, assoc = Metis_Print.LeftAssoc}, {token = "--", precedence = 6, assoc = Metis_Print.LeftAssoc}, {token = "==", precedence = 4, assoc = Metis_Print.NonAssoc}]; (* The negation symbol *) val negation : string Unsynchronized.ref = Unsynchronized.ref "~"; (* Binder symbols *) val binders : string list Unsynchronized.ref = Unsynchronized.ref ["\\","!","?","?!"]; (* Bracket symbols *) val brackets : (string * string) list Unsynchronized.ref = Unsynchronized.ref [("[","]"),("{","}")]; (* Pretty printing *) fun pp inputTerm = let val quants = !binders and iOps = !infixes and neg = !negation and bracks = !brackets val bMap = let fun f (b1,b2) = (b1 ^ b2, b1, b2) in List.map f bracks end val bTokens = op@ (unzip bracks) val iTokens = Metis_Print.tokensInfixes iOps fun destI tm = case tm of Fn (f,[a,b]) => let val f = Metis_Name.toString f in if Metis_StringSet.member f iTokens then SOME (f,a,b) else NONE end | _ => NONE fun isI tm = Option.isSome (destI tm) fun iToken (_,tok) = Metis_Print.program [(if tok = "," then Metis_Print.skip else Metis_Print.ppString " "), Metis_Print.ppString tok, Metis_Print.break]; val iPrinter = Metis_Print.ppInfixes iOps destI iToken val specialTokens = Metis_StringSet.addList iTokens (neg :: quants @ ["$","(",")"] @ bTokens) fun vName bv s = Metis_StringSet.member s bv fun checkVarName bv n = let val s = Metis_Name.toString n in if vName bv s then s else "$" ^ s end fun varName bv = Metis_Print.ppMap (checkVarName bv) Metis_Print.ppString fun checkFunctionName bv n = let val s = Metis_Name.toString n in if Metis_StringSet.member s specialTokens orelse vName bv s then "(" ^ s ^ ")" else s end fun functionName bv = Metis_Print.ppMap (checkFunctionName bv) Metis_Print.ppString fun stripNeg tm = case tm of Fn (f,[a]) => if Metis_Name.toString f <> neg then (0,tm) else let val (n,tm) = stripNeg a in (n + 1, tm) end | _ => (0,tm) val destQuant = let fun dest q (Fn (q', [Var v, body])) = if Metis_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 => Metis_Useful.first (fn q => dest q tm) quants end fun isQuant tm = Option.isSome (destQuant tm) fun destBrack (Fn (b,[tm])) = let val s = Metis_Name.toString b in case List.find (fn (n,_,_) => n = s) bMap of NONE => NONE | SOME (_,b1,b2) => SOME (b1,tm,b2) end | destBrack _ = NONE fun isBrack tm = Option.isSome (destBrack tm) fun functionArgument bv tm = Metis_Print.sequence Metis_Print.break (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)) = Metis_Print.inconsistentBlock 2 (functionName bv f :: List.map (functionArgument bv) args) and customBracket bv tm = case destBrack tm of SOME (b1,tm,b2) => Metis_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 = Metis_StringSet.addList bv (List.map Metis_Name.toString (v :: vs)) in Metis_Print.program [Metis_Print.ppString q, varName bv v, Metis_Print.program (List.map (Metis_Print.sequence Metis_Print.break o varName bv) vs), Metis_Print.ppString ".", Metis_Print.break, innerQuant bv tm] end and quantifier bv tm = if not (isQuant tm) then customBracket bv tm else Metis_Print.inconsistentBlock 2 [innerQuant bv tm] and molecule bv (tm,r) = let val (n,tm) = stripNeg tm in Metis_Print.inconsistentBlock n [Metis_Print.duplicate n (Metis_Print.ppString 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 = Metis_Print.ppBracket "(" ")" (term bv) tm in term Metis_StringSet.empty end inputTerm; val toString = Metis_Print.toString pp; (* Parsing *) local open Metis_Parse; infixr 9 >>++ infixr 8 ++ infixr 7 >> infixr 6 || val isAlphaNum = let val alphaNumChars = String.explode "_'" in fn c => mem c alphaNumChars orelse Char.isAlphaNum c end; local val alphaNumToken = atLeastOne (some isAlphaNum) >> String.implode; val symbolToken = let fun isNeg c = str c = !negation val symbolChars = String.explode "<>=-*+/\\?@|!$%&#^:;~" fun isSymbol c = mem c symbolChars fun isNonNegSymbol c = not (isNeg c) andalso isSymbol c in some isNeg >> str || (some isNonNegSymbol ++ many (some isSymbol)) >> (String.implode o op::) end; val punctToken = let val punctChars = String.explode "()[]{}.," fun isPunct c = mem c punctChars in some isPunct >> str end; val lexToken = alphaNumToken || symbolToken || punctToken; val space = many (some Char.isSpace); in val lexer = (space ++ lexToken ++ space) >> (fn (_,(tok,_)) => tok); end; fun termParser inputStream = let val quants = !binders and iOps = !infixes and neg = !negation and bracks = ("(",")") :: !brackets val bracks = List.map (fn (b1,b2) => (b1 ^ b2, b1, b2)) bracks val bTokens = List.map #2 bracks @ List.map #3 bracks fun possibleVarName "" = false | possibleVarName s = isAlphaNum (String.sub (s,0)) fun vName bv s = Metis_StringSet.member s bv val iTokens = Metis_Print.tokensInfixes iOps fun iMk (f,a,b) = Fn (Metis_Name.fromString f, [a,b]) val iParser = parseInfixes iOps iMk any val specialTokens = Metis_StringSet.addList iTokens (neg :: quants @ ["$"] @ bTokens) fun varName bv = some (vName bv) || (some (Metis_Useful.equal "$") ++ some possibleVarName) >> snd fun fName bv s = not (Metis_StringSet.member s specialTokens) andalso not (vName bv s) fun functionName bv = some (fName bv) || (some (Metis_Useful.equal "(") ++ any ++ some (Metis_Useful.equal ")")) >> (fn (_,(s,_)) => s) fun basic bv tokens = let val var = varName bv >> (Var o Metis_Name.fromString) val const = functionName bv >> (fn f => Fn (Metis_Name.fromString f, [])) fun bracket (ab,a,b) = (some (Metis_Useful.equal a) ++ term bv ++ some (Metis_Useful.equal b)) >> (fn (_,(tm,_)) => if ab = "()" then tm else Fn (Metis_Name.fromString ab, [tm])) fun quantifier q = let fun bind (v,t) = Fn (Metis_Name.fromString q, [Var (Metis_Name.fromString v), t]) in (some (Metis_Useful.equal q) ++ atLeastOne (some possibleVarName) ++ some (Metis_Useful.equal ".")) >>++ (fn (_,(vs,_)) => term (Metis_StringSet.addList bv vs) >> (fn body => List.foldr bind body vs)) end in var || const || first (List.map bracket bracks) || first (List.map quantifier quants) end tokens and molecule bv tokens = let val negations = many (some (Metis_Useful.equal neg)) >> length val function = (functionName bv ++ many (basic bv)) >> (fn (f,args) => Fn (Metis_Name.fromString f, args)) || basic bv in (negations ++ function) >> (fn (n,tm) => funpow n (fn t => Fn (Metis_Name.fromString neg, [t])) tm) end tokens and term bv tokens = iParser (molecule bv) tokens in term Metis_StringSet.empty end inputStream; in fun fromString input = let val chars = Metis_Stream.fromList (String.explode input) val tokens = everything (lexer >> singleton) chars val terms = everything (termParser >> singleton) tokens in case Metis_Stream.toList terms of [tm] => tm | _ => raise Error "Metis_Term.fromString" end; end; local val antiquotedTermToString = Metis_Print.toString (Metis_Print.ppBracket "(" ")" pp); in val parse = Metis_Parse.parseQuotation antiquotedTermToString fromString; end; end structure Metis_TermOrdered = struct type t = Metis_Term.term val compare = Metis_Term.compare end structure Metis_TermMap = Metis_KeyMap (Metis_TermOrdered); structure Metis_TermSet = Metis_ElementSet (Metis_TermMap); (**** Original file: src/Subst.sig ****) (* ========================================================================= *) (* FIRST ORDER LOGIC SUBSTITUTIONS *) (* Copyright (c) 2002 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Subst = sig (* ------------------------------------------------------------------------- *) (* A type of first order logic substitutions. *) (* ------------------------------------------------------------------------- *) type subst (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val empty : subst val null : subst -> bool val size : subst -> int val peek : subst -> Metis_Term.var -> Metis_Term.term option val insert : subst -> Metis_Term.var * Metis_Term.term -> subst val singleton : Metis_Term.var * Metis_Term.term -> subst val toList : subst -> (Metis_Term.var * Metis_Term.term) list val fromList : (Metis_Term.var * Metis_Term.term) list -> subst val foldl : (Metis_Term.var * Metis_Term.term * 'a -> 'a) -> 'a -> subst -> 'a val foldr : (Metis_Term.var * Metis_Term.term * 'a -> 'a) -> 'a -> subst -> 'a val pp : subst Metis_Print.pp val toString : subst -> string (* ------------------------------------------------------------------------- *) (* Normalizing removes identity substitutions. *) (* ------------------------------------------------------------------------- *) val normalize : subst -> subst (* ------------------------------------------------------------------------- *) (* Applying a substitution to a first order logic term. *) (* ------------------------------------------------------------------------- *) val subst : subst -> Metis_Term.term -> Metis_Term.term (* ------------------------------------------------------------------------- *) (* Restricting a substitution to a smaller set of variables. *) (* ------------------------------------------------------------------------- *) val restrict : subst -> Metis_NameSet.set -> subst val remove : subst -> Metis_NameSet.set -> subst (* ------------------------------------------------------------------------- *) (* Composing two substitutions so that the following identity holds: *) (* *) (* subst (compose sub1 sub2) tm = subst sub2 (subst sub1 tm) *) (* ------------------------------------------------------------------------- *) val compose : subst -> subst -> subst (* ------------------------------------------------------------------------- *) (* 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 *) val isRenaming : subst -> bool (* ------------------------------------------------------------------------- *) (* Creating a substitution to freshen variables. *) (* ------------------------------------------------------------------------- *) 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. *) (* ------------------------------------------------------------------------- *) val match : subst -> Metis_Term.term -> Metis_Term.term -> subst (* raises Error *) (* ------------------------------------------------------------------------- *) (* Unification for first order logic terms. *) (* ------------------------------------------------------------------------- *) val unify : subst -> Metis_Term.term -> Metis_Term.term -> subst (* raises Error *) end (**** Original file: src/Subst.sml ****) (* ========================================================================= *) (* FIRST ORDER LOGIC SUBSTITUTIONS *) (* Copyright (c) 2002 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Subst :> Metis_Subst = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* A type of first order logic substitutions. *) (* ------------------------------------------------------------------------- *) datatype subst = Metis_Subst of Metis_Term.term Metis_NameMap.map; (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val empty = Metis_Subst (Metis_NameMap.new ()); fun null (Metis_Subst m) = Metis_NameMap.null m; fun size (Metis_Subst m) = Metis_NameMap.size m; fun peek (Metis_Subst m) v = Metis_NameMap.peek m v; fun insert (Metis_Subst m) v_tm = Metis_Subst (Metis_NameMap.insert m v_tm); fun singleton v_tm = insert empty v_tm; fun toList (Metis_Subst m) = Metis_NameMap.toList m; fun fromList l = Metis_Subst (Metis_NameMap.fromList l); fun foldl f b (Metis_Subst m) = Metis_NameMap.foldl f b m; fun foldr f b (Metis_Subst m) = Metis_NameMap.foldr f b m; fun pp sub = Metis_Print.ppBracket "<[" "]>" (Metis_Print.ppOpList "," (Metis_Print.ppOp2 " |->" Metis_Name.pp Metis_Term.pp)) (toList sub); val toString = Metis_Print.toString pp; (* ------------------------------------------------------------------------- *) (* Normalizing removes identity substitutions. *) (* ------------------------------------------------------------------------- *) local fun isNotId (v,tm) = not (Metis_Term.equalVar v tm); in fun normalize (sub as Metis_Subst m) = let val m' = Metis_NameMap.filter isNotId m in if Metis_NameMap.size m = Metis_NameMap.size m' then sub else Metis_Subst m' end; end; (* ------------------------------------------------------------------------- *) (* Applying a substitution to a first order logic term. *) (* ------------------------------------------------------------------------- *) fun subst sub = let fun tmSub (tm as Metis_Term.Var v) = (case peek sub v of SOME tm' => if Metis_Portable.pointerEqual (tm,tm') then tm else tm' | NONE => tm) | tmSub (tm as Metis_Term.Fn (f,args)) = let val args' = Metis_Sharing.map tmSub args in if Metis_Portable.pointerEqual (args,args') then tm else Metis_Term.Fn (f,args') end in fn tm => if null sub then tm else tmSub tm end; (* ------------------------------------------------------------------------- *) (* Restricting a substitution to a given set of variables. *) (* ------------------------------------------------------------------------- *) fun restrict (sub as Metis_Subst m) varSet = let fun isRestrictedVar (v,_) = Metis_NameSet.member v varSet val m' = Metis_NameMap.filter isRestrictedVar m in if Metis_NameMap.size m = Metis_NameMap.size m' then sub else Metis_Subst m' end; fun remove (sub as Metis_Subst m) varSet = let fun isRestrictedVar (v,_) = not (Metis_NameSet.member v varSet) val m' = Metis_NameMap.filter isRestrictedVar m in if Metis_NameMap.size m = Metis_NameMap.size m' then sub else Metis_Subst m' end; (* ------------------------------------------------------------------------- *) (* Composing two substitutions so that the following identity holds: *) (* *) (* subst (compose sub1 sub2) tm = subst sub2 (subst sub1 tm) *) (* ------------------------------------------------------------------------- *) fun compose (sub1 as Metis_Subst m1) sub2 = let fun f (v,tm,s) = insert s (v, subst sub2 tm) in if null sub2 then sub1 else Metis_NameMap.foldl f sub2 m1 end; (* ------------------------------------------------------------------------- *) (* Creating the union of two compatible substitutions. *) (* ------------------------------------------------------------------------- *) local fun compatible ((_,tm1),(_,tm2)) = if Metis_Term.equal tm1 tm2 then SOME tm1 else raise Error "Metis_Subst.union: incompatible"; in fun union (s1 as Metis_Subst m1) (s2 as Metis_Subst m2) = if Metis_NameMap.null m1 then s2 else if Metis_NameMap.null m2 then s1 else Metis_Subst (Metis_NameMap.union compatible m1 m2); end; (* ------------------------------------------------------------------------- *) (* Substitutions can be inverted iff they are renaming substitutions. *) (* ------------------------------------------------------------------------- *) local fun inv (v, Metis_Term.Var w, s) = if Metis_NameMap.inDomain w s then raise Error "Metis_Subst.invert: non-injective" else Metis_NameMap.insert s (w, Metis_Term.Var v) | inv (_, Metis_Term.Fn _, _) = raise Error "Metis_Subst.invert: non-variable"; in fun invert (Metis_Subst m) = Metis_Subst (Metis_NameMap.foldl inv (Metis_NameMap.new ()) m); end; val isRenaming = can invert; (* ------------------------------------------------------------------------- *) (* Creating a substitution to freshen variables. *) (* ------------------------------------------------------------------------- *) val freshVars = let fun add (v,m) = insert m (v, Metis_Term.newVar ()) in Metis_NameSet.foldl add empty end; (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) val redexes = let fun add (v,_,s) = Metis_NameSet.add s v in foldl add Metis_NameSet.empty end; val residueFreeVars = let fun add (_,t,s) = Metis_NameSet.union s (Metis_Term.freeVars t) in foldl add Metis_NameSet.empty end; val freeVars = let fun add (v,t,s) = Metis_NameSet.union (Metis_NameSet.add s v) (Metis_Term.freeVars t) in foldl add Metis_NameSet.empty end; (* ------------------------------------------------------------------------- *) (* Functions. *) (* ------------------------------------------------------------------------- *) val functions = let fun add (_,t,s) = Metis_NameAritySet.union s (Metis_Term.functions t) in foldl add Metis_NameAritySet.empty end; (* ------------------------------------------------------------------------- *) (* Matching for first order logic terms. *) (* ------------------------------------------------------------------------- *) local fun matchList sub [] = sub | matchList sub ((Metis_Term.Var v, tm) :: rest) = let val sub = case peek sub v of NONE => insert sub (v,tm) | SOME tm' => if Metis_Term.equal tm tm' then sub else raise Error "Metis_Subst.match: incompatible matches" in matchList sub rest end | matchList sub ((Metis_Term.Fn (f1,args1), Metis_Term.Fn (f2,args2)) :: rest) = if Metis_Name.equal f1 f2 andalso length args1 = length args2 then matchList sub (zip args1 args2 @ rest) else raise Error "Metis_Subst.match: different structure" | matchList _ _ = raise Error "Metis_Subst.match: functions can't match vars"; in fun match sub tm1 tm2 = matchList sub [(tm1,tm2)]; end; (* ------------------------------------------------------------------------- *) (* Unification for first order logic terms. *) (* ------------------------------------------------------------------------- *) local fun solve sub [] = sub | solve sub ((tm1_tm2 as (tm1,tm2)) :: rest) = if Metis_Portable.pointerEqual tm1_tm2 then solve sub rest else solve' sub (subst sub tm1) (subst sub tm2) rest and solve' sub (Metis_Term.Var v) tm rest = if Metis_Term.equalVar v tm then solve sub rest else if Metis_Term.freeIn v tm then raise Error "Metis_Subst.unify: occurs check" else (case peek sub v of NONE => solve (compose sub (singleton (v,tm))) rest | SOME tm' => solve' sub tm' tm rest) | solve' sub tm1 (tm2 as Metis_Term.Var _) rest = solve' sub tm2 tm1 rest | solve' sub (Metis_Term.Fn (f1,args1)) (Metis_Term.Fn (f2,args2)) rest = if Metis_Name.equal f1 f2 andalso length args1 = length args2 then solve sub (zip args1 args2 @ rest) else raise Error "Metis_Subst.unify: different structure"; in fun unify sub tm1 tm2 = solve sub [(tm1,tm2)]; end; end (**** Original file: src/Atom.sig ****) (* ========================================================================= *) (* FIRST ORDER LOGIC ATOMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Atom = sig (* ------------------------------------------------------------------------- *) (* A type for storing first order logic atoms. *) (* ------------------------------------------------------------------------- *) type relationName = Metis_Name.name type relation = relationName * int type atom = relationName * Metis_Term.term list (* ------------------------------------------------------------------------- *) (* Constructors and destructors. *) (* ------------------------------------------------------------------------- *) val name : atom -> relationName val arguments : atom -> Metis_Term.term list val arity : atom -> int val relation : atom -> relation val functions : atom -> Metis_NameAritySet.set val functionNames : atom -> Metis_NameSet.set (* Binary relations *) val mkBinop : relationName -> Metis_Term.term * Metis_Term.term -> atom val destBinop : relationName -> atom -> Metis_Term.term * Metis_Term.term val isBinop : relationName -> atom -> bool (* ------------------------------------------------------------------------- *) (* The size of an atom in symbols. *) (* ------------------------------------------------------------------------- *) val symbols : atom -> int (* ------------------------------------------------------------------------- *) (* A total comparison function for atoms. *) (* ------------------------------------------------------------------------- *) val compare : atom * atom -> order val equal : atom -> atom -> bool (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) val subterm : atom -> Metis_Term.path -> Metis_Term.term val subterms : atom -> (Metis_Term.path * Metis_Term.term) list val replace : atom -> Metis_Term.path * Metis_Term.term -> atom val find : (Metis_Term.term -> bool) -> atom -> Metis_Term.path option (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) val freeIn : Metis_Term.var -> atom -> bool val freeVars : atom -> Metis_NameSet.set (* ------------------------------------------------------------------------- *) (* Substitutions. *) (* ------------------------------------------------------------------------- *) val subst : Metis_Subst.subst -> atom -> atom (* ------------------------------------------------------------------------- *) (* Matching. *) (* ------------------------------------------------------------------------- *) val match : Metis_Subst.subst -> atom -> atom -> Metis_Subst.subst (* raises Error *) (* ------------------------------------------------------------------------- *) (* Unification. *) (* ------------------------------------------------------------------------- *) val unify : Metis_Subst.subst -> atom -> atom -> Metis_Subst.subst (* raises Error *) (* ------------------------------------------------------------------------- *) (* The equality relation. *) (* ------------------------------------------------------------------------- *) val eqRelationName : relationName val eqRelation : relation val mkEq : Metis_Term.term * Metis_Term.term -> atom val destEq : atom -> Metis_Term.term * Metis_Term.term val isEq : atom -> bool val mkRefl : Metis_Term.term -> atom val destRefl : atom -> Metis_Term.term val isRefl : atom -> bool val sym : atom -> atom (* raises Error if given a refl *) val lhs : atom -> Metis_Term.term val rhs : atom -> Metis_Term.term (* ------------------------------------------------------------------------- *) (* Special support for terms with type annotations. *) (* ------------------------------------------------------------------------- *) val typedSymbols : atom -> int val nonVarTypedSubterms : atom -> (Metis_Term.path * Metis_Term.term) list (* ------------------------------------------------------------------------- *) (* Parsing and pretty printing. *) (* ------------------------------------------------------------------------- *) val pp : atom Metis_Print.pp val toString : atom -> string val fromString : string -> atom val parse : Metis_Term.term Metis_Parse.quotation -> atom end (**** Original file: src/Atom.sml ****) (* ========================================================================= *) (* FIRST ORDER LOGIC ATOMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Atom :> Metis_Atom = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* A type for storing first order logic atoms. *) (* ------------------------------------------------------------------------- *) type relationName = Metis_Name.name; type relation = relationName * int; type atom = relationName * Metis_Term.term list; (* ------------------------------------------------------------------------- *) (* Constructors and destructors. *) (* ------------------------------------------------------------------------- *) fun name ((rel,_) : atom) = rel; fun arguments ((_,args) : atom) = args; fun arity atm = length (arguments atm); fun relation atm = (name atm, arity atm); val functions = let fun f (tm,acc) = Metis_NameAritySet.union (Metis_Term.functions tm) acc in fn atm => List.foldl f Metis_NameAritySet.empty (arguments atm) end; val functionNames = let fun f (tm,acc) = Metis_NameSet.union (Metis_Term.functionNames tm) acc in fn atm => List.foldl f Metis_NameSet.empty (arguments atm) end; (* Binary relations *) fun mkBinop p (a,b) : atom = (p,[a,b]); fun destBinop p (x,[a,b]) = if Metis_Name.equal x p then (a,b) else raise Error "Metis_Atom.destBinop: wrong binop" | destBinop _ _ = raise Error "Metis_Atom.destBinop: not a binop"; fun isBinop p = can (destBinop p); (* ------------------------------------------------------------------------- *) (* The size of an atom in symbols. *) (* ------------------------------------------------------------------------- *) fun symbols atm = List.foldl (fn (tm,z) => Metis_Term.symbols tm + z) 1 (arguments atm); (* ------------------------------------------------------------------------- *) (* A total comparison function for atoms. *) (* ------------------------------------------------------------------------- *) fun compare ((p1,tms1),(p2,tms2)) = case Metis_Name.compare (p1,p2) of LESS => LESS | EQUAL => lexCompare Metis_Term.compare (tms1,tms2) | GREATER => GREATER; fun equal atm1 atm2 = compare (atm1,atm2) = EQUAL; (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) fun subterm _ [] = raise Bug "Metis_Atom.subterm: empty path" | subterm ((_,tms) : atom) (h :: t) = if h >= length tms then raise Error "Metis_Atom.subterm: bad path" else Metis_Term.subterm (List.nth (tms,h)) t; fun subterms ((_,tms) : atom) = let fun f ((n,tm),l) = List.map (fn (p,s) => (n :: p, s)) (Metis_Term.subterms tm) @ l in List.foldl f [] (enumerate tms) end; fun replace _ ([],_) = raise Bug "Metis_Atom.replace: empty path" | replace (atm as (rel,tms)) (h :: t, res) : atom = if h >= length tms then raise Error "Metis_Atom.replace: bad path" else let val tm = List.nth (tms,h) val tm' = Metis_Term.replace tm (t,res) in if Metis_Portable.pointerEqual (tm,tm') then atm else (rel, updateNth (h,tm') tms) end; fun find pred = let fun f (i,tm) = case Metis_Term.find pred tm of SOME path => SOME (i :: path) | NONE => NONE in fn (_,tms) : atom => first f (enumerate tms) end; (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) fun freeIn v atm = List.exists (Metis_Term.freeIn v) (arguments atm); val freeVars = let fun f (tm,acc) = Metis_NameSet.union (Metis_Term.freeVars tm) acc in fn atm => List.foldl f Metis_NameSet.empty (arguments atm) end; (* ------------------------------------------------------------------------- *) (* Substitutions. *) (* ------------------------------------------------------------------------- *) fun subst sub (atm as (p,tms)) : atom = let val tms' = Metis_Sharing.map (Metis_Subst.subst sub) tms in if Metis_Portable.pointerEqual (tms',tms) then atm else (p,tms') end; (* ------------------------------------------------------------------------- *) (* Matching. *) (* ------------------------------------------------------------------------- *) local fun matchArg ((tm1,tm2),sub) = Metis_Subst.match sub tm1 tm2; in fun match sub (p1,tms1) (p2,tms2) = let val _ = (Metis_Name.equal p1 p2 andalso length tms1 = length tms2) orelse raise Error "Metis_Atom.match" in List.foldl matchArg sub (zip tms1 tms2) end; end; (* ------------------------------------------------------------------------- *) (* Unification. *) (* ------------------------------------------------------------------------- *) local fun unifyArg ((tm1,tm2),sub) = Metis_Subst.unify sub tm1 tm2; in fun unify sub (p1,tms1) (p2,tms2) = let val _ = (Metis_Name.equal p1 p2 andalso length tms1 = length tms2) orelse raise Error "Metis_Atom.unify" in List.foldl unifyArg sub (zip tms1 tms2) end; end; (* ------------------------------------------------------------------------- *) (* The equality relation. *) (* ------------------------------------------------------------------------- *) val eqRelationName = Metis_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 _ = Metis_Term.equal l r orelse raise Error "Metis_Atom.destRefl" in l end; fun isRefl x = can destRefl x; fun sym atm = let val (l,r) = destEq atm val _ = not (Metis_Term.equal l r) orelse raise Error "Metis_Atom.sym: refl" in mkEq (r,l) end; fun lhs atm = fst (destEq atm); fun rhs atm = snd (destEq atm); (* ------------------------------------------------------------------------- *) (* Special support for terms with type annotations. *) (* ------------------------------------------------------------------------- *) fun typedSymbols ((_,tms) : atom) = List.foldl (fn (tm,z) => Metis_Term.typedSymbols tm + z) 1 tms; fun nonVarTypedSubterms (_,tms) = let fun addArg ((n,arg),acc) = let fun addTm ((path,tm),acc) = (n :: path, tm) :: acc in List.foldl addTm acc (Metis_Term.nonVarTypedSubterms arg) end in List.foldl addArg [] (enumerate tms) end; (* ------------------------------------------------------------------------- *) (* Parsing and pretty printing. *) (* ------------------------------------------------------------------------- *) val pp = Metis_Print.ppMap Metis_Term.Fn Metis_Term.pp; val toString = Metis_Print.toString pp; fun fromString s = Metis_Term.destFn (Metis_Term.fromString s); val parse = Metis_Parse.parseQuotation Metis_Term.toString fromString; end structure Metis_AtomOrdered = struct type t = Metis_Atom.atom val compare = Metis_Atom.compare end structure Metis_AtomMap = Metis_KeyMap (Metis_AtomOrdered); structure Metis_AtomSet = Metis_ElementSet (Metis_AtomMap); (**** Original file: src/Formula.sig ****) (* ========================================================================= *) (* FIRST ORDER LOGIC FORMULAS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Formula = sig (* ------------------------------------------------------------------------- *) (* A type of first order logic formulas. *) (* ------------------------------------------------------------------------- *) datatype formula = True | False | Metis_Atom of Metis_Atom.atom | Not of formula | And of formula * formula | Or of formula * formula | Imp of formula * formula | Iff of formula * formula | Forall of Metis_Term.var * formula | Exists of Metis_Term.var * formula (* ------------------------------------------------------------------------- *) (* Constructors and destructors. *) (* ------------------------------------------------------------------------- *) (* Booleans *) val mkBoolean : bool -> formula val destBoolean : formula -> bool val isBoolean : formula -> bool val isTrue : formula -> bool val isFalse : formula -> bool (* Functions *) val functions : formula -> Metis_NameAritySet.set val functionNames : formula -> Metis_NameSet.set (* Relations *) val relations : formula -> Metis_NameAritySet.set val relationNames : formula -> Metis_NameSet.set (* Atoms *) val destAtom : formula -> Metis_Atom.atom val isAtom : formula -> bool (* Negations *) val destNeg : formula -> formula val isNeg : formula -> bool val stripNeg : formula -> int * formula (* Conjunctions *) val listMkConj : formula list -> formula val stripConj : formula -> formula list val flattenConj : formula -> formula list (* Disjunctions *) val listMkDisj : formula list -> formula val stripDisj : formula -> formula list val flattenDisj : formula -> formula list (* Equivalences *) val listMkEquiv : formula list -> formula val stripEquiv : formula -> formula list val flattenEquiv : formula -> formula list (* Universal quantification *) val destForall : formula -> Metis_Term.var * formula val isForall : formula -> bool 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 *) val destExists : formula -> Metis_Term.var * formula val isExists : formula -> bool val listMkExists : Metis_Term.var list * formula -> formula val setMkExists : Metis_NameSet.set * formula -> formula val stripExists : formula -> Metis_Term.var list * formula (* ------------------------------------------------------------------------- *) (* The size of a formula in symbols. *) (* ------------------------------------------------------------------------- *) val symbols : formula -> int (* ------------------------------------------------------------------------- *) (* A total comparison function for formulas. *) (* ------------------------------------------------------------------------- *) val compare : formula * formula -> order val equal : formula -> formula -> bool (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) val freeIn : Metis_Term.var -> formula -> bool val freeVars : formula -> Metis_NameSet.set val freeVarsList : formula list -> Metis_NameSet.set val specialize : formula -> formula val generalize : formula -> formula (* ------------------------------------------------------------------------- *) (* Substitutions. *) (* ------------------------------------------------------------------------- *) val subst : Metis_Subst.subst -> formula -> formula (* ------------------------------------------------------------------------- *) (* The equality relation. *) (* ------------------------------------------------------------------------- *) val mkEq : Metis_Term.term * Metis_Term.term -> formula val destEq : formula -> Metis_Term.term * Metis_Term.term val isEq : formula -> bool val mkNeq : Metis_Term.term * Metis_Term.term -> formula val destNeq : formula -> Metis_Term.term * Metis_Term.term val isNeq : formula -> bool val mkRefl : Metis_Term.term -> formula val destRefl : formula -> Metis_Term.term val isRefl : formula -> bool val sym : formula -> formula (* raises Error if given a refl *) val lhs : formula -> Metis_Term.term val rhs : formula -> Metis_Term.term (* ------------------------------------------------------------------------- *) (* Splitting goals. *) (* ------------------------------------------------------------------------- *) val splitGoal : formula -> formula list (* ------------------------------------------------------------------------- *) (* Parsing and pretty-printing. *) (* ------------------------------------------------------------------------- *) type quotation = formula Metis_Parse.quotation val pp : formula Metis_Print.pp val toString : formula -> string val fromString : string -> formula val parse : quotation -> formula end (**** Original file: src/Formula.sml ****) (* ========================================================================= *) (* FIRST ORDER LOGIC FORMULAS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Formula :> Metis_Formula = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* A type of first order logic formulas. *) (* ------------------------------------------------------------------------- *) datatype formula = True | False | Metis_Atom of Metis_Atom.atom | Not of formula | And of formula * formula | Or of formula * formula | Imp of formula * formula | Iff of formula * formula | Forall of Metis_Term.var * formula | Exists of Metis_Term.var * formula; (* ------------------------------------------------------------------------- *) (* Constructors and destructors. *) (* ------------------------------------------------------------------------- *) (* Booleans *) fun mkBoolean true = True | mkBoolean false = False; fun destBoolean True = true | destBoolean False = false | destBoolean _ = raise Error "destBoolean"; val isBoolean = can destBoolean; fun isTrue fm = case fm of True => true | _ => false; fun isFalse fm = case fm of False => true | _ => false; (* Functions *) local fun funcs fs [] = fs | funcs fs (True :: fms) = funcs fs fms | funcs fs (False :: fms) = funcs fs fms | funcs fs (Metis_Atom atm :: fms) = funcs (Metis_NameAritySet.union (Metis_Atom.functions atm) fs) fms | funcs fs (Not p :: fms) = funcs fs (p :: fms) | funcs fs (And (p,q) :: fms) = funcs fs (p :: q :: fms) | funcs fs (Or (p,q) :: fms) = funcs fs (p :: q :: fms) | funcs fs (Imp (p,q) :: fms) = funcs fs (p :: q :: fms) | funcs fs (Iff (p,q) :: fms) = funcs fs (p :: q :: fms) | funcs fs (Forall (_,p) :: fms) = funcs fs (p :: fms) | funcs fs (Exists (_,p) :: fms) = funcs fs (p :: fms); in fun functions fm = funcs Metis_NameAritySet.empty [fm]; end; local fun funcs fs [] = fs | funcs fs (True :: fms) = funcs fs fms | funcs fs (False :: fms) = funcs fs fms | funcs fs (Metis_Atom atm :: fms) = funcs (Metis_NameSet.union (Metis_Atom.functionNames atm) fs) fms | funcs fs (Not p :: fms) = funcs fs (p :: fms) | funcs fs (And (p,q) :: fms) = funcs fs (p :: q :: fms) | funcs fs (Or (p,q) :: fms) = funcs fs (p :: q :: fms) | funcs fs (Imp (p,q) :: fms) = funcs fs (p :: q :: fms) | funcs fs (Iff (p,q) :: fms) = funcs fs (p :: q :: fms) | funcs fs (Forall (_,p) :: fms) = funcs fs (p :: fms) | funcs fs (Exists (_,p) :: fms) = funcs fs (p :: fms); in fun functionNames fm = funcs Metis_NameSet.empty [fm]; end; (* Relations *) local fun rels fs [] = fs | rels fs (True :: fms) = rels fs fms | rels fs (False :: fms) = rels fs fms | rels fs (Metis_Atom atm :: fms) = rels (Metis_NameAritySet.add fs (Metis_Atom.relation atm)) fms | rels fs (Not p :: fms) = rels fs (p :: fms) | rels fs (And (p,q) :: fms) = rels fs (p :: q :: fms) | rels fs (Or (p,q) :: fms) = rels fs (p :: q :: fms) | rels fs (Imp (p,q) :: fms) = rels fs (p :: q :: fms) | rels fs (Iff (p,q) :: fms) = rels fs (p :: q :: fms) | rels fs (Forall (_,p) :: fms) = rels fs (p :: fms) | rels fs (Exists (_,p) :: fms) = rels fs (p :: fms); in fun relations fm = rels Metis_NameAritySet.empty [fm]; end; local fun rels fs [] = fs | rels fs (True :: fms) = rels fs fms | rels fs (False :: fms) = rels fs fms | rels fs (Metis_Atom atm :: fms) = rels (Metis_NameSet.add fs (Metis_Atom.name atm)) fms | rels fs (Not p :: fms) = rels fs (p :: fms) | rels fs (And (p,q) :: fms) = rels fs (p :: q :: fms) | rels fs (Or (p,q) :: fms) = rels fs (p :: q :: fms) | rels fs (Imp (p,q) :: fms) = rels fs (p :: q :: fms) | rels fs (Iff (p,q) :: fms) = rels fs (p :: q :: fms) | rels fs (Forall (_,p) :: fms) = rels fs (p :: fms) | rels fs (Exists (_,p) :: fms) = rels fs (p :: fms); in fun relationNames fm = rels Metis_NameSet.empty [fm]; end; (* Atoms *) fun destAtom (Metis_Atom atm) = atm | destAtom _ = raise Error "Metis_Formula.destAtom"; val isAtom = can destAtom; (* Negations *) fun destNeg (Not p) = p | destNeg _ = raise Error "Metis_Formula.destNeg"; val isNeg = can destNeg; val stripNeg = let fun strip n (Not fm) = strip (n + 1) fm | strip n fm = (n,fm) in strip 0 end; (* Conjunctions *) fun listMkConj fms = case List.rev fms of [] => True | fm :: fms => List.foldl And fm fms; local fun strip cs (And (p,q)) = strip (p :: cs) q | strip cs fm = List.rev (fm :: cs); in fun stripConj True = [] | stripConj fm = strip [] fm; end; val flattenConj = let fun flat acc [] = acc | flat acc (And (p,q) :: fms) = flat acc (q :: p :: fms) | flat acc (True :: fms) = flat acc fms | flat acc (fm :: fms) = flat (fm :: acc) fms in fn fm => flat [] [fm] end; (* Disjunctions *) fun listMkDisj fms = case List.rev fms of [] => False | fm :: fms => List.foldl Or fm fms; local fun strip cs (Or (p,q)) = strip (p :: cs) q | strip cs fm = List.rev (fm :: cs); in fun stripDisj False = [] | stripDisj fm = strip [] fm; end; val flattenDisj = let fun flat acc [] = acc | flat acc (Or (p,q) :: fms) = flat acc (q :: p :: fms) | flat acc (False :: fms) = flat acc fms | flat acc (fm :: fms) = flat (fm :: acc) fms in fn fm => flat [] [fm] end; (* Equivalences *) fun listMkEquiv fms = case List.rev fms of [] => True | fm :: fms => List.foldl Iff fm fms; local fun strip cs (Iff (p,q)) = strip (p :: cs) q | strip cs fm = List.rev (fm :: cs); in fun stripEquiv True = [] | stripEquiv fm = strip [] fm; end; val flattenEquiv = let fun flat acc [] = acc | flat acc (Iff (p,q) :: fms) = flat acc (q :: p :: fms) | flat acc (True :: fms) = flat acc fms | flat acc (fm :: fms) = flat (fm :: acc) fms in fn fm => flat [] [fm] end; (* Universal quantifiers *) fun destForall (Forall v_f) = v_f | destForall _ = raise Error "destForall"; val isForall = can destForall; fun listMkForall ([],body) = body | listMkForall (v :: vs, body) = Forall (v, listMkForall (vs,body)); fun setMkForall (vs,body) = Metis_NameSet.foldr Forall body vs; local fun strip vs (Forall (v,b)) = strip (v :: vs) b | strip vs tm = (List.rev vs, tm); in val stripForall = strip []; end; (* Existential quantifiers *) fun destExists (Exists v_f) = v_f | destExists _ = raise Error "destExists"; val isExists = can destExists; fun listMkExists ([],body) = body | listMkExists (v :: vs, body) = Exists (v, listMkExists (vs,body)); fun setMkExists (vs,body) = Metis_NameSet.foldr Exists body vs; local fun strip vs (Exists (v,b)) = strip (v :: vs) b | strip vs tm = (List.rev vs, tm); in val stripExists = strip []; end; (* ------------------------------------------------------------------------- *) (* The size of a formula in symbols. *) (* ------------------------------------------------------------------------- *) local fun sz n [] = n | sz n (True :: fms) = sz (n + 1) fms | sz n (False :: fms) = sz (n + 1) fms | sz n (Metis_Atom atm :: fms) = sz (n + Metis_Atom.symbols atm) fms | sz n (Not p :: fms) = sz (n + 1) (p :: fms) | sz n (And (p,q) :: fms) = sz (n + 1) (p :: q :: fms) | sz n (Or (p,q) :: fms) = sz (n + 1) (p :: q :: fms) | sz n (Imp (p,q) :: fms) = sz (n + 1) (p :: q :: fms) | sz n (Iff (p,q) :: fms) = sz (n + 1) (p :: q :: fms) | sz n (Forall (_,p) :: fms) = sz (n + 1) (p :: fms) | sz n (Exists (_,p) :: fms) = sz (n + 1) (p :: fms); in fun symbols fm = sz 0 [fm]; end; (* ------------------------------------------------------------------------- *) (* A total comparison function for formulas. *) (* ------------------------------------------------------------------------- *) local fun cmp [] = EQUAL | cmp (f1_f2 :: fs) = if Metis_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 | (Metis_Atom atm1, Metis_Atom atm2) => (case Metis_Atom.compare (atm1,atm2) of LESS => LESS | EQUAL => cmp fs | GREATER => GREATER) | (Metis_Atom _, _) => LESS | (_, Metis_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 Metis_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 Metis_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. *) (* ------------------------------------------------------------------------- *) fun freeIn v = let fun f [] = false | f (True :: fms) = f fms | f (False :: fms) = f fms | f (Metis_Atom atm :: fms) = Metis_Atom.freeIn v atm orelse f fms | f (Not p :: fms) = f (p :: fms) | f (And (p,q) :: fms) = f (p :: q :: fms) | 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 Metis_Name.equal v w then f fms else f (p :: fms) | f (Exists (w,p) :: fms) = if Metis_Name.equal v w then f fms else f (p :: fms) in fn fm => f [fm] end; local fun fv vs [] = vs | fv vs ((_,True) :: fms) = fv vs fms | fv vs ((_,False) :: fms) = fv vs fms | fv vs ((bv, Metis_Atom atm) :: fms) = fv (Metis_NameSet.union vs (Metis_NameSet.difference (Metis_Atom.freeVars atm) bv)) fms | fv vs ((bv, Not p) :: fms) = fv vs ((bv,p) :: fms) | fv vs ((bv, And (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms) | fv vs ((bv, Or (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms) | fv vs ((bv, Imp (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms) | fv vs ((bv, Iff (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms) | fv vs ((bv, Forall (v,p)) :: fms) = fv vs ((Metis_NameSet.add bv v, p) :: fms) | fv vs ((bv, Exists (v,p)) :: fms) = fv vs ((Metis_NameSet.add bv v, p) :: fms); fun add (fm,vs) = fv vs [(Metis_NameSet.empty,fm)]; in fun freeVars fm = add (fm,Metis_NameSet.empty); fun freeVarsList fms = List.foldl add Metis_NameSet.empty fms; end; fun specialize fm = snd (stripForall fm); fun generalize fm = listMkForall (Metis_NameSet.toList (freeVars fm), fm); (* ------------------------------------------------------------------------- *) (* Substitutions. *) (* ------------------------------------------------------------------------- *) local fun substCheck sub fm = if Metis_Subst.null sub then fm else substFm sub fm and substFm sub fm = case fm of True => fm | False => fm | Metis_Atom (p,tms) => let val tms' = Metis_Sharing.map (Metis_Subst.subst sub) tms in if Metis_Portable.pointerEqual (tms,tms') then fm else Metis_Atom (p,tms') end | Not p => let val p' = substFm sub p in if Metis_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 | Imp (p,q) => substConn sub fm Imp p q | Iff (p,q) => substConn sub fm Iff p q | Forall (v,p) => substQuant sub fm Forall v p | Exists (v,p) => substQuant sub fm Exists v p and substConn sub fm conn p q = let val p' = substFm sub p and q' = substFm sub q in if Metis_Portable.pointerEqual (p,p') andalso Metis_Portable.pointerEqual (q,q') then fm else conn (p',q') end and substQuant sub fm quant v p = let val v' = let fun f (w,s) = if Metis_Name.equal w v then s else case Metis_Subst.peek sub w of NONE => Metis_NameSet.add s w | SOME tm => Metis_NameSet.union s (Metis_Term.freeVars tm) val vars = freeVars p val vars = Metis_NameSet.foldl f Metis_NameSet.empty vars in Metis_Term.variantPrime vars v end val sub = if Metis_Name.equal v v' then Metis_Subst.remove sub (Metis_NameSet.singleton v) else Metis_Subst.insert sub (v, Metis_Term.Var v') val p' = substCheck sub p in if Metis_Name.equal v v' andalso Metis_Portable.pointerEqual (p,p') then fm else quant (v',p') end; in val subst = substCheck; end; (* ------------------------------------------------------------------------- *) (* The equality relation. *) (* ------------------------------------------------------------------------- *) fun mkEq a_b = Metis_Atom (Metis_Atom.mkEq a_b); fun destEq fm = Metis_Atom.destEq (destAtom fm); val isEq = can destEq; fun mkNeq a_b = Not (mkEq a_b); fun destNeq (Not fm) = destEq fm | destNeq _ = raise Error "Metis_Formula.destNeq"; val isNeq = can destNeq; fun mkRefl tm = Metis_Atom (Metis_Atom.mkRefl tm); fun destRefl fm = Metis_Atom.destRefl (destAtom fm); val isRefl = can destRefl; fun sym fm = Metis_Atom (Metis_Atom.sym (destAtom fm)); fun lhs fm = fst (destEq fm); fun rhs fm = snd (destEq fm); (* ------------------------------------------------------------------------- *) (* Parsing and pretty-printing. *) (* ------------------------------------------------------------------------- *) type quotation = formula Metis_Parse.quotation; val truthName = Metis_Name.fromString "T" and falsityName = Metis_Name.fromString "F" and conjunctionName = Metis_Name.fromString "/\\" and disjunctionName = Metis_Name.fromString "\\/" and implicationName = Metis_Name.fromString "==>" and equivalenceName = Metis_Name.fromString "<=>" and universalName = Metis_Name.fromString "!" and existentialName = Metis_Name.fromString "?"; local fun demote True = Metis_Term.Fn (truthName,[]) | demote False = Metis_Term.Fn (falsityName,[]) | demote (Metis_Atom (p,tms)) = Metis_Term.Fn (p,tms) | demote (Not p) = let val Unsynchronized.ref s = Metis_Term.negation in Metis_Term.Fn (Metis_Name.fromString s, [demote p]) end | demote (And (p,q)) = Metis_Term.Fn (conjunctionName, [demote p, demote q]) | demote (Or (p,q)) = Metis_Term.Fn (disjunctionName, [demote p, demote q]) | demote (Imp (p,q)) = Metis_Term.Fn (implicationName, [demote p, demote q]) | demote (Iff (p,q)) = Metis_Term.Fn (equivalenceName, [demote p, demote q]) | demote (Forall (v,b)) = Metis_Term.Fn (universalName, [Metis_Term.Var v, demote b]) | demote (Exists (v,b)) = Metis_Term.Fn (existentialName, [Metis_Term.Var v, demote b]); in fun pp fm = Metis_Term.pp (demote fm); end; val toString = Metis_Print.toString pp; local fun isQuant [Metis_Term.Var _, _] = true | isQuant _ = false; fun promote (Metis_Term.Var v) = Metis_Atom (v,[]) | promote (Metis_Term.Fn (f,tms)) = if Metis_Name.equal f truthName andalso List.null tms then True else if Metis_Name.equal f falsityName andalso List.null tms then False else if Metis_Name.toString f = !Metis_Term.negation andalso length tms = 1 then Not (promote (hd tms)) else if Metis_Name.equal f conjunctionName andalso length tms = 2 then And (promote (hd tms), promote (List.nth (tms,1))) else if Metis_Name.equal f disjunctionName andalso length tms = 2 then Or (promote (hd tms), promote (List.nth (tms,1))) else if Metis_Name.equal f implicationName andalso length tms = 2 then Imp (promote (hd tms), promote (List.nth (tms,1))) else if Metis_Name.equal f equivalenceName andalso length tms = 2 then Iff (promote (hd tms), promote (List.nth (tms,1))) else if Metis_Name.equal f universalName andalso isQuant tms then Forall (Metis_Term.destVar (hd tms), promote (List.nth (tms,1))) else if Metis_Name.equal f existentialName andalso isQuant tms then Exists (Metis_Term.destVar (hd tms), promote (List.nth (tms,1))) else Metis_Atom (f,tms); in fun fromString s = promote (Metis_Term.fromString s); end; val parse = Metis_Parse.parseQuotation toString fromString; (* ------------------------------------------------------------------------- *) (* Splitting goals. *) (* ------------------------------------------------------------------------- *) local fun add_asms asms goal = if List.null asms then goal else Imp (listMkConj (List.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)) => List.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 (Not f2 :: asms) true f1 | (false, Exists (v,f)) => List.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 () = Metis_Print.trace pp "Metis_Formula.splitGoal: fm" fm val () = Metis_Print.trace (Metis_Print.ppList pp) "Metis_Formula.splitGoal: result" result in result end; *) end structure Metis_FormulaOrdered = struct type t = Metis_Formula.formula val compare = Metis_Formula.compare end structure Metis_FormulaMap = Metis_KeyMap (Metis_FormulaOrdered); structure Metis_FormulaSet = Metis_ElementSet (Metis_FormulaMap); (**** Original file: src/Literal.sig ****) (* ========================================================================= *) (* FIRST ORDER LOGIC LITERALS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Literal = sig (* ------------------------------------------------------------------------- *) (* A type for storing first order logic literals. *) (* ------------------------------------------------------------------------- *) type polarity = bool type literal = polarity * Metis_Atom.atom (* ------------------------------------------------------------------------- *) (* Constructors and destructors. *) (* ------------------------------------------------------------------------- *) val polarity : literal -> polarity val atom : literal -> Metis_Atom.atom val name : literal -> Metis_Atom.relationName val arguments : literal -> Metis_Term.term list val arity : literal -> int val positive : literal -> bool val negative : literal -> bool val negate : literal -> literal val relation : literal -> Metis_Atom.relation val functions : literal -> Metis_NameAritySet.set val functionNames : literal -> Metis_NameSet.set (* Binary relations *) val mkBinop : Metis_Atom.relationName -> polarity * Metis_Term.term * Metis_Term.term -> literal val destBinop : Metis_Atom.relationName -> literal -> polarity * Metis_Term.term * Metis_Term.term val isBinop : Metis_Atom.relationName -> literal -> bool (* Formulas *) val toFormula : literal -> Metis_Formula.formula val fromFormula : Metis_Formula.formula -> literal (* ------------------------------------------------------------------------- *) (* The size of a literal in symbols. *) (* ------------------------------------------------------------------------- *) val symbols : literal -> int (* ------------------------------------------------------------------------- *) (* A total comparison function for literals. *) (* ------------------------------------------------------------------------- *) val compare : literal * literal -> order (* negative < positive *) val equal : literal -> literal -> bool (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) val subterm : literal -> Metis_Term.path -> Metis_Term.term val subterms : literal -> (Metis_Term.path * Metis_Term.term) list val replace : literal -> Metis_Term.path * Metis_Term.term -> literal (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) val freeIn : Metis_Term.var -> literal -> bool val freeVars : literal -> Metis_NameSet.set (* ------------------------------------------------------------------------- *) (* Substitutions. *) (* ------------------------------------------------------------------------- *) val subst : Metis_Subst.subst -> literal -> literal (* ------------------------------------------------------------------------- *) (* Matching. *) (* ------------------------------------------------------------------------- *) val match : (* raises Error *) Metis_Subst.subst -> literal -> literal -> Metis_Subst.subst (* ------------------------------------------------------------------------- *) (* Unification. *) (* ------------------------------------------------------------------------- *) val unify : (* raises Error *) Metis_Subst.subst -> literal -> literal -> Metis_Subst.subst (* ------------------------------------------------------------------------- *) (* The equality relation. *) (* ------------------------------------------------------------------------- *) val mkEq : Metis_Term.term * Metis_Term.term -> literal val destEq : literal -> Metis_Term.term * Metis_Term.term val isEq : literal -> bool val mkNeq : Metis_Term.term * Metis_Term.term -> literal val destNeq : literal -> Metis_Term.term * Metis_Term.term val isNeq : literal -> bool val mkRefl : Metis_Term.term -> literal val destRefl : literal -> Metis_Term.term val isRefl : literal -> bool val mkIrrefl : Metis_Term.term -> literal val destIrrefl : literal -> Metis_Term.term val isIrrefl : literal -> bool (* The following work with both equalities and disequalities *) val sym : literal -> literal (* raises Error if given a refl or irrefl *) val lhs : literal -> Metis_Term.term val rhs : literal -> Metis_Term.term (* ------------------------------------------------------------------------- *) (* Special support for terms with type annotations. *) (* ------------------------------------------------------------------------- *) val typedSymbols : literal -> int val nonVarTypedSubterms : literal -> (Metis_Term.path * Metis_Term.term) list (* ------------------------------------------------------------------------- *) (* Parsing and pretty-printing. *) (* ------------------------------------------------------------------------- *) val pp : literal Metis_Print.pp val toString : literal -> string val fromString : string -> literal val parse : Metis_Term.term Metis_Parse.quotation -> literal end (**** Original file: src/Literal.sml ****) (* ========================================================================= *) (* FIRST ORDER LOGIC LITERALS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Literal :> Metis_Literal = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* A type for storing first order logic literals. *) (* ------------------------------------------------------------------------- *) type polarity = bool; type literal = polarity * Metis_Atom.atom; (* ------------------------------------------------------------------------- *) (* Constructors and destructors. *) (* ------------------------------------------------------------------------- *) fun polarity ((pol,_) : literal) = pol; fun atom ((_,atm) : literal) = atm; fun name lit = Metis_Atom.name (atom lit); fun arguments lit = Metis_Atom.arguments (atom lit); fun arity lit = Metis_Atom.arity (atom lit); fun positive lit = polarity lit; fun negative lit = not (polarity lit); fun negate (pol,atm) : literal = (not pol, atm) fun relation lit = Metis_Atom.relation (atom lit); fun functions lit = Metis_Atom.functions (atom lit); fun functionNames lit = Metis_Atom.functionNames (atom lit); (* Binary relations *) fun mkBinop rel (pol,a,b) : literal = (pol, Metis_Atom.mkBinop rel (a,b)); fun destBinop rel ((pol,atm) : literal) = case Metis_Atom.destBinop rel atm of (a,b) => (pol,a,b); fun isBinop rel = can (destBinop rel); (* Formulas *) fun toFormula (true,atm) = Metis_Formula.Metis_Atom atm | toFormula (false,atm) = Metis_Formula.Not (Metis_Formula.Metis_Atom atm); fun fromFormula (Metis_Formula.Metis_Atom atm) = (true,atm) | fromFormula (Metis_Formula.Not (Metis_Formula.Metis_Atom atm)) = (false,atm) | fromFormula _ = raise Error "Metis_Literal.fromFormula"; (* ------------------------------------------------------------------------- *) (* The size of a literal in symbols. *) (* ------------------------------------------------------------------------- *) fun symbols ((_,atm) : literal) = Metis_Atom.symbols atm; (* ------------------------------------------------------------------------- *) (* A total comparison function for literals. *) (* ------------------------------------------------------------------------- *) val compare = prodCompare boolCompare Metis_Atom.compare; fun equal (p1,atm1) (p2,atm2) = p1 = p2 andalso Metis_Atom.equal atm1 atm2; (* ------------------------------------------------------------------------- *) (* Subterms. *) (* ------------------------------------------------------------------------- *) fun subterm lit path = Metis_Atom.subterm (atom lit) path; fun subterms lit = Metis_Atom.subterms (atom lit); fun replace (lit as (pol,atm)) path_tm = let val atm' = Metis_Atom.replace atm path_tm in if Metis_Portable.pointerEqual (atm,atm') then lit else (pol,atm') end; (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) fun freeIn v lit = Metis_Atom.freeIn v (atom lit); fun freeVars lit = Metis_Atom.freeVars (atom lit); (* ------------------------------------------------------------------------- *) (* Substitutions. *) (* ------------------------------------------------------------------------- *) fun subst sub (lit as (pol,atm)) : literal = let val atm' = Metis_Atom.subst sub atm in if Metis_Portable.pointerEqual (atm',atm) then lit else (pol,atm') end; (* ------------------------------------------------------------------------- *) (* Matching. *) (* ------------------------------------------------------------------------- *) fun match sub ((pol1,atm1) : literal) (pol2,atm2) = let val _ = pol1 = pol2 orelse raise Error "Metis_Literal.match" in Metis_Atom.match sub atm1 atm2 end; (* ------------------------------------------------------------------------- *) (* Unification. *) (* ------------------------------------------------------------------------- *) fun unify sub ((pol1,atm1) : literal) (pol2,atm2) = let val _ = pol1 = pol2 orelse raise Error "Metis_Literal.unify" in Metis_Atom.unify sub atm1 atm2 end; (* ------------------------------------------------------------------------- *) (* The equality relation. *) (* ------------------------------------------------------------------------- *) fun mkEq l_r : literal = (true, Metis_Atom.mkEq l_r); fun destEq ((true,atm) : literal) = Metis_Atom.destEq atm | destEq (false,_) = raise Error "Metis_Literal.destEq"; val isEq = can destEq; fun mkNeq l_r : literal = (false, Metis_Atom.mkEq l_r); fun destNeq ((false,atm) : literal) = Metis_Atom.destEq atm | destNeq (true,_) = raise Error "Metis_Literal.destNeq"; val isNeq = can destNeq; fun mkRefl tm = (true, Metis_Atom.mkRefl tm); fun destRefl (true,atm) = Metis_Atom.destRefl atm | destRefl (false,_) = raise Error "Metis_Literal.destRefl"; val isRefl = can destRefl; fun mkIrrefl tm = (false, Metis_Atom.mkRefl tm); fun destIrrefl (true,_) = raise Error "Metis_Literal.destIrrefl" | destIrrefl (false,atm) = Metis_Atom.destRefl atm; val isIrrefl = can destIrrefl; fun sym (pol,atm) : literal = (pol, Metis_Atom.sym atm); fun lhs ((_,atm) : literal) = Metis_Atom.lhs atm; fun rhs ((_,atm) : literal) = Metis_Atom.rhs atm; (* ------------------------------------------------------------------------- *) (* Special support for terms with type annotations. *) (* ------------------------------------------------------------------------- *) fun typedSymbols ((_,atm) : literal) = Metis_Atom.typedSymbols atm; fun nonVarTypedSubterms ((_,atm) : literal) = Metis_Atom.nonVarTypedSubterms atm; (* ------------------------------------------------------------------------- *) (* Parsing and pretty-printing. *) (* ------------------------------------------------------------------------- *) val pp = Metis_Print.ppMap toFormula Metis_Formula.pp; val toString = Metis_Print.toString pp; fun fromString s = fromFormula (Metis_Formula.fromString s); val parse = Metis_Parse.parseQuotation Metis_Term.toString fromString; end structure Metis_LiteralOrdered = struct type t = Metis_Literal.literal val compare = Metis_Literal.compare end structure Metis_LiteralMap = Metis_KeyMap (Metis_LiteralOrdered); structure Metis_LiteralSet = struct local structure S = Metis_ElementSet (Metis_LiteralMap); in open S; end; fun negateMember lit set = member (Metis_Literal.negate lit) set; val negate = let fun f (lit,set) = add set (Metis_Literal.negate lit) in foldl f empty end; val relations = let fun f (lit,set) = Metis_NameAritySet.add set (Metis_Literal.relation lit) in foldl f Metis_NameAritySet.empty end; val functions = let fun f (lit,set) = Metis_NameAritySet.union set (Metis_Literal.functions lit) in foldl f Metis_NameAritySet.empty end; fun freeIn v = exists (Metis_Literal.freeIn v); val freeVars = let fun f (lit,set) = Metis_NameSet.union set (Metis_Literal.freeVars lit) in foldl f Metis_NameSet.empty end; val freeVarsList = let fun f (lits,set) = Metis_NameSet.union set (freeVars lits) in List.foldl f Metis_NameSet.empty end; val symbols = let fun f (lit,z) = Metis_Literal.symbols lit + z in foldl f 0 end; val typedSymbols = let fun f (lit,z) = Metis_Literal.typedSymbols lit + z in foldl f 0 end; fun subst sub lits = let fun substLit (lit,(eq,lits')) = let val lit' = Metis_Literal.subst sub lit val eq = eq andalso Metis_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 = Metis_Formula.listMkConj (List.map Metis_Literal.toFormula (toList set)); fun disjoin set = Metis_Formula.listMkDisj (List.map Metis_Literal.toFormula (toList set)); val pp = Metis_Print.ppMap toList (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_Literal.pp)); end structure Metis_LiteralSetOrdered = struct type t = Metis_LiteralSet.set val compare = Metis_LiteralSet.compare end structure Metis_LiteralSetMap = Metis_KeyMap (Metis_LiteralSetOrdered); structure Metis_LiteralSetSet = Metis_ElementSet (Metis_LiteralSetMap); (**** Original file: src/Thm.sig ****) (* ========================================================================= *) (* A LOGICAL KERNEL FOR FIRST ORDER CLAUSAL THEOREMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Thm = sig (* ------------------------------------------------------------------------- *) (* An abstract type of first order logic theorems. *) (* ------------------------------------------------------------------------- *) type thm (* ------------------------------------------------------------------------- *) (* Theorem destructors. *) (* ------------------------------------------------------------------------- *) type clause = Metis_LiteralSet.set datatype inferenceType = Axiom | Assume | Metis_Subst | Factor | Resolve | Refl | Equality type inference = inferenceType * thm list val clause : thm -> clause val inference : thm -> inference (* Tautologies *) val isTautology : thm -> bool (* Contradictions *) val isContradiction : thm -> bool (* Unit theorems *) val destUnit : thm -> Metis_Literal.literal val isUnit : thm -> bool (* Unit equality theorems *) val destUnitEq : thm -> Metis_Term.term * Metis_Term.term val isUnitEq : thm -> bool (* Literals *) val member : Metis_Literal.literal -> thm -> bool val negateMember : Metis_Literal.literal -> thm -> bool (* ------------------------------------------------------------------------- *) (* A total order. *) (* ------------------------------------------------------------------------- *) val compare : thm * thm -> order val equal : thm -> thm -> bool (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) val freeIn : Metis_Term.var -> thm -> bool val freeVars : thm -> Metis_NameSet.set (* ------------------------------------------------------------------------- *) (* Pretty-printing. *) (* ------------------------------------------------------------------------- *) val ppInferenceType : inferenceType Metis_Print.pp val inferenceTypeToString : inferenceType -> string val pp : thm Metis_Print.pp val toString : thm -> string (* ------------------------------------------------------------------------- *) (* Primitive rules of inference. *) (* ------------------------------------------------------------------------- *) (* ------------------------------------------------------------------------- *) (* *) (* ----- axiom C *) (* C *) (* ------------------------------------------------------------------------- *) val axiom : clause -> thm (* ------------------------------------------------------------------------- *) (* *) (* ----------- assume L *) (* L \/ ~L *) (* ------------------------------------------------------------------------- *) val assume : Metis_Literal.literal -> thm (* ------------------------------------------------------------------------- *) (* C *) (* -------- subst s *) (* C[s] *) (* ------------------------------------------------------------------------- *) val subst : Metis_Subst.subst -> thm -> thm (* ------------------------------------------------------------------------- *) (* L \/ C ~L \/ D *) (* --------------------- resolve L *) (* C \/ D *) (* *) (* The literal L must occur in the first theorem, and the literal ~L must *) (* occur in the second theorem. *) (* ------------------------------------------------------------------------- *) val resolve : Metis_Literal.literal -> thm -> thm -> thm (* ------------------------------------------------------------------------- *) (* *) (* --------- refl t *) (* t = t *) (* ------------------------------------------------------------------------- *) val refl : Metis_Term.term -> thm (* ------------------------------------------------------------------------- *) (* *) (* ------------------------ equality L p t *) (* ~(s = t) \/ ~L \/ L' *) (* *) (* where s is the subterm of L at path p, and L' is L with the subterm at *) (* path p being replaced by t. *) (* ------------------------------------------------------------------------- *) val equality : Metis_Literal.literal -> Metis_Term.path -> Metis_Term.term -> thm end (**** Original file: src/Thm.sml ****) (* ========================================================================= *) (* A LOGICAL KERNEL FOR FIRST ORDER CLAUSAL THEOREMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Thm :> Metis_Thm = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* An abstract type of first order logic theorems. *) (* ------------------------------------------------------------------------- *) type clause = Metis_LiteralSet.set; datatype inferenceType = Axiom | Assume | Metis_Subst | Factor | Resolve | Refl | Equality; datatype thm = Metis_Thm of clause * (inferenceType * thm list); type inference = inferenceType * thm list; (* ------------------------------------------------------------------------- *) (* Theorem destructors. *) (* ------------------------------------------------------------------------- *) fun clause (Metis_Thm (cl,_)) = cl; fun inference (Metis_Thm (_,inf)) = inf; (* Tautologies *) local fun chk (_,NONE) = NONE | chk ((pol,atm), SOME set) = if (pol andalso Metis_Atom.isRefl atm) orelse Metis_AtomSet.member atm set then NONE else SOME (Metis_AtomSet.add set atm); in fun isTautology th = case Metis_LiteralSet.foldl chk (SOME Metis_AtomSet.empty) (clause th) of SOME _ => false | NONE => true; end; (* Contradictions *) fun isContradiction th = Metis_LiteralSet.null (clause th); (* Unit theorems *) fun destUnit (Metis_Thm (cl,_)) = if Metis_LiteralSet.size cl = 1 then Metis_LiteralSet.pick cl else raise Error "Metis_Thm.destUnit"; val isUnit = can destUnit; (* Unit equality theorems *) fun destUnitEq th = Metis_Literal.destEq (destUnit th); val isUnitEq = can destUnitEq; (* Literals *) fun member lit (Metis_Thm (cl,_)) = Metis_LiteralSet.member lit cl; fun negateMember lit (Metis_Thm (cl,_)) = Metis_LiteralSet.negateMember lit cl; (* ------------------------------------------------------------------------- *) (* A total order. *) (* ------------------------------------------------------------------------- *) fun compare (th1,th2) = Metis_LiteralSet.compare (clause th1, clause th2); fun equal th1 th2 = Metis_LiteralSet.equal (clause th1) (clause th2); (* ------------------------------------------------------------------------- *) (* Free variables. *) (* ------------------------------------------------------------------------- *) fun freeIn v (Metis_Thm (cl,_)) = Metis_LiteralSet.freeIn v cl; fun freeVars (Metis_Thm (cl,_)) = Metis_LiteralSet.freeVars cl; (* ------------------------------------------------------------------------- *) (* Pretty-printing. *) (* ------------------------------------------------------------------------- *) fun inferenceTypeToString Axiom = "Axiom" | inferenceTypeToString Assume = "Assume" | inferenceTypeToString Metis_Subst = "Metis_Subst" | inferenceTypeToString Factor = "Factor" | inferenceTypeToString Resolve = "Resolve" | inferenceTypeToString Refl = "Refl" | inferenceTypeToString Equality = "Equality"; fun ppInferenceType inf = Metis_Print.ppString (inferenceTypeToString inf); local fun toFormula th = Metis_Formula.listMkDisj (List.map Metis_Literal.toFormula (Metis_LiteralSet.toList (clause th))); in fun pp th = Metis_Print.inconsistentBlock 3 [Metis_Print.ppString "|- ", Metis_Formula.pp (toFormula th)]; end; val toString = Metis_Print.toString pp; (* ------------------------------------------------------------------------- *) (* Primitive rules of inference. *) (* ------------------------------------------------------------------------- *) (* ------------------------------------------------------------------------- *) (* *) (* ----- axiom C *) (* C *) (* ------------------------------------------------------------------------- *) fun axiom cl = Metis_Thm (cl,(Axiom,[])); (* ------------------------------------------------------------------------- *) (* *) (* ----------- assume L *) (* L \/ ~L *) (* ------------------------------------------------------------------------- *) fun assume lit = Metis_Thm (Metis_LiteralSet.fromList [lit, Metis_Literal.negate lit], (Assume,[])); (* ------------------------------------------------------------------------- *) (* C *) (* -------- subst s *) (* C[s] *) (* ------------------------------------------------------------------------- *) fun subst sub (th as Metis_Thm (cl,inf)) = let val cl' = Metis_LiteralSet.subst sub cl in if Metis_Portable.pointerEqual (cl,cl') then th else case inf of (Metis_Subst,_) => Metis_Thm (cl',inf) | _ => Metis_Thm (cl',(Metis_Subst,[th])) end; (* ------------------------------------------------------------------------- *) (* L \/ C ~L \/ D *) (* --------------------- resolve L *) (* C \/ D *) (* *) (* The literal L must occur in the first theorem, and the literal ~L must *) (* occur in the second theorem. *) (* ------------------------------------------------------------------------- *) fun resolve lit (th1 as Metis_Thm (cl1,_)) (th2 as Metis_Thm (cl2,_)) = let val cl1' = Metis_LiteralSet.delete cl1 lit and cl2' = Metis_LiteralSet.delete cl2 (Metis_Literal.negate lit) in Metis_Thm (Metis_LiteralSet.union cl1' cl2', (Resolve,[th1,th2])) end; (*MetisDebug val resolve = fn lit => fn pos => fn neg => resolve lit pos neg handle Error err => raise Error ("Metis_Thm.resolve:\nlit = " ^ Metis_Literal.toString lit ^ "\npos = " ^ toString pos ^ "\nneg = " ^ toString neg ^ "\n" ^ err) | Bug bug => raise Bug ("Metis_Thm.resolve:\nlit = " ^ Metis_Literal.toString lit ^ "\npos = " ^ toString pos ^ "\nneg = " ^ toString neg ^ "\n" ^ bug); *) (* ------------------------------------------------------------------------- *) (* *) (* --------- refl t *) (* t = t *) (* ------------------------------------------------------------------------- *) fun refl tm = Metis_Thm (Metis_LiteralSet.singleton (true, Metis_Atom.mkRefl tm), (Refl,[])); (* ------------------------------------------------------------------------- *) (* *) (* ------------------------ equality L p t *) (* ~(s = t) \/ ~L \/ L' *) (* *) (* where s is the subterm of L at path p, and L' is L with the subterm at *) (* path p being replaced by t. *) (* ------------------------------------------------------------------------- *) fun equality lit path t = let val s = Metis_Literal.subterm lit path val lit' = Metis_Literal.replace lit (path,t) val eqLit = Metis_Literal.mkNeq (s,t) val cl = Metis_LiteralSet.fromList [eqLit, Metis_Literal.negate lit, lit'] in Metis_Thm (cl,(Equality,[])) end; end (**** Original file: src/Proof.sig ****) (* ========================================================================= *) (* PROOFS IN FIRST ORDER LOGIC *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Proof = sig (* ------------------------------------------------------------------------- *) (* A type of first order logic proofs. *) (* ------------------------------------------------------------------------- *) datatype inference = Axiom of Metis_LiteralSet.set | Assume of Metis_Atom.atom | Metis_Subst of Metis_Subst.subst * Metis_Thm.thm | Resolve of Metis_Atom.atom * Metis_Thm.thm * Metis_Thm.thm | Refl of Metis_Term.term | Equality of Metis_Literal.literal * Metis_Term.path * Metis_Term.term type proof = (Metis_Thm.thm * inference) list (* ------------------------------------------------------------------------- *) (* Reconstructing single inferences. *) (* ------------------------------------------------------------------------- *) val inferenceType : inference -> Metis_Thm.inferenceType val parents : inference -> Metis_Thm.thm list val inferenceToThm : inference -> Metis_Thm.thm val thmToInference : Metis_Thm.thm -> inference (* ------------------------------------------------------------------------- *) (* Reconstructing whole proofs. *) (* ------------------------------------------------------------------------- *) 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_Print.pp val inferenceToString : inference -> string val pp : proof Metis_Print.pp val toString : proof -> string end (**** Original file: src/Proof.sml ****) (* ========================================================================= *) (* PROOFS IN FIRST ORDER LOGIC *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Proof :> Metis_Proof = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* A type of first order logic proofs. *) (* ------------------------------------------------------------------------- *) datatype inference = Axiom of Metis_LiteralSet.set | Assume of Metis_Atom.atom | Metis_Subst of Metis_Subst.subst * Metis_Thm.thm | Resolve of Metis_Atom.atom * Metis_Thm.thm * Metis_Thm.thm | Refl of Metis_Term.term | Equality of Metis_Literal.literal * Metis_Term.path * Metis_Term.term; type proof = (Metis_Thm.thm * inference) list; (* ------------------------------------------------------------------------- *) (* Printing. *) (* ------------------------------------------------------------------------- *) fun inferenceType (Axiom _) = Metis_Thm.Axiom | inferenceType (Assume _) = Metis_Thm.Assume | inferenceType (Metis_Subst _) = Metis_Thm.Metis_Subst | inferenceType (Resolve _) = Metis_Thm.Resolve | inferenceType (Refl _) = Metis_Thm.Refl | inferenceType (Equality _) = Metis_Thm.Equality; local fun ppAssume atm = Metis_Print.sequence Metis_Print.break (Metis_Atom.pp atm); fun ppSubst ppThm (sub,thm) = Metis_Print.sequence Metis_Print.break (Metis_Print.inconsistentBlock 1 [Metis_Print.ppString "{", Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Subst.pp ("sub",sub), Metis_Print.ppString ",", Metis_Print.break, Metis_Print.ppOp2 " =" Metis_Print.ppString ppThm ("thm",thm), Metis_Print.ppString "}"]); fun ppResolve ppThm (res,pos,neg) = Metis_Print.sequence Metis_Print.break (Metis_Print.inconsistentBlock 1 [Metis_Print.ppString "{", Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Atom.pp ("res",res), Metis_Print.ppString ",", Metis_Print.break, Metis_Print.ppOp2 " =" Metis_Print.ppString ppThm ("pos",pos), Metis_Print.ppString ",", Metis_Print.break, Metis_Print.ppOp2 " =" Metis_Print.ppString ppThm ("neg",neg), Metis_Print.ppString "}"]); fun ppRefl tm = Metis_Print.sequence Metis_Print.break (Metis_Term.pp tm); fun ppEquality (lit,path,res) = Metis_Print.sequence Metis_Print.break (Metis_Print.inconsistentBlock 1 [Metis_Print.ppString "{", Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Literal.pp ("lit",lit), Metis_Print.ppString ",", Metis_Print.break, Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Term.ppPath ("path",path), Metis_Print.ppString ",", Metis_Print.break, Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Term.pp ("res",res), Metis_Print.ppString "}"]); fun ppInf ppAxiom ppThm inf = let val infString = Metis_Thm.inferenceTypeToString (inferenceType inf) in Metis_Print.inconsistentBlock 2 [Metis_Print.ppString infString, (case inf of Axiom cl => ppAxiom cl | Assume x => ppAssume x | Metis_Subst x => ppSubst ppThm x | Resolve x => ppResolve ppThm x | Refl x => ppRefl x | Equality x => ppEquality x)] end; fun ppAxiom cl = Metis_Print.sequence Metis_Print.break (Metis_Print.ppMap Metis_LiteralSet.toList (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_Literal.pp)) cl); in val ppInference = ppInf ppAxiom Metis_Thm.pp; fun pp prf = let fun thmString n = "(" ^ Int.toString n ^ ")" val prf = enumerate prf fun ppThm th = Metis_Print.ppString let val cl = Metis_Thm.clause th fun pred (_,(th',_)) = Metis_LiteralSet.equal (Metis_Thm.clause th') cl in case List.find pred prf of NONE => "(?)" | SOME (n,_) => thmString n end fun ppStep (n,(th,inf)) = let val s = thmString n in Metis_Print.sequence (Metis_Print.consistentBlock (1 + size s) [Metis_Print.ppString (s ^ " "), Metis_Thm.pp th, Metis_Print.breaks 2, Metis_Print.ppBracket "[" "]" (ppInf (K Metis_Print.skip) ppThm) inf]) Metis_Print.newline end in Metis_Print.consistentBlock 0 [Metis_Print.ppString "START OF PROOF", Metis_Print.newline, Metis_Print.program (List.map ppStep prf), Metis_Print.ppString "END OF PROOF"] end (*MetisDebug handle Error err => raise Bug ("Metis_Proof.pp: shouldn't fail:\n" ^ err); *) end; val inferenceToString = Metis_Print.toString ppInference; val toString = Metis_Print.toString pp; (* ------------------------------------------------------------------------- *) (* Reconstructing single inferences. *) (* ------------------------------------------------------------------------- *) fun parents (Axiom _) = [] | parents (Assume _) = [] | parents (Metis_Subst (_,th)) = [th] | parents (Resolve (_,th,th')) = [th,th'] | parents (Refl _) = [] | parents (Equality _) = []; fun inferenceToThm (Axiom cl) = Metis_Thm.axiom cl | inferenceToThm (Assume atm) = Metis_Thm.assume (true,atm) | inferenceToThm (Metis_Subst (sub,th)) = Metis_Thm.subst sub th | inferenceToThm (Resolve (atm,th,th')) = Metis_Thm.resolve (true,atm) th th' | inferenceToThm (Refl tm) = Metis_Thm.refl tm | inferenceToThm (Equality (lit,path,r)) = Metis_Thm.equality lit path r; local fun reconstructSubst cl cl' = let fun recon [] = let (*MetisTrace3 val () = Metis_Print.trace Metis_LiteralSet.pp "reconstructSubst: cl" cl val () = Metis_Print.trace Metis_LiteralSet.pp "reconstructSubst: cl'" cl' *) in raise Bug "can't reconstruct Metis_Subst rule" end | recon (([],sub) :: others) = if Metis_LiteralSet.equal (Metis_LiteralSet.subst sub cl) cl' then sub else recon others | recon ((lit :: lits, sub) :: others) = let fun checkLit (lit',acc) = case total (Metis_Literal.match sub lit) lit' of NONE => acc | SOME sub => (lits,sub) :: acc in recon (Metis_LiteralSet.foldl checkLit others cl') end in Metis_Subst.normalize (recon [(Metis_LiteralSet.toList cl, Metis_Subst.empty)]) end (*MetisDebug handle Error err => raise Bug ("Metis_Proof.recontructSubst: shouldn't fail:\n" ^ err); *) fun reconstructResolvant cl1 cl2 cl = (if not (Metis_LiteralSet.subset cl1 cl) then Metis_LiteralSet.pick (Metis_LiteralSet.difference cl1 cl) else if not (Metis_LiteralSet.subset cl2 cl) then Metis_Literal.negate (Metis_LiteralSet.pick (Metis_LiteralSet.difference cl2 cl)) else (* A useless resolution, but we must reconstruct it anyway *) let val cl1' = Metis_LiteralSet.negate cl1 and cl2' = Metis_LiteralSet.negate cl2 val lits = Metis_LiteralSet.intersectList [cl1,cl1',cl2,cl2'] in if not (Metis_LiteralSet.null lits) then Metis_LiteralSet.pick lits else raise Bug "can't reconstruct Resolve rule" end) (*MetisDebug handle Error err => raise Bug ("Metis_Proof.recontructResolvant: shouldn't fail:\n" ^ err); *) fun reconstructEquality cl = let (*MetisTrace3 val () = Metis_Print.trace Metis_LiteralSet.pp "Metis_Proof.reconstructEquality: cl" cl *) fun sync s t path (f,a) (f',a') = if not (Metis_Name.equal f f' andalso length a = length a') then NONE else let val itms = enumerate (zip a a') in case List.filter (not o uncurry Metis_Term.equal o snd) itms of [(i,(tm,tm'))] => let val path = i :: path in if Metis_Term.equal tm s andalso Metis_Term.equal tm' t then SOME (List.rev path) else case (tm,tm') of (Metis_Term.Fn f_a, Metis_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 else let val (s,t) = Metis_Literal.destNeq neq val path = if not (Metis_Term.equal s t) then sync s t [] atm atm' else if not (Metis_Atom.equal atm atm') then NONE else Metis_Atom.find (Metis_Term.equal s) atm in case path of SOME path => SOME ((pol',atm),path,t) | NONE => NONE end val candidates = case List.partition Metis_Literal.isNeq (Metis_LiteralSet.toList cl) of ([l1],[l2,l3]) => [(l1,l2,l3),(l1,l3,l2)] | ([l1,l2],[l3]) => [(l1,l2,l3),(l1,l3,l2),(l2,l1,l3),(l2,l3,l1)] | ([l1],[l2]) => [(l1,l1,l2),(l1,l2,l1)] | _ => raise Bug "reconstructEquality: malformed" (*MetisTrace3 val ppCands = Metis_Print.ppList (Metis_Print.ppTriple Metis_Literal.pp Metis_Literal.pp Metis_Literal.pp) val () = Metis_Print.trace ppCands "Metis_Proof.reconstructEquality: candidates" candidates *) in case first recon candidates of SOME info => info | NONE => raise Bug "can't reconstruct Equality rule" end (*MetisDebug handle Error err => raise Bug ("Metis_Proof.recontructEquality: shouldn't fail:\n" ^ err); *) fun reconstruct cl (Metis_Thm.Axiom,[]) = Axiom cl | reconstruct cl (Metis_Thm.Assume,[]) = (case Metis_LiteralSet.findl Metis_Literal.positive cl of SOME (_,atm) => Assume atm | NONE => raise Bug "malformed Assume inference") | reconstruct cl (Metis_Thm.Metis_Subst,[th]) = Metis_Subst (reconstructSubst (Metis_Thm.clause th) cl, th) | reconstruct cl (Metis_Thm.Resolve,[th1,th2]) = let val cl1 = Metis_Thm.clause th1 and cl2 = Metis_Thm.clause th2 val (pol,atm) = reconstructResolvant cl1 cl2 cl in if pol then Resolve (atm,th1,th2) else Resolve (atm,th2,th1) end | reconstruct cl (Metis_Thm.Refl,[]) = (case Metis_LiteralSet.findl (K true) cl of SOME lit => Refl (Metis_Literal.destRefl lit) | NONE => raise Bug "malformed Refl inference") | reconstruct cl (Metis_Thm.Equality,[]) = Equality (reconstructEquality cl) | reconstruct _ _ = raise Bug "malformed inference"; in fun thmToInference th = let (*MetisTrace3 val () = Metis_Print.trace Metis_Thm.pp "Metis_Proof.thmToInference: th" th *) val cl = Metis_Thm.clause th val thmInf = Metis_Thm.inference th (*MetisTrace3 val ppThmInf = Metis_Print.ppPair Metis_Thm.ppInferenceType (Metis_Print.ppList Metis_Thm.pp) val () = Metis_Print.trace ppThmInf "Metis_Proof.thmToInference: thmInf" thmInf *) val inf = reconstruct cl thmInf (*MetisTrace3 val () = Metis_Print.trace ppInference "Metis_Proof.thmToInference: inf" inf *) (*MetisDebug val () = let val th' = inferenceToThm inf in if Metis_LiteralSet.equal (Metis_Thm.clause th') cl then () else raise Bug ("Metis_Proof.thmToInference: bad inference reconstruction:" ^ "\n th = " ^ Metis_Thm.toString th ^ "\n inf = " ^ inferenceToString inf ^ "\n inf th = " ^ Metis_Thm.toString th') end *) in inf end (*MetisDebug handle Error err => raise Bug ("Metis_Proof.thmToInference: shouldn't fail:\n" ^ err); *) end; (* ------------------------------------------------------------------------- *) (* Reconstructing whole proofs. *) (* ------------------------------------------------------------------------- *) local val emptyThms : Metis_Thm.thm Metis_LiteralSetMap.map = Metis_LiteralSetMap.new (); fun addThms (th,ths) = let val cl = Metis_Thm.clause th in if Metis_LiteralSetMap.inDomain cl ths then ths else let val (_,pars) = Metis_Thm.inference th val ths = List.foldl addThms ths pars in if Metis_LiteralSetMap.inDomain cl ths then ths else Metis_LiteralSetMap.insert ths (cl,th) end end; fun mkThms th = addThms (th,emptyThms); fun addProof (th,(ths,acc)) = let val cl = Metis_Thm.clause th in case Metis_LiteralSetMap.peek ths cl of NONE => (ths,acc) | SOME th => let val (_,pars) = Metis_Thm.inference th val (ths,acc) = List.foldl addProof (ths,acc) pars val ths = Metis_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 () = Metis_Print.trace Metis_Print.ppInt "Metis_Proof.proof: unnecessary clauses" (Metis_LiteralSetMap.size ths) *) in List.rev acc end; in fun proof th = let (*MetisTrace3 val () = Metis_Print.trace Metis_Thm.pp "Metis_Proof.proof: th" th *) val ths = mkThms th val infs = mkProof ths th (*MetisTrace3 val () = Metis_Print.trace Metis_Print.ppInt "Metis_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) => Metis_LiteralSet.freeIn v lits | (_, Assume atm) => Metis_Atom.freeIn v atm | (th, Metis_Subst _) => Metis_Thm.freeIn v th | (_, Resolve _) => false | (_, Refl tm) => Metis_Term.freeIn v tm | (_, Equality (lit,_,tm)) => Metis_Literal.freeIn v lit orelse Metis_Term.freeIn v tm in List.exists free end; val freeVars = let fun inc (th_inf,set) = Metis_NameSet.union set (case th_inf of (_, Axiom lits) => Metis_LiteralSet.freeVars lits | (_, Assume atm) => Metis_Atom.freeVars atm | (th, Metis_Subst _) => Metis_Thm.freeVars th | (_, Resolve _) => Metis_NameSet.empty | (_, Refl tm) => Metis_Term.freeVars tm | (_, Equality (lit,_,tm)) => Metis_NameSet.union (Metis_Literal.freeVars lit) (Metis_Term.freeVars tm)) in List.foldl inc Metis_NameSet.empty end; end (**** Original file: src/Rule.sig ****) (* ========================================================================= *) (* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Rule = sig (* ------------------------------------------------------------------------- *) (* An equation consists of two terms (t,u) plus a theorem (stronger than) *) (* t = u \/ C. *) (* ------------------------------------------------------------------------- *) type equation = (Metis_Term.term * Metis_Term.term) * Metis_Thm.thm val ppEquation : equation Metis_Print.pp val equationToString : equation -> string (* Returns t = u if the equation theorem contains this literal *) val equationLiteral : equation -> Metis_Literal.literal option val reflEqn : Metis_Term.term -> equation val symEqn : equation -> equation val transEqn : equation -> equation -> equation (* ------------------------------------------------------------------------- *) (* A conversion takes a term t and either: *) (* 1. Returns a term u together with a theorem (stronger than) t = u \/ C. *) (* 2. Raises an Error exception. *) (* ------------------------------------------------------------------------- *) type conv = Metis_Term.term -> Metis_Term.term * Metis_Thm.thm val allConv : conv val noConv : conv val thenConv : conv -> conv -> conv val orelseConv : conv -> conv -> conv val tryConv : conv -> conv val repeatConv : conv -> conv val firstConv : conv list -> conv val everyConv : conv list -> conv val rewrConv : equation -> Metis_Term.path -> conv val pathConv : conv -> Metis_Term.path -> conv val subtermConv : conv -> int -> conv val subtermsConv : conv -> conv (* All function arguments *) (* ------------------------------------------------------------------------- *) (* Applying a conversion to every subterm, with some traversal strategy. *) (* ------------------------------------------------------------------------- *) val bottomUpConv : conv -> conv val topDownConv : conv -> conv val repeatTopDownConv : conv -> conv (* useful for rewriting *) (* ------------------------------------------------------------------------- *) (* A literule (bad pun) takes a literal L and either: *) (* 1. Returns a literal L' with a theorem (stronger than) ~L \/ L' \/ C. *) (* 2. Raises an Error exception. *) (* ------------------------------------------------------------------------- *) type literule = Metis_Literal.literal -> Metis_Literal.literal * Metis_Thm.thm val allLiterule : literule val noLiterule : literule val thenLiterule : literule -> literule -> literule val orelseLiterule : literule -> literule -> literule val tryLiterule : literule -> literule val repeatLiterule : literule -> literule val firstLiterule : literule list -> literule val everyLiterule : literule list -> literule val rewrLiterule : equation -> Metis_Term.path -> literule val pathLiterule : conv -> Metis_Term.path -> literule val argumentLiterule : conv -> int -> literule val allArgumentsLiterule : conv -> literule (* ------------------------------------------------------------------------- *) (* A rule takes one theorem and either deduces another or raises an Error *) (* exception. *) (* ------------------------------------------------------------------------- *) type rule = Metis_Thm.thm -> Metis_Thm.thm val allRule : rule val noRule : rule val thenRule : rule -> rule -> rule val orelseRule : rule -> rule -> rule val tryRule : rule -> rule val changedRule : rule -> rule val repeatRule : rule -> rule val firstRule : rule list -> rule val everyRule : rule list -> rule val literalRule : literule -> Metis_Literal.literal -> rule val rewrRule : equation -> Metis_Literal.literal -> Metis_Term.path -> rule val pathRule : conv -> Metis_Literal.literal -> Metis_Term.path -> rule val literalsRule : literule -> Metis_LiteralSet.set -> rule val allLiteralsRule : literule -> rule val convRule : conv -> rule (* All arguments of all literals *) (* ------------------------------------------------------------------------- *) (* *) (* --------- reflexivity *) (* x = x *) (* ------------------------------------------------------------------------- *) val reflexivityRule : Metis_Term.term -> Metis_Thm.thm val reflexivity : Metis_Thm.thm (* ------------------------------------------------------------------------- *) (* *) (* --------------------- symmetry *) (* ~(x = y) \/ y = x *) (* ------------------------------------------------------------------------- *) val symmetryRule : Metis_Term.term -> Metis_Term.term -> Metis_Thm.thm val symmetry : Metis_Thm.thm (* ------------------------------------------------------------------------- *) (* *) (* --------------------------------- transitivity *) (* ~(x = y) \/ ~(y = z) \/ x = z *) (* ------------------------------------------------------------------------- *) val transitivity : Metis_Thm.thm (* ------------------------------------------------------------------------- *) (* *) (* ---------------------------------------------- functionCongruence (f,n) *) (* ~(x0 = y0) \/ ... \/ ~(x{n-1} = y{n-1}) \/ *) (* f x0 ... x{n-1} = f y0 ... y{n-1} *) (* ------------------------------------------------------------------------- *) val functionCongruence : Metis_Term.function -> Metis_Thm.thm (* ------------------------------------------------------------------------- *) (* *) (* ---------------------------------------------- relationCongruence (R,n) *) (* ~(x0 = y0) \/ ... \/ ~(x{n-1} = y{n-1}) \/ *) (* ~R x0 ... x{n-1} \/ R y0 ... y{n-1} *) (* ------------------------------------------------------------------------- *) val relationCongruence : Metis_Atom.relation -> Metis_Thm.thm (* ------------------------------------------------------------------------- *) (* x = y \/ C *) (* -------------- symEq (x = y) *) (* y = x \/ C *) (* ------------------------------------------------------------------------- *) val symEq : Metis_Literal.literal -> rule (* ------------------------------------------------------------------------- *) (* ~(x = y) \/ C *) (* ----------------- symNeq ~(x = y) *) (* ~(y = x) \/ C *) (* ------------------------------------------------------------------------- *) val symNeq : Metis_Literal.literal -> rule (* ------------------------------------------------------------------------- *) (* sym (x = y) = symEq (x = y) /\ sym ~(x = y) = symNeq ~(x = y) *) (* ------------------------------------------------------------------------- *) val sym : Metis_Literal.literal -> rule (* ------------------------------------------------------------------------- *) (* ~(x = x) \/ C *) (* ----------------- removeIrrefl *) (* C *) (* *) (* where all irreflexive equalities. *) (* ------------------------------------------------------------------------- *) val removeIrrefl : rule (* ------------------------------------------------------------------------- *) (* x = y \/ y = x \/ C *) (* ----------------------- removeSym *) (* x = y \/ C *) (* *) (* where all duplicate copies of equalities and disequalities are removed. *) (* ------------------------------------------------------------------------- *) val removeSym : rule (* ------------------------------------------------------------------------- *) (* ~(v = t) \/ C *) (* ----------------- expandAbbrevs *) (* C[t/v] *) (* *) (* where t must not contain any occurrence of the variable v. *) (* ------------------------------------------------------------------------- *) val expandAbbrevs : rule (* ------------------------------------------------------------------------- *) (* simplify = isTautology + expandAbbrevs + removeSym *) (* ------------------------------------------------------------------------- *) val simplify : Metis_Thm.thm -> Metis_Thm.thm option (* ------------------------------------------------------------------------- *) (* C *) (* -------- freshVars *) (* C[s] *) (* *) (* where s is a renaming substitution chosen so that all of the variables in *) (* C are replaced by fresh variables. *) (* ------------------------------------------------------------------------- *) val freshVars : rule (* ------------------------------------------------------------------------- *) (* C *) (* ---------------------------- factor *) (* C_s_1, C_s_2, ..., C_s_n *) (* *) (* where each s_i is a substitution that factors C, meaning that the theorem *) (* *) (* C_s_i = (removeIrrefl o removeSym o Metis_Thm.subst s_i) C *) (* *) (* has fewer literals than C. *) (* *) (* Also, if s is any substitution that factors C, then one of the s_i will *) (* result in a theorem C_s_i that strictly subsumes the theorem C_s. *) (* ------------------------------------------------------------------------- *) val factor' : Metis_Thm.clause -> Metis_Subst.subst list val factor : Metis_Thm.thm -> Metis_Thm.thm list end (**** Original file: src/Rule.sml ****) (* ========================================================================= *) (* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Rule :> Metis_Rule = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* Variable names. *) (* ------------------------------------------------------------------------- *) val xVarName = Metis_Name.fromString "x"; val xVar = Metis_Term.Var xVarName; val yVarName = Metis_Name.fromString "y"; val yVar = Metis_Term.Var yVarName; val zVarName = Metis_Name.fromString "z"; val zVar = Metis_Term.Var zVarName; fun xIVarName i = Metis_Name.fromString ("x" ^ Int.toString i); fun xIVar i = Metis_Term.Var (xIVarName i); fun yIVarName i = Metis_Name.fromString ("y" ^ Int.toString i); fun yIVar i = Metis_Term.Var (yIVarName i); (* ------------------------------------------------------------------------- *) (* *) (* --------- reflexivity *) (* x = x *) (* ------------------------------------------------------------------------- *) fun reflexivityRule x = Metis_Thm.refl x; val reflexivity = reflexivityRule xVar; (* ------------------------------------------------------------------------- *) (* *) (* --------------------- symmetry *) (* ~(x = y) \/ y = x *) (* ------------------------------------------------------------------------- *) fun symmetryRule x y = let val reflTh = reflexivityRule x val reflLit = Metis_Thm.destUnit reflTh val eqTh = Metis_Thm.equality reflLit [0] y in Metis_Thm.resolve reflLit reflTh eqTh end; val symmetry = symmetryRule xVar yVar; (* ------------------------------------------------------------------------- *) (* *) (* --------------------------------- transitivity *) (* ~(x = y) \/ ~(y = z) \/ x = z *) (* ------------------------------------------------------------------------- *) val transitivity = let val eqTh = Metis_Thm.equality (Metis_Literal.mkEq (yVar,zVar)) [0] xVar in Metis_Thm.resolve (Metis_Literal.mkEq (yVar,xVar)) symmetry eqTh end; (* ------------------------------------------------------------------------- *) (* x = y \/ C *) (* -------------- symEq (x = y) *) (* y = x \/ C *) (* ------------------------------------------------------------------------- *) fun symEq lit th = let val (x,y) = Metis_Literal.destEq lit in if Metis_Term.equal x y then th else let val sub = Metis_Subst.fromList [(xVarName,x),(yVarName,y)] val symTh = Metis_Thm.subst sub symmetry in Metis_Thm.resolve lit th symTh end end; (* ------------------------------------------------------------------------- *) (* An equation consists of two terms (t,u) plus a theorem (stronger than) *) (* t = u \/ C. *) (* ------------------------------------------------------------------------- *) type equation = (Metis_Term.term * Metis_Term.term) * Metis_Thm.thm; fun ppEquation ((_,th) : equation) = Metis_Thm.pp th; val equationToString = Metis_Print.toString ppEquation; fun equationLiteral (t_u,th) = let val lit = Metis_Literal.mkEq t_u in if Metis_LiteralSet.member lit (Metis_Thm.clause th) then SOME lit else NONE end; fun reflEqn t = ((t,t), Metis_Thm.refl t); fun symEqn (eqn as ((t,u), th)) = if Metis_Term.equal t u then eqn else ((u,t), case equationLiteral eqn of SOME t_u => symEq t_u th | NONE => th); fun transEqn (eqn1 as ((x,y), th1)) (eqn2 as ((_,z), th2)) = if Metis_Term.equal x y then eqn2 else if Metis_Term.equal y z then eqn1 else if Metis_Term.equal x z then reflEqn x else ((x,z), case equationLiteral eqn1 of NONE => th1 | SOME x_y => case equationLiteral eqn2 of NONE => th2 | SOME y_z => let val sub = Metis_Subst.fromList [(xVarName,x),(yVarName,y),(zVarName,z)] val th = Metis_Thm.subst sub transitivity val th = Metis_Thm.resolve x_y th1 th val th = Metis_Thm.resolve y_z th2 th in th end); (*MetisDebug val transEqn = fn eqn1 => fn eqn2 => transEqn eqn1 eqn2 handle Error err => raise Error ("Metis_Rule.transEqn:\neqn1 = " ^ equationToString eqn1 ^ "\neqn2 = " ^ equationToString eqn2 ^ "\n" ^ err); *) (* ------------------------------------------------------------------------- *) (* A conversion takes a term t and either: *) (* 1. Returns a term u together with a theorem (stronger than) t = u \/ C. *) (* 2. Raises an Error exception. *) (* ------------------------------------------------------------------------- *) type conv = Metis_Term.term -> Metis_Term.term * Metis_Thm.thm; fun allConv tm = (tm, Metis_Thm.refl tm); val noConv : conv = fn _ => raise Error "noConv"; (*MetisDebug fun traceConv s conv tm = let val res as (tm',th) = conv tm val () = trace (s ^ ": " ^ Metis_Term.toString tm ^ " --> " ^ Metis_Term.toString tm' ^ " " ^ Metis_Thm.toString th ^ "\n") in res end handle Error err => (trace (s ^ ": " ^ Metis_Term.toString tm ^ " --> Error: " ^ err ^ "\n"); raise Error (s ^ ": " ^ err)); *) fun thenConvTrans tm (tm',th1) (tm'',th2) = let val eqn1 = ((tm,tm'),th1) and eqn2 = ((tm',tm''),th2) val (_,th) = transEqn eqn1 eqn2 in (tm'',th) end; fun thenConv conv1 conv2 tm = let val res1 as (tm',_) = conv1 tm val res2 = conv2 tm' in thenConvTrans tm res1 res2 end; fun orelseConv (conv1 : conv) conv2 tm = conv1 tm handle Error _ => conv2 tm; fun tryConv conv = orelseConv conv allConv; fun changedConv conv tm = let val res as (tm',_) = conv tm in if tm = tm' then raise Error "changedConv" else res end; fun repeatConv conv tm = tryConv (thenConv conv (repeatConv conv)) tm; fun firstConv [] _ = raise Error "firstConv" | firstConv [conv] tm = conv tm | firstConv (conv :: convs) tm = orelseConv conv (firstConv convs) tm; fun everyConv [] tm = allConv tm | everyConv [conv] tm = conv tm | everyConv (conv :: convs) tm = thenConv conv (everyConv convs) tm; fun rewrConv (eqn as ((x,y), eqTh)) path tm = if Metis_Term.equal x y then allConv tm else if List.null path then (y,eqTh) else let val reflTh = Metis_Thm.refl tm val reflLit = Metis_Thm.destUnit reflTh val th = Metis_Thm.equality reflLit (1 :: path) y val th = Metis_Thm.resolve reflLit reflTh th val th = case equationLiteral eqn of NONE => th | SOME x_y => Metis_Thm.resolve x_y eqTh th val tm' = Metis_Term.replace tm (path,y) in (tm',th) end; (*MetisDebug val rewrConv = fn eqn as ((x,y),eqTh) => fn path => fn tm => rewrConv eqn path tm handle Error err => raise Error ("Metis_Rule.rewrConv:\nx = " ^ Metis_Term.toString x ^ "\ny = " ^ Metis_Term.toString y ^ "\neqTh = " ^ Metis_Thm.toString eqTh ^ "\npath = " ^ Metis_Term.pathToString path ^ "\ntm = " ^ Metis_Term.toString tm ^ "\n" ^ err); *) fun pathConv conv path tm = let val x = Metis_Term.subterm tm path val (y,th) = conv x in rewrConv ((x,y),th) path tm end; fun subtermConv conv i = pathConv conv [i]; fun subtermsConv _ (tm as Metis_Term.Var _) = allConv tm | subtermsConv conv (tm as Metis_Term.Fn (_,a)) = everyConv (List.map (subtermConv conv) (interval 0 (length a))) tm; (* ------------------------------------------------------------------------- *) (* Applying a conversion to every subterm, with some traversal strategy. *) (* ------------------------------------------------------------------------- *) fun bottomUpConv conv tm = thenConv (subtermsConv (bottomUpConv conv)) (repeatConv conv) tm; fun topDownConv conv tm = thenConv (repeatConv conv) (subtermsConv (topDownConv conv)) tm; fun repeatTopDownConv conv = let fun f tm = thenConv (repeatConv conv) g tm and g tm = thenConv (subtermsConv f) h tm and h tm = tryConv (thenConv conv f) tm in f end; (*MetisDebug val repeatTopDownConv = fn conv => fn tm => repeatTopDownConv conv tm handle Error err => raise Error ("repeatTopDownConv: " ^ err); *) (* ------------------------------------------------------------------------- *) (* A literule (bad pun) takes a literal L and either: *) (* 1. Returns a literal L' with a theorem (stronger than) ~L \/ L' \/ C. *) (* 2. Raises an Error exception. *) (* ------------------------------------------------------------------------- *) type literule = Metis_Literal.literal -> Metis_Literal.literal * Metis_Thm.thm; fun allLiterule lit = (lit, Metis_Thm.assume lit); val noLiterule : literule = fn _ => raise Error "noLiterule"; fun thenLiterule literule1 literule2 lit = let val res1 as (lit',th1) = literule1 lit val res2 as (lit'',th2) = literule2 lit' in if Metis_Literal.equal lit lit' then res2 else if Metis_Literal.equal lit' lit'' then res1 else if Metis_Literal.equal lit lit'' then allLiterule lit else (lit'', if not (Metis_Thm.member lit' th1) then th1 else if not (Metis_Thm.negateMember lit' th2) then th2 else Metis_Thm.resolve lit' th1 th2) end; fun orelseLiterule (literule1 : literule) literule2 lit = literule1 lit handle Error _ => literule2 lit; fun tryLiterule literule = orelseLiterule literule allLiterule; fun changedLiterule literule lit = let val res as (lit',_) = literule lit in if lit = lit' then raise Error "changedLiterule" else res end; fun repeatLiterule literule lit = tryLiterule (thenLiterule literule (repeatLiterule literule)) lit; fun firstLiterule [] _ = raise Error "firstLiterule" | firstLiterule [literule] lit = literule lit | firstLiterule (literule :: literules) lit = orelseLiterule literule (firstLiterule literules) lit; fun everyLiterule [] lit = allLiterule lit | everyLiterule [literule] lit = literule lit | everyLiterule (literule :: literules) lit = thenLiterule literule (everyLiterule literules) lit; fun rewrLiterule (eqn as ((x,y),eqTh)) path lit = if Metis_Term.equal x y then allLiterule lit else let val th = Metis_Thm.equality lit path y val th = case equationLiteral eqn of NONE => th | SOME x_y => Metis_Thm.resolve x_y eqTh th val lit' = Metis_Literal.replace lit (path,y) in (lit',th) end; (*MetisDebug val rewrLiterule = fn eqn => fn path => fn lit => rewrLiterule eqn path lit handle Error err => raise Error ("Metis_Rule.rewrLiterule:\neqn = " ^ equationToString eqn ^ "\npath = " ^ Metis_Term.pathToString path ^ "\nlit = " ^ Metis_Literal.toString lit ^ "\n" ^ err); *) fun pathLiterule conv path lit = let val tm = Metis_Literal.subterm lit path val (tm',th) = conv tm in rewrLiterule ((tm,tm'),th) path lit end; fun argumentLiterule conv i = pathLiterule conv [i]; fun allArgumentsLiterule conv lit = everyLiterule (List.map (argumentLiterule conv) (interval 0 (Metis_Literal.arity lit))) lit; (* ------------------------------------------------------------------------- *) (* A rule takes one theorem and either deduces another or raises an Error *) (* exception. *) (* ------------------------------------------------------------------------- *) type rule = Metis_Thm.thm -> Metis_Thm.thm; val allRule : rule = fn th => th; val noRule : rule = fn _ => raise Error "noRule"; fun thenRule (rule1 : rule) (rule2 : rule) th = rule1 (rule2 th); fun orelseRule (rule1 : rule) rule2 th = rule1 th handle Error _ => rule2 th; fun tryRule rule = orelseRule rule allRule; fun changedRule rule th = let val th' = rule th in if not (Metis_LiteralSet.equal (Metis_Thm.clause th) (Metis_Thm.clause th')) then th' else raise Error "changedRule" end; fun repeatRule rule lit = tryRule (thenRule rule (repeatRule rule)) lit; fun firstRule [] _ = raise Error "firstRule" | firstRule [rule] th = rule th | firstRule (rule :: rules) th = orelseRule rule (firstRule rules) th; fun everyRule [] th = allRule th | everyRule [rule] th = rule th | everyRule (rule :: rules) th = thenRule rule (everyRule rules) th; fun literalRule literule lit th = let val (lit',litTh) = literule lit in if Metis_Literal.equal lit lit' then th else if not (Metis_Thm.negateMember lit litTh) then litTh else Metis_Thm.resolve lit th litTh end; (*MetisDebug val literalRule = fn literule => fn lit => fn th => literalRule literule lit th handle Error err => raise Error ("Metis_Rule.literalRule:\nlit = " ^ Metis_Literal.toString lit ^ "\nth = " ^ Metis_Thm.toString th ^ "\n" ^ err); *) fun rewrRule eqTh lit path = literalRule (rewrLiterule eqTh path) lit; fun pathRule conv lit path = literalRule (pathLiterule conv path) lit; fun literalsRule literule = let fun f (lit,th) = if Metis_Thm.member lit th then literalRule literule lit th else th in fn lits => fn th => Metis_LiteralSet.foldl f th lits end; fun allLiteralsRule literule th = literalsRule literule (Metis_Thm.clause th) th; fun convRule conv = allLiteralsRule (allArgumentsLiterule conv); (* ------------------------------------------------------------------------- *) (* *) (* ---------------------------------------------- functionCongruence (f,n) *) (* ~(x0 = y0) \/ ... \/ ~(x{n-1} = y{n-1}) \/ *) (* f x0 ... x{n-1} = f y0 ... y{n-1} *) (* ------------------------------------------------------------------------- *) fun functionCongruence (f,n) = let val xs = List.tabulate (n,xIVar) and ys = List.tabulate (n,yIVar) fun cong ((i,yi),(th,lit)) = let val path = [1,i] val th = Metis_Thm.resolve lit th (Metis_Thm.equality lit path yi) val lit = Metis_Literal.replace lit (path,yi) in (th,lit) end val reflTh = Metis_Thm.refl (Metis_Term.Fn (f,xs)) val reflLit = Metis_Thm.destUnit reflTh in fst (List.foldl cong (reflTh,reflLit) (enumerate ys)) end; (* ------------------------------------------------------------------------- *) (* *) (* ---------------------------------------------- relationCongruence (R,n) *) (* ~(x0 = y0) \/ ... \/ ~(x{n-1} = y{n-1}) \/ *) (* ~R x0 ... x{n-1} \/ R y0 ... y{n-1} *) (* ------------------------------------------------------------------------- *) fun relationCongruence (R,n) = let val xs = List.tabulate (n,xIVar) and ys = List.tabulate (n,yIVar) fun cong ((i,yi),(th,lit)) = let val path = [i] val th = Metis_Thm.resolve lit th (Metis_Thm.equality lit path yi) val lit = Metis_Literal.replace lit (path,yi) in (th,lit) end val assumeLit = (false,(R,xs)) val assumeTh = Metis_Thm.assume assumeLit in fst (List.foldl cong (assumeTh,assumeLit) (enumerate ys)) end; (* ------------------------------------------------------------------------- *) (* ~(x = y) \/ C *) (* ----------------- symNeq ~(x = y) *) (* ~(y = x) \/ C *) (* ------------------------------------------------------------------------- *) fun symNeq lit th = let val (x,y) = Metis_Literal.destNeq lit in if Metis_Term.equal x y then th else let val sub = Metis_Subst.fromList [(xVarName,y),(yVarName,x)] val symTh = Metis_Thm.subst sub symmetry in Metis_Thm.resolve lit th symTh end end; (* ------------------------------------------------------------------------- *) (* sym (x = y) = symEq (x = y) /\ sym ~(x = y) = symNeq ~(x = y) *) (* ------------------------------------------------------------------------- *) fun sym (lit as (pol,_)) th = if pol then symEq lit th else symNeq lit th; (* ------------------------------------------------------------------------- *) (* ~(x = x) \/ C *) (* ----------------- removeIrrefl *) (* C *) (* *) (* where all irreflexive equalities. *) (* ------------------------------------------------------------------------- *) local fun irrefl ((true,_),th) = th | irrefl (lit as (false,atm), th) = case total Metis_Atom.destRefl atm of SOME x => Metis_Thm.resolve lit th (Metis_Thm.refl x) | NONE => th; in fun removeIrrefl th = Metis_LiteralSet.foldl irrefl th (Metis_Thm.clause th); end; (* ------------------------------------------------------------------------- *) (* x = y \/ y = x \/ C *) (* ----------------------- removeSym *) (* x = y \/ C *) (* *) (* where all duplicate copies of equalities and disequalities are removed. *) (* ------------------------------------------------------------------------- *) local fun rem (lit as (pol,atm), eqs_th as (eqs,th)) = case total Metis_Atom.sym atm of NONE => eqs_th | SOME atm' => if Metis_LiteralSet.member lit eqs then (eqs, if pol then symEq lit th else symNeq lit th) else (Metis_LiteralSet.add eqs (pol,atm'), th); in fun removeSym th = snd (Metis_LiteralSet.foldl rem (Metis_LiteralSet.empty,th) (Metis_Thm.clause th)); end; (* ------------------------------------------------------------------------- *) (* ~(v = t) \/ C *) (* ----------------- expandAbbrevs *) (* C[t/v] *) (* *) (* where t must not contain any occurrence of the variable v. *) (* ------------------------------------------------------------------------- *) local fun expand lit = let val (x,y) = Metis_Literal.destNeq lit val _ = Metis_Term.isTypedVar x orelse Metis_Term.isTypedVar y orelse raise Error "Metis_Rule.expandAbbrevs: no vars" val _ = not (Metis_Term.equal x y) orelse raise Error "Metis_Rule.expandAbbrevs: equal vars" in Metis_Subst.unify Metis_Subst.empty x y end; in fun expandAbbrevs th = case Metis_LiteralSet.firstl (total expand) (Metis_Thm.clause th) of NONE => removeIrrefl th | SOME sub => expandAbbrevs (Metis_Thm.subst sub th); end; (* ------------------------------------------------------------------------- *) (* simplify = isTautology + expandAbbrevs + removeSym *) (* ------------------------------------------------------------------------- *) fun simplify th = if Metis_Thm.isTautology th then NONE else let val th' = th val th' = expandAbbrevs th' val th' = removeSym th' in if Metis_Thm.equal th th' then SOME th else simplify th' end; (* ------------------------------------------------------------------------- *) (* C *) (* -------- freshVars *) (* C[s] *) (* *) (* where s is a renaming substitution chosen so that all of the variables in *) (* C are replaced by fresh variables. *) (* ------------------------------------------------------------------------- *) fun freshVars th = Metis_Thm.subst (Metis_Subst.freshVars (Metis_Thm.freeVars th)) th; (* ------------------------------------------------------------------------- *) (* C *) (* ---------------------------- factor *) (* C_s_1, C_s_2, ..., C_s_n *) (* *) (* where each s_i is a substitution that factors C, meaning that the theorem *) (* *) (* C_s_i = (removeIrrefl o removeSym o Metis_Thm.subst s_i) C *) (* *) (* has fewer literals than C. *) (* *) (* Also, if s is any substitution that factors C, then one of the s_i will *) (* result in a theorem C_s_i that strictly subsumes the theorem C_s. *) (* ------------------------------------------------------------------------- *) local datatype edge = FactorEdge of Metis_Atom.atom * Metis_Atom.atom | ReflEdge of Metis_Term.term * Metis_Term.term; fun ppEdge (FactorEdge atm_atm') = Metis_Print.ppPair Metis_Atom.pp Metis_Atom.pp atm_atm' | ppEdge (ReflEdge tm_tm') = Metis_Print.ppPair Metis_Term.pp Metis_Term.pp tm_tm'; datatype joinStatus = Joined | Joinable of Metis_Subst.subst | Apart; fun joinEdge sub edge = let val result = case edge of FactorEdge (atm,atm') => total (Metis_Atom.unify sub atm) atm' | ReflEdge (tm,tm') => total (Metis_Subst.unify sub tm) tm' in case result of NONE => Apart | SOME sub' => if Metis_Portable.pointerEqual (sub,sub') then Joined else Joinable sub' end; fun updateApart sub = let fun update acc [] = SOME acc | update acc (edge :: edges) = case joinEdge sub edge of Joined => NONE | Joinable _ => update (edge :: acc) edges | Apart => update acc edges in update [] end; fun addFactorEdge (pol,atm) ((pol',atm'),acc) = if pol <> pol' then acc else let val edge = FactorEdge (atm,atm') in case joinEdge Metis_Subst.empty edge of Joined => raise Bug "addFactorEdge: joined" | Joinable sub => (sub,edge) :: acc | Apart => acc end; fun addReflEdge (false,_) acc = acc | addReflEdge (true,atm) acc = let val edge = ReflEdge (Metis_Atom.destEq atm) in case joinEdge Metis_Subst.empty edge of Joined => raise Bug "addRefl: joined" | Joinable _ => edge :: acc | Apart => acc end; fun addIrreflEdge (true,_) acc = acc | addIrreflEdge (false,atm) acc = let val edge = ReflEdge (Metis_Atom.destEq atm) in case joinEdge Metis_Subst.empty edge of Joined => raise Bug "addRefl: joined" | Joinable sub => (sub,edge) :: acc | Apart => acc end; fun init_edges acc _ [] = let fun init ((apart,sub,edge),(edges,acc)) = (edge :: edges, (apart,sub,edges) :: acc) in snd (List.foldl init ([],[]) acc) end | init_edges acc apart ((sub,edge) :: sub_edges) = let (*MetisDebug val () = if not (Metis_Subst.null sub) then () else raise Bug "Metis_Rule.factor.init_edges: empty subst" *) val (acc,apart) = case updateApart sub apart of SOME apart' => ((apart',sub,edge) :: acc, edge :: apart) | NONE => (acc,apart) in init_edges acc apart sub_edges end; fun mk_edges apart sub_edges [] = init_edges [] apart sub_edges | mk_edges apart sub_edges (lit :: lits) = let val sub_edges = List.foldl (addFactorEdge lit) sub_edges lits val (apart,sub_edges) = case total Metis_Literal.sym lit of NONE => (apart,sub_edges) | SOME lit' => let val apart = addReflEdge lit apart val sub_edges = addIrreflEdge lit sub_edges val sub_edges = List.foldl (addFactorEdge lit') sub_edges lits in (apart,sub_edges) end in mk_edges apart sub_edges lits end; fun fact acc [] = acc | fact acc ((_,sub,[]) :: others) = fact (sub :: acc) others | fact acc ((apart, sub, edge :: edges) :: others) = let val others = case joinEdge sub edge of Joinable sub' => let val others = (edge :: apart, sub, edges) :: others in case updateApart sub' apart of NONE => others | SOME apart' => (apart',sub',edges) :: others end | _ => (apart,sub,edges) :: others in fact acc others end; in fun factor' cl = let (*MetisTrace6 val () = Metis_Print.trace Metis_LiteralSet.pp "Metis_Rule.factor': cl" cl *) val edges = mk_edges [] [] (Metis_LiteralSet.toList cl) (*MetisTrace6 val ppEdgesSize = Metis_Print.ppMap length Metis_Print.ppInt val ppEdgel = Metis_Print.ppList ppEdge val ppEdges = Metis_Print.ppList (Metis_Print.ppTriple ppEdgel Metis_Subst.pp ppEdgel) val () = Metis_Print.trace ppEdgesSize "Metis_Rule.factor': |edges|" edges val () = Metis_Print.trace ppEdges "Metis_Rule.factor': edges" edges *) val result = fact [] edges (*MetisTrace6 val ppResult = Metis_Print.ppList Metis_Subst.pp val () = Metis_Print.trace ppResult "Metis_Rule.factor': result" result *) in result end; end; fun factor th = let fun fact sub = removeIrrefl (removeSym (Metis_Thm.subst sub th)) in List.map fact (factor' (Metis_Thm.clause th)) end; end (**** Original file: src/Normalize.sig ****) (* ========================================================================= *) (* NORMALIZING FORMULAS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Normalize = sig (* ------------------------------------------------------------------------- *) (* Negation normal form. *) (* ------------------------------------------------------------------------- *) 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. *) (* ------------------------------------------------------------------------- *) 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 (**** Original file: src/Normalize.sml ****) (* ========================================================================= *) (* NORMALIZING FORMULAS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Normalize :> Metis_Normalize = struct open Metis_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. *) (* ------------------------------------------------------------------------- *) 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 = 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 = addLogReal p1 p2 and n = multiplyLogReal n1 n2 in Count {positive = p, negative = n} end; fun countOr2 (count1,count2) = let val Count {positive = p1, negative = n1} = count1 and Count {positive = p2, negative = n2} = count2 val p = multiplyLogReal p1 p2 and n = addLogReal n1 n2 in Count {positive = p, negative = n} end; (* 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 = 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 = Metis_Print.ppMap countToString Metis_Print.ppString; (* ------------------------------------------------------------------------- *) (* A type of normalized formula. *) (* ------------------------------------------------------------------------- *) datatype formula = True | False | Metis_Literal of Metis_NameSet.set * Metis_Literal.literal | And of Metis_NameSet.set * count * formula Metis_Set.set | Or of Metis_NameSet.set * count * formula Metis_Set.set | Xor of Metis_NameSet.set * count * bool * formula Metis_Set.set | Exists of Metis_NameSet.set * count * Metis_NameSet.set * formula | Forall of Metis_NameSet.set * count * Metis_NameSet.set * formula; fun compare f1_f2 = if Metis_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 | (Metis_Literal (_,l1), Metis_Literal (_,l2)) => Metis_Literal.compare (l1,l2) | (Metis_Literal _, _) => LESS | (_, Metis_Literal _) => GREATER | (And (_,_,s1), And (_,_,s2)) => Metis_Set.compare (s1,s2) | (And _, _) => LESS | (_, And _) => GREATER | (Or (_,_,s1), Or (_,_,s2)) => Metis_Set.compare (s1,s2) | (Or _, _) => LESS | (_, Or _) => GREATER | (Xor (_,_,p1,s1), Xor (_,_,p2,s2)) => (case boolCompare (p1,p2) of LESS => LESS | EQUAL => Metis_Set.compare (s1,s2) | GREATER => GREATER) | (Xor _, _) => LESS | (_, Xor _) => GREATER | (Exists (_,_,n1,f1), Exists (_,_,n2,f2)) => (case Metis_NameSet.compare (n1,n2) of LESS => LESS | EQUAL => compare (f1,f2) | GREATER => GREATER) | (Exists _, _) => LESS | (_, Exists _) => GREATER | (Forall (_,_,n1,f1), Forall (_,_,n2,f2)) => (case Metis_NameSet.compare (n1,n2) of LESS => LESS | EQUAL => compare (f1,f2) | GREATER => GREATER); val empty = Metis_Set.empty compare; val singleton = Metis_Set.singleton compare; local fun neg True = False | neg False = True | neg (Metis_Literal (fv,lit)) = Metis_Literal (fv, Metis_Literal.negate lit) | neg (And (fv,c,s)) = Or (fv, countNegate c, neg_set s) | neg (Or (fv,c,s)) = And (fv, countNegate c, neg_set s) | neg (Xor (fv,c,p,s)) = Xor (fv, c, not p, s) | neg (Exists (fv,c,n,f)) = Forall (fv, countNegate c, n, neg f) | neg (Forall (fv,c,n,f)) = Exists (fv, countNegate c, n, neg f) and neg_set s = Metis_Set.foldl neg_elt empty s and neg_elt (f,s) = Metis_Set.add s (neg f); in val negate = neg; val negateSet = neg_set; end; fun negateMember x s = Metis_Set.member (negate x) s; local fun member s x = negateMember x s; in fun negateDisjoint s1 s2 = if Metis_Set.size s1 < Metis_Set.size s2 then not (Metis_Set.exists (member s2) s1) else not (Metis_Set.exists (member s1) s2); end; fun polarity True = true | polarity False = false | polarity (Metis_Literal (_,(pol,_))) = not pol | polarity (And _) = true | polarity (Or _) = false | polarity (Xor (_,_,pol,_)) = pol | polarity (Exists _) = true | polarity (Forall _) = false; (*MetisDebug val polarity = fn f => let val res1 = compare (f, negate f) = LESS val res2 = polarity f val _ = res1 = res2 orelse raise Bug "polarity" in res2 end; *) fun applyPolarity true fm = fm | applyPolarity false fm = negate fm; fun freeVars True = Metis_NameSet.empty | freeVars False = Metis_NameSet.empty | freeVars (Metis_Literal (fv,_)) = fv | freeVars (And (fv,_,_)) = fv | freeVars (Or (fv,_,_)) = fv | freeVars (Xor (fv,_,_,_)) = fv | freeVars (Exists (fv,_,_,_)) = fv | freeVars (Forall (fv,_,_,_)) = fv; fun freeIn v fm = Metis_NameSet.member v (freeVars fm); val freeVarsSet = let fun free (fm,acc) = Metis_NameSet.union (freeVars fm) acc in Metis_Set.foldl free Metis_NameSet.empty end; fun count True = countTrue | count False = countFalse | count (Metis_Literal _) = countLiteral | count (And (_,c,_)) = c | count (Or (_,c,_)) = c | count (Xor (_,c,p,_)) = if p then c else countNegate c | count (Exists (_,c,_,_)) = c | count (Forall (_,c,_,_)) = c; val countAndSet = let fun countAnd (fm,c) = countAnd2 (count fm, c) in Metis_Set.foldl countAnd countTrue end; val countOrSet = let fun countOr (fm,c) = countOr2 (count fm, c) in Metis_Set.foldl countOr countFalse end; val countXorSet = let fun countXor (fm,c) = countXor2 (count fm, c) in Metis_Set.foldl countXor countFalse end; fun And2 (False,_) = False | And2 (_,False) = False | And2 (True,f2) = f2 | And2 (f1,True) = f1 | And2 (f1,f2) = let val (fv1,c1,s1) = case f1 of And fv_c_s => fv_c_s | _ => (freeVars f1, count f1, singleton f1) and (fv2,c2,s2) = case f2 of And fv_c_s => fv_c_s | _ => (freeVars f2, count f2, singleton f2) in if not (negateDisjoint s1 s2) then False else let val s = Metis_Set.union s1 s2 in case Metis_Set.size s of 0 => True | 1 => Metis_Set.pick s | n => if n = Metis_Set.size s1 + Metis_Set.size s2 then And (Metis_NameSet.union fv1 fv2, countAnd2 (c1,c2), s) else And (freeVarsSet s, countAndSet s, s) end end; val AndList = List.foldl And2 True; val AndSet = Metis_Set.foldl And2 True; fun Or2 (True,_) = True | Or2 (_,True) = True | Or2 (False,f2) = f2 | Or2 (f1,False) = f1 | Or2 (f1,f2) = let val (fv1,c1,s1) = case f1 of Or fv_c_s => fv_c_s | _ => (freeVars f1, count f1, singleton f1) and (fv2,c2,s2) = case f2 of Or fv_c_s => fv_c_s | _ => (freeVars f2, count f2, singleton f2) in if not (negateDisjoint s1 s2) then True else let val s = Metis_Set.union s1 s2 in case Metis_Set.size s of 0 => False | 1 => Metis_Set.pick s | n => if n = Metis_Set.size s1 + Metis_Set.size s2 then Or (Metis_NameSet.union fv1 fv2, countOr2 (c1,c2), s) else Or (freeVarsSet s, countOrSet s, s) end end; val OrList = List.foldl Or2 False; val OrSet = Metis_Set.foldl Or2 False; fun pushOr2 (f1,f2) = let val s1 = case f1 of And (_,_,s) => s | _ => singleton f1 and s2 = case f2 of And (_,_,s) => s | _ => singleton f2 fun g x1 (x2,acc) = And2 (Or2 (x1,x2), acc) fun f (x1,acc) = Metis_Set.foldl (g x1) acc s2 in Metis_Set.foldl f True s1 end; val pushOrList = List.foldl pushOr2 False; local fun normalize fm = let val p = polarity fm val fm = applyPolarity p fm in (freeVars fm, count fm, p, singleton fm) end; in fun Xor2 (False,f2) = f2 | Xor2 (f1,False) = f1 | Xor2 (True,f2) = negate f2 | Xor2 (f1,True) = negate f1 | Xor2 (f1,f2) = let val (fv1,c1,p1,s1) = case f1 of Xor x => x | _ => normalize f1 and (fv2,c2,p2,s2) = case f2 of Xor x => x | _ => normalize f2 val s = Metis_Set.symmetricDifference s1 s2 val fm = case Metis_Set.size s of 0 => False | 1 => Metis_Set.pick s | n => if n = Metis_Set.size s1 + Metis_Set.size s2 then Xor (Metis_NameSet.union fv1 fv2, countXor2 (c1,c2), true, s) else Xor (freeVarsSet s, countXorSet s, true, s) val p = p1 = p2 in applyPolarity p fm end; end; val XorList = List.foldl Xor2 False; val XorSet = Metis_Set.foldl Xor2 False; fun XorPolarityList (p,l) = applyPolarity p (XorList l); fun XorPolaritySet (p,s) = applyPolarity p (XorSet s); fun destXor (Xor (_,_,p,s)) = let val (fm1,s) = Metis_Set.deletePick s val fm2 = if Metis_Set.size s = 1 then applyPolarity p (Metis_Set.pick s) else Xor (freeVarsSet s, countXorSet s, p, s) in (fm1,fm2) end | destXor _ = raise Error "destXor"; fun pushXor fm = let val (f1,f2) = destXor fm val f1' = negate f1 and f2' = negate f2 in And2 (Or2 (f1,f2), Or2 (f1',f2')) end; fun Exists1 (v,init_fm) = let fun exists_gen fm = let val fv = Metis_NameSet.delete (freeVars fm) v val c = count fm val n = Metis_NameSet.singleton v in Exists (fv,c,n,fm) end fun exists fm = if freeIn v fm then exists_free fm else fm and exists_free (Or (_,_,s)) = OrList (Metis_Set.transform exists s) | exists_free (fm as And (_,_,s)) = let val sv = Metis_Set.filter (freeIn v) s in if Metis_Set.size sv <> 1 then exists_gen fm else let val fm = Metis_Set.pick sv val s = Metis_Set.delete s fm in And2 (exists_free fm, AndSet s) end end | exists_free (Exists (fv,c,n,f)) = Exists (Metis_NameSet.delete fv v, c, Metis_NameSet.add n v, f) | exists_free fm = exists_gen fm in exists init_fm end; fun ExistsList (vs,f) = List.foldl Exists1 f vs; fun ExistsSet (n,f) = Metis_NameSet.foldl Exists1 f n; fun Forall1 (v,init_fm) = let fun forall_gen fm = let val fv = Metis_NameSet.delete (freeVars fm) v val c = count fm val n = Metis_NameSet.singleton v in Forall (fv,c,n,fm) end fun forall fm = if freeIn v fm then forall_free fm else fm and forall_free (And (_,_,s)) = AndList (Metis_Set.transform forall s) | forall_free (fm as Or (_,_,s)) = let val sv = Metis_Set.filter (freeIn v) s in if Metis_Set.size sv <> 1 then forall_gen fm else let val fm = Metis_Set.pick sv val s = Metis_Set.delete s fm in Or2 (forall_free fm, OrSet s) end end | forall_free (Forall (fv,c,n,f)) = Forall (Metis_NameSet.delete fv v, c, Metis_NameSet.add n v, f) | forall_free fm = forall_gen fm in forall init_fm end; fun ForallList (vs,f) = List.foldl Forall1 f vs; fun ForallSet (n,f) = Metis_NameSet.foldl Forall1 f n; fun generalize f = ForallSet (freeVars f, f); local fun subst_fv fvSub = let fun add_fv (v,s) = Metis_NameSet.union (Metis_NameMap.get fvSub v) s in Metis_NameSet.foldl add_fv Metis_NameSet.empty end; fun subst_rename (v,(avoid,bv,sub,domain,fvSub)) = let val v' = Metis_Term.variantPrime avoid v val avoid = Metis_NameSet.add avoid v' val bv = Metis_NameSet.add bv v' val sub = Metis_Subst.insert sub (v, Metis_Term.Var v') val domain = Metis_NameSet.add domain v val fvSub = Metis_NameMap.insert fvSub (v, Metis_NameSet.singleton v') in (avoid,bv,sub,domain,fvSub) end; fun subst_check sub domain fvSub fm = let val domain = Metis_NameSet.intersect domain (freeVars fm) in if Metis_NameSet.null domain then fm else subst_domain sub domain fvSub fm end and subst_domain sub domain fvSub fm = case fm of Metis_Literal (fv,lit) => let val fv = Metis_NameSet.difference fv domain val fv = Metis_NameSet.union fv (subst_fv fvSub domain) val lit = Metis_Literal.subst sub lit in Metis_Literal (fv,lit) end | And (_,_,s) => AndList (Metis_Set.transform (subst_check sub domain fvSub) s) | Or (_,_,s) => OrList (Metis_Set.transform (subst_check sub domain fvSub) s) | Xor (_,_,p,s) => XorPolarityList (p, Metis_Set.transform (subst_check sub domain fvSub) s) | Exists fv_c_n_f => subst_quant Exists sub domain fvSub fv_c_n_f | Forall fv_c_n_f => subst_quant Forall sub domain fvSub fv_c_n_f | _ => raise Bug "subst_domain" and subst_quant quant sub domain fvSub (fv,c,bv,fm) = let val sub_fv = subst_fv fvSub domain val fv = Metis_NameSet.union sub_fv (Metis_NameSet.difference fv domain) val captured = Metis_NameSet.intersect bv sub_fv val bv = Metis_NameSet.difference bv captured val avoid = Metis_NameSet.union fv bv val (_,bv,sub,domain,fvSub) = Metis_NameSet.foldl subst_rename (avoid,bv,sub,domain,fvSub) captured val fm = subst_domain sub domain fvSub fm in quant (fv,c,bv,fm) end; in fun subst sub = let fun mk_dom (v,tm,(d,fv)) = (Metis_NameSet.add d v, Metis_NameMap.insert fv (v, Metis_Term.freeVars tm)) val domain_fvSub = (Metis_NameSet.empty, Metis_NameMap.new ()) val (domain,fvSub) = Metis_Subst.foldl mk_dom domain_fvSub sub in subst_check sub domain fvSub end; end; fun fromFormula fm = case fm of Metis_Formula.True => True | Metis_Formula.False => False | Metis_Formula.Metis_Atom atm => Metis_Literal (Metis_Atom.freeVars atm, (true,atm)) | Metis_Formula.Not p => negateFromFormula p | Metis_Formula.And (p,q) => And2 (fromFormula p, fromFormula q) | Metis_Formula.Or (p,q) => Or2 (fromFormula p, fromFormula q) | Metis_Formula.Imp (p,q) => Or2 (negateFromFormula p, fromFormula q) | Metis_Formula.Iff (p,q) => Xor2 (negateFromFormula p, fromFormula q) | Metis_Formula.Forall (v,p) => Forall1 (v, fromFormula p) | Metis_Formula.Exists (v,p) => Exists1 (v, fromFormula p) and negateFromFormula fm = case fm of Metis_Formula.True => False | Metis_Formula.False => True | Metis_Formula.Metis_Atom atm => Metis_Literal (Metis_Atom.freeVars atm, (false,atm)) | Metis_Formula.Not p => fromFormula p | Metis_Formula.And (p,q) => Or2 (negateFromFormula p, negateFromFormula q) | Metis_Formula.Or (p,q) => And2 (negateFromFormula p, negateFromFormula q) | Metis_Formula.Imp (p,q) => And2 (fromFormula p, negateFromFormula q) | Metis_Formula.Iff (p,q) => Xor2 (fromFormula p, fromFormula q) | Metis_Formula.Forall (v,p) => Exists1 (v, negateFromFormula p) | Metis_Formula.Exists (v,p) => Forall1 (v, negateFromFormula p); local fun lastElt (s : formula Metis_Set.set) = case Metis_Set.findr (K true) s of NONE => raise Bug "lastElt: empty set" | SOME fm => fm; fun negateLastElt s = let val fm = lastElt s in Metis_Set.add (Metis_Set.delete s fm) (negate fm) end; fun form fm = case fm of True => Metis_Formula.True | False => Metis_Formula.False | Metis_Literal (_,lit) => Metis_Literal.toFormula lit | And (_,_,s) => Metis_Formula.listMkConj (Metis_Set.transform form s) | Or (_,_,s) => Metis_Formula.listMkDisj (Metis_Set.transform form s) | Xor (_,_,p,s) => xorForm p s | Exists (_,_,n,f) => Metis_Formula.listMkExists (Metis_NameSet.toList n, form f) | Forall (_,_,n,f) => Metis_Formula.listMkForall (Metis_NameSet.toList n, form f) (* To convert a Xor set to an Iff list we need to know *) (* whether the size of the set is even or odd: *) (* *) (* b XOR a = b <=> ~a *) (* c XOR b XOR a = c <=> b <=> a *) (* d XOR c XOR b XOR a = d <=> c <=> b <=> ~a *) (* e XOR d XOR c XOR b XOR a = e <=> d <=> c <=> b <=> a *) and xorForm p s = let val p = if Metis_Set.size s mod 2 = 0 then not p else p val s = if p then s else negateLastElt s in Metis_Formula.listMkEquiv (Metis_Set.transform form s) end; in val toFormula = form; end; fun toLiteral (Metis_Literal (_,lit)) = lit | toLiteral _ = raise Error "Metis_Normalize.toLiteral"; local fun addLiteral (l,s) = Metis_LiteralSet.add s (toLiteral l); in fun toClause False = Metis_LiteralSet.empty | toClause (Or (_,_,s)) = Metis_Set.foldl addLiteral Metis_LiteralSet.empty s | toClause l = Metis_LiteralSet.singleton (toLiteral l); end; val pp = Metis_Print.ppMap toFormula Metis_Formula.pp; val toString = Metis_Print.toString pp; (* ------------------------------------------------------------------------- *) (* Negation normal form. *) (* ------------------------------------------------------------------------- *) fun nnf fm = toFormula (fromFormula fm); (* ------------------------------------------------------------------------- *) (* Basic conjunctive normal form. *) (* ------------------------------------------------------------------------- *) local val counter : int Metis_StringMap.map Unsynchronized.ref = Unsynchronized.ref (Metis_StringMap.new ()); fun new n () = let val Unsynchronized.ref m = counter val s = Metis_Name.toString n val i = Option.getOpt (Metis_StringMap.peek m s, 0) val () = counter := Metis_StringMap.insert m (s, i + 1) val i = if i = 0 then "" else "_" ^ Int.toString i val s = skolemPrefix ^ "_" ^ s ^ i in Metis_Name.fromString s end; in fun newSkolemFunction n = Metis_Portable.critical (new n) (); end; fun skolemize fv bv fm = let val fv = Metis_NameSet.transform Metis_Term.Var fv fun mk (v,s) = Metis_Subst.insert s (v, Metis_Term.Fn (newSkolemFunction v, fv)) in subst (Metis_NameSet.foldl mk Metis_Subst.empty bv) fm end; local fun rename avoid fv bv fm = let val captured = Metis_NameSet.intersect avoid bv in if Metis_NameSet.null captured then fm else let fun ren (v,(a,s)) = let val v' = Metis_Term.variantPrime a v in (Metis_NameSet.add a v', Metis_Subst.insert s (v, Metis_Term.Var v')) end val avoid = Metis_NameSet.union (Metis_NameSet.union avoid fv) bv val (_,sub) = Metis_NameSet.foldl ren (avoid,Metis_Subst.empty) captured in subst sub fm end end; fun cnfFm avoid fm = (*MetisTrace5 let val fm' = cnfFm' avoid fm val () = Metis_Print.trace pp "Metis_Normalize.cnfFm: fm" fm val () = Metis_Print.trace pp "Metis_Normalize.cnfFm: fm'" fm' in fm' end and cnfFm' avoid fm = *) case fm of True => True | False => False | Metis_Literal _ => fm | And (_,_,s) => AndList (Metis_Set.transform (cnfFm avoid) s) | Or (fv,_,s) => let val avoid = Metis_NameSet.union avoid fv val (fms,_) = Metis_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 = Metis_NameSet.union avoid (freeVars fm) in (fms,avoid) end; in val basicCnf = cnfFm Metis_NameSet.empty; end; (* ------------------------------------------------------------------------- *) (* Finding the formula definition that minimizes the number of clauses. *) (* ------------------------------------------------------------------------- *) local type best = count * formula option; fun minBreak countClauses fm best = case fm of True => best | False => best | Metis_Literal _ => best | And (_,_,s) => minBreakSet countClauses countAnd2 countTrue AndSet s best | Or (_,_,s) => minBreakSet countClauses countOr2 countFalse OrSet s best | Xor (_,_,_,s) => 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 = let fun fwd (fm,(c1,s1,l)) = let val c1' = count2 (count fm, c1) and s1' = Metis_Set.add s1 fm in (c1', s1', (c1,s1,fm) :: l) end fun bwd ((c1,s1,fm),(c2,s2,l)) = let val c2' = count2 (count fm, c2) and s2' = Metis_Set.add s2 fm in (c2', s2', (c1,s1,fm,c2,s2) :: l) end val (c1,_,fms) = List.foldl fwd (count0,empty,[]) fms val (c2,_,fms) = List.foldl bwd (count0,empty,[]) fms (*MetisDebug val _ = countEquivish c1 c2 orelse raise Bug ("cumulativeCounts: c1 = " ^ countToString c1 ^ ", c2 = " ^ countToString c2) *) in fms end fun breakSing ((c1,_,fm,c2,_),best) = let val cFms = count2 (c1,c2) fun countCls cFm = countClauses (count2 (cFms,cFm)) in minBreak countCls fm best end val breakSet1 = let fun break c1 s1 fm c2 (best as (bcl,_)) = if Metis_Set.null s1 then best else let val cDef = countDefinition (countXor2 (c1, count fm)) val cFm = count2 (countLiteral,c2) val cl = countAnd2 (cDef, countClauses cFm) val noBetter = countLeqish bcl cl in if noBetter then best else (cl, SOME (mkSet (Metis_Set.add s1 fm))) end in fn ((c1,s1,fm,c2,s2),best) => break c1 s1 fm c2 (break c2 s2 fm c1 best) end val fms = Metis_Set.toList fmSet fun breakSet measure best = let val fms = sortMap (measure o count) countCompare fms in List.foldl breakSet1 best (cumulatives fms) end val best = List.foldl breakSing best (cumulatives fms) val best = breakSet I best val best = breakSet countNegate best val best = breakSet countClauses best in best end in fun minimumDefinition fm = let val cl = count fm in if countLeqish cl countLiteral then NONE else let val (cl',def) = minBreak I fm (cl,NONE) (*MetisTrace1 val () = case def of NONE => () | SOME d => Metis_Print.trace pp ("defCNF: before = " ^ countToString cl ^ ", after = " ^ countToString cl' ^ ", definition") d *) in def end end; end; (* ------------------------------------------------------------------------- *) (* Conjunctive normal form derivations. *) (* ------------------------------------------------------------------------- *) datatype thm = Metis_Thm of formula * inference and 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; 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 (Metis_Thm (fm1,_), Metis_Thm (fm2,_)) = compare (fm1,fm2); fun parentsThm (Metis_Thm (_,inf)) = parentsInference inf; fun mkAxiom fm = Metis_Thm (fromFormula fm, Axiom fm); fun destThm (Metis_Thm (fm,inf)) = (toFormula fm, inf); local val emptyProved : (thm,Metis_Formula.formula) Metis_Map.map = Metis_Map.new compareThm; fun isProved proved th = Metis_Map.inDomain th proved; fun isUnproved proved th = not (isProved proved th); fun lookupProved proved th = case Metis_Map.peek proved th of SOME fm => fm | NONE => raise Bug "Metis_Normalize.lookupProved"; fun prove acc proved ths = case ths of [] => List.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 List.null deps then let val (fm,inf) = destThm th val fms = List.map (lookupProved proved) pars val acc = (fm,inf,fms) :: acc val proved = Metis_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 = Metis_Print.ppMap toStringInference Metis_Print.ppString; (* ------------------------------------------------------------------------- *) (* Simplifying with definitions. *) (* ------------------------------------------------------------------------- *) datatype simplify = Simp of {formula : (formula, formula * thm) Metis_Map.map, andSet : (formula Metis_Set.set * formula * thm) list, orSet : (formula Metis_Set.set * formula * thm) list, xorSet : (formula Metis_Set.set * formula * thm) list}; val simplifyEmpty = Simp {formula = Metis_Map.new compare, andSet = [], orSet = [], xorSet = []}; local fun simpler fm s = Metis_Set.size s <> 1 orelse case Metis_Set.pick s of True => false | False => false | Metis_Literal _ => false | _ => true; fun addSet set_defs body_def = let fun def_body_size (body,_,_) = Metis_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 Metis_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 Metis_Literal _, simp) = let val s = Metis_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 Metis_Set.foldl addXorLiteral simp s end | _ => simp end | add (simp as Simp {formula,andSet,orSet,xorSet}) (body,def,th) = if Metis_Map.inDomain body formula then simp else let val formula = Metis_Map.insert formula (body,(def,th)) val formula = Metis_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 Metis_Set.size s = 1 then add simp (Metis_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 Metis_Thm (fm,_)) = add simp (fm,True,th); end; local fun simplifySet set_defs set = let fun pred (s,_,_) = Metis_Set.subset s set in case List.find pred set_defs of NONE => NONE | SOME (s,f,th) => let val set = Metis_Set.add (Metis_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 Metis_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) = Metis_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 Metis_Thm (fm,_) => case simp fm [] of SOME (fm,ths) => let val inf = Simplify (th,ths) in Metis_Thm (fm,inf) end | NONE => th end; end; (*MetisTrace2 val simplify = fn simp => fn th as Metis_Thm (fm,_) => let val th' as Metis_Thm (fm',_) = simplify simp th val () = if compare (fm,fm') = EQUAL then () else (Metis_Print.trace pp "Metis_Normalize.simplify: fm" fm; Metis_Print.trace pp "Metis_Normalize.simplify: fm'" fm') in th' end; *) (* ------------------------------------------------------------------------- *) (* Definitions. *) (* ------------------------------------------------------------------------- *) local val counter : int Unsynchronized.ref = Unsynchronized.ref 0; fun new () = let val Unsynchronized.ref i = counter val () = counter := i + 1 in definitionPrefix ^ "_" ^ Int.toString i end; in fun newDefinitionRelation () = Metis_Portable.critical new (); end; fun newDefinition def = let val fv = freeVars def val rel = newDefinitionRelation () val atm = (Metis_Name.fromString rel, Metis_NameSet.transform Metis_Term.Var fv) val fm = Metis_Formula.Iff (Metis_Formula.Metis_Atom atm, toFormula def) val fm = Metis_Formula.setMkForall (fv,fm) val inf = Definition (rel,fm) val lit = Metis_Literal (fv,(false,atm)) val fm = Xor2 (lit,def) in Metis_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 = [(Metis_LiteralSet.empty,th)] in (cls,InconsistentCnf) end; fun def_cnf_clause inf (fm,acc) = let val cl = toClause fm val th = Metis_Thm (fm,inf) in (cl,th) :: acc end (*MetisDebug handle Error err => (Metis_Print.trace pp "Metis_Normalize.addCnf.def_cnf_clause: fm" fm; raise Bug ("Metis_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 Metis_Thm (fm,_)) ths = case fm of True => def_cnf cls simp ths | False => def_cnf_inconsistent th | And (_,_,s) => let fun add (f,z) = Metis_Thm (f, Conjunct th) :: z in def_cnf cls simp (Metis_Set.foldr add ths s) end | Exists (fv,_,n,f) => let val th = Metis_Thm (skolemize fv n f, Skolemize th) in def_cnf_formula cls simp th ths end | Forall (_,_,_,f) => let val th = Metis_Thm (f, Specialize th) in def_cnf_formula cls simp th ths end | _ => case minimumDefinition fm of 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 (Metis_Thm (fm,inf)) | And (_,_,s) => let val inf = Conjunct (Metis_Thm (fm,inf)) val cls = Metis_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; 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 List.rev cls end; end; fun cnf fm = let val cls = proveCnf [mkAxiom fm] in List.map fst cls end; end (**** Original file: src/Model.sig ****) (* ========================================================================= *) (* RANDOM FINITE MODELS *) (* Copyright (c) 2003 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Model = sig (* ------------------------------------------------------------------------- *) (* Metis_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. *) (* ------------------------------------------------------------------------- *) 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. *) (* ------------------------------------------------------------------------- *) type parameters = {size : int, fixed : fixed} type model val default : parameters val new : parameters -> model val size : model -> int (* ------------------------------------------------------------------------- *) (* Interpreting terms and formulas in the model. *) (* ------------------------------------------------------------------------- *) 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 val interpretFormula : model -> valuation -> Metis_Formula.formula -> bool val interpretLiteral : model -> valuation -> Metis_Literal.literal -> bool val interpretClause : model -> valuation -> Metis_Thm.clause -> bool (* ------------------------------------------------------------------------- *) (* Check whether random groundings of a formula are true in the model. *) (* 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 option} -> model -> Metis_Atom.atom -> {T : int, F : int} val checkFormula : {maxChecks : int option} -> model -> Metis_Formula.formula -> {T : int, F : int} val checkLiteral : {maxChecks : int option} -> model -> Metis_Literal.literal -> {T : int, F : int} val checkClause : {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 (**** Original file: src/Model.sml ****) (* ========================================================================= *) (* RANDOM FINITE MODELS *) (* Copyright (c) 2003 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Model :> Metis_Model = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* Constants. *) (* ------------------------------------------------------------------------- *) val maxSpace = 1000; (* ------------------------------------------------------------------------- *) (* Helper functions. *) (* ------------------------------------------------------------------------- *) 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 "Metis_Model.intToBool"; fun minMaxInterval i j = interval i (1 + j - i); (* ------------------------------------------------------------------------- *) (* Metis_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 | 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. *) (* ------------------------------------------------------------------------- *) 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 uselessFixedFunction : fixedFunction = K (K NONE); val uselessFixedRelation : fixedRelation = K (K NONE); val emptyFunctions : fixedFunction Metis_NameArityMap.map = Metis_NameArityMap.new (); val emptyRelations : fixedRelation Metis_NameArityMap.map = Metis_NameArityMap.new (); fun fixed0 f sz elts = case elts of [] => f sz | _ => raise Bug "Metis_Model.fixed0: wrong arity"; fun fixed1 f sz elts = case elts of [x] => f sz x | _ => raise Bug "Metis_Model.fixed1: wrong arity"; fun fixed2 f sz elts = case elts of [x,y] => f sz x y | _ => raise Bug "Metis_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 Metis_NameArityMap.peek fns name_arity end; fun peekRelationFixed fix name_arity = let val Fixed {relations = rels, ...} = fix in Metis_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 = Metis_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 = Metis_NameArityMap.insert rels name_arity_rel in Fixed {functions = fns, relations = rels} end; local fun union _ = raise Bug "Metis_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 = Metis_NameArityMap.union union fns1 fns2 val rels = Metis_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 "Metis_Model.hasTypeFn: wrong arity"; fun eqRel _ elts = case elts of [x,y] => SOME (x = y) | _ => raise Bug "Metis_Model.eqRel: wrong arity"; in val basicFixed = let val fns = Metis_NameArityMap.singleton (Metis_Term.hasTypeFunction,hasTypeFn) val rels = Metis_NameArityMap.singleton (Metis_Atom.eqRelation,eqRel) in Fixed {functions = fns, relations = rels} end; end; (* ------------------------------------------------------------------------- *) (* Renaming fixed model parts. *) (* ------------------------------------------------------------------------- *) type fixedMap = {functionMap : Metis_Name.name Metis_NameArityMap.map, relationMap : Metis_Name.name Metis_NameArityMap.map}; fun mapFixed fixMap fix = let val {functionMap = fnMap, relationMap = relMap} = fixMap and Fixed {functions = fns, relations = rels} = fix val fns = Metis_NameArityMap.compose fnMap fns val rels = Metis_NameArityMap.compose relMap rels in Fixed {functions = fns, relations = rels} end; local fun mkEntry tag (na,n) = (tag,na,n); fun mkList tag m = List.map (mkEntry tag) (Metis_NameArityMap.toList m); fun ppEntry (tag,source_arity,target) = Metis_Print.inconsistentBlock 2 [Metis_Print.ppString tag, Metis_Print.break, Metis_NameArity.pp source_arity, Metis_Print.ppString " ->", Metis_Print.break, Metis_Name.pp target]; in fun ppFixedMap fixMap = let val {functionMap = fnMap, relationMap = relMap} = fixMap in case mkList "function" fnMap @ mkList "relation" relMap of [] => Metis_Print.skip | entry :: entries => Metis_Print.consistentBlock 0 (ppEntry entry :: List.map (Metis_Print.sequence Metis_Print.newline 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 "Metis_Model.projectionName: less than projectionMin" val _ = i <= projectionMax orelse raise Bug "Metis_Model.projectionName: greater than projectionMax" in Metis_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) (Metis_NameArityMap.insert acc (mkProj i)) val fns = addProj projectionMin emptyFunctions val rels = emptyRelations in Fixed {functions = fns, relations = rels} end; val projectionFixed = unionListFixed (List.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 "Metis_Model.numeralName: less than numeralMin" val _ = i <= numeralMax orelse raise Bug "Metis_Model.numeralName: greater than numeralMax" val s = if i < 0 then "negative" ^ Int.toString (~i) else Int.toString i in Metis_Name.fromString s end; val addName = Metis_Name.fromString "+" and divName = Metis_Name.fromString "div" and dividesName = Metis_Name.fromString "divides" and evenName = Metis_Name.fromString "even" and expName = Metis_Name.fromString "exp" and geName = Metis_Name.fromString ">=" and gtName = Metis_Name.fromString ">" and isZeroName = Metis_Name.fromString "isZero" and leName = Metis_Name.fromString "<=" and ltName = Metis_Name.fromString "<" and modName = Metis_Name.fromString "mod" and multName = Metis_Name.fromString "*" and negName = Metis_Name.fromString "~" and oddName = Metis_Name.fromString "odd" and preName = Metis_Name.fromString "pre" and subName = Metis_Name.fromString "-" and sucName = Metis_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 = Metis_NameArityMap.fromList (List.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 = Metis_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 = Metis_NameArityMap.fromList (List.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 = Metis_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 = Metis_Name.fromString "card" and complementName = Metis_Name.fromString "complement" and differenceName = Metis_Name.fromString "difference" and emptyName = Metis_Name.fromString "empty" and memberName = Metis_Name.fromString "member" and insertName = Metis_Name.fromString "insert" and intersectName = Metis_Name.fromString "intersect" and singletonName = Metis_Name.fromString "singleton" and subsetName = Metis_Name.fromString "subset" and symmetricDifferenceName = Metis_Name.fromString "symmetricDifference" and unionName = Metis_Name.fromString "union" and universeName = Metis_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 = Metis_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 = Metis_NameArityMap.fromList [((memberName,2), fixed2 memberRel), ((subsetName,2), fixed2 subsetRel)] in Fixed {functions = fns, relations = rels} end; end; (* Lists *) val appendName = Metis_Name.fromString "@" and consName = Metis_Name.fromString "::" and lengthName = Metis_Name.fromString "length" and nilName = Metis_Name.fromString "nil" and nullName = Metis_Name.fromString "null" and tailName = Metis_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 = Metis_NameArityMap.fromList [((appendName,2),addName), ((consName,2),sucName), ((lengthName,1), projectionName 1), ((nilName,0), numeralName 0), ((tailName,1),preName)], relationMap = Metis_NameArityMap.fromList [((nullName,1),isZeroName)]}; in val listFixed = mapFixed fixMap baseFix; end; (* ------------------------------------------------------------------------- *) (* Valuations. *) (* ------------------------------------------------------------------------- *) datatype valuation = Valuation of element Metis_NameMap.map; val emptyValuation = Valuation (Metis_NameMap.new ()); fun insertValuation (Valuation m) v_i = Valuation (Metis_NameMap.insert m v_i); fun peekValuation (Valuation m) v = Metis_NameMap.peek m v; fun constantValuation i = let fun add (v,V) = insertValuation V (v,i) in Metis_NameSet.foldl add emptyValuation end; val zeroValuation = constantValuation zeroElement; fun getValuation V v = case peekValuation V v of SOME i => i | NONE => raise Error "Metis_Model.getValuation: incomplete valuation"; fun randomValuation {size = N} vs = let fun f (v,V) = insertValuation V (v, Metis_Portable.randomInt N) in Metis_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 (Metis_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 = Metis_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 Metis_NameArityMap.map Unsynchronized.ref}; fun newTables N R = Tables {domainSize = N, rangeSize = R, tableMap = Unsynchronized.ref (Metis_NameArityMap.new ())}; fun getTables tables n_a = let val Tables {domainSize = N, rangeSize = _, tableMap = tm} = tables val Unsynchronized.ref m = tm in case Metis_NameArityMap.peek m n_a of SOME t => t | NONE => let val (_,a) = n_a val t = newTable N a val m = Metis_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; (* ------------------------------------------------------------------------- *) (* A type of random finite models. *) (* ------------------------------------------------------------------------- *) type parameters = {size : int, fixed : fixed}; datatype model = Metis_Model of {size : int, fixedFunctions : (element list -> element option) Metis_NameArityMap.map, fixedRelations : (element list -> bool option) Metis_NameArityMap.map, randomFunctions : tables, randomRelations : tables}; fun new {size = N, fixed} = let val Fixed {functions = fns, relations = rels} = fixed val fixFns = Metis_NameArityMap.transform (fn f => f {size = N}) fns and fixRels = Metis_NameArityMap.transform (fn r => r {size = N}) rels val rndFns = newTables N N and rndRels = newTables N 2 in Metis_Model {size = N, fixedFunctions = fixFns, fixedRelations = fixRels, randomFunctions = rndFns, randomRelations = rndRels} end; fun size (Metis_Model {size = N, ...}) = N; fun peekFixedFunction M (n,elts) = let val Metis_Model {fixedFunctions = fixFns, ...} = M in case Metis_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 Metis_Model {fixedRelations = fixRels, ...} = M in case Metis_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 Metis_Term.Var _ => tm | Metis_Term.Fn f_tms => case Metis_Term.stripApp tm of (_,[]) => tm | (v as Metis_Term.Var _, tms) => Metis_Term.Fn (Metis_Term.appName, v :: tms) | (Metis_Term.Fn (f,tms), tms') => Metis_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 Metis_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 Metis_Model {randomRelations = rndRels, ...} = M in intToBool (lookupTables rndRels n_elts) end; fun interpretTerm M V = let fun interpret tm = case destTerm tm of Metis_Term.Var v => getValuation V v | Metis_Term.Fn (f,tms) => interpretFunction M (f, List.map interpret tms) in interpret end; fun interpretAtom M V (r,tms) = interpretRelation M (r, List.map (interpretTerm M V) tms); fun interpretFormula M = let val N = size M fun interpret V fm = case fm of Metis_Formula.True => true | Metis_Formula.False => false | Metis_Formula.Metis_Atom atm => interpretAtom M V atm | Metis_Formula.Not p => not (interpret V p) | Metis_Formula.Or (p,q) => interpret V p orelse interpret V q | Metis_Formula.And (p,q) => interpret V p andalso interpret V q | Metis_Formula.Imp (p,q) => interpret V (Metis_Formula.Or (Metis_Formula.Not p, q)) | Metis_Formula.Iff (p,q) => interpret V p = interpret V q | Metis_Formula.Forall (v,p) => interpret' V p v N | Metis_Formula.Exists (v,p) => interpret V (Metis_Formula.Not (Metis_Formula.Forall (v, Metis_Formula.Not p))) and interpret' V fm v i = i = 0 orelse let val i = i - 1 val V' = insertValuation V (v,i) in interpret V' fm andalso interpret' V fm v i end in interpret end; 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 = Metis_LiteralSet.exists (interpretLiteral M V) cl; (* ------------------------------------------------------------------------- *) (* Check whether random groundings of a formula are true in the model. *) (* Note: if it's cheaper, a systematic check will be performed instead. *) (* ------------------------------------------------------------------------- *) 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 (Metis_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 (Metis_Atom.freeVars atm) atm; fun checkFormula maxChecks M fm = check interpretFormula maxChecks M (Metis_Formula.freeVars fm) fm; fun checkLiteral maxChecks M lit = check interpretLiteral maxChecks M (Metis_Literal.freeVars lit) lit; fun checkClause maxChecks M cl = check interpretClause maxChecks M (Metis_LiteralSet.freeVars cl) cl; (* ------------------------------------------------------------------------- *) (* Updating the model. *) (* ------------------------------------------------------------------------- *) fun updateFunction M func_elts_elt = let val Metis_Model {randomFunctions = rndFns, ...} = M val () = updateTables rndFns func_elts_elt in () end; fun updateRelation M (rel_elts,pol) = let val Metis_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 Metis_Term.functionName * modelTerm list * int list; fun modelTerm M V = let fun modelTm tm = case destTerm tm of Metis_Term.Var v => (ModelVar, getValuation V v) | Metis_Term.Fn (f,tms) => let val (tms,xs) = unzip (List.map modelTm tms) in (ModelFn (f,tms,xs), interpretFunction M (f,xs)) end in modelTm end; (* ------------------------------------------------------------------------- *) (* Perturbing the model. *) (* ------------------------------------------------------------------------- *) datatype perturbation = FunctionPerturbation of (Metis_Term.functionName * element list) * element | RelationPerturbation of (Metis_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 List.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 "Metis_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 (List.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 = Metis_LiteralSet.foldl (pertLiteral M V) acc cl; fun pickPerturb M perts = if List.null perts then () else perturb M (List.nth (perts, Metis_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 = Metis_Print.program [Metis_Print.ppString "Metis_Model{", Metis_Print.ppInt (size M), Metis_Print.ppString "}"]; end (**** Original file: src/Problem.sig ****) (* ========================================================================= *) (* CNF PROBLEMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Problem = sig (* ------------------------------------------------------------------------- *) (* Problems. *) (* ------------------------------------------------------------------------- *) type problem = {axioms : Metis_Thm.clause list, conjecture : Metis_Thm.clause list} val size : problem -> {clauses : int, literals : int, symbols : int, typedSymbols : int} 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 (* ------------------------------------------------------------------------- *) (* Categorizing problems. *) (* ------------------------------------------------------------------------- *) datatype propositional = Propositional | EffectivelyPropositional | NonPropositional datatype equality = NonEquality | Equality | PureEquality datatype horn = Trivial | Unit | DoubleHorn | Horn | NegativeHorn | NonHorn type category = {propositional : propositional, equality : equality, horn : horn} val categorize : problem -> category val categoryToString : category -> string end (**** Original file: src/Problem.sml ****) (* ========================================================================= *) (* CNF PROBLEMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Problem :> Metis_Problem = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* Problems. *) (* ------------------------------------------------------------------------- *) type problem = {axioms : Metis_Thm.clause list, conjecture : Metis_Thm.clause list}; fun toClauses {axioms,conjecture} = axioms @ conjecture; fun size prob = let fun lits (cl,n) = n + Metis_LiteralSet.size cl fun syms (cl,n) = n + Metis_LiteralSet.symbols cl fun typedSyms (cl,n) = n + Metis_LiteralSet.typedSymbols cl val cls = toClauses prob in {clauses = length cls, literals = List.foldl lits 0 cls, symbols = List.foldl syms 0 cls, typedSymbols = List.foldl typedSyms 0 cls} end; fun freeVars {axioms,conjecture} = Metis_NameSet.union (Metis_LiteralSet.freeVarsList axioms) (Metis_LiteralSet.freeVarsList conjecture); local fun clauseToFormula cl = Metis_Formula.listMkDisj (Metis_LiteralSet.transform Metis_Literal.toFormula cl); in fun toFormula prob = Metis_Formula.listMkConj (List.map clauseToFormula (toClauses prob)); fun toGoal {axioms,conjecture} = let val clToFm = Metis_Formula.generalize o clauseToFormula val clsToFm = Metis_Formula.listMkConj o List.map clToFm val fm = Metis_Formula.False val fm = if List.null conjecture then fm else Metis_Formula.Imp (clsToFm conjecture, fm) val fm = Metis_Formula.Imp (clsToFm axioms, fm) in fm end; end; fun toString prob = Metis_Formula.toString (toFormula prob); (* ------------------------------------------------------------------------- *) (* Categorizing problems. *) (* ------------------------------------------------------------------------- *) datatype propositional = Propositional | EffectivelyPropositional | NonPropositional; datatype equality = NonEquality | Equality | PureEquality; datatype horn = Trivial | Unit | DoubleHorn | Horn | NegativeHorn | NonHorn; type category = {propositional : propositional, equality : equality, horn : horn}; fun categorize prob = let val cls = toClauses prob val rels = let fun f (cl,set) = Metis_NameAritySet.union set (Metis_LiteralSet.relations cl) in List.foldl f Metis_NameAritySet.empty cls end val funs = let fun f (cl,set) = Metis_NameAritySet.union set (Metis_LiteralSet.functions cl) in List.foldl f Metis_NameAritySet.empty cls end val propositional = if Metis_NameAritySet.allNullary rels then Propositional else if Metis_NameAritySet.allNullary funs then EffectivelyPropositional else NonPropositional val equality = if not (Metis_NameAritySet.member Metis_Atom.eqRelation rels) then NonEquality else if Metis_NameAritySet.size rels = 1 then PureEquality else Equality val horn = if List.exists Metis_LiteralSet.null cls then Trivial else if List.all (fn cl => Metis_LiteralSet.size cl = 1) cls then Unit else let fun pos cl = Metis_LiteralSet.count Metis_Literal.positive cl <= 1 fun neg cl = Metis_LiteralSet.count Metis_Literal.negative cl <= 1 in case (List.all pos cls, List.all neg cls) of (true,true) => DoubleHorn | (true,false) => Horn | (false,true) => NegativeHorn | (false,false) => NonHorn end in {propositional = propositional, equality = equality, horn = horn} end; fun categoryToString {propositional,equality,horn} = (case propositional of Propositional => "propositional" | EffectivelyPropositional => "effectively propositional" | NonPropositional => "non-propositional") ^ ", " ^ (case equality of NonEquality => "non-equality" | Equality => "equality" | PureEquality => "pure equality") ^ ", " ^ (case horn of Trivial => "trivial" | Unit => "unit" | DoubleHorn => "horn (and negative horn)" | Horn => "horn" | NegativeHorn => "negative horn" | NonHorn => "non-horn"); end (**** Original file: src/TermNet.sig ****) (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_TermNet = sig (* ------------------------------------------------------------------------- *) (* A type of term sets that can be efficiently matched and unified. *) (* ------------------------------------------------------------------------- *) type parameters = {fifo : bool} type 'a termNet (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val new : parameters -> 'a termNet val null : 'a termNet -> bool val size : 'a termNet -> int val insert : 'a termNet -> Metis_Term.term * 'a -> 'a termNet val fromList : parameters -> (Metis_Term.term * 'a) list -> 'a termNet val filter : ('a -> bool) -> 'a termNet -> 'a termNet val toString : 'a termNet -> string val pp : 'a Metis_Print.pp -> 'a termNet Metis_Print.pp (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) (* *) (* These function return OVER-APPROXIMATIONS! *) (* Filter afterwards to get the precise set of satisfying values. *) (* ------------------------------------------------------------------------- *) val match : 'a termNet -> Metis_Term.term -> 'a list val matched : 'a termNet -> Metis_Term.term -> 'a list val unify : 'a termNet -> Metis_Term.term -> 'a list end (**** Original file: src/TermNet.sml ****) (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_TermNet :> Metis_TermNet = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* Anonymous variables. *) (* ------------------------------------------------------------------------- *) val anonymousName = Metis_Name.fromString "_"; val anonymousVar = Metis_Term.Var anonymousName; (* ------------------------------------------------------------------------- *) (* Quotient terms. *) (* ------------------------------------------------------------------------- *) datatype qterm = Var | Fn of Metis_NameArity.nameArity * qterm list; local fun cmp [] = EQUAL | cmp (q1_q2 :: qs) = if Metis_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 Metis_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 (Metis_Term.Var _) = Var | termToQterm (Metis_Term.Fn (f,l)) = Fn ((f, length l), List.map termToQterm l); local fun qm [] = true | qm ((Var,_) :: rest) = qm rest | qm ((Fn _, Var) :: _) = false | qm ((Fn (f,a), Fn (g,b)) :: rest) = Metis_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 _, Metis_Term.Var _) :: _) = false | qm ((Fn ((f,n),a), Metis_Term.Fn (g,b)) :: rest) = Metis_Name.equal f g andalso n = length b andalso qm (zip a b @ rest); in fun matchQtermTerm qtm tm = qm [(qtm,tm)]; end; local fun qn qsub [] = SOME qsub | qn qsub ((Metis_Term.Var v, qtm) :: rest) = (case Metis_NameMap.peek qsub v of NONE => qn (Metis_NameMap.insert qsub (v,qtm)) rest | SOME qtm' => if equalQterm qtm qtm' then qn qsub rest else NONE) | qn _ ((Metis_Term.Fn _, Var) :: _) = NONE | qn qsub ((Metis_Term.Fn (f,a), Fn ((g,n),b)) :: rest) = if Metis_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 _ = Metis_NameArity.equal f g orelse raise Error "Metis_TermNet.qv" in Fn (f, zipWith qv a b) end; fun qu qsub [] = qsub | qu qsub ((Var, _) :: rest) = qu qsub rest | qu qsub ((qtm, Metis_Term.Var v) :: rest) = let val qtm = case Metis_NameMap.peek qsub v of NONE => qtm | SOME qtm' => qv qtm qtm' in qu (Metis_NameMap.insert qsub (v,qtm)) rest end | qu qsub ((Fn ((f,n),a), Metis_Term.Fn (g,b)) :: rest) = if Metis_Name.equal f g andalso n = length b then qu qsub (zip a b @ rest) else raise Error "Metis_TermNet.qu"; in fun unifyQtermQterm qtm qtm' = total (qv qtm) qtm'; fun unifyQtermTerm qsub qtm tm = total (qu qsub) [(qtm,tm)]; end; local fun qtermToTerm Var = anonymousVar | qtermToTerm (Fn ((f,_),l)) = Metis_Term.Fn (f, List.map qtermToTerm l); in val ppQterm = Metis_Print.ppMap qtermToTerm Metis_Term.pp; end; (* ------------------------------------------------------------------------- *) (* A type of term sets that can be efficiently matched and unified. *) (* ------------------------------------------------------------------------- *) type parameters = {fifo : bool}; datatype 'a net = Result of 'a list | Single of qterm * 'a net | Multiple of 'a net option * 'a net Metis_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)) = Metis_NameArityMap.foldl (fn (_,n,acc) => acc + computeSize n) (case vs of SOME n => computeSize n | NONE => 0) fs; in fun netSize NONE = NONE | netSize (SOME n) = SOME (computeSize n, n); end; fun size (Net (_,_,NONE)) = 0 | size (Net (_, _, SOME (i,_))) = i; fun null net = size net = 0; fun singles qtms a = List.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 equalQterm qtm qtm' then Single (qtm, add a qtms n) else add a input1 (add n [qtm'] (Multiple (NONE, Metis_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 = Metis_NameArityMap.peek fs f in Multiple (vs, Metis_NameArityMap.insert fs (f, oadd a (l @ qtms) n)) end | add _ _ _ = raise Bug "Metis_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)) handle Error _ => raise Bug "Metis_TermNet.insert: should never fail"; end; fun fromList parm l = List.foldl (fn (tm_a,n) => insert n tm_a) (new parm) l; fun filter pred = let fun filt (Result l) = (case List.filter (fn (_,a) => pred a) l of [] => NONE | l => SOME (Result l)) | filt (Single (qtm,n)) = (case filt n of NONE => NONE | SOME n => SOME (Single (qtm,n))) | filt (Multiple (vs,fs)) = let val vs = Option.mapPartial filt vs val fs = Metis_NameArityMap.mapPartial (fn (_,n) => filt n) fs in if not (Option.isSome vs) andalso Metis_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)) end handle Error _ => raise Bug "Metis_TermNet.filter: should never fail"; fun toString net = "Metis_TermNet[" ^ Int.toString (size net) ^ "]"; (* ------------------------------------------------------------------------- *) (* Specialized fold operations to support matching and unification. *) (* ------------------------------------------------------------------------- *) local fun norm (0 :: ks, (f as (_,n)) :: fs, qtms) = let val (a,qtms) = revDivide qtms n in addQterm (Fn (f,a)) (ks,fs,qtms) end | norm stack = stack and addQterm qtm (ks,fs,qtms) = let val ks = case ks of [] => [] | k :: ks => (k - 1) :: ks in norm (ks, fs, qtm :: qtms) end and addFn (f as (_,n)) (ks,fs,qtms) = norm (n :: ks, f :: fs, qtms); in val stackEmpty = ([],[],[]); val stackAddQterm = addQterm; val stackAddFn = addFn; fun stackValue ([],[],[qtm]) = qtm | stackValue _ = raise Bug "Metis_TermNet.stackValue"; end; local 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 - 1, stackAddQterm qtm stack, net) :: 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 fun getFns (f as (_,k), net, x) = (k + n, stackAddFn f stack, net) :: x in fold inc acc (Metis_NameArityMap.foldr getFns rest fns) end | fold _ _ _ = raise Bug "Metis_TermNet.foldTerms.fold"; in fun foldTerms inc acc net = fold inc acc [(1,stackEmpty,net)]; end; fun foldEqualTerms pat inc acc = let fun fold ([],net) = inc (pat,net,acc) | 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)) = (case Metis_NameArityMap.peek fns f of NONE => acc | SOME net => fold (a @ pats, net)) | fold _ = raise Bug "Metis_TermNet.foldEqualTerms.fold"; in fn net => fold ([pat],net) end; local 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) = 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) = (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) = let val rest = case v of NONE => rest | SOME net => (pats, stackAddQterm pat stack, net) :: rest val rest = case Metis_NameArityMap.peek fns f of NONE => rest | SOME net => (a @ pats, stackAddFn f stack, net) :: rest in fold inc acc rest end | fold _ _ _ = raise Bug "Metis_TermNet.foldUnifiableTerms.fold"; in fun foldUnifiableTerms pat inc acc net = fold inc acc [([pat],stackEmpty,net)]; end; (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) (* *) (* These function return OVER-APPROXIMATIONS! *) (* Filter afterwards to get the precise set of satisfying values. *) (* ------------------------------------------------------------------------- *) local fun idwise ((m,_),(n,_)) = Int.compare (m,n); fun fifoize ({fifo, ...} : parameters) l = if fifo then sort idwise l else l; in fun finally parm l = List.map snd (fifoize parm l); end; local fun mat acc [] = acc | 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) = let val rest = case vs of NONE => rest | SOME n => (n,tms) :: rest val rest = case tm of Metis_Term.Var _ => rest | Metis_Term.Fn (f,l) => case Metis_NameArityMap.peek fs (f, length l) of NONE => rest | SOME n => (n, l @ tms) :: rest in mat acc rest end | mat _ _ = raise Bug "Metis_TermNet.match: Match"; in fun match (Net (_,_,NONE)) _ = [] | match (Net (p, _, SOME (_,n))) tm = finally p (mat [] [(n,[tm])]) handle Error _ => raise Bug "Metis_TermNet.match: should never fail"; end; local fun unseenInc qsub v tms (qtm,net,rest) = (Metis_NameMap.insert qsub (v,qtm), net, tms) :: rest; 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) = (case matchTermQterm qsub tm qtm of NONE => mat acc rest | SOME qsub => mat acc ((qsub,net,tms) :: rest)) | mat acc ((qsub, net as Multiple _, Metis_Term.Var v :: tms) :: rest) = (case Metis_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), Metis_Term.Fn (f,a) :: tms) :: rest) = let val rest = case Metis_NameArityMap.peek fns (f, length a) of NONE => rest | SOME net => (qsub, net, a @ tms) :: rest in mat acc rest end | mat _ _ = raise Bug "Metis_TermNet.matched.mat"; in fun matched (Net (_,_,NONE)) _ = [] | matched (Net (parm, _, SOME (_,net))) tm = finally parm (mat [] [(Metis_NameMap.new (), net, [tm])]) handle Error _ => raise Bug "Metis_TermNet.matched: should never fail"; end; local fun inc qsub v tms (qtm,net,rest) = (Metis_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) = (case unifyQtermTerm qsub qtm tm of NONE => mat acc rest | SOME qsub => mat acc ((qsub,net,tms) :: rest)) | mat acc ((qsub, net as Multiple _, Metis_Term.Var v :: tms) :: rest) = (case Metis_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), Metis_Term.Fn (f,a) :: tms) :: rest) = let val rest = case v of NONE => rest | SOME net => (qsub,net,tms) :: rest val rest = case Metis_NameArityMap.peek fns (f, length a) of NONE => rest | SOME net => (qsub, net, a @ tms) :: rest in mat acc rest end | mat _ _ = raise Bug "Metis_TermNet.unify.mat"; in fun unify (Net (_,_,NONE)) _ = [] | unify (Net (parm, _, SOME (_,net))) tm = finally parm (mat [] [(Metis_NameMap.new (), net, [tm])]) handle Error _ => raise Bug "Metis_TermNet.unify: should never fail"; end; (* ------------------------------------------------------------------------- *) (* Pretty printing. *) (* ------------------------------------------------------------------------- *) local fun inc (qtm, Result l, acc) = List.foldl (fn ((n,a),acc) => (n,(qtm,a)) :: acc) acc l | inc _ = raise Bug "Metis_TermNet.pp.inc"; fun toList (Net (_,_,NONE)) = [] | toList (Net (parm, _, SOME (_,net))) = finally parm (foldTerms inc [] net); in fun pp ppA = Metis_Print.ppMap toList (Metis_Print.ppList (Metis_Print.ppOp2 " |->" ppQterm ppA)); end; end (**** Original file: src/AtomNet.sig ****) (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_AtomNet = sig (* ------------------------------------------------------------------------- *) (* A type of atom sets that can be efficiently matched and unified. *) (* ------------------------------------------------------------------------- *) type parameters = {fifo : bool} type 'a atomNet (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val new : parameters -> 'a atomNet val size : 'a atomNet -> int val insert : 'a atomNet -> Metis_Atom.atom * 'a -> 'a atomNet val fromList : parameters -> (Metis_Atom.atom * 'a) list -> 'a atomNet val filter : ('a -> bool) -> 'a atomNet -> 'a atomNet val toString : 'a atomNet -> string val pp : 'a Metis_Print.pp -> 'a atomNet Metis_Print.pp (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) (* *) (* These function return OVER-APPROXIMATIONS! *) (* Filter afterwards to get the precise set of satisfying values. *) (* ------------------------------------------------------------------------- *) val match : 'a atomNet -> Metis_Atom.atom -> 'a list val matched : 'a atomNet -> Metis_Atom.atom -> 'a list val unify : 'a atomNet -> Metis_Atom.atom -> 'a list end (**** Original file: src/AtomNet.sml ****) (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_AtomNet :> Metis_AtomNet = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* Helper functions. *) (* ------------------------------------------------------------------------- *) fun atomToTerm atom = Metis_Term.Fn atom; fun termToAtom (Metis_Term.Var _) = raise Bug "Metis_AtomNet.termToAtom" | termToAtom (Metis_Term.Fn atom) = atom; (* ------------------------------------------------------------------------- *) (* A type of atom sets that can be efficiently matched and unified. *) (* ------------------------------------------------------------------------- *) type parameters = Metis_TermNet.parameters; type 'a atomNet = 'a Metis_TermNet.termNet; (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val new = Metis_TermNet.new; val size = Metis_TermNet.size; fun insert net (atm,a) = Metis_TermNet.insert net (atomToTerm atm, a); fun fromList parm l = List.foldl (fn (atm_a,n) => insert n atm_a) (new parm) l; val filter = Metis_TermNet.filter; fun toString net = "Metis_AtomNet[" ^ Int.toString (size net) ^ "]"; val pp = Metis_TermNet.pp; (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) (* *) (* These function return OVER-APPROXIMATIONS! *) (* Filter afterwards to get the precise set of satisfying values. *) (* ------------------------------------------------------------------------- *) fun match net atm = Metis_TermNet.match net (atomToTerm atm); fun matched net atm = Metis_TermNet.matched net (atomToTerm atm); fun unify net atm = Metis_TermNet.unify net (atomToTerm atm); end (**** Original file: src/LiteralNet.sig ****) (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_LiteralNet = sig (* ------------------------------------------------------------------------- *) (* A type of literal sets that can be efficiently matched and unified. *) (* ------------------------------------------------------------------------- *) type parameters = {fifo : bool} type 'a literalNet (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val new : parameters -> 'a literalNet val size : 'a literalNet -> int val profile : 'a literalNet -> {positive : int, negative : int} val insert : 'a literalNet -> Metis_Literal.literal * 'a -> 'a literalNet val fromList : parameters -> (Metis_Literal.literal * 'a) list -> 'a literalNet val filter : ('a -> bool) -> 'a literalNet -> 'a literalNet val toString : 'a literalNet -> string val pp : 'a Metis_Print.pp -> 'a literalNet Metis_Print.pp (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) (* *) (* These function return OVER-APPROXIMATIONS! *) (* Filter afterwards to get the precise set of satisfying values. *) (* ------------------------------------------------------------------------- *) val match : 'a literalNet -> Metis_Literal.literal -> 'a list val matched : 'a literalNet -> Metis_Literal.literal -> 'a list val unify : 'a literalNet -> Metis_Literal.literal -> 'a list end (**** Original file: src/LiteralNet.sml ****) (* ========================================================================= *) (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_LiteralNet :> Metis_LiteralNet = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* A type of literal sets that can be efficiently matched and unified. *) (* ------------------------------------------------------------------------- *) type parameters = Metis_AtomNet.parameters; type 'a literalNet = {positive : 'a Metis_AtomNet.atomNet, negative : 'a Metis_AtomNet.atomNet}; (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) fun new parm = {positive = Metis_AtomNet.new parm, negative = Metis_AtomNet.new parm}; local fun pos ({positive,...} : 'a literalNet) = Metis_AtomNet.size positive; fun neg ({negative,...} : 'a literalNet) = Metis_AtomNet.size negative; in fun size net = pos net + neg net; fun profile net = {positive = pos net, negative = neg net}; end; fun insert {positive,negative} ((true,atm),a) = {positive = Metis_AtomNet.insert positive (atm,a), negative = negative} | insert {positive,negative} ((false,atm),a) = {positive = positive, negative = Metis_AtomNet.insert negative (atm,a)}; fun fromList parm l = List.foldl (fn (lit_a,n) => insert n lit_a) (new parm) l; fun filter pred {positive,negative} = {positive = Metis_AtomNet.filter pred positive, negative = Metis_AtomNet.filter pred negative}; fun toString net = "Metis_LiteralNet[" ^ Int.toString (size net) ^ "]"; fun pp ppA = Metis_Print.ppMap (fn {positive,negative} => (positive,negative)) (Metis_Print.ppOp2 " + NEGATIVE" (Metis_AtomNet.pp ppA) (Metis_AtomNet.pp ppA)); (* ------------------------------------------------------------------------- *) (* Matching and unification queries. *) (* *) (* These function return OVER-APPROXIMATIONS! *) (* Filter afterwards to get the precise set of satisfying values. *) (* ------------------------------------------------------------------------- *) fun match ({positive,...} : 'a literalNet) (true,atm) = Metis_AtomNet.match positive atm | match {negative,...} (false,atm) = Metis_AtomNet.match negative atm; fun matched ({positive,...} : 'a literalNet) (true,atm) = Metis_AtomNet.matched positive atm | matched {negative,...} (false,atm) = Metis_AtomNet.matched negative atm; fun unify ({positive,...} : 'a literalNet) (true,atm) = Metis_AtomNet.unify positive atm | unify {negative,...} (false,atm) = Metis_AtomNet.unify negative atm; end (**** Original file: src/Subsume.sig ****) (* ========================================================================= *) (* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES *) (* Copyright (c) 2002 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Subsume = sig (* ------------------------------------------------------------------------- *) (* A type of clause sets that supports efficient subsumption checking. *) (* ------------------------------------------------------------------------- *) type 'a subsume val new : unit -> 'a subsume val size : 'a subsume -> int val insert : 'a subsume -> Metis_Thm.clause * 'a -> 'a subsume val filter : ('a -> bool) -> 'a subsume -> 'a subsume val pp : 'a subsume Metis_Print.pp val toString : 'a subsume -> string (* ------------------------------------------------------------------------- *) (* Subsumption checking. *) (* ------------------------------------------------------------------------- *) val subsumes : (Metis_Thm.clause * Metis_Subst.subst * 'a -> bool) -> 'a subsume -> Metis_Thm.clause -> (Metis_Thm.clause * Metis_Subst.subst * 'a) option val isSubsumed : 'a subsume -> Metis_Thm.clause -> bool val strictlySubsumes : (* exclude subsuming clauses with more literals *) (Metis_Thm.clause * Metis_Subst.subst * 'a -> bool) -> 'a subsume -> Metis_Thm.clause -> (Metis_Thm.clause * Metis_Subst.subst * 'a) option val isStrictlySubsumed : 'a subsume -> Metis_Thm.clause -> bool (* ------------------------------------------------------------------------- *) (* Single clause versions. *) (* ------------------------------------------------------------------------- *) val clauseSubsumes : Metis_Thm.clause -> Metis_Thm.clause -> Metis_Subst.subst option val clauseStrictlySubsumes : Metis_Thm.clause -> Metis_Thm.clause -> Metis_Subst.subst option end (**** Original file: src/Subsume.sml ****) (* ========================================================================= *) (* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES *) (* Copyright (c) 2002 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Subsume :> Metis_Subsume = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* Helper functions. *) (* ------------------------------------------------------------------------- *) fun findRest pred = let fun f _ [] = NONE | f ys (x :: xs) = if pred x then SOME (x, List.revAppend (ys,xs)) else f (x :: ys) xs in f [] end; local fun addSym (lit,acc) = case total Metis_Literal.sym lit of NONE => acc | SOME lit => lit :: acc in fun clauseSym lits = List.foldl addSym lits lits; end; fun sortClause cl = let val lits = Metis_LiteralSet.toList cl in sortMap Metis_Literal.typedSymbols (revCompare Int.compare) lits end; fun incompatible lit = let val lits = clauseSym [lit] in fn lit' => not (List.exists (can (Metis_Literal.unify Metis_Subst.empty lit')) lits) end; (* ------------------------------------------------------------------------- *) (* Metis_Clause ids and lengths. *) (* ------------------------------------------------------------------------- *) type clauseId = int; type clauseLength = int; local type idSet = (clauseId * clauseLength) Metis_Set.set; fun idCompare ((id1,len1),(id2,len2)) = case Int.compare (len1,len2) of LESS => LESS | EQUAL => Int.compare (id1,id2) | GREATER => GREATER; in val idSetEmpty : idSet = Metis_Set.empty idCompare; fun idSetAdd (id_len,set) : idSet = Metis_Set.add set id_len; fun idSetAddMax max (id_len as (_,len), set) : idSet = if len <= max then Metis_Set.add set id_len else set; fun idSetIntersect set1 set2 : idSet = Metis_Set.intersect set1 set2; end; (* ------------------------------------------------------------------------- *) (* A type of clause sets that supports efficient subsumption checking. *) (* ------------------------------------------------------------------------- *) datatype 'a subsume = Metis_Subsume of {empty : (Metis_Thm.clause * Metis_Subst.subst * 'a) list, unit : (Metis_Literal.literal * Metis_Thm.clause * 'a) Metis_LiteralNet.literalNet, nonunit : {nextId : clauseId, clauses : (Metis_Literal.literal list * Metis_Thm.clause * 'a) Metis_IntMap.map, fstLits : (clauseId * clauseLength) Metis_LiteralNet.literalNet, sndLits : (clauseId * clauseLength) Metis_LiteralNet.literalNet}}; fun new () = Metis_Subsume {empty = [], unit = Metis_LiteralNet.new {fifo = false}, nonunit = {nextId = 0, clauses = Metis_IntMap.new (), fstLits = Metis_LiteralNet.new {fifo = false}, sndLits = Metis_LiteralNet.new {fifo = false}}}; fun size (Metis_Subsume {empty, unit, nonunit = {clauses,...}}) = length empty + Metis_LiteralNet.size unit + Metis_IntMap.size clauses; fun insert (Metis_Subsume {empty,unit,nonunit}) (cl',a) = case sortClause cl' of [] => let val empty = (cl',Metis_Subst.empty,a) :: empty in Metis_Subsume {empty = empty, unit = unit, nonunit = nonunit} end | [lit] => let val unit = Metis_LiteralNet.insert unit (lit,(lit,cl',a)) in Metis_Subsume {empty = empty, unit = unit, nonunit = nonunit} end | fstLit :: (nonFstLits as sndLit :: otherLits) => let val {nextId,clauses,fstLits,sndLits} = nonunit val id_length = (nextId, Metis_LiteralSet.size cl') val fstLits = Metis_LiteralNet.insert fstLits (fstLit,id_length) val (sndLit,otherLits) = case findRest (incompatible fstLit) nonFstLits of SOME sndLit_otherLits => sndLit_otherLits | NONE => (sndLit,otherLits) val sndLits = Metis_LiteralNet.insert sndLits (sndLit,id_length) val lits' = otherLits @ [fstLit,sndLit] val clauses = Metis_IntMap.insert clauses (nextId,(lits',cl',a)) val nextId = nextId + 1 val nonunit = {nextId = nextId, clauses = clauses, fstLits = fstLits, sndLits = sndLits} in Metis_Subsume {empty = empty, unit = unit, nonunit = nonunit} end; fun filter pred (Metis_Subsume {empty,unit,nonunit}) = let val empty = List.filter (pred o #3) empty val unit = Metis_LiteralNet.filter (pred o #3) unit val nonunit = let val {nextId,clauses,fstLits,sndLits} = nonunit val clauses' = Metis_IntMap.filter (pred o #3 o snd) clauses in if Metis_IntMap.size clauses = Metis_IntMap.size clauses' then nonunit else let fun predId (id,_) = Metis_IntMap.inDomain id clauses' val fstLits = Metis_LiteralNet.filter predId fstLits and sndLits = Metis_LiteralNet.filter predId sndLits in {nextId = nextId, clauses = clauses', fstLits = fstLits, sndLits = sndLits} end end in Metis_Subsume {empty = empty, unit = unit, nonunit = nonunit} end; fun toString subsume = "Metis_Subsume{" ^ Int.toString (size subsume) ^ "}"; fun pp subsume = Metis_Print.ppMap toString Metis_Print.ppString subsume; (* ------------------------------------------------------------------------- *) (* Subsumption checking. *) (* ------------------------------------------------------------------------- *) local fun matchLit lit' (lit,acc) = case total (Metis_Literal.match Metis_Subst.empty lit') lit of SOME sub => sub :: acc | NONE => acc; in fun genClauseSubsumes pred cl' lits' cl a = let fun mkSubsl acc sub [] = SOME (sub, sortMap length Int.compare acc) | mkSubsl acc sub (lit' :: lits') = case List.foldl (matchLit lit') [] cl of [] => NONE | [sub'] => (case total (Metis_Subst.union sub) sub' of NONE => NONE | SOME sub => mkSubsl acc sub lits') | subs => mkSubsl (subs :: acc) sub lits' fun search [] = NONE | search ((sub,[]) :: others) = let val x = (cl',sub,a) in if pred x then SOME x else search others end | search ((_, [] :: _) :: others) = search others | search ((sub, (sub' :: subs) :: subsl) :: others) = let val others = (sub, subs :: subsl) :: others in case total (Metis_Subst.union sub) sub' of NONE => search others | SOME sub => search ((sub,subsl) :: others) end in case mkSubsl [] Metis_Subst.empty lits' of NONE => NONE | SOME sub_subsl => search [sub_subsl] end; end; local fun emptySubsumes pred empty = List.find pred empty; fun unitSubsumes pred unit = let fun subLit lit = let fun subUnit (lit',cl',a) = case total (Metis_Literal.match Metis_Subst.empty lit') lit of NONE => NONE | SOME sub => let val x = (cl',sub,a) in if pred x then SOME x else NONE end in first subUnit (Metis_LiteralNet.match unit lit) end in first subLit end; fun nonunitSubsumes pred nonunit max cl = let val addId = case max of NONE => idSetAdd | SOME n => idSetAddMax n fun subLit lits (lit,acc) = List.foldl addId acc (Metis_LiteralNet.match lits lit) val {nextId = _, clauses, fstLits, sndLits} = nonunit fun subCl' (id,_) = let val (lits',cl',a) = Metis_IntMap.get clauses id in genClauseSubsumes pred cl' lits' cl a end val fstCands = List.foldl (subLit fstLits) idSetEmpty cl val sndCands = List.foldl (subLit sndLits) idSetEmpty cl val cands = idSetIntersect fstCands sndCands in Metis_Set.firstl subCl' cands end; fun genSubsumes pred (Metis_Subsume {empty,unit,nonunit}) max cl = case emptySubsumes pred empty of s as SOME _ => s | NONE => if max = SOME 0 then NONE else let val cl = clauseSym (Metis_LiteralSet.toList cl) in case unitSubsumes pred unit cl of s as SOME _ => s | NONE => if max = SOME 1 then NONE else nonunitSubsumes pred nonunit max cl end; in fun subsumes pred subsume cl = genSubsumes pred subsume NONE cl; fun strictlySubsumes pred subsume cl = genSubsumes pred subsume (SOME (Metis_LiteralSet.size cl)) cl; end; (*MetisTrace4 val subsumes = fn pred => fn subsume => fn cl => let val ppCl = Metis_LiteralSet.pp val ppSub = Metis_Subst.pp val () = Metis_Print.trace ppCl "Metis_Subsume.subsumes: cl" cl val result = subsumes pred subsume cl val () = case result of NONE => trace "Metis_Subsume.subsumes: not subsumed\n" | SOME (cl,sub,_) => (Metis_Print.trace ppCl "Metis_Subsume.subsumes: subsuming cl" cl; Metis_Print.trace ppSub "Metis_Subsume.subsumes: subsuming sub" sub) in result end; val strictlySubsumes = fn pred => fn subsume => fn cl => let val ppCl = Metis_LiteralSet.pp val ppSub = Metis_Subst.pp val () = Metis_Print.trace ppCl "Metis_Subsume.strictlySubsumes: cl" cl val result = strictlySubsumes pred subsume cl val () = case result of NONE => trace "Metis_Subsume.subsumes: not subsumed\n" | SOME (cl,sub,_) => (Metis_Print.trace ppCl "Metis_Subsume.subsumes: subsuming cl" cl; Metis_Print.trace ppSub "Metis_Subsume.subsumes: subsuming sub" sub) in result end; *) fun isSubsumed subs cl = Option.isSome (subsumes (K true) subs cl); fun isStrictlySubsumed subs cl = Option.isSome (strictlySubsumes (K true) subs cl); (* ------------------------------------------------------------------------- *) (* Single clause versions. *) (* ------------------------------------------------------------------------- *) fun clauseSubsumes cl' cl = let val lits' = sortClause cl' and lits = clauseSym (Metis_LiteralSet.toList cl) in case genClauseSubsumes (K true) cl' lits' lits () of SOME (_,sub,()) => SOME sub | NONE => NONE end; fun clauseStrictlySubsumes cl' cl = if Metis_LiteralSet.size cl' > Metis_LiteralSet.size cl then NONE else clauseSubsumes cl' cl; end (**** Original file: src/KnuthBendixOrder.sig ****) (* ========================================================================= *) (* THE KNUTH-BENDIX TERM ORDERING *) (* Copyright (c) 2002 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_KnuthBendixOrder = sig (* ------------------------------------------------------------------------- *) (* The weight of all constants must be at least 1, and there must be at most *) (* one unary function with weight 0. *) (* ------------------------------------------------------------------------- *) type kbo = {weight : Metis_Term.function -> int, precedence : Metis_Term.function * Metis_Term.function -> order} val default : kbo val compare : kbo -> Metis_Term.term * Metis_Term.term -> order option end (**** Original file: src/KnuthBendixOrder.sml ****) (* ========================================================================= *) (* KNUTH-BENDIX TERM ORDERING CONSTRAINTS *) (* Copyright (c) 2002 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_KnuthBendixOrder :> Metis_KnuthBendixOrder = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* Helper functions. *) (* ------------------------------------------------------------------------- *) fun notEqualTerm (x,y) = not (Metis_Term.equal x y); fun firstNotEqualTerm f l = case List.find notEqualTerm l of SOME (x,y) => f x y | NONE => raise Bug "firstNotEqualTerm"; (* ------------------------------------------------------------------------- *) (* The weight of all constants must be at least 1, and there must be at most *) (* one unary function with weight 0. *) (* ------------------------------------------------------------------------- *) type kbo = {weight : Metis_Term.function -> int, precedence : Metis_Term.function * Metis_Term.function -> order}; (* Default weight = uniform *) val uniformWeight : Metis_Term.function -> int = K 1; (* Default precedence = by arity *) val arityPrecedence : Metis_Term.function * Metis_Term.function -> order = fn ((f1,n1),(f2,n2)) => case Int.compare (n1,n2) of LESS => LESS | EQUAL => Metis_Name.compare (f1,f2) | GREATER => GREATER; (* The default order *) val default = {weight = uniformWeight, precedence = arityPrecedence}; (* ------------------------------------------------------------------------- *) (* Metis_Term weight-1 represented as a linear function of the weight-1 of the *) (* variables in the term (plus a constant). *) (* *) (* Note that the conditions on weight functions ensure that all weights are *) (* at least 1, so all weight-1s are at least 0. *) (* ------------------------------------------------------------------------- *) datatype weight = Weight of int Metis_NameMap.map * int; val weightEmpty : int Metis_NameMap.map = Metis_NameMap.new (); val weightZero = Weight (weightEmpty,0); fun weightIsZero (Weight (m,c)) = c = 0 andalso Metis_NameMap.null m; fun weightNeg (Weight (m,c)) = Weight (Metis_NameMap.transform ~ m, ~c); local fun add ((_,n1),(_,n2)) = let val n = n1 + n2 in if n = 0 then NONE else SOME n end; in fun weightAdd (Weight (m1,c1)) (Weight (m2,c2)) = Weight (Metis_NameMap.union add m1 m2, c1 + c2); end; fun weightSubtract w1 w2 = weightAdd w1 (weightNeg w2); fun weightTerm weight = let fun wt m c [] = Weight (m,c) | wt m c (Metis_Term.Var v :: tms) = let val n = Option.getOpt (Metis_NameMap.peek m v, 0) in wt (Metis_NameMap.insert m (v, n + 1)) (c + 1) tms end | wt m c (Metis_Term.Fn (f,a) :: tms) = wt m (c + weight (f, length a)) (a @ tms) in fn tm => wt weightEmpty ~1 [tm] end; fun weightLowerBound (w as Weight (m,c)) = if Metis_NameMap.exists (fn (_,n) => n < 0) m then NONE else SOME c; (*MetisDebug val ppWeightList = let fun ppCoeff n = if n < 0 then Metis_Print.sequence (Metis_Print.ppString "~") (ppCoeff (~n)) else if n = 1 then Metis_Print.skip else Metis_Print.ppInt n fun pp_tm (NONE,n) = Metis_Print.ppInt n | pp_tm (SOME v, n) = Metis_Print.sequence (ppCoeff n) (Metis_Name.pp v) in fn [] => Metis_Print.ppInt 0 | tms => Metis_Print.ppOpList " +" pp_tm tms end; fun ppWeight (Weight (m,c)) = let val l = Metis_NameMap.toList m val l = List.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 = Metis_Print.toString ppWeight; *) (* ------------------------------------------------------------------------- *) (* The Knuth-Bendix term order. *) (* ------------------------------------------------------------------------- *) fun compare {weight,precedence} = let fun weightDifference tm1 tm2 = let val w1 = weightTerm weight tm1 and w2 = weightTerm weight tm2 in weightSubtract w2 w1 end fun weightLess tm1 tm2 = let val w = weightDifference tm1 tm2 in if weightIsZero w then precedenceLess tm1 tm2 else weightDiffLess w tm1 tm2 end and weightDiffLess w tm1 tm2 = case weightLowerBound w of NONE => false | SOME 0 => precedenceLess tm1 tm2 | SOME n => n > 0 and precedenceLess (Metis_Term.Fn (f1,a1)) (Metis_Term.Fn (f2,a2)) = (case precedence ((f1, length a1), (f2, length a2)) of LESS => true | EQUAL => firstNotEqualTerm weightLess (zip a1 a2) | GREATER => false) | precedenceLess _ _ = false fun weightDiffGreater w tm1 tm2 = weightDiffLess (weightNeg w) tm2 tm1 fun weightCmp tm1 tm2 = let val w = weightDifference tm1 tm2 in if weightIsZero w then precedenceCmp tm1 tm2 else if weightDiffLess w tm1 tm2 then SOME LESS else if weightDiffGreater w tm1 tm2 then SOME GREATER else NONE end and precedenceCmp (Metis_Term.Fn (f1,a1)) (Metis_Term.Fn (f2,a2)) = (case precedence ((f1, length a1), (f2, length a2)) of LESS => SOME LESS | EQUAL => firstNotEqualTerm weightCmp (zip a1 a2) | GREATER => SOME GREATER) | precedenceCmp _ _ = raise Bug "kboOrder.precendenceCmp" in fn (tm1,tm2) => if Metis_Term.equal tm1 tm2 then SOME EQUAL else weightCmp tm1 tm2 end; (*MetisTrace7 val compare = fn kbo => fn (tm1,tm2) => let val () = Metis_Print.trace Metis_Term.pp "Metis_KnuthBendixOrder.compare: tm1" tm1 val () = Metis_Print.trace Metis_Term.pp "Metis_KnuthBendixOrder.compare: tm2" tm2 val result = compare kbo (tm1,tm2) val () = case result of NONE => trace "Metis_KnuthBendixOrder.compare: result = Incomparable\n" | SOME x => Metis_Print.trace Metis_Print.ppOrder "Metis_KnuthBendixOrder.compare: result" x in result end; *) end (**** Original file: src/Rewrite.sig ****) (* ========================================================================= *) (* ORDERED REWRITING FOR FIRST ORDER TERMS *) (* Copyright (c) 2003 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Rewrite = sig (* ------------------------------------------------------------------------- *) (* 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 type equation = Metis_Rule.equation type rewrite (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val new : reductionOrder -> rewrite val peek : rewrite -> equationId -> (equation * orient option) option val size : rewrite -> int val equations : rewrite -> equation list val toString : rewrite -> string val pp : rewrite Metis_Print.pp (* ------------------------------------------------------------------------- *) (* Add equations into the system. *) (* ------------------------------------------------------------------------- *) val add : rewrite -> equationId * equation -> rewrite val addList : rewrite -> (equationId * equation) list -> rewrite (* ------------------------------------------------------------------------- *) (* Rewriting (the order must be a refinement of the rewrite order). *) (* ------------------------------------------------------------------------- *) val rewrConv : rewrite -> reductionOrder -> Metis_Rule.conv val rewriteConv : rewrite -> reductionOrder -> Metis_Rule.conv val rewriteLiteralsRule : rewrite -> reductionOrder -> Metis_LiteralSet.set -> Metis_Rule.rule val rewriteRule : rewrite -> reductionOrder -> Metis_Rule.rule val rewrIdConv : rewrite -> reductionOrder -> equationId -> Metis_Rule.conv val rewriteIdConv : rewrite -> reductionOrder -> equationId -> Metis_Rule.conv val rewriteIdLiteralsRule : rewrite -> reductionOrder -> equationId -> Metis_LiteralSet.set -> Metis_Rule.rule val rewriteIdRule : rewrite -> reductionOrder -> equationId -> Metis_Rule.rule (* ------------------------------------------------------------------------- *) (* Inter-reduce the equations in the system. *) (* ------------------------------------------------------------------------- *) val reduce' : rewrite -> rewrite * equationId list val reduce : rewrite -> rewrite val isReduced : rewrite -> bool (* ------------------------------------------------------------------------- *) (* Rewriting as a derived rule. *) (* ------------------------------------------------------------------------- *) val rewrite : equation list -> Metis_Thm.thm -> Metis_Thm.thm val orderedRewrite : reductionOrder -> equation list -> Metis_Thm.thm -> Metis_Thm.thm end (**** Original file: src/Rewrite.sml ****) (* ========================================================================= *) (* ORDERED REWRITING FOR FIRST ORDER TERMS *) (* Copyright (c) 2003 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Rewrite :> Metis_Rewrite = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* Orientations of equations. *) (* ------------------------------------------------------------------------- *) datatype orient = LeftToRight | RightToLeft; fun toStringOrient ort = case ort of LeftToRight => "-->" | RightToLeft => "<--"; val ppOrient = Metis_Print.ppMap toStringOrient Metis_Print.ppString; fun toStringOrientOption orto = case orto of SOME ort => toStringOrient ort | NONE => "<->"; val ppOrientOption = Metis_Print.ppMap toStringOrientOption Metis_Print.ppString; (* ------------------------------------------------------------------------- *) (* A type of rewrite systems. *) (* ------------------------------------------------------------------------- *) type reductionOrder = Metis_Term.term * Metis_Term.term -> order option; type equationId = int; type equation = Metis_Rule.equation; datatype rewrite = Metis_Rewrite of {order : reductionOrder, known : (equation * orient option) Metis_IntMap.map, redexes : (equationId * orient) Metis_TermNet.termNet, subterms : (equationId * bool * Metis_Term.path) Metis_TermNet.termNet, waiting : Metis_IntSet.set}; fun updateWaiting rw waiting = let val Metis_Rewrite {order, known, redexes, subterms, waiting = _} = rw in Metis_Rewrite {order = order, known = known, redexes = redexes, subterms = subterms, waiting = waiting} end; fun deleteWaiting (rw as Metis_Rewrite {waiting,...}) id = updateWaiting rw (Metis_IntSet.delete waiting id); (* ------------------------------------------------------------------------- *) (* Basic operations *) (* ------------------------------------------------------------------------- *) fun new order = Metis_Rewrite {order = order, known = Metis_IntMap.new (), redexes = Metis_TermNet.new {fifo = false}, subterms = Metis_TermNet.new {fifo = false}, waiting = Metis_IntSet.empty}; fun peek (Metis_Rewrite {known,...}) id = Metis_IntMap.peek known id; fun size (Metis_Rewrite {known,...}) = Metis_IntMap.size known; fun equations (Metis_Rewrite {known,...}) = Metis_IntMap.foldr (fn (_,(eqn,_),eqns) => eqn :: eqns) [] known; val pp = Metis_Print.ppMap equations (Metis_Print.ppList Metis_Rule.ppEquation); (*MetisTrace1 local fun ppEq ((x_y,_),ort) = Metis_Print.ppOp2 (" " ^ toStringOrientOption ort) Metis_Term.pp Metis_Term.pp x_y; fun ppField f ppA a = Metis_Print.inconsistentBlock 2 [Metis_Print.ppString (f ^ " ="), Metis_Print.break, ppA a]; val ppKnown = ppField "known" (Metis_Print.ppMap Metis_IntMap.toList (Metis_Print.ppList (Metis_Print.ppPair Metis_Print.ppInt ppEq))); val ppRedexes = ppField "redexes" (Metis_TermNet.pp (Metis_Print.ppPair Metis_Print.ppInt ppOrient)); val ppSubterms = ppField "subterms" (Metis_TermNet.pp (Metis_Print.ppMap (fn (i,l,p) => (i, (if l then 0 else 1) :: p)) (Metis_Print.ppPair Metis_Print.ppInt Metis_Term.ppPath))); val ppWaiting = ppField "waiting" (Metis_Print.ppMap (Metis_IntSet.toList) (Metis_Print.ppList Metis_Print.ppInt)); in fun pp (Metis_Rewrite {known,redexes,subterms,waiting,...}) = Metis_Print.inconsistentBlock 2 [Metis_Print.ppString "Metis_Rewrite", Metis_Print.break, Metis_Print.inconsistentBlock 1 [Metis_Print.ppString "{", ppKnown known, (*MetisTrace5 Metis_Print.ppString ",", Metis_Print.break, ppRedexes redexes, Metis_Print.ppString ",", Metis_Print.break, ppSubterms subterms, Metis_Print.ppString ",", Metis_Print.break, ppWaiting waiting, *) Metis_Print.skip], Metis_Print.ppString "}"] end; *) val toString = Metis_Print.toString pp; (* ------------------------------------------------------------------------- *) (* Debug functions. *) (* ------------------------------------------------------------------------- *) fun termReducible order known id = let fun eqnRed ((l,r),_) tm = case total (Metis_Subst.match Metis_Subst.empty l) tm of NONE => false | SOME sub => order (tm, Metis_Subst.subst (Metis_Subst.normalize sub) r) = SOME GREATER fun knownRed tm (eqnId,(eqn,ort)) = eqnId <> id andalso ((ort <> SOME RightToLeft andalso eqnRed eqn tm) orelse (ort <> SOME LeftToRight andalso eqnRed (Metis_Rule.symEqn eqn) tm)) fun termRed tm = Metis_IntMap.exists (knownRed tm) known orelse subtermRed tm and subtermRed (Metis_Term.Var _) = false | subtermRed (Metis_Term.Fn (_,tms)) = List.exists termRed tms in termRed end; fun literalReducible order known id lit = List.exists (termReducible order known id) (Metis_Literal.arguments lit); fun literalsReducible order known id lits = Metis_LiteralSet.exists (literalReducible order known id) lits; fun thmReducible order known id th = literalsReducible order known id (Metis_Thm.clause th); (* ------------------------------------------------------------------------- *) (* Add equations into the system. *) (* ------------------------------------------------------------------------- *) fun orderToOrient (SOME EQUAL) = raise Error "Metis_Rewrite.orient: reflexive" | orderToOrient (SOME GREATER) = SOME LeftToRight | orderToOrient (SOME LESS) = SOME RightToLeft | orderToOrient NONE = NONE; local fun ins redexes redex id ort = Metis_TermNet.insert redexes (redex,(id,ort)); in fun addRedexes id (((l,r),_),ort) redexes = case ort of SOME LeftToRight => ins redexes l id LeftToRight | SOME RightToLeft => ins redexes r id RightToLeft | NONE => ins (ins redexes l id LeftToRight) r id RightToLeft; end; fun add (rw as Metis_Rewrite {known,...}) (id,eqn) = if Metis_IntMap.inDomain id known then rw else let val Metis_Rewrite {order,redexes,subterms,waiting, ...} = rw val ort = orderToOrient (order (fst eqn)) val known = Metis_IntMap.insert known (id,(eqn,ort)) val redexes = addRedexes id (eqn,ort) redexes val waiting = Metis_IntSet.add waiting id val rw = Metis_Rewrite {order = order, known = known, redexes = redexes, subterms = subterms, waiting = waiting} (*MetisTrace5 val () = Metis_Print.trace pp "Metis_Rewrite.add: result" rw *) in rw end; local fun uncurriedAdd (eqn,rw) = add rw eqn; in fun addList rw = List.foldl uncurriedAdd rw; end; (* ------------------------------------------------------------------------- *) (* Rewriting (the supplied order must be a refinement of the rewrite order). *) (* ------------------------------------------------------------------------- *) local fun reorder ((i,_),(j,_)) = Int.compare (j,i); in fun matchingRedexes redexes tm = sort reorder (Metis_TermNet.match redexes tm); end; fun wellOriented NONE _ = true | wellOriented (SOME LeftToRight) LeftToRight = true | wellOriented (SOME RightToLeft) RightToLeft = true | wellOriented _ _ = false; fun redexResidue LeftToRight ((l_r,_) : equation) = l_r | redexResidue RightToLeft ((l,r),_) = (r,l); fun orientedEquation LeftToRight eqn = eqn | orientedEquation RightToLeft eqn = Metis_Rule.symEqn eqn; fun rewrIdConv' order known redexes id tm = let fun rewr (id',lr) = let val _ = id <> id' orelse raise Error "same theorem" val (eqn,ort) = Metis_IntMap.get known id' val _ = wellOriented ort lr orelse raise Error "orientation" val (l,r) = redexResidue lr eqn val sub = Metis_Subst.normalize (Metis_Subst.match Metis_Subst.empty l tm) val tm' = Metis_Subst.subst sub r val _ = Option.isSome ort orelse order (tm,tm') = SOME GREATER orelse raise Error "order" val (_,th) = orientedEquation lr eqn in (tm', Metis_Thm.subst sub th) end in case first (total rewr) (matchingRedexes redexes tm) of NONE => raise Error "Metis_Rewrite.rewrIdConv: no matching rewrites" | SOME res => res end; fun rewriteIdConv' order known redexes id = if Metis_IntMap.null known then Metis_Rule.allConv else Metis_Rule.repeatTopDownConv (rewrIdConv' order known redexes id); fun mkNeqConv order lit = let val (l,r) = Metis_Literal.destNeq lit in case order (l,r) of NONE => raise Error "incomparable" | SOME LESS => let val th = Metis_Rule.symmetryRule l r in fn tm => if Metis_Term.equal tm r then (l,th) else raise Error "mkNeqConv: RL" end | SOME EQUAL => raise Error "irreflexive" | SOME GREATER => let val th = Metis_Thm.assume lit in fn tm => if Metis_Term.equal tm l then (r,th) else raise Error "mkNeqConv: LR" end end; datatype neqConvs = NeqConvs of Metis_Rule.conv list; val neqConvsEmpty = NeqConvs []; fun neqConvsNull (NeqConvs l) = List.null l; fun neqConvsAdd order (neq as NeqConvs l) lit = case total (mkNeqConv order) lit of NONE => neq | SOME conv => NeqConvs (conv :: l); fun mkNeqConvs order = let fun add (lit,neq) = neqConvsAdd order neq lit in Metis_LiteralSet.foldl add neqConvsEmpty end; fun buildNeqConvs order lits = let fun add (lit,(neq,neqs)) = (neqConvsAdd order neq lit, (lit,neq) :: neqs) in snd (Metis_LiteralSet.foldl add (neqConvsEmpty,[]) lits) end; fun neqConvsToConv (NeqConvs l) = Metis_Rule.firstConv l; fun neqConvsUnion (NeqConvs l1) (NeqConvs l2) = NeqConvs (List.revAppend (l1,l2)); fun neqConvsRewrIdLiterule order known redexes id neq = if Metis_IntMap.null known andalso neqConvsNull neq then Metis_Rule.allLiterule else let val neq_conv = neqConvsToConv neq val rewr_conv = rewrIdConv' order known redexes id val conv = Metis_Rule.orelseConv neq_conv rewr_conv val conv = Metis_Rule.repeatTopDownConv conv in Metis_Rule.allArgumentsLiterule conv end; fun rewriteIdEqn' order known redexes id (eqn as (l_r,th)) = let val neq = mkNeqConvs order (Metis_Thm.clause th) val literule = neqConvsRewrIdLiterule order known redexes id neq val (strongEqn,lit) = case Metis_Rule.equationLiteral eqn of NONE => (true, Metis_Literal.mkEq l_r) | SOME lit => (false,lit) val (lit',litTh) = literule lit in if Metis_Literal.equal lit lit' then eqn else (Metis_Literal.destEq lit', if strongEqn then th else if not (Metis_Thm.negateMember lit litTh) then litTh else Metis_Thm.resolve lit th litTh) end (*MetisDebug handle Error err => raise Error ("Metis_Rewrite.rewriteIdEqn':\n" ^ err); *) fun rewriteIdLiteralsRule' order known redexes id lits th = let val mk_literule = neqConvsRewrIdLiterule order known redexes id fun rewr_neq_lit ((lit,rneq),(changed,lneq,lits,th)) = let val neq = neqConvsUnion lneq rneq val (lit',litTh) = mk_literule neq lit val lneq = neqConvsAdd order lneq lit' val lits = Metis_LiteralSet.add lits lit' in if Metis_Literal.equal lit lit' then (changed,lneq,lits,th) else (true, lneq, lits, if Metis_Thm.member lit th then Metis_Thm.resolve lit th litTh else th) end fun rewr_neq_lits lits th = let val neqs = buildNeqConvs order lits val neq = neqConvsEmpty val lits = Metis_LiteralSet.empty val (changed,neq,lits,th) = List.foldl rewr_neq_lit (false,neq,lits,th) neqs in if changed then rewr_neq_lits lits th else (neq,th) end val (neq,lits) = Metis_LiteralSet.partition Metis_Literal.isNeq lits val (neq,th) = rewr_neq_lits neq th val rewr_literule = mk_literule neq fun rewr_lit (lit,th) = if not (Metis_Thm.member lit th) then th else Metis_Rule.literalRule rewr_literule lit th in Metis_LiteralSet.foldl rewr_lit th lits end; fun rewriteIdRule' order known redexes id th = rewriteIdLiteralsRule' order known redexes id (Metis_Thm.clause th) th; (*MetisDebug val rewriteIdRule' = fn order => fn known => fn redexes => fn id => fn th => let (*MetisTrace6 val () = Metis_Print.trace Metis_Thm.pp "Metis_Rewrite.rewriteIdRule': th" th *) val result = rewriteIdRule' order known redexes id th (*MetisTrace6 val () = Metis_Print.trace Metis_Thm.pp "Metis_Rewrite.rewriteIdRule': result" result *) val () = if not (thmReducible order known id result) then () else raise Bug "Metis_Rewrite.rewriteIdRule': should be normalized" in result end handle Error err => raise Error ("Metis_Rewrite.rewriteIdRule':\n" ^ "th = " ^ Metis_Thm.toString th ^ "\n" ^ err) | Bug bug => raise Bug ("Metis_Rewrite.rewriteIdRule':\n" ^ "th = " ^ Metis_Thm.toString th ^ "\n" ^ bug); *) fun rewrIdConv (Metis_Rewrite {known,redexes,...}) order = rewrIdConv' order known redexes; fun rewrConv rewrite order = rewrIdConv rewrite order ~1; fun rewriteIdConv (Metis_Rewrite {known,redexes,...}) order = rewriteIdConv' order known redexes; fun rewriteConv rewrite order = rewriteIdConv rewrite order ~1; fun rewriteIdLiteralsRule (Metis_Rewrite {known,redexes,...}) order = rewriteIdLiteralsRule' order known redexes; fun rewriteLiteralsRule rewrite order = rewriteIdLiteralsRule rewrite order ~1; fun rewriteIdRule (Metis_Rewrite {known,redexes,...}) order = rewriteIdRule' order known redexes; (*MetisDebug val rewriteIdRule = fn rewr => fn order => fn id => fn th => let val result = rewriteIdRule rewr order id th val th' = rewriteIdRule rewr order id result val () = if Metis_Thm.equal th' result then () else let fun trace p s = Metis_Print.trace p ("Metis_Rewrite.rewriteIdRule: "^s) val () = trace pp "rewr" rewr val () = trace Metis_Thm.pp "th" th val () = trace Metis_Thm.pp "result" result val () = trace Metis_Thm.pp "th'" th' in raise Bug "Metis_Rewrite.rewriteIdRule: should be idempotent" end in result end handle Error err => raise Error ("Metis_Rewrite.rewriteIdRule:\n" ^ err); *) fun rewriteRule rewrite order = rewriteIdRule rewrite order ~1; (* ------------------------------------------------------------------------- *) (* Inter-reduce the equations in the system. *) (* ------------------------------------------------------------------------- *) fun addSubterms id (((l,r),_) : equation) subterms = let fun addSubterm b ((path,tm),net) = Metis_TermNet.insert net (tm,(id,b,path)) val subterms = List.foldl (addSubterm true) subterms (Metis_Term.subterms l) val subterms = List.foldl (addSubterm false) subterms (Metis_Term.subterms r) in subterms end; fun sameRedexes NONE _ _ = false | sameRedexes (SOME LeftToRight) (l0,_) (l,_) = Metis_Term.equal l0 l | sameRedexes (SOME RightToLeft) (_,r0) (_,r) = Metis_Term.equal r0 r; fun redexResidues NONE (l,r) = [(l,r,false),(r,l,false)] | redexResidues (SOME LeftToRight) (l,r) = [(l,r,true)] | redexResidues (SOME RightToLeft) (l,r) = [(r,l,true)]; fun findReducibles order known subterms id = let fun checkValidRewr (l,r,ord) id' left path = let val (((x,y),_),_) = Metis_IntMap.get known id' val tm = Metis_Term.subterm (if left then x else y) path val sub = Metis_Subst.match Metis_Subst.empty l tm in if ord then () else let val tm' = Metis_Subst.subst (Metis_Subst.normalize sub) r in if order (tm,tm') = SOME GREATER then () else raise Error "order" end end fun addRed lr ((id',left,path),todo) = if id <> id' andalso not (Metis_IntSet.member id' todo) andalso can (checkValidRewr lr id' left) path then Metis_IntSet.add todo id' else todo fun findRed (lr as (l,_,_), todo) = List.foldl (addRed lr) todo (Metis_TermNet.matched subterms l) in List.foldl findRed end; fun reduce1 new id (eqn0,ort0) (rpl,spl,todo,rw,changed) = let val (eq0,_) = eqn0 val Metis_Rewrite {order,known,redexes,subterms,waiting} = rw val eqn as (eq,_) = rewriteIdEqn' order known redexes id eqn0 val identical = let val (l0,r0) = eq0 and (l,r) = eq in Metis_Term.equal l l0 andalso Metis_Term.equal r r0 end val same_redexes = identical orelse sameRedexes ort0 eq0 eq val rpl = if same_redexes then rpl else Metis_IntSet.add rpl id val spl = if new orelse identical then spl else Metis_IntSet.add spl id val changed = if not new andalso identical then changed else Metis_IntSet.add changed id val ort = if same_redexes then SOME ort0 else total orderToOrient (order eq) in case ort of NONE => let val known = Metis_IntMap.delete known id val rw = Metis_Rewrite {order = order, known = known, redexes = redexes, subterms = subterms, waiting = waiting} in (rpl,spl,todo,rw,changed) end | SOME ort => let val todo = if not new andalso same_redexes then todo else findReducibles order known subterms id todo (redexResidues ort eq) val known = if identical then known else Metis_IntMap.insert known (id,(eqn,ort)) val redexes = if same_redexes then redexes else addRedexes id (eqn,ort) redexes val subterms = if new orelse not identical then addSubterms id eqn subterms else subterms val rw = Metis_Rewrite {order = order, known = known, redexes = redexes, subterms = subterms, waiting = waiting} in (rpl,spl,todo,rw,changed) end end; fun pick known set = let fun oriented id = case Metis_IntMap.peek known id of SOME (x as (_, SOME _)) => SOME (id,x) | _ => NONE fun any id = case Metis_IntMap.peek known id of SOME x => SOME (id,x) | _ => NONE in case Metis_IntSet.firstl oriented set of x as SOME _ => x | NONE => Metis_IntSet.firstl any set end; local fun cleanRedexes known redexes rpl = if Metis_IntSet.null rpl then redexes else let fun filt (id,_) = not (Metis_IntSet.member id rpl) fun addReds (id,reds) = case Metis_IntMap.peek known id of NONE => reds | SOME eqn_ort => addRedexes id eqn_ort reds val redexes = Metis_TermNet.filter filt redexes val redexes = Metis_IntSet.foldl addReds redexes rpl in redexes end; fun cleanSubterms known subterms spl = if Metis_IntSet.null spl then subterms else let fun filt (id,_,_) = not (Metis_IntSet.member id spl) fun addSubtms (id,subtms) = case Metis_IntMap.peek known id of NONE => subtms | SOME (eqn,_) => addSubterms id eqn subtms val subterms = Metis_TermNet.filter filt subterms val subterms = Metis_IntSet.foldl addSubtms subterms spl in subterms end; in fun rebuild rpl spl rw = let (*MetisTrace5 val ppPl = Metis_Print.ppMap Metis_IntSet.toList (Metis_Print.ppList Metis_Print.ppInt) val () = Metis_Print.trace ppPl "Metis_Rewrite.rebuild: rpl" rpl val () = Metis_Print.trace ppPl "Metis_Rewrite.rebuild: spl" spl *) val Metis_Rewrite {order,known,redexes,subterms,waiting} = rw val redexes = cleanRedexes known redexes rpl val subterms = cleanSubterms known subterms spl in Metis_Rewrite {order = order, known = known, redexes = redexes, subterms = subterms, waiting = waiting} end; end; fun reduceAcc (rpl, spl, todo, rw as Metis_Rewrite {known,waiting,...}, changed) = case pick known todo of SOME (id,eqn_ort) => let val todo = Metis_IntSet.delete todo id in reduceAcc (reduce1 false id eqn_ort (rpl,spl,todo,rw,changed)) end | NONE => case pick known waiting of SOME (id,eqn_ort) => let val rw = deleteWaiting rw id in reduceAcc (reduce1 true id eqn_ort (rpl,spl,todo,rw,changed)) end | NONE => (rebuild rpl spl rw, Metis_IntSet.toList changed); fun isReduced (Metis_Rewrite {waiting,...}) = Metis_IntSet.null waiting; fun reduce' rw = if isReduced rw then (rw,[]) else reduceAcc (Metis_IntSet.empty,Metis_IntSet.empty,Metis_IntSet.empty,rw,Metis_IntSet.empty); (*MetisDebug val reduce' = fn rw => let (*MetisTrace4 val () = Metis_Print.trace pp "Metis_Rewrite.reduce': rw" rw *) val Metis_Rewrite {known,order,...} = rw val result as (Metis_Rewrite {known = known', ...}, _) = reduce' rw (*MetisTrace4 val ppResult = Metis_Print.ppPair pp (Metis_Print.ppList Metis_Print.ppInt) val () = Metis_Print.trace ppResult "Metis_Rewrite.reduce': result" result *) val ths = List.map (fn (id,((_,th),_)) => (id,th)) (Metis_IntMap.toList known') val _ = not (List.exists (uncurry (thmReducible order known')) ths) orelse raise Bug "Metis_Rewrite.reduce': not fully reduced" in result end handle Error err => raise Bug ("Metis_Rewrite.reduce': shouldn't fail\n" ^ err); *) fun reduce rw = fst (reduce' rw); (* ------------------------------------------------------------------------- *) (* Rewriting as a derived rule. *) (* ------------------------------------------------------------------------- *) local fun addEqn (id_eqn,rw) = add rw id_eqn; in fun orderedRewrite order ths = let val rw = List.foldl addEqn (new order) (enumerate ths) in rewriteRule rw order end; end; local val order : reductionOrder = K (SOME GREATER); in val rewrite = orderedRewrite order; end; end (**** Original file: src/Units.sig ****) (* ========================================================================= *) (* A STORE FOR UNIT THEOREMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Units = sig (* ------------------------------------------------------------------------- *) (* A type of unit store. *) (* ------------------------------------------------------------------------- *) type unitThm = Metis_Literal.literal * Metis_Thm.thm type units (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val empty : units val size : units -> int val toString : units -> string val pp : units Metis_Print.pp (* ------------------------------------------------------------------------- *) (* Add units into the store. *) (* ------------------------------------------------------------------------- *) val add : units -> unitThm -> units val addList : units -> unitThm list -> units (* ------------------------------------------------------------------------- *) (* Matching. *) (* ------------------------------------------------------------------------- *) val match : units -> Metis_Literal.literal -> (unitThm * Metis_Subst.subst) option (* ------------------------------------------------------------------------- *) (* Reducing by repeated matching and resolution. *) (* ------------------------------------------------------------------------- *) val reduce : units -> Metis_Rule.rule end (**** Original file: src/Units.sml ****) (* ========================================================================= *) (* A STORE FOR UNIT THEOREMS *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Units :> Metis_Units = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* A type of unit store. *) (* ------------------------------------------------------------------------- *) type unitThm = Metis_Literal.literal * Metis_Thm.thm; datatype units = Metis_Units of unitThm Metis_LiteralNet.literalNet; (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val empty = Metis_Units (Metis_LiteralNet.new {fifo = false}); fun size (Metis_Units net) = Metis_LiteralNet.size net; fun toString units = "U{" ^ Int.toString (size units) ^ "}"; val pp = Metis_Print.ppMap toString Metis_Print.ppString; (* ------------------------------------------------------------------------- *) (* Add units into the store. *) (* ------------------------------------------------------------------------- *) fun add (units as Metis_Units net) (uTh as (lit,th)) = let val net = Metis_LiteralNet.insert net (lit,uTh) in case total Metis_Literal.sym lit of NONE => Metis_Units net | SOME (lit' as (pol,_)) => let val th' = (if pol then Metis_Rule.symEq else Metis_Rule.symNeq) lit th val net = Metis_LiteralNet.insert net (lit',(lit',th')) in Metis_Units net end end; val addList = List.foldl (fn (th,u) => add u th); (* ------------------------------------------------------------------------- *) (* Matching. *) (* ------------------------------------------------------------------------- *) fun match (Metis_Units net) lit = let fun check (uTh as (lit',_)) = case total (Metis_Literal.match Metis_Subst.empty lit') lit of NONE => NONE | SOME sub => SOME (uTh,sub) in first check (Metis_LiteralNet.match net lit) end; (* ------------------------------------------------------------------------- *) (* Reducing by repeated matching and resolution. *) (* ------------------------------------------------------------------------- *) fun reduce units = let fun red1 (lit,news_th) = case total Metis_Literal.destIrrefl lit of SOME tm => let val (news,th) = news_th val th = Metis_Thm.resolve lit th (Metis_Thm.refl tm) in (news,th) end | NONE => let val lit' = Metis_Literal.negate lit in case match units lit' of NONE => news_th | SOME ((_,rth),sub) => let val (news,th) = news_th val rth = Metis_Thm.subst sub rth val th = Metis_Thm.resolve lit th rth val new = Metis_LiteralSet.delete (Metis_Thm.clause rth) lit' val news = Metis_LiteralSet.union new news in (news,th) end end fun red (news,th) = if Metis_LiteralSet.null news then th else red (Metis_LiteralSet.foldl red1 (Metis_LiteralSet.empty,th) news) in fn th => Metis_Rule.removeSym (red (Metis_Thm.clause th, th)) end; end (**** Original file: src/Clause.sig ****) (* ========================================================================= *) (* CLAUSE = ID + THEOREM *) (* Copyright (c) 2002 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Clause = sig (* ------------------------------------------------------------------------- *) (* A type of clause. *) (* ------------------------------------------------------------------------- *) datatype literalOrder = NoLiteralOrder | UnsignedLiteralOrder | PositiveLiteralOrder type parameters = {ordering : Metis_KnuthBendixOrder.kbo, orderLiterals : literalOrder, orderTerms : bool} type clauseId = int type clauseInfo = {parameters : parameters, id : clauseId, thm : Metis_Thm.thm} type clause (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val default : parameters val newId : unit -> clauseId val mk : clauseInfo -> clause val dest : clause -> clauseInfo val id : clause -> clauseId val thm : clause -> Metis_Thm.thm val equalThms : clause -> clause -> bool val literals : clause -> Metis_Thm.clause val isTautology : clause -> bool val isContradiction : clause -> bool (* ------------------------------------------------------------------------- *) (* The term ordering is used to cut down inferences. *) (* ------------------------------------------------------------------------- *) val largestLiterals : clause -> Metis_LiteralSet.set val largestEquations : clause -> (Metis_Literal.literal * Metis_Rewrite.orient * Metis_Term.term) list val largestSubterms : clause -> (Metis_Literal.literal * Metis_Term.path * Metis_Term.term) list val allSubterms : clause -> (Metis_Literal.literal * Metis_Term.path * Metis_Term.term) list (* ------------------------------------------------------------------------- *) (* Subsumption. *) (* ------------------------------------------------------------------------- *) val subsumes : clause Metis_Subsume.subsume -> clause -> bool (* ------------------------------------------------------------------------- *) (* Simplifying rules: these preserve the clause id. *) (* ------------------------------------------------------------------------- *) val freshVars : clause -> clause val simplify : clause -> clause option val reduce : Metis_Units.units -> clause -> clause val rewrite : Metis_Rewrite.rewrite -> clause -> clause (* ------------------------------------------------------------------------- *) (* Inference rules: these generate new clause ids. *) (* ------------------------------------------------------------------------- *) val factor : clause -> clause list val resolve : clause * Metis_Literal.literal -> clause * Metis_Literal.literal -> clause val paramodulate : clause * Metis_Literal.literal * Metis_Rewrite.orient * Metis_Term.term -> clause * Metis_Literal.literal * Metis_Term.path * Metis_Term.term -> clause (* ------------------------------------------------------------------------- *) (* Pretty printing. *) (* ------------------------------------------------------------------------- *) val showId : bool Unsynchronized.ref val pp : clause Metis_Print.pp val toString : clause -> string end (**** Original file: src/Clause.sml ****) (* ========================================================================= *) (* CLAUSE = ID + THEOREM *) (* Copyright (c) 2002 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Clause :> Metis_Clause = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* Helper functions. *) (* ------------------------------------------------------------------------- *) val newId = let val r = Unsynchronized.ref 0 fun new () = let val Unsynchronized.ref n = r val () = r := n + 1 in n end in fn () => Metis_Portable.critical new () end; (* ------------------------------------------------------------------------- *) (* A type of clause. *) (* ------------------------------------------------------------------------- *) datatype literalOrder = NoLiteralOrder | UnsignedLiteralOrder | PositiveLiteralOrder; type parameters = {ordering : Metis_KnuthBendixOrder.kbo, orderLiterals : literalOrder, orderTerms : bool}; type clauseId = int; type clauseInfo = {parameters : parameters, id : clauseId, thm : Metis_Thm.thm}; datatype clause = Metis_Clause of clauseInfo; (* ------------------------------------------------------------------------- *) (* Pretty printing. *) (* ------------------------------------------------------------------------- *) val showId = Unsynchronized.ref false; local val ppIdThm = Metis_Print.ppPair Metis_Print.ppInt Metis_Thm.pp; in fun pp (Metis_Clause {id,thm,...}) = if !showId then ppIdThm (id,thm) else Metis_Thm.pp thm; end; fun toString cl = Metis_Print.toString pp cl; (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val default : parameters = {ordering = Metis_KnuthBendixOrder.default, orderLiterals = PositiveLiteralOrder, orderTerms = true}; fun mk info = Metis_Clause info fun dest (Metis_Clause info) = info; fun id (Metis_Clause {id = i, ...}) = i; fun thm (Metis_Clause {thm = th, ...}) = th; fun equalThms cl cl' = Metis_Thm.equal (thm cl) (thm cl'); fun new parameters thm = Metis_Clause {parameters = parameters, id = newId (), thm = thm}; fun literals cl = Metis_Thm.clause (thm cl); fun isTautology (Metis_Clause {thm,...}) = Metis_Thm.isTautology thm; fun isContradiction (Metis_Clause {thm,...}) = Metis_Thm.isContradiction thm; (* ------------------------------------------------------------------------- *) (* The term ordering is used to cut down inferences. *) (* ------------------------------------------------------------------------- *) fun strictlyLess ordering x_y = case Metis_KnuthBendixOrder.compare ordering x_y of SOME LESS => true | _ => false; fun isLargerTerm ({ordering,orderTerms,...} : parameters) l_r = not orderTerms orelse not (strictlyLess ordering l_r); local fun atomToTerms atm = case total Metis_Atom.destEq atm of NONE => [Metis_Term.Fn atm] | SOME (l,r) => [l,r]; fun notStrictlyLess ordering (xs,ys) = let fun less x = List.exists (fn y => strictlyLess ordering (x,y)) ys in not (List.all less xs) end; in fun isLargerLiteral ({ordering,orderLiterals,...} : parameters) lits = case orderLiterals of NoLiteralOrder => K true | UnsignedLiteralOrder => let fun addLit ((_,atm),acc) = atomToTerms atm @ acc val tms = Metis_LiteralSet.foldl addLit [] lits in fn (_,atm') => notStrictlyLess ordering (atomToTerms atm', tms) end | PositiveLiteralOrder => case Metis_LiteralSet.findl (K true) lits of NONE => K true | SOME (pol,_) => let fun addLit ((p,atm),acc) = if p = pol then atomToTerms atm @ acc else acc val tms = Metis_LiteralSet.foldl addLit [] lits in fn (pol',atm') => if pol <> pol' then pol else notStrictlyLess ordering (atomToTerms atm', tms) end; end; fun largestLiterals (Metis_Clause {parameters,thm,...}) = let val litSet = Metis_Thm.clause thm val isLarger = isLargerLiteral parameters litSet fun addLit (lit,s) = if isLarger lit then Metis_LiteralSet.add s lit else s in Metis_LiteralSet.foldr addLit Metis_LiteralSet.empty litSet end; (*MetisTrace6 val largestLiterals = fn cl => let val ppResult = Metis_LiteralSet.pp val () = Metis_Print.trace pp "Metis_Clause.largestLiterals: cl" cl val result = largestLiterals cl val () = Metis_Print.trace ppResult "Metis_Clause.largestLiterals: result" result in result end; *) fun largestEquations (cl as Metis_Clause {parameters,...}) = let fun addEq lit ort (l_r as (l,_)) acc = if isLargerTerm parameters l_r then (lit,ort,l) :: acc else acc fun addLit (lit,acc) = case total Metis_Literal.destEq lit of NONE => acc | SOME (l,r) => let val acc = addEq lit Metis_Rewrite.RightToLeft (r,l) acc val acc = addEq lit Metis_Rewrite.LeftToRight (l,r) acc in acc end in Metis_LiteralSet.foldr addLit [] (largestLiterals cl) end; local fun addLit (lit,acc) = let fun addTm ((path,tm),acc) = (lit,path,tm) :: acc in List.foldl addTm acc (Metis_Literal.nonVarTypedSubterms lit) end; in fun largestSubterms cl = Metis_LiteralSet.foldl addLit [] (largestLiterals cl); fun allSubterms cl = Metis_LiteralSet.foldl addLit [] (literals cl); end; (* ------------------------------------------------------------------------- *) (* Subsumption. *) (* ------------------------------------------------------------------------- *) fun subsumes (subs : clause Metis_Subsume.subsume) cl = Metis_Subsume.isStrictlySubsumed subs (literals cl); (* ------------------------------------------------------------------------- *) (* Simplifying rules: these preserve the clause id. *) (* ------------------------------------------------------------------------- *) fun freshVars (Metis_Clause {parameters,id,thm}) = Metis_Clause {parameters = parameters, id = id, thm = Metis_Rule.freshVars thm}; fun simplify (Metis_Clause {parameters,id,thm}) = case Metis_Rule.simplify thm of NONE => NONE | SOME thm => SOME (Metis_Clause {parameters = parameters, id = id, thm = thm}); fun reduce units (Metis_Clause {parameters,id,thm}) = Metis_Clause {parameters = parameters, id = id, thm = Metis_Units.reduce units thm}; local fun simp rewr (parm : parameters) id th = let val {ordering,...} = parm val cmp = Metis_KnuthBendixOrder.compare ordering in Metis_Rewrite.rewriteIdRule rewr cmp id th end; in fun rewrite rewr cl = let val Metis_Clause {parameters = parm, id, thm = th} = cl (*MetisTrace4 val () = Metis_Print.trace Metis_Rewrite.pp "Metis_Clause.rewrite: rewr" rewr val () = Metis_Print.trace Metis_Print.ppInt "Metis_Clause.rewrite: id" id val () = Metis_Print.trace pp "Metis_Clause.rewrite: cl" cl *) val th = case Metis_Rewrite.peek rewr id of NONE => simp rewr parm id th | SOME ((_,th),_) => if Metis_Rewrite.isReduced rewr then th else simp rewr parm id th val result = Metis_Clause {parameters = parm, id = id, thm = th} (*MetisTrace4 val () = Metis_Print.trace pp "Metis_Clause.rewrite: result" result *) in result end (*MetisDebug handle Error err => raise Error ("Metis_Clause.rewrite:\n" ^ err); *) end; (* ------------------------------------------------------------------------- *) (* Inference rules: these generate new clause ids. *) (* ------------------------------------------------------------------------- *) fun factor (cl as Metis_Clause {parameters,thm,...}) = let val lits = largestLiterals cl fun apply sub = new parameters (Metis_Thm.subst sub thm) in List.map apply (Metis_Rule.factor' lits) end; (*MetisTrace5 val factor = fn cl => let val () = Metis_Print.trace pp "Metis_Clause.factor: cl" cl val result = factor cl val () = Metis_Print.trace (Metis_Print.ppList pp) "Metis_Clause.factor: result" result in result end; *) fun resolve (cl1,lit1) (cl2,lit2) = let (*MetisTrace5 val () = Metis_Print.trace pp "Metis_Clause.resolve: cl1" cl1 val () = Metis_Print.trace Metis_Literal.pp "Metis_Clause.resolve: lit1" lit1 val () = Metis_Print.trace pp "Metis_Clause.resolve: cl2" cl2 val () = Metis_Print.trace Metis_Literal.pp "Metis_Clause.resolve: lit2" lit2 *) val Metis_Clause {parameters, thm = th1, ...} = cl1 and Metis_Clause {thm = th2, ...} = cl2 val sub = Metis_Literal.unify Metis_Subst.empty lit1 (Metis_Literal.negate lit2) (*MetisTrace5 val () = Metis_Print.trace Metis_Subst.pp "Metis_Clause.resolve: sub" sub *) val lit1 = Metis_Literal.subst sub lit1 val lit2 = Metis_Literal.negate lit1 val th1 = Metis_Thm.subst sub th1 and th2 = Metis_Thm.subst sub th2 val _ = isLargerLiteral parameters (Metis_Thm.clause th1) lit1 orelse (*MetisTrace5 (trace "Metis_Clause.resolve: th1 violates ordering\n"; false) orelse *) raise Error "resolve: clause1: ordering constraints" val _ = isLargerLiteral parameters (Metis_Thm.clause th2) lit2 orelse (*MetisTrace5 (trace "Metis_Clause.resolve: th2 violates ordering\n"; false) orelse *) raise Error "resolve: clause2: ordering constraints" val th = Metis_Thm.resolve lit1 th1 th2 (*MetisTrace5 val () = Metis_Print.trace Metis_Thm.pp "Metis_Clause.resolve: th" th *) val cl = Metis_Clause {parameters = parameters, id = newId (), thm = th} (*MetisTrace5 val () = Metis_Print.trace pp "Metis_Clause.resolve: cl" cl *) in cl end; fun paramodulate (cl1,lit1,ort1,tm1) (cl2,lit2,path2,tm2) = let (*MetisTrace5 val () = Metis_Print.trace pp "Metis_Clause.paramodulate: cl1" cl1 val () = Metis_Print.trace Metis_Literal.pp "Metis_Clause.paramodulate: lit1" lit1 val () = Metis_Print.trace Metis_Rewrite.ppOrient "Metis_Clause.paramodulate: ort1" ort1 val () = Metis_Print.trace Metis_Term.pp "Metis_Clause.paramodulate: tm1" tm1 val () = Metis_Print.trace pp "Metis_Clause.paramodulate: cl2" cl2 val () = Metis_Print.trace Metis_Literal.pp "Metis_Clause.paramodulate: lit2" lit2 val () = Metis_Print.trace Metis_Term.ppPath "Metis_Clause.paramodulate: path2" path2 val () = Metis_Print.trace Metis_Term.pp "Metis_Clause.paramodulate: tm2" tm2 *) val Metis_Clause {parameters, thm = th1, ...} = cl1 and Metis_Clause {thm = th2, ...} = cl2 val sub = Metis_Subst.unify Metis_Subst.empty tm1 tm2 val lit1 = Metis_Literal.subst sub lit1 and lit2 = Metis_Literal.subst sub lit2 and th1 = Metis_Thm.subst sub th1 and th2 = Metis_Thm.subst sub th2 val _ = isLargerLiteral parameters (Metis_Thm.clause th1) lit1 orelse raise Error "Metis_Clause.paramodulate: with clause: ordering" val _ = isLargerLiteral parameters (Metis_Thm.clause th2) lit2 orelse raise Error "Metis_Clause.paramodulate: into clause: ordering" val eqn = (Metis_Literal.destEq lit1, th1) val eqn as (l_r,_) = case ort1 of Metis_Rewrite.LeftToRight => eqn | Metis_Rewrite.RightToLeft => Metis_Rule.symEqn eqn (*MetisTrace6 val () = Metis_Print.trace Metis_Rule.ppEquation "Metis_Clause.paramodulate: eqn" eqn *) val _ = isLargerTerm parameters l_r orelse raise Error "Metis_Clause.paramodulate: equation: ordering constraints" val th = Metis_Rule.rewrRule eqn lit2 path2 th2 (*MetisTrace5 val () = Metis_Print.trace Metis_Thm.pp "Metis_Clause.paramodulate: th" th *) in Metis_Clause {parameters = parameters, id = newId (), thm = th} end (*MetisTrace5 handle Error err => let val () = trace ("Metis_Clause.paramodulate: failed: " ^ err ^ "\n") in raise Error err end; *) end (**** Original file: src/Active.sig ****) (* ========================================================================= *) (* THE ACTIVE SET OF CLAUSES *) (* Copyright (c) 2002 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Active = sig (* ------------------------------------------------------------------------- *) (* A type of active clause sets. *) (* ------------------------------------------------------------------------- *) type simplify = {subsume : bool, reduce : bool, rewrite : bool} type parameters = {clause : Metis_Clause.parameters, prefactor : simplify, postfactor : simplify} type active (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val default : parameters val size : active -> int val saturation : active -> Metis_Clause.clause list (* ------------------------------------------------------------------------- *) (* Create a new active clause set and initialize clauses. *) (* ------------------------------------------------------------------------- *) 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. *) (* ------------------------------------------------------------------------- *) val add : active -> Metis_Clause.clause -> active * Metis_Clause.clause list (* ------------------------------------------------------------------------- *) (* Pretty printing. *) (* ------------------------------------------------------------------------- *) val pp : active Metis_Print.pp end (**** Original file: src/Active.sml ****) (* ========================================================================= *) (* THE ACTIVE SET OF CLAUSES *) (* Copyright (c) 2002 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Active :> Metis_Active = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* Helper functions. *) (* ------------------------------------------------------------------------- *) (*MetisDebug local fun mkRewrite ordering = let fun add (cl,rw) = let val {id, thm = th, ...} = Metis_Clause.dest cl in case total Metis_Thm.destUnitEq th of SOME l_r => Metis_Rewrite.add rw (id,(l_r,th)) | NONE => rw end in List.foldl add (Metis_Rewrite.new (Metis_KnuthBendixOrder.compare ordering)) end; fun allFactors red = let fun allClause cl = List.all red (cl :: Metis_Clause.factor cl) orelse let val () = Metis_Print.trace Metis_Clause.pp "Metis_Active.isSaturated.allFactors: cl" cl in false end in List.all allClause end; fun allResolutions red = let fun allClause2 cl_lit cl = let fun allLiteral2 lit = case total (Metis_Clause.resolve cl_lit) (cl,lit) of NONE => true | SOME cl => allFactors red [cl] in Metis_LiteralSet.all allLiteral2 (Metis_Clause.literals cl) end orelse let val () = Metis_Print.trace Metis_Clause.pp "Metis_Active.isSaturated.allResolutions: cl2" cl in false end fun allClause1 allCls cl = let val cl = Metis_Clause.freshVars cl fun allLiteral1 lit = List.all (allClause2 (cl,lit)) allCls in Metis_LiteralSet.all allLiteral1 (Metis_Clause.literals cl) end orelse let val () = Metis_Print.trace Metis_Clause.pp "Metis_Active.isSaturated.allResolutions: cl1" cl in false end in fn [] => true | allCls as cl :: cls => allClause1 allCls cl andalso allResolutions red cls end; fun allParamodulations red cls = let fun allClause2 cl_lit_ort_tm cl = let fun allLiteral2 lit = let val para = Metis_Clause.paramodulate cl_lit_ort_tm fun allSubterms (path,tm) = case total para (cl,lit,path,tm) of NONE => true | SOME cl => allFactors red [cl] in List.all allSubterms (Metis_Literal.nonVarTypedSubterms lit) end orelse let val () = Metis_Print.trace Metis_Literal.pp "Metis_Active.isSaturated.allParamodulations: lit2" lit in false end in Metis_LiteralSet.all allLiteral2 (Metis_Clause.literals cl) end orelse let val () = Metis_Print.trace Metis_Clause.pp "Metis_Active.isSaturated.allParamodulations: cl2" cl val (_,_,ort,_) = cl_lit_ort_tm val () = Metis_Print.trace Metis_Rewrite.ppOrient "Metis_Active.isSaturated.allParamodulations: ort1" ort in false end fun allClause1 cl = let val cl = Metis_Clause.freshVars cl fun allLiteral1 lit = let fun allCl2 x = List.all (allClause2 x) cls in case total Metis_Literal.destEq lit of NONE => true | SOME (l,r) => allCl2 (cl,lit,Metis_Rewrite.LeftToRight,l) andalso allCl2 (cl,lit,Metis_Rewrite.RightToLeft,r) end orelse let val () = Metis_Print.trace Metis_Literal.pp "Metis_Active.isSaturated.allParamodulations: lit1" lit in false end in Metis_LiteralSet.all allLiteral1 (Metis_Clause.literals cl) end orelse let val () = Metis_Print.trace Metis_Clause.pp "Metis_Active.isSaturated.allParamodulations: cl1" cl in false end in List.all allClause1 cls end; fun redundant {subsume,reduce,rewrite} = let fun simp cl = case Metis_Clause.simplify cl of NONE => true (* tautology case *) | SOME cl => Metis_Subsume.isSubsumed subsume (Metis_Clause.literals cl) orelse let val cl' = cl val cl' = Metis_Clause.reduce reduce cl' val cl' = Metis_Clause.rewrite rewrite cl' in not (Metis_Clause.equalThms cl cl') andalso (simp cl' orelse let val () = Metis_Print.trace Metis_Clause.pp "Metis_Active.isSaturated.redundant: cl'" cl' in false end) end in fn cl => simp cl orelse let val () = Metis_Print.trace Metis_Clause.pp "Metis_Active.isSaturated.redundant: cl" cl in false end end; in fun isSaturated ordering subs cls = let val rd = Metis_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 () = Metis_Print.trace Metis_Rewrite.pp "Metis_Active.isSaturated: rw" rw val () = Metis_Print.trace (Metis_Print.ppList Metis_Clause.pp) "Metis_Active.isSaturated: clauses" cls in false end end; end; fun checkSaturated ordering subs cls = if isSaturated ordering subs cls then () else raise Bug "Metis_Active.checkSaturated"; *) (* ------------------------------------------------------------------------- *) (* A type of active clause sets. *) (* ------------------------------------------------------------------------- *) type simplify = {subsume : bool, reduce : bool, rewrite : bool}; type parameters = {clause : Metis_Clause.parameters, prefactor : simplify, postfactor : simplify}; datatype active = Metis_Active of {parameters : parameters, clauses : Metis_Clause.clause Metis_IntMap.map, units : Metis_Units.units, rewrite : Metis_Rewrite.rewrite, subsume : Metis_Clause.clause Metis_Subsume.subsume, literals : (Metis_Clause.clause * Metis_Literal.literal) Metis_LiteralNet.literalNet, equations : (Metis_Clause.clause * Metis_Literal.literal * Metis_Rewrite.orient * Metis_Term.term) Metis_TermNet.termNet, subterms : (Metis_Clause.clause * Metis_Literal.literal * Metis_Term.path * Metis_Term.term) Metis_TermNet.termNet, allSubterms : (Metis_Clause.clause * Metis_Term.term) Metis_TermNet.termNet}; fun getSubsume (Metis_Active {subsume = s, ...}) = s; fun setRewrite active rewrite = let val Metis_Active {parameters,clauses,units,subsume,literals,equations, subterms,allSubterms,...} = active in Metis_Active {parameters = parameters, clauses = clauses, units = units, rewrite = rewrite, subsume = subsume, literals = literals, equations = equations, subterms = subterms, allSubterms = allSubterms} end; (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val maxSimplify : simplify = {subsume = true, reduce = true, rewrite = true}; val default : parameters = {clause = Metis_Clause.default, prefactor = maxSimplify, postfactor = maxSimplify}; fun empty parameters = let val {clause,...} = parameters val {ordering,...} = clause in Metis_Active {parameters = parameters, clauses = Metis_IntMap.new (), units = Metis_Units.empty, rewrite = Metis_Rewrite.new (Metis_KnuthBendixOrder.compare ordering), subsume = Metis_Subsume.new (), literals = Metis_LiteralNet.new {fifo = false}, equations = Metis_TermNet.new {fifo = false}, subterms = Metis_TermNet.new {fifo = false}, allSubterms = Metis_TermNet.new {fifo = false}} end; fun size (Metis_Active {clauses,...}) = Metis_IntMap.size clauses; fun clauses (Metis_Active {clauses = cls, ...}) = let fun add (_,cl,acc) = cl :: acc in Metis_IntMap.foldr add [] cls end; fun saturation active = let fun remove (cl,(cls,subs)) = let val lits = Metis_Clause.literals cl in if Metis_Subsume.isStrictlySubsumed subs lits then (cls,subs) else (cl :: cls, Metis_Subsume.insert subs (lits,())) end val cls = clauses active val (cls,_) = List.foldl remove ([], Metis_Subsume.new ()) cls val (cls,subs) = List.foldl remove ([], Metis_Subsume.new ()) cls (*MetisDebug val Metis_Active {parameters,...} = active val {clause,...} = parameters val {ordering,...} = clause val () = checkSaturated ordering subs cls *) in cls end; (* ------------------------------------------------------------------------- *) (* Pretty printing. *) (* ------------------------------------------------------------------------- *) val pp = let fun toStr active = "Metis_Active{" ^ Int.toString (size active) ^ "}" in Metis_Print.ppMap toStr Metis_Print.ppString end; (*MetisDebug local fun ppField f ppA a = Metis_Print.inconsistentBlock 2 [Metis_Print.ppString (f ^ " ="), Metis_Print.break, ppA a]; val ppClauses = ppField "clauses" (Metis_Print.ppMap Metis_IntMap.toList (Metis_Print.ppList (Metis_Print.ppPair Metis_Print.ppInt Metis_Clause.pp))); val ppRewrite = ppField "rewrite" Metis_Rewrite.pp; val ppSubterms = ppField "subterms" (Metis_TermNet.pp (Metis_Print.ppMap (fn (c,l,p,t) => ((Metis_Clause.id c, l, p), t)) (Metis_Print.ppPair (Metis_Print.ppTriple Metis_Print.ppInt Metis_Literal.pp Metis_Term.ppPath) Metis_Term.pp))); in fun pp (Metis_Active {clauses,rewrite,subterms,...}) = Metis_Print.inconsistentBlock 2 [Metis_Print.ppString "Metis_Active", Metis_Print.break, Metis_Print.inconsistentBlock 1 [Metis_Print.ppString "{", ppClauses clauses, Metis_Print.ppString ",", Metis_Print.break, ppRewrite rewrite, (*MetisTrace5 Metis_Print.ppString ",", Metis_Print.break, ppSubterms subterms, *) Metis_Print.skip], Metis_Print.ppString "}"]; end; *) val toString = Metis_Print.toString pp; (* ------------------------------------------------------------------------- *) (* Simplify clauses. *) (* ------------------------------------------------------------------------- *) fun simplify simp units rewr subs = let val {subsume = s, reduce = r, rewrite = w} = simp fun rewrite cl = let val cl' = Metis_Clause.rewrite rewr cl in if Metis_Clause.equalThms cl cl' then SOME cl else case Metis_Clause.simplify cl' of NONE => NONE | SOME cl'' => (* *) (* Post-rewrite simplification can enable more rewrites: *) (* *) (* ~(X = f(X)) \/ ~(g(Y) = f(X)) \/ ~(c = f(X)) *) (* ---------------------------------------------- rewrite *) (* ~(X = f(X)) \/ ~(g(Y) = X) \/ ~(c = X) *) (* ---------------------------------------------- simplify *) (* ~(g(Y) = f(g(Y))) \/ ~(c = g(Y)) *) (* ---------------------------------------------- rewrite *) (* ~(c = f(c)) \/ ~(c = g(Y)) *) (* *) (* This was first observed in a bug discovered by Martin *) (* Desharnais and Jasmin Blanchett *) (* *) if Metis_Clause.equalThms cl' cl'' then SOME cl' else rewrite cl'' end in fn cl => case Metis_Clause.simplify cl of NONE => NONE | SOME cl => case (if w then rewrite cl else SOME cl) of NONE => NONE | SOME cl => let val cl = if r then Metis_Clause.reduce units cl else cl in if s andalso Metis_Clause.subsumes subs cl then NONE else SOME cl end end; (*MetisDebug val simplify = fn simp => fn units => fn rewr => fn subs => fn cl => let val ppClOpt = Metis_Print.ppOption Metis_Clause.pp fun trace pp s = Metis_Print.trace pp ("Metis_Active.simplify: " ^ s) val traceCl = trace Metis_Clause.pp val traceClOpt = trace ppClOpt (*MetisTrace4 val () = traceCl "cl" cl *) val cl' = simplify simp units rewr subs cl (*MetisTrace4 val () = Metis_Print.trace ppClOpt "Metis_Active.simplify: cl'" cl' *) val () = case cl' of NONE => () | SOME cl' => case (case simplify simp units rewr subs cl' of NONE => SOME ("away", K ()) | SOME cl'' => if Metis_Clause.equalThms cl' cl'' then NONE else SOME ("further", fn () => traceCl "cl''" cl'')) of NONE => () | SOME (e,f) => let val () = Metis_Print.trace Metis_Rewrite.pp "Metis_Active.simplify: rewr" rewr val () = Metis_Print.trace Metis_Units.pp "Metis_Active.simplify: units" units val () = Metis_Print.trace Metis_Subsume.pp "Metis_Active.simplify: subs" subs val () = traceCl "cl" cl val () = traceCl "cl'" cl' val () = traceClOpt "simplify cl'" (Metis_Clause.simplify cl') val () = traceCl "rewrite cl'" (Metis_Clause.rewrite rewr cl') val () = traceCl "reduce cl'" (Metis_Clause.reduce units cl') val () = f () in raise Bug ("Metis_Active.simplify: clause should have been simplified "^e) end in cl' end; *) fun simplifyActive simp active = let val Metis_Active {units,rewrite,subsume,...} = active in simplify simp units rewrite subsume end; (* ------------------------------------------------------------------------- *) (* Add a clause into the active set. *) (* ------------------------------------------------------------------------- *) fun addUnit units cl = let val th = Metis_Clause.thm cl in case total Metis_Thm.destUnit th of SOME lit => Metis_Units.add units (lit,th) | NONE => units end; fun addRewrite rewrite cl = let val th = Metis_Clause.thm cl in case total Metis_Thm.destUnitEq th of SOME l_r => Metis_Rewrite.add rewrite (Metis_Clause.id cl, (l_r,th)) | NONE => rewrite end; fun addSubsume subsume cl = Metis_Subsume.insert subsume (Metis_Clause.literals cl, cl); fun addLiterals literals cl = let fun add (lit as (_,atm), literals) = if Metis_Atom.isEq atm then literals else Metis_LiteralNet.insert literals (lit,(cl,lit)) in Metis_LiteralSet.foldl add literals (Metis_Clause.largestLiterals cl) end; fun addEquations equations cl = let fun add ((lit,ort,tm),equations) = Metis_TermNet.insert equations (tm,(cl,lit,ort,tm)) in List.foldl add equations (Metis_Clause.largestEquations cl) end; fun addSubterms subterms cl = let fun add ((lit,path,tm),subterms) = Metis_TermNet.insert subterms (tm,(cl,lit,path,tm)) in List.foldl add subterms (Metis_Clause.largestSubterms cl) end; fun addAllSubterms allSubterms cl = let fun add ((_,_,tm),allSubterms) = Metis_TermNet.insert allSubterms (tm,(cl,tm)) in List.foldl add allSubterms (Metis_Clause.allSubterms cl) end; fun addClause active cl = let val Metis_Active {parameters,clauses,units,rewrite,subsume,literals, equations,subterms,allSubterms} = active val clauses = Metis_IntMap.insert clauses (Metis_Clause.id cl, cl) and subsume = addSubsume subsume cl and literals = addLiterals literals cl and equations = addEquations equations cl and subterms = addSubterms subterms cl and allSubterms = addAllSubterms allSubterms cl in Metis_Active {parameters = parameters, clauses = clauses, units = units, rewrite = rewrite, subsume = subsume, literals = literals, equations = equations, subterms = subterms, allSubterms = allSubterms} end; fun addFactorClause active cl = let val Metis_Active {parameters,clauses,units,rewrite,subsume,literals, equations,subterms,allSubterms} = active val units = addUnit units cl and rewrite = addRewrite rewrite cl in Metis_Active {parameters = parameters, clauses = clauses, units = units, rewrite = rewrite, subsume = subsume, literals = literals, equations = equations, subterms = subterms, allSubterms = allSubterms} end; (* ------------------------------------------------------------------------- *) (* Derive (unfactored) consequences of a clause. *) (* ------------------------------------------------------------------------- *) fun deduceResolution literals cl (lit as (_,atm), acc) = let fun resolve (cl_lit,acc) = case total (Metis_Clause.resolve cl_lit) (cl,lit) of SOME cl' => cl' :: acc | NONE => acc (*MetisTrace4 val () = Metis_Print.trace Metis_Literal.pp "Metis_Active.deduceResolution: lit" lit *) in if Metis_Atom.isEq atm then acc else List.foldl resolve acc (Metis_LiteralNet.unify literals (Metis_Literal.negate lit)) end; fun deduceParamodulationWith subterms cl ((lit,ort,tm),acc) = let fun para (cl_lit_path_tm,acc) = case total (Metis_Clause.paramodulate (cl,lit,ort,tm)) cl_lit_path_tm of SOME cl' => cl' :: acc | NONE => acc in List.foldl para acc (Metis_TermNet.unify subterms tm) end; fun deduceParamodulationInto equations cl ((lit,path,tm),acc) = let fun para (cl_lit_ort_tm,acc) = case total (Metis_Clause.paramodulate cl_lit_ort_tm) (cl,lit,path,tm) of SOME cl' => cl' :: acc | NONE => acc in List.foldl para acc (Metis_TermNet.unify equations tm) end; fun deduce active cl = let val Metis_Active {parameters,literals,equations,subterms,...} = active val lits = Metis_Clause.largestLiterals cl val eqns = Metis_Clause.largestEquations cl val subtms = if Metis_TermNet.null equations then [] else Metis_Clause.largestSubterms cl (*MetisTrace5 val () = Metis_Print.trace Metis_LiteralSet.pp "Metis_Active.deduce: lits" lits val () = Metis_Print.trace (Metis_Print.ppList (Metis_Print.ppMap (fn (lit,ort,_) => (lit,ort)) (Metis_Print.ppPair Metis_Literal.pp Metis_Rewrite.ppOrient))) "Metis_Active.deduce: eqns" eqns val () = Metis_Print.trace (Metis_Print.ppList (Metis_Print.ppTriple Metis_Literal.pp Metis_Term.ppPath Metis_Term.pp)) "Metis_Active.deduce: subtms" subtms *) val acc = [] val acc = Metis_LiteralSet.foldl (deduceResolution literals cl) acc lits val acc = List.foldl (deduceParamodulationWith subterms cl) acc eqns val acc = List.foldl (deduceParamodulationInto equations cl) acc subtms val acc = List.rev acc (*MetisTrace5 val () = Metis_Print.trace (Metis_Print.ppList Metis_Clause.pp) "Metis_Active.deduce: acc" acc *) in acc end; (* ------------------------------------------------------------------------- *) (* Extract clauses from the active set that can be simplified. *) (* ------------------------------------------------------------------------- *) local fun clause_rewritables active = let val Metis_Active {clauses,rewrite,...} = active fun rewr (id,cl,ids) = let val cl' = Metis_Clause.rewrite rewrite cl in if Metis_Clause.equalThms cl cl' then ids else Metis_IntSet.add ids id end in Metis_IntMap.foldr rewr Metis_IntSet.empty clauses end; fun orderedRedexResidues (((l,r),_),ort) = case ort of NONE => [] | SOME Metis_Rewrite.LeftToRight => [(l,r,true)] | SOME Metis_Rewrite.RightToLeft => [(r,l,true)]; fun unorderedRedexResidues (((l,r),_),ort) = case ort of NONE => [(l,r,false),(r,l,false)] | SOME _ => []; fun rewrite_rewritables active rewr_ids = let val Metis_Active {parameters,rewrite,clauses,allSubterms,...} = active val {clause = {ordering,...}, ...} = parameters val order = Metis_KnuthBendixOrder.compare ordering fun addRewr (id,acc) = if Metis_IntMap.inDomain id clauses then Metis_IntSet.add acc id else acc fun addReduce ((l,r,ord),acc) = let fun isValidRewr tm = case total (Metis_Subst.match Metis_Subst.empty l) tm of NONE => false | SOME sub => ord orelse let val tm' = Metis_Subst.subst (Metis_Subst.normalize sub) r in order (tm,tm') = SOME GREATER end fun addRed ((cl,tm),acc) = let (*MetisTrace5 val () = Metis_Print.trace Metis_Clause.pp "Metis_Active.addRed: cl" cl val () = Metis_Print.trace Metis_Term.pp "Metis_Active.addRed: tm" tm *) val id = Metis_Clause.id cl in if Metis_IntSet.member id acc then acc else if not (isValidRewr tm) then acc else Metis_IntSet.add acc id end (*MetisTrace5 val () = Metis_Print.trace Metis_Term.pp "Metis_Active.addReduce: l" l val () = Metis_Print.trace Metis_Term.pp "Metis_Active.addReduce: r" r val () = Metis_Print.trace Metis_Print.ppBool "Metis_Active.addReduce: ord" ord *) in List.foldl addRed acc (Metis_TermNet.matched allSubterms l) end fun addEquation redexResidues (id,acc) = case Metis_Rewrite.peek rewrite id of NONE => acc | SOME eqn_ort => List.foldl addReduce acc (redexResidues eqn_ort) val addOrdered = addEquation orderedRedexResidues val addUnordered = addEquation unorderedRedexResidues val ids = Metis_IntSet.empty val ids = List.foldl addRewr ids rewr_ids val ids = List.foldl addOrdered ids rewr_ids val ids = List.foldl addUnordered ids rewr_ids in ids end; fun choose_clause_rewritables active ids = size active <= length ids fun rewritables active ids = if choose_clause_rewritables active ids then clause_rewritables active else rewrite_rewritables active ids; (*MetisDebug val rewritables = fn active => fn ids => let val clause_ids = clause_rewritables active val rewrite_ids = rewrite_rewritables active ids val () = if Metis_IntSet.equal rewrite_ids clause_ids then () else let val ppIdl = Metis_Print.ppList Metis_Print.ppInt val ppIds = Metis_Print.ppMap Metis_IntSet.toList ppIdl val () = Metis_Print.trace pp "Metis_Active.rewritables: active" active val () = Metis_Print.trace ppIdl "Metis_Active.rewritables: ids" ids val () = Metis_Print.trace ppIds "Metis_Active.rewritables: clause_ids" clause_ids val () = Metis_Print.trace ppIds "Metis_Active.rewritables: rewrite_ids" rewrite_ids in raise Bug "Metis_Active.rewritables: ~(rewrite_ids SUBSET clause_ids)" end in if choose_clause_rewritables active ids then clause_ids else rewrite_ids end; *) fun delete active ids = if Metis_IntSet.null ids then active else let fun idPred id = not (Metis_IntSet.member id ids) fun clausePred cl = idPred (Metis_Clause.id cl) val Metis_Active {parameters, clauses, units, rewrite, subsume, literals, equations, subterms, allSubterms} = active val clauses = Metis_IntMap.filter (idPred o fst) clauses and subsume = Metis_Subsume.filter clausePred subsume and literals = Metis_LiteralNet.filter (clausePred o #1) literals and equations = Metis_TermNet.filter (clausePred o #1) equations and subterms = Metis_TermNet.filter (clausePred o #1) subterms and allSubterms = Metis_TermNet.filter (clausePred o fst) allSubterms in Metis_Active {parameters = parameters, clauses = clauses, units = units, rewrite = rewrite, subsume = subsume, literals = literals, equations = equations, subterms = subterms, allSubterms = allSubterms} end; in fun extract_rewritables (active as Metis_Active {clauses,rewrite,...}) = if Metis_Rewrite.isReduced rewrite then (active,[]) else let (*MetisTrace3 val () = trace "Metis_Active.extract_rewritables: inter-reducing\n" *) val (rewrite,ids) = Metis_Rewrite.reduce' rewrite val active = setRewrite active rewrite val ids = rewritables active ids val cls = Metis_IntSet.transform (Metis_IntMap.get clauses) ids (*MetisTrace3 val ppCls = Metis_Print.ppList Metis_Clause.pp val () = Metis_Print.trace ppCls "Metis_Active.extract_rewritables: cls" cls *) in (delete active ids, cls) end (*MetisDebug handle Error err => raise Bug ("Metis_Active.extract_rewritables: shouldn't fail\n" ^ err); *) end; (* ------------------------------------------------------------------------- *) (* Factor clauses. *) (* ------------------------------------------------------------------------- *) local fun prefactor_simplify active subsume = let val Metis_Active {parameters,units,rewrite,...} = active val {prefactor,...} = parameters in simplify prefactor units rewrite subsume end; fun postfactor_simplify active subsume = let val Metis_Active {parameters,units,rewrite,...} = active val {postfactor,...} = parameters in simplify postfactor units rewrite subsume end; val sort_utilitywise = let fun utility cl = case Metis_LiteralSet.size (Metis_Clause.literals cl) of 0 => ~1 | 1 => if Metis_Thm.isUnitEq (Metis_Clause.thm cl) then 0 else 1 | n => n in sortMap utility Int.compare end; fun post_factor (cl, active_subsume_acc as (active,subsume,acc)) = case postfactor_simplify active subsume cl of NONE => active_subsume_acc | SOME cl' => if Metis_Clause.equalThms cl' cl then let val active = addFactorClause active cl and subsume = addSubsume subsume cl and acc = cl :: acc in (active,subsume,acc) end else (* If the clause was changed in the post-factor simplification *) (* step, then it may have altered the set of largest literals *) (* in the clause. The safest thing to do is to factor again. *) factor1 cl' active_subsume_acc and factor1 cl active_subsume_acc = let val cls = sort_utilitywise (cl :: Metis_Clause.factor cl) in List.foldl post_factor active_subsume_acc cls end; fun pre_factor (cl, active_subsume_acc as (active,subsume,_)) = case prefactor_simplify active subsume cl of NONE => active_subsume_acc | SOME cl => factor1 cl active_subsume_acc; fun factor' active acc [] = (active, List.rev acc) | factor' active acc cls = let val cls = sort_utilitywise cls val subsume = getSubsume active val (active,_,acc) = List.foldl pre_factor (active,subsume,acc) cls val (active,cls) = extract_rewritables active in factor' active acc cls end; in fun factor active cls = factor' active [] cls; end; (*MetisTrace4 val factor = fn active => fn cls => let val ppCls = Metis_Print.ppList Metis_Clause.pp val () = Metis_Print.trace ppCls "Metis_Active.factor: cls" cls val (active,cls') = factor active cls val () = Metis_Print.trace ppCls "Metis_Active.factor: cls'" cls' in (active,cls') end; *) (* ------------------------------------------------------------------------- *) (* Create a new active clause set and initialize clauses. *) (* ------------------------------------------------------------------------- *) fun new parameters {axioms,conjecture} = let val {clause,...} = parameters fun mk_clause th = Metis_Clause.mk {parameters = clause, id = Metis_Clause.newId (), thm = th} val active = empty parameters val (active,axioms) = factor active (List.map mk_clause axioms) val (active,conjecture) = factor active (List.map mk_clause conjecture) in (active, {axioms = axioms, conjecture = conjecture}) end; (* ------------------------------------------------------------------------- *) (* Add a clause into the active set and deduce all consequences. *) (* ------------------------------------------------------------------------- *) fun add active cl = case simplifyActive maxSimplify active cl of NONE => (active,[]) | SOME cl' => if Metis_Clause.isContradiction cl' then (active,[cl']) else if not (Metis_Clause.equalThms cl cl') then factor active [cl'] else let (*MetisTrace2 val () = Metis_Print.trace Metis_Clause.pp "Metis_Active.add: cl" cl *) val active = addClause active cl val cl = Metis_Clause.freshVars cl val cls = deduce active cl val (active,cls) = factor active cls (*MetisTrace2 val ppCls = Metis_Print.ppList Metis_Clause.pp val () = Metis_Print.trace ppCls "Metis_Active.add: cls" cls *) in (active,cls) end; end (**** Original file: src/Waiting.sig ****) (* ========================================================================= *) (* THE WAITING SET OF CLAUSES *) (* Copyright (c) 2002 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Waiting = sig (* ------------------------------------------------------------------------- *) (* 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 : weight, variablesWeight : weight, literalsWeight : weight, models : modelParameters list} (* ------------------------------------------------------------------------- *) (* A type of waiting sets of clauses. *) (* ------------------------------------------------------------------------- *) type waiting type distance (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val default : parameters val new : parameters -> {axioms : Metis_Clause.clause list, conjecture : Metis_Clause.clause list} -> waiting val size : waiting -> int val pp : waiting Metis_Print.pp (* ------------------------------------------------------------------------- *) (* Adding new clauses. *) (* ------------------------------------------------------------------------- *) val add : waiting -> distance * Metis_Clause.clause list -> waiting (* ------------------------------------------------------------------------- *) (* Removing the lightest clause. *) (* ------------------------------------------------------------------------- *) val remove : waiting -> ((distance * Metis_Clause.clause) * waiting) option end (**** Original file: src/Waiting.sml ****) (* ========================================================================= *) (* THE WAITING SET OF CLAUSES *) (* Copyright (c) 2002 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Waiting :> Metis_Waiting = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* A type of waiting sets of clauses. *) (* ------------------------------------------------------------------------- *) type weight = real; type modelParameters = {model : Metis_Model.parameters, initialPerturbations : int, maxChecks : int option, perturbations : int, weight : weight} type parameters = {symbolsWeight : weight, variablesWeight : weight, literalsWeight : weight, models : modelParameters list}; type distance = real; datatype waiting = Metis_Waiting of {parameters : parameters, clauses : (weight * (distance * Metis_Clause.clause)) Metis_Heap.heap, models : Metis_Model.model list}; (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val defaultModels : modelParameters list = [{model = Metis_Model.default, initialPerturbations = 100, maxChecks = SOME 20, perturbations = 0, weight = 1.0}]; val default : parameters = {symbolsWeight = 1.0, literalsWeight = 1.0, variablesWeight = 1.0, models = defaultModels}; fun size (Metis_Waiting {clauses,...}) = Metis_Heap.size clauses; val pp = Metis_Print.ppMap (fn w => "Metis_Waiting{" ^ Int.toString (size w) ^ "}") Metis_Print.ppString; (*MetisDebug val pp = Metis_Print.ppMap (fn Metis_Waiting {clauses,...} => List.map (fn (w,(_,cl)) => (w, Metis_Clause.id cl, cl)) (Metis_Heap.toList clauses)) (Metis_Print.ppList (Metis_Print.ppTriple Metis_Print.ppReal Metis_Print.ppInt Metis_Clause.pp)); *) (* ------------------------------------------------------------------------- *) (* Perturbing the models. *) (* ------------------------------------------------------------------------- *) type modelClause = Metis_NameSet.set * Metis_Thm.clause; fun mkModelClause cl = let val lits = Metis_Clause.literals cl val fvs = Metis_LiteralSet.freeVars lits in (fvs,lits) end; val mkModelClauses = List.map mkModelClause; fun perturbModel M cls = if List.null cls then K () else let val N = {size = Metis_Model.size M} fun perturbClause (fv,cl) = let val V = Metis_Model.randomValuation N fv in if Metis_Model.interpretClause M V cl then () else Metis_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 = Metis_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} = Metis_Model.check Metis_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; (* ------------------------------------------------------------------------- *) (* Metis_Clause weights. *) (* ------------------------------------------------------------------------- *) local fun clauseSymbols cl = Real.fromInt (Metis_LiteralSet.typedSymbols cl); fun clauseVariables cl = Real.fromInt (Metis_NameSet.size (Metis_LiteralSet.freeVars cl) + 1); fun clauseLiterals cl = Real.fromInt (Metis_LiteralSet.size cl); fun clausePriority cl = 1e~12 * Real.fromInt (Metis_Clause.id cl); in fun clauseWeight (parm : parameters) mods dist mcl cl = let (*MetisTrace3 val () = Metis_Print.trace Metis_Clause.pp "Metis_Waiting.clauseWeight: cl" cl *) val {symbolsWeight,variablesWeight,literalsWeight,models,...} = parm val lits = Metis_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 = checkModels models mods mcl (*MetisTrace4 val () = trace ("Metis_Waiting.clauseWeight: dist = " ^ Real.toString dist ^ "\n") val () = trace ("Metis_Waiting.clauseWeight: symbolsW = " ^ Real.toString symbolsW ^ "\n") val () = trace ("Metis_Waiting.clauseWeight: variablesW = " ^ Real.toString variablesW ^ "\n") val () = trace ("Metis_Waiting.clauseWeight: literalsW = " ^ Real.toString literalsW ^ "\n") val () = trace ("Metis_Waiting.clauseWeight: modelsW = " ^ Real.toString modelsW ^ "\n") *) val weight = dist * symbolsW * variablesW * literalsW * modelsW val weight = weight + clausePriority cl (*MetisTrace3 val () = trace ("Metis_Waiting.clauseWeight: weight = " ^ Real.toString weight ^ "\n") *) in weight end; end; (* ------------------------------------------------------------------------- *) (* Adding new clauses. *) (* ------------------------------------------------------------------------- *) fun add' waiting dist mcls cls = let val Metis_Waiting {parameters,clauses,models} = waiting val {models = modelParameters, ...} = parameters (*MetisDebug val _ = not (List.null cls) orelse raise Bug "Metis_Waiting.add': null" val _ = length mcls = length cls orelse raise Bug "Metis_Waiting.add': different lengths" *) val dist = dist + Math.ln (Real.fromInt (length cls)) fun addCl ((mcl,cl),acc) = let val weight = clauseWeight parameters models dist mcl cl in Metis_Heap.add acc (weight,(dist,cl)) end val clauses = List.foldl addCl clauses (zip mcls cls) val () = perturbModels modelParameters models mcls in Metis_Waiting {parameters = parameters, clauses = clauses, models = models} end; fun add waiting (dist,cls) = if List.null cls then waiting else let (*MetisTrace3 val () = Metis_Print.trace pp "Metis_Waiting.add: waiting" waiting val () = Metis_Print.trace (Metis_Print.ppList Metis_Clause.pp) "Metis_Waiting.add: cls" cls *) val waiting = add' waiting dist (mkModelClauses cls) cls (*MetisTrace3 val () = Metis_Print.trace pp "Metis_Waiting.add: waiting" waiting *) in waiting end; local fun cmp ((w1,_),(w2,_)) = Real.compare (w1,w2); fun empty parameters axioms conjecture = let val {models = modelParameters, ...} = parameters val clauses = Metis_Heap.new cmp and models = List.map (initialModel axioms conjecture) modelParameters in Metis_Waiting {parameters = parameters, clauses = clauses, models = models} end; in fun new parameters {axioms,conjecture} = let val mAxioms = mkModelClauses axioms and mConjecture = mkModelClauses conjecture val waiting = empty parameters mAxioms mConjecture in if List.null axioms andalso List.null conjecture then waiting else add' waiting 0.0 (mAxioms @ mConjecture) (axioms @ conjecture) end (*MetisDebug handle e => let val () = Metis_Print.trace Metis_Print.ppException "Metis_Waiting.new: exception" e in raise e end; *) end; (* ------------------------------------------------------------------------- *) (* Removing the lightest clause. *) (* ------------------------------------------------------------------------- *) fun remove (Metis_Waiting {parameters,clauses,models}) = if Metis_Heap.null clauses then NONE else let val ((_,dcl),clauses) = Metis_Heap.remove clauses val waiting = Metis_Waiting {parameters = parameters, clauses = clauses, models = models} in SOME (dcl,waiting) end; end (**** Original file: src/Resolution.sig ****) (* ========================================================================= *) (* THE RESOLUTION PROOF PROCEDURE *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) signature Metis_Resolution = sig (* ------------------------------------------------------------------------- *) (* A type of resolution proof procedures. *) (* ------------------------------------------------------------------------- *) type parameters = {active : Metis_Active.parameters, waiting : Metis_Waiting.parameters} type resolution (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val default : parameters 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_Print.pp (* ------------------------------------------------------------------------- *) (* The main proof loop. *) (* ------------------------------------------------------------------------- *) datatype decision = Contradiction of Metis_Thm.thm | Satisfiable of Metis_Thm.thm list datatype state = Decided of decision | Undecided of resolution val iterate : resolution -> state val loop : resolution -> decision end (**** Original file: src/Resolution.sml ****) (* ========================================================================= *) (* THE RESOLUTION PROOF PROCEDURE *) (* Copyright (c) 2001 Joe Leslie-Hurd, distributed under the BSD License *) (* ========================================================================= *) structure Metis_Resolution :> Metis_Resolution = struct open Metis_Useful; (* ------------------------------------------------------------------------- *) (* A type of resolution proof procedures. *) (* ------------------------------------------------------------------------- *) type parameters = {active : Metis_Active.parameters, waiting : Metis_Waiting.parameters}; datatype resolution = Metis_Resolution of {parameters : parameters, active : Metis_Active.active, waiting : Metis_Waiting.waiting}; (* ------------------------------------------------------------------------- *) (* Basic operations. *) (* ------------------------------------------------------------------------- *) val default : parameters = {active = Metis_Active.default, waiting = Metis_Waiting.default}; fun new parameters ths = let val {active = activeParm, waiting = waitingParm} = parameters val (active,cls) = Metis_Active.new activeParm ths (* cls = factored ths *) val waiting = Metis_Waiting.new waitingParm cls in Metis_Resolution {parameters = parameters, active = active, waiting = waiting} end (*MetisDebug handle e => let val () = Metis_Print.trace Metis_Print.ppException "Metis_Resolution.new: exception" e in raise e end; *) fun active (Metis_Resolution {active = a, ...}) = a; fun waiting (Metis_Resolution {waiting = w, ...}) = w; val pp = Metis_Print.ppMap (fn Metis_Resolution {active,waiting,...} => "Metis_Resolution(" ^ Int.toString (Metis_Active.size active) ^ "<-" ^ Int.toString (Metis_Waiting.size waiting) ^ ")") Metis_Print.ppString; (* ------------------------------------------------------------------------- *) (* The main proof loop. *) (* ------------------------------------------------------------------------- *) datatype decision = Contradiction of Metis_Thm.thm | Satisfiable of Metis_Thm.thm list; datatype state = Decided of decision | Undecided of resolution; fun iterate res = let val Metis_Resolution {parameters,active,waiting} = res (*MetisTrace2 val () = Metis_Print.trace Metis_Active.pp "Metis_Resolution.iterate: active" active val () = Metis_Print.trace Metis_Waiting.pp "Metis_Resolution.iterate: waiting" waiting *) in case Metis_Waiting.remove waiting of NONE => let val sat = Satisfiable (List.map Metis_Clause.thm (Metis_Active.saturation active)) in Decided sat end | SOME ((d,cl),waiting) => if Metis_Clause.isContradiction cl then Decided (Contradiction (Metis_Clause.thm cl)) else let (*MetisTrace1 val () = Metis_Print.trace Metis_Clause.pp "Metis_Resolution.iterate: cl" cl *) val (active,cls) = Metis_Active.add active cl val waiting = Metis_Waiting.add waiting (d,cls) val res = Metis_Resolution {parameters = parameters, active = active, waiting = waiting} in Undecided res end end; fun loop res = case iterate res of Decided dec => dec | Undecided res => loop res; end ;