1234567891011121314151617181920212223242526272829303132333435363738394041424344454647484950515253545556575859606162636465666768697071727374757677787980818283848586878889909192939495969798991001011021031041051061071081091101111121131141151161171181191201211221231241251261271281291301311321331341351361371381391401411421431441451461471481491501511521531541551561571581591601611621631641651661671681691701711721731741751761771781791801811821831841851861871881891901911921931941951961971981992002012022032042052062072082092102112122132142152162172182192202212222232242252262272282292302312322332342352362372382392402412422432442452462472482492502512522532542552562572582592602612622632642652662672682692702712722732742752762772782792802812822832842852862872882892902912922932942952962972982993003013023033043053063073083093103113123133143153163173183193203213223233243253263273283293303313323333343353363373383393403413423433443453463473483493503513523533543553563573583593603613623633643653663673683693703713723733743753763773783793803813823833843853863873883893903913923933943953963973983994004014024034044054064074084094104114124134144154164174184194204214224234244254264274284294304314324334344354364374384394404414424434444454464474484494504514524534544554564574584594604614624634644654664674684694704714724734744754764774784794804814824834844854864874884894904914924934944954964974984995005015025035045055065075085095105115125135145155165175185195205215225235245255265275285295305315325335345355365375385395405415425435445455465475485495505515525535545555565575585595605615625635645655665675685695705715725735745755765775785795805815825835845855865875885895905915925935945955965975985996006016026036046056066076086096106116126136146156166176186196206216226236246256266276286296306316326336346356366376386396406416426436446456466476486496506516526536546556566576586596606616626636646656666676686696706716726736746756766776786796806816826836846856866876886896906916926936946956966976986997007017027037047057067077087097107117127137147157167177187197207217227237247257267277287297307317327337347357367377387397407417427437447457467477487497507517527537547557567577587597607617627637647657667677687697707717727737747757767777787797807817827837847857867877887897907917927937947957967977987998008018028038048058068078088098108118128138148158168178188198208218228238248258268278288298308318328338348358368378388398408418428438448458468478488498508518528538548558568578588598608618628638648658668678688698708718728738748758768778788798808818828838848858868878888898908918928938948958968978988999009019029039049059069079089099109119129139149159169179189199209219229239249259269279289299309319329339349359369379389399409419429439449459469479489499509519529539549559569579589599609619629639649659669679689699709719729739749759769779789799809819829839849859869879889899909919929939949959969979989991000100110021003100410051006100710081009101010111012101310141015101610171018101910201021102210231024102510261027102810291030103110321033103410351036103710381039104010411042104310441045104610471048104910501051105210531054105510561057105810591060106110621063106410651066106710681069107010711072107310741075107610771078107910801081108210831084108510861087108810891090109110921093109410951096109710981099110011011102110311041105110611071108110911101111111211131114111511161117111811191120112111221123112411251126112711281129113011311132113311341135113611371138113911401141114211431144114511461147114811491150115111521153115411551156115711581159116011611162116311641165116611671168116911701171117211731174117511761177117811791180118111821183118411851186118711881189119011911192119311941195119611971198119912001201120212031204120512061207120812091210121112121213121412151216121712181219122012211222122312241225122612271228122912301231123212331234123512361237123812391240124112421243124412451246124712481249125012511252125312541255125612571258125912601261126212631264126512661267126812691270127112721273127412751276127712781279128012811282128312841285128612871288128912901291129212931294129512961297129812991300130113021303130413051306130713081309131013111312131313141315131613171318131913201321132213231324132513261327132813291330133113321333133413351336133713381339134013411342134313441345134613471348134913501351135213531354135513561357135813591360136113621363136413651366136713681369137013711372137313741375137613771378137913801381138213831384138513861387138813891390139113921393139413951396139713981399140014011402140314041405140614071408140914101411141214131414141514161417141814191420142114221423142414251426142714281429143014311432143314341435143614371438143914401441144214431444144514461447144814491450145114521453145414551456145714581459146014611462146314641465146614671468146914701471147214731474147514761477147814791480148114821483148414851486148714881489149014911492149314941495149614971498149915001501150215031504150515061507150815091510151115121513151415151516151715181519152015211522152315241525152615271528152915301531153215331534153515361537153815391540154115421543154415451546154715481549155015511552155315541555155615571558155915601561156215631564156515661567156815691570157115721573157415751576157715781579158015811582158315841585158615871588158915901591159215931594159515961597159815991600160116021603160416051606160716081609161016111612161316141615161616171618161916201621162216231624162516261627162816291630163116321633163416351636163716381639164016411642164316441645164616471648164916501651165216531654165516561657165816591660166116621663166416651666166716681669167016711672167316741675167616771678167916801681168216831684168516861687168816891690169116921693169416951696169716981699170017011702170317041705170617071708170917101711171217131714171517161717171817191720172117221723172417251726172717281729(************************************************************************)(* * 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) *)(************************************************************************)openPpopenUtilopenCErrorsopenNamesopenConstropenContextopenEvdopenEConstropenVarsopenEnvironopenMod_substopenGlobnamesopenLibobjectopenNamegenopenLibnamesopenTermopsopenInductiveopsopenTypeclassesopenPatternopenPatternopsopenClenvopenTacredopenPrintermoduleNamedDecl=Context.Named.Declaration(****************************************)(* General functions *)(****************************************)typedebug=Debug|Info|OffexceptionBoundletrechead_boundsigmat=matchEConstr.kindsigmatwith|Prod(_,_,b)->head_boundsigmab|LetIn(_,_,_,b)->head_boundsigmab|App(c,_)->head_boundsigmac|Case(_,_,_,_,_,c,_)->head_boundsigmac|Ind(ind,_)->GlobRef.IndRefind|Const(c,_)->GlobRef.ConstRefc|Construct(c,_)->GlobRef.ConstructRefc|Varid->GlobRef.VarRefid|Proj(p,_)->GlobRef.ConstRef(Projection.constantp)|Cast(c,_,_)->head_boundsigmac|Evar_|Rel_|Meta_|Sort_|Fix_|Lambda_|CoFix_|Int_|Float_|Array_->raiseBoundlethead_constrsigmac=tryhead_boundsigmacwithBound->user_err(Pp.str"Head identifier must be a constant, section variable, \
(co)inductive type, (co)inductive type constructor, or projection.")letdecompose_app_boundsigmat=lett=strip_outer_castsigmatinlet_,ccl=decompose_prod_assumsigmatinlethd,args=decompose_app_vectsigmacclinletopenGlobRefinmatchEConstr.kindsigmahdwith|Const(c,u)->ConstRefc,args|Ind(i,u)->IndRefi,args|Construct(c,u)->ConstructRefc,args|Varid->VarRefid,args|Proj(p,c)->ConstRef(Projection.constantp),Array.conscargs|_->raiseBound(** Compute the set of section variables that remain in the named context.
Starts from the top to the bottom of the context, stops at the first
different declaration between the named hyps and the section context. *)letsecvars_of_hypshyps=letsecctx=Global.named_context()inletopenContext.Named.Declarationinletpred,all=List.fold_left(fun(pred,all)decl->trylet_=Context.Named.lookup(get_iddecl)hypsin(* Approximation, it might be an hypothesis reintroduced with same name and unconvertible types,
we must allow it currently, as comparing the declarations for syntactic equality is too
strong a check (e.g. an unfold in a section variable would make it unusable). *)(Id.Pred.add(get_iddecl)pred,all)withNot_found->(pred,false))(Id.Pred.empty,true)secctxinifallthenId.Pred.full(* If the whole section context is available *)elsepredletempty_hint_info={hint_priority=None;hint_pattern=None}(************************************************************************)(* The Type of Constructions Autotactic Hints *)(************************************************************************)type'ahint_ast=|Res_pfof'a(* Hint Apply *)|ERes_pfof'a(* Hint EApply *)|Give_exactof'a|Res_pf_THEN_trivial_failof'a(* Hint Immediate *)|Unfold_nthofevaluable_global_reference(* Hint Unfold *)|ExternofPattern.constr_patternoption*Genarg.glob_generic_argument(* Hint Extern *)type'ahints_path_atom_gen=|PathHintsof'alist(* For forward hints, their names is the list of projections *)|PathAnytypehints_path_atom=GlobRef.thints_path_atom_gentype'ahints_path_gen=|PathAtomof'ahints_path_atom_gen|PathStarof'ahints_path_gen|PathSeqof'ahints_path_gen*'ahints_path_gen|PathOrof'ahints_path_gen*'ahints_path_gen|PathEmpty|PathEpsilontypepre_hints_path=Libnames.qualidhints_path_gentypehints_path=GlobRef.thints_path_gentypehint_term=|IsGlobRefofGlobRef.t|IsConstrofconstr*Univ.ContextSet.toption(* None if monomorphic *)type'awith_uid={obj:'a;uid:KerName.t;}typeraw_hint=constr*types*Univ.ContextSet.toptiontypehint={hint_term:constr;hint_type:types;hint_uctx:Univ.ContextSet.toption;(* None if monomorphic *)hint_clnv:clausenv;}type'awith_metadata={pri:int(** A number lower is higher priority *);pat:constr_patternoption(** A pattern for the concl of the Goal *);name:hints_path_atom(** A potential name to refer to the hint *);db:stringoption(** The database from which the hint comes *);secvars:Id.Pred.t(** The set of section variables the hint depends on *);code:'a(** the tactic to apply when the concl matches pat *)}typefull_hint=hinthint_astwith_uidwith_metadatatypehint_entry=GlobRef.toption*raw_hinthint_astwith_uidwith_metadatatypehint_mode=|ModeInput(* No evars *)|ModeNoHeadEvar(* No evar at the head *)|ModeOutput(* Anything *)type'ahints_transparency_target=|HintsVariables|HintsConstants|HintsReferencesof'alisttypeimport_level=HintLax|HintWarn|HintStrictletwarn_hint_to_string=function|HintLax->"Lax"|HintWarn->"Warn"|HintStrict->"Strict"letstring_to_warn_hint=function|"Lax"->HintLax|"Warn"->HintWarn|"Strict"->HintStrict|_->user_errPp.(str"Only the following values are accepted: Lax, Warn, Strict.")letwarn_hint=Goptions.declare_interpreted_string_option_and_ref~depr:false~key:["Loose";"Hint";"Behavior"]~value:HintLaxstring_to_warn_hintwarn_hint_to_stringletfresh_key=letid=Summary.ref~name:"HINT-COUNTER"0infun()->letcur=incrid;!idinletlbl=Id.of_string("_"^string_of_intcur)inletkn=Lib.make_knlblinlet(mp,_)=KerName.reprknin(* We embed the full path of the kernel name in the label so that
the identifier should be unique. This ensures that including
two modules together won't confuse the corresponding labels. *)letlbl=Id.of_string_soft(Printf.sprintf"%s#%i"(ModPath.to_stringmp)cur)inKerName.makemp(Label.of_idlbl)letpri_order_int(id1,{pri=pri1})(id2,{pri=pri2})=letd=pri1-pri2inifInt.equald0thenid2-id1elsedletpri_ordert1t2=pri_order_intt1t2<=0(* Nov 98 -- Papageno *)(* Les Hints sont ré-organisés en plusieurs databases.
La table impérative "searchtable", de type "hint_db_table",
associe une database (hint_db) à chaque nom.
Une hint_db est une table d'association fonctionelle constr -> search_entry
Le constr correspond à la constante de tête de la conclusion.
Une search_entry est un triplet comprenant :
- la liste des tactiques qui n'ont pas de pattern associé
- la liste des tactiques qui ont un pattern
- un discrimination net borné (Btermdn.t) constitué de tous les
patterns de la seconde liste de tactiques *)typestored_data=int*full_hint(* First component is the index of insertion in the table, to keep most recent first semantics. *)moduleBounded_net:sigtypetvalempty:tvaladd:TransparentState.toption->t->Pattern.constr_pattern->stored_data->tvallookup:Environ.env->Evd.evar_map->TransparentState.toption->t->EConstr.constr->stored_datalistend=structmoduleData=structtypet=stored_dataletcompare=pri_order_intendmoduleBnet=Btermdn.Make(Data)typediff=Pattern.constr_pattern*stored_datatypedata=BnetofBnet.t|Diffofdiff*datareftypet=datarefletempty=ref(BnetBnet.empty)letadd_stnetpv=ref(Diff((p,v),net))letrecforceenvstnet=match!netwith|Bnetdn->dn|Diff((p,v),rem)->letdn=forceenvstreminletp=Bnet.patternenvstpinletdn=Bnet.adddnpvinlet()=net:=(Bnetdn)indnletlookupenvsigmastnetp=letdn=forceenvstnetinBnet.lookupenvsigmastdnpendtypesearch_entry={sentry_nopat:stored_datalist;sentry_pat:stored_datalist;sentry_bnet:Bounded_net.t;sentry_mode:hint_modearraylist;}letempty_se={sentry_nopat=[];sentry_pat=[];sentry_bnet=Bounded_net.empty;sentry_mode=[];}leteq_pri_auto_tactic(_,x)(_,y)=KerName.equalx.code.uidy.code.uidletadd_tacpattse=matchpatwith|None->ifList.exists(eq_pri_auto_tactict)se.sentry_nopatthenseelse{sewithsentry_nopat=List.insertpri_ordertse.sentry_nopat}|Some(st,pat)->ifList.exists(eq_pri_auto_tactict)se.sentry_patthenseelse{sewithsentry_pat=List.insertpri_ordertse.sentry_pat;sentry_bnet=Bounded_net.addstse.sentry_bnetpatt;}letrebuild_dnstse=letdn'=List.fold_left(fundn(id,t)->Bounded_net.add(Somest)dn(Option.gett.pat)(id,t))Bounded_net.emptyse.sentry_patin{sewithsentry_bnet=dn'}letlookup_tacsenvsigmaconclstse=letl'=Bounded_net.lookupenvsigmastse.sentry_bnetconclinletsl'=List.stable_sortpri_order_intl'inList.mergepri_order_intse.sentry_nopatsl'letis_transparent_grts=letopenGlobRefinfunction|VarRefid->TransparentState.is_transparent_variabletsid|ConstRefcst->TransparentState.is_transparent_constanttscst|IndRef_|ConstructRef_->falseletstrip_paramsenvsigmac=matchEConstr.kindsigmacwith|App(f,args)->(matchEConstr.kindsigmafwith|Const(cst,_)->(matchStructures.PrimitiveProjections.find_optcstwith|Somep->letp=Projection.makepfalseinletnpars=Projection.nparspinifArray.lengthargs>nparsthenmkApp(mkProj(p,args.(npars)),Array.subargs(npars+1)(Array.lengthargs-(npars+1)))elsec|None->c)|_->c)|_->cletmerge_context_set_optsigmactx=matchctxwith|None->sigma|Somectx->Evd.merge_context_setEvd.univ_flexiblesigmactxletinstantiate_hintenvsigmap=letmk_clenv(c,cty,ctx)=letsigma=merge_context_set_optsigmactxinletcl=mk_clenv_from_envenvsigmaNone(c,cty)inlettemplval={cl.templvalwithrebus=strip_paramsenvsigmacl.templval.rebus}inletcl=mk_clausenvempty_envcl.evdtemplvalcl.templtypin{hint_term=c;hint_type=cty;hint_uctx=ctx;hint_clnv=cl;}inletcode=matchp.code.objwith|Res_pfc->Res_pf(mk_clenvc)|ERes_pfc->ERes_pf(mk_clenvc)|Res_pf_THEN_trivial_failc->Res_pf_THEN_trivial_fail(mk_clenvc)|Give_exactc->Give_exact(mk_clenvc)|(Unfold_nth_|Extern_)ash->hin{pwithcode={p.codewithobj=code}}lethints_path_atom_eqh1h2=matchh1,h2with|PathHintsl1,PathHintsl2->List.equalGlobRef.equall1l2|PathAny,PathAny->true|_->falseletrechints_path_eqh1h2=matchh1,h2with|PathAtomh1,PathAtomh2->hints_path_atom_eqh1h2|PathStarh1,PathStarh2->hints_path_eqh1h2|PathSeq(l1,r1),PathSeq(l2,r2)->hints_path_eql1l2&&hints_path_eqr1r2|PathOr(l1,r1),PathOr(l2,r2)->hints_path_eql1l2&&hints_path_eqr1r2|PathEmpty,PathEmpty->true|PathEpsilon,PathEpsilon->true|_->falseletpath_matcheshphints=letrecauxhphintsk=matchhp,hintswith|PathAtom_,[]->false|PathAtomPathAny,(_::hints')->khints'|PathAtomp,(h::hints')->ifhints_path_atom_eqphthenkhints'elsefalse|PathStarhp',hints->khints||auxhp'hints(funhints'->auxhphints'k)|PathSeq(hp,hp'),hints->auxhphints(funhints'->auxhp'hints'k)|PathOr(hp,hp'),hints->auxhphintsk||auxhp'hintsk|PathEmpty,_->false|PathEpsilon,hints->khintsinauxhphints(funhints'->true)letrecmatches_epsilon=function|PathAtom_->false|PathStar_->true|PathSeq(p,p')->matches_epsilonp&&matches_epsilonp'|PathOr(p,p')->matches_epsilonp||matches_epsilonp'|PathEmpty->false|PathEpsilon->trueletrecis_empty=function|PathAtom_->false|PathStar_->false|PathSeq(p,p')->is_emptyp||is_emptyp'|PathOr(p,p')->matches_epsilonp&&matches_epsilonp'|PathEmpty->true|PathEpsilon->falseletpath_seqpp'=matchp,p'with|PathEpsilon,p'->p'|p,PathEpsilon->p|p,p'->PathSeq(p,p')letrecpath_derivatehphint=letrecderivate_atomshintshints'=matchhints,hints'with|gr::grs,gr'::grs'whenGlobRef.equalgrgr'->derivate_atomsgrsgrs'|[],[]->PathEpsilon|[],hints->PathEmpty|grs,[]->PathAtom(PathHintsgrs)|_,_->PathEmptyinmatchhpwith|PathAtomPathAny->PathEpsilon|PathAtom(PathHintsgrs)->(matchgrs,hintwith|h::_,PathAny->PathEmpty|hints,PathHintshints'->derivate_atomshintshints'|_,_->assertfalse)|PathStarp->ifpath_matchesp[hint]thenhpelsePathEpsilon|PathSeq(hp,hp')->lethpder=path_derivatehphintinifmatches_epsilonhpthenPathOr(path_seqhpderhp',path_derivatehp'hint)elseifis_emptyhpderthenPathEmptyelsepath_seqhpderhp'|PathOr(hp,hp')->PathOr(path_derivatehphint,path_derivatehp'hint)|PathEmpty->PathEmpty|PathEpsilon->PathEmptyletrecnormalize_pathh=matchhwith|PathStarPathEpsilon->PathEpsilon|PathSeq(PathEmpty,_)|PathSeq(_,PathEmpty)->PathEmpty|PathSeq(PathEpsilon,p)|PathSeq(p,PathEpsilon)->normalize_pathp|PathOr(PathEmpty,p)|PathOr(p,PathEmpty)->normalize_pathp|PathOr(p,q)->letp',q'=normalize_pathp,normalize_pathqinifhints_path_eqpp'&&hints_path_eqqq'thenhelsenormalize_path(PathOr(p',q'))|PathSeq(p,q)->letp',q'=normalize_pathp,normalize_pathqinifhints_path_eqpp'&&hints_path_eqqq'thenhelsenormalize_path(PathSeq(p',q'))|_->hletpath_derivatehphint=normalize_path(path_derivatehphint)letpp_hints_path_atomprga=matchawith|PathAny->str"_"|PathHintsgrs->pr_sequenceprggrsletpp_hints_path_genprg=letrecaux=function|PathAtompa->pp_hints_path_atomprgpa|PathStar(PathAtomPathAny)->str"_*"|PathStarp->str"("++auxp++str")*"|PathSeq(p,p')->auxp++spc()++auxp'|PathOr(p,p')->str"("++auxp++spc()++str"|"++cut()++spc()++auxp'++str")"|PathEmpty->str"emp"|PathEpsilon->str"eps"inauxletpp_hints_path=pp_hints_path_genpr_globalletglob_hints_path_atomp=matchpwith|PathHintsg->PathHints(List.mapNametab.globalg)|PathAny->PathAnyletglob_hints_path=letrecaux=function|PathAtompa->PathAtom(glob_hints_path_atompa)|PathStarp->PathStar(auxp)|PathSeq(p,p')->PathSeq(auxp,auxp')|PathOr(p,p')->PathOr(auxp,auxp')|PathEmpty->PathEmpty|PathEpsilon->PathEpsiloninauxletsubst_path_atomsubstp=matchpwith|PathAny->p|PathHintsgrs->letgr'gr=fst(subst_globalsubstgr)inletgrs'=List.Smart.mapgr'grsinifgrs'==grsthenpelsePathHintsgrs'letrecsubst_hints_pathsubsthp=matchhpwith|PathAtomp->letp'=subst_path_atomsubstpinifp'==pthenhpelsePathAtomp'|PathStarp->letp'=subst_hints_pathsubstpinifp'==pthenhpelsePathStarp'|PathSeq(p,q)->letp'=subst_hints_pathsubstpinletq'=subst_hints_pathsubstqinifp'==p&&q'==qthenhpelsePathSeq(p',q')|PathOr(p,q)->letp'=subst_hints_pathsubstpinletq'=subst_hints_pathsubstqinifp'==p&&q'==qthenhpelsePathOr(p',q')|_->hptypehint_db_name=stringtype'awith_mode=|ModeMatchof'a|ModeMismatchmoduleHint_db:sigtypetvalempty:?name:hint_db_name->TransparentState.t->bool->tvalmap_none:secvars:Id.Pred.t->t->full_hintlistvalmap_all:secvars:Id.Pred.t->GlobRef.t->t->full_hintlistvalmap_existential:evar_map->secvars:Id.Pred.t->(GlobRef.t*constrarray)->constr->t->full_hintlistwith_modevalmap_eauto:Environ.env->evar_map->secvars:Id.Pred.t->(GlobRef.t*constrarray)->constr->t->full_hintlistwith_modevalmap_auto:Environ.env->evar_map->secvars:Id.Pred.t->(GlobRef.t*constrarray)->constr->t->full_hintlistvaladd_list:env->evar_map->hint_entrylist->t->tvalremove_one:Environ.env->GlobRef.t->t->tvalremove_list:Environ.env->GlobRef.tlist->t->tvaliter:(GlobRef.toption->hint_modearraylist->full_hintlist->unit)->t->unitvaluse_dn:t->boolvaltransparent_state:t->TransparentState.tvalset_transparent_state:t->TransparentState.t->tvaladd_cut:hints_path->t->tvaladd_mode:GlobRef.t->hint_modearray->t->tvalcut:t->hints_pathvalunfolds:t->Id.Set.t*Cset.tvaladd_modes:hint_modearraylistGlobRef.Map.t->t->tvalmodes:t->hint_modearraylistGlobRef.Map.tvalfold:(GlobRef.toption->hint_modearraylist->full_hintlist->'a->'a)->t->'a->'aend=structtypet={hintdb_state:TransparentState.t;hintdb_cut:hints_path;hintdb_unfolds:Id.Set.t*Cset.t;hintdb_max_id:int;use_dn:bool;hintdb_map:search_entryGlobRef.Map.t;(* A list of unindexed entries starting with an unfoldable constant
or with no associated pattern. *)hintdb_nopat:(GlobRef.toption*stored_data)list;hintdb_name:stringoption;}letnext_hint_iddb=leth=db.hintdb_max_idin{dbwithhintdb_max_id=succdb.hintdb_max_id},hletempty?namestuse_dn={hintdb_state=st;hintdb_cut=PathEmpty;hintdb_unfolds=(Id.Set.empty,Cset.empty);hintdb_max_id=0;use_dn=use_dn;hintdb_map=GlobRef.Map.empty;hintdb_nopat=[];hintdb_name=name;}letfindkeydb=tryGlobRef.Map.findkeydb.hintdb_mapwithNot_found->empty_seletrealize_tacsecvars(id,tac)=ifId.Pred.subsettac.secvarssecvarsthenSometacelse(* Warn about no longer typable hint? *)Nonelethead_evarsigmac=letrechrecc=matchEConstr.kindsigmacwith|Evar(evk,_)->evk|App(c,_)->hrecc|Cast(c,_,_)->hrecc|_->raiseEvarutil.NoHeadEvarinhreccletmatch_modesigmamarg=matchmwith|ModeInput->not(occur_existentialsigmaarg)|ModeNoHeadEvar->(tryignore(head_evarsigmaarg);falsewithEvarutil.NoHeadEvar->true)|ModeOutput->trueletmatches_modesigmaargsmode=Array.lengthmode==Array.lengthargs&&Array.for_all2(match_modesigma)modeargsletmatches_modessigmaargsmodes=ifList.is_emptymodesthentrueelseList.exists(matches_modesigmaargs)modesletmerge_entrysecvarsdbnopatpat=leth=List.sortpri_order_int(List.mapsnddb.hintdb_nopat)inleth=List.mergepri_order_inthnopatinleth=List.mergepri_order_inthpatinList.map_filter(realize_tacsecvars)hletmap_none~secvarsdb=merge_entrysecvarsdb[][]letmap_all~secvarskdb=letse=findkdbinmerge_entrysecvarsdbse.sentry_nopatse.sentry_pat(* Precondition: concl has no existentials *)letmap_autoenvsigma~secvars(k,args)concldb=letse=findkdbinletst=ifdb.use_dnthen(Somedb.hintdb_state)elseNoneinletpat=lookup_tacsenvsigmaconclstseinmerge_entrysecvarsdb[]patletmap_existentialsigma~secvars(k,args)concldb=letse=findkdbinifmatches_modessigmaargsse.sentry_modethenModeMatch(merge_entrysecvarsdbse.sentry_nopatse.sentry_pat)elseModeMismatch(* [c] contains an existential *)letmap_eautoenvsigma~secvars(k,args)concldb=letse=findkdbinifmatches_modessigmaargsse.sentry_modethenletst=ifdb.use_dnthenSomedb.hintdb_stateelseNoneinletpat=lookup_tacsenvsigmaconclstseinModeMatch(merge_entrysecvarsdb[]pat)elseModeMismatchletis_exact=function|Give_exact_->true|_->falseletis_unfold=function|Unfold_nth_->true|_->falseletaddkvgridvdb=letidv=id,{vwithdb=db.hintdb_name}inletk=matchgrwith|Somegr->ifdb.use_dn&&is_transparent_grdb.hintdb_stategr&&is_unfoldv.code.objthenNoneelseSomegr|None->Noneinmatchkwith|None->letis_present(_,(_,v'))=KerName.equalv.code.uidv'.code.uidinifnot(List.existsis_presentdb.hintdb_nopat)then(* FIXME *){dbwithhintdb_nopat=(gr,idv)::db.hintdb_nopat}elsedb|Somegr->letpat=ifnotdb.use_dn&&is_exactv.code.objthenNoneelseletdnst=ifdb.use_dnthenSomedb.hintdb_stateelseNoneinOption.map(funp->(dnst,p))v.patinletoval=findgrdbin{dbwithhintdb_map=GlobRef.Map.addgr(add_tacpatidvoval)db.hintdb_map}letrebuild_dbst'db=letdb'={dbwithhintdb_map=GlobRef.Map.map(rebuild_dnst')db.hintdb_map;hintdb_state=st';hintdb_nopat=[]}inList.fold_left(fundb(gr,(id,v))->addkvgridvdb)db'db.hintdb_nopatletadd_oneenvsigma(k,v)db=letv=instantiate_hintenvsigmavinletst',db,rebuild=matchv.code.objwith|Unfold_nthegr->letaddunfts(ids,csts)=letopenTransparentStateinmatchegrwith|EvalVarRefid->{tswithtr_var=Id.Pred.addidts.tr_var},(Id.Set.addidids,csts)|EvalConstRefcst->{tswithtr_cst=Cpred.addcstts.tr_cst},(ids,Cset.addcstcsts)inletstate,unfs=addunfdb.hintdb_statedb.hintdb_unfoldsinstate,{dbwithhintdb_unfolds=unfs},true|_->db.hintdb_state,db,falseinletdb=ifdb.use_dn&&rebuildthenrebuild_dbst'dbelsedbinletdb,id=next_hint_iddbinaddkvkidvdbletadd_listenvsigmaldb=List.fold_left(fundbk->add_oneenvsigmakdb)dblletremove_sdlpsdl=List.filterpsdlletremove_hestpse=letsl1'=remove_sdlpse.sentry_nopatinletsl2'=remove_sdlpse.sentry_patinifsl1'==se.sentry_nopat&&sl2'==se.sentry_patthenseelserebuild_dnst{sewithsentry_nopat=sl1';sentry_pat=sl2'}letremove_listenvgrsdb=letfilter(_,h)=matchh.namewithPathHints[gr]->not(List.mem_fGlobRef.equalgrgrs)|_->trueinlethintmap=GlobRef.Map.map(remove_hedb.hintdb_statefilter)db.hintdb_mapinlethintnopat=List.filter(fun(ge,sd)->filtersd)db.hintdb_nopatin{dbwithhintdb_map=hintmap;hintdb_nopat=hintnopat}letremove_oneenvgrdb=remove_listenv[gr]dbletget_entryse=leth=List.mergepri_order_intse.sentry_nopatse.sentry_patinList.mapsndhletiterfdb=letiter_sekse=f(Somek)se.sentry_mode(get_entryse)infNone[](List.map(funx->snd(sndx))db.hintdb_nopat);GlobRef.Map.iteriter_sedb.hintdb_mapletfoldfdbaccu=letaccu=fNone[](List.map(funx->snd(sndx))db.hintdb_nopat)accuinGlobRef.Map.fold(funkse->f(Somek)se.sentry_mode(get_entryse))db.hintdb_mapacculettransparent_statedb=db.hintdb_stateletset_transparent_statedbst=ifdb.use_dnthenrebuild_dbstdbelse{dbwithhintdb_state=st}letadd_cutpathdb={dbwithhintdb_cut=normalize_path(PathOr(db.hintdb_cut,path))}letadd_modegrmdb=letse=findgrdbinletse={sewithsentry_mode=m::se.sentry_mode}in{dbwithhintdb_map=GlobRef.Map.addgrsedb.hintdb_map}letcutdb=db.hintdb_cutletunfoldsdb=db.hintdb_unfoldsletadd_modesmodesdb=letfgreme=Some{ewithsentry_mode=me.sentry_mode@e.sentry_mode}inletmode_entries=GlobRef.Map.map(funm->{empty_sewithsentry_mode=m})modesin{dbwithhintdb_map=GlobRef.Map.unionfdb.hintdb_mapmode_entries}letmodesdb=GlobRef.Map.map(funse->se.sentry_mode)db.hintdb_mapletuse_dndb=db.use_dnendmoduleHintdbmap=String.Maptypehint_db=Hint_db.tletsearchtable=Summary.ref~name:"searchtable"Hintdbmap.emptyletstatustable=Summary.ref~name:"statustable"KNmap.emptyletsearchtable_mapname=Hintdbmap.findname!searchtableletsearchtable_add(name,db)=searchtable:=Hintdbmap.addnamedb!searchtableletcurrent_db_names()=Hintdbmap.domain!searchtableletcurrent_db()=Hintdbmap.bindings!searchtableletcurrent_pure_db()=List.mapsnd(current_db())leterror_no_such_hint_databasex=user_err~hdr:"Hints"(str"No such Hint database: "++strx++str".")(**************************************************************************)(* Auxiliary functions to prepare AUTOHINT objects *)(**************************************************************************)letrecnb_hypsigmac=matchEConstr.kindsigmacwith|Prod(_,_,c2)->ifnoccurnsigma1c2then1+(nb_hypsigmac2)elsenb_hypsigmac2|_->0(* adding and removing tactics in the search table *)letwith_uidc={obj=c;uid=fresh_key()}letsecvars_of_idsets=Id.Set.fold(funidp->ifis_section_variableidthenId.Pred.addidpelsep)sId.Pred.emptyletsecvars_of_constrenvsigmac=secvars_of_idset(Termops.global_vars_setenvsigmac)letsecvars_of_globalenvgr=secvars_of_idset(vars_of_globalenvgr)letmake_exact_entryenvsigmainfo?(name=PathAny)(c,cty,ctx)=letsecvars=secvars_of_constrenvsigmacinletcty=strip_outer_castsigmactyinmatchEConstr.kindsigmactywith|Prod_->failwith"make_exact_entry"|_->lethd=tryhead_boundsigmactywithBound->failwith"make_exact_entry"inletpri=matchinfo.hint_prioritywithNone->0|Somep->pinletpat=matchinfo.hint_patternwith|Somepat->sndpat|None->Patternops.pattern_of_constrenvsigma(EConstr.to_constr~abort_on_undefined_evars:falsesigmacty)in(Somehd,{pri;pat=Somepat;name;db=None;secvars;code=with_uid(Give_exact(c,cty,ctx));})letmake_apply_entryenvsigmahnfinfo?(name=PathAny)(c,cty,ctx)=letcty=ifhnfthenhnf_constrenvsigmactyelsectyinmatchEConstr.kindsigmactywith|Prod_->letsigma'=merge_context_set_optsigmactxinletce=mk_clenv_from_envenvsigma'None(c,cty)inletc'=clenv_type(* ~reduce:false *)ceinlethd=tryhead_boundce.evdc'withBound->failwith"make_apply_entry"inletmiss=clenv_missingceinletnmiss=List.lengthmissinletsecvars=secvars_of_constrenvsigmacinletpri=matchinfo.hint_prioritywithNone->nb_hypsigma'cty+nmiss|Somep->pinletpat=matchinfo.hint_patternwith|Somep->sndp|None->Patternops.pattern_of_constrenvce.evd(EConstr.to_constr~abort_on_undefined_evars:falsesigmac')inifInt.equalnmiss0then(Somehd,{pri;pat=Somepat;name;db=None;secvars;code=with_uid(Res_pf(c,cty,ctx));})else(Somehd,{pri;pat=Somepat;name;db=None;secvars;code=with_uid(ERes_pf(c,cty,ctx));})|_->failwith"make_apply_entry"(* flags is (e,h,v) with e=true if eapply and h=true if hnf and v=true if verbose
c is a constr
cty is the type of constr *)letfresh_global_or_constrenvsigmacr=matchcrwith|IsGlobRefgr->let(c,ctx)=UnivGen.fresh_global_instanceenvgrinletctx=ifEnviron.is_polymorphicenvgrthenSomectxelseNonein(EConstr.of_constrc,ctx)|IsConstr(c,ctx)->(c,ctx)letmake_resolvesenvsigma(eapply,hnf)info~check?namecr=letc,ctx=fresh_global_or_constrenvsigmacrinletcty=Retyping.get_type_ofenvsigmacinlettry_applyf=trylet(_,hint)asans=f(c,cty,ctx)inmatchhint.code.objwith|ERes_pf_->ifnoteapplythenNoneelseSomeans|_->SomeanswithFailure_->Noneinletents=List.map_filtertry_apply[make_exact_entryenvsigmainfo?name;make_apply_entryenvsigmahnfinfo?name]inifcheck&&List.is_emptyentsthenuser_err~hdr:"Hint"(pr_leconstr_envenvsigmac++spc()++(ifeapplythenstr"cannot be used as a hint."elsestr"can be used as a hint only for eauto."));ents(* used to add an hypothesis to the local hint database *)letmake_resolve_hypenvsigmahname=letdecl=EConstr.lookup_namedhnameenvinletc=mkVarhnameintry[make_apply_entryenvsigmatrueempty_hint_info~name:(PathHints[GlobRef.VarRefhname])(c,NamedDecl.get_typedecl,None)]with|Failure_->[]|ewhennoncriticale->anomaly(Pp.str"make_resolve_hyp.")(* REM : in most cases hintname = id *)letmake_unfolderef=letg=global_of_evaluable_referenceerefin(Someg,{pri=4;pat=None;name=PathHints[g];db=None;secvars=secvars_of_global(Global.env())g;code=with_uid(Unfold_ntheref)})letmake_externpripattacast=lethdconstr=matchpatwith|None->None|Somec->trySome(head_pattern_boundc)withBoundPattern->user_err(Pp.str"Head pattern or sub-pattern must be a global constant, a section variable, \
an if, case, or let expression, an application, or a projection.")in(hdconstr,{pri=pri;pat=pat;name=PathAny;db=None;secvars=Id.Pred.empty;(* Approximation *)code=with_uid(Extern(pat,tacast))})letmake_moderefm=letopenTerminletty,_=Typeops.type_of_global_in_context(Global.env())refinletctx,t=decompose_prodtyinletn=List.lengthctxinletm'=Array.of_listminifnot(n==Array.lengthm')thenuser_err~hdr:"Hint"(pr_globalref++str" has "++intn++str" arguments while the mode declares "++int(Array.lengthm'))elsem'letmake_trivialenvsigma?(name=PathAny)r=letc,ctx=fresh_global_or_constrenvsigmarinletsigma=merge_context_set_optsigmactxinlett=hnf_constrenvsigma(Retyping.get_type_ofenvsigmac)inlethd=head_constrsigmatinletce=mk_clenv_from_envenvsigmaNone(c,t)in(Somehd,{pri=1;pat=Some(Patternops.pattern_of_constrenvce.evd(EConstr.to_constrsigma(clenv_typece)));name=name;db=None;secvars=secvars_of_constrenvsigmac;code=with_uid(Res_pf_THEN_trivial_fail(c,t,ctx))})(**************************************************************************)(* declaration of the AUTOHINT library object *)(**************************************************************************)(* If the database does not exist, it is created *)(* TODO: should a warning be printed in this case ?? *)letget_dbdbname=trysearchtable_mapdbnamewithNot_found->Hint_db.empty~name:dbnameTransparentState.emptyfalseletadd_hintdbnamehintlist=letcheck(_,h)=let()=ifKNmap.memh.code.uid!statustablethenuser_errPp.(str"Conflicting hint keys. This can happen when including \
twice the same module.")instatustable:=KNmap.addh.code.uidfalse!statustableinlet()=List.itercheckhintlistinletdb=get_dbdbnameinletenv=Global.env()inletsigma=Evd.from_envenvinletdb'=Hint_db.add_listenvsigmahintlistdbinsearchtable_add(dbname,db')letadd_transparencydbnametargetb=letopenTransparentStateinletdb=get_dbdbnameinletst=Hint_db.transparent_statedbinletst'=matchtargetwith|HintsVariables->{stwithtr_var=(ifbthenId.Pred.fullelseId.Pred.empty)}|HintsConstants->{stwithtr_cst=(ifbthenCpred.fullelseCpred.empty)}|HintsReferencesgrs->List.fold_left(funstgr->matchgrwith|EvalConstRefc->{stwithtr_cst=(ifbthenCpred.addelseCpred.remove)cst.tr_cst}|EvalVarRefv->{stwithtr_var=(ifbthenId.Pred.addelseId.Pred.remove)vst.tr_var})stgrsinsearchtable_add(dbname,Hint_db.set_transparent_statedbst')letremove_hintdbnamegrs=letenv=Global.env()inletdb=get_dbdbnameinletdb'=Hint_db.remove_listenvgrsdbinsearchtable_add(dbname,db')letadd_cutdbnamepath=letdb=get_dbdbnameinletdb'=Hint_db.add_cutpathdbinsearchtable_add(dbname,db')letadd_modedbnamelm=letdb=get_dbdbnameinletdb'=Hint_db.add_modelmdbinsearchtable_add(dbname,db')typedb_obj={db_local:bool;db_name:string;db_use_dn:bool;db_ts:TransparentState.t;}letcache_db(_,{db_name=name;db_use_dn=b;db_ts=ts})=searchtable_add(name,Hint_db.empty~nametsb)letload_db_x=cache_dbxletclassify_dbdb=ifdb.db_localthenDisposeelseSubstitutedbletinDB:db_obj->obj=declare_object{(default_object"AUTOHINT_DB")withcache_function=cache_db;load_function=load_db;subst_function=(fun(_,x)->x);classify_function=classify_db;}letcreate_hint_dblntsb=lethint={db_local=l;db_name=n;db_use_dn=b;db_ts=ts}inLib.add_anonymous_leaf(inDBhint)typehint_action=|AddTransparencyof{grefs:evaluable_global_referencehints_transparency_target;state:bool;}|AddHintsofhint_entrylist|RemoveHintsofGlobRef.tlist|AddCutofhints_path|AddModeof{gref:GlobRef.t;mode:hint_modearray}typehint_locality=Local|Export|SuperGlobaltypehint_obj={hint_local:hint_locality;hint_name:string;hint_action:hint_action;}letsuperglobalh=matchh.hint_localwith|SuperGlobal->true|Local|Export->falseletload_autohint_(kn,h)=letname=h.hint_nameinletsuperglobal=superglobalhinmatchh.hint_actionwith|AddTransparency{grefs;state}->ifsuperglobalthenadd_transparencynamegrefsstate|AddHintshints->ifsuperglobalthenadd_hintnamehints|RemoveHintshints->ifsuperglobalthenremove_hintnamehints|AddCutpaths->ifsuperglobalthenadd_cutnamepaths|AddMode{gref;mode}->ifsuperglobalthenadd_modenamegrefmodeletopen_autohinti(kn,h)=letsuperglobal=superglobalhinifInt.equali1thenmatchh.hint_actionwith|AddHintshints->let()=ifnotsuperglobalthen(* Import-bound hints must be declared when not imported yet *)letfilter(_,h)=not@@KNmap.memh.code.uid!statustableinadd_hinth.hint_name(List.filterfilterhints)inletadd(_,hint)=statustable:=KNmap.addhint.code.uidtrue!statustableinList.iteraddhints|AddCutpaths->ifnotsuperglobalthenadd_cuth.hint_namepaths|AddTransparency{grefs;state}->ifnotsuperglobalthenadd_transparencyh.hint_namegrefsstate|RemoveHintshints->ifnotsuperglobalthenremove_hinth.hint_namehints|AddMode{gref;mode}->ifnotsuperglobalthenadd_modeh.hint_namegrefmodeletcache_autohint(kn,obj)=load_autohint1(kn,obj);open_autohint1(kn,obj)letsubst_autohint(subst,obj)=letsubst_keygr=let(gr',t)=subst_globalsubstgrinmatchtwith|None->gr'|Somet->(tryhead_boundEvd.empty(EConstr.of_constrt.Univ.univ_abstracted_value)withBound->gr')inletsubst_mpssubstc=EConstr.of_constr(subst_mpssubst(EConstr.Unsafe.to_constrc))inletsubst_aux((c,t,ctx)ash)=letc'=subst_mpssubstcinlett'=subst_mpssubsttinifc==c'&&t'==tthenhelse(c',t',ctx)inletsubst_hint(k,dataashint)=letk'=Option.Smart.mapsubst_keykinletenv=Global.env()inletsigma=Evd.from_envenvinletpat'=Option.Smart.map(subst_patternenvsigmasubst)data.patinletcode'=matchdata.code.objwith|Res_pfh->leth'=subst_auxhinifh==h'thendata.code.objelseRes_pfh'|ERes_pfh->leth'=subst_auxhinifh==h'thendata.code.objelseERes_pfh'|Give_exacth->leth'=subst_auxhinifh==h'thendata.code.objelseGive_exacth'|Res_pf_THEN_trivial_failh->leth'=subst_auxhinifh==h'thendata.code.objelseRes_pf_THEN_trivial_failh'|Unfold_nthref->letref'=subst_evaluable_referencesubstrefinifref==ref'thendata.code.objelseUnfold_nthref'|Extern(pat,tac)->letpat'=Option.Smart.map(subst_patternenvsigmasubst)data.patinlettac'=Genintern.generic_substitutesubsttacinifpat==pat'&&tac==tac'thendata.code.objelseExtern(pat',tac')inletname'=subst_path_atomsubstdata.nameinletuid'=subst_knsubstdata.code.uidinletdata'=ifdata.code.uid==uid'&&data.pat==pat'&&data.name==name'&&data.code.obj==code'thendataelse{datawithpat=pat';name=name';code={obj=code';uid=uid'}}inifk'==k&&data'==datathenhintelse(k',data')inletaction=matchobj.hint_actionwith|AddTransparency{grefs=target;state=b}->lettarget'=matchtargetwith|HintsVariables->target|HintsConstants->target|HintsReferencesgrs->letgrs'=List.Smart.map(subst_evaluable_referencesubst)grsinifgrs==grs'thentargetelseHintsReferencesgrs'iniftarget'==targetthenobj.hint_actionelseAddTransparency{grefs=target';state=b}|AddHintshints->lethints'=List.Smart.mapsubst_hinthintsinifhints'==hintsthenobj.hint_actionelseAddHintshints'|RemoveHintsgrs->letgrs'=List.Smart.map(subst_global_referencesubst)grsinifgrs==grs'thenobj.hint_actionelseRemoveHintsgrs'|AddCutpath->letpath'=subst_hints_pathsubstpathinifpath'==paththenobj.hint_actionelseAddCutpath'|AddMode{gref=l;mode=m}->letl'=subst_global_referencesubstlinifl'==lthenobj.hint_actionelseAddMode{gref=l';mode=m}inifaction==obj.hint_actionthenobjelse{objwithhint_action=action}letclassify_autohintobj=matchobj.hint_actionwith|AddHints[]->Dispose|_->matchobj.hint_localwith|Local->Dispose|Export|SuperGlobal->SubstituteobjletinAutoHint:hint_obj->obj=declare_object{(default_object"AUTOHINT")withcache_function=cache_autohint;load_function=load_autohint;open_function=simple_openopen_autohint;subst_function=subst_autohint;classify_function=classify_autohint;}letmake_hint~localnameaction={hint_local=local;hint_name=name;hint_action=action;}letwarn_deprecated_hint_without_locality=CWarnings.create~name:"deprecated-hint-without-locality"~category:"deprecated"(fun()->strbrk"The default value for hint locality is currently \
\"local\" in a section and \"global\" otherwise, but is scheduled to change \
in a future release. For the time being, adding hints outside of sections \
without specifying an explicit locality attribute is therefore deprecated. It is \
recommended to use \"export\" whenever possible. Use the attributes \
#[local], #[global] and #[export] depending on your choice. For example: \
\"#[export] Hint Unfold foo : bar.\"")letcheck_hint_locality=letopenGoptionsinfunction|OptGlobal->ifGlobal.sections_are_opened()thenCErrors.user_errPp.(str"This command does not support the global attribute in sections.");|OptExport->ifGlobal.sections_are_opened()thenCErrors.user_errPp.(str"This command does not support the export attribute in sections.");|OptDefault->ifnot@@Global.sections_are_opened()thenwarn_deprecated_hint_without_locality()|OptLocal->()letinterp_locality=function|Goptions.OptDefault|Goptions.OptGlobal->SuperGlobal|Goptions.OptExport->Export|Goptions.OptLocal->Localletremove_hints~localitydbnamesgrs=letlocal=interp_localitylocalityinletdbnames=ifList.is_emptydbnamesthen["core"]elsedbnamesinList.iter(fundbname->lethint=make_hint~localdbname(RemoveHintsgrs)inLib.add_anonymous_leaf(inAutoHinthint))dbnames(**************************************************************************)(* The "Hint" vernacular command *)(**************************************************************************)letadd_resolvesenvsigmaclist~localdbnames=List.iter(fundbname->letr=List.flatten(List.map(fun(pri,hnf,path,gr)->make_resolvesenvsigma(true,hnf)pri~check:true~name:pathgr)clist)inletcheck(_,hint)=matchhint.code.objwith|ERes_pf(c,cty,ctx)->letsigma'=merge_context_set_optsigmactxinletce=mk_clenv_from_envenvsigma'None(c,cty)inletmiss=clenv_missingceinletnmiss=List.lengthmissinletvariables=str(CString.pluralnmiss"variable")inFeedback.msg_info(strbrk"The hint "++pr_leconstr_envenvsigma'c++strbrk" will only be used by eauto, because applying "++pr_leconstr_envenvsigma'c++strbrk" would leave "++variables++Pp.spc()++Pp.prlist_with_sepPp.pr_commaName.print(List.map(Evd.meta_namece.evd)miss)++strbrk" as unresolved existential "++variables++str".")|_->()inlet()=ifnot!Flags.quietthenList.itercheckrinlethint=make_hint~localdbname(AddHintsr)inLib.add_anonymous_leaf(inAutoHinthint))dbnamesletadd_unfoldsl~localdbnames=List.iter(fundbname->lethint=make_hint~localdbname(AddHints(List.mapmake_unfoldl))inLib.add_anonymous_leaf(inAutoHinthint))dbnamesletadd_cutsl~localdbnames=List.iter(fundbname->lethint=make_hint~localdbname(AddCutl)inLib.add_anonymous_leaf(inAutoHinthint))dbnamesletadd_modelm~localdbnames=List.iter(fundbname->letm'=make_modelminlethint=make_hint~localdbname(AddMode{gref=l;mode=m'})inLib.add_anonymous_leaf(inAutoHinthint))dbnamesletadd_transparencylb~localdbnames=List.iter(fundbname->lethint=make_hint~localdbname(AddTransparency{grefs=l;state=b})inLib.add_anonymous_leaf(inAutoHinthint))dbnamesletadd_externinfotacast~localdbname=letpat=matchinfo.hint_patternwith|None->None|Some(_,pat)->Somepatinlethint=make_hint~localdbname(AddHints[make_extern(Option.getinfo.hint_priority)pattacast])inLib.add_anonymous_leaf(inAutoHinthint)letadd_externsinfotacast~localdbnames=List.iter(add_externinfotacast~local)dbnamesletadd_trivialsenvsigmal~localdbnames=List.iter(fundbname->letl=List.map(fun(name,c)->make_trivialenvsigma~namec)linlethint=make_hint~localdbname(AddHintsl)inLib.add_anonymous_leaf(inAutoHinthint))dbnamestypehnf=booltypenonrechint_info=hint_infotypehints_entry=|HintsResolveEntryof(hint_info*hnf*hints_path_atom*hint_term)list|HintsImmediateEntryof(hints_path_atom*hint_term)list|HintsCutEntryofhints_path|HintsUnfoldEntryofevaluable_global_referencelist|HintsTransparencyEntryofevaluable_global_referencehints_transparency_target*bool|HintsModeEntryofGlobRef.t*hint_modelist|HintsExternEntryofhint_info*Genarg.glob_generic_argumentletdefault_prepare_hint_ident=Id.of_string"H"exceptionFoundofconstr*typesletprepare_hintcheckenvinit(sigma,c)=letsigma=Typeclasses.resolve_typeclasses~fail:falseenvsigmain(* We re-abstract over uninstantiated evars and universes.
It is actually a bit stupid to generalize over evars since the first
thing make_resolves will do is to re-instantiate the products *)letsigma=Evd.nf_univ_variablessigmainletc=Evarutil.nf_evarsigmacinletc=drop_extra_implicit_argssigmacinletvars=ref(collect_varssigmac)inletsubst=ref[]inletrecfind_next_evarc=matchEConstr.kindsigmacwith|Evar(evk,argsasev)->(* We skip the test whether args is the identity or not *)lett=Evarutil.nf_evarsigma(existential_typesigmaev)inlett=List.fold_right(fun(e,id)c->replace_termsigmaeidc)!substtinifnot(closed0sigmac)thenuser_errPp.(str"Hints with holes dependent on a bound variable not supported.");ifoccur_existentialsigmatthen(* Not clever enough to construct dependency graph of evars *)user_errPp.(str"Not clever enough to deal with evars dependent in other evars.");raise(Found(c,t))|_->EConstr.itersigmafind_next_evarcinletreciterc=tryfind_next_evarc;cwithFound(evar,t)->letid=next_ident_away_fromdefault_prepare_hint_ident(funid->Id.Set.memid!vars)invars:=Id.Set.addid!vars;subst:=(evar,mkVarid)::!subst;mkNamedLambda(make_annotidSorts.Relevant)t(iter(replace_termsigmaevar(mkVarid)c))inletc'=itercinletenv=Global.env()inifcheckthenPretyping.check_evarsenvsigmac';letdiff=Univ.ContextSet.diff(Evd.universe_context_setsigma)(Evd.universe_context_setinit)in(c',diff)letadd_hints~localitydbnamesh=letlocal=interp_localitylocalityinifString.List.mem"nocore"dbnamesthenuser_errPp.(str"The hint database \"nocore\" is meant to stay empty.");assert(not(List.is_emptydbnames));letenv=Global.env()inletsigma=Evd.from_envenvinmatchhwith|HintsResolveEntrylhints->add_resolvesenvsigmalhints~localdbnames|HintsImmediateEntrylhints->add_trivialsenvsigmalhints~localdbnames|HintsCutEntrylhints->add_cutslhints~localdbnames|HintsModeEntry(l,m)->add_modelm~localdbnames|HintsUnfoldEntrylhints->add_unfoldslhints~localdbnames|HintsTransparencyEntry(lhints,b)->add_transparencylhintsb~localdbnames|HintsExternEntry(info,tacexp)->add_externsinfotacexp~localdbnameslethint_globrefgr=IsGlobRefgrlethint_constr(c,diff)=IsConstr(c,diff)letexpand_constructor_hintsenvsigmalems=List.map_append(fun(evd,lem)->matchEConstr.kindsigmalemwith|Ind(ind,u)->List.init(nconstructorsenvind)(funi->IsGlobRef(GlobRef.ConstructRef((ind,i+1))))|_->let(c,ctx)=prepare_hintfalseenvsigma(evd,lem)inletctx=ifUniv.ContextSet.is_emptyctxthenNoneelseSomectxin[IsConstr(c,ctx)])lems(* builds a hint database from a constr signature *)(* typically used with (lid, ltyp) = pf_hyps_types <some goal> *)letconstructor_hintsenvsigmaeapplylems=letlems=expand_constructor_hintsenvsigmalemsinList.map_append(funlem->make_resolvesenvsigma(eapply,true)empty_hint_info~check:truelem)lemsletmake_local_hint_dbenvsigmatseapplylems=letmapc=cenvsigmainletlems=List.mapmaplemsinletsign=EConstr.named_contextenvinletts=matchtswith|None->Hint_db.transparent_state(searchtable_map"core")|Somets->tsinlethintlist=List.map_append(fundecl->make_resolve_hypenvsigma(Named.Declaration.get_iddecl))signinHint_db.emptytsfalse|>Hint_db.add_listenvsigmahintlist|>Hint_db.add_listenvsigma(constructor_hintsenvsigmaeapplylems)letmake_local_hint_dbenvsigma?tseapplylems=make_local_hint_dbenvsigmatseapplylemsletmake_db_listdbnames=letuse_core=not(List.mem"nocore"dbnames)inletdbnames=List.removeString.equal"nocore"dbnamesinletdbnames=ifuse_corethen"core"::dbnameselsedbnamesinletlookupdb=trysearchtable_mapdbwithNot_found->error_no_such_hint_databasedbinList.maplookupdbnamesletpush_resolvesenvsigmahintdb=letname=PathHints[hint]inletentries=make_resolvesenvsigma(true,false)empty_hint_info~check:false~name(IsGlobRefhint)inHint_db.add_listenvsigmaentriesdbletpush_resolve_hypenvsigmadecldb=letentries=make_resolve_hypenvsigmadeclinHint_db.add_listenvsigmaentriesdb(**************************************************************************)(* Functions for printing the hints *)(**************************************************************************)letpr_hint_eltenvsigmah=pr_econstr_envenvsigmah.hint_termletpr_hintenvsigmah=matchh.objwith|Res_pfc->(str"simple apply "++pr_hint_eltenvsigmac)|ERes_pfc->(str"simple eapply "++pr_hint_eltenvsigmac)|Give_exactc->(str"exact "++pr_hint_eltenvsigmac)|Res_pf_THEN_trivial_failc->(str"simple apply "++pr_hint_eltenvsigmac++str" ; trivial")|Unfold_nthc->str"unfold "++pr_evaluable_referencec|Extern(_,tac)->str"(*external*) "++Pputils.pr_glb_genericenvsigmatacletpr_id_hintenvsigma(id,v)=letpr_patp=str", pattern "++pr_lconstr_pattern_envenvsigmapin(pr_hintenvsigmav.code++str"(level "++intv.pri++pr_opt_no_spcpr_patv.pat++str", id "++intid++str")"++spc())letpr_hint_listenvsigmahintlist=(str" "++hov0(prlist(pr_id_hintenvsigma)hintlist)++fnl())letpr_hints_dbenvsigma(name,db,hintlist)=(str"In the database "++strname++str":"++ifList.is_emptyhintlistthen(str" nothing"++fnl())else(fnl()++pr_hint_listenvsigmahintlist))(* Print all hints associated to head c in any database *)letpr_hint_list_for_headenvsigmac=letdbs=current_db()inletvalidate(name,db)=lethints=List.map(funv->0,v)(Hint_db.map_all~secvars:Id.Pred.fullcdb)in(name,db,hints)inletvalid_dbs=List.mapvalidatedbsinifList.is_emptyvalid_dbsthen(str"No hint declared for :"++pr_globalc)elsehov0(str"For "++pr_globalc++str" -> "++fnl()++hov0(prlist(pr_hints_dbenvsigma)valid_dbs))letpr_hint_refref=pr_hint_list_for_headref(* Print all hints associated to head id in any database *)letpr_hint_termenvsigmacl=tryletdbs=current_db()inletvalid_dbs=letfn=trylethdc=decompose_app_boundsigmaclinifoccur_existentialsigmaclthen(fundb->matchHint_db.map_existentialsigma~secvars:Id.Pred.fullhdccldbwith|ModeMatchl->l|ModeMismatch->[])elseHint_db.map_autoenvsigma~secvars:Id.Pred.fullhdcclwithBound->Hint_db.map_none~secvars:Id.Pred.fullinletfndb=List.map(funx->0,x)(fndb)inList.map(fun(name,db)->(name,db,fndb))dbsinifList.is_emptyvalid_dbsthen(str"No hint applicable for current goal")else(str"Applicable Hints :"++fnl()++hov0(prlist(pr_hints_dbenvsigma)valid_dbs))withMatch_failure_|Failure_->(str"No hint applicable for current goal")(* print all hints that apply to the concl of the current goal *)letpr_applicable_hintpf=letenv=Global.env()inletProof.{goals;sigma}=Proof.datapfinmatchgoalswith|[]->CErrors.user_errPp.(str"No focused goal.")|g::_->pr_hint_termenvsigma(Goal.V82.conclsigmag)letpp_hint_mode=function|ModeInput->str"+"|ModeNoHeadEvar->str"!"|ModeOutput->str"-"(* displays the whole hint database db *)letpr_hint_db_envenvsigmadb=letpr_mode=prvect_with_sepspcpp_hint_modeinletpr_modesl=ifList.is_emptylthenmt()elsestr" (modes "++prlist_with_seppr_commapr_model++str")"inletcontent=letfoldheadmodeshintlistaccu=letgoal_descr=matchheadwith|None->str"For any goal"|Somehead->str"For "++pr_globalhead++pr_modesmodesinlethints=pr_hint_listenvsigma(List.map(funx->(0,x))hintlist)inlethint_descr=hov0(goal_descr++str" -> "++hints)inaccu++hint_descrinHint_db.foldfolddb(mt())inlet{TransparentState.tr_var=ids;tr_cst=csts}=Hint_db.transparent_statedbinhov0((ifHint_db.use_dndbthenstr"Discriminated database"elsestr"Non-discriminated database"))++fnl()++hov2(str"Unfoldable variable definitions: "++pr_idpredids)++fnl()++hov2(str"Unfoldable constant definitions: "++pr_cpredcsts)++fnl()++hov2(str"Cut: "++pp_hints_path(Hint_db.cutdb))++fnl()++contentletpr_hint_db_by_nameenvsigmadbname=tryletdb=searchtable_mapdbnameinpr_hint_db_envenvsigmadbwithNot_found->error_no_such_hint_databasedbname(* displays all the hints of all databases *)letpr_searchtableenvsigma=letfoldnamedbaccu=accu++str"In the database "++strname++str":"++fnl()++pr_hint_db_envenvsigmadb++fnl()inHintdbmap.foldfold!searchtable(mt())letprint_mpmp=tryletqid=Nametab.shortest_qualid_of_modulempinstr" from "++pr_qualidqidwithNot_found->mt()letis_importedh=tryKNmap.findh.uid!statustablewithNot_found->truelethint_trace=Evd.Store.field()letlog_hinth=letopenProofview.NotationsinProofview.tclEVARMAP>>=funsigma->letstore=get_extra_datasigmainmatchStore.getstorehint_tracewith|None->(* All calls to hint logging should be well-scoped *)assertfalse|Sometrace->lettrace=KNmap.addh.uidhtraceinletstore=Store.setstorehint_tracetraceinProofview.Unsafe.tclEVARS(set_extra_datastoresigma)letwarn_non_imported_hint=CWarnings.create~name:"non-imported-hint"~category:"automation"(fun(hint,mp)->strbrk"Hint used but not imported: "++hint++print_mpmp)letwarnenvsigmah=lethint=pr_hintenvsigmahinletmp=KerName.modpathh.uidinwarn_non_imported_hint(hint,mp)letwrap_hint_warningt=letopenProofview.NotationsinProofview.tclEVARMAP>>=funsigma->letstore=get_extra_datasigmainletold=Store.getstorehint_traceinletstore=Store.setstorehint_traceKNmap.emptyinProofview.Unsafe.tclEVARS(set_extra_datastoresigma)>>=fun()->t>>=funans->Proofview.tclENV>>=funenv->Proofview.tclEVARMAP>>=funsigma->letstore=get_extra_datasigmainlethints=matchStore.getstorehint_tracewith|None->assertfalse|Somehints->hintsinlet()=KNmap.iter(fun_h->warnenvsigmah)hintsinletstore=matcholdwith|None->Store.removestorehint_trace|Somev->Store.setstorehint_tracevinProofview.Unsafe.tclEVARS(set_extra_datastoresigma)>>=fun()->Proofview.tclUNITansletwrap_hint_warning_funenvsigmat=letstore=get_extra_datasigmainletold=Store.getstorehint_traceinletstore=Store.setstorehint_traceKNmap.emptyinlet(ans,sigma)=t(set_extra_datastoresigma)inletstore=get_extra_datasigmainlethints=matchStore.getstorehint_tracewith|None->assertfalse|Somehints->hintsinlet()=KNmap.iter(fun_h->warnenvsigmah)hintsinletstore=matcholdwith|None->Store.removestorehint_trace|Somev->Store.setstorehint_tracevin(ans,set_extra_datastoresigma)letrun_hinttack=matchwarn_hint()with|HintLax->ktac.obj|HintWarn->ifis_importedtacthenktac.objelseProofview.tclTHEN(log_hinttac)(ktac.obj)|HintStrict->ifis_importedtacthenktac.objelseletinfo=Exninfo.reify()inProofview.tclZERO~info(UserError(None,(str"Tactic failure.")))moduleFullHint=structtypet=full_hintletpriority(h:t)=h.priletpattern(h:t)=h.patletdatabase(h:t)=h.dbletrun(h:t)k=run_hinth.codekletprintenvsigma(h:t)=pr_hintenvsigmah.codeletname(h:t)=h.nameletrepr(h:t)=h.code.objendletconnect_hint_clenvhgl=let{hint_uctx=ctx;hint_clnv=clenv}=hin(* [clenv] has been generated by a hint-making function, so the only relevant
data in its evarmap is the set of metas. The [evar_reset_evd] function
below just replaces the metas of sigma by those coming from the clenv. *)letsigma=Tacmach.New.projectglinletevd=Evd.evars_reset_evd~with_conv_pbs:true~with_univs:falsesigmaclenv.evdin(* Still, we need to update the universes *)matchh.hint_uctxwith|Somectx->(* Refresh the instance of the hint *)let(subst,ctx)=UnivGen.fresh_universe_context_set_instancectxinletemapc=Vars.subst_univs_level_constrsubstcinletevd=Evd.merge_context_setEvd.univ_flexibleevdctxin(* Only metas are mentioning the old universes. *)Clenv.mk_clausenv(Proofview.Goal.envgl)(Evd.map_metasemapevd)(Evd.map_flemapclenv.templval)(Evd.map_flemapclenv.templtyp)|None->Clenv.mk_clausenv(Proofview.Goal.envgl)evdclenv.templvalclenv.templtypletfresh_hintenvsigmah=let{hint_term=c;hint_uctx=ctx}=hinmatchh.hint_uctxwith|None->sigma,c|Somectx->(* Refresh the instance of the hint *)let(subst,ctx)=UnivGen.fresh_universe_context_set_instancectxinletc=Vars.subst_univs_level_constrsubstcinletsigma=Evd.merge_context_setEvd.univ_flexiblesigmactxinsigma,clethint_res_pf?with_evars?with_classes?flagsh=Proofview.Goal.enterbeginfungl->letclenv=connect_hint_clenvhglinClenv.res_pf?with_evars?with_classes?flagsclenvend