aboutsummaryrefslogtreecommitdiff
path: root/lib
diff options
context:
space:
mode:
authorfilliatr2001-03-05 14:11:40 +0000
committerfilliatr2001-03-05 14:11:40 +0000
commitc034197e869cdc418b225b9abf25dee63a47e237 (patch)
tree42b5c9ecc4840c00ac1d31e5caf38b432953f7da /lib
parent9c88d6021a178cb64360bca75adbb6db030c480f (diff)
module Explore générique et réécriture EAuto avec ce module; occur check dans clenv_merge
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@1425 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'lib')
-rw-r--r--lib/explore.ml87
-rw-r--r--lib/explore.mli43
2 files changed, 130 insertions, 0 deletions
diff --git a/lib/explore.ml b/lib/explore.ml
new file mode 100644
index 0000000000..c336a4b3b8
--- /dev/null
+++ b/lib/explore.ml
@@ -0,0 +1,87 @@
+
+(*i $Id$ i*)
+
+open Format
+
+(*s Definition of a search problem. *)
+
+module type SearchProblem = sig
+ type state
+ val branching : state -> state list
+ val success : state -> bool
+ val pp : state -> unit
+end
+
+module Make = functor(S : SearchProblem) -> struct
+
+ type position = int list
+
+ let rec pp_position = function
+ | [] -> ()
+ | [i] -> printf "%d" i
+ | i :: l -> pp_position l; printf ".%d" i
+
+ (*s Depth first search. *)
+
+ let rec depth_first s =
+ if S.success s then s else depth_first_many (S.branching s)
+ and depth_first_many = function
+ | [] -> raise Not_found
+ | s :: l -> try depth_first s with Not_found -> depth_first_many l
+
+ let debug_depth_first s =
+ let rec explore p s =
+ pp_position p; S.pp s;
+ if S.success s then s else explore_many 1 p (S.branching s)
+ and explore_many i p = function
+ | [] ->
+ raise Not_found
+ | s :: l ->
+ try explore (i::p) s with Not_found -> explore_many (succ i) p l
+ in
+ explore [1] s
+
+ (*s Breadth first search. We use functional FIFOS à la Okasaki. *)
+
+ type 'a queue = 'a list * 'a list
+
+ exception Empty
+
+ let empty = [],[]
+
+ let push x (h,t) = (x::h,t)
+
+ let pop = function
+ | h, x::t -> x, (h,t)
+ | h, [] -> match List.rev h with x::t -> x, ([],t) | [] -> raise Empty
+
+ let breadth_first s =
+ let rec explore q =
+ try
+ let (s, q') = pop q in enqueue q' (S.branching s)
+ with Empty ->
+ raise Not_found
+ and enqueue q = function
+ | [] -> explore q
+ | s :: l -> if S.success s then s else enqueue (push s q) l
+ in
+ enqueue empty [s]
+
+ let debug_breadth_first s =
+ let rec explore q =
+ try
+ let ((p,s), q') = pop q in
+ enqueue 1 p q' (S.branching s)
+ with Empty ->
+ raise Not_found
+ and enqueue i p q = function
+ | [] ->
+ explore q
+ | s :: l ->
+ let ps = i::p in
+ pp_position ps; S.pp s;
+ if S.success s then s else enqueue (succ i) p (push (ps,s) q) l
+ in
+ enqueue 1 [] empty [s]
+
+end
diff --git a/lib/explore.mli b/lib/explore.mli
new file mode 100644
index 0000000000..5f484f8e6a
--- /dev/null
+++ b/lib/explore.mli
@@ -0,0 +1,43 @@
+
+(*i $Id$ i*)
+
+(*s Search strategies. *)
+
+(*s A search problem implements the following signature [SearchProblem].
+ [state] is the type of states of the search tree.
+ [branching] is the branching function; if [branching s] returns an
+ empty list, then search from [s] is aborted; successors of [s] are
+ recursively searched in the order they appear in the list.
+ [success] determines whether a given state is a success.
+
+ [pp] is a pretty-printer for states used in debugging versions of the
+ search functions. *)
+
+module type SearchProblem = sig
+
+ type state
+
+ val branching : state -> state list
+
+ val success : state -> bool
+
+ val pp : state -> unit
+
+end
+
+(*s Functor [Make] returns some search functions given a search problem.
+ Search functions raise [Not_found] if no success is found.
+ States are always visited in the order they appear in the
+ output of [branching] (whatever the search method is).
+ Debugging versions of the search functions print the position of the
+ visited state together with the state it-self (using [S.pp]). *)
+
+module Make : functor(S : SearchProblem) -> sig
+
+ val depth_first : S.state -> S.state
+ val debug_depth_first : S.state -> S.state
+
+ val breadth_first : S.state -> S.state
+ val debug_breadth_first : S.state -> S.state
+
+end