123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197(************************************************************************)(* * The Coq Proof Assistant / The Coq Development Team *)(* v * Copyright INRIA, CNRS and contributors *)(* <O___,, * (see version control and CREDITS file for authors & dates) *)(* \VV/ **************************************************************)(* // * This file is distributed under the terms of the *)(* * GNU Lesser General Public License Version 2.1 *)(* * (see LICENSE file for the text of the license) *)(************************************************************************)openUtilopenConstropenEConstropenNamesopenPattern(* Discrimination nets with bounded depth.
See the module dn.ml for further explanations.
Eduardo (5/8/97). *)letdnet_depth=ref8typeterm_label=|GRLabelofGlobRef.t|ProdLabel|LambdaLabel|SortLabelletcompare_term_labelt1t2=matcht1,t2with|GRLabelgr1,GRLabelgr2->GlobRef.CanOrd.comparegr1gr2|_->pervasives_comparet1t2(** OK *)type'reslookup_res='resDn.lookup_res=Labelof'res|Nothing|Everythingletdecomp_patp=letrecdecrecacc=function|PApp(f,args)->decrec(Array.to_listargs@acc)f|PProj(p,c)->lethole=PMetaNoneinletparams=List.make(Projection.nparsp)holein(PRef(GlobRef.ConstRef(Projection.constantp)),params@c::acc)|c->(c,acc)indecrec[]pletdecompsigmat=letrecdecrecaccc=matchEConstr.kindsigmacwith|App(f,l)->decrec(Array.fold_right(funal->a::l)lacc)f|Proj(p,c)->(* Hack: fake evar to generate [Everything] in the functions below *)lethole=mkEvar(Evar.unsafe_of_int(-1),[])inletparams=List.make(Projection.nparsp)holein(mkConst(Projection.constantp),params@c::acc)|Cast(c1,_,_)->decrecaccc1|_->(c,acc)indecrec[]tletevaluable_constantcenv=(* This is a hack to work around a broken Print Module implementation, see
bug #2668. *)ifEnviron.mem_constantcenvthenEnviron.evaluable_constantcenvelsetrueletconstr_val_discrenvsigmat=letopenGlobRefinletc,l=decompsigmatinmatchEConstr.kindsigmacwith|Ind(ind_sp,u)->Label(GRLabel(IndRefind_sp),l)|Construct(cstr_sp,u)->Label(GRLabel(ConstructRefcstr_sp),l)|Varid->Label(GRLabel(VarRefid),l)|Const(c,_)->ifevaluable_constantcenvthenEverythingelseLabel(GRLabel(ConstRefc),l)|_->Nothingletconstr_pat_discrenvt=ifnot(Patternops.occur_meta_patternt)thenNoneelseletopenGlobRefinmatchdecomp_pattwith|PRef((IndRef_)asref),args|PRef((ConstructRef_)asref),args->Some(GRLabelref,args)|PRef((VarRefv)asref),args->Some(GRLabelref,args)|PRef((ConstRefc)asref),args->ifevaluable_constantcenvthenNoneelseSome(GRLabelref,args)|_->Noneletconstr_val_discr_stenvsigmatst=letc,l=decompsigmatinletopenGlobRefinmatchEConstr.kindsigmacwith|Const(c,u)->ifevaluable_constantcenv&&TransparentState.is_transparent_constanttscthenEverythingelseLabel(GRLabel(ConstRefc),l)|Ind(ind_sp,u)->Label(GRLabel(IndRefind_sp),l)|Construct(cstr_sp,u)->Label(GRLabel(ConstructRefcstr_sp),l)|Varid->ifEnviron.evaluable_namedidenv&&TransparentState.is_transparent_variabletsidthenEverythingelseLabel(GRLabel(VarRefid),l)|Prod(n,d,c)->Label(ProdLabel,[d;c])|Lambda(n,d,c)->ifList.is_emptylthenLabel(LambdaLabel,[d;c]@l)elseEverything|Sort_->Label(SortLabel,[])|Evar_->Everything|Rel_|Meta_|Cast_|LetIn_|App_|Case_|Fix_|CoFix_|Proj_|Int_|Float_|Array_->Nothingletconstr_pat_discr_stenvtst=letopenGlobRefinmatchdecomp_pattwith|PRef((IndRef_)asref),args|PRef((ConstructRef_)asref),args->Some(GRLabelref,args)|PRef((VarRefv)asref),args->ifEnviron.evaluable_namedvenv&&(TransparentState.is_transparent_variabletsv)thenNoneelseSome(GRLabelref,args)|PRef((ConstRefc)asref),args->ifevaluable_constantcenv&&TransparentState.is_transparent_constanttscthenNoneelseSome(GRLabelref,args)|PVarv,argswhennot(TransparentState.is_transparent_variabletsv)->Some(GRLabel(VarRefv),args)|PProd(_,d,c),[]->Some(ProdLabel,[d;c])|PLambda(_,d,c),[]->Some(LambdaLabel,[d;c])|PSorts,[]->Some(SortLabel,[])|_->Noneletbounded_constr_pat_discr_stenvst(t,depth)=ifInt.equaldepth0thenNoneelsematchconstr_pat_discr_stenvsttwith|None->None|Some(c,l)->Some(c,List.map(func->(c,depth-1))l)letbounded_constr_val_discr_stenvsigmast(t,depth)=ifInt.equaldepth0thenNothingelsematchconstr_val_discr_stenvsigmasttwith|Label(c,l)->Label(c,List.map(func->(c,depth-1))l)|Nothing->Nothing|Everything->Everythingletbounded_constr_pat_discrenv(t,depth)=ifInt.equaldepth0thenNoneelsematchconstr_pat_discrenvtwith|None->None|Some(c,l)->Some(c,List.map(func->(c,depth-1))l)letbounded_constr_val_discrenvsigma(t,depth)=ifInt.equaldepth0thenNothingelsematchconstr_val_discrenvsigmatwith|Label(c,l)->Label(c,List.map(func->(c,depth-1))l)|Nothing->Nothing|Everything->EverythingmoduleMake=functor(Z:Map.OrderedType)->structmoduleY=structtypet=term_labelletcompare=compare_term_labelendmoduleDn=Dn.Make(Y)(Z)typet=Dn.ttypepattern=Dn.patternletpatternenvstpat=matchstwith|None->Dn.pattern(bounded_constr_pat_discrenv)(pat,!dnet_depth)|Somest->Dn.pattern(bounded_constr_pat_discr_stenvst)(pat,!dnet_depth)letempty=Dn.emptyletadd=Dn.addletrmv=Dn.rmvletlookupenvsigma=function|None->(fundnt->Dn.lookupdn(bounded_constr_val_discrenvsigma)(t,!dnet_depth))|Somest->(fundnt->Dn.lookupdn(bounded_constr_val_discr_stenvsigmast)(t,!dnet_depth))end