123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467(************************************************************************)(* * 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) *)(************************************************************************)(*i*)openUtilopenConstropenNamesopenGlobnamesopenMod_substopenPp(* debug *)(*i*)(* Representation/approximation of terms to use in the dnet:
*
* - no meta or evar (use ['a pattern] for that)
*
* - [Rel]s and [Sort]s are not taken into account (that's why we need
* a second pass of linear filterin on the results - it's not a perfect
* term indexing structure)
* - Foralls and LetIns are represented by a context DCtx (a list of
* generalization, similar to rel_context, and coded with DCons and
* DNil). This allows for matching under an unfinished context
*)moduleDTerm=structtype'tt=|DRel|DSort|DRefofGlobRef.t|DCtxof't*'t(* (binding list, subterm) = Prods and LetIns *)|DLambdaof't*'t|DAppof't*'t(* binary app *)|DCaseofcase_info*'t*'t*'tarray|DFixofintarray*int*'tarray*'tarray|DCoFixofint*'tarray*'tarray|DIntofUint63.t|DFloatofFloat64.t|DArrayof'tarray*'t*'t(* special constructors only inside the left-hand side of DCtx or
DApp. Used to encode lists of foralls/letins/apps as contexts *)|DConsof('t*'toption)*'t|DNil(* debug *)let_pr_dconstrf:'at->Pp.t=function|DRel->str"*"|DSort->str"Sort"|DRef_->str"Ref"|DCtx(ctx,t)->fctx++spc()++str"|-"++spc()++ft|DLambda(t1,t2)->str"fun"++spc()++ft1++spc()++str"->"++spc()++ft2|DApp(t1,t2)->ft1++spc()++ft2|DCase(_,t1,t2,ta)->str"case"|DFix_->str"fix"|DCoFix_->str"cofix"|DInt_->str"INT"|DFloat_->str"FLOAT"|DCons((t,dopt),tl)->ft++(matchdoptwithSomet'->str":="++ft'|None->str"")++spc()++str"::"++spc()++ftl|DNil->str"[]"|DArray_->str"ARRAY"(*
* Functional iterators for the t datatype
* a.k.a boring and error-prone boilerplate code
*)letmapf=function|(DRel|DSort|DNil|DRef_|DInt_|DFloat_)asc->c|DCtx(ctx,c)->DCtx(fctx,fc)|DLambda(t,c)->DLambda(ft,fc)|DApp(t,u)->DApp(ft,fu)|DCase(ci,p,c,bl)->DCase(ci,fp,fc,Array.mapfbl)|DFix(ia,i,ta,ca)->DFix(ia,i,Array.mapfta,Array.mapfca)|DCoFix(i,ta,ca)->DCoFix(i,Array.mapfta,Array.mapfca)|DCons((t,topt),u)->DCons((ft,Option.mapftopt),fu)|DArray(t,def,ty)->DArray(Array.mapft,fdef,fty)letcompare_cici1ci2=letc=Ind.CanOrd.compareci1.ci_indci2.ci_indinifc=0thenletc=Int.compareci1.ci_nparci2.ci_nparinifc=0thenletc=Array.compareInt.compareci1.ci_cstr_ndeclsci2.ci_cstr_ndeclsinifc=0thenArray.compareInt.compareci1.ci_cstr_nargsci2.ci_cstr_nargselsecelsecelsecletcomparecmpt1t2=matcht1,t2with|DRel,DRel->0|DRel,_->-1|_,DRel->1|DSort,DSort->0|DSort,_->-1|_,DSort->1|DRefgr1,DRefgr2->GlobRef.CanOrd.comparegr1gr2|DRef_,_->-1|_,DRef_->1|DCtx(tl1,tr1),DCtx(tl2,tr2)|DLambda(tl1,tr1),DLambda(tl2,tr2)|DApp(tl1,tr1),DApp(tl2,tr2)->letc=cmptl1tl2inifc=0thencmptr1tr2elsec|DCtx_,_->-1|_,DCtx_->1|DLambda_,_->-1|_,DLambda_->1|DApp_,_->-1|_,DApp_->1|DCase(ci1,c1,t1,p1),DCase(ci2,c2,t2,p2)->letc=cmpc1c2inifc=0thenletc=cmpt1t2inifc=0thenletc=Array.comparecmpp1p2inifc=0thencompare_cici1ci2elsecelsecelsec|DCase_,_->-1|_,DCase_->1|DFix(i1,j1,tl1,pl1),DFix(i2,j2,tl2,pl2)->letc=Int.comparej1j2inifc=0thenletc=Array.compareInt.comparei1i2inifc=0thenletc=Array.comparecmptl1tl2inifc=0thenArray.comparecmppl1pl2elsecelsecelsec|DFix_,_->-1|_,DFix_->1|DCoFix(i1,tl1,pl1),DCoFix(i2,tl2,pl2)->letc=Int.comparei1i2inifc=0thenletc=Array.comparecmptl1tl2inifc=0thenArray.comparecmppl1pl2elsecelsec|DCoFix_,_->-1|_,DCoFix_->1|DInti1,DInti2->Uint63.comparei1i2|DInt_,_->-1|_,DInt_->1|DFloatf1,DFloatf2->Float64.total_comparef1f2|DFloat_,_->-1|_,DFloat_->1|DArray(t1,def1,ty1),DArray(t2,def2,ty2)->letc=Array.comparecmpt1t2inifc=0thenletc=cmpdef1def2inifc=0thencmpty1ty2elsecelsec|DArray_,_->-1|_,DArray_->1|DCons((t1,ot1),u1),DCons((t2,ot2),u2)->letc=cmpt1t2inifInt.equalc0thenletc=Option.comparecmpot1ot2inifInt.equalc0thencmpu1u2elsecelsec|DCons_,_->-1|_,DCons_->1|DNil,DNil->0letfoldfacc=function|(DRel|DNil|DSort|DRef_|DInt_|DFloat_)->acc|DCtx(ctx,c)->f(faccctx)c|DLambda(t,c)->f(facct)c|DApp(t,u)->f(facct)u|DCase(ci,p,c,bl)->Array.fold_leftf(f(faccp)c)bl|DFix(ia,i,ta,ca)->Array.fold_leftf(Array.fold_leftfaccta)ca|DCoFix(i,ta,ca)->Array.fold_leftf(Array.fold_leftfaccta)ca|DArray(t,def,ty)->f(f(Array.fold_leftfacct)def)ty|DCons((t,topt),u)->f(Option.fold_leftf(facct)topt)uletchoosef=function|(DRel|DSort|DNil|DRef_|DInt_|DFloat_)->invalid_arg"choose"|DCtx(ctx,c)->fctx|DLambda(t,c)->ft|DApp(t,u)->fu|DCase(ci,p,c,bl)->fc|DFix(ia,i,ta,ca)->fta.(0)|DCoFix(i,ta,ca)->fta.(0)|DCons((t,topt),u)->fu|DArray(t,def,ty)->ft.(0)letdummy_cmp()()=0letfold2(f:'a->'b->'c->'a)(acc:'a)(c1:'bt)(c2:'ct):'a=letheadw=map(fun_->())winifnot(Int.equal(comparedummy_cmp(headc1)(headc2))0)theninvalid_arg"fold2:compare"elsematchc1,c2with|(DRel,DRel|DNil,DNil|DSort,DSort|DRef_,DRef_|DInt_,DInt_|DFloat_,DFloat_)->acc|(DCtx(c1,t1),DCtx(c2,t2)|DApp(c1,t1),DApp(c2,t2)|DLambda(c1,t1),DLambda(c2,t2))->f(faccc1c2)t1t2|DCase(ci,p1,c1,bl1),DCase(_,p2,c2,bl2)->Array.fold_left2f(f(faccp1p2)c1c2)bl1bl2|DFix(ia,i,ta1,ca1),DFix(_,_,ta2,ca2)->Array.fold_left2f(Array.fold_left2faccta1ta2)ca1ca2|DCoFix(i,ta1,ca1),DCoFix(_,ta2,ca2)->Array.fold_left2f(Array.fold_left2faccta1ta2)ca1ca2|DArray(t1,def1,ty1),DArray(t2,def2,ty2)->f(f(Array.fold_left2facct1t2)def1def2)ty1ty2|DCons((t1,topt1),u1),DCons((t2,topt2),u2)->f(Option.fold_left2f(facct1t2)topt1topt2)u1u2|(DRel|DNil|DSort|DRef_|DCtx_|DApp_|DLambda_|DCase_|DFix_|DCoFix_|DCons_|DInt_|DFloat_|DArray_),_->assertfalseletmap2(f:'a->'b->'c)(c1:'at)(c2:'bt):'ct=letheadw=map(fun_->())winifnot(Int.equal(comparedummy_cmp(headc1)(headc2))0)theninvalid_arg"map2_t:compare"elsematchc1,c2with|(DRel,DRel|DSort,DSort|DNil,DNil|DRef_,DRef_|DInt_,DInt_|DFloat_,DFloat_)ascc->let(c,_)=ccinc|DCtx(c1,t1),DCtx(c2,t2)->DCtx(fc1c2,ft1t2)|DLambda(t1,c1),DLambda(t2,c2)->DLambda(ft1t2,fc1c2)|DApp(t1,u1),DApp(t2,u2)->DApp(ft1t2,fu1u2)|DCase(ci,p1,c1,bl1),DCase(_,p2,c2,bl2)->DCase(ci,fp1p2,fc1c2,Array.map2fbl1bl2)|DFix(ia,i,ta1,ca1),DFix(_,_,ta2,ca2)->DFix(ia,i,Array.map2fta1ta2,Array.map2fca1ca2)|DCoFix(i,ta1,ca1),DCoFix(_,ta2,ca2)->DCoFix(i,Array.map2fta1ta2,Array.map2fca1ca2)|DArray(t1,def1,ty1),DArray(t2,def2,ty2)->DArray(Array.map2ft1t2,fdef1def2,fty1ty2)|DCons((t1,topt1),u1),DCons((t2,topt2),u2)->DCons((ft1t2,Option.lift2ftopt1topt2),fu1u2)|(DRel|DNil|DSort|DRef_|DCtx_|DApp_|DLambda_|DCase_|DFix_|DCoFix_|DCons_|DInt_|DFloat_|DArray_),_->assertfalseletterminal=function|(DRel|DSort|DNil|DRef_|DInt_|DFloat_)->true|DLambda_|DApp_|DCase_|DFix_|DCoFix_|DCtx_|DCons_|DArray_->falseletcomparet1t2=comparedummy_cmpt1t2end(*
* Terms discrimination nets
* Uses the general dnet datatype on DTerm.t
* (here you can restart reading)
*)(*
* Construction of the module
*)moduletypeIDENT=sigtypetvalcompare:t->t->intvalsubst:substitution->t->tvalconstr_of:t->constrendmoduletypeOPT=sigvalreduce:constr->constrvaldirection:boolendmoduleMake=functor(Ident:IDENT)->functor(Opt:OPT)->structmoduleTDnet:Dnet.Swithtypeident=Ident.tandtype'astructure='aDTerm.tandtypemeta=int=Dnet.Make(DTerm)(Ident)(Int)typet=TDnet.ttypeident=TDnet.ident(** We will freshen metas on the fly, to cope with the implementation defect
of Term_dnet which requires metas to be all distinct. *)letfresh_meta=letindex=ref0infun()->letans=!indexinlet()=index:=succansinansopenDTermopenTDnetletpat_of_constrc:term_pattern=letopenGlobRefin(* To each evar we associate a unique identifier. *)letmetas=refEvar.Map.emptyinletrecpat_of_constrc=matchConstr.kindcwith|Rel_->TermDRel|Sort_->TermDSort|Vari->Term(DRef(VarRefi))|Const(c,u)->Term(DRef(ConstRefc))|Ind(i,u)->Term(DRef(IndRefi))|Construct(c,u)->Term(DRef(ConstructRefc))|Meta_->assertfalse|Evar(i,_)->letmeta=tryEvar.Map.findi!metaswithNot_found->letmeta=fresh_meta()inlet()=metas:=Evar.Map.addimeta!metasinmetainMetameta|Case(ci,u1,pms1,c1,_iv,c2,ca)->letf_ctx(_,p)=pat_of_constrpinTerm(DCase(ci,f_ctxc1,pat_of_constrc2,Array.mapf_ctxca))|Fix((ia,i),(_,ta,ca))->Term(DFix(ia,i,Array.mappat_of_constrta,Array.mappat_of_constrca))|CoFix(i,(_,ta,ca))->Term(DCoFix(i,Array.mappat_of_constrta,Array.mappat_of_constrca))|Cast(c,_,_)->pat_of_constrc|Lambda(_,t,c)->Term(DLambda(pat_of_constrt,pat_of_constrc))|(Prod_|LetIn_)->let(ctx,c)=ctx_of_constr(TermDNil)cinTerm(DCtx(ctx,c))|App(f,ca)->Array.fold_left(funca->Term(DApp(c,a)))(pat_of_constrf)(Array.mappat_of_constrca)|Proj(p,c)->Term(DApp(Term(DRef(ConstRef(Projection.constantp))),pat_of_constrc))|Inti->Term(DInti)|Floatf->Term(DFloatf)|Array(_u,t,def,ty)->Term(DArray(Array.mappat_of_constrt,pat_of_constrdef,pat_of_constrty))andctx_of_constrctxc=matchConstr.kindcwith|Prod(_,t,c)->ctx_of_constr(Term(DCons((pat_of_constrt,None),ctx)))c|LetIn(_,d,t,c)->ctx_of_constr(Term(DCons((pat_of_constrt,Some(pat_of_constrd)),ctx)))c|_->ctx,pat_of_constrcinpat_of_constrcletempty_ctx:term_pattern->term_pattern=function|Meta_asc->c|Term(DCtx(_,_))asc->c|c->Term(DCtx(TermDNil,c))(*
* Basic primitives
*)letempty=TDnet.emptyletsubstst=letsleafid=Ident.substsidinletsnode=function|DTerm.DRefgr->DTerm.DRef(fst(subst_globalsgr))|n->ninTDnet.mapsleafsnodetletunion=TDnet.unionletadd(c:constr)(id:Ident.t)(dn:t)=letc=Opt.reducecinletc=empty_ctx(pat_of_constrc)inTDnet.adddncidletnew_meta()=Meta(fresh_meta())letrecremove_cap:term_pattern->term_pattern=function|Term(DCons(t,u))->Term(DCons(t,remove_capu))|TermDNil->new_meta()|Meta_asm->m|_->assertfalseletunder_prod:term_pattern->term_pattern=function|Term(DCtx(t,u))->Term(DCtx(remove_capt,u))|Metam->Term(DCtx(new_meta(),Metam))|_->assertfalse(* debug *)(* let rec pr_term_pattern p =
(fun pr_t -> function
| Term t -> pr_t t
| Meta m -> str"["++Pp.int (Obj.magic m)++str"]"
) (pr_dconstr pr_term_pattern) p*)letsearch_patcpatdpatdn=letwhole_c=EConstr.of_constrcpatin(* if we are at the root, add an empty context *)letdpat=under_prod(empty_ctxdpat)inTDnet.Idset.fold(funidacc->letc_id=Opt.reduce(Ident.constr_ofid)inletc_id=EConstr.of_constrc_idinlet(ctx,wc)=tryTermops.align_prod_letinEvd.emptywhole_cc_id(* FIXME *)withInvalid_argument_->[],c_idinletwc,whole_c=ifOpt.directionthenwhole_c,wcelsewc,whole_cintrylet_=Termops.filteringEvd.emptyctxReduction.CUMULwcwhole_cinid::accwithTermops.CannotFilter->(* msgnl(str"recon "++Termops.print_constr_env (Global.env()) wc); *)acc)(TDnet.find_matchdpatdn)[](*
* High-level primitives describing specific search problems
*)letsearch_patterndnpat=letpat=Opt.reducepatinsearch_patpat(empty_ctx(pat_of_constrpat))dnletfind_alldn=Idset.elements(TDnet.find_alldn)letmapfdn=TDnet.mapf(funx->x)dnletrefresh_metasdn=letnew_metas=refInt.Map.emptyinletrefresh_one_metai=tryInt.Map.findi!new_metaswithNot_found->letnew_meta=fresh_meta()inlet()=new_metas:=Int.Map.addinew_meta!new_metasinnew_metainTDnet.map_metasrefresh_one_metadnendmoduletypeS=sigtypettypeidentvalempty:tvaladd:constr->ident->t->tvalunion:t->t->tvalsubst:substitution->t->tvalsearch_pattern:t->constr->identlistvalfind_all:t->identlistvalmap:(ident->ident)->t->tvalrefresh_metas:t->tend