diff options
| author | filliatr | 1999-11-19 09:11:35 +0000 |
|---|---|---|
| committer | filliatr | 1999-11-19 09:11:35 +0000 |
| commit | 5865d4d79f052cfb190e728e8618cb05e2ac845f (patch) | |
| tree | 530892b92888f132d18c799d2a3ad221dcaa965c | |
| parent | f9f2c2bc695033f93a0b7352027678c4ca305ccd (diff) | |
discriminations nets
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@123 85f007b7-540e-0410-9357-904b9bb8a0f7
| -rw-r--r-- | .depend | 22 | ||||
| -rw-r--r-- | Makefile | 6 | ||||
| -rw-r--r-- | dev/changements.txt | 2 | ||||
| -rw-r--r-- | lib/gmap.ml | 9 | ||||
| -rw-r--r-- | lib/gmap.mli | 6 | ||||
| -rw-r--r-- | lib/tlm.ml | 112 | ||||
| -rw-r--r-- | lib/tlm.mli | 7 | ||||
| -rw-r--r-- | tactics/btermdn.ml | 52 | ||||
| -rw-r--r-- | tactics/btermdn.mli | 21 | ||||
| -rw-r--r-- | tactics/dn.ml | 2 | ||||
| -rw-r--r-- | tactics/nbtermdn.ml | 75 | ||||
| -rw-r--r-- | tactics/nbtermdn.mli | 30 | ||||
| -rw-r--r-- | tactics/termdn.ml | 68 | ||||
| -rw-r--r-- | tactics/termdn.mli | 27 |
14 files changed, 358 insertions, 81 deletions
@@ -75,7 +75,11 @@ proofs/tacmach.cmi: kernel/environ.cmi proofs/evar_refiner.cmi \ proofs/tacred.cmi: kernel/closure.cmi kernel/environ.cmi kernel/evd.cmi \ kernel/names.cmi kernel/reduction.cmi kernel/term.cmi proofs/typing_ev.cmi: kernel/environ.cmi kernel/evd.cmi kernel/term.cmi +tactics/btermdn.cmi: kernel/generic.cmi kernel/term.cmi tactics/dn.cmi: lib/tlm.cmi +tactics/nbtermdn.cmi: tactics/btermdn.cmi kernel/generic.cmi kernel/term.cmi \ + tactics/termdn.cmi +tactics/termdn.cmi: tactics/dn.cmi kernel/generic.cmi kernel/term.cmi toplevel/errors.cmi: parsing/coqast.cmi lib/pp.cmi toplevel/himsg.cmi: kernel/environ.cmi kernel/names.cmi lib/pp.cmi \ kernel/sign.cmi kernel/term.cmi kernel/type_errors.cmi @@ -203,8 +207,8 @@ lib/stamps.cmo: lib/stamps.cmi lib/stamps.cmx: lib/stamps.cmi lib/system.cmo: lib/pp.cmi lib/util.cmi lib/system.cmi lib/system.cmx: lib/pp.cmx lib/util.cmx lib/system.cmi -lib/tlm.cmo: lib/tlm.cmi -lib/tlm.cmx: lib/tlm.cmi +lib/tlm.cmo: lib/gmap.cmi lib/gset.cmi lib/tlm.cmi +lib/tlm.cmx: lib/gmap.cmx lib/gset.cmx lib/tlm.cmi lib/util.cmo: lib/pp.cmi lib/util.cmi lib/util.cmx: lib/pp.cmx lib/util.cmi library/declare.cmo: kernel/constant.cmi kernel/generic.cmi \ @@ -361,8 +365,22 @@ proofs/typing_ev.cmo: kernel/environ.cmi kernel/generic.cmi kernel/names.cmi \ proofs/typing_ev.cmx: kernel/environ.cmx kernel/generic.cmx kernel/names.cmx \ kernel/reduction.cmx kernel/term.cmx kernel/type_errors.cmx \ kernel/typeops.cmx lib/util.cmx proofs/typing_ev.cmi +tactics/btermdn.cmo: tactics/dn.cmi kernel/term.cmi tactics/termdn.cmi \ + tactics/btermdn.cmi +tactics/btermdn.cmx: tactics/dn.cmx kernel/term.cmx tactics/termdn.cmx \ + tactics/btermdn.cmi tactics/dn.cmo: lib/tlm.cmi tactics/dn.cmi tactics/dn.cmx: lib/tlm.cmx tactics/dn.cmi +tactics/nbtermdn.cmo: tactics/btermdn.cmi kernel/generic.cmi lib/gmap.cmi \ + library/libobject.cmi library/library.cmi kernel/names.cmi \ + kernel/term.cmi tactics/termdn.cmi lib/util.cmi tactics/nbtermdn.cmi +tactics/nbtermdn.cmx: tactics/btermdn.cmx kernel/generic.cmx lib/gmap.cmx \ + library/libobject.cmx library/library.cmx kernel/names.cmx \ + kernel/term.cmx tactics/termdn.cmx lib/util.cmx tactics/nbtermdn.cmi +tactics/termdn.cmo: tactics/dn.cmi kernel/generic.cmi kernel/term.cmi \ + lib/util.cmi tactics/termdn.cmi +tactics/termdn.cmx: tactics/dn.cmx kernel/generic.cmx kernel/term.cmx \ + lib/util.cmx tactics/termdn.cmi toplevel/errors.cmo: parsing/ast.cmi lib/options.cmi lib/pp.cmi lib/util.cmi \ toplevel/errors.cmi toplevel/errors.cmx: parsing/ast.cmx lib/options.cmx lib/pp.cmx lib/util.cmx \ @@ -33,7 +33,8 @@ CONFIG=config/coq_config.cmo LIB=lib/pp_control.cmo lib/pp.cmo lib/util.cmo \ lib/hashcons.cmo lib/dyn.cmo lib/system.cmo lib/options.cmo \ - lib/bstack.cmo lib/edit.cmo lib/stamps.cmo lib/gset.cmo lib/gmap.cmo + lib/bstack.cmo lib/edit.cmo lib/stamps.cmo lib/gset.cmo lib/gmap.cmo \ + lib/tlm.cmo KERNEL=kernel/names.cmo kernel/generic.cmo kernel/univ.cmo kernel/term.cmo \ kernel/sign.cmo kernel/constant.cmo \ @@ -53,7 +54,8 @@ PROOFS=proofs/typing_ev.cmo proofs/tacred.cmo \ proofs/refiner.cmo proofs/evar_refiner.cmo \ proofs/macros.cmo proofs/tacinterp.cmo # proofs/clenv.cmo -TACTICS=tactics/dn.cmo +TACTICS=tactics/dn.cmo tactics/termdn.cmo tactics/btermdn.cmo \ + tactics/nbtermdn.cmo PRETYPING=pretyping/astterm.cmo diff --git a/dev/changements.txt b/dev/changements.txt index 2bdcfbc4a3..543b0fa699 100644 --- a/dev/changements.txt +++ b/dev/changements.txt @@ -7,6 +7,8 @@ Changements d'organisation / modules : Names -> kernel/names.ml et kernel/sign.ml (les parties noms et signatures ont été séparées) + Avm,Mavm,Fmavm,Mhm -> utiliser plutôt Map (et freeze alors gratuit) + Changements dans les types de données : --------------------------------------- diff --git a/lib/gmap.ml b/lib/gmap.ml index beb0a1994e..52af882b35 100644 --- a/lib/gmap.ml +++ b/lib/gmap.ml @@ -107,3 +107,12 @@ Empty -> accu | Node(l, v, d, r, _) -> fold f l (f v d (fold f r accu)) + +(* Added with respect to ocaml standard library. *) + + let dom m = fold (fun x _ acc -> x::acc) m [] + + let rng m = fold (fun _ y acc -> y::acc) m [] + + let to_list m = fold (fun x y acc -> (x,y)::acc) m [] + diff --git a/lib/gmap.mli b/lib/gmap.mli index 861e730e26..908bad1bcf 100644 --- a/lib/gmap.mli +++ b/lib/gmap.mli @@ -14,3 +14,9 @@ val mem : 'a -> ('a,'b) t -> bool val iter : ('a -> 'b -> unit) -> ('a,'b) t -> unit val map : ('b -> 'c) -> ('a,'b) t -> ('a,'c) t val fold : ('a -> 'b -> 'c -> 'c) -> ('a,'b) t -> 'c -> 'c + +(* Additions with respect to ocaml standard library. *) + +val dom : ('a,'b) t -> 'a list +val rng : ('a,'b) t -> 'b list +val to_list : ('a,'b) t -> ('a * 'b) list diff --git a/lib/tlm.ml b/lib/tlm.ml index 9b61de9336..d16353f1ec 100644 --- a/lib/tlm.ml +++ b/lib/tlm.ml @@ -1,92 +1,56 @@ (* $Id$ *) -(* 1er choix : une liste -module MySet = struct - type 'a t = 'a list - let mt = [] - let add = add_set - let rmv = rmv_set - let toList l = l - let app = List.map -end -*) +type ('a,'b) t = Node of 'b Gset.t * ('a, ('a,'b) t) Gmap.t -(* 2 ème choix : un arbre *) -module MySet = struct - type 'a t = 'a Coq_set.t - let mt = Coq_set.empty - let add = Coq_set.add - let rmv = Coq_set.remove - let toList = Coq_set.elements - let app f l = Coq_set.fold (fun a b -> add (f a) b) l mt -end +let empty = Node (Gset.empty, Gmap.empty) -module type MyMapType = sig - type ('a, 'b) t - val create : unit -> ('a,'b) t - val map : ('a,'b) t -> 'a -> 'b - val dom : ('a,'b) t -> 'a list - val rng : ('a,'b) t -> 'b list - val in_dom : ('a,'b) t -> 'a -> bool - val add : ('a,'b) t -> 'a * 'b -> ('a,'b) t - val remap : ('a,'b) t -> 'a -> 'b -> ('a,'b) t - val app : (('a * 'c) -> unit) -> ('a,'c) t -> unit - val toList : ('a,'b) t -> ('a * 'b) list -end;; +let map (Node (_,m)) lbl = Gmap.find lbl m -module MyMap = (Listmap : MyMapType);; +let xtract (Node (hereset,_)) = Gset.elements hereset -type ('a,'b) t = - NODE of 'b MySet.t * ('a, ('a,'b) t) MyMap.t;; +let dom (Node (_,m)) = Gmap.dom m -let create () = NODE(MySet.mt,MyMap.create());; +let in_dom (Node (_,m)) lbl = Gmap.mem lbl m -let map (NODE (_,m)) lbl = MyMap.map m lbl;; -let xtract (NODE (hereset,_)) = MySet.toList hereset;; -let dom (NODE (_,m)) = MyMap.dom m;; -let in_dom (NODE (_,m)) lbl = MyMap.in_dom m lbl;; - -let is_empty_node (NODE(a,b)) = (MySet.toList a = []) & (MyMap.toList b = []);; +let is_empty_node (Node(a,b)) = (Gset.elements a = []) & (Gmap.to_list b = []) let assure_arc m lbl = - if MyMap.in_dom m lbl then m - else MyMap.add m (lbl,NODE (MySet.mt,MyMap.create())) -;; - -let cleanse_arcs (NODE (hereset,m)) = -let l = MyMap.rng m -in NODE(hereset,if List.for_all is_empty_node l then MyMap.create() else m) -;; - -let rec at_path f (NODE (hereset,m)) = function - [] -> cleanse_arcs(NODE(f hereset,m)) + if Gmap.mem lbl m then + m + else + Gmap.add lbl (Node (Gset.empty,Gmap.empty)) m + +let cleanse_arcs (Node (hereset,m)) = + let l = Gmap.rng m in + Node(hereset, if List.for_all is_empty_node l then Gmap.empty else m) + +let rec at_path f (Node (hereset,m)) = function + | [] -> + cleanse_arcs (Node(f hereset,m)) | h::t -> - let m = assure_arc m h - in cleanse_arcs(NODE(hereset, - MyMap.remap m h (at_path f (MyMap.map m h) t))) -;; + let m = assure_arc m h in + cleanse_arcs (Node(hereset, + Gmap.add h (at_path f (Gmap.find h m) t) m)) let add tm (path,v) = - at_path (fun hereset -> MySet.add v hereset) tm path -;; - + at_path (fun hereset -> Gset.add v hereset) tm path + let rmv tm (path,v) = - at_path (fun hereset -> MySet.rmv v hereset) tm path -;; + at_path (fun hereset -> Gset.remove v hereset) tm path let app f tlm = - let rec apprec pfx (NODE(hereset,m)) = - let path = List.rev pfx - in (MySet.app (fun v -> f(path,v)) hereset; - MyMap.app (fun (l,tm) -> apprec (l::pfx) tm) m) - in apprec [] tlm -;; + let rec apprec pfx (Node(hereset,m)) = + let path = List.rev pfx in + Gset.iter (fun v -> f(path,v)) hereset; + Gmap.iter (fun l tm -> apprec (l::pfx) tm) m + in + apprec [] tlm -let toList tlm = - let rec torec pfx (NODE(hereset,m)) = - let path = List.rev pfx - in List.flatten((List.map (fun v -> (path,v)) (MySet.toList hereset)):: - (List.map (fun (l,tm) -> torec (l::pfx) tm) (MyMap.toList m))) - in torec [] tlm -;; +let to_list tlm = + let rec torec pfx (Node(hereset,m)) = + let path = List.rev pfx in + List.flatten((List.map (fun v -> (path,v)) (Gset.elements hereset)):: + (List.map (fun (l,tm) -> torec (l::pfx) tm) (Gmap.to_list m))) + in + torec [] tlm diff --git a/lib/tlm.mli b/lib/tlm.mli index 41af7c1aba..44a2adc8bf 100644 --- a/lib/tlm.mli +++ b/lib/tlm.mli @@ -1,9 +1,12 @@ (* $Id$ *) +(* Tries. This module implements a data structure [('a,'b) t] mapping lists + of values of type ['a] to sets (as lists) of values of type ['b]. *) + type ('a,'b) t -val create : unit -> ('a,'b) t +val empty : ('a,'b) t (* Work on labels, not on paths. *) @@ -17,6 +20,6 @@ val in_dom : ('a,'b) t -> 'a -> bool val add : ('a,'b) t -> 'a list * 'b -> ('a,'b) t val rmv : ('a,'b) t -> ('a list * 'b) -> ('a,'b) t -val app : (('a list * 'c) -> unit) -> ('a,'c) t -> unit +val app : (('a list * 'b) -> unit) -> ('a,'b) t -> unit val to_list : ('a,'b) t -> ('a list * 'b) list diff --git a/tactics/btermdn.ml b/tactics/btermdn.ml new file mode 100644 index 0000000000..3015e651b1 --- /dev/null +++ b/tactics/btermdn.ml @@ -0,0 +1,52 @@ + +(* $Id$ *) + +open Term +open Termdn + +(* Discrimination nets with bounded depth. + See the module dn.ml for further explanations. + Eduardo (5/8/97). *) + +let dnet_depth = ref 8 + +let bounded_constr_pat_discr (t,depth) = + if depth = 0 then + None + else + match constr_pat_discr t with + | None -> None + | Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l) + +let bounded_constr_val_discr (t,depth) = + if depth = 0 then + None + else + match constr_val_discr t with + | None -> None + | Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l) + +type 'a t = (lbl,constr * int,'a) Dn.under_t + +let newdn () = Dn.create bounded_constr_pat_discr + +let ex_termdn = newdn() + +let inDN tdn = { + Dn.args = ex_termdn.Dn.args; + Dn.tm = tdn } + +let outDN dn = dn.Dn.tm + +let create () = outDN (newdn()) + +let add dn (c,v) = outDN (Dn.add (inDN dn) ((c,!dnet_depth),v)) + +let rmv dn (c,v) = outDN (Dn.rmv (inDN dn) ((c,!dnet_depth),v)) + +let lookup dn t = + List.map + (fun ((c,_),v) -> (c,v)) + (Dn.lookup (inDN dn) bounded_constr_val_discr (t,!dnet_depth)) + +let app f dn = Dn.app (fun ((c,_),v) -> f(c,v)) (inDN dn) diff --git a/tactics/btermdn.mli b/tactics/btermdn.mli new file mode 100644 index 0000000000..882f0bbb9c --- /dev/null +++ b/tactics/btermdn.mli @@ -0,0 +1,21 @@ + +(* $Id$ *) + +(*i*) +open Generic +open Term +(*i*) + +(* Discrimination nets with bounded depth. *) + +type 'a t + +val create : unit -> 'a t + +val add : 'a t -> (constr * 'a) -> 'a t +val rmv : 'a t -> (constr * 'a) -> 'a t + +val lookup : 'a t -> constr -> (constr * 'a) list +val app : ((constr * 'a) -> unit) -> 'a t -> unit + +val dnet_depth : int ref diff --git a/tactics/dn.ml b/tactics/dn.ml index 55112ba798..33ce0df5db 100644 --- a/tactics/dn.ml +++ b/tactics/dn.ml @@ -31,7 +31,7 @@ type ('lbl,'pat,'inf) t = { tm : ('lbl,'pat,'inf) under_t; args : ('lbl,'pat) dn_args } -let create dna = {tm = Tlm.create(); args = dna} +let create dna = { tm = Tlm.empty; args = dna } let path_of dna = let rec path_of_deferred = function diff --git a/tactics/nbtermdn.ml b/tactics/nbtermdn.ml new file mode 100644 index 0000000000..ec5de39c03 --- /dev/null +++ b/tactics/nbtermdn.ml @@ -0,0 +1,75 @@ + +(* $Id$ *) + +open Util +open Names +open Generic +open Term +open Libobject +open Library + +(* Named, bounded-depth, term-discrimination nets. + Implementation: + Term-patterns are stored in discrimination-nets, which are + themselves stored in a hash-table, indexed by the first label. + They are also stored by name in a table on-the-side, so that we can + override them if needed. *) + +(* The former comments are from Chet. + See the module dn.ml for further explanations. + Eduardo (5/8/97) *) + +type ('na,'a) t = { + mutable table : ('na,constr * 'a) Gmap.t; + mutable patterns : (Termdn.lbl option,'a Btermdn.t) Gmap.t } + +type ('na,'a) frozen_t = + ('na,constr * 'a) Gmap.t * (Termdn.lbl option,'a Btermdn.t) Gmap.t + +let create () = + { table = Gmap.empty; + patterns = Gmap.empty } + +let get_dn dnm hkey = + try Gmap.find hkey dnm with Not_found -> Btermdn.create () + +let add dn (na,(pat,valu)) = + let hkey = option_app fst (Termdn.constr_pat_discr pat) in + dn.table <- Gmap.add na (pat,valu) dn.table; + let dnm = dn.patterns in + dn.patterns <- Gmap.add hkey (Btermdn.add (get_dn dnm hkey) (pat,valu)) dnm + +let rmv dn na = + let (pat,valu) = Gmap.find na dn.table in + let hkey = option_app fst (Termdn.constr_pat_discr pat) in + dn.table <- Gmap.remove na dn.table; + let dnm = dn.patterns in + dn.patterns <- Gmap.add hkey (Btermdn.rmv (get_dn dnm hkey) (pat,valu)) dnm + +let in_dn dn na = Gmap.mem na dn.table + +let remap ndn na (pat,valu) = + rmv ndn na; + add ndn (na,(pat,valu)) + +let lookup dn valu = + let hkey = option_app fst (Termdn.constr_val_discr valu) in + try Btermdn.lookup (Gmap.find hkey dn.patterns) valu with Not_found -> [] + +let app f dn = Gmap.iter f dn.table + +let dnet_depth = Btermdn.dnet_depth + +let freeze dn = (dn.table, dn.patterns) + +let unfreeze (fnm,fdnm) dn = + dn.table <- fnm; + dn.patterns <- fdnm + +let empty dn = + dn.table <- Gmap.empty; + dn.patterns <- Gmap.empty + +let to2lists dn = + (Gmap.to_list dn.table, Gmap.to_list dn.patterns) + diff --git a/tactics/nbtermdn.mli b/tactics/nbtermdn.mli new file mode 100644 index 0000000000..e94a21e9d5 --- /dev/null +++ b/tactics/nbtermdn.mli @@ -0,0 +1,30 @@ + +(* $Id$ *) + +(*i*) +open Generic +open Term +(*i*) + +(* Named, bounded-depth, term-discrimination nets. *) + +type ('na,'a) t +type ('na,'a) frozen_t + +val create : unit -> ('na,'a) t + +val add : ('na,'a) t -> ('na * (constr * 'a)) -> unit +val rmv : ('na,'a) t -> 'na -> unit +val in_dn : ('na,'a) t -> 'na -> bool +val remap : ('na,'a) t -> 'na -> (constr * 'a) -> unit + +val lookup : ('na,'a) t -> constr -> (constr * 'a) list +val app : ('na -> (constr * 'a) -> unit) -> ('na,'a) t -> unit + +val dnet_depth : int ref + +val freeze : ('na,'a) t -> ('na,'a) frozen_t +val unfreeze : ('na,'a) frozen_t -> ('na,'a) t -> unit +val empty : ('na,'a) t -> unit +val to2lists : ('na,'a) t -> ('na * (Term.constr * 'a)) list * + (Termdn.lbl option * 'a Btermdn.t) list diff --git a/tactics/termdn.ml b/tactics/termdn.ml new file mode 100644 index 0000000000..f1f8374c9b --- /dev/null +++ b/tactics/termdn.ml @@ -0,0 +1,68 @@ + +(* $Id$ *) + +open Util +open Generic +open Term + +(* Discrimination nets of terms. + See the module dn.ml for further explanations. + Eduardo (5/8/97) *) + +type lbl = + | TERM of constr + | DOPER of sorts oper + | DLAMBDA + +type 'a t = (lbl,constr,'a) Dn.t + +(*If we have: f a b c ..., decomp gives: (f,[a;b;c;...])*) + +let decomp = + let rec decrec acc = function + | DOPN(AppL,cl) -> decrec (array_app_tl cl acc) (array_hd cl) + | DOP2(Cast,c1,_) -> decrec acc c1 + | c -> (c,acc) + in + decrec [] + +let constr_pat_discr t = + if not(occur_meta t) then + None + else + match decomp t with + (* | DOPN(Const _,_) as c,l -> Some(TERM c,l) *) + | DOPN(MutInd _,_) as c,l -> Some(TERM c,l) + | DOPN(MutConstruct _,_) as c,l -> Some(TERM c,l) + | VAR _ as c,l -> Some(TERM c,l) + | c -> None + +let constr_val_discr t = + match decomp t with + (* DOPN(Const _,_) as c,l -> Some(TERM c,l) *) + | DOPN(MutInd _,_) as c,l -> Some(TERM c,l) + | DOPN(MutConstruct _,_) as c,l -> Some(TERM c,l) + | VAR _ as c,l -> Some(TERM c,l) + | c -> None + +(* Les deux fonctions suivantes ecrasaient les precedentes, + ajout d'un suffixe _nil CP 16/08 *) + +let constr_pat_discr_nil t = + match constr_pat_discr t with + | None -> None + | Some (c,_) -> Some(c,[]) + +let constr_val_discr_nil t = + match constr_val_discr t with + | None -> None + | Some (c,_) -> Some(c,[]) + +let create () = Dn.create constr_pat_discr + +let add = Dn.add +let rmv = Dn.rmv + +let lookup dn t = Dn.lookup dn constr_val_discr t + +let app f dn = Dn.app f dn diff --git a/tactics/termdn.mli b/tactics/termdn.mli new file mode 100644 index 0000000000..95cfc62438 --- /dev/null +++ b/tactics/termdn.mli @@ -0,0 +1,27 @@ + +(* $Id$ *) + +(*i*) +open Generic +open Term +(*i*) + +(* Discrimination nets of terms. *) + +type lbl = + | TERM of constr + | DOPER of sorts oper + | DLAMBDA + +type 'a t = (lbl,constr,'a) Dn.t + +val create : unit -> 'a t + +val add : 'a t -> (constr * 'a) -> 'a t +val rmv : 'a t -> (constr * 'a) -> 'a t + +val lookup : 'a t -> constr -> (constr * 'a) list +val app : ((constr * 'a) -> unit) -> 'a t -> unit + +val constr_pat_discr : constr -> (lbl * constr list) option +val constr_val_discr : constr -> (lbl * constr list) option |
